diff --git a/thys/Algebraic_Numbers/Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Algebraic_Numbers.thy @@ -1,1338 +1,1349 @@ (* Author: RenĂ© Thiemann Akihisa Yamada Contributors: Manuel Eberl (algebraic integers) License: BSD *) section \Algebraic Numbers: Addition and Multiplication\ text \This theory contains the remaining field operations for algebraic numbers, namely addition and multiplication.\ theory Algebraic_Numbers imports Algebraic_Numbers_Prelim Resultant Polynomial_Factorization.Polynomial_Divisibility begin interpretation coeff_hom: monoid_add_hom "\p. coeff p i" by (unfold_locales, auto) interpretation coeff_hom: comm_monoid_add_hom "\p. coeff p i".. interpretation coeff_hom: group_add_hom "\p. coeff p i".. interpretation coeff_hom: ab_group_add_hom "\p. coeff p i".. interpretation coeff_0_hom: monoid_mult_hom "\p. coeff p 0" by (unfold_locales, auto simp: coeff_mult) interpretation coeff_0_hom: semiring_hom "\p. coeff p 0".. interpretation coeff_0_hom: comm_monoid_mult_hom "\p. coeff p 0".. interpretation coeff_0_hom: comm_semiring_hom "\p. coeff p 0".. subsection \Addition of Algebraic Numbers\ definition "x_y \ [: [: 0, 1 :], -1 :]" definition "poly_x_minus_y p = poly_lift p \\<^sub>p x_y" lemma coeff_xy_power: assumes "k \ n" shows "coeff (x_y ^ n :: 'a :: comm_ring_1 poly poly) k = monom (of_nat (n choose (n - k)) * (- 1) ^ k) (n - k)" proof - define X :: "'a poly poly" where "X = monom (monom 1 1) 0" define Y :: "'a poly poly" where "Y = monom (-1) 1" have [simp]: "monom 1 b * (-1) ^ k = monom ((-1)^k :: 'a) b" for b k by (auto simp: monom_altdef minus_one_power_iff) have "(X + Y) ^ n = (\i\n. of_nat (n choose i) * X ^ i * Y ^ (n - i))" by (subst binomial_ring) auto also have "\ = (\i\n. of_nat (n choose i) * monom (monom ((-1) ^ (n - i)) i) (n - i))" by (simp add: X_def Y_def monom_power mult_monom mult.assoc) also have "\ = (\i\n. monom (monom (of_nat (n choose i) * (-1) ^ (n - i)) i) (n - i))" by (simp add: of_nat_poly smult_monom) also have "coeff \ k = (\i\n. if n - i = k then monom (of_nat (n choose i) * (- 1) ^ (n - i)) i else 0)" by (simp add: of_nat_poly coeff_sum) also have "\ = (\i\{n-k}. monom (of_nat (n choose i) * (- 1) ^ (n - i)) i)" using \k \ n\ by (intro sum.mono_neutral_cong_right) auto also have "X + Y = x_y" by (simp add: X_def Y_def x_y_def monom_altdef) finally show ?thesis using \k \ n\ by simp qed text \The following polynomial represents the sum of two algebraic numbers.\ definition poly_add :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_add p q = resultant (poly_x_minus_y p) (poly_lift q)" subsubsection \@{term poly_add} has desired root\ interpretation poly_x_minus_y_hom: comm_ring_hom poly_x_minus_y by (unfold_locales; simp add: poly_x_minus_y_def hom_distribs) lemma poly2_x_y[simp]: fixes x :: "'a :: comm_ring_1" shows "poly2 x_y x y = x - y" unfolding poly2_def by (simp add: x_y_def) lemma degree_poly_x_minus_y[simp]: fixes p :: "'a::idom poly" shows "degree (poly_x_minus_y p) = degree p" unfolding poly_x_minus_y_def x_y_def by auto lemma poly_x_minus_y_pCons[simp]: "poly_x_minus_y (pCons a p) = [:[: a :]:] + poly_x_minus_y p * x_y" unfolding poly_x_minus_y_def x_y_def by simp lemma poly_poly_poly_x_minus_y[simp]: fixes p :: "'a :: comm_ring_1 poly" shows "poly (poly (poly_x_minus_y p) q) x = poly p (x - poly q x)" by (induct p; simp add: ring_distribs x_y_def) lemma poly2_poly_x_minus_y[simp]: fixes p :: "'a :: comm_ring_1 poly" shows "poly2 (poly_x_minus_y p) x y = poly p (x-y)" unfolding poly2_def by simp interpretation x_y_mult_hom: zero_hom_0 "\p :: 'a :: comm_ring_1 poly poly. x_y * p" proof (unfold_locales) fix p :: "'a poly poly" assume "x_y * p = 0" then show "p = 0" apply (simp add: x_y_def) by (metis eq_neg_iff_add_eq_0 minus_equation_iff minus_pCons synthetic_div_unique_lemma) qed lemma x_y_nonzero[simp]: "x_y \ 0" by (simp add: x_y_def) lemma degree_x_y[simp]: "degree x_y = 1" by (simp add: x_y_def) interpretation x_y_mult_hom: inj_comm_monoid_add_hom "\p :: 'a :: idom poly poly. x_y * p" proof (unfold_locales) show "x_y * p = x_y * q \ p = q" for p q :: "'a poly poly" proof (induct p arbitrary:q) case 0 then show ?case by simp next case p: (pCons a p) from p(3)[unfolded mult_pCons_right] have "x_y * (monom a 0 + pCons 0 1 * p) = x_y * q" apply (subst(asm) pCons_0_as_mult) apply (subst(asm) smult_prod) by (simp only: field_simps distrib_left) then have "monom a 0 + pCons 0 1 * p = q" by simp then show "pCons a p = q" using pCons_as_add by (simp add: monom_0 monom_Suc) qed qed interpretation poly_x_minus_y_hom: inj_idom_hom poly_x_minus_y proof fix p :: "'a poly" assume 0: "poly_x_minus_y p = 0" then have "poly_lift p \\<^sub>p x_y = 0" by (simp add: poly_x_minus_y_def) then show "p = 0" proof (induct p) case 0 then show ?case by simp next case (pCons a p) note p = this[unfolded poly_lift_pCons pcompose_pCons] show ?case proof (cases "a=0") case a0: True with p have "x_y * poly_lift p \\<^sub>p x_y = 0" by simp then have "poly_lift p \\<^sub>p x_y = 0" by simp then show ?thesis using p by simp next case a0: False with p have p0: "p \ 0" by auto from p have "[:[:a:]:] = - x_y * poly_lift p \\<^sub>p x_y" by (simp add: eq_neg_iff_add_eq_0) then have "degree [:[:a:]:] = degree (x_y * poly_lift p \\<^sub>p x_y)" by simp also have "... = degree (x_y::'a poly poly) + degree (poly_lift p \\<^sub>p x_y)" apply (subst degree_mult_eq) apply simp apply (subst pcompose_eq_0) apply (simp add: x_y_def) apply (simp add: p0) apply simp done finally have False by simp then show ?thesis.. qed qed qed lemma poly_add: fixes p q :: "'a ::comm_ring_1 poly" assumes q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" shows "poly (poly_add p q) (x+y) = 0" proof (unfold poly_add_def, rule poly_resultant_zero[OF disjI2]) have "degree q > 0" using poly_zero q0 y by auto thus degq: "degree (poly_lift q) > 0" by auto qed (insert x y, simp_all) subsubsection \@{const poly_add} is nonzero\ text \ We first prove that @{const poly_lift} preserves factorization. The result will be essential also in the next section for division of algebraic numbers. \ interpretation poly_lift_hom: unit_preserving_hom "poly_lift :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly \ _" proof fix x :: "'a poly" assume "poly_lift x dvd 1" then have "poly_y_x (poly_lift x) dvd poly_y_x 1" by simp then show "x dvd 1" by (auto simp add: poly_y_x_poly_lift) qed interpretation poly_lift_hom: factor_preserving_hom "poly_lift::'a::idom poly \ 'a poly poly" proof unfold_locales fix p :: "'a poly" assume p: "irreducible p" show "irreducible (poly_lift p)" proof(rule ccontr) from p have p0: "p \ 0" and "\ p dvd 1" by (auto dest: irreducible_not_unit) with poly_lift_hom.hom_dvd[of p 1] have p1: "\ poly_lift p dvd 1" by auto assume "\ irreducible (poly_lift p)" from this[unfolded irreducible_altdef,simplified] p0 p1 obtain q where "q dvd poly_lift p" and pq: "\ poly_lift p dvd q" and q: "\ q dvd 1" by auto then obtain r where "q * r = poly_lift p" by (elim dvdE, auto) then have "poly_y_x (q * r) = poly_y_x (poly_lift p)" by auto also have "... = [:p:]" by (auto simp: poly_y_x_poly_lift monom_0) also have "poly_y_x (q * r) = poly_y_x q * poly_y_x r" by (auto simp: hom_distribs) finally have "... = [:p:]" by auto then have qp: "poly_y_x q dvd [:p:]" by (metis dvdI) from dvd_const[OF this] p0 have "degree (poly_y_x q) = 0" by auto from degree_0_id[OF this,symmetric] obtain s where qs: "poly_y_x q = [:s:]" by auto have "poly_lift s = poly_y_x (poly_y_x (poly_lift s))" by auto also have "... = poly_y_x [:s:]" by (auto simp: poly_y_x_poly_lift monom_0) also have "... = q" by (auto simp: qs[symmetric]) finally have sq: "poly_lift s = q" by auto from qp[unfolded qs] have sp: "s dvd p" by (auto simp: const_poly_dvd) from irreducibleD'[OF p this] sq q pq show False by auto qed qed text \ We now show that @{const poly_x_minus_y} is a factor-preserving homomorphism. This is essential for this section. This is easy since @{const poly_x_minus_y} can be represented as the composition of two factor-preserving homomorphisms. \ lemma poly_x_minus_y_as_comp: "poly_x_minus_y = (\p. p \\<^sub>p x_y) \ poly_lift" by (intro ext, unfold poly_x_minus_y_def, auto) context idom_isom begin sublocale comm_semiring_isom.. end + interpretation poly_x_minus_y_hom: factor_preserving_hom "poly_x_minus_y :: 'a :: idom poly \ 'a poly poly" -proof- - interpret x_y_hom: bijective "\p :: 'a poly poly. p \\<^sub>p x_y" - proof (unfold bijective_eq_bij, rule id_imp_bij) - fix p :: "'a poly poly" show "p \\<^sub>p x_y \\<^sub>p x_y = p" - apply (induct p,simp) - apply(unfold x_y_def hom_distribs pcompose_pCons) by (simp) +proof - + have \p \\<^sub>p x_y \\<^sub>p x_y = p\ for p :: \'a poly poly\ + proof (induction p) + case 0 + show ?case + by simp + next + case (pCons a p) + then show ?case + by (unfold x_y_def hom_distribs pcompose_pCons) simp qed - interpret x_y_hom: idom_isom "\p :: 'a poly poly. p \\<^sub>p x_y" by (unfold_locales, auto) - show "factor_preserving_hom (poly_x_minus_y :: 'a poly \ _)" - by (unfold poly_x_minus_y_as_comp, rule factor_preserving_hom_comp, unfold_locales) + then interpret x_y_hom: bijective "\p :: 'a poly poly. p \\<^sub>p x_y" + by (unfold bijective_eq_bij) (rule involuntory_imp_bij) + interpret x_y_hom: idom_isom "\p :: 'a poly poly. p \\<^sub>p x_y" + by standard simp_all + have \factor_preserving_hom (\p :: 'a poly poly. p \\<^sub>p x_y)\ + and \factor_preserving_hom (poly_lift :: 'a poly \ 'a poly poly)\ + .. + then show "factor_preserving_hom (poly_x_minus_y :: 'a poly \ _)" + by (unfold poly_x_minus_y_as_comp) (rule factor_preserving_hom_comp) qed text \ Now we show that results of @{const poly_x_minus_y} and @{const poly_lift} are coprime. \ lemma poly_y_x_const[simp]: "poly_y_x [:[:a:]:] = [:[:a:]:]" by (simp add: poly_y_x_def monom_0) context begin private abbreviation "y_x == [: [: 0, -1 :], 1 :]" lemma poly_y_x_x_y[simp]: "poly_y_x x_y = y_x" by (simp add: x_y_def poly_y_x_def monom_Suc monom_0) private lemma y_x[simp]: fixes x :: "'a :: comm_ring_1" shows "poly2 y_x x y = y - x" unfolding poly2_def by simp private definition "poly_y_minus_x p \ poly_lift p \\<^sub>p y_x" private lemma poly_y_minus_x_0[simp]: "poly_y_minus_x 0 = 0" by (simp add: poly_y_minus_x_def) private lemma poly_y_minus_x_pCons[simp]: "poly_y_minus_x (pCons a p) = [:[: a :]:] + poly_y_minus_x p * y_x" by (simp add: poly_y_minus_x_def) private lemma poly_y_x_poly_x_minus_y: fixes p :: "'a :: idom poly" shows "poly_y_x (poly_x_minus_y p) = poly_y_minus_x p" apply (induct p, simp) apply (unfold poly_x_minus_y_pCons hom_distribs) by simp lemma degree_poly_y_minus_x[simp]: fixes p :: "'a :: idom poly" shows "degree (poly_y_x (poly_x_minus_y p)) = degree p" by (simp add: poly_y_minus_x_def poly_y_x_poly_x_minus_y) end lemma dvd_all_coeffs_iff: fixes x :: "'a :: comm_semiring_1" (* No addition needed! *) shows "(\pi \ set (coeffs p). x dvd pi) \ (\i. x dvd coeff p i)" (is "?l = ?r") proof- have "?r = (\i\{..degree p} \ {Suc (degree p)..}. x dvd coeff p i)" by auto also have "... = (\i\degree p. x dvd coeff p i)" by (auto simp add: ball_Un coeff_eq_0) also have "... = ?l" by (auto simp: coeffs_def) finally show ?thesis.. qed lemma primitive_imp_no_constant_factor: fixes p :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} poly" assumes pr: "primitive p" and F: "mset_factors F p" and fF: "f \# F" shows "degree f \ 0" proof from F fF have irr: "irreducible f" and fp: "f dvd p" by (auto dest: mset_factors_imp_dvd) assume deg: "degree f = 0" then obtain f0 where f0: "f = [:f0:]" by (auto dest: degree0_coeffs) with fp have "[:f0:] dvd p" by simp then have "f0 dvd coeff p i" for i by (simp add: const_poly_dvd_iff) with primitiveD[OF pr] dvd_all_coeffs_iff have "f0 dvd 1" by (auto simp: coeffs_def) with f0 irr show False by auto qed lemma coprime_poly_x_minus_y_poly_lift: fixes p q :: "'a :: ufd poly" assumes degp: "degree p > 0" and degq: "degree q > 0" and pr: "primitive p" shows "coprime (poly_x_minus_y p) (poly_lift q)" proof(rule ccontr) from degp have p: "\ p dvd 1" by (auto simp: dvd_const) from degp have p0: "p \ 0" by auto from mset_factors_exist[of p, OF p0 p] obtain F where F: "mset_factors F p" by auto with poly_x_minus_y_hom.hom_mset_factors have pF: "mset_factors (image_mset poly_x_minus_y F) (poly_x_minus_y p)" by auto from degq have q: "\ q dvd 1" by (auto simp: dvd_const) from degq have q0: "q \ 0" by auto from mset_factors_exist[OF q0 q] obtain G where G: "mset_factors G q" by auto with poly_lift_hom.hom_mset_factors have pG: "mset_factors (image_mset poly_lift G) (poly_lift q)" by auto assume "\ coprime (poly_x_minus_y p) (poly_lift q)" from this[unfolded not_coprime_iff_common_factor] obtain r where rp: "r dvd (poly_x_minus_y p)" and rq: "r dvd (poly_lift q)" and rU: "\ r dvd 1" by auto note poly_lift_hom.hom_dvd from rp p0 have r0: "r \ 0" by auto from mset_factors_exist[OF r0 rU] obtain H where H: "mset_factors H r" by auto then have "H \ {#}" by auto then obtain h where hH: "h \# H" by fastforce with H mset_factors_imp_dvd have hr: "h dvd r" and h: "irreducible h" by auto from irreducible_not_unit[OF h] have hU: "\ h dvd 1" by auto from hr rp have "h dvd (poly_x_minus_y p)" by (rule dvd_trans) from irreducible_dvd_imp_factor[OF this h pF] p0 obtain f where f: "f \# F" and fh: "poly_x_minus_y f ddvd h" by auto from hr rq have "h dvd (poly_lift q)" by (rule dvd_trans) from irreducible_dvd_imp_factor[OF this h pG] q0 obtain g where g: "g \# G" and gh: "poly_lift g ddvd h" by auto from fh gh have "poly_x_minus_y f ddvd poly_lift g" using ddvd_trans by auto then have "poly_y_x (poly_x_minus_y f) ddvd poly_y_x (poly_lift g)" by simp also have "poly_y_x (poly_lift g) = [:g:]" unfolding poly_y_x_poly_lift monom_0 by auto finally have ddvd: "poly_y_x (poly_x_minus_y f) ddvd [:g:]" by auto then have "degree (poly_y_x (poly_x_minus_y f)) = 0" by (metis degree_pCons_0 dvd_0_left_iff dvd_const) then have "degree f = 0" by simp with primitive_imp_no_constant_factor[OF pr F f] show False by auto qed lemma poly_add_nonzero: fixes p q :: "'a :: ufd poly" assumes p0: "p \ 0" and q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and pr: "primitive p" shows "poly_add p q \ 0" proof have degp: "degree p > 0" using le_0_eq order_degree order_root p0 x by (metis gr0I) have degq: "degree q > 0" using le_0_eq order_degree order_root q0 y by (metis gr0I) assume 0: "poly_add p q = 0" from resultant_zero_imp_common_factor[OF _ this[unfolded poly_add_def]] degp and coprime_poly_x_minus_y_poly_lift[OF degp degq pr] show False by auto qed subsubsection \Summary for addition\ text \Now we lift the results to one that uses @{const ipoly}, by showing some homomorphism lemmas.\ lemma (in comm_ring_hom) map_poly_x_minus_y: "map_poly (map_poly hom) (poly_x_minus_y p) = poly_x_minus_y (map_poly hom p)" proof- interpret mp: map_poly_comm_ring_hom hom.. interpret mmp: map_poly_comm_ring_hom "map_poly hom".. show ?thesis apply (induct p, simp) apply(unfold x_y_def hom_distribs poly_x_minus_y_pCons, simp) done qed lemma (in comm_ring_hom) hom_poly_lift[simp]: "map_poly (map_poly hom) (poly_lift q) = poly_lift (map_poly hom q)" proof - show ?thesis unfolding poly_lift_def unfolding map_poly_map_poly[of coeff_lift,OF coeff_lift_hom.hom_zero] unfolding map_poly_coeff_lift_hom by simp qed lemma lead_coeff_poly_x_minus_y: fixes p :: "'a::idom poly" shows "lead_coeff (poly_x_minus_y p) = [:lead_coeff p * ((- 1) ^ degree p):]" (is "?l = ?r") proof- have "?l = Polynomial.smult (lead_coeff p) ((- 1) ^ degree p)" by (unfold poly_x_minus_y_def, subst lead_coeff_comp; simp add: x_y_def) also have "... = ?r" by (unfold hom_distribs, simp add: smult_as_map_poly[symmetric]) finally show ?thesis. qed lemma degree_coeff_poly_x_minus_y: fixes p q :: "'a :: {idom, semiring_char_0} poly" shows "degree (coeff (poly_x_minus_y p) i) = degree p - i" proof - consider "i = degree p" | "i > degree p" | "i < degree p" by force thus ?thesis proof cases assume "i > degree p" thus ?thesis by (subst coeff_eq_0) auto next assume "i = degree p" thus ?thesis using lead_coeff_poly_x_minus_y[of p] by (simp add: lead_coeff_poly_x_minus_y) next assume "i < degree p" define n where "n = degree p" have "degree (coeff (poly_x_minus_y p) i) = degree (\j\n. [:coeff p j:] * coeff (x_y ^ j) i)" (is "_ = degree (sum ?f _)") by (simp add: poly_x_minus_y_def pcompose_conv_poly poly_altdef coeff_sum n_def) also have "{..n} = insert n {.. = ?f n + sum ?f {.. = n - i" proof - have "degree (?f n) = n - i" using \i < degree p\ by (simp add: n_def coeff_xy_power degree_monom_eq) moreover have "degree (sum ?f {.. {.. j - i" proof (cases "i \ j") case True thus ?thesis by (auto simp: n_def coeff_xy_power degree_monom_eq) next case False hence "coeff (x_y ^ j :: 'a poly poly) i = 0" by (subst coeff_eq_0) (auto simp: degree_power_eq) thus ?thesis by simp qed also have "\ < n - i" using \j \ {.. \i < degree p\ by (auto simp: n_def) finally show "degree ([:coeff p j:] * coeff (x_y ^ j) i) < n - i" . qed (use \i < degree p\ in \auto simp: n_def\) ultimately show ?thesis by (subst degree_add_eq_left) auto qed finally show ?thesis by (simp add: n_def) qed qed lemma coeff_0_poly_x_minus_y [simp]: "coeff (poly_x_minus_y p) 0 = p" by (induction p) (auto simp: poly_x_minus_y_def x_y_def) lemma (in idom_hom) poly_add_hom: assumes p0: "hom (lead_coeff p) \ 0" and q0: "hom (lead_coeff q) \ 0" shows "map_poly hom (poly_add p q) = poly_add (map_poly hom p) (map_poly hom q)" proof - interpret mh: map_poly_idom_hom.. show ?thesis unfolding poly_add_def apply (subst mh.resultant_map_poly(1)[symmetric]) apply (subst degree_map_poly_2) apply (unfold lead_coeff_poly_x_minus_y, unfold hom_distribs, simp add: p0) apply simp apply (subst degree_map_poly_2) apply (simp_all add: q0 map_poly_x_minus_y) done qed lemma(in zero_hom) hom_lead_coeff_nonzero_imp_map_poly_hom: assumes "hom (lead_coeff p) \ 0" shows "map_poly hom p \ 0" proof assume "map_poly hom p = 0" then have "coeff (map_poly hom p) (degree p) = 0" by simp with assms show False by simp qed lemma ipoly_poly_add: fixes x y :: "'a :: idom" assumes p0: "(of_int (lead_coeff p) :: 'a) \ 0" and q0: "(of_int (lead_coeff q) :: 'a) \ 0" and x: "ipoly p x = 0" and y: "ipoly q y = 0" shows "ipoly (poly_add p q) (x+y) = 0" using assms of_int_hom.hom_lead_coeff_nonzero_imp_map_poly_hom[OF q0] by (auto intro: poly_add simp: of_int_hom.poly_add_hom[OF p0 q0]) lemma (in comm_monoid_gcd) gcd_list_eq_0_iff[simp]: "listgcd xs = 0 \ (\x \ set xs. x = 0)" by (induct xs, auto) lemma primitive_field_poly[simp]: "primitive (p :: 'a :: field poly) \ p \ 0" by (unfold primitive_iff_some_content_dvd_1,auto simp: dvd_field_iff coeffs_def) lemma ipoly_poly_add_nonzero: fixes x y :: "'a :: field" assumes "p \ 0" and "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "(of_int (lead_coeff p) :: 'a) \ 0" and "(of_int (lead_coeff q) :: 'a) \ 0" shows "poly_add p q \ 0" proof- from assms have "(of_int_poly (poly_add p q) :: 'a poly) \ 0" apply (subst of_int_hom.poly_add_hom,simp,simp) by (rule poly_add_nonzero, auto dest:of_int_hom.hom_lead_coeff_nonzero_imp_map_poly_hom) then show ?thesis by auto qed lemma represents_add: assumes x: "p represents x" and y: "q represents y" shows "(poly_add p q) represents (x + y)" using assms by (intro representsI ipoly_poly_add ipoly_poly_add_nonzero, auto) subsection \Division of Algebraic Numbers\ definition poly_x_mult_y where [code del]: "poly_x_mult_y p \ (\ i \ degree p. monom (monom (coeff p i) i) i)" lemma coeff_poly_x_mult_y: shows "coeff (poly_x_mult_y p) i = monom (coeff p i) i" (is "?l = ?r") proof(cases "degree p < i") case i: False have "?l = sum (\j. if j = i then (monom (coeff p j) j) else 0) {..degree p}" (is "_ = sum ?f ?A") by (simp add: poly_x_mult_y_def coeff_sum) also have "... = sum ?f {i}" using i by (intro sum.mono_neutral_right, auto) also have "... = ?f i" by simp also have "... = ?r" by auto finally show ?thesis. next case True then show ?thesis by (auto simp: poly_x_mult_y_def coeff_eq_0 coeff_sum) qed lemma poly_x_mult_y_code[code]: "poly_x_mult_y p = (let cs = coeffs p in poly_of_list (map (\ (i, ai). monom ai i) (zip [0 ..< length cs] cs)))" unfolding Let_def poly_of_list_def proof (rule poly_eqI, unfold coeff_poly_x_mult_y) fix n let ?xs = "zip [0.. degree p \ p = 0" unfolding degree_eq_length_coeffs by (cases n, auto) hence "monom (coeff p n) n = 0" using coeff_eq_0[of p n] by auto thus ?thesis unfolding id by simp qed qed definition poly_div :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_div p q = resultant (poly_x_mult_y p) (poly_lift q)" text \@{const poly_div} has desired roots.\ lemma poly2_poly_x_mult_y: fixes p :: "'a :: comm_ring_1 poly" shows "poly2 (poly_x_mult_y p) x y = poly p (x * y)" apply (subst(3) poly_as_sum_of_monoms[symmetric]) apply (unfold poly_x_mult_y_def hom_distribs) by (auto simp: poly2_monom poly_monom power_mult_distrib ac_simps) lemma poly_div: fixes p q :: "'a ::field poly" assumes q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and y0: "y \ 0" shows "poly (poly_div p q) (x/y) = 0" proof (unfold poly_div_def, rule poly_resultant_zero[OF disjI2]) have "degree q > 0" using poly_zero q0 y by auto thus degq: "degree (poly_lift q) > 0" by auto qed (insert x y y0, simp_all add: poly2_poly_x_mult_y) text \@{const poly_div} is nonzero.\ interpretation poly_x_mult_y_hom: ring_hom "poly_x_mult_y :: 'a :: {idom,ring_char_0} poly \ _" by (unfold_locales, auto intro: poly2_ext simp: poly2_poly_x_mult_y hom_distribs) interpretation poly_x_mult_y_hom: inj_ring_hom "poly_x_mult_y :: 'a :: {idom,ring_char_0} poly \ _" proof let ?h = poly_x_mult_y fix f :: "'a poly" assume "?h f = 0" then have "poly2 (?h f) x 1 = 0" for x by simp from this[unfolded poly2_poly_x_mult_y] show "f = 0" by auto qed lemma degree_poly_x_mult_y[simp]: fixes p :: "'a :: {idom, ring_char_0} poly" shows "degree (poly_x_mult_y p) = degree p" (is "?l = ?r") proof(rule antisym) show "?r \ ?l" by (cases "p=0", auto intro: le_degree simp: coeff_poly_x_mult_y) show "?l \ ?r" unfolding poly_x_mult_y_def by (auto intro: degree_sum_le le_trans[OF degree_monom_le]) qed interpretation poly_x_mult_y_hom: unit_preserving_hom "poly_x_mult_y :: 'a :: field_char_0 poly \ _" proof(unfold_locales) let ?h = "poly_x_mult_y :: 'a poly \ _" fix f :: "'a poly" assume unit: "?h f dvd 1" then have "degree (?h f) = 0" and "coeff (?h f) 0 dvd 1" unfolding poly_dvd_1 by auto then have deg: "degree f = 0" by (auto simp add: degree_monom_eq) with unit show "f dvd 1" by(cases "f = 0", auto) qed lemmas poly_y_x_o_poly_lift = o_def[of poly_y_x poly_lift, unfolded poly_y_x_poly_lift] lemma irreducible_dvd_degree: assumes "(f::'a::field poly) dvd g" "irreducible g" "degree f > 0" shows "degree f = degree g" using assms by (metis irreducible_altdef degree_0 dvd_refl is_unit_field_poly linorder_neqE_nat poly_divides_conv0) lemma coprime_poly_x_mult_y_poly_lift: fixes p q :: "'a :: field_char_0 poly" assumes degp: "degree p > 0" and degq: "degree q > 0" and nz: "poly p 0 \ 0 \ poly q 0 \ 0" shows "coprime (poly_x_mult_y p) (poly_lift q)" proof(rule ccontr) from degp have p: "\ p dvd 1" by (auto simp: dvd_const) from degp have p0: "p \ 0" by auto from mset_factors_exist[of p, OF p0 p] obtain F where F: "mset_factors F p" by auto then have pF: "prod_mset (image_mset poly_x_mult_y F) = poly_x_mult_y p" by (auto simp: hom_distribs) from degq have q: "\ is_unit q" by (auto simp: dvd_const) from degq have q0: "q \ 0" by auto from mset_factors_exist[OF q0 q] obtain G where G: "mset_factors G q" by auto with poly_lift_hom.hom_mset_factors have pG: "mset_factors (image_mset poly_lift G) (poly_lift q)" by auto from poly_y_x_hom.hom_mset_factors[OF this] have pG: "mset_factors (image_mset coeff_lift G) [:q:]" by (auto simp: poly_y_x_poly_lift monom_0 image_mset.compositionality poly_y_x_o_poly_lift) assume "\ coprime (poly_x_mult_y p) (poly_lift q)" then have "\ coprime (poly_y_x (poly_x_mult_y p)) (poly_y_x (poly_lift q))" by (simp del: coprime_iff_coprime) from this[unfolded not_coprime_iff_common_factor] obtain r where rp: "r dvd poly_y_x (poly_x_mult_y p)" and rq: "r dvd poly_y_x (poly_lift q)" and rU: "\ r dvd 1" by auto from rp p0 have r0: "r \ 0" by auto from mset_factors_exist[OF r0 rU] obtain H where H: "mset_factors H r" by auto then have "H \ {#}" by auto then obtain h where hH: "h \# H" by fastforce with H mset_factors_imp_dvd have hr: "h dvd r" and h: "irreducible h" by auto from irreducible_not_unit[OF h] have hU: "\ h dvd 1" by auto from hr rp have "h dvd poly_y_x (poly_x_mult_y p)" by (rule dvd_trans) note this[folded pF,unfolded poly_y_x_hom.hom_prod_mset image_mset.compositionality] from prime_elem_dvd_prod_mset[OF h[folded prime_elem_iff_irreducible] this] obtain f where f: "f \# F" and hf: "h dvd poly_y_x (poly_x_mult_y f)" by auto have irrF: "irreducible f" using f F by blast from dvd_trans[OF hr rq] have "h dvd [:q:]" by (simp add: poly_y_x_poly_lift monom_0) from irreducible_dvd_imp_factor[OF this h pG] q0 obtain g where g: "g \# G" and gh: "[:g:] dvd h" by auto from dvd_trans[OF gh hf] have *: "[:g:] dvd poly_y_x (poly_x_mult_y f)" using dvd_trans by auto show False proof (cases "poly f 0 = 0") case f_0: False from poly_hom.hom_dvd[OF *] have "g dvd poly (poly_y_x (poly_x_mult_y f)) [:0:]" by simp also have "... = [:poly f 0:]" by (intro poly_ext, fold poly2_def, simp add: poly2_poly_x_mult_y) also have "... dvd 1" using f_0 by auto finally have "g dvd 1". with g G show False by (auto elim!: mset_factorsE dest!: irreducible_not_unit) next case True hence "[:0,1:] dvd f" by (unfold dvd_iff_poly_eq_0, simp) from irreducible_dvd_degree[OF this irrF] have "degree f = 1" by auto from degree1_coeffs[OF this] True obtain c where c: "c \ 0" and f: "f = [:0,c:]" by auto from g G have irrG: "irreducible g" by auto from poly_hom.hom_dvd[OF *] have "g dvd poly (poly_y_x (poly_x_mult_y f)) 1" by simp also have "\ = f" by (auto simp: f poly_x_mult_y_code Let_def c poly_y_x_pCons map_poly_monom poly_monom poly_lift_def) also have "\ dvd [:0,1:]" unfolding f dvd_def using c by (intro exI[of _ "[: inverse c :]"], auto) finally have g01: "g dvd [:0,1:]" . from divides_degree[OF this] irrG have "degree g = 1" by auto from degree1_coeffs[OF this] obtain a b where g: "g = [:b,a:]" and a: "a \ 0" by auto from g01[unfolded dvd_def] g obtain k where id: "[:0,1:] = g * k" by auto from id have 0: "g \ 0" "k \ 0" by auto from arg_cong[OF id, of degree] have "degree k = 0" unfolding degree_mult_eq[OF 0] unfolding g using a by auto from degree0_coeffs[OF this] obtain kk where k: "k = [:kk:]" by auto from id[unfolded g k] a have "b = 0" by auto hence "poly g 0 = 0" by (auto simp: g) from True this nz \f \# F\ \g \# G\ F G show False by (auto dest!:mset_factors_imp_dvd elim:dvdE) qed qed lemma poly_div_nonzero: fixes p q :: "'a :: field_char_0 poly" assumes p0: "p \ 0" and q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and p_0: "poly p 0 \ 0 \ poly q 0 \ 0" shows "poly_div p q \ 0" proof have degp: "degree p > 0" using le_0_eq order_degree order_root p0 x by (metis gr0I) have degq: "degree q > 0" using le_0_eq order_degree order_root q0 y by (metis gr0I) assume 0: "poly_div p q = 0" from resultant_zero_imp_common_factor[OF _ this[unfolded poly_div_def]] degp and coprime_poly_x_mult_y_poly_lift[OF degp degq] p_0 show False by auto qed subsubsection \Summary for division\ text \Now we lift the results to one that uses @{const ipoly}, by showing some homomorphism lemmas.\ lemma (in inj_comm_ring_hom) poly_x_mult_y_hom: "poly_x_mult_y (map_poly hom p) = map_poly (map_poly hom) (poly_x_mult_y p)" proof - interpret mh: map_poly_inj_comm_ring_hom.. interpret mmh: map_poly_inj_comm_ring_hom "map_poly hom".. show ?thesis unfolding poly_x_mult_y_def by (simp add: hom_distribs) qed lemma (in inj_comm_ring_hom) poly_div_hom: "map_poly hom (poly_div p q) = poly_div (map_poly hom p) (map_poly hom q)" proof - have zero: "\x. hom x = 0 \ x = 0" by simp interpret mh: map_poly_inj_comm_ring_hom.. show ?thesis unfolding poly_div_def mh.resultant_hom[symmetric] by (simp add: poly_x_mult_y_hom) qed lemma ipoly_poly_div: fixes x y :: "'a :: field_char_0" assumes "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "y \ 0" shows "ipoly (poly_div p q) (x/y) = 0" by (unfold of_int_hom.poly_div_hom, rule poly_div, insert assms, auto) lemma ipoly_poly_div_nonzero: fixes x y :: "'a :: field_char_0" assumes "p \ 0" and "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "poly p 0 \ 0 \ poly q 0 \ 0" shows "poly_div p q \ 0" proof- from assms have "(of_int_poly (poly_div p q) :: 'a poly) \ 0" using of_int_hom.poly_map_poly[of p] by (subst of_int_hom.poly_div_hom, subst poly_div_nonzero, auto) then show ?thesis by auto qed lemma represents_div: fixes x y :: "'a :: field_char_0" assumes "p represents x" and "q represents y" and "poly q 0 \ 0" shows "(poly_div p q) represents (x / y)" using assms by (intro representsI ipoly_poly_div ipoly_poly_div_nonzero, auto) subsection \Multiplication of Algebraic Numbers\ definition poly_mult where "poly_mult p q \ poly_div p (reflect_poly q)" lemma represents_mult: assumes px: "p represents x" and qy: "q represents y" and q_0: "poly q 0 \ 0" shows "(poly_mult p q) represents (x * y)" proof- from q_0 qy have y0: "y \ 0" by auto from represents_inverse[OF y0 qy] y0 px q_0 have "poly_mult p q represents x / (inverse y)" unfolding poly_mult_def by (intro represents_div, auto) with y0 show ?thesis by (simp add: field_simps) qed subsection \Summary: Closure Properties of Algebraic Numbers\ lemma algebraic_representsI: "p represents x \ algebraic x" unfolding represents_def algebraic_altdef_ipoly by auto lemma algebraic_of_rat: "algebraic (of_rat x)" by (rule algebraic_representsI[OF poly_rat_represents_of_rat]) lemma algebraic_uminus: "algebraic x \ algebraic (-x)" by (auto dest: algebraic_imp_represents_irreducible intro: algebraic_representsI represents_uminus) lemma algebraic_inverse: "algebraic x \ algebraic (inverse x)" using algebraic_of_rat[of 0] by (cases "x = 0", auto dest: algebraic_imp_represents_irreducible intro: algebraic_representsI represents_inverse) lemma algebraic_plus: "algebraic x \ algebraic y \ algebraic (x + y)" by (auto dest!: algebraic_imp_represents_irreducible_cf_pos intro!: algebraic_representsI[OF represents_add]) lemma algebraic_div: assumes x: "algebraic x" and y: "algebraic y" shows "algebraic (x/y)" proof(cases "y = 0 \ x = 0") case True then show ?thesis using algebraic_of_rat[of 0] by auto next case False then have x0: "x \ 0" and y0: "y \ 0" by auto from x y obtain p q where px: "p represents x" and irr: "irreducible q" and qy: "q represents y" by (auto dest!: algebraic_imp_represents_irreducible) show ?thesis using False px represents_irr_non_0[OF irr qy] by (auto intro!: algebraic_representsI[OF represents_div] qy) qed lemma algebraic_times: "algebraic x \ algebraic y \ algebraic (x * y)" using algebraic_div[OF _ algebraic_inverse, of x y] by (simp add: field_simps) lemma algebraic_root: "algebraic x \ algebraic (root n x)" proof - assume "algebraic x" then obtain p where p: "p represents x" by (auto dest: algebraic_imp_represents_irreducible_cf_pos) from algebraic_representsI[OF represents_nth_root_neg_real[OF _ this, of n]] algebraic_representsI[OF represents_nth_root_pos_real[OF _ this, of n]] algebraic_of_rat[of 0] show ?thesis by (cases "n = 0", force, cases "n > 0", force, cases "n < 0", auto) qed lemma algebraic_nth_root: "n \ 0 \ algebraic x \ y^n = x \ algebraic y" by (auto dest: algebraic_imp_represents_irreducible_cf_pos intro: algebraic_representsI represents_nth_root) subsection \More on algebraic integers\ (* TODO: this is actually equal to @{term "(-1)^(m*n)"}, but we need a bit more theory on permutations to show this with a reasonable amount of effort. *) definition poly_add_sign :: "nat \ nat \ 'a :: comm_ring_1" where "poly_add_sign m n = signof (\i. if i < n then m + i else if i < m + n then i - n else i)" lemma lead_coeff_poly_add: fixes p q :: "'a :: {idom, semiring_char_0} poly" defines "m \ degree p" and "n \ degree q" assumes "lead_coeff p = 1" "lead_coeff q = 1" "m > 0" "n > 0" shows "lead_coeff (poly_add p q :: 'a poly) = poly_add_sign m n" proof - from assms have [simp]: "p \ 0" "q \ 0" by auto define M where "M = sylvester_mat (poly_x_minus_y p) (poly_lift q)" define \ :: "nat \ nat" where "\ = (\i. if i < n then m + i else if i < m + n then i - n else i)" have \: "\ permutes {0.._def inj_on_def) have nz: "M $$ (i, \ i) \ 0" if "i < m + n" for i using that by (auto simp: M_def \_def sylvester_index_mat m_def n_def) (* have "{(i,j). i \ {.. j \ {.. i < j \ \ i > \ j} = {.. {n.. ?lhs" thus "ij \ ?rhs" by (simp add: \_def split: prod.splits if_splits) auto qed (auto simp: \_def) hence "inversions_on {.. = n * m" by (simp add: inversions_on_def) hence "signof \ = (-1)^(m*n)" using \ by (simp add: signof_def sign_def evenperm_iff_even_inversions) *) have indices_eq: "{0.. (+) n ` {.. \. signof \ * (\i=0.. i)))" have "degree (f \) = degree (\i=0.. i))" using nz by (auto simp: f_def degree_mult_eq signof_def) also have "\ = (\i=0.. i)))" using nz by (subst degree_prod_eq_sum_degree) auto also have "\ = (\i i))) + (\i (n + i))))" by (subst indices_eq, subst sum.union_disjoint) (auto simp: sum.reindex) also have "(\i i))) = (\i_def m_def n_def) also have "(\i (n + i)))) = (\i_def m_def n_def) finally have deg_f1: "degree (f \) = m * n" by simp have deg_f2: "degree (f \) < m * n" if "\ permutes {0.. \ \" for \ proof (cases "\i\{0.. i) = 0") case True hence *: "(\i = 0.. i)) = 0" by auto show ?thesis using \m > 0\ \n > 0\ by (simp add: f_def *) next case False note nz = this from that have \_less: "\ i < m + n" if "i < m + n" for i using permutes_in_image[OF \\ permutes _\] that by auto have "degree (f \) = degree (\i=0.. i))" using nz by (auto simp: f_def degree_mult_eq signof_def) also have "\ = (\i=0.. i)))" using nz by (subst degree_prod_eq_sum_degree) auto also have "\ = (\i i))) + (\i (n + i))))" by (subst indices_eq, subst sum.union_disjoint) (auto simp: sum.reindex) also have "(\i (n + i)))) = (\i_less by (intro sum.cong) (auto simp: M_def sylvester_index_mat \_def m_def n_def) also have "(\i i))) < (\ix\{.. x)) \ m" using \_less by (auto simp: M_def sylvester_index_mat \_def m_def n_def degree_coeff_poly_x_minus_y) next have "\i i \ \ i" proof (rule ccontr) assume nex: "~(\i i \ \ i)" have "\i\m+n-k. \ i = \ i" if "k \ m" for k using that proof (induction k) case 0 thus ?case using \\ permutes _\ \\ permutes _\ by (fastforce simp: permutes_def) next case (Suc k) have IH: "\ i = \ i" if "i \ m+n-k" for i using Suc.prems Suc.IH that by auto from nz have "M $$ (m + n - Suc k, \ (m + n - Suc k)) \ 0" using Suc.prems by auto moreover have "m + n - Suc k \ n" using Suc.prems by auto ultimately have "\ (m+n-Suc k) \ m-Suc k" using assms \_less[of "m+n-Suc k"] Suc.prems by (auto simp: M_def sylvester_index_mat m_def n_def split: if_splits) have "\(\ (m+n-Suc k) > m - Suc k)" proof assume *: "\ (m+n-Suc k) > m - Suc k" have less: "\ (m+n-Suc k) < m" proof (rule ccontr) assume *: "\\ (m + n - Suc k) < m" define j where "j = \ (m + n - Suc k) - m" have "\ (m + n - Suc k) = m + j" using * by (simp add: j_def) moreover { have "j < n" using \_less[of "m+n-Suc k"] \m > 0\ \n > 0\ by (simp add: j_def) hence "\ j = \ j" using nex by auto with \j < n\ have "\ j = m + j" by (auto simp: \_def) } ultimately have "\ (m + n - Suc k) = \ j" by simp hence "m + n - Suc k = j" using permutes_inj[OF \\ permutes _\] unfolding inj_def by blast thus False using \n \ m + n - Suc k\ \_less[of "m+n-Suc k"] \n > 0\ unfolding j_def by linarith qed define j where "j = \ (m+n-Suc k) - (m - Suc k)" from * have j: "\ (m+n-Suc k) = m - Suc k + j" "j > 0" by (auto simp: j_def) have "\ (m+n-Suc k + j) = \ (m+n - Suc k + j)" using * by (intro IH) (auto simp: j_def) also { have "j < Suc k" using less by (auto simp: j_def algebra_simps) hence "m + n - Suc k + j < m + n" using \m > 0\ \n > 0\ Suc.prems by linarith hence "\ (m +n - Suc k + j) = m - Suc k + j" unfolding \_def using Suc.prems by (simp add: \_def) } finally have "\ (m + n - Suc k + j) = \ (m + n - Suc k)" using j by simp hence "m + n - Suc k + j = m + n - Suc k" using permutes_inj[OF \\ permutes _\] unfolding inj_def by blast thus False using \j > 0\ by simp qed with \\ (m+n-Suc k) \ m-Suc k\ have eq: "\ (m+n-Suc k) = m - Suc k" by linarith show ?case proof safe fix i :: nat assume i: "i \ m + n - Suc k" show "\ i = \ i" using eq Suc.prems \m > 0\ IH i proof (cases "i = m + n - Suc k") case True thus ?thesis using eq Suc.prems \m > 0\ by (auto simp: \_def) qed (use IH i in auto) qed qed from this[of m] and nex have "\ i = \ i" for i by (cases "i \ n") auto hence "\ = \" by force thus False using \\ \ \\ by contradiction qed then obtain i where i: "i < n" "\ i \ \ i" by auto have "\ i < m + n" using i by (intro \_less) auto moreover have "\ i = m + i" using i by (auto simp: \_def) ultimately have "degree (M $$ (i, \ i)) < m" using i \m > 0\ by (auto simp: M_def m_def n_def sylvester_index_mat degree_coeff_poly_x_minus_y) thus "\i\{.. i)) < m" using i by blast qed auto finally show "degree (f \) < m * n" by (simp add: mult_ac) qed have "lead_coeff (f \) = poly_add_sign m n" proof - have "lead_coeff (f \) = signof \ * (\i=0.. i)))" by (simp add: f_def signof_def lead_coeff_prod) also have "(\i=0.. i))) = (\i i))) * (\i (n + i))))" by (subst indices_eq, subst prod.union_disjoint) (auto simp: prod.reindex) also have "(\i i))) = (\i_def sylvester_index_mat) also have "(\i (n + i)))) = (\i_def sylvester_index_mat) also have "signof \ = poly_add_sign m n" by (simp add: \_def poly_add_sign_def m_def n_def cong: if_cong) finally show ?thesis using assms by simp qed have "lead_coeff (poly_add p q) = lead_coeff (det (sylvester_mat (poly_x_minus_y p) (poly_lift q)))" by (simp add: poly_add_def resultant_def) also have "det (sylvester_mat (poly_x_minus_y p) (poly_lift q)) = (\\ | \ permutes {0..)" by (simp add: det_def m_def n_def M_def f_def) also have "{\. \ permutes {0.. ({\. \ permutes {0..})" using \ by auto also have "(\\\\. f \) = (\\\{\. \ permutes {0..}. f \) + f \" by (subst sum.insert) (auto simp: finite_permutations) also have "lead_coeff \ = lead_coeff (f \)" proof - have "degree (\\\{\. \ permutes {0..}. f \) < m * n" using assms by (intro degree_sum_smaller deg_f2) (auto simp: m_def n_def finite_permutations) with deg_f1 show ?thesis by (subst lead_coeff_add_le) auto qed finally show ?thesis using \lead_coeff (f \) = _\ by simp qed lemma lead_coeff_poly_mult: fixes p q :: "'a :: {idom, ring_char_0} poly" defines "m \ degree p" and "n \ degree q" assumes "lead_coeff p = 1" "lead_coeff q = 1" "m > 0" "n > 0" assumes "coeff q 0 \ 0" shows "lead_coeff (poly_mult p q :: 'a poly) = 1" proof - from assms have [simp]: "p \ 0" "q \ 0" by auto have [simp]: "degree (reflect_poly q) = n" using assms by (subst degree_reflect_poly_eq) (auto simp: n_def) define M where "M = sylvester_mat (poly_x_mult_y p) (poly_lift (reflect_poly q))" have nz: "M $$ (i, i) \ 0" if "i < m + n" for i using that by (auto simp: M_def sylvester_index_mat m_def n_def coeff_poly_x_mult_y) have indices_eq: "{0.. (+) n ` {.. \. signof \ * (\i=0.. i)))" have "degree (f id) = degree (\i=0.. = (\i=0.. = (\iiiiii) < m * n" if "\ permutes {0.. \ id" for \ proof (cases "\i\{0.. i) = 0") case True hence *: "(\i = 0.. i)) = 0" by auto show ?thesis using \m > 0\ \n > 0\ by (simp add: f_def *) next case False note nz = this from that have \_less: "\ i < m + n" if "i < m + n" for i using permutes_in_image[OF \\ permutes _\] that by auto have "degree (f \) = degree (\i=0.. i))" using nz by (auto simp: f_def degree_mult_eq signof_def) also have "\ = (\i=0.. i)))" using nz by (subst degree_prod_eq_sum_degree) auto also have "\ = (\i i))) + (\i (n + i))))" by (subst indices_eq, subst sum.union_disjoint) (auto simp: sum.reindex) also have "(\i (n + i)))) = (\i_less by (intro sum.cong) (auto simp: M_def sylvester_index_mat m_def n_def) also have "(\i i))) < (\ix\{.. x)) \ m" using \_less by (auto simp: M_def sylvester_index_mat m_def n_def degree_coeff_poly_x_minus_y coeff_poly_x_mult_y intro: order.trans[OF degree_monom_le]) next have "\i i \ i" proof (rule ccontr) assume nex: "\(\i i \ i)" have "\ i = i" for i using that proof (induction i rule: less_induct) case (less i) consider "i < n" | "i \ {n.. m + n" by force thus ?case proof cases assume "i < n" thus ?thesis using nex by auto next assume "i \ m + n" thus ?thesis using \\ permutes _\ by (auto simp: permutes_def) next assume i: "i \ {n.. j = j" if "j < i" for j using that less.prems by (intro less.IH) auto from nz have "M $$ (i, \ i) \ 0" using i by auto hence "\ i \ i" using i \_less[of i] by (auto simp: M_def sylvester_index_mat m_def n_def) moreover have "\ i \ i" proof (rule ccontr) assume *: "\\ i \ i" from * have "\ (\ i) = \ i" by (subst IH) auto hence "\ i = i" using permutes_inj[OF \\ permutes _\] unfolding inj_def by blast with * show False by simp qed ultimately show ?case by simp qed qed hence "\ = id" by force with \\ \ id\ show False by contradiction qed then obtain i where i: "i < n" "\ i \ i" by auto have "\ i < m + n" using i by (intro \_less) auto hence "degree (M $$ (i, \ i)) < m" using i \m > 0\ by (auto simp: M_def m_def n_def sylvester_index_mat degree_coeff_poly_x_minus_y coeff_poly_x_mult_y intro: le_less_trans[OF degree_monom_le]) thus "\i\{.. i)) < m" using i by blast qed auto finally show "degree (f \) < m * n" by (simp add: mult_ac) qed have "lead_coeff (f id) = 1" proof - have "lead_coeff (f id) = (\i=0..i=0..iiiiii\ | \ permutes {0..)" by (simp add: det_def m_def n_def M_def f_def) also have "{\. \ permutes {0... \ permutes {0..\\\. f \) = (\\\{\. \ permutes {0..) + f id" by (subst sum.insert) (auto simp: finite_permutations) also have "lead_coeff \ = lead_coeff (f id)" proof - have "degree (\\\{\. \ permutes {0..) < m * n" using assms by (intro degree_sum_smaller deg_f2) (auto simp: m_def n_def finite_permutations) with deg_f1 show ?thesis by (subst lead_coeff_add_le) auto qed finally show ?thesis using \lead_coeff (f id) = 1\ by simp qed lemma algebraic_int_plus [intro]: fixes x y :: "'a :: field_char_0" assumes "algebraic_int x" "algebraic_int y" shows "algebraic_int (x + y)" proof - from assms(1) obtain p where p: "lead_coeff p = 1" "ipoly p x = 0" by (auto simp: algebraic_int_altdef_ipoly) from assms(2) obtain q where q: "lead_coeff q = 1" "ipoly q y = 0" by (auto simp: algebraic_int_altdef_ipoly) have deg_pos: "degree p > 0" "degree q > 0" using p q by (auto intro!: Nat.gr0I elim!: degree_eq_zeroE) define r where "r = poly_add_sign (degree p) (degree q) * poly_add p q" have "lead_coeff r = 1" using p q deg_pos by (simp add: r_def lead_coeff_mult poly_add_sign_def signof_def lead_coeff_poly_add) moreover have "ipoly r (x + y) = 0" using p q by (simp add: ipoly_poly_add r_def of_int_poly_hom.hom_mult) ultimately show ?thesis by (auto simp: algebraic_int_altdef_ipoly) qed lemma algebraic_int_times [intro]: fixes x y :: "'a :: field_char_0" assumes "algebraic_int x" "algebraic_int y" shows "algebraic_int (x * y)" proof (cases "y = 0") case [simp]: False from assms(1) obtain p where p: "lead_coeff p = 1" "ipoly p x = 0" by (auto simp: algebraic_int_altdef_ipoly) from assms(2) obtain q where q: "lead_coeff q = 1" "ipoly q y = 0" by (auto simp: algebraic_int_altdef_ipoly) have deg_pos: "degree p > 0" "degree q > 0" using p q by (auto intro!: Nat.gr0I elim!: degree_eq_zeroE) have [simp]: "q \ 0" using q by auto define n where "n = Polynomial.order 0 q" have "monom 1 n dvd q" by (simp add: n_def monom_1_dvd_iff) then obtain q' where q_split: "q = q' * monom 1 n" by auto have "Polynomial.order 0 q = Polynomial.order 0 q' + n" using \q \ 0\ unfolding q_split by (subst order_mult) auto hence "poly q' 0 \ 0" unfolding n_def using \q \ 0\ by (simp add: q_split order_root) have q': "ipoly q' y = 0" "lead_coeff q' = 1" using q_split q by (auto simp: of_int_poly_hom.hom_mult poly_monom lead_coeff_mult degree_monom_eq) from this have deg_pos': "degree q' > 0" by (intro Nat.gr0I) (auto elim!: degree_eq_zeroE) from \poly q' 0 \ 0\ have [simp]: "coeff q' 0 \ 0" by (auto simp: monom_1_dvd_iff' poly_0_coeff_0) have "p represents x" "q' represents y" using p q' by (auto simp: represents_def) hence "poly_mult p q' represents x * y" by (rule represents_mult) (simp add: poly_0_coeff_0) moreover have "lead_coeff (poly_mult p q') = 1" using p deg_pos q' deg_pos' by (simp add: lead_coeff_mult lead_coeff_poly_mult) ultimately show ?thesis by (auto simp: algebraic_int_altdef_ipoly represents_def) qed auto lemma algebraic_int_power [intro]: "algebraic_int (x :: 'a :: field_char_0) \ algebraic_int (x ^ n)" by (induction n) auto lemma algebraic_int_diff [intro]: fixes x y :: "'a :: field_char_0" assumes "algebraic_int x" "algebraic_int y" shows "algebraic_int (x - y)" using algebraic_int_plus[OF assms(1) algebraic_int_minus[OF assms(2)]] by simp lemma algebraic_int_sum [intro]: "(\x. x \ A \ algebraic_int (f x :: 'a :: field_char_0)) \ algebraic_int (sum f A)" by (induction A rule: infinite_finite_induct) auto lemma algebraic_int_prod [intro]: "(\x. x \ A \ algebraic_int (f x :: 'a :: field_char_0)) \ algebraic_int (prod f A)" by (induction A rule: infinite_finite_induct) auto lemma algebraic_int_nth_root_real_iff: "algebraic_int (root n x) \ n = 0 \ algebraic_int x" proof - have "algebraic_int x" if "algebraic_int (root n x)" "n \ 0" proof - from that(1) have "algebraic_int (root n x ^ n)" by auto also have "root n x ^ n = (if even n then \x\ else x)" using sgn_power_root[of n x] that(2) by (auto simp: sgn_if split: if_splits) finally show ?thesis by (auto split: if_splits) qed thus ?thesis by auto qed lemma algebraic_int_power_iff: "algebraic_int (x ^ n :: 'a :: field_char_0) \ n = 0 \ algebraic_int x" proof - have "algebraic_int x" if "algebraic_int (x ^ n)" "n > 0" proof (rule algebraic_int_root) show "poly (monom 1 n) x = x ^ n" by (auto simp: poly_monom) qed (use that in \auto simp: degree_monom_eq\) thus ?thesis by auto qed lemma algebraic_int_power_iff' [simp]: "n > 0 \ algebraic_int (x ^ n :: 'a :: field_char_0) \ algebraic_int x" by (subst algebraic_int_power_iff) auto lemma algebraic_int_sqrt_iff [simp]: "algebraic_int (sqrt x) \ algebraic_int x" by (simp add: sqrt_def algebraic_int_nth_root_real_iff) lemma algebraic_int_csqrt_iff [simp]: "algebraic_int (csqrt x) \ algebraic_int x" proof assume "algebraic_int (csqrt x)" hence "algebraic_int (csqrt x ^ 2)" by (rule algebraic_int_power) thus "algebraic_int x" by simp qed auto lemma algebraic_int_norm_complex [intro]: assumes "algebraic_int (z :: complex)" shows "algebraic_int (norm z)" proof - from assms have "algebraic_int (z * cnj z)" by auto also have "z * cnj z = of_real (norm z ^ 2)" by (rule complex_norm_square [symmetric]) finally show ?thesis by simp qed hide_const (open) x_y end diff --git a/thys/Berlekamp_Zassenhaus/Missing_Multiset2.thy b/thys/Berlekamp_Zassenhaus/More_Missing_Multiset.thy rename from thys/Berlekamp_Zassenhaus/Missing_Multiset2.thy rename to thys/Berlekamp_Zassenhaus/More_Missing_Multiset.thy --- a/thys/Berlekamp_Zassenhaus/Missing_Multiset2.thy +++ b/thys/Berlekamp_Zassenhaus/More_Missing_Multiset.thy @@ -1,270 +1,149 @@ -theory Missing_Multiset2 - imports "HOL-Combinatorics.Permutations" - Containers.Containers_Auxiliary (* only for a lemma *) +theory More_Missing_Multiset + imports + "HOL-Combinatorics.Permutations" + Polynomial_Factorization.Missing_Multiset begin -subsubsection \Missing muiltiset\ - -lemma id_imp_bij: - assumes id: "\x. f (f x) = x" shows "bij f" -proof (intro bijI injI surjI[of f, OF id]) - fix x y assume "f x = f y" - then have "f (f x) = f (f y)" by auto - with id show "x = y" by auto -qed - -lemma rel_mset_Zero_iff[simp]: - shows "rel_mset rel {#} Y \ Y = {#}" and "rel_mset rel X {#} \ X = {#}" - using rel_mset_Zero rel_mset_size by (fastforce, fastforce) - -definition "is_mset_set X \ \x \# X. count X x = 1" - -lemma is_mset_setD[dest]: "is_mset_set X \ x \# X \ count X x = 1" - unfolding is_mset_set_def by auto - -lemma is_mset_setI[intro]: - assumes "\x. x \# X \ count X x = 1" - shows "is_mset_set X" - using assms unfolding is_mset_set_def by auto - -lemma is_mset_set[simp]: "is_mset_set (mset_set X)" - unfolding is_mset_set_def - by (meson count_mset_set(1) count_mset_set(2) count_mset_set(3) not_in_iff) - -lemma is_mset_set_add[simp]: - "is_mset_set (X + {#x#}) \ is_mset_set X \ x \# X" (is "?L \ ?R") -proof(intro iffI conjI) - assume L: ?L - with count_eq_zero_iff count_single show "is_mset_set X" - unfolding is_mset_set_def - by (metis (no_types, hide_lams) add_mset_add_single count_add_mset nat.inject set_mset_add_mset_insert union_single_eq_member) - show "x \# X" - proof - assume "x \# X" - then have "count (X + {#x#}) x > 1" by auto - with L show False by (auto simp: is_mset_set_def) - qed -next - assume R: ?R show ?L - proof - fix x' assume x': "x' \# X + {#x#}" - show "count (X + {#x#}) x' = 1" - proof(cases "x' \# X") - case True with R have "count X x' = 1" by auto - moreover from True R have "count {#x#} x' = 0" by auto - ultimately show ?thesis by auto - next - case False then have "count X x' = 0" by (simp add: not_in_iff) - with R x' show ?thesis by auto - qed - qed +lemma rel_mset_free: + assumes rel: "rel_mset rel X Y" and xs: "mset xs = X" + shows "\ys. mset ys = Y \ list_all2 rel xs ys" +proof- + from rel[unfolded rel_mset_def] obtain xs' ys' + where xs': "mset xs' = X" and ys': "mset ys' = Y" and xsys': "list_all2 rel xs' ys'" by auto + from xs' xs have "mset xs = mset xs'" by auto + from mset_eq_permutation[OF this] + obtain f where perm: "f permutes {..i. i < length xs \ xs ! i = xs' ! f i" by auto + note [simp] = list_all2_lengthD[OF xsys',symmetric] + note [simp] = atLeast0LessThan[symmetric] + note bij = permutes_bij[OF perm] + define ys where "ys \ map (nth ys' \ f) [0..x | x \# X \ y = f x. count X x)" -proof(induct X) - case empty show ?case by auto -next - case (add x X) - define X' where "X' \ X + {#x#}" - have "(\z | z \# X' \ y = f z. count (X + {#x#}) z) = - (\z | z \# X' \ y = f z. count X z) + (\z | z \# X' \ y = f z. count {#x#} z)" - unfolding plus_multiset.rep_eq sum.distrib.. - also have split: - "{z. z \# X' \ y = f z} = - {z. z \# X' \ y = f z \ z \ x} \ {z. z \# X' \ y = f z \ z = x}" by blast - then have "(\z | z \# X' \ y = f z. count {#x#} z) = - (\z | z \# X' \ y = f z \ z = x. count {#x#} z)" - unfolding split by (subst sum.union_disjoint, auto) - also have "... = (if y = f x then 1 else 0)" using card_eq_Suc_0_ex1 by (auto simp: X'_def) - also have "(\z | z \# X' \ y = f z. count X z) = (\z | z \# X \ y = f z. count X z)" - proof(cases "x \# X") - case True then have "z \# X' \ z \# X" for z by (auto simp: X'_def) - then show ?thesis by auto - next - case False - have split: "{z. z \# X' \ y = f z} = {z. z \# X \ y = f z} \ {z. z = x \ y = f z}" - by (auto simp: X'_def) - also have "sum (count X) ... = (\z | z \# X \ y = f z. count X z) + (\z | z = x \ y = f z. count X z)" - by (subst sum.union_disjoint, auto simp: False) - also with False have "\z. z = x \ y = f z \ count X z = 0" by (meson count_inI) - with sum.neutral_const have "(\z | z = x \ y = f z. count X z) = 0" by auto - finally show ?thesis by auto - qed - also have "... = count (image_mset f X) y" using add by auto - finally show ?case by (simp add: X'_def) +lemma rel_mset_split: + assumes rel: "rel_mset rel (X1+X2) Y" + shows "\Y1 Y2. Y = Y1 + Y2 \ rel_mset rel X1 Y1 \ rel_mset rel X2 Y2" +proof- + obtain xs1 where xs1: "mset xs1 = X1" using ex_mset by auto + obtain xs2 where xs2: "mset xs2 = X2" using ex_mset by auto + from xs1 xs2 have "mset (xs1 @ xs2) = X1 + X2" by auto + from rel_mset_free[OF rel this] obtain ys + where ys: "mset ys = Y" "list_all2 rel (xs1 @ xs2) ys" by auto + then obtain ys1 ys2 + where ys12: "ys = ys1 @ ys2" + and xs1ys1: "list_all2 rel xs1 ys1" + and xs2ys2: "list_all2 rel xs2 ys2" + using list_all2_append1 by blast + from ys12 ys have "Y = mset ys1 + mset ys2" by auto + moreover from xs1 xs1ys1 have "rel_mset rel X1 (mset ys1)" unfolding rel_mset_def by auto + moreover from xs2 xs2ys2 have "rel_mset rel X2 (mset ys2)" unfolding rel_mset_def by auto + ultimately show ?thesis by (subst exI[of _ "mset ys1"], subst exI[of _ "mset ys2"],auto) qed -lemma is_mset_set_image: - assumes "inj_on f (set_mset X)" and "is_mset_set X" - shows "is_mset_set (image_mset f X)" -proof (cases X) - case empty then show ?thesis by auto -next - case (add x X) - define X' where "X' \ add_mset x X" - with assms add have inj:"inj_on f (set_mset X')" - and X': "is_mset_set X'" by auto - show ?thesis - proof(unfold add, intro is_mset_setI, fold X'_def) - fix y assume "y \# image_mset f X'" - then have "y \ f ` set_mset X'" by auto - with inj have "\!x'. x' \# X' \ y = f x'" by (meson imageE inj_onD) - then obtain x' where x': "{x'. x' \# X' \ y = f x'} = {x'}" by auto - then have "count (image_mset f X') y = count X' x'" - unfolding count_image_mset by auto - also from X' x' have "... = 1" by auto - finally show "count (image_mset f X') y = 1". - qed +lemma rel_mset_OO: + assumes AB: "rel_mset R A B" and BC: "rel_mset S B C" + shows "rel_mset (R OO S) A C" +proof- + from AB obtain as bs where A_as: "A = mset as" and B_bs: "B = mset bs" and as_bs: "list_all2 R as bs" + by (auto simp: rel_mset_def) + from rel_mset_free[OF BC] B_bs obtain cs where C_cs: "C = mset cs" and bs_cs: "list_all2 S bs cs" + by auto + from list_all2_trans[OF _ as_bs bs_cs, of "R OO S"] A_as C_cs + show ?thesis by (auto simp: rel_mset_def) qed (* a variant for "right" *) lemma ex_mset_zip_right: assumes "length xs = length ys" "mset ys' = mset ys" shows "\xs'. length ys' = length xs' \ mset (zip xs' ys') = mset (zip xs ys)" using assms proof (induct xs ys arbitrary: ys' rule: list_induct2) case Nil thus ?case by auto next case (Cons x xs y ys ys') obtain j where j_len: "j < length ys'" and nth_j: "ys' ! j = y" by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) define ysa where "ysa = take j ys' @ drop (Suc j) ys'" have "mset ys' = {#y#} + mset ysa" unfolding ysa_def using j_len nth_j by (metis Cons_nth_drop_Suc union_mset_add_mset_right add_mset_remove_trivial add_diff_cancel_left' append_take_drop_id mset.simps(2) mset_append) hence ms_y: "mset ysa = mset ys" by (simp add: Cons.prems) then obtain xsa where len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" using Cons.hyps(2) by blast define xs' where "xs' = take j xsa @ x # drop j xsa" have ys': "ys' = take j ysa @ y # drop j ysa" using ms_y j_len nth_j Cons.prems ysa_def by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons length_drop size_mset) have j_len': "j \ length ysa" using j_len ys' ysa_def by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) have "length ys' = length xs'" unfolding xs'_def using Cons.prems len_a ms_y by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" unfolding ys' xs'_def apply (rule HOL.trans[OF mset_zip_take_Cons_drop_twice]) using j_len' by (auto simp: len_a ms_a) ultimately show ?case by blast qed lemma list_all2_reorder_right_invariance: assumes rel: "list_all2 R xs ys" and ms_y: "mset ys' = mset ys" shows "\xs'. list_all2 R xs' ys' \ mset xs' = mset xs" proof - have len: "length xs = length ys" using rel list_all2_conv_all_nth by auto obtain xs' where len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" using len ms_y by (metis ex_mset_zip_right) have "list_all2 R xs' ys'" using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) moreover have "mset xs' = mset xs" using len len' ms_xy map_fst_zip mset_map by metis ultimately show ?thesis by blast qed lemma rel_mset_via_perm: "rel_mset rel (mset xs) (mset ys) \ (\zs. mset xs = mset zs \ list_all2 rel zs ys)" proof (unfold rel_mset_def, intro iffI, goal_cases) case 1 then obtain zs ws where zs: "mset zs = mset xs" and ws: "mset ws = mset ys" and zsws: "list_all2 rel zs ws" by auto note list_all2_reorder_right_invariance[OF zsws ws[symmetric], unfolded zs] then show ?case by (auto dest: sym) next case 2 from this show ?case by force qed -lemma rel_mset_free: - assumes rel: "rel_mset rel X Y" and xs: "mset xs = X" - shows "\ys. mset ys = Y \ list_all2 rel xs ys" -proof- - from rel[unfolded rel_mset_def] obtain xs' ys' - where xs': "mset xs' = X" and ys': "mset ys' = Y" and xsys': "list_all2 rel xs' ys'" by auto - from xs' xs have "mset xs = mset xs'" by auto - from mset_eq_permutation[OF this] - obtain f where perm: "f permutes {..i. i < length xs \ xs ! i = xs' ! f i" by auto - note [simp] = list_all2_lengthD[OF xsys',symmetric] - note [simp] = atLeast0LessThan[symmetric] - note bij = permutes_bij[OF perm] - define ys where "ys \ map (nth ys' \ f) [0..Y1 Y2. Y = Y1 + Y2 \ rel_mset rel X1 Y1 \ rel_mset rel X2 Y2" -proof- - obtain xs1 where xs1: "mset xs1 = X1" using ex_mset by auto - obtain xs2 where xs2: "mset xs2 = X2" using ex_mset by auto - from xs1 xs2 have "mset (xs1 @ xs2) = X1 + X2" by auto - from rel_mset_free[OF rel this] obtain ys - where ys: "mset ys = Y" "list_all2 rel (xs1 @ xs2) ys" by auto - then obtain ys1 ys2 - where ys12: "ys = ys1 @ ys2" - and xs1ys1: "list_all2 rel xs1 ys1" - and xs2ys2: "list_all2 rel xs2 ys2" - using list_all2_append1 by blast - from ys12 ys have "Y = mset ys1 + mset ys2" by auto - moreover from xs1 xs1ys1 have "rel_mset rel X1 (mset ys1)" unfolding rel_mset_def by auto - moreover from xs2 xs2ys2 have "rel_mset rel X2 (mset ys2)" unfolding rel_mset_def by auto - ultimately show ?thesis by (subst exI[of _ "mset ys1"], subst exI[of _ "mset ys2"],auto) -qed - -lemma rel_mset_OO: - assumes AB: "rel_mset R A B" and BC: "rel_mset S B C" - shows "rel_mset (R OO S) A C" -proof- - from AB obtain as bs where A_as: "A = mset as" and B_bs: "B = mset bs" and as_bs: "list_all2 R as bs" - by (auto simp: rel_mset_def) - from rel_mset_free[OF BC] B_bs obtain cs where C_cs: "C = mset cs" and bs_cs: "list_all2 S bs cs" - by auto - from list_all2_trans[OF _ as_bs bs_cs, of "R OO S"] A_as C_cs - show ?thesis by (auto simp: rel_mset_def) -qed - end diff --git a/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field.thy b/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field.thy --- a/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field.thy +++ b/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field.thy @@ -1,615 +1,615 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) subsection \Polynomials in a Finite Field\ text \We connect polynomials in a prime field with integer polynomials modulo some prime.\ theory Poly_Mod_Finite_Field imports Finite_Field Polynomial_Interpolation.Ring_Hom_Poly "HOL-Types_To_Sets.Types_To_Sets" - Missing_Multiset2 + More_Missing_Multiset Poly_Mod begin (* TODO: Move -- General transfer rule *) declare rel_mset_Zero[transfer_rule] lemma mset_transfer[transfer_rule]: "(list_all2 rel ===> rel_mset rel) mset mset" proof (intro rel_funI) show "list_all2 rel xs ys \ rel_mset rel (mset xs) (mset ys)" for xs ys proof (induct xs arbitrary: ys) case Nil then show ?case by auto next case IH: (Cons x xs) then show ?case by (auto dest!:msed_rel_invL simp: list_all2_Cons1 intro!:rel_mset_Plus) qed qed abbreviation to_int_poly :: "'a :: finite mod_ring poly \ int poly" where "to_int_poly \ map_poly to_int_mod_ring" interpretation to_int_poly_hom: map_poly_inj_zero_hom to_int_mod_ring .. lemma irreducible\<^sub>d_def_0: fixes f :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" shows "irreducible\<^sub>d f = (degree f \ 0 \ (\ g h. degree g \ 0 \ degree h \ 0 \ f \ g * h))" proof- have "degree g \ 0 \ g \ 0" for g :: "'a poly" by auto note 1 = degree_mult_eq[OF this this, simplified] then show ?thesis by (force elim!: irreducible\<^sub>dE) qed subsection \Transferring to class-based mod-ring\ locale poly_mod_type = poly_mod m for m and ty :: "'a :: nontriv itself" + assumes m: "m = CARD('a)" begin lemma m1: "m > 1" using nontriv[where 'a = 'a] by (auto simp:m) sublocale poly_mod_2 using m1 by unfold_locales definition MP_Rel :: "int poly \ 'a mod_ring poly \ bool" where "MP_Rel f f' \ (Mp f = to_int_poly f')" definition M_Rel :: "int \ 'a mod_ring \ bool" where "M_Rel x x' \ (M x = to_int_mod_ring x')" definition "MF_Rel \ rel_prod M_Rel (rel_mset MP_Rel)" lemma to_int_mod_ring_plus: "to_int_mod_ring ((x :: 'a mod_ring) + y) = M (to_int_mod_ring x + to_int_mod_ring y)" unfolding M_def using m by (transfer, auto) lemma to_int_mod_ring_times: "to_int_mod_ring ((x :: 'a mod_ring) * y) = M (to_int_mod_ring x * to_int_mod_ring y)" unfolding M_def using m by (transfer, auto) lemma degree_MP_Rel [transfer_rule]: "(MP_Rel ===> (=)) degree_m degree" unfolding MP_Rel_def rel_fun_def by (auto intro!: degree_map_poly) lemma eq_M_Rel[transfer_rule]: "(M_Rel ===> M_Rel ===> (=)) (\ x y. M x = M y) (=)" unfolding M_Rel_def rel_fun_def by auto interpretation to_int_mod_ring_hom: map_poly_inj_zero_hom to_int_mod_ring.. lemma eq_MP_Rel[transfer_rule]: "(MP_Rel ===> MP_Rel ===> (=)) (=m) (=)" unfolding MP_Rel_def rel_fun_def by auto lemma eq_Mf_Rel[transfer_rule]: "(MF_Rel ===> MF_Rel ===> (=)) (\ x y. Mf x = Mf y) (=)" proof (intro rel_funI, goal_cases) case (1 cfs Cfs dgs Dgs) obtain c fs where cfs: "cfs = (c,fs)" by force obtain C Fs where Cfs: "Cfs = (C,Fs)" by force obtain d gs where dgs: "dgs = (d,gs)" by force obtain D Gs where Dgs: "Dgs = (D,Gs)" by force note pairs = cfs Cfs dgs Dgs from 1[unfolded pairs MF_Rel_def rel_prod.simps] have *[transfer_rule]: "M_Rel c C" "M_Rel d D" "rel_mset MP_Rel fs Fs" "rel_mset MP_Rel gs Gs" by auto have eq1: "(M c = M d) = (C = D)" by transfer_prover from *(3)[unfolded rel_mset_def] obtain fs' Fs' where fs_eq: "mset fs' = fs" "mset Fs' = Fs" and rel_f: "list_all2 MP_Rel fs' Fs'" by auto from *(4)[unfolded rel_mset_def] obtain gs' Gs' where gs_eq: "mset gs' = gs" "mset Gs' = Gs" and rel_g: "list_all2 MP_Rel gs' Gs'" by auto have eq2: "(image_mset Mp fs = image_mset Mp gs) = (Fs = Gs)" using *(3-4) proof (induct fs arbitrary: Fs gs Gs) case (empty Fs gs Gs) from empty(1) have Fs: "Fs = {#}" unfolding rel_mset_def by auto with empty show ?case by (cases gs; cases Gs; auto simp: rel_mset_def) next case (add f fs Fs' gs' Gs') note [transfer_rule] = add(3) from msed_rel_invL[OF add(2)] obtain Fs F where Fs': "Fs' = Fs + {#F#}" and rel[transfer_rule]: "MP_Rel f F" "rel_mset MP_Rel fs Fs" by auto note IH = add(1)[OF rel(2)] { from add(3)[unfolded rel_mset_def] obtain gs Gs where id: "mset gs = gs'" "mset Gs = Gs'" and rel: "list_all2 MP_Rel gs Gs" by auto have "Mp f \# image_mset Mp gs' \ F \# Gs'" proof - have "?thesis = ((Mp f \ Mp ` set gs) = (F \ set Gs))" unfolding id[symmetric] by simp also have \ using rel proof (induct gs Gs rule: list_all2_induct) case (Cons g gs G Gs) note [transfer_rule] = Cons(1-2) have id: "(Mp g = Mp f) = (F = G)" by (transfer, auto) show ?case using id Cons(3) by auto qed auto finally show ?thesis by simp qed } note id = this show ?case proof (cases "Mp f \# image_mset Mp gs'") case False have "Mp f \# image_mset Mp (fs + {#f#})" by auto with False have F: "image_mset Mp (fs + {#f#}) \ image_mset Mp gs'" by metis with False[unfolded id] show ?thesis unfolding Fs' by auto next case True then obtain g where fg: "Mp f = Mp g" and g: "g \# gs'" by auto from g obtain gs where gs': "gs' = add_mset g gs" by (rule mset_add) from msed_rel_invL[OF add(3)[unfolded gs']] obtain Gs G where Gs': "Gs' = Gs + {# G #}" and gG[transfer_rule]: "MP_Rel g G" and gsGs: "rel_mset MP_Rel gs Gs" by auto have FG: "F = G" by (transfer, simp add: fg) note IH = IH[OF gsGs] show ?thesis unfolding gs' Fs' Gs' by (simp add: fg IH FG) qed qed show "(Mf cfs = Mf dgs) = (Cfs = Dgs)" unfolding pairs Mf_def split by (simp add: eq1 eq2) qed lemmas coeff_map_poly_of_int = coeff_map_poly[of of_int, OF of_int_0] lemma plus_MP_Rel[transfer_rule]: "(MP_Rel ===> MP_Rel ===> MP_Rel) (+) (+)" unfolding MP_Rel_def proof (intro rel_funI, goal_cases) case (1 x f y g) have "Mp (x + y) = Mp (Mp x + Mp y)" by simp also have "\ = Mp (map_poly to_int_mod_ring f + map_poly to_int_mod_ring g)" unfolding 1 .. also have "\ = map_poly to_int_mod_ring (f + g)" unfolding poly_eq_iff Mp_coeff by (auto simp: to_int_mod_ring_plus) finally show ?case . qed lemma times_MP_Rel[transfer_rule]: "(MP_Rel ===> MP_Rel ===> MP_Rel) ((*)) ((*))" unfolding MP_Rel_def proof (intro rel_funI, goal_cases) case (1 x f y g) have "Mp (x * y) = Mp (Mp x * Mp y)" by simp also have "\ = Mp (map_poly to_int_mod_ring f * map_poly to_int_mod_ring g)" unfolding 1 .. also have "\ = map_poly to_int_mod_ring (f * g)" proof - { fix n :: nat define A where "A = {.. n}" have "finite A" unfolding A_def by auto then have "M (\i\n. to_int_mod_ring (coeff f i) * to_int_mod_ring (coeff g (n - i))) = to_int_mod_ring (\i\n. coeff f i * coeff g (n - i))" unfolding A_def[symmetric] proof (induct A) case (insert a A) have "?case = ?case" (is "(?l = ?r) = _") by simp have "?r = to_int_mod_ring (coeff f a * coeff g (n - a) + (\i\ A. coeff f i * coeff g (n - i)))" using insert(1-2) by auto note r = this[unfolded to_int_mod_ring_plus to_int_mod_ring_times] from insert(1-2) have "?l = M (to_int_mod_ring (coeff f a) * to_int_mod_ring (coeff g (n - a)) + M (\i\A. to_int_mod_ring (coeff f i) * to_int_mod_ring (coeff g (n - i))))" by simp also have "M (\i\A. to_int_mod_ring (coeff f i) * to_int_mod_ring (coeff g (n - i))) = to_int_mod_ring (\i\A. coeff f i * coeff g (n - i))" unfolding insert .. finally show ?case unfolding r by simp qed auto } then show ?thesis by (auto intro!:poly_eqI simp: coeff_mult Mp_coeff) qed finally show ?case . qed lemma smult_MP_Rel[transfer_rule]: "(M_Rel ===> MP_Rel ===> MP_Rel) smult smult" unfolding MP_Rel_def M_Rel_def proof (intro rel_funI, goal_cases) case (1 x x' f f') thus ?case unfolding poly_eq_iff coeff Mp_coeff coeff_smult M_def proof (intro allI, goal_cases) case (1 n) have "x * coeff f n mod m = (x mod m) * (coeff f n mod m) mod m" by (simp add: mod_simps) also have "\ = to_int_mod_ring x' * (to_int_mod_ring (coeff f' n)) mod m" using 1 by auto also have " \ = to_int_mod_ring (x' * coeff f' n)" unfolding to_int_mod_ring_times M_def by simp finally show ?case by auto qed qed lemma one_M_Rel[transfer_rule]: "M_Rel 1 1" unfolding M_Rel_def M_def unfolding m by auto lemma one_MP_Rel[transfer_rule]: "MP_Rel 1 1" unfolding MP_Rel_def poly_eq_iff Mp_coeff M_def unfolding m by auto lemma zero_M_Rel[transfer_rule]: "M_Rel 0 0" unfolding M_Rel_def M_def unfolding m by auto lemma zero_MP_Rel[transfer_rule]: "MP_Rel 0 0" unfolding MP_Rel_def poly_eq_iff Mp_coeff M_def unfolding m by auto lemma listprod_MP_Rel[transfer_rule]: "(list_all2 MP_Rel ===> MP_Rel) prod_list prod_list" proof (intro rel_funI, goal_cases) case (1 xs ys) thus ?case proof (induct xs ys rule: list_all2_induct) case (Cons x xs y ys) note [transfer_rule] = this show ?case by simp transfer_prover qed (simp add: one_MP_Rel) qed lemma prod_mset_MP_Rel[transfer_rule]: "(rel_mset MP_Rel ===> MP_Rel) prod_mset prod_mset" proof (intro rel_funI, goal_cases) case (1 xs ys) have "(MP_Rel ===> MP_Rel ===> MP_Rel) ((*)) ((*))" "MP_Rel 1 1" by transfer_prover+ from 1 this show ?case proof (induct xs ys rule: rel_mset_induct) case (add R x xs y ys) note [transfer_rule] = this show ?case by simp transfer_prover qed (simp add: one_MP_Rel) qed lemma right_unique_MP_Rel[transfer_rule]: "right_unique MP_Rel" unfolding right_unique_def MP_Rel_def by auto lemma M_to_int_mod_ring: "M (to_int_mod_ring (x :: 'a mod_ring)) = to_int_mod_ring x" unfolding M_def unfolding m by (transfer, auto) lemma Mp_to_int_poly: "Mp (to_int_poly (f :: 'a mod_ring poly)) = to_int_poly f" by (auto simp: poly_eq_iff Mp_coeff M_to_int_mod_ring) lemma right_total_M_Rel[transfer_rule]: "right_total M_Rel" unfolding right_total_def M_Rel_def using M_to_int_mod_ring by blast lemma left_total_M_Rel[transfer_rule]: "left_total M_Rel" unfolding left_total_def M_Rel_def[abs_def] proof fix x show "\ x' :: 'a mod_ring. M x = to_int_mod_ring x'" unfolding M_def unfolding m by (rule exI[of _ "of_int x"], transfer, simp) qed lemma bi_total_M_Rel[transfer_rule]: "bi_total M_Rel" using right_total_M_Rel left_total_M_Rel by (metis bi_totalI) lemma right_total_MP_Rel[transfer_rule]: "right_total MP_Rel" unfolding right_total_def MP_Rel_def proof fix f :: "'a mod_ring poly" show "\x. Mp x = to_int_poly f" by (intro exI[of _ "to_int_poly f"], simp add: Mp_to_int_poly) qed lemma to_int_mod_ring_of_int_M: "to_int_mod_ring (of_int x :: 'a mod_ring) = M x" unfolding M_def unfolding m by transfer auto lemma Mp_f_representative: "Mp f = to_int_poly (map_poly of_int f :: 'a mod_ring poly)" unfolding Mp_def by (auto intro: poly_eqI simp: coeff_map_poly to_int_mod_ring_of_int_M) lemma left_total_MP_Rel[transfer_rule]: "left_total MP_Rel" unfolding left_total_def MP_Rel_def[abs_def] using Mp_f_representative by blast lemma bi_total_MP_Rel[transfer_rule]: "bi_total MP_Rel" using right_total_MP_Rel left_total_MP_Rel by (metis bi_totalI) lemma bi_total_MF_Rel[transfer_rule]: "bi_total MF_Rel" unfolding MF_Rel_def[abs_def] by (intro prod.bi_total_rel multiset.bi_total_rel bi_total_MP_Rel bi_total_M_Rel) lemma right_total_MF_Rel[transfer_rule]: "right_total MF_Rel" using bi_total_MF_Rel unfolding bi_total_alt_def by auto lemma left_total_MF_Rel[transfer_rule]: "left_total MF_Rel" using bi_total_MF_Rel unfolding bi_total_alt_def by auto lemma domain_RT_rel[transfer_domain_rule]: "Domainp MP_Rel = (\ f. True)" proof fix f :: "int poly" show "Domainp MP_Rel f = True" unfolding MP_Rel_def[abs_def] Domainp.simps by (auto simp: Mp_f_representative) qed lemma mem_MP_Rel[transfer_rule]: "(MP_Rel ===> rel_set MP_Rel ===> (=)) (\ x Y. \y \ Y. eq_m x y) (\)" proof (intro rel_funI iffI) fix x y X Y assume xy: "MP_Rel x y" and XY: "rel_set MP_Rel X Y" { assume "\x' \ X. x =m x'" then obtain x' where x'X: "x' \ X" and xx': "x =m x'" by auto with xy have x'y: "MP_Rel x' y" by (auto simp: MP_Rel_def) from rel_setD1[OF XY x'X] obtain y' where "MP_Rel x' y'" and "y' \ Y" by auto with x'y show "y \ Y" by (auto simp: MP_Rel_def) } assume "y \ Y" from rel_setD2[OF XY this] obtain x' where x'X: "x' \ X" and x'y: "MP_Rel x' y" by auto from xy x'y have "x =m x'" by (auto simp: MP_Rel_def) with x'X show "\x' \ X. x =m x'" by auto qed lemma conversep_MP_Rel_OO_MP_Rel [simp]: "MP_Rel\\ OO MP_Rel = (=)" using Mp_to_int_poly by (intro ext, auto simp: OO_def MP_Rel_def) lemma MP_Rel_OO_conversep_MP_Rel [simp]: "MP_Rel OO MP_Rel\\ = eq_m" by (intro ext, auto simp: OO_def MP_Rel_def Mp_f_representative) lemma conversep_MP_Rel_OO_eq_m [simp]: "MP_Rel\\ OO eq_m = MP_Rel\\" by (intro ext, auto simp: OO_def MP_Rel_def) lemma eq_m_OO_MP_Rel [simp]: "eq_m OO MP_Rel = MP_Rel" by (intro ext, auto simp: OO_def MP_Rel_def) lemma eq_mset_MP_Rel [transfer_rule]: "(rel_mset MP_Rel ===> rel_mset MP_Rel ===> (=)) (rel_mset eq_m) (=)" proof (intro rel_funI iffI) fix A B X Y assume AX: "rel_mset MP_Rel A X" and BY: "rel_mset MP_Rel B Y" { assume AB: "rel_mset eq_m A B" from AX have "rel_mset MP_Rel\\ X A" by (simp add: multiset.rel_flip) note rel_mset_OO[OF this AB] note rel_mset_OO[OF this BY] then show "X = Y" by (simp add: multiset.rel_eq) } assume "X = Y" with BY have "rel_mset MP_Rel\\ X B" by (simp add: multiset.rel_flip) from rel_mset_OO[OF AX this] show "rel_mset eq_m A B" by simp qed lemma dvd_MP_Rel[transfer_rule]: "(MP_Rel ===> MP_Rel ===> (=)) (dvdm) (dvd)" unfolding dvdm_def[abs_def] dvd_def[abs_def] by transfer_prover lemma irreducible_MP_Rel [transfer_rule]: "(MP_Rel ===> (=)) irreducible_m irreducible" unfolding irreducible_m_def irreducible_def by transfer_prover lemma irreducible\<^sub>d_MP_Rel [transfer_rule]: "(MP_Rel ===> (=)) irreducible\<^sub>d_m irreducible\<^sub>d" unfolding irreducible\<^sub>d_m_def[abs_def] irreducible\<^sub>d_def[abs_def] by transfer_prover lemma UNIV_M_Rel[transfer_rule]: "rel_set M_Rel {0.. (=) ===> M_Rel) coeff coeff" unfolding rel_fun_def M_Rel_def MP_Rel_def Mp_coeff[symmetric] by auto lemma M_1_1: "M 1 = 1" unfolding M_def unfolding m by simp lemma square_free_MP_Rel [transfer_rule]: "(MP_Rel ===> (=)) square_free_m square_free" unfolding square_free_m_def[abs_def] square_free_def[abs_def] by (transfer_prover_start, transfer_step+, auto) lemma mset_factors_m_MP_Rel [transfer_rule]: "(rel_mset MP_Rel ===> MP_Rel ===> (=)) mset_factors_m mset_factors" unfolding mset_factors_def mset_factors_m_def by (transfer_prover_start, transfer_step+, auto dest:eq_m_irreducible_m) lemma coprime_MP_Rel [transfer_rule]: "(MP_Rel ===> MP_Rel ===> (=)) coprime_m coprime" unfolding coprime_m_def[abs_def] coprime_def' [abs_def] by (transfer_prover_start, transfer_step+, auto) lemma prime_elem_MP_Rel [transfer_rule]: "(MP_Rel ===> (=)) prime_elem_m prime_elem" unfolding prime_elem_m_def prime_elem_def by transfer_prover end context poly_mod_2 begin lemma non_empty: "{0.. {}" using m1 by auto lemma type_to_set: assumes type_def: "\(Rep :: 'b \ int) Abs. type_definition Rep Abs {0 ..< m :: int}" shows "class.nontriv (TYPE('b))" (is ?a) and "m = int CARD('b)" (is ?b) proof - from type_def obtain rep :: "'b \ int" and abs :: "int \ 'b" where t: "type_definition rep abs {0 ..< m}" by auto have "card (UNIV :: 'b set) = card {0 ..< m}" using t by (rule type_definition.card) also have "\ = m" using m1 by auto finally show ?b .. then show ?a unfolding class.nontriv_def using m1 by auto qed end locale poly_mod_prime_type = poly_mod_type m ty for m :: int and ty :: "'a :: prime_card itself" begin lemma factorization_MP_Rel [transfer_rule]: "(MP_Rel ===> MF_Rel ===> (=)) factorization_m (factorization Irr_Mon)" unfolding rel_fun_def proof (intro allI impI, goal_cases) case (1 f F cfs Cfs) note [transfer_rule] = 1(1) obtain c fs where cfs: "cfs = (c,fs)" by force obtain C Fs where Cfs: "Cfs = (C,Fs)" by force from 1(2)[unfolded rel_prod.simps cfs Cfs MF_Rel_def] have tr[transfer_rule]: "M_Rel c C" "rel_mset MP_Rel fs Fs" by auto have eq: "(f =m smult c (prod_mset fs) = (F = smult C (prod_mset Fs)))" by transfer_prover have "set_mset Fs \ Irr_Mon = (\ x \# Fs. irreducible\<^sub>d x \ monic x)" unfolding Irr_Mon_def by auto also have "\ = (\f\#fs. irreducible\<^sub>d_m f \ monic (Mp f))" proof (rule sym, transfer_prover_start, transfer_step+) { fix f assume "f \# fs" have "monic (Mp f) \ M (coeff f (degree_m f)) = M 1" unfolding Mp_coeff[symmetric] by simp } thus "(\f\#fs. irreducible\<^sub>d_m f \ monic (Mp f)) = (\x\#fs. irreducible\<^sub>d_m x \ M (coeff x (degree_m x)) = M 1)" by auto qed finally show "factorization_m f cfs = factorization Irr_Mon F Cfs" unfolding cfs Cfs factorization_m_def factorization_def split eq by simp qed lemma unique_factorization_MP_Rel [transfer_rule]: "(MP_Rel ===> MF_Rel ===> (=)) unique_factorization_m (unique_factorization Irr_Mon)" unfolding rel_fun_def proof (intro allI impI, goal_cases) case (1 f F cfs Cfs) note [transfer_rule] = 1(1,2) let ?F = "factorization Irr_Mon F" let ?f = "factorization_m f" let ?R = "Collect ?F" let ?L = "Mf ` Collect ?f" note X_to_x = right_total_MF_Rel[unfolded right_total_def, rule_format] { fix X assume "X \ ?R" hence F: "?F X" by simp from X_to_x[of X] obtain x where rel[transfer_rule]: "MF_Rel x X" by blast from F[untransferred] have "Mf x \ ?L" by blast with rel have "\ x. Mf x \ ?L \ MF_Rel x X" by blast } note R_to_L = this show "unique_factorization_m f cfs = unique_factorization Irr_Mon F Cfs" unfolding unique_factorization_m_def unique_factorization_def proof - have fF: "?F Cfs = ?f cfs" by transfer simp have "(?L = {Mf cfs}) = (?L \ {Mf cfs} \ Mf cfs \ ?L)" by blast also have "?L \ {Mf cfs} = (\ dfs. ?f dfs \ Mf dfs = Mf cfs)" by blast also have "\ = (\ y. ?F y \ y = Cfs)" (is "?left = ?right") proof (rule; intro allI impI) fix Dfs assume *: ?left and F: "?F Dfs" from X_to_x[of Dfs] obtain dfs where [transfer_rule]: "MF_Rel dfs Dfs" by auto from F[untransferred] have f: "?f dfs" . from *[rule_format, OF f] have eq: "Mf dfs = Mf cfs" by simp have "(Mf dfs = Mf cfs) = (Dfs = Cfs)" by (transfer_prover_start, transfer_step+, simp) thus "Dfs = Cfs" using eq by simp next fix dfs assume *: ?right and f: "?f dfs" from left_total_MF_Rel obtain Dfs where rel[transfer_rule]: "MF_Rel dfs Dfs" unfolding left_total_def by blast have "?F Dfs" by (transfer, rule f) from *[rule_format, OF this] have eq: "Dfs = Cfs" . have "(Mf dfs = Mf cfs) = (Dfs = Cfs)" by (transfer_prover_start, transfer_step+, simp) thus "Mf dfs = Mf cfs" using eq by simp qed also have "Mf cfs \ ?L = (\ dfs. ?f dfs \ Mf cfs = Mf dfs)" by auto also have "\ = ?F Cfs" unfolding fF proof assume "\ dfs. ?f dfs \ Mf cfs = Mf dfs" then obtain dfs where f: "?f dfs" and id: "Mf dfs = Mf cfs" by auto from f have "?f (Mf dfs)" by simp from this[unfolded id] show "?f cfs" by simp qed blast finally show "(?L = {Mf cfs}) = (?R = {Cfs})" by auto qed qed end context begin private lemma 1: "poly_mod_type TYPE('a :: nontriv) m = (m = int CARD('a))" and 2: "class.nontriv TYPE('a) = (CARD('a) \ 2)" unfolding poly_mod_type_def class.prime_card_def class.nontriv_def poly_mod_prime_type_def by auto private lemma 3: "poly_mod_prime_type TYPE('b) m = (m = int CARD('b))" and 4: "class.prime_card TYPE('b :: prime_card) = prime CARD('b :: prime_card)" unfolding poly_mod_type_def class.prime_card_def class.nontriv_def poly_mod_prime_type_def by auto lemmas poly_mod_type_simps = 1 2 3 4 end lemma remove_duplicate_premise: "(PROP P \ PROP P \ PROP Q) \ (PROP P \ PROP Q)" (is "?l \ ?r") proof (intro Pure.equal_intr_rule) assume p: "PROP P" and ppq: "PROP ?l" from ppq[OF p p] show "PROP Q". next assume p: "PROP P" and pq: "PROP ?r" from pq[OF p] show "PROP Q". qed context poly_mod_prime begin lemma type_to_set: assumes type_def: "\(Rep :: 'b \ int) Abs. type_definition Rep Abs {0 ..< p :: int}" shows "class.prime_card (TYPE('b))" (is ?a) and "p = int CARD('b)" (is ?b) proof - from prime have p2: "p \ 2" by (rule prime_ge_2_int) from type_def obtain rep :: "'b \ int" and abs :: "int \ 'b" where t: "type_definition rep abs {0 ..< p}" by auto have "card (UNIV :: 'b set) = card {0 ..< p}" using t by (rule type_definition.card) also have "\ = p" using p2 by auto finally show ?b .. then show ?a unfolding class.prime_card_def using prime p2 by auto qed end (* it will be nice to be able to automate this *) lemmas (in poly_mod_type) prime_elem_m_dvdm_multD = prime_elem_dvd_multD [where 'a = "'a mod_ring poly",untransferred] lemmas (in poly_mod_2) prime_elem_m_dvdm_multD = poly_mod_type.prime_elem_m_dvdm_multD [unfolded poly_mod_type_simps, internalize_sort "'a :: nontriv", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas(in poly_mod_prime_type) degree_m_mult_eq = degree_mult_eq [where 'a = "'a mod_ring", untransferred] lemmas(in poly_mod_prime) degree_m_mult_eq = poly_mod_prime_type.degree_m_mult_eq [unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemma(in poly_mod_prime) irreducible\<^sub>d_lifting: assumes n: "n \ 0" and deg: "poly_mod.degree_m (p^n) f = degree_m f" and irr: "irreducible\<^sub>d_m f" shows "poly_mod.irreducible\<^sub>d_m (p^n) f" proof - interpret q: poly_mod_2 "p^n" unfolding poly_mod_2_def using n m1 by auto show "q.irreducible\<^sub>d_m f" proof (rule q.irreducible\<^sub>d_mI) from deg irr show "q.degree_m f > 0" by (auto elim: irreducible\<^sub>d_mE) then have pdeg_f: "degree_m f \ 0" by (simp add: deg) note pMp_Mp = Mp_Mp_pow_is_Mp[OF n m1] fix g h assume deg_g: "degree g < q.degree_m f" and deg_h: "degree h < q.degree_m f" and eq: "q.eq_m f (g * h)" from eq have p_f: "f =m (g * h)" using pMp_Mp by metis have "\g =m 0" and "\h =m 0" apply (metis degree_0 mult_zero_left Mp_0 p_f pdeg_f poly_mod.mult_Mp(1)) by (metis degree_0 mult_eq_0_iff Mp_0 mult_Mp(2) p_f pdeg_f) note [simp] = degree_m_mult_eq[OF this] from degree_m_le[of g] deg_g have 2: "degree_m g < degree_m f" by (fold deg, auto) from degree_m_le[of h] deg_h have 3: "degree_m h < degree_m f" by (fold deg, auto) from irreducible\<^sub>d_mD(2)[OF irr 2 3] p_f show False by auto qed qed (* Lifting UFD properties *) lemmas (in poly_mod_prime_type) mset_factors_exist = mset_factors_exist[where 'a = "'a mod_ring poly",untransferred] lemmas (in poly_mod_prime) mset_factors_exist = poly_mod_prime_type.mset_factors_exist [unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas (in poly_mod_prime_type) mset_factors_unique = mset_factors_unique[where 'a = "'a mod_ring poly",untransferred] lemmas (in poly_mod_prime) mset_factors_unique = poly_mod_prime_type.mset_factors_unique [unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas (in poly_mod_prime_type) prime_elem_iff_irreducible = prime_elem_iff_irreducible[where 'a = "'a mod_ring poly",untransferred] lemmas (in poly_mod_prime) prime_elem_iff_irreducible[simp] = poly_mod_prime_type.prime_elem_iff_irreducible [unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas (in poly_mod_prime_type) irreducible_connect = irreducible_connect_field[where 'a = "'a mod_ring", untransferred] lemmas (in poly_mod_prime) irreducible_connect[simp] = poly_mod_prime_type.irreducible_connect [unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas (in poly_mod_prime_type) irreducible_degree = irreducible_degree_field[where 'a = "'a mod_ring", untransferred] lemmas (in poly_mod_prime) irreducible_degree = poly_mod_prime_type.irreducible_degree [unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] end diff --git a/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy b/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy --- a/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy +++ b/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy @@ -1,971 +1,971 @@ theory Unique_Factorization imports Polynomial_Interpolation.Ring_Hom_Poly Polynomial_Factorization.Polynomial_Divisibility "HOL-Combinatorics.Permutations" "HOL-Computational_Algebra.Euclidean_Algorithm" Containers.Containers_Auxiliary (* only for a lemma *) - Missing_Multiset2 + More_Missing_Multiset "HOL-Algebra.Divisibility" begin hide_const(open) Divisibility.prime Divisibility.irreducible hide_fact(open) Divisibility.irreducible_def Divisibility.irreducibleI Divisibility.irreducibleD Divisibility.irreducibleE hide_const (open) Rings.coprime lemma irreducible_uminus [simp]: fixes a::"'a::idom" shows "irreducible (-a) \ irreducible a" using irreducible_mult_unit_left[of "-1::'a"] by auto context comm_monoid_mult begin definition coprime :: "'a \ 'a \ bool" where coprime_def': "coprime p q \ \r. r dvd p \ r dvd q \ r dvd 1" lemma coprimeI: assumes "\r. r dvd p \ r dvd q \ r dvd 1" shows "coprime p q" using assms by (auto simp: coprime_def') lemma coprimeE: assumes "coprime p q" and "(\r. r dvd p \ r dvd q \ r dvd 1) \ thesis" shows thesis using assms by (auto simp: coprime_def') lemma coprime_commute [ac_simps]: "coprime p q \ coprime q p" by (auto simp add: coprime_def') lemma not_coprime_iff_common_factor: "\ coprime p q \ (\r. r dvd p \ r dvd q \ \ r dvd 1)" by (auto simp add: coprime_def') end lemma (in algebraic_semidom) coprime_iff_coprime [simp, code]: "coprime = Rings.coprime" by (simp add: fun_eq_iff coprime_def coprime_def') lemma (in comm_semiring_1) coprime_0 [simp]: "coprime p 0 \ p dvd 1" "coprime 0 p \ p dvd 1" by (auto intro: coprimeI elim: coprimeE dest: dvd_trans) (**** until here ****) (* TODO: move or...? *) lemma dvd_rewrites: "dvd.dvd ((*)) = (dvd)" by (unfold dvd.dvd_def dvd_def, rule) subsection \Interfacing UFD properties\ hide_const (open) Divisibility.irreducible context comm_monoid_mult_isom begin lemma coprime_hom[simp]: "coprime (hom x) y' \ coprime x (Hilbert_Choice.inv hom y')" proof- show ?thesis by (unfold coprime_def', fold ball_UNIV, subst surj[symmetric], simp) qed lemma coprime_inv_hom[simp]: "coprime (Hilbert_Choice.inv hom x') y \ coprime x' (hom y)" proof- interpret inv: comm_monoid_mult_isom "Hilbert_Choice.inv hom".. show ?thesis by simp qed end subsubsection \Original part\ lemma dvd_dvd_imp_smult: fixes p q :: "'a :: idom poly" assumes pq: "p dvd q" and qp: "q dvd p" shows "\c. p = smult c q" proof (cases "p = 0") case True then show ?thesis by auto next case False from qp obtain r where r: "p = q * r" by (elim dvdE, auto) with False qp have r0: "r \ 0" and q0: "q \ 0" by auto with divides_degree[OF pq] divides_degree[OF qp] False have "degree p = degree q" by auto with r degree_mult_eq[OF q0 r0] have "degree r = 0" by auto from degree_0_id[OF this] obtain c where "r = [:c:]" by metis from r[unfolded this] show ?thesis by auto qed lemma dvd_const: assumes pq: "(p::'a::semidom poly) dvd q" and q0: "q \ 0" and degq: "degree q = 0" shows "degree p = 0" proof- from dvdE[OF pq] obtain r where *: "q = p * r". with q0 have "p \ 0" "r \ 0" by auto from degree_mult_eq[OF this] degq * show "degree p = 0" by auto qed context Rings.dvd begin abbreviation ddvd (infix "ddvd" 40) where "x ddvd y \ x dvd y \ y dvd x" lemma ddvd_sym[sym]: "x ddvd y \ y ddvd x" by auto end context comm_monoid_mult begin lemma ddvd_trans[trans]: "x ddvd y \ y ddvd z \ x ddvd z" using dvd_trans by auto lemma ddvd_transp: "transp (ddvd)" by (intro transpI, fact ddvd_trans) end context comm_semiring_1 begin definition mset_factors where "mset_factors F p \ F \ {#} \ (\f. f \# F \ irreducible f) \ p = prod_mset F" lemma mset_factorsI[intro!]: assumes "\f. f \# F \ irreducible f" and "F \ {#}" and "prod_mset F = p" shows "mset_factors F p" unfolding mset_factors_def using assms by auto lemma mset_factorsD: assumes "mset_factors F p" shows "f \# F \ irreducible f" and "F \ {#}" and "prod_mset F = p" using assms[unfolded mset_factors_def] by auto lemma mset_factorsE[elim]: assumes "mset_factors F p" and "(\f. f \# F \ irreducible f) \ F \ {#} \ prod_mset F = p \ thesis" shows thesis using assms[unfolded mset_factors_def] by auto lemma mset_factors_imp_not_is_unit: assumes "mset_factors F p" shows "\ p dvd 1" proof(cases F) case empty with assms show ?thesis by auto next case (add f F) with assms have "\ f dvd 1" "p = f * prod_mset F" by (auto intro!: irreducible_not_unit) then show ?thesis by auto qed definition primitive_poly where "primitive_poly f \ \d. (\i. d dvd coeff f i) \ d dvd 1" end lemma(in semidom) mset_factors_imp_nonzero: assumes "mset_factors F p" shows "p \ 0" proof assume "p = 0" moreover from assms have "prod_mset F = p" by auto ultimately obtain f where "f \# F" "f = 0" by auto with assms show False by auto qed class ufd = idom + assumes mset_factors_exist: "\x. x \ 0 \ \ x dvd 1 \ \F. mset_factors F x" and mset_factors_unique: "\x F G. mset_factors F x \ mset_factors G x \ rel_mset (ddvd) F G" subsubsection \Connecting to HOL/Divisibility\ context comm_semiring_1 begin abbreviation "mk_monoid \ \carrier = UNIV - {0}, mult = (*), one = 1\" lemma carrier_0[simp]: "x \ carrier mk_monoid \ x \ 0" by auto lemmas mk_monoid_simps = carrier_0 monoid.simps abbreviation irred where "irred \ Divisibility.irreducible mk_monoid" abbreviation factor where "factor \ Divisibility.factor mk_monoid" abbreviation factors where "factors \ Divisibility.factors mk_monoid" abbreviation properfactor where "properfactor \ Divisibility.properfactor mk_monoid" lemma factors: "factors fs y \ prod_list fs = y \ Ball (set fs) irred" proof - have "prod_list fs = foldr (*) fs 1" by (induct fs, auto) thus ?thesis unfolding factors_def by auto qed lemma factor: "factor x y \ (\z. z \ 0 \ x * z = y)" unfolding factor_def by auto lemma properfactor_nz: shows "(y :: 'a) \ 0 \ properfactor x y \ x dvd y \ \ y dvd x" by (auto simp: properfactor_def factor_def dvd_def) lemma mem_Units[simp]: "y \ Units mk_monoid \ y dvd 1" unfolding dvd_def Units_def by (auto simp: ac_simps) end context idom begin lemma irred_0[simp]: "irred (0::'a)" by (unfold Divisibility.irreducible_def, auto simp: factor properfactor_def) lemma factor_idom[simp]: "factor (x::'a) y \ (if y = 0 then x = 0 else x dvd y)" by (cases "y = 0"; auto intro: exI[of _ 1] elim: dvdE simp: factor) lemma associated_connect[simp]: "(\\<^bsub>mk_monoid\<^esub>) = (ddvd)" by (intro ext, unfold associated_def, auto) lemma essentially_equal_connect[simp]: "essentially_equal mk_monoid fs gs \ rel_mset (ddvd) (mset fs) (mset gs)" by (auto simp: essentially_equal_def rel_mset_via_perm perm_iff_eq_mset) lemma irred_idom_nz: assumes x0: "(x::'a) \ 0" shows "irred x \ irreducible x" using x0 by (auto simp: irreducible_altdef Divisibility.irreducible_def properfactor_nz) lemma dvd_dvd_imp_unit_mult: assumes xy: "x dvd y" and yx: "y dvd x" shows "\z. z dvd 1 \ y = x * z" proof(cases "x = 0") case True with xy show ?thesis by (auto intro: exI[of _ 1]) next case x0: False from xy obtain z where z: "y = x * z" by (elim dvdE, auto) from yx obtain w where w: "x = y * w" by (elim dvdE, auto) from z w have "x * (z * w) = x" by (auto simp: ac_simps) then have "z * w = 1" using x0 by auto with z show ?thesis by (auto intro: exI[of _ z]) qed lemma irred_inner_nz: assumes x0: "x \ 0" shows "(\b. b dvd x \ \ x dvd b \ b dvd 1) \ (\a b. x = a * b \ a dvd 1 \ b dvd 1)" (is "?l \ ?r") proof (intro iffI allI impI) assume l: ?l fix a b assume xab: "x = a * b" then have ax: "a dvd x" and bx: "b dvd x" by auto { assume a1: "\ a dvd 1" with l ax have xa: "x dvd a" by auto from dvd_dvd_imp_unit_mult[OF ax xa] obtain z where z1: "z dvd 1" and xaz: "x = a * z" by auto from xab x0 have "a \ 0" by auto with xab xaz have "b = z" by auto with z1 have "b dvd 1" by auto } then show "a dvd 1 \ b dvd 1" by auto next assume r: ?r fix b assume bx: "b dvd x" and xb: "\ x dvd b" then obtain a where xab: "x = a * b" by (elim dvdE, auto simp: ac_simps) with r consider "a dvd 1" | "b dvd 1" by auto then show "b dvd 1" proof(cases) case 2 then show ?thesis by auto next case 1 then obtain c where ac1: "a * c = 1" by (elim dvdE, auto) from xab have "x * c = b * (a * c)" by (auto simp: ac_simps) with ac1 have "x * c = b" by auto then have "x dvd b" by auto with xb show ?thesis by auto qed qed lemma irred_idom[simp]: "irred x \ x = 0 \ irreducible x" by (cases "x = 0"; simp add: irred_idom_nz irred_inner_nz irreducible_def) lemma assumes "x \ 0" and "factors fs x" and "f \ set fs" shows "f \ 0" using assms by (auto simp: factors) lemma factors_as_mset_factors: assumes x0: "x \ 0" and x1: "x \ 1" shows "factors fs x \ mset_factors (mset fs) x" using assms by (auto simp: factors prod_mset_prod_list) end context ufd begin interpretation comm_monoid_cancel: comm_monoid_cancel "mk_monoid::'a monoid" apply (unfold_locales) apply simp_all using mult_left_cancel apply (auto simp: ac_simps) done lemma factors_exist: assumes "a \ 0" and "\ a dvd 1" shows "\fs. set fs \ UNIV - {0} \ factors fs a" proof- from mset_factors_exist[OF assms] obtain F where "mset_factors F a" by auto also from ex_mset obtain fs where "F = mset fs" by metis finally have fs: "mset_factors (mset fs) a". then have "factors fs a" using assms by (subst factors_as_mset_factors, auto) moreover have "set fs \ UNIV - {0}" using fs by (auto elim!: mset_factorsE) ultimately show ?thesis by auto qed lemma factors_unique: assumes fs: "factors fs a" and gs: "factors gs a" and a0: "a \ 0" and a1: "\ a dvd 1" shows "rel_mset (ddvd) (mset fs) (mset gs)" proof- from a1 have "a \ 1" by auto with a0 fs gs have "mset_factors (mset fs) a" "mset_factors (mset gs) a" by (unfold factors_as_mset_factors) from mset_factors_unique[OF this] show ?thesis. qed lemma factorial_monoid: "factorial_monoid (mk_monoid :: 'a monoid)" by (unfold_locales; auto simp add: factors_exist factors_unique) end lemma (in idom) factorial_monoid_imp_ufd: assumes "factorial_monoid (mk_monoid :: 'a monoid)" shows "class.ufd ((*) :: 'a \ _) 1 (+) 0 (-) uminus" proof (unfold_locales) interpret factorial_monoid "mk_monoid :: 'a monoid" by (fact assms) { fix x assume x: "x \ 0" "\ x dvd 1" note * = factors_exist[simplified, OF this] with x show "\F. mset_factors F x" by (subst(asm) factors_as_mset_factors, auto) } fix x F G assume FG: "mset_factors F x" "mset_factors G x" with mset_factors_imp_not_is_unit have x1: "\ x dvd 1" by auto from FG(1) have x0: "x \ 0" by (rule mset_factors_imp_nonzero) obtain fs gs where fsgs: "F = mset fs" "G = mset gs" using ex_mset by metis note FG = FG[unfolded this] then have 0: "0 \ set fs" "0 \ set gs" by (auto elim!: mset_factorsE) from x1 have "x \ 1" by auto note FG[folded factors_as_mset_factors[OF x0 this]] from factors_unique[OF this, simplified, OF x0 x1, folded fsgs] 0 show "rel_mset (ddvd) F G" by auto qed subsection \Preservation of Irreducibility\ locale comm_semiring_1_hom = comm_monoid_mult_hom hom + zero_hom hom for hom :: "'a :: comm_semiring_1 \ 'b :: comm_semiring_1" locale irreducibility_hom = comm_semiring_1_hom + assumes irreducible_imp_irreducible_hom: "irreducible a \ irreducible (hom a)" begin lemma hom_mset_factors: assumes F: "mset_factors F p" shows "mset_factors (image_mset hom F) (hom p)" proof (unfold mset_factors_def, intro conjI allI impI) from F show "hom p = prod_mset (image_mset hom F)" "image_mset hom F \ {#}" by (auto simp: hom_distribs) fix f' assume "f' \# image_mset hom F" then obtain f where f: "f \# F" and f'f: "f' = hom f" by auto with F irreducible_imp_irreducible_hom show "irreducible f'" unfolding f'f by auto qed end locale unit_preserving_hom = comm_semiring_1_hom + assumes is_unit_hom_if: "\x. hom x dvd 1 \ x dvd 1" begin lemma is_unit_hom_iff[simp]: "hom x dvd 1 \ x dvd 1" using is_unit_hom_if hom_dvd by force lemma irreducible_hom_imp_irreducible: assumes irr: "irreducible (hom a)" shows "irreducible a" proof (intro irreducibleI) from irr show "a \ 0" by auto from irr show "\ a dvd 1" by (auto dest: irreducible_not_unit) fix b c assume "a = b * c" then have "hom a = hom b * hom c" by (simp add: hom_distribs) with irr have "hom b dvd 1 \ hom c dvd 1" by (auto dest: irreducibleD) then show "b dvd 1 \ c dvd 1" by simp qed end locale factor_preserving_hom = unit_preserving_hom + irreducibility_hom begin lemma irreducible_hom[simp]: "irreducible (hom a) \ irreducible a" using irreducible_hom_imp_irreducible irreducible_imp_irreducible_hom by metis end lemma factor_preserving_hom_comp: assumes f: "factor_preserving_hom f" and g: "factor_preserving_hom g" shows "factor_preserving_hom (f o g)" proof- interpret f: factor_preserving_hom f by (rule f) interpret g: factor_preserving_hom g by (rule g) show ?thesis by (unfold_locales, auto simp: hom_distribs) qed context comm_semiring_isom begin sublocale unit_preserving_hom by (unfold_locales, auto) sublocale factor_preserving_hom proof (standard) fix a :: 'a assume "irreducible a" note a = this[unfolded irreducible_def] show "irreducible (hom a)" proof (rule ccontr) assume "\ irreducible (hom a)" from this[unfolded Factorial_Ring.irreducible_def,simplified] a obtain hb hc where eq: "hom a = hb * hc" and nu: "\ hb dvd 1" "\ hc dvd 1" by auto from bij obtain b where hb: "hb = hom b" by (elim bij_pointE) from bij obtain c where hc: "hc = hom c" by (elim bij_pointE) from eq[unfolded hb hc, folded hom_mult] have "a = b * c" by auto with nu hb hc have "a = b * c" "\ b dvd 1" "\ c dvd 1" by auto with a show False by auto qed qed end subsubsection\Back to divisibility\ lemma(in comm_semiring_1) mset_factors_mult: assumes F: "mset_factors F a" and G: "mset_factors G b" shows "mset_factors (F+G) (a*b)" proof(intro mset_factorsI) fix f assume "f \# F + G" then consider "f \# F" | "f \# G" by auto then show "irreducible f" by(cases, insert F G, auto) qed (insert F G, auto) lemma(in ufd) dvd_imp_subset_factors: assumes ab: "a dvd b" and F: "mset_factors F a" and G: "mset_factors G b" shows "\G'. G' \# G \ rel_mset (ddvd) F G'" proof- from F G have a0: "a \ 0" and b0: "b \ 0" by (simp_all add: mset_factors_imp_nonzero) from ab obtain c where c: "b = a * c" by (elim dvdE, auto) with b0 have c0: "c \ 0" by auto show ?thesis proof(cases "c dvd 1") case True show ?thesis proof(cases F) case empty with F show ?thesis by auto next case (add f F') with F have a: "f * prod_mset F' = a" and F': "\f. f \# F' \ irreducible f" and irrf: "irreducible f" by auto from irrf have f0: "f \ 0" and f1: "\f dvd 1" by (auto dest: irreducible_not_unit) from a c have "(f * c) * prod_mset F' = b" by (auto simp: ac_simps) moreover { have "irreducible (f * c)" using True irrf by (subst irreducible_mult_unit_right) with F' irrf have "\f'. f' \# F' + {#f * c#} \ irreducible f'" by auto } ultimately have "mset_factors (F' + {#f * c#}) b" by (intro mset_factorsI, auto) from mset_factors_unique[OF this G] have F'G: "rel_mset (ddvd) (F' + {#f * c#}) G". from True add have FF': "rel_mset (ddvd) F (F' + {#f * c#})" by (auto simp add: multiset.rel_refl intro!: rel_mset_Plus) have "rel_mset (ddvd) F G" apply(rule transpD[OF multiset.rel_transp[OF transpI] FF' F'G]) using ddvd_trans. then show ?thesis by auto qed next case False from mset_factors_exist[OF c0 this] obtain H where H: "mset_factors H c" by auto from c mset_factors_mult[OF F H] have "mset_factors (F + H) b" by auto note mset_factors_unique[OF this G] from rel_mset_split[OF this] obtain G1 G2 where "G = G1 + G2" "rel_mset (ddvd) F G1" "rel_mset (ddvd) H G2" by auto then show ?thesis by (intro exI[of _ "G1"], auto) qed qed lemma(in idom) irreducible_factor_singleton: assumes a: "irreducible a" shows "mset_factors F a \ F = {#a#}" proof(cases F) case empty with mset_factorsD show ?thesis by auto next case (add f F') show ?thesis proof assume F: "mset_factors F a" from add mset_factorsD[OF F] have *: "a = f * prod_mset F'" by auto then have fa: "f dvd a" by auto from * a have f0: "f \ 0" by auto from add have "f \# F" by auto with F have f: "irreducible f" by auto from add have "F' \# F" by auto then have unitemp: "prod_mset F' dvd 1 \ F' = {#}" proof(induct F') case empty then show ?case by auto next case (add f F') from add have "f \# F" by (simp add: mset_subset_eq_insertD) with F irreducible_not_unit have "\ f dvd 1" by auto then have "\ (prod_mset F' * f) dvd 1" by simp with add show ?case by auto qed show "F = {#a#}" proof(cases "a dvd f") case True then obtain r where "f = a * r" by (elim dvdE, auto) with * have "f = (r * prod_mset F') * f" by (auto simp: ac_simps) with f0 have "r * prod_mset F' = 1" by auto then have "prod_mset F' dvd 1" by (metis dvd_triv_right) with unitemp * add show ?thesis by auto next case False with fa a f show ?thesis by (auto simp: irreducible_altdef) qed qed (insert a, auto) qed lemma(in ufd) irreducible_dvd_imp_factor: assumes ab: "a dvd b" and a: "irreducible a" and G: "mset_factors G b" shows "\g \# G. a ddvd g" proof- from a have "mset_factors {#a#} a" by auto from dvd_imp_subset_factors[OF ab this G] obtain G' where G'G: "G' \# G" and rel: "rel_mset (ddvd) {#a#} G'" by auto with rel_mset_size size_1_singleton_mset size_single obtain g where gG': "G' = {#g#}" by fastforce from rel[unfolded this rel_mset_def] have "a ddvd g" by auto with gG' G'G show ?thesis by auto qed lemma(in idom) prod_mset_remove_units: "prod_mset F ddvd prod_mset {# f \# F. \f dvd 1 #}" proof(induct F) case (add f F) then show ?case by (cases "f = 0", auto) qed auto lemma(in comm_semiring_1) mset_factors_imp_dvd: assumes "mset_factors F x" and "f \# F" shows "f dvd x" using assms by (simp add: dvd_prod_mset mset_factors_def) lemma(in ufd) prime_elem_iff_irreducible[iff]: "prime_elem x \ irreducible x" proof (intro iffI, fact prime_elem_imp_irreducible, rule prime_elemI) assume r: "irreducible x" then show x0: "x \ 0" and x1: "\ x dvd 1" by (auto dest: irreducible_not_unit) from irreducible_factor_singleton[OF r] have *: "mset_factors {#x#} x" by auto fix a b assume "x dvd a * b" then obtain c where abxc: "a * b = x * c" by (elim dvdE, auto) show "x dvd a \ x dvd b" proof(cases "c = 0 \ a = 0 \ b = 0") case True with abxc show ?thesis by auto next case False then have a0: "a \ 0" and b0: "b \ 0" and c0: "c \ 0" by auto from x0 c0 have xc0: "x * c \ 0" by auto from x1 have xc1: "\ x * c dvd 1" by auto show ?thesis proof (cases "a dvd 1 \ b dvd 1") case False then have a1: "\ a dvd 1" and b1: "\ b dvd 1" by auto from mset_factors_exist[OF a0 a1] obtain F where Fa: "mset_factors F a" by auto then have F0: "F \ {#}" by auto from mset_factors_exist[OF b0 b1] obtain G where Gb: "mset_factors G b" by auto then have G0: "G \ {#}" by auto from mset_factors_mult[OF Fa Gb] have FGxc: "mset_factors (F + G) (x * c)" by (simp add: abxc) show ?thesis proof (cases "c dvd 1") case True from r irreducible_mult_unit_right[OF this] have "irreducible (x*c)" by simp note irreducible_factor_singleton[OF this] FGxc with F0 G0 have False by (cases F; cases G; auto) then show ?thesis by auto next case False from mset_factors_exist[OF c0 this] obtain H where "mset_factors H c" by auto with * have xHxc: "mset_factors (add_mset x H) (x * c)" by force note rel = mset_factors_unique[OF this FGxc] obtain hs where "mset hs = H" using ex_mset by auto then have "mset (x#hs) = add_mset x H" by auto from rel_mset_free[OF rel this] obtain jjs where jjsGH: "mset jjs = F + G" and rel: "list_all2 (ddvd) (x # hs) jjs" by auto then obtain j js where jjs: "jjs = j # js" by (cases jjs, auto) with rel have xj: "x ddvd j" by auto from jjs jjsGH have j: "j \ set_mset (F + G)" by (intro union_single_eq_member, auto) from j consider "j \# F" | "j \# G" by auto then show ?thesis proof(cases) case 1 with Fa have "j dvd a" by (auto intro: mset_factors_imp_dvd) with xj dvd_trans have "x dvd a" by auto then show ?thesis by auto next case 2 with Gb have "j dvd b" by (auto intro: mset_factors_imp_dvd) with xj dvd_trans have "x dvd b" by auto then show ?thesis by auto qed qed next case True then consider "a dvd 1" | "b dvd 1" by auto then show ?thesis proof(cases) case 1 then obtain d where ad: "a * d = 1" by (elim dvdE, auto) from abxc have "x * (c * d) = a * b * d" by (auto simp: ac_simps) also have "... = a * d * b" by (auto simp: ac_simps) finally have "x dvd b" by (intro dvdI, auto simp: ad) then show ?thesis by auto next case 2 then obtain d where bd: "b * d = 1" by (elim dvdE, auto) from abxc have "x * (c * d) = a * b * d" by (auto simp: ac_simps) also have "... = (b * d) * a" by (auto simp: ac_simps) finally have "x dvd a" by (intro dvdI, auto simp:bd) then show ?thesis by auto qed qed qed qed subsection\Results for GCDs etc.\ lemma prod_list_remove1: "(x :: 'b :: comm_monoid_mult) \ set xs \ prod_list (remove1 x xs) * x = prod_list xs" by (induct xs, auto simp: ac_simps) (* Isabelle 2015-style and generalized gcd-class without normalization and factors *) class comm_monoid_gcd = gcd + comm_semiring_1 + assumes gcd_dvd1[iff]: "gcd a b dvd a" and gcd_dvd2[iff]: "gcd a b dvd b" and gcd_greatest: "c dvd a \ c dvd b \ c dvd gcd a b" begin lemma gcd_0_0[simp]: "gcd 0 0 = 0" using gcd_greatest[OF dvd_0_right dvd_0_right, of 0] by auto lemma gcd_zero_iff[simp]: "gcd a b = 0 \ a = 0 \ b = 0" proof assume "gcd a b = 0" from gcd_dvd1[of a b, unfolded this] gcd_dvd2[of a b, unfolded this] show "a = 0 \ b = 0" by auto qed auto lemma gcd_zero_iff'[simp]: "0 = gcd a b \ a = 0 \ b = 0" using gcd_zero_iff by metis lemma dvd_gcd_0_iff[simp]: shows "x dvd gcd 0 a \ x dvd a" (is ?g1) and "x dvd gcd a 0 \ x dvd a" (is ?g2) proof- have "a dvd gcd a 0" "a dvd gcd 0 a" by (auto intro: gcd_greatest) with dvd_refl show ?g1 ?g2 by (auto dest: dvd_trans) qed lemma gcd_dvd_1[simp]: "gcd a b dvd 1 \ coprime a b" using dvd_trans[OF gcd_greatest[of _ a b], of _ 1] by (cases "a = 0 \ b = 0") (auto intro!: coprimeI elim: coprimeE) lemma dvd_imp_gcd_dvd_gcd: "b dvd c \ gcd a b dvd gcd a c" by (meson gcd_dvd1 gcd_dvd2 gcd_greatest dvd_trans) definition listgcd :: "'a list \ 'a" where "listgcd xs = foldr gcd xs 0" lemma listgcd_simps[simp]: "listgcd [] = 0" "listgcd (x # xs) = gcd x (listgcd xs)" by (auto simp: listgcd_def) lemma listgcd: "x \ set xs \ listgcd xs dvd x" proof (induct xs) case (Cons y ys) show ?case proof (cases "x = y") case False with Cons have dvd: "listgcd ys dvd x" by auto thus ?thesis unfolding listgcd_simps using dvd_trans by blast next case True thus ?thesis unfolding listgcd_simps using dvd_trans by blast qed qed simp lemma listgcd_greatest: "(\ x. x \ set xs \ y dvd x) \ y dvd listgcd xs" by (induct xs arbitrary:y, auto intro: gcd_greatest) end context Rings.dvd begin definition "is_gcd x a b \ x dvd a \ x dvd b \ (\y. y dvd a \ y dvd b \ y dvd x)" definition "some_gcd a b \ SOME x. is_gcd x a b" lemma is_gcdI[intro!]: assumes "x dvd a" "x dvd b" "\y. y dvd a \ y dvd b \ y dvd x" shows "is_gcd x a b" by (insert assms, auto simp: is_gcd_def) lemma is_gcdE[elim!]: assumes "is_gcd x a b" and "x dvd a \ x dvd b \ (\y. y dvd a \ y dvd b \ y dvd x) \ thesis" shows thesis by (insert assms, auto simp: is_gcd_def) lemma is_gcd_some_gcdI: assumes "\x. is_gcd x a b" shows "is_gcd (some_gcd a b) a b" by (unfold some_gcd_def, rule someI_ex[OF assms]) end context comm_semiring_1 begin lemma some_gcd_0[intro!]: "is_gcd (some_gcd a 0) a 0" "is_gcd (some_gcd 0 b) 0 b" by (auto intro!: is_gcd_some_gcdI intro: exI[of _ a] exI[of _ b]) lemma some_gcd_0_dvd[intro!]: "some_gcd a 0 dvd a" "some_gcd 0 b dvd b" using some_gcd_0 by auto lemma dvd_some_gcd_0[intro!]: "a dvd some_gcd a 0" "b dvd some_gcd 0 b" using some_gcd_0[of a] some_gcd_0[of b] by auto end context idom begin lemma is_gcd_connect: assumes "a \ 0" "b \ 0" shows "isgcd mk_monoid x a b \ is_gcd x a b" using assms by (force simp: isgcd_def) lemma some_gcd_connect: assumes "a \ 0" and "b \ 0" shows "somegcd mk_monoid a b = some_gcd a b" using assms by (auto intro!: arg_cong[of _ _ Eps] simp: is_gcd_connect some_gcd_def somegcd_def) end context comm_monoid_gcd begin lemma is_gcd_gcd: "is_gcd (gcd a b) a b" using gcd_greatest by auto lemma is_gcd_some_gcd: "is_gcd (some_gcd a b) a b" by (insert is_gcd_gcd, auto intro!: is_gcd_some_gcdI) lemma gcd_dvd_some_gcd: "gcd a b dvd some_gcd a b" using is_gcd_some_gcd by auto lemma some_gcd_dvd_gcd: "some_gcd a b dvd gcd a b" using is_gcd_some_gcd by (auto intro: gcd_greatest) lemma some_gcd_ddvd_gcd: "some_gcd a b ddvd gcd a b" by (auto intro: gcd_dvd_some_gcd some_gcd_dvd_gcd) lemma some_gcd_dvd: "some_gcd a b dvd d \ gcd a b dvd d" "d dvd some_gcd a b \ d dvd gcd a b" using some_gcd_ddvd_gcd[of a b] by (auto dest:dvd_trans) end class idom_gcd = comm_monoid_gcd + idom begin interpretation raw: comm_monoid_cancel "mk_monoid :: 'a monoid" by (unfold_locales, auto intro: mult_commute mult_assoc) interpretation raw: gcd_condition_monoid "mk_monoid :: 'a monoid" by (unfold_locales, auto simp: is_gcd_connect intro!: exI[of _ "gcd _ _"] dest: gcd_greatest) lemma gcd_mult_ddvd: "d * gcd a b ddvd gcd (d * a) (d * b)" proof (cases "d = 0") case True then show ?thesis by auto next case d0: False show ?thesis proof (cases "a = 0 \ b = 0") case False note some_gcd_ddvd_gcd[of a b] with d0 have "d * gcd a b ddvd d * some_gcd a b" by auto also have "d * some_gcd a b ddvd some_gcd (d * a) (d * b)" using False d0 raw.gcd_mult by (simp add: some_gcd_connect) also note some_gcd_ddvd_gcd finally show ?thesis. next case True with d0 show ?thesis apply (elim disjE) apply (rule ddvd_trans[of _ "d * b"]; force) apply (rule ddvd_trans[of _ "d * a"]; force) done qed qed lemma gcd_greatest_mult: assumes cad: "c dvd a * d" and cbd: "c dvd b * d" shows "c dvd gcd a b * d" proof- from gcd_greatest[OF assms] have c: "c dvd gcd (d * a) (d * b)" by (auto simp: ac_simps) note gcd_mult_ddvd[of d a b] then have "gcd (d * a) (d * b) dvd gcd a b * d" by (auto simp: ac_simps) from dvd_trans[OF c this] show ?thesis . qed lemma listgcd_greatest_mult: "(\ x :: 'a. x \ set xs \ y dvd x * z) \ y dvd listgcd xs * z" proof (induct xs) case (Cons x xs) from Cons have "y dvd x * z" "y dvd listgcd xs * z" by auto thus ?case unfolding listgcd_simps by (rule gcd_greatest_mult) qed (simp) lemma dvd_factor_mult_gcd: assumes dvd: "k dvd p * q" "k dvd p * r" and q0: "q \ 0" and r0: "r \ 0" shows "k dvd p * gcd q r" proof - from dvd gcd_greatest[of k "p * q" "p * r"] have "k dvd gcd (p * q) (p * r)" by simp also from gcd_mult_ddvd[of p q r] have "... dvd (p * gcd q r)" by auto finally show ?thesis . qed lemma coprime_mult_cross_dvd: assumes coprime: "coprime p q" and eq: "p' * p = q' * q" shows "p dvd q'" (is ?g1) and "q dvd p'" (is ?g2) proof (atomize(full), cases "p = 0 \ q = 0") case True then show "?g1 \ ?g2" proof assume p0: "p = 0" with coprime have "q dvd 1" by auto with eq p0 show ?thesis by auto next assume q0: "q = 0" with coprime have "p dvd 1" by auto with eq q0 show ?thesis by auto qed next case False { fix p q r p' q' :: 'a assume cop: "coprime p q" and eq: "p' * p = q' * q" and p: "p \ 0" and q: "q \ 0" and r: "r dvd p" "r dvd q" let ?gcd = "gcd q p" from eq have "p' * p dvd q' * q" by auto hence d1: "p dvd q' * q" by (rule dvd_mult_right) have d2: "p dvd q' * p" by auto from dvd_factor_mult_gcd[OF d1 d2 q p] have 1: "p dvd q' * ?gcd" . from q p have 2: "?gcd dvd q" by auto from q p have 3: "?gcd dvd p" by auto from cop[unfolded coprime_def', rule_format, OF 3 2] have "?gcd dvd 1" . from 1 dvd_mult_unit_iff[OF this] have "p dvd q'" by auto } note main = this from main[OF coprime eq,of 1] False coprime coprime_commute main[OF _ eq[symmetric], of 1] show "?g1 \ ?g2" by auto qed end subclass (in ring_gcd) idom_gcd by (unfold_locales, auto) lemma coprime_rewrites: "comm_monoid_mult.coprime ((*)) 1 = coprime" apply (intro ext) apply (subst comm_monoid_mult.coprime_def') apply (unfold_locales) apply (unfold dvd_rewrites) apply (fold coprime_def') .. (* TODO: incorporate into the default class hierarchy *) locale gcd_condition = fixes ty :: "'a :: idom itself" assumes gcd_exists: "\a b :: 'a. \x. is_gcd x a b" begin sublocale idom_gcd "(*)" "1 :: 'a" "(+)" 0 "(-)" uminus some_gcd rewrites "dvd.dvd ((*)) = (dvd)" and "comm_monoid_mult.coprime ((*) ) 1 = Unique_Factorization.coprime" proof- have "is_gcd (some_gcd a b) a b" for a b :: 'a by (intro is_gcd_some_gcdI gcd_exists) from this[unfolded is_gcd_def] show "class.idom_gcd (*) (1 :: 'a) (+) 0 (-) uminus some_gcd" by (unfold_locales, auto simp: dvd_rewrites) qed (simp_all add: dvd_rewrites coprime_rewrites) end instance semiring_gcd \ comm_monoid_gcd by (intro_classes, auto) lemma listgcd_connect: "listgcd = gcd_list" proof (intro ext) fix xs :: "'a list" show "listgcd xs = gcd_list xs" by(induct xs, auto) qed interpretation some_gcd: gcd_condition "TYPE('a::ufd)" proof(unfold_locales, intro exI) interpret factorial_monoid "mk_monoid :: 'a monoid" by (fact factorial_monoid) note d = dvd.dvd_def some_gcd_def carrier_0 fix a b :: 'a show "is_gcd (some_gcd a b) a b" proof (cases "a = 0 \ b = 0") case True thus ?thesis using some_gcd_0 by auto next case False with gcdof_exists[of a b] show ?thesis by (auto intro!: is_gcd_some_gcdI simp add: is_gcd_connect some_gcd_connect) qed qed lemma some_gcd_listgcd_dvd_listgcd: "some_gcd.listgcd xs dvd listgcd xs" by (induct xs, auto simp:some_gcd_dvd intro:dvd_imp_gcd_dvd_gcd) lemma listgcd_dvd_some_gcd_listgcd: "listgcd xs dvd some_gcd.listgcd xs" by (induct xs, auto simp:some_gcd_dvd intro:dvd_imp_gcd_dvd_gcd) context factorial_ring_gcd begin text \Do not declare the following as subclass, to avoid conflict in \field \ gcd_condition\ vs. \factorial_ring_gcd \ gcd_condition\. \ sublocale as_ufd: ufd proof(unfold_locales, goal_cases) case (1 x) from prime_factorization_exists[OF \x \ 0\] obtain F where f: "\f. f \# F \ prime_elem f" and Fx: "normalize (prod_mset F) = normalize x" by auto from associatedE2[OF Fx] obtain u where u: "is_unit u" "x = u * prod_mset F" by blast from \\ is_unit x\ Fx have "F \ {#}" by auto then obtain g G where F: "F = add_mset g G" by (cases F, auto) then have "g \# F" by auto with f[OF this]prime_elem_iff_irreducible irreducible_mult_unit_left[OF unit_factor_is_unit[OF \x \ 0\]] have g: "irreducible (u * g)" using u(1) by (subst irreducible_mult_unit_left) simp_all show ?case proof (intro exI conjI mset_factorsI) show "prod_mset (add_mset (u * g) G) = x" using \x \ 0\ by (simp add: F ac_simps u) fix f assume "f \# add_mset (u * g) G" with f[unfolded F] g prime_elem_iff_irreducible show "irreducible f" by auto qed auto next case (2 x F G) note transpD[OF multiset.rel_transp[OF ddvd_transp],trans] obtain fs where F: "F = mset fs" by (metis ex_mset) have "list_all2 (ddvd) fs (map normalize fs)" by (intro list_all2_all_nthI, auto) then have FH: "rel_mset (ddvd) F (image_mset normalize F)" by (unfold rel_mset_def F, force) also have FG: "image_mset normalize F = image_mset normalize G" proof (intro prime_factorization_unique'') from 2 have xF: "x = prod_mset F" and xG: "x = prod_mset G" by auto from xF have "normalize x = normalize (prod_mset (image_mset normalize F))" by (simp add: normalize_prod_mset_normalize) with xG have nFG: "\ = normalize (prod_mset (image_mset normalize G))" by (simp_all add: normalize_prod_mset_normalize) then show "normalize (\i\#image_mset normalize F. i) = normalize (\i\#image_mset normalize G. i)" by auto next from 2 prime_elem_iff_irreducible have "f \# F \ prime_elem f" "g \# G \ prime_elem g" for f g by (auto intro: prime_elemI) then show " Multiset.Ball (image_mset normalize F) prime" "Multiset.Ball (image_mset normalize G) prime" by auto qed also obtain gs where G: "G = mset gs" by (metis ex_mset) have "list_all2 ((ddvd)\\) gs (map normalize gs)" by (intro list_all2_all_nthI, auto) then have "rel_mset (ddvd) (image_mset normalize G) G" by (subst multiset.rel_flip[symmetric], unfold rel_mset_def G, force) finally show ?case. qed end instance int :: ufd by (intro class.ufd.of_class.intro as_ufd.ufd_axioms) instance int :: idom_gcd by (intro_classes, auto) instance field \ ufd by (intro_classes, auto simp: dvd_field_iff) end diff --git a/thys/PAC_Checker/Duplicate_Free_Multiset.thy b/thys/Nested_Multisets_Ordinals/Duplicate_Free_Multiset.thy rename from thys/PAC_Checker/Duplicate_Free_Multiset.thy rename to thys/Nested_Multisets_Ordinals/Duplicate_Free_Multiset.thy --- a/thys/PAC_Checker/Duplicate_Free_Multiset.thy +++ b/thys/Nested_Multisets_Ordinals/Duplicate_Free_Multiset.thy @@ -1,157 +1,238 @@ (* - File: Duplicate_Free_Multiset.thy - Author: Mathias Fleury, Daniela Kaufmann, JKU - Maintainer: Mathias Fleury, JKU + File: Duplicate_Free_Multiset.thy + Authors and contributors: Mathias Fleury, Daniela Kaufmann, JKU; + Jose DivasĂ³n, Sebastiaan Joosten, RenĂ© Thiemann, Akihisa Yamada *) + theory Duplicate_Free_Multiset -imports Nested_Multisets_Ordinals.Multiset_More +imports Multiset_More begin - section \Duplicate Free Multisets\ -(* TODO Move Multiset_More *) -lemma remove_diff_multiset[simp]: \x13 \# A \ A - add_mset x13 B = A - B\ - by (metis diff_intersect_left_idem inter_add_right1) - - text \Duplicate free multisets are isomorphic to finite sets, but it can be useful to reason about duplication to speak about intermediate execution steps in the refinements. \ + +definition distinct_mset :: "'a multiset \ bool" where + "distinct_mset S \ (\a. a \# S \ count S a = 1)" + +lemma distinct_mset_count_less_1: "distinct_mset S \ (\a. count S a \ 1)" + using eq_iff nat_le_linear unfolding distinct_mset_def by fastforce + +lemma distinct_mset_empty[simp]: "distinct_mset {#}" + unfolding distinct_mset_def by auto + +lemma distinct_mset_singleton: "distinct_mset {#a#}" + unfolding distinct_mset_def by auto + +lemma distinct_mset_union: + assumes dist: "distinct_mset (A + B)" + shows "distinct_mset A" + unfolding distinct_mset_count_less_1 +proof (rule allI) + fix a + have \count A a \ count (A + B) a\ by auto + moreover have \count (A + B) a \ 1\ + using dist unfolding distinct_mset_count_less_1 by auto + ultimately show \count A a \ 1\ + by simp +qed + +lemma distinct_mset_minus[simp]: "distinct_mset A \ distinct_mset (A - B)" + by (metis diff_subset_eq_self mset_subset_eq_exists_conv distinct_mset_union) + +lemma distinct_mset_rempdups_union_mset: + assumes "distinct_mset A" and "distinct_mset B" + shows "A \# B = remdups_mset (A + B)" + using assms nat_le_linear unfolding remdups_mset_def + by (force simp add: multiset_eq_iff max_def count_mset_set_if distinct_mset_def not_in_iff) + +lemma distinct_mset_add_mset[simp]: "distinct_mset (add_mset a L) \ a \# L \ distinct_mset L" + unfolding distinct_mset_def + apply (rule iffI) + apply (auto split: if_split_asm; fail)[] + by (auto simp: not_in_iff; fail) + +lemma distinct_mset_size_eq_card: "distinct_mset C \ size C = card (set_mset C)" + by (induction C) auto + +lemma distinct_mset_add: + "distinct_mset (L + L') \ distinct_mset L \ distinct_mset L' \ L \# L' = {#}" + by (induction L arbitrary: L') auto + +lemma distinct_mset_set_mset_ident[simp]: "distinct_mset M \ mset_set (set_mset M) = M" + by (induction M) auto + +lemma distinct_finite_set_mset_subseteq_iff[iff]: + assumes "distinct_mset M" "finite N" + shows "set_mset M \ N \ M \# mset_set N" + by (metis assms distinct_mset_set_mset_ident finite_set_mset msubset_mset_set_iff) + +lemma distinct_mem_diff_mset: + assumes dist: "distinct_mset M" and mem: "x \ set_mset (M - N)" + shows "x \ set_mset N" +proof - + have "count M x = 1" + using dist mem by (meson distinct_mset_def in_diffD) + then show ?thesis + using mem by (metis count_greater_eq_one_iff in_diff_count not_less) +qed + +lemma distinct_set_mset_eq: + assumes "distinct_mset M" "distinct_mset N" "set_mset M = set_mset N" + shows "M = N" + using assms distinct_mset_set_mset_ident by fastforce + +lemma distinct_mset_union_mset[simp]: + \distinct_mset (D \# C) \ distinct_mset D \ distinct_mset C\ + unfolding distinct_mset_count_less_1 by force + +lemma distinct_mset_inter_mset: + "distinct_mset C \ distinct_mset (C \# D)" + "distinct_mset D \ distinct_mset (C \# D)" + by (auto simp add: distinct_mset_def min_def count_eq_zero_iff elim!: le_SucE) + +lemma distinct_mset_remove1_All: "distinct_mset C \ remove1_mset L C = removeAll_mset L C" + by (auto simp: multiset_eq_iff distinct_mset_count_less_1) + +lemma distinct_mset_size_2: "distinct_mset {#a, b#} \ a \ b" + by auto + +lemma distinct_mset_filter: "distinct_mset M \ distinct_mset {# L \# M. P L#}" + by (simp add: distinct_mset_def) + +lemma distinct_mset_mset_distinct[simp]: \distinct_mset (mset xs) = distinct xs\ + by (induction xs) auto + +lemma distinct_image_mset_inj: + \inj_on f (set_mset M) \ distinct_mset (image_mset f M) \ distinct_mset M\ + by (induction M) (auto simp: inj_on_def) + lemma distinct_mset_remdups_mset_id: \distinct_mset C \ remdups_mset C = C\ by (induction C) auto -lemma notin_add_mset_remdups_mset: - \a \# A \ add_mset a (remdups_mset A) = remdups_mset (add_mset a A)\ - by auto - lemma distinct_mset_image_mset: \distinct_mset (image_mset f (mset xs)) \ distinct (map f xs)\ apply (subst mset_map[symmetric]) apply (subst distinct_mset_mset_distinct) .. lemma distinct_mset_mono: \D' \# D \ distinct_mset D \ distinct_mset D'\ by (metis distinct_mset_union subset_mset.le_iff_add) lemma distinct_mset_mono_strict: \D' \# D \ distinct_mset D \ distinct_mset D'\ using distinct_mset_mono by auto lemma distinct_set_mset_eq_iff: assumes \distinct_mset M\ \distinct_mset N\ shows \set_mset M = set_mset N \ M = N\ using assms distinct_mset_set_mset_ident by fastforce lemma distinct_mset_union2: \distinct_mset (A + B) \ distinct_mset B\ using distinct_mset_union[of B A] by (auto simp: ac_simps) lemma distinct_mset_mset_set: \distinct_mset (mset_set A)\ unfolding distinct_mset_def count_mset_set_if by (auto simp: not_in_iff) lemma distinct_mset_inter_remdups_mset: assumes dist: \distinct_mset A\ shows \A \# remdups_mset B = A \# B\ proof - have [simp]: \A' \# remove1_mset a (remdups_mset Aa) = A' \# Aa\ if \A' \# remdups_mset Aa = A' \# Aa\ and \a \# A'\ and \a \# Aa\ for A' Aa :: \'a multiset\ and a by (metis insert_DiffM inter_add_right1 set_mset_remdups_mset that) show ?thesis using dist apply (induction A) subgoal by auto subgoal for a A' by (cases \a \# B\) (use multi_member_split[of a \B\] multi_member_split[of a \A\] in \auto simp: mset_set.insert_remove\) done qed -lemma finite_mset_set_inter: - \finite A \ finite B \ mset_set (A \ B) = mset_set A \# mset_set B\ - apply (induction A rule: finite_induct) - subgoal by auto - subgoal for a A - by (cases \a \ B\; cases \a \# mset_set B\) - (use multi_member_split[of a \mset_set B\] in - \auto simp: mset_set.insert_remove\) - done - -lemma removeAll_notin: \a \# A \ removeAll_mset a A = A\ - using count_inI by force - -lemma same_mset_distinct_iff: - \mset M = mset M' \ distinct M \ distinct M'\ - by (auto simp: distinct_mset_mset_distinct[symmetric] simp del: distinct_mset_mset_distinct) - +abbreviation (input) is_mset_set :: \'a multiset \ bool\ + where \is_mset_set \ distinct_mset\ -subsection \More Lists\ -lemma in_set_conv_iff: - \x \ set (take n xs) \ (\i < n. i < length xs \ xs ! i = x)\ - apply (induction n) - subgoal by auto - subgoal for n - apply (cases \Suc n < length xs\) - subgoal by (auto simp: take_Suc_conv_app_nth less_Suc_eq dest: in_set_takeD) - subgoal - apply (cases \n < length xs\) - subgoal - apply (auto simp: in_set_conv_nth) - by (rename_tac i, rule_tac x=i in exI; auto; fail)+ - subgoal - apply (auto simp: take_Suc_conv_app_nth dest: in_set_takeD) - by (rename_tac i, rule_tac x=i in exI; auto; fail)+ - done - done - done - -lemma in_set_take_conv_nth: - \x \ set (take n xs) \ (\m - by (metis in_set_conv_nth length_take min.commute min.strict_boundedE nth_take) +lemma is_mset_set_def: + \is_mset_set X \ (\x \# X. count X x = 1)\ + by (auto simp add: distinct_mset_def) -lemma in_set_remove1D: - \a \ set (remove1 x xs) \ a \ set xs\ - by (meson notin_set_remove1) - +lemma is_mset_setD[dest]: "is_mset_set X \ x \# X \ count X x = 1" + unfolding is_mset_set_def by auto -subsection \Generic Multiset\ +lemma is_mset_setI[intro]: + assumes "\x. x \# X \ count X x = 1" + shows "is_mset_set X" + using assms unfolding is_mset_set_def by auto -lemma mset_drop_upto: \mset (drop a N) = {#N!i. i \# mset_set {a.. -proof (induction N arbitrary: a) - case Nil - then show ?case by simp +lemma is_mset_set[simp]: "is_mset_set (mset_set X)" + by (fact distinct_mset_mset_set) + +lemma is_mset_set_add[simp]: + "is_mset_set (X + {#x#}) \ is_mset_set X \ x \# X" (is "?L \ ?R") +proof(intro iffI conjI) + assume L: ?L + with count_eq_zero_iff count_single show "is_mset_set X" + unfolding is_mset_set_def + by (metis (no_types, hide_lams) add_mset_add_single count_add_mset nat.inject set_mset_add_mset_insert union_single_eq_member) + show "x \# X" + proof + assume "x \# X" + then have "count (X + {#x#}) x > 1" by auto + with L show False by (auto simp: is_mset_set_def) + qed next - case (Cons c N) - have upt: \{0.. - by auto - then have H: \mset_set {0.. - unfolding upt by auto - have mset_case_Suc: \{#case x of 0 \ c | Suc x \ N ! x . x \# mset_set {Suc a..# mset_set {Suc a.. for a b - by (rule image_mset_cong) (auto split: nat.splits) - have Suc_Suc: \{Suc a.. for a b - by auto - then have mset_set_Suc_Suc: \mset_set {Suc a..# mset_set {a.. for a b - unfolding Suc_Suc by (subst image_mset_mset_set[symmetric]) auto - have *: \{#N ! (x-Suc 0) . x \# mset_set {Suc a..# mset_set {a.. - for a b - by (auto simp add: mset_set_Suc_Suc) - show ?case - apply (cases a) - using Cons[of 0] Cons by (auto simp: nth_Cons drop_Cons H mset_case_Suc *) + assume R: ?R show ?L + proof + fix x' assume x': "x' \# X + {#x#}" + show "count (X + {#x#}) x' = 1" + proof(cases "x' \# X") + case True with R have "count X x' = 1" by auto + moreover from True R have "count {#x#} x' = 0" by auto + ultimately show ?thesis by auto + next + case False then have "count X x' = 0" by (simp add: not_in_iff) + with R x' show ?thesis by auto + qed + qed qed - -subsection \Other\ +lemma mset_set_id: + assumes "is_mset_set X" + shows "mset_set (set_mset X) = X" + using assms by (fact distinct_mset_set_mset_ident) -text \I believe this should be added to the simplifier by default...\ -lemma Collect_eq_comp': \ {(x, y). P x y} O {(c, a). c = f a} = {(x, a). P x (f a)}\ - by auto +lemma is_mset_set_image: + assumes "inj_on f (set_mset X)" and "is_mset_set X" + shows "is_mset_set (image_mset f X)" +proof (cases X) + case empty then show ?thesis by auto +next + case (add x X) + define X' where "X' \ add_mset x X" + with assms add have inj:"inj_on f (set_mset X')" + and X': "is_mset_set X'" by auto + show ?thesis + proof(unfold add, intro is_mset_setI, fold X'_def) + fix y assume "y \# image_mset f X'" + then have "y \ f ` set_mset X'" by auto + with inj have "\!x'. x' \# X' \ y = f x'" by (meson imageE inj_onD) + then obtain x' where x': "{x'. x' \# X' \ y = f x'} = {x'}" by auto + then have "count (image_mset f X') y = count X' x'" + by (simp add: count_image_mset') + also from X' x' have "... = 1" by auto + finally show "count (image_mset f X') y = 1". + qed +qed end diff --git a/thys/Nested_Multisets_Ordinals/Multiset_More.thy b/thys/Nested_Multisets_Ordinals/Multiset_More.thy --- a/thys/Nested_Multisets_Ordinals/Multiset_More.thy +++ b/thys/Nested_Multisets_Ordinals/Multiset_More.thy @@ -1,1025 +1,979 @@ (* Title: More about Multisets Author: Mathias Fleury , 2015 Author: Jasmin Blanchette , 2014, 2015 Author: Anders Schlichtkrull , 2017 Author: Dmitriy Traytel , 2014 Maintainer: Mathias Fleury *) section \More about Multisets\ theory Multiset_More imports "HOL-Library.Multiset_Order" "HOL-Library.Sublist" begin text \ Isabelle's theory of finite multisets is not as developed as other areas, such as lists and sets. The present theory introduces some missing concepts and lemmas. Some of it is expected to move to Isabelle's library. \ subsection \Basic Setup\ declare diff_single_trivial [simp] in_image_mset [iff] image_mset.compositionality [simp] (*To have the same rules as the set counter-part*) mset_subset_eqD[dest, intro?] (*@{thm subsetD}*) Multiset.in_multiset_in_set[simp] inter_add_left1[simp] inter_add_left2[simp] inter_add_right1[simp] inter_add_right2[simp] sum_mset_sum_list[simp] subsection \Lemmas about Intersection, Union and Pointwise Inclusion\ lemma subset_mset_imp_subset_add_mset: "A \# B \ A \# add_mset x B" by (auto simp add: subseteq_mset_def le_SucI) lemma subset_add_mset_notin_subset_mset: \A \# add_mset b B \ b \# A \ A \# B\ by (simp add: subset_mset.le_iff_sup) lemma subset_msetE: "\A \# B; \A \# B; \ B \# A\ \ R\ \ R" by (simp add: subset_mset.less_le_not_le) lemma Diff_triv_mset: "M \# N = {#} \ M - N = M" by (metis diff_intersect_left_idem diff_zero) lemma diff_intersect_sym_diff: "(A - B) \# (B - A) = {#}" by (rule multiset_eqI) simp declare subset_msetE [elim!] lemma subseq_mset_subseteq_mset: "subseq xs ys \ mset xs \# mset ys" proof (induct xs arbitrary: ys) case (Cons x xs) note Outer_Cons = this then show ?case proof (induct ys) case (Cons y ys) have "subseq xs ys" by (metis Cons.prems(2) subseq_Cons' subseq_Cons2_iff) then show ?case using Cons by (metis mset.simps(2) mset_subset_eq_add_mset_cancel subseq_Cons2_iff subset_mset_imp_subset_add_mset) qed simp qed simp +lemma finite_mset_set_inter: + \finite A \ finite B \ mset_set (A \ B) = mset_set A \# mset_set B\ + apply (induction A rule: finite_induct) + subgoal by auto + subgoal for a A + by (cases \a \ B\; cases \a \# mset_set B\) + (use multi_member_split[of a \mset_set B\] in + \auto simp: mset_set.insert_remove\) + done + subsection \Lemmas about Filter and Image\ lemma count_image_mset_ge_count: "count (image_mset f A) (f b) \ count A b" by (induction A) auto lemma count_image_mset_inj: assumes \inj f\ shows \count (image_mset f M) (f x) = count M x\ by (induct M) (use assms in \auto simp: inj_on_def\) lemma count_image_mset_le_count_inj_on: "inj_on f (set_mset M) \ count (image_mset f M) y \ count M (inv_into (set_mset M) f y)" proof (induct M) case (add x M) note ih = this(1) and inj_xM = this(2) have inj_M: "inj_on f (set_mset M)" using inj_xM by simp show ?case proof (cases "x \# M") case x_in_M: True show ?thesis proof (cases "y = f x") case y_eq_fx: True show ?thesis using x_in_M ih[OF inj_M] unfolding y_eq_fx by (simp add: inj_M insert_absorb) next case y_ne_fx: False show ?thesis using x_in_M ih[OF inj_M] y_ne_fx insert_absorb by fastforce qed next case x_ni_M: False show ?thesis proof (cases "y = f x") case y_eq_fx: True have "f x \# image_mset f M" using x_ni_M inj_xM by force thus ?thesis unfolding y_eq_fx by (metis (no_types) inj_xM count_add_mset count_greater_eq_Suc_zero_iff count_inI image_mset_add_mset inv_into_f_f union_single_eq_member) next case y_ne_fx: False show ?thesis proof (rule ccontr) assume neg_conj: "\ count (image_mset f (add_mset x M)) y \ count (add_mset x M) (inv_into (set_mset (add_mset x M)) f y)" have cnt_y: "count (add_mset (f x) (image_mset f M)) y = count (image_mset f M) y" using y_ne_fx by simp have "inv_into (set_mset M) f y \# add_mset x M \ inv_into (set_mset (add_mset x M)) f (f (inv_into (set_mset M) f y)) = inv_into (set_mset M) f y" by (meson inj_xM inv_into_f_f) hence "0 < count (image_mset f (add_mset x M)) y \ count M (inv_into (set_mset M) f y) = 0 \ x = inv_into (set_mset M) f y" using neg_conj cnt_y ih[OF inj_M] by (metis (no_types) count_add_mset count_greater_zero_iff count_inI f_inv_into_f image_mset_add_mset set_image_mset) thus False using neg_conj cnt_y x_ni_M ih[OF inj_M] by (metis (no_types) count_greater_zero_iff count_inI eq_iff image_mset_add_mset less_imp_le) qed qed qed qed simp lemma mset_filter_compl: "mset (filter p xs) + mset (filter (Not \ p) xs) = mset xs" by (induction xs) (auto simp: ac_simps) text \Near duplicate of @{thm [source] filter_eq_replicate_mset}: @{thm filter_eq_replicate_mset}.\ lemma filter_mset_eq: "filter_mset ((=) L) A = replicate_mset (count A L) L" by (auto simp: multiset_eq_iff) lemma filter_mset_cong[fundef_cong]: assumes "M = M'" "\a. a \# M \ P a = Q a" shows "filter_mset P M = filter_mset Q M" proof - have "M - filter_mset Q M = filter_mset (\a. \Q a) M" by (metis multiset_partition add_diff_cancel_left') then show ?thesis by (auto simp: filter_mset_eq_conv assms) qed lemma image_mset_filter_swap: "image_mset f {# x \# M. P (f x)#} = {# x \# image_mset f M. P x#}" by (induction M) auto lemma image_mset_cong2: "(\x. x \# M \ f x = g x) \ M = N \ image_mset f M = image_mset g N" by (hypsubst, rule image_mset_cong) lemma filter_mset_empty_conv: \(filter_mset P M = {#}) = (\L\#M. \ P L)\ by (induction M) auto lemma multiset_filter_mono2: \filter_mset P A \# filter_mset Q A \ (\a\#A. P a \ Q a)\ by (induction A) (auto intro: subset_mset.trans) lemma image_filter_cong: assumes \\C. C \# M \ P C \ f C = g C\ shows \{#f C. C \# {#C \# M. P C#}#} = {#g C | C\# M. P C#}\ using assms by (induction M) auto lemma image_mset_filter_swap2: \{#C \# {#P x. x \# D#}. Q C #} = {#P x. x \# {#C| C \# D. Q (P C)#}#}\ by (simp add: image_mset_filter_swap) declare image_mset_cong2 [cong] lemma filter_mset_empty_if_finite_and_filter_set_empty: assumes "{x \ X. P x} = {}" and "finite X" shows "{#x \# mset_set X. P x#} = {#}" proof - have empty_empty: "\Y. set_mset Y = {} \ Y = {#}" by auto from assms have "set_mset {#x \# mset_set X. P x#} = {}" by auto then show ?thesis by (rule empty_empty) qed subsection \Lemmas about Sum\ lemma sum_image_mset_sum_map[simp]: "sum_mset (image_mset f (mset xs)) = sum_list (map f xs)" by (metis mset_map sum_mset_sum_list) lemma sum_image_mset_mono: fixes f :: "'a \ 'b::canonically_ordered_monoid_add" assumes sub: "A \# B" shows "(\m \# A. f m) \ (\m \# B. f m)" by (metis image_mset_union le_iff_add sub subset_mset.add_diff_inverse sum_mset.union) lemma sum_image_mset_mono_mem: "n \# M \ f n \ (\m \# M. f m)" for f :: "'a \ 'b::canonically_ordered_monoid_add" using le_iff_add multi_member_split by fastforce lemma count_sum_mset_if_1_0: \count M a = (\x\#M. if x = a then 1 else 0)\ by (induction M) auto lemma sum_mset_dvd: fixes k :: "'a::comm_semiring_1_cancel" assumes "\m \# M. k dvd f m" shows "k dvd (\m \# M. f m)" using assms by (induct M) auto lemma sum_mset_distrib_div_if_dvd: fixes k :: "'a::unique_euclidean_semiring" assumes "\m \# M. k dvd f m" shows "(\m \# M. f m) div k = (\m \# M. f m div k)" using assms by (induct M) (auto simp: div_plus_div_distrib_dvd_left) subsection \Lemmas about Remove\ lemma set_mset_minus_replicate_mset[simp]: "n \ count A a \ set_mset (A - replicate_mset n a) = set_mset A - {a}" "n < count A a \ set_mset (A - replicate_mset n a) = set_mset A" unfolding set_mset_def by (auto split: if_split simp: not_in_iff) abbreviation removeAll_mset :: "'a \ 'a multiset \ 'a multiset" where "removeAll_mset C M \ M - replicate_mset (count M C) C" lemma mset_removeAll[simp, code]: "removeAll_mset C (mset L) = mset (removeAll C L)" by (induction L) (auto simp: ac_simps multiset_eq_iff split: if_split_asm) lemma removeAll_mset_filter_mset: "removeAll_mset C M = filter_mset ((\) C) M" by (induction M) (auto simp: ac_simps multiset_eq_iff) abbreviation remove1_mset :: "'a \ 'a multiset \ 'a multiset" where "remove1_mset C M \ M - {#C#}" lemma removeAll_subseteq_remove1_mset: "removeAll_mset x M \# remove1_mset x M" by (auto simp: subseteq_mset_def) lemma in_remove1_mset_neq: assumes ab: "a \ b" shows "a \# remove1_mset b C \ a \# C" by (metis assms diff_single_trivial in_diffD insert_DiffM insert_noteq_member) lemma size_mset_removeAll_mset_le_iff: "size (removeAll_mset x M) < size M \ x \# M" by (auto intro: count_inI mset_subset_size simp: subset_mset_def multiset_eq_iff) lemma size_remove1_mset_If: \size (remove1_mset x M) = size M - (if x \# M then 1 else 0)\ by (auto simp: size_Diff_subset_Int) lemma size_mset_remove1_mset_le_iff: "size (remove1_mset x M) < size M \ x \# M" using less_irrefl by (fastforce intro!: mset_subset_size elim: in_countE simp: subset_mset_def multiset_eq_iff) lemma remove_1_mset_id_iff_notin: "remove1_mset a M = M \ a \# M" by (meson diff_single_trivial multi_drop_mem_not_eq) lemma id_remove_1_mset_iff_notin: "M = remove1_mset a M \ a \# M" using remove_1_mset_id_iff_notin by metis lemma remove1_mset_eqE: "remove1_mset L x1 = M \ (L \# x1 \ x1 = M + {#L#} \ P) \ (L \# x1 \ x1 = M \ P) \ P" by (cases "L \# x1") auto lemma image_filter_ne_mset[simp]: "image_mset f {#x \# M. f x \ y#} = removeAll_mset y (image_mset f M)" by (induction M) simp_all lemma image_mset_remove1_mset_if: "image_mset f (remove1_mset a M) = (if a \# M then remove1_mset (f a) (image_mset f M) else image_mset f M)" by (auto simp: image_mset_Diff) lemma filter_mset_neq: "{#x \# M. x \ y#} = removeAll_mset y M" by (metis add_diff_cancel_left' filter_eq_replicate_mset multiset_partition) lemma filter_mset_neq_cond: "{#x \# M. P x \ x \ y#} = removeAll_mset y {# x\#M. P x#}" by (metis filter_filter_mset filter_mset_neq) lemma remove1_mset_add_mset_If: "remove1_mset L (add_mset L' C) = (if L = L' then C else remove1_mset L C + {#L'#})" by (auto simp: multiset_eq_iff) lemma minus_remove1_mset_if: "A - remove1_mset b B = (if b \# B \ b \# A \ count A b \ count B b then {#b#} + (A - B) else A - B)" by (auto simp: multiset_eq_iff count_greater_zero_iff[symmetric] simp del: count_greater_zero_iff) lemma add_mset_eq_add_mset_ne: "a \ b \ add_mset a A = add_mset b B \ a \# B \ b \# A \ A = add_mset b (B - {#a#})" by (metis (no_types, lifting) diff_single_eq_union diff_union_swap multi_self_add_other_not_self remove_1_mset_id_iff_notin union_single_eq_diff) lemma add_mset_eq_add_mset: \add_mset a M = add_mset b M' \ (a = b \ M = M') \ (a \ b \ b \# M \ add_mset a (M - {#b#}) = M')\ by (metis add_mset_eq_add_mset_ne add_mset_remove_trivial union_single_eq_member) (* TODO move to Multiset: could replace add_mset_remove_trivial_eq? *) lemma add_mset_remove_trivial_iff: \N = add_mset a (N - {#b#}) \ a \# N \ a = b\ by (metis add_left_cancel add_mset_remove_trivial insert_DiffM2 single_eq_single size_mset_remove1_mset_le_iff union_single_eq_member) lemma trivial_add_mset_remove_iff: \add_mset a (N - {#b#}) = N \ a \# N \ a = b\ by (subst eq_commute) (fact add_mset_remove_trivial_iff) lemma remove1_single_empty_iff[simp]: \remove1_mset L {#L'#} = {#} \ L = L'\ using add_mset_remove_trivial_iff by fastforce lemma add_mset_less_imp_less_remove1_mset: assumes xM_lt_N: "add_mset x M < N" shows "M < remove1_mset x N" proof - have "M < N" using assms le_multiset_right_total mset_le_trans by blast then show ?thesis by (metis add_less_cancel_right add_mset_add_single diff_single_trivial insert_DiffM2 xM_lt_N) qed +lemma remove_diff_multiset[simp]: \x13 \# A \ A - add_mset x13 B = A - B\ + by (metis diff_intersect_left_idem inter_add_right1) + +lemma removeAll_notin: \a \# A \ removeAll_mset a A = A\ + using count_inI by force + +lemma mset_drop_upto: \mset (drop a N) = {#N!i. i \# mset_set {a.. +proof (induction N arbitrary: a) + case Nil + then show ?case by simp +next + case (Cons c N) + have upt: \{0.. + by auto + then have H: \mset_set {0.. + unfolding upt by auto + have mset_case_Suc: \{#case x of 0 \ c | Suc x \ N ! x . x \# mset_set {Suc a..# mset_set {Suc a.. for a b + by (rule image_mset_cong) (auto split: nat.splits) + have Suc_Suc: \{Suc a.. for a b + by auto + then have mset_set_Suc_Suc: \mset_set {Suc a..# mset_set {a.. for a b + unfolding Suc_Suc by (subst image_mset_mset_set[symmetric]) auto + have *: \{#N ! (x-Suc 0) . x \# mset_set {Suc a..# mset_set {a.. + for a b + by (auto simp add: mset_set_Suc_Suc) + show ?case + apply (cases a) + using Cons[of 0] Cons by (auto simp: nth_Cons drop_Cons H mset_case_Suc *) +qed + subsection \Lemmas about Replicate\ lemma replicate_mset_minus_replicate_mset_same[simp]: "replicate_mset m x - replicate_mset n x = replicate_mset (m - n) x" by (induct m arbitrary: n, simp, metis left_diff_repeat_mset_distrib' repeat_mset_replicate_mset) lemma replicate_mset_subset_iff_lt[simp]: "replicate_mset m x \# replicate_mset n x \ m < n" by (induct n m rule: diff_induct) (auto intro: subset_mset.gr_zeroI) lemma replicate_mset_subseteq_iff_le[simp]: "replicate_mset m x \# replicate_mset n x \ m \ n" by (induct n m rule: diff_induct) auto lemma replicate_mset_lt_iff_lt[simp]: "replicate_mset m x < replicate_mset n x \ m < n" by (induct n m rule: diff_induct) (auto intro: subset_mset.gr_zeroI gr_zeroI) lemma replicate_mset_le_iff_le[simp]: "replicate_mset m x \ replicate_mset n x \ m \ n" by (induct n m rule: diff_induct) auto lemma replicate_mset_eq_iff[simp]: "replicate_mset m x = replicate_mset n y \ m = n \ (m \ 0 \ x = y)" by (cases m; cases n; simp) (metis in_replicate_mset insert_noteq_member size_replicate_mset union_single_eq_diff) lemma replicate_mset_plus: "replicate_mset (a + b) C = replicate_mset a C + replicate_mset b C" by (induct a) (auto simp: ac_simps) lemma mset_replicate_replicate_mset: "mset (replicate n L) = replicate_mset n L" by (induction n) auto lemma set_mset_single_iff_replicate_mset: "set_mset U = {a} \ (\n > 0. U = replicate_mset n a)" by (rule, metis count_greater_zero_iff count_replicate_mset insertI1 multi_count_eq singletonD zero_less_iff_neq_zero, force) lemma ex_replicate_mset_if_all_elems_eq: assumes "\x \# M. x = y" shows "\n. M = replicate_mset n y" using assms by (metis count_replicate_mset mem_Collect_eq multiset_eqI neq0_conv set_mset_def) subsection \Multiset and Set Conversions\ lemma count_mset_set_if: "count (mset_set A) a = (if a \ A \ finite A then 1 else 0)" by auto lemma mset_set_set_mset_empty_mempty[iff]: "mset_set (set_mset D) = {#} \ D = {#}" by (simp add: mset_set_empty_iff) lemma count_mset_set_le_one: "count (mset_set A) x \ 1" by (simp add: count_mset_set_if) lemma mset_set_set_mset_subseteq[simp]: "mset_set (set_mset A) \# A" by (simp add: mset_set_set_mset_msubset) lemma mset_sorted_list_of_set[simp]: "mset (sorted_list_of_set A) = mset_set A" by (metis mset_sorted_list_of_multiset sorted_list_of_mset_set) lemma sorted_sorted_list_of_multiset[simp]: "sorted (sorted_list_of_multiset (M :: 'a::linorder multiset))" by (metis mset_sorted_list_of_multiset sorted_list_of_multiset_mset sorted_sort) lemma mset_take_subseteq: "mset (take n xs) \# mset xs" apply (induct xs arbitrary: n) apply simp by (case_tac n) simp_all lemma sorted_list_of_multiset_eq_Nil[simp]: "sorted_list_of_multiset M = [] \ M = {#}" by (metis mset_sorted_list_of_multiset sorted_list_of_multiset_empty) subsection \Duplicate Removal\ (* TODO: use abbreviation? *) definition remdups_mset :: "'v multiset \ 'v multiset" where "remdups_mset S = mset_set (set_mset S)" lemma set_mset_remdups_mset[simp]: \set_mset (remdups_mset A) = set_mset A\ unfolding remdups_mset_def by auto lemma count_remdups_mset_eq_1: "a \# remdups_mset A \ count (remdups_mset A) a = 1" unfolding remdups_mset_def by (auto simp: count_eq_zero_iff intro: count_inI) lemma remdups_mset_empty[simp]: "remdups_mset {#} = {#}" unfolding remdups_mset_def by auto lemma remdups_mset_singleton[simp]: "remdups_mset {#a#} = {#a#}" unfolding remdups_mset_def by auto lemma remdups_mset_eq_empty[iff]: "remdups_mset D = {#} \ D = {#}" unfolding remdups_mset_def by blast lemma remdups_mset_singleton_sum[simp]: "remdups_mset (add_mset a A) = (if a \# A then remdups_mset A else add_mset a (remdups_mset A))" unfolding remdups_mset_def by (simp_all add: insert_absorb) lemma mset_remdups_remdups_mset[simp]: "mset (remdups D) = remdups_mset (mset D)" by (induction D) (auto simp add: ac_simps) declare mset_remdups_remdups_mset[symmetric, code] -definition distinct_mset :: "'a multiset \ bool" where - "distinct_mset S \ (\a. a \# S \ count S a = 1)" - -lemma distinct_mset_count_less_1: "distinct_mset S \ (\a. count S a \ 1)" - using eq_iff nat_le_linear unfolding distinct_mset_def by fastforce - -lemma distinct_mset_empty[simp]: "distinct_mset {#}" - unfolding distinct_mset_def by auto - -lemma distinct_mset_singleton: "distinct_mset {#a#}" - unfolding distinct_mset_def by auto - -lemma distinct_mset_union: - assumes dist: "distinct_mset (A + B)" - shows "distinct_mset A" - unfolding distinct_mset_count_less_1 -proof (rule allI) - fix a - have \count A a \ count (A + B) a\ by auto - moreover have \count (A + B) a \ 1\ - using dist unfolding distinct_mset_count_less_1 by auto - ultimately show \count A a \ 1\ - by simp -qed - -lemma distinct_mset_minus[simp]: "distinct_mset A \ distinct_mset (A - B)" - by (metis diff_subset_eq_self mset_subset_eq_exists_conv distinct_mset_union) - lemma count_remdups_mset_If: \count (remdups_mset A) a = (if a \# A then 1 else 0)\ unfolding remdups_mset_def by auto -lemma distinct_mset_rempdups_union_mset: - assumes "distinct_mset A" and "distinct_mset B" - shows "A \# B = remdups_mset (A + B)" - using assms nat_le_linear unfolding remdups_mset_def - by (force simp add: multiset_eq_iff max_def count_mset_set_if distinct_mset_def not_in_iff) - -lemma distinct_mset_add_mset[simp]: "distinct_mset (add_mset a L) \ a \# L \ distinct_mset L" - unfolding distinct_mset_def - apply (rule iffI) - apply (auto split: if_split_asm; fail)[] - by (auto simp: not_in_iff; fail) - -lemma distinct_mset_size_eq_card: "distinct_mset C \ size C = card (set_mset C)" - by (induction C) auto - -lemma distinct_mset_add: - "distinct_mset (L + L') \ distinct_mset L \ distinct_mset L' \ L \# L' = {#}" - by (induction L arbitrary: L') auto - -lemma distinct_mset_set_mset_ident[simp]: "distinct_mset M \ mset_set (set_mset M) = M" - by (induction M) auto - -lemma distinct_finite_set_mset_subseteq_iff[iff]: - assumes "distinct_mset M" "finite N" - shows "set_mset M \ N \ M \# mset_set N" - by (metis assms distinct_mset_set_mset_ident finite_set_mset msubset_mset_set_iff) - -lemma distinct_mem_diff_mset: - assumes dist: "distinct_mset M" and mem: "x \ set_mset (M - N)" - shows "x \ set_mset N" -proof - - have "count M x = 1" - using dist mem by (meson distinct_mset_def in_diffD) - then show ?thesis - using mem by (metis count_greater_eq_one_iff in_diff_count not_less) -qed - -lemma distinct_set_mset_eq: - assumes "distinct_mset M" "distinct_mset N" "set_mset M = set_mset N" - shows "M = N" - using assms distinct_mset_set_mset_ident by fastforce - -lemma distinct_mset_union_mset[simp]: - \distinct_mset (D \# C) \ distinct_mset D \ distinct_mset C\ - unfolding distinct_mset_count_less_1 by force - -lemma distinct_mset_inter_mset: - "distinct_mset C \ distinct_mset (C \# D)" - "distinct_mset D \ distinct_mset (C \# D)" - by (auto simp add: distinct_mset_def min_def count_eq_zero_iff elim!: le_SucE) - -lemma distinct_mset_remove1_All: "distinct_mset C \ remove1_mset L C = removeAll_mset L C" - by (auto simp: multiset_eq_iff distinct_mset_count_less_1) - -lemma distinct_mset_size_2: "distinct_mset {#a, b#} \ a \ b" +lemma notin_add_mset_remdups_mset: + \a \# A \ add_mset a (remdups_mset A) = remdups_mset (add_mset a A)\ by auto -lemma distinct_mset_filter: "distinct_mset M \ distinct_mset {# L \# M. P L#}" - by (simp add: distinct_mset_def) - -lemma distinct_mset_mset_distinct[simp]: \distinct_mset (mset xs) = distinct xs\ - by (induction xs) auto - -lemma distinct_image_mset_inj: - \inj_on f (set_mset M) \ distinct_mset (image_mset f M) \ distinct_mset M\ - by (induction M) (auto simp: inj_on_def) - subsection \Repeat Operation\ lemma repeat_mset_compower: "repeat_mset n A = (((+) A) ^^ n) {#}" by (induction n) auto lemma repeat_mset_prod: "repeat_mset (m * n) A = (((+) (repeat_mset n A)) ^^ m) {#}" by (induction m) (auto simp: repeat_mset_distrib) subsection \Cartesian Product\ text \Definition of the cartesian products over multisets. The construction mimics of the cartesian product on sets and use the same theorem names (adding only the suffix \_mset\ to Sigma and Times). See file @{file \~~/src/HOL/Product_Type.thy\}\ definition Sigma_mset :: "'a multiset \ ('a \ 'b multiset) \ ('a \ 'b) multiset" where "Sigma_mset A B \ \\<^sub># {#{#(a, b). b \# B a#}. a \# A #}" abbreviation Times_mset :: "'a multiset \ 'b multiset \ ('a \ 'b) multiset" (infixr "\#" 80) where "Times_mset A B \ Sigma_mset A (\_. B)" hide_const (open) Times_mset text \Contrary to the set version @{term \SIGMA x:A. B\}, we use the non-ASCII symbol \\#\.\ syntax "_Sigma_mset" :: "[pttrn, 'a multiset, 'b multiset] => ('a * 'b) multiset" ("(3SIGMAMSET _\#_./ _)" [0, 0, 10] 10) translations "SIGMAMSET x\#A. B" == "CONST Sigma_mset A (\x. B)" text \Link between the multiset and the set cartesian product:\ lemma Times_mset_Times: "set_mset (A \# B) = set_mset A \ set_mset B" unfolding Sigma_mset_def by auto lemma Sigma_msetI [intro!]: "\a \# A; b \# B a\ \ (a, b) \# Sigma_mset A B" by (unfold Sigma_mset_def) auto lemma Sigma_msetE[elim!]: "\c \# Sigma_mset A B; \x y. \x \# A; y \# B x; c = (x, y)\ \ P\ \ P" by (unfold Sigma_mset_def) auto text \Elimination of @{term "(a, b) \# A \# B"} -- introduces no eigenvariables.\ lemma Sigma_msetD1: "(a, b) \# Sigma_mset A B \ a \# A" by blast lemma Sigma_msetD2: "(a, b) \# Sigma_mset A B \ b \# B a" by blast lemma Sigma_msetE2: "\(a, b) \# Sigma_mset A B; \a \# A; b \# B a\ \ P\ \ P" by blast lemma Sigma_mset_cong: "\A = B; \x. x \# B \ C x = D x\ \ (SIGMAMSET x \# A. C x) = (SIGMAMSET x \# B. D x)" by (metis (mono_tags, lifting) Sigma_mset_def image_mset_cong) lemma count_sum_mset: "count (\\<^sub># M) b = (\P \# M. count P b)" by (induction M) auto lemma Sigma_mset_plus_distrib1[simp]: "Sigma_mset (A + B) C = Sigma_mset A C + Sigma_mset B C" unfolding Sigma_mset_def by auto lemma Sigma_mset_plus_distrib2[simp]: "Sigma_mset A (\i. B i + C i) = Sigma_mset A B + Sigma_mset A C" unfolding Sigma_mset_def by (induction A) (auto simp: multiset_eq_iff) lemma Times_mset_single_left: "{#a#} \# B = image_mset (Pair a) B" unfolding Sigma_mset_def by auto lemma Times_mset_single_right: "A \# {#b#} = image_mset (\a. Pair a b) A" unfolding Sigma_mset_def by (induction A) auto lemma Times_mset_single_single[simp]: "{#a#} \# {#b#} = {#(a, b)#}" unfolding Sigma_mset_def by simp lemma count_image_mset_Pair: "count (image_mset (Pair a) B) (x, b) = (if x = a then count B b else 0)" by (induction B) auto lemma count_Sigma_mset: "count (Sigma_mset A B) (a, b) = count A a * count (B a) b" by (induction A) (auto simp: Sigma_mset_def count_image_mset_Pair) lemma Sigma_mset_empty1[simp]: "Sigma_mset {#} B = {#}" unfolding Sigma_mset_def by auto lemma Sigma_mset_empty2[simp]: "A \# {#} = {#}" by (auto simp: multiset_eq_iff count_Sigma_mset) lemma Sigma_mset_mono: assumes "A \# C" and "\x. x \# A \ B x \# D x" shows "Sigma_mset A B \# Sigma_mset C D" proof - have "count A a * count (B a) b \ count C a * count (D a) b" for a b using assms unfolding subseteq_mset_def by (metis count_inI eq_iff mult_eq_0_iff mult_le_mono) then show ?thesis by (auto simp: subseteq_mset_def count_Sigma_mset) qed lemma mem_Sigma_mset_iff[iff]: "((a,b) \# Sigma_mset A B) = (a \# A \ b \# B a)" by blast lemma mem_Times_mset_iff: "x \# A \# B \ fst x \# A \ snd x \# B" by (induct x) simp lemma Sigma_mset_empty_iff: "(SIGMAMSET i\#I. X i) = {#} \ (\i\#I. X i = {#})" by (auto simp: Sigma_mset_def) lemma Times_mset_subset_mset_cancel1: "x \# A \ (A \# B \# A \# C) = (B \# C)" by (auto simp: subseteq_mset_def count_Sigma_mset) lemma Times_mset_subset_mset_cancel2: "x \# C \ (A \# C \# B \# C) = (A \# B)" by (auto simp: subseteq_mset_def count_Sigma_mset) lemma Times_mset_eq_cancel2: "x \# C \ (A \# C = B \# C) = (A = B)" by (auto simp: multiset_eq_iff count_Sigma_mset dest!: in_countE) lemma split_paired_Ball_mset_Sigma_mset[simp]: "(\z\#Sigma_mset A B. P z) \ (\x\#A. \y\#B x. P (x, y))" by blast lemma split_paired_Bex_mset_Sigma_mset[simp]: "(\z\#Sigma_mset A B. P z) \ (\x\#A. \y\#B x. P (x, y))" by blast lemma sum_mset_if_eq_constant: "(\x\#M. if a = x then (f x) else 0) = (((+) (f a)) ^^ (count M a)) 0" by (induction M) (auto simp: ac_simps) lemma iterate_op_plus: "(((+) k) ^^ m) 0 = k * m" by (induction m) auto lemma untion_image_mset_Pair_distribute: "\\<^sub>#{#image_mset (Pair x) (C x). x \# J - I#} = \\<^sub># {#image_mset (Pair x) (C x). x \# J#} - \\<^sub>#{#image_mset (Pair x) (C x). x \# I#}" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant iterate_op_plus diff_mult_distrib2) lemma Sigma_mset_Un_distrib1: "Sigma_mset (I \# J) C = Sigma_mset I C \# Sigma_mset J C" by (auto simp add: Sigma_mset_def union_mset_def untion_image_mset_Pair_distribute) lemma Sigma_mset_Un_distrib2: "(SIGMAMSET i\#I. A i \# B i) = Sigma_mset I A \# Sigma_mset I B" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def diff_mult_distrib2 iterate_op_plus max_def not_in_iff) lemma Sigma_mset_Int_distrib1: "Sigma_mset (I \# J) C = Sigma_mset I C \# Sigma_mset J C" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff) lemma Sigma_mset_Int_distrib2: "(SIGMAMSET i\#I. A i \# B i) = Sigma_mset I A \# Sigma_mset I B" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff) lemma Sigma_mset_Diff_distrib1: "Sigma_mset (I - J) C = Sigma_mset I C - Sigma_mset J C" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff diff_mult_distrib2) lemma Sigma_mset_Diff_distrib2: "(SIGMAMSET i\#I. A i - B i) = Sigma_mset I A - Sigma_mset I B" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff diff_mult_distrib) lemma Sigma_mset_Union: "Sigma_mset (\\<^sub>#X) B = (\\<^sub># (image_mset (\A. Sigma_mset A B) X))" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff sum_mset_distrib_left) lemma Times_mset_Un_distrib1: "(A \# B) \# C = A \# C \# B \# C" by (fact Sigma_mset_Un_distrib1) lemma Times_mset_Int_distrib1: "(A \# B) \# C = A \# C \# B \# C" by (fact Sigma_mset_Int_distrib1) lemma Times_mset_Diff_distrib1: "(A - B) \# C = A \# C - B \# C" by (fact Sigma_mset_Diff_distrib1) lemma Times_mset_empty[simp]: "A \# B = {#} \ A = {#} \ B = {#}" by (auto simp: Sigma_mset_empty_iff) lemma Times_insert_left: "A \# add_mset x B = A \# B + image_mset (\a. Pair a x) A" unfolding add_mset_add_single[of x B] Sigma_mset_plus_distrib2 by (simp add: Times_mset_single_right) lemma Times_insert_right: "add_mset a A \# B = A \# B + image_mset (Pair a) B" unfolding add_mset_add_single[of a A] Sigma_mset_plus_distrib1 by (simp add: Times_mset_single_left) lemma fst_image_mset_times_mset [simp]: "image_mset fst (A \# B) = (if B = {#} then {#} else repeat_mset (size B) A)" by (induct B) (auto simp: Times_mset_single_right ac_simps Times_insert_left) lemma snd_image_mset_times_mset [simp]: "image_mset snd (A \# B) = (if A = {#} then {#} else repeat_mset (size A) B)" by (induct B) (auto simp add: Times_mset_single_right Times_insert_left image_mset_const_eq) lemma product_swap_mset: "image_mset prod.swap (A \# B) = B \# A" by (induction A) (auto simp add: Times_mset_single_left Times_mset_single_right Times_insert_right Times_insert_left) context begin qualified definition product_mset :: "'a multiset \ 'b multiset \ ('a \ 'b) multiset" where [code_abbrev]: "product_mset A B = A \# B" lemma member_product_mset: "x \# product_mset A B \ x \# A \# B" by (simp add: Multiset_More.product_mset_def) end lemma count_Sigma_mset_abs_def: "count (Sigma_mset A B) = (\(a, b) \ count A a * count (B a) b)" by (auto simp: fun_eq_iff count_Sigma_mset) lemma Times_mset_image_mset1: "image_mset f A \# B = image_mset (\(a, b). (f a, b)) (A \# B)" by (induct B) (auto simp: Times_insert_left) lemma Times_mset_image_mset2: "A \# image_mset f B = image_mset (\(a, b). (a, f b)) (A \# B)" by (induct A) (auto simp: Times_insert_right) lemma sum_le_singleton: "A \ {x} \ sum f A = (if x \ A then f x else 0)" by (auto simp: subset_singleton_iff elim: finite_subset) lemma Times_mset_assoc: "(A \# B) \# C = image_mset (\(a, b, c). ((a, b), c)) (A \# B \# C)" by (auto simp: multiset_eq_iff count_Sigma_mset count_image_mset vimage_def Times_mset_Times Int_commute count_eq_zero_iff intro!: trans[OF _ sym[OF sum_le_singleton[of _ "(_, _, _)"]]] cong: sum.cong if_cong) subsection \Transfer Rules\ lemma plus_multiset_transfer[transfer_rule]: "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (rel_mset R))) (+) (+)" by (unfold rel_fun_def rel_mset_def) (force dest: list_all2_appendI intro: exI[of _ "_ @ _"] conjI[rotated]) lemma minus_multiset_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique R" shows "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (rel_mset R))) (-) (-)" proof (unfold rel_fun_def rel_mset_def, safe) fix xs ys xs' ys' assume [transfer_rule]: "list_all2 R xs ys" "list_all2 R xs' ys'" have "list_all2 R (fold remove1 xs' xs) (fold remove1 ys' ys)" by transfer_prover moreover have "mset (fold remove1 xs' xs) = mset xs - mset xs'" by (induct xs' arbitrary: xs) auto moreover have "mset (fold remove1 ys' ys) = mset ys - mset ys'" by (induct ys' arbitrary: ys) auto ultimately show "\xs'' ys''. mset xs'' = mset xs - mset xs' \ mset ys'' = mset ys - mset ys' \ list_all2 R xs'' ys''" by blast qed declare rel_mset_Zero[transfer_rule] lemma count_transfer[transfer_rule]: assumes "bi_unique R" shows "(rel_fun (rel_mset R) (rel_fun R (=))) count count" unfolding rel_fun_def rel_mset_def proof safe fix x y xs ys assume "list_all2 R xs ys" "R x y" then show "count (mset xs) x = count (mset ys) y" proof (induct xs ys rule: list.rel_induct) case (Cons x' xs y' ys) then show ?case using assms unfolding bi_unique_alt_def2 by (auto simp: rel_fun_def) qed simp qed lemma subseteq_multiset_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique R" "right_total R" shows "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (=))) (\M N. filter_mset (Domainp R) M \# filter_mset (Domainp R) N) (\#)" proof - have count_filter_mset_less: "(\a. count (filter_mset (Domainp R) M) a \ count (filter_mset (Domainp R) N) a) \ (\a \ {x. Domainp R x}. count M a \ count N a)" for M and N by auto show ?thesis unfolding subseteq_mset_def count_filter_mset_less by transfer_prover qed lemma sum_mset_transfer[transfer_rule]: "R 0 0 \ rel_fun R (rel_fun R R) (+) (+) \ (rel_fun (rel_mset R) R) sum_mset sum_mset" using sum_list_transfer[of R] unfolding rel_fun_def rel_mset_def by auto lemma Sigma_mset_transfer[transfer_rule]: "(rel_fun (rel_mset R) (rel_fun (rel_fun R (rel_mset S)) (rel_mset (rel_prod R S)))) Sigma_mset Sigma_mset" by (unfold Sigma_mset_def) transfer_prover subsection \Even More about Multisets\ subsubsection \Multisets and Functions\ lemma range_image_mset: assumes "set_mset Ds \ range f" shows "Ds \ range (image_mset f)" proof - have "\D. D \# Ds \ (\C. f C = D)" using assms by blast then obtain f_i where f_p: "\D. D \# Ds \ (f (f_i D) = D)" by metis define Cs where "Cs \ image_mset f_i Ds" from f_p Cs_def have "image_mset f Cs = Ds" by auto then show ?thesis by blast qed subsubsection \Multisets and Lists\ lemma length_sorted_list_of_multiset[simp]: "length (sorted_list_of_multiset A) = size A" by (metis mset_sorted_list_of_multiset size_mset) definition list_of_mset :: "'a multiset \ 'a list" where "list_of_mset m = (SOME l. m = mset l)" lemma list_of_mset_exi: "\l. m = mset l" using ex_mset by metis lemma mset_list_of_mset[simp]: "mset (list_of_mset m) = m" by (metis (mono_tags, lifting) ex_mset list_of_mset_def someI_ex) lemma length_list_of_mset[simp]: "length (list_of_mset A) = size A" unfolding list_of_mset_def by (metis (mono_tags) ex_mset size_mset someI_ex) lemma range_mset_map: assumes "set_mset Ds \ range f" shows "Ds \ range (\Cl. mset (map f Cl))" proof - have "Ds \ range (image_mset f)" by (simp add: assms range_image_mset) then obtain Cs where Cs_p: "image_mset f Cs = Ds" by auto define Cl where "Cl = list_of_mset Cs" then have "mset Cl = Cs" by auto then have "image_mset f (mset Cl) = Ds" using Cs_p by auto then have "mset (map f Cl) = Ds" by auto then show ?thesis by auto qed lemma list_of_mset_empty[iff]: "list_of_mset m = [] \ m = {#}" by (metis (mono_tags, lifting) ex_mset list_of_mset_def mset_zero_iff_right someI_ex) lemma in_mset_conv_nth: "(x \# mset xs) = (\i# LL" assumes "LL \ set Ci" shows "L \# sum_list Ci" using assms by (induction Ci) auto lemma in_mset_sum_list2: assumes "L \# sum_list Ci" obtains LL where "LL \ set Ci" "L \# LL" using assms by (induction Ci) auto (* TODO: Make [simp]. *) lemma in_mset_sum_list_iff: "a \# sum_list \ \ (\A \ set \. a \# A)" by (metis in_mset_sum_list in_mset_sum_list2) lemma subseteq_list_Union_mset: assumes "length Ci = n" assumes "length CAi = n" assumes "\i# CAi ! i " shows "\\<^sub># (mset Ci) \# \\<^sub># (mset CAi)" using assms proof (induction n arbitrary: Ci CAi) case 0 then show ?case by auto next case (Suc n) from Suc have "\i# tl CAi ! i" by (simp add: nth_tl) hence "\\<^sub>#(mset (tl Ci)) \# \\<^sub>#(mset (tl CAi))" using Suc by auto moreover have "hd Ci \# hd CAi" using Suc by (metis hd_conv_nth length_greater_0_conv zero_less_Suc) ultimately show "\\<^sub>#(mset Ci) \# \\<^sub>#(mset CAi)" using Suc by (cases Ci; cases CAi) (auto intro: subset_mset.add_mono) qed +lemma same_mset_distinct_iff: + \mset M = mset M' \ distinct M \ distinct M'\ + by (fact mset_eq_imp_distinct_iff) + subsubsection \More on Multisets and Functions\ lemma subseteq_mset_size_eql: "X \# Y \ size Y = size X \ X = Y" using mset_subset_size subset_mset_def by fastforce lemma image_mset_of_subset_list: assumes "image_mset \ C' = mset lC" shows "\qC'. map \ qC' = lC \ mset qC' = C'" using assms apply (induction lC arbitrary: C') subgoal by simp subgoal by (fastforce dest!: msed_map_invR intro: exI[of _ \_ # _\]) done lemma image_mset_of_subset: assumes "A \# image_mset \ C'" shows "\A'. image_mset \ A' = A \ A' \# C'" proof - define C where "C = image_mset \ C'" define lA where "lA = list_of_mset A" define lD where "lD = list_of_mset (C-A)" define lC where "lC = lA @ lD" have "mset lC = C" using C_def assms unfolding lD_def lC_def lA_def by auto then have "\qC'. map \ qC' = lC \ mset qC' = C'" using assms image_mset_of_subset_list unfolding C_def by metis then obtain qC' where qC'_p: "map \ qC' = lC \ mset qC' = C'" by auto let ?lA' = "take (length lA) qC'" have m: "map \ ?lA' = lA" using qC'_p lC_def by (metis append_eq_conv_conj take_map) let ?A' = "mset ?lA'" have "image_mset \ ?A' = A" using m using lA_def by (metis (full_types) ex_mset list_of_mset_def mset_map someI_ex) moreover have "?A' \# C'" using qC'_p unfolding lA_def using mset_take_subseteq by blast ultimately show ?thesis by blast qed lemma all_the_same: "\x \# X. x = y \ card (set_mset X) \ Suc 0" by (metis card.empty card.insert card_mono finite.intros(1) finite_insert le_SucI singletonI subsetI) lemma Melem_subseteq_Union_mset[simp]: assumes "x \# T" shows "x \# \\<^sub>#T" using assms sum_mset.remove by force lemma Melem_subset_eq_sum_list[simp]: assumes "x \# mset T" shows "x \# sum_list T" using assms by (metis mset_subset_eq_add_left sum_mset.remove sum_mset_sum_list) lemma less_subset_eq_Union_mset[simp]: assumes "i < length CAi" shows "CAi ! i \# \\<^sub>#(mset CAi)" proof - from assms have "CAi ! i \# mset CAi" by auto then show ?thesis by auto qed lemma less_subset_eq_sum_list[simp]: assumes "i < length CAi" shows "CAi ! i \# sum_list CAi" proof - from assms have "CAi ! i \# mset CAi" by auto then show ?thesis by auto qed subsubsection \More on Multiset Order\ lemma less_multiset_doubletons: assumes "y < t \ y < s" "x < t \ x < s" shows "{#y, x#} < {#t, s#}" unfolding less_multiset\<^sub>D\<^sub>M proof (intro exI) let ?X = "{#t, s#}" let ?Y = "{#y, x#}" show "?X \ {#} \ ?X \# {#t, s#} \ {#y, x#} = {#t, s#} - ?X + ?Y \ (\k. k \# ?Y \ (\a. a \# ?X \ k < a))" using add_eq_conv_diff assms by auto qed end diff --git a/thys/Ordered_Resolution_Prover/Standard_Redundancy.thy b/thys/Ordered_Resolution_Prover/Standard_Redundancy.thy --- a/thys/Ordered_Resolution_Prover/Standard_Redundancy.thy +++ b/thys/Ordered_Resolution_Prover/Standard_Redundancy.thy @@ -1,339 +1,339 @@ (* Title: The Standard Redundancy Criterion Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2014 Author: Anders Schlichtkrull , 2017 Maintainer: Anders Schlichtkrull *) section \The Standard Redundancy Criterion\ theory Standard_Redundancy imports Proving_Process begin text \ This material is based on Section 4.2.2 (``The Standard Redundancy Criterion'') of Bachmair and Ganzinger's chapter. \ locale standard_redundancy_criterion = inference_system \ for \ :: "('a :: wellorder) inference set" begin definition redundant_infer :: "'a clause set \ 'a inference \ bool" where "redundant_infer N \ \ (\DD. set_mset DD \ N \ (\I. I \m DD + side_prems_of \ \ I \ concl_of \) \ (\D. D \# DD \ D < main_prem_of \))" definition Rf :: "'a clause set \ 'a clause set" where "Rf N = {C. \DD. set_mset DD \ N \ (\I. I \m DD \ I \ C) \ (\D. D \# DD \ D < C)}" definition Ri :: "'a clause set \ 'a inference set" where "Ri N = {\ \ \. redundant_infer N \}" lemma tautology_Rf: assumes "Pos A \# C" assumes "Neg A \# C" shows "C \ Rf N" proof - have "set_mset {#} \ N \ (\I. I \m {#} \ I \ C) \ (\D. D \# {#} \ D < C)" using assms by auto then show "C \ Rf N" unfolding Rf_def by blast qed lemma tautology_redundant_infer: assumes pos: "Pos A \# concl_of \" and neg: "Neg A \# concl_of \" shows "redundant_infer N \" by (metis empty_iff empty_subsetI neg pos pos_neg_in_imp_true redundant_infer_def set_mset_empty) lemma contradiction_Rf: "{#} \ N \ Rf N = UNIV - {{#}}" unfolding Rf_def by force text \ The following results correspond to Lemma 4.5. The lemma \wlog_non_Rf\ generalizes the core of the argument. \ lemma Rf_mono: "N \ N' \ Rf N \ Rf N'" unfolding Rf_def by auto lemma wlog_non_Rf: assumes ex: "\DD. set_mset DD \ N \ (\I. I \m DD + CC \ I \ E) \ (\D'. D' \# DD \ D' < D)" shows "\DD. set_mset DD \ N - Rf N \ (\I. I \m DD + CC \ I \ E) \ (\D'. D' \# DD \ D' < D)" proof - from ex obtain DD0 where dd0: "DD0 \ {DD. set_mset DD \ N \ (\I. I \m DD + CC \ I \ E) \ (\D'. D' \# DD \ D' < D)}" by blast have "\DD. set_mset DD \ N \ (\I. I \m DD + CC \ I \ E) \ (\D'. D' \# DD \ D' < D) \ (\DD'. set_mset DD' \ N \ (\I. I \m DD' + CC \ I \ E) \ (\D'. D' \# DD' \ D' < D) \ DD \ DD')" using wf_eq_minimal[THEN iffD1, rule_format, OF wf_less_multiset dd0] unfolding not_le[symmetric] by blast then obtain DD where dd_subs_n: "set_mset DD \ N" and ddcc_imp_e: "\I. I \m DD + CC \ I \ E" and dd_lt_d: "\D'. D' \# DD \ D' < D" and d_min: "\DD'. set_mset DD' \ N \ (\I. I \m DD' + CC \ I \ E) \ (\D'. D' \# DD' \ D' < D) \ DD \ DD'" by blast have "\Da. Da \# DD \ Da \ Rf N" proof clarify fix Da assume da_in_dd: "Da \# DD" and da_rf: "Da \ Rf N" from da_rf obtain DD' where dd'_subs_n: "set_mset DD' \ N" and dd'_imp_da: "\I. I \m DD' \ I \ Da" and dd'_lt_da: "\D'. D' \# DD' \ D' < Da" unfolding Rf_def by blast define DDa where "DDa = DD - {#Da#} + DD'" have "set_mset DDa \ N" unfolding DDa_def using dd_subs_n dd'_subs_n by (meson contra_subsetD in_diffD subsetI union_iff) moreover have "\I. I \m DDa + CC \ I \ E" using dd'_imp_da ddcc_imp_e da_in_dd unfolding DDa_def true_cls_mset_def by (metis in_remove1_mset_neq union_iff) moreover have "\D'. D' \# DDa \ D' < D" using dd_lt_d dd'_lt_da da_in_dd unfolding DDa_def by (metis insert_DiffM2 order.strict_trans union_iff) moreover have "DDa < DD" unfolding DDa_def by (meson da_in_dd dd'_lt_da mset_lt_single_right_iff single_subset_iff union_le_diff_plus) ultimately show False using d_min unfolding less_eq_multiset_def by (auto intro!: antisym) qed then show ?thesis using dd_subs_n ddcc_imp_e dd_lt_d by auto qed lemma Rf_imp_ex_non_Rf: assumes "C \ Rf N" shows "\CC. set_mset CC \ N - Rf N \ (\I. I \m CC \ I \ C) \ (\C'. C' \# CC \ C' < C)" using assms by (auto simp: Rf_def intro: wlog_non_Rf[of _ "{#}", simplified]) lemma Rf_subs_Rf_diff_Rf: "Rf N \ Rf (N - Rf N)" proof fix C assume c_rf: "C \ Rf N" then obtain CC where cc_subs: "set_mset CC \ N - Rf N" and cc_imp_c: "\I. I \m CC \ I \ C" and cc_lt_c: "\C'. C' \# CC \ C' < C" using Rf_imp_ex_non_Rf by blast have "\D. D \# CC \ D \ Rf N" using cc_subs by (simp add: subset_iff) then have cc_nr: "\C DD. C \# CC \ set_mset DD \ N \ \I. I \m DD \ I \ C \ \D. D \# DD \ \ D < C" unfolding Rf_def by auto metis have "set_mset CC \ N" using cc_subs by auto then have "set_mset CC \ N - {C. \DD. set_mset DD \ N \ (\I. I \m DD \ I \ C) \ (\D. D \# DD \ D < C)}" - using cc_nr by auto + using cc_nr by blast then show "C \ Rf (N - Rf N)" using cc_imp_c cc_lt_c unfolding Rf_def by auto qed lemma Rf_eq_Rf_diff_Rf: "Rf N = Rf (N - Rf N)" by (metis Diff_subset Rf_mono Rf_subs_Rf_diff_Rf subset_antisym) text \ The following results correspond to Lemma 4.6. \ lemma Ri_mono: "N \ N' \ Ri N \ Ri N'" unfolding Ri_def redundant_infer_def by auto lemma Ri_subs_Ri_diff_Rf: "Ri N \ Ri (N - Rf N)" proof fix \ assume \_ri: "\ \ Ri N" then obtain CC D E where \: "\ = Infer CC D E" by (cases \) have cc: "CC = side_prems_of \" and d: "D = main_prem_of \" and e: "E = concl_of \" unfolding \ by simp_all obtain DD where "set_mset DD \ N" and "\I. I \m DD + CC \ I \ E" and "\C. C \# DD \ C < D" using \_ri unfolding Ri_def redundant_infer_def cc d e by blast then obtain DD' where "set_mset DD' \ N - Rf N" and "\I. I \m DD' + CC \ I \ E" and "\D'. D' \# DD' \ D' < D" using wlog_non_Rf by atomize_elim blast then show "\ \ Ri (N - Rf N)" using \_ri unfolding Ri_def redundant_infer_def d cc e by blast qed lemma Ri_eq_Ri_diff_Rf: "Ri N = Ri (N - Rf N)" by (metis Diff_subset Ri_mono Ri_subs_Ri_diff_Rf subset_antisym) lemma Ri_subset_\: "Ri N \ \" unfolding Ri_def by blast lemma Rf_indep: "N' \ Rf N \ Rf N \ Rf (N - N')" by (metis Diff_cancel Diff_eq_empty_iff Diff_mono Rf_eq_Rf_diff_Rf Rf_mono) lemma Ri_indep: "N' \ Rf N \ Ri N \ Ri (N - N')" by (metis Diff_mono Ri_eq_Ri_diff_Rf Ri_mono order_refl) lemma Rf_model: assumes "I \s N - Rf N" shows "I \s N" proof - have "I \s Rf (N - Rf N)" unfolding true_clss_def by (subst Rf_def, simp add: true_cls_mset_def, metis assms subset_eq true_clss_def) then have "I \s Rf N" using Rf_subs_Rf_diff_Rf true_clss_mono by blast then show ?thesis using assms by (metis Un_Diff_cancel true_clss_union) qed lemma Rf_sat: "satisfiable (N - Rf N) \ satisfiable N" by (metis Rf_model) text \ The following corresponds to Theorem 4.7: \ sublocale redundancy_criterion \ Rf Ri by unfold_locales (rule Ri_subset_\, (elim Rf_mono Ri_mono Rf_indep Ri_indep Rf_sat)+) end locale standard_redundancy_criterion_reductive = standard_redundancy_criterion + reductive_inference_system begin text \ The following corresponds to Theorem 4.8: \ lemma Ri_effective: assumes in_\: "\ \ \" and concl_of_in_n_un_rf_n: "concl_of \ \ N \ Rf N" shows "\ \ Ri N" proof - obtain CC D E where \: "\ = Infer CC D E" by (cases \) then have cc: "CC = side_prems_of \" and d: "D = main_prem_of \" and e: "E = concl_of \" unfolding \ by simp_all note e_in_n_un_rf_n = concl_of_in_n_un_rf_n[folded e] { assume "E \ N" moreover have "E < D" using \_reductive e d in_\ by auto ultimately have "set_mset {#E#} \ N" and "\I. I \m {#E#} + CC \ I \ E" and "\D'. D' \# {#E#} \ D' < D" by simp_all then have "redundant_infer N \" using redundant_infer_def cc d e by blast } moreover { assume "E \ Rf N" then obtain DD where dd_sset: "set_mset DD \ N" and dd_imp_e: "\I. I \m DD \ I \ E" and dd_lt_e: "\C'. C' \# DD \ C' < E" unfolding Rf_def by blast from dd_lt_e have "\Da. Da \# DD \ Da < D" using d e in_\ \_reductive less_trans by blast then have "redundant_infer N \" using redundant_infer_def dd_sset dd_imp_e cc d e by blast } ultimately show "\ \ Ri N" using in_\ e_in_n_un_rf_n unfolding Ri_def by blast qed sublocale effective_redundancy_criterion \ Rf Ri unfolding effective_redundancy_criterion_def by (intro conjI redundancy_criterion_axioms, unfold_locales, rule Ri_effective) lemma contradiction_Rf: "{#} \ N \ Ri N = \" unfolding Ri_def redundant_infer_def using \_reductive le_multiset_empty_right by (force intro: exI[of _ "{#{#}#}"] le_multiset_empty_left) end locale standard_redundancy_criterion_counterex_reducing = standard_redundancy_criterion + counterex_reducing_inference_system begin text \ The following result corresponds to Theorem 4.9. \ lemma saturated_upto_complete_if: assumes satur: "saturated_upto N" and unsat: "\ satisfiable N" shows "{#} \ N" proof (rule ccontr) assume ec_ni_n: "{#} \ N" define M where "M = N - Rf N" have ec_ni_m: "{#} \ M" unfolding M_def using ec_ni_n by fast have "I_of M \s M" proof (rule ccontr) assume "\ I_of M \s M" then obtain D where d_in_m: "D \ M" and d_cex: "\ I_of M \ D" and d_min: "\C. C \ M \ C < D \ I_of M \ C" using ex_min_counterex by meson then obtain \ CC E where \: "\ = Infer CC D E" and cc_subs_m: "set_mset CC \ M" and cc_true: "I_of M \m CC" and \_in: "\ \ \" and e_cex: "\ I_of M \ E" and e_lt_d: "E < D" using \_counterex_reducing[OF ec_ni_m] not_less by metis have cc: "CC = side_prems_of \" and d: "D = main_prem_of \" and e: "E = concl_of \" unfolding \ by simp_all have "\ \ Ri N" by (rule subsetD[OF satur[unfolded saturated_upto_def inferences_from_def infer_from_def]]) (simp add: \_in d_in_m cc_subs_m cc[symmetric] d[symmetric] M_def[symmetric]) then have "\ \ Ri M" unfolding M_def using Ri_indep by fast then obtain DD where dd_subs_m: "set_mset DD \ M" and dd_cc_imp_d: "\I. I \m DD + CC \ I \ E" and dd_lt_d: "\C. C \# DD \ C < D" unfolding Ri_def redundant_infer_def cc d e by blast from dd_subs_m dd_lt_d have "I_of M \m DD" using d_min unfolding true_cls_mset_def by (metis contra_subsetD) then have "I_of M \ E" using dd_cc_imp_d cc_true by auto then show False using e_cex by auto qed then have "I_of M \s N" using M_def Rf_model by blast then show False using unsat by blast qed theorem saturated_upto_complete: assumes "saturated_upto N" shows "\ satisfiable N \ {#} \ N" using assms saturated_upto_complete_if true_clss_def by auto end end diff --git a/thys/PAC_Checker/Finite_Map_Multiset.thy b/thys/PAC_Checker/Finite_Map_Multiset.thy --- a/thys/PAC_Checker/Finite_Map_Multiset.thy +++ b/thys/PAC_Checker/Finite_Map_Multiset.thy @@ -1,227 +1,229 @@ (* File: Finite_Map_Multiset.thy Author: Mathias Fleury, Daniela Kaufmann, JKU Maintainer: Mathias Fleury, JKU *) theory Finite_Map_Multiset -imports "HOL-Library.Finite_Map" Duplicate_Free_Multiset +imports + "HOL-Library.Finite_Map" + Nested_Multisets_Ordinals.Duplicate_Free_Multiset begin notation image_mset (infixr "`#" 90) section \Finite maps and multisets\ subsection \Finite sets and multisets\ abbreviation mset_fset :: \'a fset \ 'a multiset\ where \mset_fset N \ mset_set (fset N)\ definition fset_mset :: \'a multiset \ 'a fset\ where \fset_mset N \ Abs_fset (set_mset N)\ lemma fset_mset_mset_fset: \fset_mset (mset_fset N) = N\ by (auto simp: fset.fset_inverse fset_mset_def) lemma mset_fset_fset_mset[simp]: \mset_fset (fset_mset N) = remdups_mset N\ by (auto simp: fset.fset_inverse fset_mset_def Abs_fset_inverse remdups_mset_def) lemma in_mset_fset_fmember[simp]: \x \# mset_fset N \ x |\| N\ by (auto simp: fmember.rep_eq) lemma in_fset_mset_mset[simp]: \x |\| fset_mset N \ x \# N\ by (auto simp: fmember.rep_eq fset_mset_def Abs_fset_inverse) subsection \Finite map and multisets\ text \Roughly the same as \<^term>\ran\ and \<^term>\dom\, but with duplication in the content (unlike their finite sets counterpart) while still working on finite domains (unlike a function mapping). Remark that \<^term>\dom_m\ (the keys) does not contain duplicates, but we keep for symmetry (and for easier use of multiset operators as in the definition of \<^term>\ran_m\). \ definition dom_m where \dom_m N = mset_fset (fmdom N)\ definition ran_m where \ran_m N = the `# fmlookup N `# dom_m N\ lemma dom_m_fmdrop[simp]: \dom_m (fmdrop C N) = remove1_mset C (dom_m N)\ unfolding dom_m_def by (cases \C |\| fmdom N\) (auto simp: mset_set.remove fmember.rep_eq) lemma dom_m_fmdrop_All: \dom_m (fmdrop C N) = removeAll_mset C (dom_m N)\ unfolding dom_m_def by (cases \C |\| fmdom N\) (auto simp: mset_set.remove fmember.rep_eq) lemma dom_m_fmupd[simp]: \dom_m (fmupd k C N) = add_mset k (remove1_mset k (dom_m N))\ unfolding dom_m_def by (cases \k |\| fmdom N\) (auto simp: mset_set.remove fmember.rep_eq mset_set.insert_remove) lemma distinct_mset_dom: \distinct_mset (dom_m N)\ by (simp add: distinct_mset_mset_set dom_m_def) lemma in_dom_m_lookup_iff: \C \# dom_m N' \ fmlookup N' C \ None\ by (auto simp: dom_m_def fmdom.rep_eq fmlookup_dom'_iff) lemma in_dom_in_ran_m[simp]: \i \# dom_m N \ the (fmlookup N i) \# ran_m N\ by (auto simp: ran_m_def) lemma fmupd_same[simp]: \x1 \# dom_m x1aa \ fmupd x1 (the (fmlookup x1aa x1)) x1aa = x1aa\ by (metis fmap_ext fmupd_lookup in_dom_m_lookup_iff option.collapse) lemma ran_m_fmempty[simp]: \ran_m fmempty = {#}\ and dom_m_fmempty[simp]: \dom_m fmempty = {#}\ by (auto simp: ran_m_def dom_m_def) lemma fmrestrict_set_fmupd: \a \ xs \ fmrestrict_set xs (fmupd a C N) = fmupd a C (fmrestrict_set xs N)\ \a \ xs \ fmrestrict_set xs (fmupd a C N) = fmrestrict_set xs N\ by (auto simp: fmfilter_alt_defs) lemma fset_fmdom_fmrestrict_set: \fset (fmdom (fmrestrict_set xs N)) = fset (fmdom N) \ xs\ by (auto simp: fmfilter_alt_defs) lemma dom_m_fmrestrict_set: \dom_m (fmrestrict_set (set xs) N) = mset xs \# dom_m N\ using fset_fmdom_fmrestrict_set[of \set xs\ N] distinct_mset_dom[of N] distinct_mset_inter_remdups_mset[of \mset_fset (fmdom N)\ \mset xs\] by (auto simp: dom_m_def fset_mset_mset_fset finite_mset_set_inter multiset_inter_commute remdups_mset_def) lemma dom_m_fmrestrict_set': \dom_m (fmrestrict_set xs N) = mset_set (xs \ set_mset (dom_m N))\ using fset_fmdom_fmrestrict_set[of \xs\ N] distinct_mset_dom[of N] by (auto simp: dom_m_def fset_mset_mset_fset finite_mset_set_inter multiset_inter_commute remdups_mset_def) lemma indom_mI: \fmlookup m x = Some y \ x \# dom_m m\ by (drule fmdomI) (auto simp: dom_m_def fmember.rep_eq) lemma fmupd_fmdrop_id: assumes \k |\| fmdom N'\ shows \fmupd k (the (fmlookup N' k)) (fmdrop k N') = N'\ proof - have [simp]: \map_upd k (the (fmlookup N' k)) (\x. if x \ k then fmlookup N' x else None) = map_upd k (the (fmlookup N' k)) (fmlookup N')\ by (auto intro!: ext simp: map_upd_def) have [simp]: \map_upd k (the (fmlookup N' k)) (fmlookup N') = fmlookup N'\ using assms by (auto intro!: ext simp: map_upd_def) have [simp]: \finite (dom (\x. if x = k then None else fmlookup N' x))\ by (subst dom_if) auto show ?thesis apply (auto simp: fmupd_def fmupd.abs_eq[symmetric]) unfolding fmlookup_drop apply (simp add: fmlookup_inverse) done qed lemma fm_member_split: \k |\| fmdom N' \ \N'' v. N' = fmupd k v N'' \ the (fmlookup N' k) = v \ k |\| fmdom N''\ by (rule exI[of _ \fmdrop k N'\]) (auto simp: fmupd_fmdrop_id) lemma \fmdrop k (fmupd k va N'') = fmdrop k N''\ by (simp add: fmap_ext) lemma fmap_ext_fmdom: \(fmdom N = fmdom N') \ (\ x. x |\| fmdom N \ fmlookup N x = fmlookup N' x) \ N = N'\ by (rule fmap_ext) (case_tac \x |\| fmdom N\, auto simp: fmdom_notD) lemma fmrestrict_set_insert_in: \xa \ fset (fmdom N) \ fmrestrict_set (insert xa l1) N = fmupd xa (the (fmlookup N xa)) (fmrestrict_set l1 N)\ apply (rule fmap_ext_fmdom) apply (auto simp: fset_fmdom_fmrestrict_set fmember.rep_eq notin_fset; fail)[] apply (auto simp: fmlookup_dom_iff; fail) done lemma fmrestrict_set_insert_notin: \xa \ fset (fmdom N) \ fmrestrict_set (insert xa l1) N = fmrestrict_set l1 N\ by (rule fmap_ext_fmdom) (auto simp: fset_fmdom_fmrestrict_set fmember.rep_eq notin_fset) lemma fmrestrict_set_insert_in_dom_m[simp]: \xa \# dom_m N \ fmrestrict_set (insert xa l1) N = fmupd xa (the (fmlookup N xa)) (fmrestrict_set l1 N)\ by (simp add: fmrestrict_set_insert_in dom_m_def) lemma fmrestrict_set_insert_notin_dom_m[simp]: \xa \# dom_m N \ fmrestrict_set (insert xa l1) N = fmrestrict_set l1 N\ by (simp add: fmrestrict_set_insert_notin dom_m_def) lemma fmlookup_restrict_set_id: \fset (fmdom N) \ A \ fmrestrict_set A N = N\ by (metis fmap_ext fmdom'_alt_def fmdom'_notD fmlookup_restrict_set subset_iff) lemma fmlookup_restrict_set_id': \set_mset (dom_m N) \ A \ fmrestrict_set A N = N\ by (rule fmlookup_restrict_set_id) (auto simp: dom_m_def) lemma ran_m_mapsto_upd: assumes NC: \C \# dom_m N\ shows \ran_m (fmupd C C' N) = add_mset C' (remove1_mset (the (fmlookup N C)) (ran_m N))\ proof - define N' where \N' = fmdrop C N\ have N_N': \dom_m N = add_mset C (dom_m N')\ using NC unfolding N'_def by auto have \C \# dom_m N'\ using NC distinct_mset_dom[of N] unfolding N_N' by auto then show ?thesis by (auto simp: N_N' ran_m_def mset_set.insert_remove image_mset_remove1_mset_if intro!: image_mset_cong) qed lemma ran_m_mapsto_upd_notin: assumes NC: \C \# dom_m N\ shows \ran_m (fmupd C C' N) = add_mset C' (ran_m N)\ using NC by (auto simp: ran_m_def mset_set.insert_remove image_mset_remove1_mset_if intro!: image_mset_cong split: if_splits) lemma image_mset_If_eq_notin: \C \# A \ {#f (if x = C then a x else b x). x \# A#} = {# f(b x). x \# A #}\ by (induction A) auto lemma filter_mset_cong2: "(\x. x \# M \ f x = g x) \ M = N \ filter_mset f M = filter_mset g N" by (hypsubst, rule filter_mset_cong, simp) lemma ran_m_fmdrop: \C \# dom_m N \ ran_m (fmdrop C N) = remove1_mset (the (fmlookup N C)) (ran_m N)\ using distinct_mset_dom[of N] by (cases \fmlookup N C\) (auto simp: ran_m_def image_mset_If_eq_notin[of C _ \\x. fst (the x)\] dest!: multi_member_split intro!: filter_mset_cong2 image_mset_cong2) lemma ran_m_fmdrop_notin: \C \# dom_m N \ ran_m (fmdrop C N) = ran_m N\ using distinct_mset_dom[of N] by (auto simp: ran_m_def image_mset_If_eq_notin[of C _ \\x. fst (the x)\] dest!: multi_member_split intro!: filter_mset_cong2 image_mset_cong2) lemma ran_m_fmdrop_If: \ran_m (fmdrop C N) = (if C \# dom_m N then remove1_mset (the (fmlookup N C)) (ran_m N) else ran_m N)\ using distinct_mset_dom[of N] by (auto simp: ran_m_def image_mset_If_eq_notin[of C _ \\x. fst (the x)\] dest!: multi_member_split intro!: filter_mset_cong2 image_mset_cong2) lemma dom_m_empty_iff[iff]: \dom_m NU = {#} \ NU = fmempty\ by (cases NU) (auto simp: dom_m_def mset_set.insert_remove) end \ No newline at end of file diff --git a/thys/PAC_Checker/PAC_Checker.thy b/thys/PAC_Checker/PAC_Checker.thy --- a/thys/PAC_Checker/PAC_Checker.thy +++ b/thys/PAC_Checker/PAC_Checker.thy @@ -1,1382 +1,1383 @@ (* File: PAC_Checker.thy Author: Mathias Fleury, Daniela Kaufmann, JKU Maintainer: Mathias Fleury, JKU *) theory PAC_Checker imports PAC_Polynomials_Operations PAC_Checker_Specification PAC_Map_Rel Show.Show Show.Show_Instances + PAC_Misc begin section \Executable Checker\ text \In this layer we finally refine the checker to executable code.\ subsection \Definitions\ text \Compared to the previous layer, we add an error message when an error is discovered. We do not attempt to prove anything on the error message (neither that there really is an error, nor that the error message is correct). \ paragraph \Extended error message\ datatype 'a code_status = is_cfailed: CFAILED (the_error: 'a) | CSUCCESS | is_cfound: CFOUND text \In the following function, we merge errors. We will never merge an error message with an another error message; hence we do not attempt to concatenate error messages. \ fun merge_cstatus where \merge_cstatus (CFAILED a) _ = CFAILED a\ | \merge_cstatus _ (CFAILED a) = CFAILED a\ | \merge_cstatus CFOUND _ = CFOUND\ | \merge_cstatus _ CFOUND = CFOUND\ | \merge_cstatus _ _ = CSUCCESS\ definition code_status_status_rel :: \('a code_status \ status) set\ where \code_status_status_rel = {(CFOUND, FOUND), (CSUCCESS, SUCCESS)} \ {(CFAILED a, FAILED)| a. True}\ lemma in_code_status_status_rel_iff[simp]: \(CFOUND, b) \ code_status_status_rel \ b = FOUND\ \(a, FOUND) \ code_status_status_rel \ a = CFOUND\ \(CSUCCESS, b) \ code_status_status_rel \ b = SUCCESS\ \(a, SUCCESS) \ code_status_status_rel \ a = CSUCCESS\ \(a, FAILED) \ code_status_status_rel \ is_cfailed a\ \(CFAILED C, b) \ code_status_status_rel \ b = FAILED\ by (cases a; cases b; auto simp: code_status_status_rel_def; fail)+ paragraph \Refinement relation\ fun pac_step_rel_raw :: \('olbl \ 'lbl) set \ ('a \ 'b) set \ ('c \ 'd) set \ ('a, 'c, 'olbl) pac_step \ ('b, 'd, 'lbl) pac_step \ bool\ where \pac_step_rel_raw R1 R2 R3 (Add p1 p2 i r) (Add p1' p2' i' r') \ (p1, p1') \ R1 \ (p2, p2') \ R1 \ (i, i') \ R1 \ (r, r') \ R2\ | \pac_step_rel_raw R1 R2 R3 (Mult p1 p2 i r) (Mult p1' p2' i' r') \ (p1, p1') \ R1 \ (p2, p2') \ R2 \ (i, i') \ R1 \ (r, r') \ R2\ | \pac_step_rel_raw R1 R2 R3 (Del p1) (Del p1') \ (p1, p1') \ R1\ | \pac_step_rel_raw R1 R2 R3 (Extension i x p1) (Extension j x' p1') \ (i, j) \ R1 \ (x, x') \ R3 \ (p1, p1') \ R2\ | \pac_step_rel_raw R1 R2 R3 _ _ \ False\ fun pac_step_rel_assn :: \('olbl \ 'lbl \ assn) \ ('a \ 'b \ assn) \ ('c \ 'd \ assn) \ ('a, 'c, 'olbl) pac_step \ ('b, 'd, 'lbl) pac_step \ assn\ where \pac_step_rel_assn R1 R2 R3 (Add p1 p2 i r) (Add p1' p2' i' r') = R1 p1 p1' * R1 p2 p2' * R1 i i' * R2 r r'\ | \pac_step_rel_assn R1 R2 R3 (Mult p1 p2 i r) (Mult p1' p2' i' r') = R1 p1 p1' * R2 p2 p2' * R1 i i' * R2 r r'\ | \pac_step_rel_assn R1 R2 R3 (Del p1) (Del p1') = R1 p1 p1'\ | \pac_step_rel_assn R1 R2 R3 (Extension i x p1) (Extension i' x' p1') = R1 i i' * R3 x x' * R2 p1 p1'\ | \pac_step_rel_assn R1 R2 _ _ _ = false\ lemma pac_step_rel_assn_alt_def: \pac_step_rel_assn R1 R2 R3 x y = ( case (x, y) of (Add p1 p2 i r, Add p1' p2' i' r') \ R1 p1 p1' * R1 p2 p2' * R1 i i' * R2 r r' | (Mult p1 p2 i r, Mult p1' p2' i' r') \ R1 p1 p1' * R2 p2 p2' * R1 i i' * R2 r r' | (Del p1, Del p1') \ R1 p1 p1' | (Extension i x p1, Extension i' x' p1') \ R1 i i' * R3 x x' * R2 p1 p1' | _ \ false)\ by (auto split: pac_step.splits) paragraph \Addition checking\ definition error_msg where \error_msg i msg = CFAILED (''s CHECKING failed at line '' @ show i @ '' with error '' @ msg)\ definition error_msg_notin_dom_err where \error_msg_notin_dom_err = '' notin domain''\ definition error_msg_notin_dom :: \nat \ string\ where \error_msg_notin_dom i = show i @ error_msg_notin_dom_err\ definition error_msg_reused_dom where \error_msg_reused_dom i = show i @ '' already in domain''\ definition error_msg_not_equal_dom where \error_msg_not_equal_dom p q pq r = show p @ '' + '' @ show q @ '' = '' @ show pq @ '' not equal'' @ show r\ definition check_not_equal_dom_err :: \llist_polynomial \ llist_polynomial \ llist_polynomial \ llist_polynomial \ string nres\ where \check_not_equal_dom_err p q pq r = SPEC (\_. True)\ definition vars_llist :: \llist_polynomial \ string set\ where \vars_llist xs = \(set ` fst ` set xs)\ definition check_addition_l :: \_ \ _ \ string set \ nat \ nat \ nat \ llist_polynomial \ string code_status nres\ where \check_addition_l spec A \ p q i r = do { let b = p \# dom_m A \ q \# dom_m A \ i \# dom_m A \ vars_llist r \ \; if \b then RETURN (error_msg i ((if p \# dom_m A then error_msg_notin_dom p else []) @ (if q \# dom_m A then error_msg_notin_dom p else []) @ (if i \# dom_m A then error_msg_reused_dom p else []))) else do { ASSERT (p \# dom_m A); let p = the (fmlookup A p); ASSERT (q \# dom_m A); let q = the (fmlookup A q); pq \ add_poly_l (p, q); b \ weak_equality_l pq r; b' \ weak_equality_l r spec; if b then (if b' then RETURN CFOUND else RETURN CSUCCESS) else do { c \ check_not_equal_dom_err p q pq r; RETURN (error_msg i c)} } }\ paragraph \Multiplication checking\ definition check_mult_l_dom_err :: \bool \ nat \ bool \ nat \ string nres\ where \check_mult_l_dom_err p_notin p i_already i = SPEC (\_. True)\ definition check_mult_l_mult_err :: \llist_polynomial \ llist_polynomial \ llist_polynomial \ llist_polynomial \ string nres\ where \check_mult_l_mult_err p q pq r = SPEC (\_. True)\ definition check_mult_l :: \_ \ _ \ _ \ nat \llist_polynomial \ nat \ llist_polynomial \ string code_status nres\ where \check_mult_l spec A \ p q i r = do { let b = p \# dom_m A \ i \# dom_m A \ vars_llist q \ \\ vars_llist r \ \; if \b then do { c \ check_mult_l_dom_err (p \# dom_m A) p (i \# dom_m A) i; RETURN (error_msg i c)} else do { ASSERT (p \# dom_m A); let p = the (fmlookup A p); pq \ mult_poly_full p q; b \ weak_equality_l pq r; b' \ weak_equality_l r spec; if b then (if b' then RETURN CFOUND else RETURN CSUCCESS) else do { c \ check_mult_l_mult_err p q pq r; RETURN (error_msg i c) } } }\ paragraph \Deletion checking\ definition check_del_l :: \_ \ _ \ nat \ string code_status nres\ where \check_del_l spec A p = RETURN CSUCCESS\ paragraph \Extension checking\ definition check_extension_l_dom_err :: \nat \ string nres\ where \check_extension_l_dom_err p = SPEC (\_. True)\ definition check_extension_l_no_new_var_err :: \llist_polynomial \ string nres\ where \check_extension_l_no_new_var_err p = SPEC (\_. True)\ definition check_extension_l_new_var_multiple_err :: \string \ llist_polynomial \ string nres\ where \check_extension_l_new_var_multiple_err v p = SPEC (\_. True)\ definition check_extension_l_side_cond_err :: \string \ llist_polynomial \ llist_polynomial \ llist_polynomial \ string nres\ where \check_extension_l_side_cond_err v p p' q = SPEC (\_. True)\ definition check_extension_l :: \_ \ _ \ string set \ nat \ string \ llist_polynomial \ (string code_status) nres\ where \check_extension_l spec A \ i v p = do { let b = i \# dom_m A \ v \ \ \ ([v], -1) \ set p; if \b then do { c \ check_extension_l_dom_err i; RETURN (error_msg i c) } else do { let p' = remove1 ([v], -1) p; let b = vars_llist p' \ \; if \b then do { c \ check_extension_l_new_var_multiple_err v p'; RETURN (error_msg i c) } else do { p2 \ mult_poly_full p' p'; let p' = map (\(a,b). (a, -b)) p'; q \ add_poly_l (p2, p'); eq \ weak_equality_l q []; if eq then do { RETURN (CSUCCESS) } else do { c \ check_extension_l_side_cond_err v p p' q; RETURN (error_msg i c) } } } }\ lemma check_extension_alt_def: \check_extension A \ i v p \ do { b \ SPEC(\b. b \ i \# dom_m A \ v \ \); if \b then RETURN (False) else do { p' \ RETURN (p + Var v); b \ SPEC(\b. b \ vars p' \ \); if \b then RETURN (False) else do { pq \ mult_poly_spec p' p'; let p' = - p'; p \ add_poly_spec pq p'; eq \ weak_equality p 0; if eq then RETURN(True) else RETURN (False) } } }\ proof - have [intro]: \ab \ \ \ vars ba \ \ \ MPoly_Type.coeff (ba + Var ab) (monomial (Suc 0) ab) = 1\ for ab ba by (subst coeff_add[symmetric], subst not_in_vars_coeff0) (auto simp flip: coeff_add monom.abs_eq simp: not_in_vars_coeff0 MPoly_Type.coeff_def Var.abs_eq Var\<^sub>0_def lookup_single_eq monom.rep_eq) have [simp]: \MPoly_Type.coeff p (monomial (Suc 0) ab) = -1\ if \vars (p + Var ab) \ \\ \ab \ \\ for ab proof - define q where \q \ p + Var ab\ then have p: \p = q - Var ab\ by auto show ?thesis unfolding p apply (subst coeff_minus[symmetric], subst not_in_vars_coeff0) using that unfolding q_def[symmetric] by (auto simp flip: coeff_minus simp: not_in_vars_coeff0 Var.abs_eq Var\<^sub>0_def simp flip: monom.abs_eq simp: not_in_vars_coeff0 MPoly_Type.coeff_def Var.abs_eq Var\<^sub>0_def lookup_single_eq monom.rep_eq) qed have [simp]: \vars (p - Var ab) = vars (Var ab - p)\ for ab using vars_uminus[of \p - Var ab\] by simp show ?thesis unfolding check_extension_def apply (auto 5 5 simp: check_extension_def weak_equality_def mult_poly_spec_def field_simps add_poly_spec_def power2_eq_square cong: if_cong intro!: intro_spec_refine[where R=Id, simplified] split: option.splits dest: ideal.span_add) done qed (* Copy of WB_More_Refinement *) lemma RES_RES_RETURN_RES: \RES A \ (\T. RES (f T)) = RES (\(f ` A))\ by (auto simp: pw_eq_iff refine_pw_simps) lemma check_add_alt_def: \check_add A \ p q i r \ do { b \ SPEC(\b. b \ p \# dom_m A \ q \# dom_m A \ i \# dom_m A \ vars r \ \); if \b then RETURN False else do { ASSERT (p \# dom_m A); let p = the (fmlookup A p); ASSERT (q \# dom_m A); let q = the (fmlookup A q); pq \ add_poly_spec p q; eq \ weak_equality pq r; RETURN eq } }\ (is \_ \ ?A\) proof - have check_add_alt_def: \check_add A \ p q i r = do { b \ SPEC(\b. b \ p \# dom_m A \ q \# dom_m A \ i \# dom_m A \ vars r \ \); if \b then SPEC(\b. b \ p \# dom_m A \ q \# dom_m A \ i \# dom_m A \ vars r \ \ \ the (fmlookup A p) + the (fmlookup A q) - r \ ideal polynomial_bool) else SPEC(\b. b \ p \# dom_m A \ q \# dom_m A \ i \# dom_m A \ vars r \ \ \ the (fmlookup A p) + the (fmlookup A q) - r \ ideal polynomial_bool)}\ (is \_ = ?B\) by (auto simp: check_add_def RES_RES_RETURN_RES) have \?A \ \ Id (check_add A \ p q i r)\ apply refine_vcg apply ((auto simp: check_add_alt_def weak_equality_def add_poly_spec_def RES_RES_RETURN_RES summarize_ASSERT_conv cong: if_cong intro!: ideal.span_zero;fail)+) done then show ?thesis unfolding check_add_alt_def[symmetric] by simp qed lemma check_mult_alt_def: \check_mult A \ p q i r \ do { b \ SPEC(\b. b \ p \# dom_m A \ i \# dom_m A \ vars q \ \ \ vars r \ \); if \b then RETURN False else do { ASSERT (p \# dom_m A); let p = the (fmlookup A p); pq \ mult_poly_spec p q; p \ weak_equality pq r; RETURN p } }\ unfolding check_mult_def apply (rule refine_IdD) by refine_vcg (auto simp: check_mult_def weak_equality_def mult_poly_spec_def RES_RES_RETURN_RES intro!: ideal.span_zero exI[of _ \the (fmlookup A p) * q\]) primrec insort_key_rel :: "('b \ 'b \ bool) \ 'b \ 'b list \ 'b list" where "insort_key_rel f x [] = [x]" | "insort_key_rel f x (y#ys) = (if f x y then (x#y#ys) else y#(insort_key_rel f x ys))" lemma set_insort_key_rel[simp]: \set (insort_key_rel R x xs) = insert x (set xs)\ by (induction xs) auto lemma sorted_wrt_insort_key_rel: \total_on R (insert x (set xs)) \ transp R \ reflp R \ sorted_wrt R xs \ sorted_wrt R (insort_key_rel R x xs)\ by (induction xs) (auto dest: transpD reflpD simp: Restricted_Predicates.total_on_def) lemma sorted_wrt_insort_key_rel2: \total_on R (insert x (set xs)) \ transp R \ x \ set xs \ sorted_wrt R xs \ sorted_wrt R (insort_key_rel R x xs)\ by (induction xs) (auto dest: transpD simp: Restricted_Predicates.total_on_def in_mono) paragraph \Step checking\ definition PAC_checker_l_step :: \_ \ string code_status \ string set \ _ \ (llist_polynomial, string, nat) pac_step \ _\ where \PAC_checker_l_step = (\spec (st', \, A) st. case st of Add _ _ _ _ \ do { r \ full_normalize_poly (pac_res st); eq \ check_addition_l spec A \ (pac_src1 st) (pac_src2 st) (new_id st) r; let _ = eq; if \is_cfailed eq then RETURN (merge_cstatus st' eq, \, fmupd (new_id st) r A) else RETURN (eq, \, A) } | Del _ \ do { eq \ check_del_l spec A (pac_src1 st); let _ = eq; if \is_cfailed eq then RETURN (merge_cstatus st' eq, \, fmdrop (pac_src1 st) A) else RETURN (eq, \, A) } | Mult _ _ _ _ \ do { r \ full_normalize_poly (pac_res st); q \ full_normalize_poly (pac_mult st); eq \ check_mult_l spec A \ (pac_src1 st) q (new_id st) r; let _ = eq; if \is_cfailed eq then RETURN (merge_cstatus st' eq, \, fmupd (new_id st) r A) else RETURN (eq, \, A) } | Extension _ _ _ \ do { r \ full_normalize_poly (([new_var st], -1) # (pac_res st)); (eq) \ check_extension_l spec A \ (new_id st) (new_var st) r; if \is_cfailed eq then do { RETURN (st', insert (new_var st) \, fmupd (new_id st) r A)} else RETURN (eq, \, A) } )\ lemma pac_step_rel_raw_def: \\K, V, R\ pac_step_rel_raw = pac_step_rel_raw K V R\ by (auto intro!: ext simp: relAPP_def) definition mononoms_equal_up_to_reorder where \mononoms_equal_up_to_reorder xs ys \ map (\(a, b). (mset a, b)) xs = map (\(a, b). (mset a, b)) ys\ definition normalize_poly_l where \normalize_poly_l p = SPEC (\p'. normalize_poly_p\<^sup>*\<^sup>* ((\(a, b). (mset a, b)) `# mset p) ((\(a, b). (mset a, b)) `# mset p') \ 0 \# snd `# mset p' \ sorted_wrt (rel2p (term_order_rel \\<^sub>r int_rel)) p' \ (\ x \ mononoms p'. sorted_wrt (rel2p var_order_rel) x))\ definition remap_polys_l_dom_err :: \string nres\ where \remap_polys_l_dom_err = SPEC (\_. True)\ definition remap_polys_l :: \llist_polynomial \ string set \ (nat, llist_polynomial) fmap \ (_ code_status \ string set \ (nat, llist_polynomial) fmap) nres\ where \remap_polys_l spec = (\\ A. do{ dom \ SPEC(\dom. set_mset (dom_m A) \ dom \ finite dom); failed \ SPEC(\_::bool. True); if failed then do { c \ remap_polys_l_dom_err; RETURN (error_msg (0 :: nat) c, \, fmempty) } else do { (b, \, A) \ FOREACH dom (\i (b, \, A'). if i \# dom_m A then do { p \ full_normalize_poly (the (fmlookup A i)); eq \ weak_equality_l p spec; \ \ RETURN(\ \ vars_llist (the (fmlookup A i))); RETURN(b \ eq, \, fmupd i p A') } else RETURN (b, \, A')) (False, \, fmempty); RETURN (if b then CFOUND else CSUCCESS, \, A) }})\ definition PAC_checker_l where \PAC_checker_l spec A b st = do { (S, _) \ WHILE\<^sub>T (\((b, A), n). \is_cfailed b \ n \ []) (\((bA), n). do { ASSERT(n \ []); S \ PAC_checker_l_step spec bA (hd n); RETURN (S, tl n) }) ((b, A), st); RETURN S }\ subsection \Correctness\ text \We now enter the locale to reason about polynomials directly.\ context poly_embed begin abbreviation pac_step_rel where \pac_step_rel \ p2rel (\Id, fully_unsorted_poly_rel O mset_poly_rel, var_rel\ pac_step_rel_raw)\ abbreviation fmap_polys_rel where \fmap_polys_rel \ \nat_rel, sorted_poly_rel O mset_poly_rel\fmap_rel\ lemma \normalize_poly_p s0 s \ (s0, p) \ mset_poly_rel \ (s, p) \ mset_poly_rel\ by (auto simp: mset_poly_rel_def normalize_poly_p_poly_of_mset) lemma vars_poly_of_vars: \vars (poly_of_vars a :: int mpoly) \ (\ ` set_mset a)\ by (induction a) (auto simp: vars_mult_Var) lemma vars_polynomial_of_mset: \vars (polynomial_of_mset za) \ \(image \ ` (set_mset o fst) ` set_mset za)\ apply (induction za) using vars_poly_of_vars by (fastforce elim!: in_vars_addE simp: vars_mult_Const split: if_splits)+ lemma fully_unsorted_poly_rel_vars_subset_vars_llist: \(A, B) \ fully_unsorted_poly_rel O mset_poly_rel \ vars B \ \ ` vars_llist A\ by (auto simp: fully_unsorted_poly_list_rel_def mset_poly_rel_def set_rel_def var_rel_def br_def vars_llist_def list_rel_append2 list_rel_append1 list_rel_split_right_iff list_mset_rel_def image_iff unsorted_term_poly_list_rel_def list_rel_split_left_iff dest!: set_rev_mp[OF _ vars_polynomial_of_mset] split_list dest: multi_member_split dest: arg_cong[of \mset _\ \add_mset _ _\ set_mset]) lemma fully_unsorted_poly_rel_extend_vars: \(A, B) \ fully_unsorted_poly_rel O mset_poly_rel \ (x1c, x1a) \ \var_rel\set_rel \ RETURN (x1c \ vars_llist A) \ \ (\var_rel\set_rel) (SPEC ((\) (x1a \ vars (B))))\ using fully_unsorted_poly_rel_vars_subset_vars_llist[of A B] apply (subst RETURN_RES_refine_iff) apply clarsimp apply (rule exI[of _ \x1a \ \ ` vars_llist A\]) apply (auto simp: set_rel_def var_rel_def br_def dest: fully_unsorted_poly_rel_vars_subset_vars_llist) done lemma remap_polys_l_remap_polys: assumes AB: \(A, B) \ \nat_rel, fully_unsorted_poly_rel O mset_poly_rel\fmap_rel\ and spec: \(spec, spec') \ sorted_poly_rel O mset_poly_rel\ and V: \(\, \') \ \var_rel\set_rel\ shows \remap_polys_l spec \ A \ \(code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel) (remap_polys spec' \' B)\ (is \_ \ \ ?R _\) proof - have 1: \inj_on id (dom :: nat set)\ for dom by auto have H: \x \# dom_m A \ (\p. (the (fmlookup A x), p) \ fully_unsorted_poly_rel \ (p, the (fmlookup B x)) \ mset_poly_rel \ thesis) \ thesis\ for x thesis using fmap_rel_nat_the_fmlookup[OF AB, of x x] fmap_rel_nat_rel_dom_m[OF AB] by auto have full_normalize_poly: \full_normalize_poly (the (fmlookup A x)) \ \ (sorted_poly_rel O mset_poly_rel) (SPEC (\p. the (fmlookup B x') - p \ More_Modules.ideal polynomial_bool \ vars p \ vars (the (fmlookup B x'))))\ if x_dom: \x \# dom_m A\ and \(x, x') \ Id\ for x x' apply (rule H[OF x_dom]) subgoal for p apply (rule full_normalize_poly_normalize_poly_p[THEN order_trans]) apply assumption subgoal using that(2) apply - unfolding conc_fun_chain[symmetric] by (rule ref_two_step', rule RES_refine) (auto simp: rtranclp_normalize_poly_p_poly_of_mset mset_poly_rel_def ideal.span_zero) done done have H': \(p, pa) \ sorted_poly_rel O mset_poly_rel \ weak_equality_l p spec \ SPEC (\eqa. eqa \ pa = spec')\ for p pa using spec by (auto simp: weak_equality_l_def weak_equality_spec_def list_mset_rel_def br_def mset_poly_rel_def dest: list_rel_term_poly_list_rel_same_rightD sorted_poly_list_relD) have emp: \(\, \') \ \var_rel\set_rel \ ((False, \, fmempty), False, \', fmempty) \ bool_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel\ for \ \' by auto show ?thesis using assms unfolding remap_polys_l_def remap_polys_l_dom_err_def remap_polys_def prod.case apply (refine_rcg full_normalize_poly fmap_rel_fmupd_fmap_rel) subgoal by auto subgoal by auto subgoal by (auto simp: error_msg_def) apply (rule 1) subgoal by auto apply (rule emp) subgoal using V by auto subgoal by auto subgoal by auto subgoal by (rule H') apply (rule fully_unsorted_poly_rel_extend_vars) subgoal by (auto intro!: fmap_rel_nat_the_fmlookup) subgoal by (auto intro!: fmap_rel_fmupd_fmap_rel) subgoal by (auto intro!: fmap_rel_fmupd_fmap_rel) subgoal by auto subgoal by auto done qed lemma fref_to_Down_curry: \(uncurry f, uncurry g) \ [P]\<^sub>f A \ \B\nres_rel \ (\x x' y y'. P (x', y') \ ((x, y), (x', y')) \ A \ f x y \ \ B (g x' y'))\ unfolding fref_def uncurry_def nres_rel_def by auto lemma weak_equality_spec_weak_equality: \(p, p') \ mset_poly_rel \ (r, r') \ mset_poly_rel \ weak_equality_spec p r \ weak_equality p' r'\ unfolding weak_equality_spec_def weak_equality_def by (auto simp: mset_poly_rel_def) lemma weak_equality_l_weak_equality_l'[refine]: \weak_equality_l p q \ \ bool_rel (weak_equality p' q')\ if \(p, p') \ sorted_poly_rel O mset_poly_rel\ \(q, q') \ sorted_poly_rel O mset_poly_rel\ for p p' q q' using that by (auto intro!: weak_equality_l_weak_equality_spec[THEN fref_to_Down_curry, THEN order_trans] ref_two_step' weak_equality_spec_weak_equality simp flip: conc_fun_chain) lemma error_msg_ne_SUCCES[iff]: \error_msg i m \ CSUCCESS\ \error_msg i m \ CFOUND\ \is_cfailed (error_msg i m)\ \\is_cfound (error_msg i m)\ by (auto simp: error_msg_def) lemma sorted_poly_rel_vars_llist: \(r, r') \ sorted_poly_rel O mset_poly_rel \ vars r' \ \ ` vars_llist r\ apply (auto simp: mset_poly_rel_def set_rel_def var_rel_def br_def vars_llist_def list_rel_append2 list_rel_append1 list_rel_split_right_iff list_mset_rel_def image_iff sorted_poly_list_rel_wrt_def dest!: set_rev_mp[OF _ vars_polynomial_of_mset] dest!: split_list) apply (auto dest!: multi_member_split simp: list_rel_append1 term_poly_list_rel_def eq_commute[of _ \mset _\] list_rel_split_right_iff list_rel_append2 list_rel_split_left_iff dest: arg_cong[of \mset _\ \add_mset _ _\ set_mset]) done lemma check_addition_l_check_add: assumes \(A, B) \ fmap_polys_rel\ and \(r, r') \ sorted_poly_rel O mset_poly_rel\ \(p, p') \ Id\ \(q, q') \ Id\ \(i, i') \ nat_rel\ \(\', \) \ \var_rel\set_rel\ shows \check_addition_l spec A \' p q i r \ \ {(st, b). (\is_cfailed st \ b) \ (is_cfound st \ spec = r)} (check_add B \ p' q' i' r')\ proof - have [refine]: \add_poly_l (p, q) \ \ (sorted_poly_rel O mset_poly_rel) (add_poly_spec p' q')\ if \(p, p') \ sorted_poly_rel O mset_poly_rel\ \(q, q') \ sorted_poly_rel O mset_poly_rel\ for p p' q q' using that by (auto intro!: add_poly_l_add_poly_p'[THEN order_trans] ref_two_step' add_poly_p'_add_poly_spec simp flip: conc_fun_chain) show ?thesis using assms unfolding check_addition_l_def check_not_equal_dom_err_def apply - apply (rule order_trans) defer apply (rule ref_two_step') apply (rule check_add_alt_def) apply refine_rcg subgoal by (drule sorted_poly_rel_vars_llist) (auto simp: set_rel_def var_rel_def br_def) subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by (auto simp: weak_equality_l_def bind_RES_RETURN_eq) done qed lemma check_del_l_check_del: \(A, B) \ fmap_polys_rel \ (x3, x3a) \ Id \ check_del_l spec A (pac_src1 (Del x3)) \ \ {(st, b). (\is_cfailed st \ b) \ (b \ st = CSUCCESS)} (check_del B (pac_src1 (Del x3a)))\ unfolding check_del_l_def check_del_def by (refine_vcg lhs_step_If RETURN_SPEC_refine) (auto simp: fmap_rel_nat_rel_dom_m bind_RES_RETURN_eq) lemma check_mult_l_check_mult: assumes \(A, B) \ fmap_polys_rel\ and \(r, r') \ sorted_poly_rel O mset_poly_rel\ and \(q, q') \ sorted_poly_rel O mset_poly_rel\ \(p, p') \ Id\ \(i, i') \ nat_rel\ \(\, \') \ \var_rel\set_rel\ shows \check_mult_l spec A \ p q i r \ \ {(st, b). (\is_cfailed st \ b) \ (is_cfound st \ spec = r)} (check_mult B \' p' q' i' r')\ proof - have [refine]: \mult_poly_full p q \ \ (sorted_poly_rel O mset_poly_rel) (mult_poly_spec p' q')\ if \(p, p') \ sorted_poly_rel O mset_poly_rel\ \(q, q') \ sorted_poly_rel O mset_poly_rel\ for p p' q q' using that by (auto intro!: mult_poly_full_mult_poly_p'[THEN order_trans] ref_two_step' mult_poly_p'_mult_poly_spec simp flip: conc_fun_chain) show ?thesis using assms unfolding check_mult_l_def check_mult_l_mult_err_def check_mult_l_dom_err_def apply - apply (rule order_trans) defer apply (rule ref_two_step') apply (rule check_mult_alt_def) apply refine_rcg subgoal by (drule sorted_poly_rel_vars_llist)+ (fastforce simp: set_rel_def var_rel_def br_def) subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by (auto simp: weak_equality_l_def bind_RES_RETURN_eq) done qed lemma normalize_poly_normalize_poly_spec: assumes \(r, t) \ unsorted_poly_rel O mset_poly_rel\ shows \normalize_poly r \ \(sorted_poly_rel O mset_poly_rel) (normalize_poly_spec t)\ proof - obtain s where rs: \(r, s) \ unsorted_poly_rel\ and st: \(s, t) \ mset_poly_rel\ using assms by auto show ?thesis by (rule normalize_poly_normalize_poly_p[THEN order_trans, OF rs]) (use st in \auto dest!: rtranclp_normalize_poly_p_poly_of_mset intro!: ref_two_step' RES_refine exI[of _ t] simp: normalize_poly_spec_def ideal.span_zero mset_poly_rel_def simp flip: conc_fun_chain\) qed lemma remove1_list_rel: \(xs, ys) \ \R\ list_rel \ (a, b) \ R \ IS_RIGHT_UNIQUE R \ IS_LEFT_UNIQUE R \ (remove1 a xs, remove1 b ys) \ \R\list_rel\ by (induction xs ys rule: list_rel_induct) (auto simp: single_valued_def IS_LEFT_UNIQUE_def) lemma remove1_list_rel2: \(xs, ys) \ \R\ list_rel \ (a, b) \ R \ (\c. (a, c) \ R \ c = b) \ (\c. (c, b) \ R \ c = a) \ (remove1 a xs, remove1 b ys) \ \R\list_rel\ apply (induction xs ys rule: list_rel_induct) apply (solves \simp (no_asm)\) by (smt list_rel_simp(4) remove1.simps(2)) lemma remove1_sorted_poly_rel_mset_poly_rel: assumes \(r, r') \ sorted_poly_rel O mset_poly_rel\ and \([a], 1) \ set r\ shows \(remove1 ([a], 1) r, r' - Var (\ a)) \ sorted_poly_rel O mset_poly_rel\ proof - have [simp]: \([a], {#a#}) \ term_poly_list_rel\ \\aa. ([a], aa) \ term_poly_list_rel \ aa = {#a#}\ by (auto simp: term_poly_list_rel_def) have H: \\aa. ([a], aa) \ term_poly_list_rel \ aa = {#a#}\ \\aa. (aa, {#a#}) \ term_poly_list_rel \ aa = [a]\ by (auto simp: single_valued_def IS_LEFT_UNIQUE_def term_poly_list_rel_def) have [simp]: \Const (1 :: int) = (1 :: int mpoly)\ by (simp add: Const.abs_eq Const\<^sub>0_one one_mpoly.abs_eq) have [simp]: \sorted_wrt term_order (map fst r) \ sorted_wrt term_order (map fst (remove1 ([a], 1) r))\ by (induction r) auto have [intro]: \distinct (map fst r) \ distinct (map fst (remove1 x r))\ for x - by (induction r) (auto dest: in_set_remove1D) + by (induction r) (auto dest: notin_set_remove1) have [simp]: \(r, ya) \ \term_poly_list_rel \\<^sub>r int_rel\list_rel \ polynomial_of_mset (mset ya) - Var (\ a) = polynomial_of_mset (remove1_mset ({#a#}, 1) (mset ya))\ for ya using assms by (auto simp: list_rel_append1 list_rel_split_right_iff dest!: split_list) show ?thesis using assms apply (elim relcompEpair) apply (rename_tac za, rule_tac b = \remove1_mset ({#a#}, 1) za\ in relcompI) apply (auto simp: mset_poly_rel_def sorted_poly_list_rel_wrt_def Collect_eq_comp' intro!: relcompI[of _ \remove1 ({#a#}, 1) ya\ for ya :: \(string multiset \ int) list\] remove1_list_rel2 intro: H simp: list_mset_rel_def br_def dest: in_diffD) done qed lemma remove1_sorted_poly_rel_mset_poly_rel_minus: assumes \(r, r') \ sorted_poly_rel O mset_poly_rel\ and \([a], -1) \ set r\ shows \(remove1 ([a], -1) r, r' + Var (\ a)) \ sorted_poly_rel O mset_poly_rel\ proof - have [simp]: \([a], {#a#}) \ term_poly_list_rel\ \\aa. ([a], aa) \ term_poly_list_rel \ aa = {#a#}\ by (auto simp: term_poly_list_rel_def) have H: \\aa. ([a], aa) \ term_poly_list_rel \ aa = {#a#}\ \\aa. (aa, {#a#}) \ term_poly_list_rel \ aa = [a]\ by (auto simp: single_valued_def IS_LEFT_UNIQUE_def term_poly_list_rel_def) have [simp]: \Const (1 :: int) = (1 :: int mpoly)\ by (simp add: Const.abs_eq Const\<^sub>0_one one_mpoly.abs_eq) have [simp]: \sorted_wrt term_order (map fst r) \ sorted_wrt term_order (map fst (remove1 ([a], -1) r))\ by (induction r) auto have [intro]: \distinct (map fst r) \ distinct (map fst (remove1 x r))\ for x apply (induction r) apply auto by (meson img_fst in_set_remove1D) have [simp]: \(r, ya) \ \term_poly_list_rel \\<^sub>r int_rel\list_rel \ polynomial_of_mset (mset ya) + Var (\ a) = polynomial_of_mset (remove1_mset ({#a#}, -1) (mset ya))\ for ya using assms by (auto simp: list_rel_append1 list_rel_split_right_iff dest!: split_list) show ?thesis using assms apply (elim relcompEpair) apply (rename_tac za, rule_tac b = \remove1_mset ({#a#}, -1) za\ in relcompI) by (auto simp: mset_poly_rel_def sorted_poly_list_rel_wrt_def Collect_eq_comp' dest: in_diffD intro!: relcompI[of _ \remove1 ({#a#}, -1) ya\ for ya :: \(string multiset \ int) list\] remove1_list_rel2 intro: H simp: list_mset_rel_def br_def) qed lemma insert_var_rel_set_rel: \(\, \') \ \var_rel\set_rel \ (yb, x2) \ var_rel \ (insert yb \, insert x2 \') \ \var_rel\set_rel\ by (auto simp: var_rel_def set_rel_def) lemma var_rel_set_rel_iff: \(\, \') \ \var_rel\set_rel \ (yb, x2) \ var_rel \ yb \ \ \ x2 \ \'\ using \_inj inj_eq by (fastforce simp: var_rel_def set_rel_def br_def) lemma check_extension_l_check_extension: assumes \(A, B) \ fmap_polys_rel\ and \(r, r') \ sorted_poly_rel O mset_poly_rel\ and \(i, i') \ nat_rel\ \(\, \') \ \var_rel\set_rel\ \(x, x') \ var_rel\ shows \check_extension_l spec A \ i x r \ \{((st), (b)). (\is_cfailed st \ b) \ (is_cfound st \ spec = r)} (check_extension B \' i' x' r')\ proof - have \x' = \ x\ using assms(5) by (auto simp: var_rel_def br_def) have [refine]: \mult_poly_full p q \ \ (sorted_poly_rel O mset_poly_rel) (mult_poly_spec p' q')\ if \(p, p') \ sorted_poly_rel O mset_poly_rel\ \(q, q') \ sorted_poly_rel O mset_poly_rel\ for p p' q q' using that by (auto intro!: mult_poly_full_mult_poly_p'[THEN order_trans] ref_two_step' mult_poly_p'_mult_poly_spec simp flip: conc_fun_chain) have [refine]: \add_poly_l (p, q) \ \ (sorted_poly_rel O mset_poly_rel) (add_poly_spec p' q')\ if \(p, p') \ sorted_poly_rel O mset_poly_rel\ \(q, q') \ sorted_poly_rel O mset_poly_rel\ for p p' q q' using that by (auto intro!: add_poly_l_add_poly_p'[THEN order_trans] ref_two_step' add_poly_p'_add_poly_spec simp flip: conc_fun_chain) have [simp]: \(l, l') \ \term_poly_list_rel \\<^sub>r int_rel\list_rel \ (map (\(a, b). (a, - b)) l, map (\(a, b). (a, - b)) l') \ \term_poly_list_rel \\<^sub>r int_rel\list_rel\ for l l' by (induction l l' rule: list_rel_induct) (auto simp: list_mset_rel_def br_def) have [intro!]: \(x2c, za) \ \term_poly_list_rel \\<^sub>r int_rel\list_rel O list_mset_rel \ (map (\(a, b). (a, - b)) x2c, {#case x of (a, b) \ (a, - b). x \# za#}) \ \term_poly_list_rel \\<^sub>r int_rel\list_rel O list_mset_rel\ for x2c za apply (auto) subgoal for y apply (induction x2c y rule: list_rel_induct) apply (auto simp: list_mset_rel_def br_def) apply (rename_tac a ba aa l l', rule_tac b = \(aa, - ba) # map (\(a, b). (a, - b)) l'\ in relcompI) by auto done have [simp]: \(\x. fst (case x of (a, b) \ (a, - b))) = fst\ by (auto intro: ext) have uminus: \(x2c, x2a) \ sorted_poly_rel O mset_poly_rel \ (map (\(a, b). (a, - b)) x2c, - x2a) \ sorted_poly_rel O mset_poly_rel\ for x2c x2a x1c x1a apply (clarsimp simp: sorted_poly_list_rel_wrt_def mset_poly_rel_def) apply (rule_tac b = \(\(a, b). (a, - b)) `# za\ in relcompI) by (auto simp: sorted_poly_list_rel_wrt_def mset_poly_rel_def comp_def polynomial_of_mset_uminus) have [simp]: \([], 0) \ sorted_poly_rel O mset_poly_rel\ by (auto simp: sorted_poly_list_rel_wrt_def mset_poly_rel_def list_mset_rel_def br_def intro!: relcompI[of _ \{#}\]) show ?thesis unfolding check_extension_l_def check_extension_l_dom_err_def check_extension_l_no_new_var_err_def check_extension_l_new_var_multiple_err_def check_extension_l_side_cond_err_def apply (rule order_trans) defer apply (rule ref_two_step') apply (rule check_extension_alt_def) apply (refine_vcg ) subgoal using assms(1,3,4,5) by (auto simp: var_rel_set_rel_iff) subgoal using assms(1,3,4,5) by (auto simp: var_rel_set_rel_iff) subgoal by auto subgoal by auto apply (subst \x' = \ x\, rule remove1_sorted_poly_rel_mset_poly_rel_minus) subgoal using assms by auto subgoal using assms by auto subgoal using sorted_poly_rel_vars_llist[of \r\ \r'\] assms by (force simp: set_rel_def var_rel_def br_def dest!: sorted_poly_rel_vars_llist) subgoal by auto subgoal by auto subgoal using assms by auto subgoal using assms by auto apply (rule uminus) subgoal using assms by auto subgoal using assms by auto subgoal using assms by auto subgoal using assms by auto subgoal using assms by auto done qed lemma full_normalize_poly_diff_ideal: fixes dom assumes \(p, p') \ fully_unsorted_poly_rel O mset_poly_rel\ shows \full_normalize_poly p \ \ (sorted_poly_rel O mset_poly_rel) (normalize_poly_spec p')\ proof - obtain q where pq: \(p, q) \ fully_unsorted_poly_rel\ and qp':\(q, p') \ mset_poly_rel\ using assms by auto show ?thesis unfolding normalize_poly_spec_def apply (rule full_normalize_poly_normalize_poly_p[THEN order_trans]) apply (rule pq) unfolding conc_fun_chain[symmetric] by (rule ref_two_step', rule RES_refine) (use qp' in \auto dest!: rtranclp_normalize_poly_p_poly_of_mset simp: mset_poly_rel_def ideal.span_zero\) qed lemma insort_key_rel_decomp: \\ys zs. xs = ys @ zs \ insort_key_rel R x xs = ys @ x # zs\ apply (induction xs) subgoal by auto subgoal for a xs by (force intro: exI[of _ \a # _\]) done lemma list_rel_append_same_length: \length xs = length xs' \ (xs @ ys, xs' @ ys') \ \R\list_rel \ (xs, xs') \ \R\list_rel \ (ys, ys') \ \R\list_rel\ by (auto simp: list_rel_def list_all2_append2 dest: list_all2_lengthD) lemma term_poly_list_rel_list_relD: \(ys, cs) \ \term_poly_list_rel \\<^sub>r int_rel\list_rel \ cs = map (\(a, y). (mset a, y)) ys\ by (induction ys arbitrary: cs) (auto simp: term_poly_list_rel_def list_rel_def list_all2_append list_all2_Cons1 list_all2_Cons2) lemma term_poly_list_rel_single: \([x32], {#x32#}) \ term_poly_list_rel\ by (auto simp: term_poly_list_rel_def) lemma unsorted_poly_rel_list_rel_list_rel_uminus: \(map (\(a, b). (a, - b)) r, yc) \ \unsorted_term_poly_list_rel \\<^sub>r int_rel\list_rel \ (r, map (\(a, b). (a, - b)) yc) \ \unsorted_term_poly_list_rel \\<^sub>r int_rel\list_rel\ by (induction r arbitrary: yc) (auto simp: elim!: list_relE3) lemma mset_poly_rel_minus: \({#(a, b)#}, v') \ mset_poly_rel \ (mset yc, r') \ mset_poly_rel \ (r, yc) \ \unsorted_term_poly_list_rel \\<^sub>r int_rel\list_rel \ (add_mset (a, b) (mset yc), v' + r') \ mset_poly_rel\ by (induction r arbitrary: r') (auto simp: mset_poly_rel_def polynomial_of_mset_uminus) lemma fully_unsorted_poly_rel_diff: \([v], v') \ fully_unsorted_poly_rel O mset_poly_rel \ (r, r') \ fully_unsorted_poly_rel O mset_poly_rel \ (v # r, v' + r') \ fully_unsorted_poly_rel O mset_poly_rel\ apply auto apply (rule_tac b = \y + ya\ in relcompI) apply (auto simp: fully_unsorted_poly_list_rel_def list_mset_rel_def br_def) apply (rule_tac b = \yb @ yc\ in relcompI) apply (auto elim!: list_relE3 simp: unsorted_poly_rel_list_rel_list_rel_uminus mset_poly_rel_minus) done lemma PAC_checker_l_step_PAC_checker_step: assumes \(Ast, Bst) \ code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel\ and \(st, st') \ pac_step_rel\ and spec: \(spec, spec') \ sorted_poly_rel O mset_poly_rel\ shows \PAC_checker_l_step spec Ast st \ \ (code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel) (PAC_checker_step spec' Bst st')\ proof - obtain A \ cst B \' cst' where Ast: \Ast = (cst, \, A)\ and Bst: \Bst = (cst', \', B)\ and \[intro]: \(\, \') \ \var_rel\set_rel\ and AB: \(A, B) \ fmap_polys_rel\ \(cst, cst') \ code_status_status_rel\ using assms(1) by (cases Ast; cases Bst; auto) have [refine]: \(r, ra) \ sorted_poly_rel O mset_poly_rel \ (eqa, eqaa) \ {(st, b). (\ is_cfailed st \ b) \ (is_cfound st \ spec = r)} \ RETURN eqa \ \ code_status_status_rel (SPEC (\st'. (\ is_failed st' \ is_found st' \ ra - spec' \ More_Modules.ideal polynomial_bool)))\ for r ra eqa eqaa using spec by (cases eqa) (auto intro!: RETURN_RES_refine dest!: sorted_poly_list_relD simp: mset_poly_rel_def ideal.span_zero) have [simp]: \(eqa, st'a) \ code_status_status_rel \ (merge_cstatus cst eqa, merge_status cst' st'a) \ code_status_status_rel\ for eqa st'a using AB by (cases eqa; cases st'a) (auto simp: code_status_status_rel_def) have [simp]: \(merge_cstatus cst CSUCCESS, cst') \ code_status_status_rel\ using AB by (cases st) (auto simp: code_status_status_rel_def) have [simp]: \(x32, x32a) \ var_rel \ ([([x32], -1::int)], -Var x32a) \ fully_unsorted_poly_rel O mset_poly_rel\ for x32 x32a by (auto simp: mset_poly_rel_def fully_unsorted_poly_list_rel_def list_mset_rel_def br_def unsorted_term_poly_list_rel_def var_rel_def Const_1_eq_1 intro!: relcompI[of _ \{#({#x32#}, -1 :: int)#}\] relcompI[of _ \[({#x32#}, -1)]\]) have H3: \p - Var a = (-Var a) + p\ for p :: \int mpoly\ and a by auto show ?thesis using assms(2) unfolding PAC_checker_l_step_def PAC_checker_step_def Ast Bst prod.case apply (cases st; cases st'; simp only: p2rel_def pac_step.case pac_step_rel_raw_def mem_Collect_eq prod.case pac_step_rel_raw.simps) subgoal apply (refine_rcg normalize_poly_normalize_poly_spec check_mult_l_check_mult check_addition_l_check_add full_normalize_poly_diff_ideal) subgoal using AB by auto subgoal using AB by auto subgoal by auto subgoal by auto subgoal by auto subgoal by (auto intro: \) apply assumption+ subgoal by (auto simp: code_status_status_rel_def) subgoal by (auto intro!: fmap_rel_fmupd_fmap_rel fmap_rel_fmdrop_fmap_rel AB) subgoal using AB by auto done subgoal apply (refine_rcg normalize_poly_normalize_poly_spec check_mult_l_check_mult check_addition_l_check_add full_normalize_poly_diff_ideal[unfolded normalize_poly_spec_def[symmetric]]) subgoal using AB by auto subgoal using AB by auto subgoal using AB by auto subgoal by auto subgoal by auto subgoal by auto apply assumption+ subgoal by (auto simp: code_status_status_rel_def) subgoal by (auto intro!: fmap_rel_fmupd_fmap_rel fmap_rel_fmdrop_fmap_rel AB) subgoal using AB by auto done subgoal apply (refine_rcg full_normalize_poly_diff_ideal check_extension_l_check_extension) subgoal using AB by (auto intro!: fully_unsorted_poly_rel_diff[of _ \-Var _ :: int mpoly\, unfolded H3[symmetric]] simp: comp_def case_prod_beta) subgoal using AB by auto subgoal using AB by auto subgoal by auto subgoal by auto subgoal by (auto simp: code_status_status_rel_def) subgoal by (auto simp: AB intro!: fmap_rel_fmupd_fmap_rel insert_var_rel_set_rel) subgoal by (auto simp: code_status_status_rel_def AB code_status.is_cfailed_def) done subgoal apply (refine_rcg normalize_poly_normalize_poly_spec check_del_l_check_del check_addition_l_check_add full_normalize_poly_diff_ideal[unfolded normalize_poly_spec_def[symmetric]]) subgoal using AB by auto subgoal using AB by auto subgoal by (auto intro!: fmap_rel_fmupd_fmap_rel fmap_rel_fmdrop_fmap_rel code_status_status_rel_def AB) subgoal by (auto intro!: fmap_rel_fmupd_fmap_rel fmap_rel_fmdrop_fmap_rel AB) done done qed lemma code_status_status_rel_discrim_iff: \(x1a, x1c) \ code_status_status_rel \ is_cfailed x1a \ is_failed x1c\ \(x1a, x1c) \ code_status_status_rel \ is_cfound x1a \ is_found x1c\ by (cases x1a; cases x1c; auto; fail)+ lemma PAC_checker_l_PAC_checker: assumes \(A, B) \ \var_rel\set_rel \\<^sub>r fmap_polys_rel\ and \(st, st') \ \pac_step_rel\list_rel\ and \(spec, spec') \ sorted_poly_rel O mset_poly_rel\ and \(b, b') \ code_status_status_rel\ shows \PAC_checker_l spec A b st \ \ (code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel) (PAC_checker spec' B b' st')\ proof - have [refine0]: \(((b, A), st), (b', B), st') \ ((code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel) \\<^sub>r \pac_step_rel\list_rel)\ using assms by (auto simp: code_status_status_rel_def) show ?thesis using assms unfolding PAC_checker_l_def PAC_checker_def apply (refine_rcg PAC_checker_l_step_PAC_checker_step WHILEIT_refine[where R = \((bool_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel) \\<^sub>r \pac_step_rel\list_rel)\]) subgoal by (auto simp: code_status_status_rel_discrim_iff) subgoal by auto subgoal by (auto simp: neq_Nil_conv) subgoal by (auto simp: neq_Nil_conv intro!: param_nth) subgoal by (auto simp: neq_Nil_conv) subgoal by auto done qed end lemma less_than_char_of_char[code_unfold]: \(x, y) \ less_than_char \ (of_char x :: nat) < of_char y\ by (auto simp: less_than_char_def less_char_def) lemmas [code] = add_poly_l'.simps[unfolded var_order_rel_def] export_code add_poly_l' in SML module_name test definition full_checker_l :: \llist_polynomial \ (nat, llist_polynomial) fmap \ (_, string, nat) pac_step list \ (string code_status \ _) nres\ where \full_checker_l spec A st = do { spec' \ full_normalize_poly spec; (b, \, A) \ remap_polys_l spec' {} A; if is_cfailed b then RETURN (b, \, A) else do { let \ = \ \ vars_llist spec; PAC_checker_l spec' (\, A) b st } }\ context poly_embed begin term normalize_poly_spec thm full_normalize_poly_diff_ideal[unfolded normalize_poly_spec_def[symmetric]] abbreviation unsorted_fmap_polys_rel where \unsorted_fmap_polys_rel \ \nat_rel, fully_unsorted_poly_rel O mset_poly_rel\fmap_rel\ lemma full_checker_l_full_checker: assumes \(A, B) \ unsorted_fmap_polys_rel\ and \(st, st') \ \pac_step_rel\list_rel\ and \(spec, spec') \ fully_unsorted_poly_rel O mset_poly_rel\ shows \full_checker_l spec A st \ \ (code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel) (full_checker spec' B st')\ proof - have [refine]: \(spec, spec') \ sorted_poly_rel O mset_poly_rel \ (\, \') \ \var_rel\set_rel \ remap_polys_l spec \ A \ \(code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel) (remap_polys_change_all spec' \' B)\ for spec spec' \ \' apply (rule remap_polys_l_remap_polys[THEN order_trans, OF assms(1)]) apply assumption+ apply (rule ref_two_step[OF order.refl]) apply(rule remap_polys_spec[THEN order_trans]) by (rule remap_polys_polynomial_bool_remap_polys_change_all) show ?thesis unfolding full_checker_l_def full_checker_def apply (refine_rcg remap_polys_l_remap_polys full_normalize_poly_diff_ideal[unfolded normalize_poly_spec_def[symmetric]] PAC_checker_l_PAC_checker) subgoal using assms(3) . subgoal by auto subgoal by (auto simp: is_cfailed_def is_failed_def) subgoal by auto apply (rule fully_unsorted_poly_rel_extend_vars) subgoal using assms(3) . subgoal by auto subgoal by auto subgoal using assms(2) by (auto simp: p2rel_def) subgoal by auto done qed lemma full_checker_l_full_checker': \(uncurry2 full_checker_l, uncurry2 full_checker) \ ((fully_unsorted_poly_rel O mset_poly_rel) \\<^sub>r unsorted_fmap_polys_rel) \\<^sub>r \pac_step_rel\list_rel \\<^sub>f \(code_status_status_rel \\<^sub>r \var_rel\set_rel \\<^sub>r fmap_polys_rel)\nres_rel\ apply (intro frefI nres_relI) using full_checker_l_full_checker by force end definition remap_polys_l2 :: \llist_polynomial \ string set \ (nat, llist_polynomial) fmap \ _ nres\ where \remap_polys_l2 spec = (\\ A. do{ n \ upper_bound_on_dom A; b \ RETURN (n \ 2^64); if b then do { c \ remap_polys_l_dom_err; RETURN (error_msg (0 ::nat) c, \, fmempty) } else do { (b, \, A) \ nfoldli ([0.._. True) (\i (b, \, A'). if i \# dom_m A then do { ASSERT(fmlookup A i \ None); p \ full_normalize_poly (the (fmlookup A i)); eq \ weak_equality_l p spec; \ \ RETURN (\ \ vars_llist (the (fmlookup A i))); RETURN(b \ eq, \, fmupd i p A') } else RETURN (b, \, A') ) (False, \, fmempty); RETURN (if b then CFOUND else CSUCCESS, \, A) } })\ lemma remap_polys_l2_remap_polys_l: \remap_polys_l2 spec \ A \ \ Id (remap_polys_l spec \ A)\ proof - have [refine]: \(A, A') \ Id \ upper_bound_on_dom A \ \ {(n, dom). dom = set [0..dom. set_mset (dom_m A') \ dom \ finite dom))\ for A A' unfolding upper_bound_on_dom_def apply (rule RES_refine) apply (auto simp: upper_bound_on_dom_def) done have 1: \inj_on id dom\ for dom by auto have 2: \x \# dom_m A \ x' \# dom_m A' \ (x, x') \ nat_rel \ (A, A') \ Id \ full_normalize_poly (the (fmlookup A x)) \ \ Id (full_normalize_poly (the (fmlookup A' x')))\ for A A' x x' by (auto) have 3: \(n, dom) \ {(n, dom). dom = set [0.. ([0.. \nat_rel\list_set_rel\ for n dom by (auto simp: list_set_rel_def br_def) have 4: \(p,q) \ Id \ weak_equality_l p spec \ \Id (weak_equality_l q spec)\ for p q spec by auto have 6: \a = b \ (a, b) \ Id\ for a b by auto show ?thesis unfolding remap_polys_l2_def remap_polys_l_def apply (refine_rcg LFO_refine[where R= \Id \\<^sub>r \Id\set_rel \\<^sub>r Id\]) subgoal by auto subgoal by auto subgoal by auto apply (rule 3) subgoal by auto subgoal by (simp add: in_dom_m_lookup_iff) subgoal by (simp add: in_dom_m_lookup_iff) apply (rule 2) subgoal by auto subgoal by auto subgoal by auto subgoal by auto apply (rule 4; assumption) apply (rule 6) subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto done qed end diff --git a/thys/PAC_Checker/PAC_Misc.thy b/thys/PAC_Checker/PAC_Misc.thy new file mode 100644 --- /dev/null +++ b/thys/PAC_Checker/PAC_Misc.thy @@ -0,0 +1,22 @@ +theory PAC_Misc + imports Main +begin + +text "I believe this should be added to the simplifier by default..." +lemma Collect_eq_comp': + "{(x, y). P x y} O {(c, a). c = f a} = {(x, a). P x (f a)}" + by auto + +lemma in_set_conv_iff: + "x \ set (take n xs) \ (\i xs ! i = x)" + by (metis in_set_conv_nth length_take min_less_iff_conj nth_take) + +lemma in_set_take_conv_nth: + "x \ set (take n xs) \ (\i set (remove1 x xs) \ a \ set xs" + by (meson notin_set_remove1) + +end diff --git a/thys/PAC_Checker/ROOT b/thys/PAC_Checker/ROOT --- a/thys/PAC_Checker/ROOT +++ b/thys/PAC_Checker/ROOT @@ -1,38 +1,37 @@ chapter AFP session PAC_Checker (AFP) = "Sepref_IICF" + description \PAC proof checker\ options [timeout = 2700] sessions "HOL-Library" "HOL-Algebra" "Polynomials" Nested_Multisets_Ordinals theories PAC_More_Poly - Duplicate_Free_Multiset Finite_Map_Multiset WB_Sort More_Loops PAC_Specification PAC_Map_Rel PAC_Checker_Specification PAC_Checker_Relation PAC_Polynomials PAC_Polynomials_Term PAC_Polynomials_Operations PAC_Assoc_Map_Rel PAC_Map_Rel PAC_Checker PAC_Checker_Init PAC_Version PAC_Checker_Synthesis theories [condition=ISABELLE_MLTON] PAC_Checker_MLton document_files "root.tex" "root.bib" export_files (in "code") [1] "PAC_Checker.PAC_Checker_MLton:**" diff --git a/thys/PAC_Checker/WB_Sort.thy b/thys/PAC_Checker/WB_Sort.thy --- a/thys/PAC_Checker/WB_Sort.thy +++ b/thys/PAC_Checker/WB_Sort.thy @@ -1,1688 +1,1688 @@ (* File: WB_Sort.thy Author: Mathias Fleury, Daniela Kaufmann, JKU Author: Maximilian Wuttke, Saarland University Maintainer: Mathias Fleury, JKU Correctness proof contributed by Maximilian Wuttke *) theory WB_Sort - imports Refine_Imperative_HOL.IICF "HOL-Library.Rewrite" Duplicate_Free_Multiset + imports + Refine_Imperative_HOL.IICF + "HOL-Library.Rewrite" + Nested_Multisets_Ordinals.Duplicate_Free_Multiset begin text \This a complete copy-paste of the IsaFoL version because sharing is too hard.\ - text \Every element between \<^term>\lo\ and \<^term>\hi\ can be chosen as pivot element.\ definition choose_pivot :: \('b \ 'b \ bool) \ ('a \ 'b) \ 'a list \ nat \ nat \ nat nres\ where \choose_pivot _ _ _ lo hi = SPEC(\k. k \ lo \ k \ hi)\ text \The element at index \p\ partitions the subarray \lo..hi\. This means that every element \ definition isPartition_wrt :: \('b \ 'b \ bool) \ 'b list \ nat \ nat \ nat \ bool\ where \isPartition_wrt R xs lo hi p \ (\ i. i \ lo \ i < p \ R (xs!i) (xs!p)) \ (\ j. j > p \ j \ hi \ R (xs!p) (xs!j))\ lemma isPartition_wrtI: \(\ i. \i \ lo; i < p\ \ R (xs!i) (xs!p)) \ (\ j. \j > p; j \ hi\ \ R (xs!p) (xs!j)) \ isPartition_wrt R xs lo hi p\ by (simp add: isPartition_wrt_def) definition isPartition :: \'a :: order list \ nat \ nat \ nat \ bool\ where \isPartition xs lo hi p \ isPartition_wrt (\) xs lo hi p\ abbreviation isPartition_map :: \('b \ 'b \ bool) \ ('a \ 'b) \ 'a list \ nat \ nat \ nat \ bool\ where \isPartition_map R h xs i j k \ isPartition_wrt (\a b. R (h a) (h b)) xs i j k\ lemma isPartition_map_def': \lo \ p \ p \ hi \ hi < length xs \ isPartition_map R h xs lo hi p = isPartition_wrt R (map h xs) lo hi p\ by (auto simp add: isPartition_wrt_def conjI) text \Example: 6 is the pivot element (with index 4); \<^term>\7\ is equal to the \<^term>\length xs - 1\.\ lemma \isPartition [0,5,3,4,6,9,8,10::nat] 0 7 4\ by (auto simp add: isPartition_def isPartition_wrt_def nth_Cons') definition sublist :: \'a list \ nat \ nat \ 'a list\ where \sublist xs i j \ take (Suc j - i) (drop i xs)\ (*take from HashMap *) lemma take_Suc0: "l\[] \ take (Suc 0) l = [l!0]" "0 < length l \ take (Suc 0) l = [l!0]" "Suc n \ length l \ take (Suc 0) l = [l!0]" by (cases l, auto)+ lemma sublist_single: \i < length xs \ sublist xs i i = [xs!i]\ by (cases xs) (auto simp add: sublist_def take_Suc0) lemma insert_eq: \insert a b = b \ {a}\ by auto lemma sublist_nth: \\lo \ hi; hi < length xs; k+lo \ hi\ \ (sublist xs lo hi)!k = xs!(lo+k)\ by (simp add: sublist_def) lemma sublist_length: \\i \ j; j < length xs\ \ length (sublist xs i j) = 1 + j - i\ by (simp add: sublist_def) lemma sublist_not_empty: \\i \ j; j < length xs; xs \ []\ \ sublist xs i j \ []\ apply simp apply (rewrite List.length_greater_0_conv[symmetric]) apply (rewrite sublist_length) by auto lemma sublist_app: \\i1 \ i2; i2 \ i3\ \ sublist xs i1 i2 @ sublist xs (Suc i2) i3 = sublist xs i1 i3\ unfolding sublist_def by (smt Suc_eq_plus1_left Suc_le_mono append.assoc le_SucI le_add_diff_inverse le_trans same_append_eq take_add) definition sorted_sublist_wrt :: \('b \ 'b \ bool) \ 'b list \ nat \ nat \ bool\ where \sorted_sublist_wrt R xs lo hi = sorted_wrt R (sublist xs lo hi)\ definition sorted_sublist :: \'a :: linorder list \ nat \ nat \ bool\ where \sorted_sublist xs lo hi = sorted_sublist_wrt (\) xs lo hi\ abbreviation sorted_sublist_map :: \('b \ 'b \ bool) \ ('a \ 'b) \ 'a list \ nat \ nat \ bool\ where \sorted_sublist_map R h xs lo hi \ sorted_sublist_wrt (\a b. R (h a) (h b)) xs lo hi\ lemma sorted_sublist_map_def': \lo < length xs \ sorted_sublist_map R h xs lo hi \ sorted_sublist_wrt R (map h xs) lo hi\ apply (simp add: sorted_sublist_wrt_def) by (simp add: drop_map sorted_wrt_map sublist_def take_map) lemma sorted_sublist_wrt_refl: \i < length xs \ sorted_sublist_wrt R xs i i\ by (auto simp add: sorted_sublist_wrt_def sublist_single) lemma sorted_sublist_refl: \i < length xs \ sorted_sublist xs i i\ by (auto simp add: sorted_sublist_def sorted_sublist_wrt_refl) lemma sublist_map: \sublist (map f xs) i j = map f (sublist xs i j)\ apply (auto simp add: sublist_def) by (simp add: drop_map take_map) lemma take_set: \j \ length xs \ x \ set (take j xs) \ (\ k. k < j \ xs!k = x)\ - apply (induction xs) - apply simp - by (meson in_set_conv_iff less_le_trans) + by (rule eq_reflection) (auto simp add: take_set) lemma drop_set: \j \ length xs \ x \ set (drop j xs) \ (\k. j\k\k xs!k=x)\ by (smt Misc.in_set_drop_conv_nth) (* lemma found by sledgehammer *) lemma sublist_el: \i \ j \ j < length xs \ x \ set (sublist xs i j) \ (\ k. k < Suc j-i \ xs!(i+k)=x)\ by (auto simp add: take_set sublist_def) lemma sublist_el': \i \ j \ j < length xs \ x \ set (sublist xs i j) \ (\ k. i\k\k\j \ xs!k=x)\ apply (subst sublist_el, assumption, assumption) by (smt Groups.add_ac(2) le_add1 le_add_diff_inverse less_Suc_eq less_diff_conv nat_less_le order_refl) lemma sublist_lt: \hi < lo \ sublist xs lo hi = []\ by (auto simp add: sublist_def) lemma nat_le_eq_or_lt: \(a :: nat) \ b = (a = b \ a < b)\ by linarith lemma sorted_sublist_wrt_le: \hi \ lo \ hi < length xs \ sorted_sublist_wrt R xs lo hi\ apply (auto simp add: nat_le_eq_or_lt) unfolding sorted_sublist_wrt_def subgoal apply (rewrite sublist_single) by auto subgoal by (auto simp add: sublist_lt) done text \Elements in a sorted sublists are actually sorted\ lemma sorted_sublist_wrt_nth_le: assumes \sorted_sublist_wrt R xs lo hi\ and \lo \ hi\ and \hi < length xs\ and \lo \ i\ and \i < j\ and \j \ hi\ shows \R (xs!i) (xs!j)\ proof - have A: \lo < length xs\ using assms(2) assms(3) by linarith obtain i' where I: \i = lo + i'\ using assms(4) le_Suc_ex by auto obtain j' where J: \j = lo + j'\ by (meson assms(4) assms(5) dual_order.trans le_iff_add less_imp_le_nat) show ?thesis using assms(1) apply (simp add: sorted_sublist_wrt_def I J) apply (rewrite sublist_nth[symmetric, where k=i', where lo=lo, where hi=hi]) using assms apply auto subgoal using I by linarith apply (rewrite sublist_nth[symmetric, where k=j', where lo=lo, where hi=hi]) using assms apply auto subgoal using J by linarith apply (rule sorted_wrt_nth_less) apply auto subgoal using I J nat_add_left_cancel_less by blast subgoal apply (simp add: sublist_length) using J by linarith done qed text \We can make the assumption \<^term>\i < j\ weaker if we have a reflexivie relation.\ lemma sorted_sublist_wrt_nth_le': assumes ref: \\ x. R x x\ and \sorted_sublist_wrt R xs lo hi\ and \lo \ hi\ and \hi < length xs\ and \lo \ i\ and \i \ j\ and \j \ hi\ shows \R (xs!i) (xs!j)\ proof - have \i < j \ i = j\ using \i \ j\ by linarith then consider (a) \i < j\ | (b) \i = j\ by blast then show ?thesis proof cases case a then show ?thesis using assms(2-5,7) sorted_sublist_wrt_nth_le by blast next case b then show ?thesis by (simp add: ref) qed qed (* lemma sorted_sublist_map_nth_le: assumes \sorted_sublist_map R h xs lo hi\ and \lo \ hi\ and \hi < length xs\ and \lo \ i\ and \i < j\ and \j \ hi\ shows \R (h (xs!i)) (h (xs!j))\ proof - show ?thesis using assms by (rule sorted_sublist_wrt_nth_le) qed *) lemma sorted_sublist_le: \hi \ lo \ hi < length xs \ sorted_sublist xs lo hi\ by (auto simp add: sorted_sublist_def sorted_sublist_wrt_le) lemma sorted_sublist_map_le: \hi \ lo \ hi < length xs \ sorted_sublist_map R h xs lo hi\ by (auto simp add: sorted_sublist_wrt_le) lemma sublist_cons: \lo < hi \ hi < length xs \ sublist xs lo hi = xs!lo # sublist xs (Suc lo) hi\ by (metis Cons_eq_appendI append_self_conv2 less_imp_le_nat less_or_eq_imp_le less_trans sublist_app sublist_single) lemma sorted_sublist_wrt_cons': \sorted_sublist_wrt R xs (lo+1) hi \ lo \ hi \ hi < length xs \ (\j. loj\hi \ R (xs!lo) (xs!j)) \ sorted_sublist_wrt R xs lo hi\ apply (auto simp add: nat_le_eq_or_lt sorted_sublist_wrt_def) apply (auto 5 4 simp add: sublist_cons sublist_el less_diff_conv add.commute[of _ lo] dest: Suc_lessI sublist_single) done lemma sorted_sublist_wrt_cons: assumes trans: \(\ x y z. \R x y; R y z\ \ R x z)\ and \sorted_sublist_wrt R xs (lo+1) hi\ and \lo \ hi\ and \hi < length xs \ and \R (xs!lo) (xs!(lo+1))\ shows \sorted_sublist_wrt R xs lo hi\ proof - show ?thesis apply (rule sorted_sublist_wrt_cons') using assms apply auto subgoal premises assms' for j proof - have A: \j=lo+1 \ j>lo+1\ using assms'(5) by linarith show ?thesis using A proof assume A: \j=lo+1\ show ?thesis by (simp add: A assms') next assume A: \j>lo+1\ show ?thesis apply (rule trans) apply (rule assms(5)) apply (rule sorted_sublist_wrt_nth_le[OF assms(2), where i=\lo+1\, where j=j]) subgoal using A assms'(6) by linarith subgoal using assms'(3) less_imp_diff_less by blast subgoal using assms'(5) by auto subgoal using A by linarith subgoal by (simp add: assms'(6)) done qed qed done qed lemma sorted_sublist_map_cons: \(\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)) \ sorted_sublist_map R h xs (lo+1) hi \ lo \ hi \ hi < length xs \ R (h (xs!lo)) (h (xs!(lo+1))) \ sorted_sublist_map R h xs lo hi\ by (blast intro: sorted_sublist_wrt_cons) lemma sublist_snoc: \lo < hi \ hi < length xs \ sublist xs lo hi = sublist xs lo (hi-1) @ [xs!hi]\ apply (simp add: sublist_def) proof - assume a1: "lo < hi" assume "hi < length xs" then have "take lo xs @ take (Suc hi - lo) (drop lo xs) = (take lo xs @ take (hi - lo) (drop lo xs)) @ [xs ! hi]" using a1 by (metis (no_types) Suc_diff_le add_Suc_right hd_drop_conv_nth le_add_diff_inverse less_imp_le_nat take_add take_hd_drop) then show "take (Suc hi - lo) (drop lo xs) = take (hi - lo) (drop lo xs) @ [xs ! hi]" by simp qed lemma sorted_sublist_wrt_snoc': \sorted_sublist_wrt R xs lo (hi-1) \ lo \ hi \ hi < length xs \ (\j. lo\j\j R (xs!j) (xs!hi)) \ sorted_sublist_wrt R xs lo hi\ apply (simp add: sorted_sublist_wrt_def) apply (auto simp add: nat_le_eq_or_lt) subgoal by (simp add: sublist_single) by (auto simp add: sublist_snoc sublist_el sorted_wrt_append add.commute[of lo] less_diff_conv simp: leI simp flip:nat_le_eq_or_lt) lemma sorted_sublist_wrt_snoc: assumes trans: \(\ x y z. \R x y; R y z\ \ R x z)\ and \sorted_sublist_wrt R xs lo (hi-1)\ and \lo \ hi\ and \hi < length xs\ and \(R (xs!(hi-1)) (xs!hi))\ shows \sorted_sublist_wrt R xs lo hi\ proof - show ?thesis apply (rule sorted_sublist_wrt_snoc') using assms apply auto subgoal premises assms' for j proof - have A: \j=hi-1 \ j using assms'(6) by linarith show ?thesis using A proof assume A: \j=hi-1\ show ?thesis by (simp add: A assms') next assume A: \j show ?thesis apply (rule trans) apply (rule sorted_sublist_wrt_nth_le[OF assms(2), where i=j, where j=\hi-1\]) prefer 6 apply (rule assms(5)) apply auto subgoal using A assms'(5) by linarith subgoal using assms'(3) less_imp_diff_less by blast subgoal using assms'(5) by auto subgoal using A by linarith done qed qed done qed lemma sublist_split: \lo \ hi \ lo < p \ p < hi \ hi < length xs \ sublist xs lo p @ sublist xs (p+1) hi = sublist xs lo hi\ by (simp add: sublist_app) lemma sublist_split_part: \lo \ hi \ lo < p \ p < hi \ hi < length xs \ sublist xs lo (p-1) @ xs!p # sublist xs (p+1) hi = sublist xs lo hi\ by (auto simp add: sublist_split[symmetric] sublist_snoc[where xs=xs,where lo=lo,where hi=p]) text \A property for partitions (we always assume that \<^term>\R\ is transitive.\ lemma isPartition_wrt_trans: \(\ x y z. \R x y; R y z\ \ R x z) \ isPartition_wrt R xs lo hi p \ (\i j. lo \ i \ i < p \ p < j \ j \ hi \ R (xs!i) (xs!j))\ by (auto simp add: isPartition_wrt_def) lemma isPartition_map_trans: \(\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)) \ hi < length xs \ isPartition_map R h xs lo hi p \ (\i j. lo \ i \ i < p \ p < j \ j \ hi \ R (h (xs!i)) (h (xs!j)))\ by (auto simp add: isPartition_wrt_def) lemma merge_sorted_wrt_partitions_between': \lo \ hi \ lo < p \ p < hi \ hi < length xs \ isPartition_wrt R xs lo hi p \ sorted_sublist_wrt R xs lo (p-1) \ sorted_sublist_wrt R xs (p+1) hi \ (\i j. lo \ i \ i < p \ p < j \ j \ hi \ R (xs!i) (xs!j)) \ sorted_sublist_wrt R xs lo hi\ apply (auto simp add: isPartition_def isPartition_wrt_def sorted_sublist_def sorted_sublist_wrt_def sublist_map) apply (simp add: sublist_split_part[symmetric]) apply (auto simp add: List.sorted_wrt_append) subgoal by (auto simp add: sublist_el) subgoal by (auto simp add: sublist_el) subgoal by (auto simp add: sublist_el') done lemma merge_sorted_wrt_partitions_between: \(\ x y z. \R x y; R y z\ \ R x z) \ isPartition_wrt R xs lo hi p \ sorted_sublist_wrt R xs lo (p-1) \ sorted_sublist_wrt R xs (p+1) hi \ lo \ hi \ hi < length xs \ lo < p \ p < hi \ hi < length xs \ sorted_sublist_wrt R xs lo hi\ by (simp add: merge_sorted_wrt_partitions_between' isPartition_wrt_trans) (* lemma merge_sorted_map_partitions_between: \(\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)) \ isPartition_map R h xs lo hi p \ sorted_sublist_map R h xs lo (p-1) \ sorted_sublist_map R h xs (p+1) hi \ lo \ hi \ hi < length xs \ lo < p \ p < hi \ hi < length xs \ sorted_sublist_map R h xs lo hi\ by (simp add: merge_sorted_wrt_partitions_between' isPartition_map_trans) *) text \The main theorem to merge sorted lists\ lemma merge_sorted_wrt_partitions: \isPartition_wrt R xs lo hi p \ sorted_sublist_wrt R xs lo (p - Suc 0) \ sorted_sublist_wrt R xs (Suc p) hi \ lo \ hi \ lo \ p \ p \ hi \ hi < length xs \ (\i j. lo \ i \ i < p \ p < j \ j \ hi \ R (xs!i) (xs!j)) \ sorted_sublist_wrt R xs lo hi\ subgoal premises assms proof - have C: \lo=p\p=hi \ lo=p\p lop=hi \ lop using assms by linarith show ?thesis using C apply auto subgoal \ \lo=p=hi\ apply (rule sorted_sublist_wrt_refl) using assms by auto subgoal \ \lo=p using assms by (simp add: isPartition_def isPartition_wrt_def sorted_sublist_wrt_cons') subgoal \ \lo using assms by (simp add: isPartition_def isPartition_wrt_def sorted_sublist_wrt_snoc') subgoal \ \lo using assms apply (rewrite merge_sorted_wrt_partitions_between'[where p=p]) by auto done qed done theorem merge_sorted_map_partitions: \(\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)) \ isPartition_map R h xs lo hi p \ sorted_sublist_map R h xs lo (p - Suc 0) \ sorted_sublist_map R h xs (Suc p) hi \ lo \ hi \ lo \ p \ p \ hi \ hi < length xs \ sorted_sublist_map R h xs lo hi\ apply (rule merge_sorted_wrt_partitions) apply auto by (simp add: merge_sorted_wrt_partitions isPartition_map_trans) lemma partition_wrt_extend: \isPartition_wrt R xs lo' hi' p \ hi < length xs \ lo \ lo' \ lo' \ hi \ hi' \ hi \ lo' \ p \ p \ hi' \ (\ i. lo\i \ i R (xs!i) (xs!p)) \ (\ j. hi' j\hi \ R (xs!p) (xs!j)) \ isPartition_wrt R xs lo hi p\ unfolding isPartition_wrt_def apply (intro conjI) subgoal by (force simp: not_le) subgoal using leI by blast done lemma partition_map_extend: \isPartition_map R h xs lo' hi' p \ hi < length xs \ lo \ lo' \ lo' \ hi \ hi' \ hi \ lo' \ p \ p \ hi' \ (\ i. lo\i \ i R (h (xs!i)) (h (xs!p))) \ (\ j. hi' j\hi \ R (h (xs!p)) (h (xs!j))) \ isPartition_map R h xs lo hi p\ by (auto simp add: partition_wrt_extend) lemma isPartition_empty: \(\ j. \lo < j; j \ hi\ \ R (xs ! lo) (xs ! j)) \ isPartition_wrt R xs lo hi lo\ by (auto simp add: isPartition_wrt_def) lemma take_ext: \(\i k < length xs \ k < length xs' \ take k xs' = take k xs\ by (simp add: nth_take_lemma) lemma drop_ext': \(\i. i\k \ i xs'!i=xs!i) \ 0 xs\[] \ \ \These corner cases will be dealt with in the next lemma\ length xs'=length xs \ drop k xs' = drop k xs\ apply (rewrite in \drop _ \ = _\ List.rev_rev_ident[symmetric]) apply (rewrite in \_ = drop _ \\ List.rev_rev_ident[symmetric]) apply (rewrite in \\ = _\ List.drop_rev) apply (rewrite in \_ = \\ List.drop_rev) apply simp apply (rule take_ext) by (auto simp add: rev_nth) lemma drop_ext: \(\i. i\k \ i xs'!i=xs!i) \ length xs'=length xs \ drop k xs' = drop k xs\ apply (cases xs) apply auto apply (cases k) subgoal by (simp add: nth_equalityI) subgoal apply (rule drop_ext') by auto done lemma sublist_ext': \(\i. lo\i\i\hi \ xs'!i=xs!i) \ length xs' = length xs \ lo \ hi \ Suc hi < length xs \ sublist xs' lo hi = sublist xs lo hi\ apply (simp add: sublist_def) apply (rule take_ext) by auto lemma lt_Suc: \(a < b) = (Suc a = b \ Suc a < b)\ by auto lemma sublist_until_end_eq_drop: \Suc hi = length xs \ sublist xs lo hi = drop lo xs\ by (simp add: sublist_def) lemma sublist_ext: \(\i. lo\i\i\hi \ xs'!i=xs!i) \ length xs' = length xs \ lo \ hi \ hi < length xs \ sublist xs' lo hi = sublist xs lo hi\ apply (auto simp add: lt_Suc[where a=hi]) subgoal by (auto simp add: sublist_until_end_eq_drop drop_ext) subgoal by (auto simp add: sublist_ext') done lemma sorted_wrt_lower_sublist_still_sorted: assumes \sorted_sublist_wrt R xs lo (lo' - Suc 0)\ and \lo \ lo'\ and \lo' < length xs\ and \(\ i. lo\i\i xs'!i=xs!i)\ and \length xs' = length xs\ shows \sorted_sublist_wrt R xs' lo (lo' - Suc 0)\ proof - have l: \lo < lo' - 1 \ lo \ lo'-1\ by linarith show ?thesis using l apply auto subgoal \ \lo < lo' - 1\ apply (auto simp add: sorted_sublist_wrt_def) apply (rewrite sublist_ext[where xs=xs]) using assms by (auto simp add: sorted_sublist_wrt_def) subgoal \ \lo >= lo' - 1\ using assms by (auto simp add: sorted_sublist_wrt_le) done qed lemma sorted_map_lower_sublist_still_sorted: assumes \sorted_sublist_map R h xs lo (lo' - Suc 0)\ and \lo \ lo'\ and \lo' < length xs\ and \(\ i. lo\i\i xs'!i=xs!i)\ and \length xs' = length xs\ shows \sorted_sublist_map R h xs' lo (lo' - Suc 0)\ using assms by (rule sorted_wrt_lower_sublist_still_sorted) lemma sorted_wrt_upper_sublist_still_sorted: assumes \sorted_sublist_wrt R xs (hi'+1) hi\ and \lo \ lo'\ and \hi < length xs\ and \\ j. hi'j\hi \ xs'!j=xs!j\ and \length xs' = length xs\ shows \sorted_sublist_wrt R xs' (hi'+1) hi\ proof - have l: \hi' + 1 < hi \ hi' + 1 \ hi\ by linarith show ?thesis using l apply auto subgoal \ \hi' + 1 < h\ apply (auto simp add: sorted_sublist_wrt_def) apply (rewrite sublist_ext[where xs=xs]) using assms by (auto simp add: sorted_sublist_wrt_def) subgoal \ \\<^term>\hi' + 1 \ hi\\ using assms by (auto simp add: sorted_sublist_wrt_le) done qed lemma sorted_map_upper_sublist_still_sorted: assumes \sorted_sublist_map R h xs (hi'+1) hi\ and \lo \ lo'\ and \hi < length xs\ and \\ j. hi'j\hi \ xs'!j=xs!j\ and \length xs' = length xs\ shows \sorted_sublist_map R h xs' (hi'+1) hi\ using assms by (rule sorted_wrt_upper_sublist_still_sorted) text \The specification of the partition function\ definition partition_spec :: \('b \ 'b \ bool) \ ('a \ 'b) \ 'a list \ nat \ nat \ 'a list \ nat \ bool\ where \partition_spec R h xs lo hi xs' p \ mset xs' = mset xs \ \ \The list is a permutation\ isPartition_map R h xs' lo hi p \ \ \We have a valid partition on the resulting list\ lo \ p \ p \ hi \ \ \The partition index is in bounds\ (\ i. i xs'!i=xs!i) \ (\ i. hii xs'!i=xs!i)\ \ \Everything else is unchanged.\ lemma in_set_take_conv_nth: \x \ set (take n xs) \ (\m by (metis in_set_conv_nth length_take min.commute min.strict_boundedE nth_take) lemma mset_drop_upto: \mset (drop a N) = {#N!i. i \# mset_set {a.. proof (induction N arbitrary: a) case Nil then show ?case by simp next case (Cons c N) have upt: \{0.. by auto then have H: \mset_set {0.. unfolding upt by auto have mset_case_Suc: \{#case x of 0 \ c | Suc x \ N ! x . x \# mset_set {Suc a..# mset_set {Suc a.. for a b by (rule image_mset_cong) (auto split: nat.splits) have Suc_Suc: \{Suc a.. for a b by auto then have mset_set_Suc_Suc: \mset_set {Suc a..# mset_set {a.. for a b unfolding Suc_Suc by (subst image_mset_mset_set[symmetric]) auto have *: \{#N ! (x-Suc 0) . x \# mset_set {Suc a..# mset_set {a.. for a b by (auto simp add: mset_set_Suc_Suc multiset.map_comp comp_def) show ?case apply (cases a) using Cons[of 0] Cons by (auto simp: nth_Cons drop_Cons H mset_case_Suc *) qed (* Actually, I only need that \set (sublist xs' lo hi) = set (sublist xs lo hi)\ *) lemma mathias: assumes Perm: \mset xs' = mset xs\ and I: \lo\i\ \i\hi\ \xs'!i=x\ and Bounds: \hi < length xs\ and Fix: \\ i. i xs'!i = xs!i\ \\ j. \hi \ xs'!j = xs!j\ shows \\j. lo\j\j\hi \ xs!j = x\ proof - define xs1 xs2 xs3 xs1' xs2' xs3' where \xs1 = take lo xs\ and \xs2 = take (Suc hi - lo) (drop lo xs)\ and \xs3 = drop (Suc hi) xs\ and \xs1' = take lo xs'\ and \xs2' = take (Suc hi - lo) (drop lo xs')\ and \xs3' = drop (Suc hi) xs'\ have [simp]: \length xs' = length xs\ using Perm by (auto dest: mset_eq_length) have [simp]: \mset xs1 = mset xs1'\ using Fix(1) unfolding xs1_def xs1'_def by (metis Perm le_cases mset_eq_length nth_take_lemma take_all) have [simp]: \mset xs3 = mset xs3'\ using Fix(2) unfolding xs3_def xs3'_def mset_drop_upto by (auto intro: image_mset_cong) have \xs = xs1 @ xs2 @ xs3\ \xs' = xs1' @ xs2' @ xs3'\ using I unfolding xs1_def xs2_def xs3_def xs1'_def xs2'_def xs3'_def by (metis append.assoc append_take_drop_id le_SucI le_add_diff_inverse order_trans take_add)+ moreover have \xs ! i = xs2 ! (i - lo)\ \i \ length xs1\ using I Bounds unfolding xs2_def xs1_def by (auto simp: nth_take min_def) moreover have \x \ set xs2'\ using I Bounds unfolding xs2'_def by (auto simp: in_set_take_conv_nth intro!: exI[of _ \i - lo\]) ultimately have \x \ set xs2\ using Perm I by (auto dest: mset_eq_setD) then obtain j where \xs ! (lo + j) = x\ \j \ hi - lo\ unfolding in_set_conv_nth xs2_def by auto then show ?thesis using Bounds I by (auto intro: exI[of _ \lo+j\]) qed text \If we fix the left and right rest of two permutated lists, then the sublists are also permutations.\ text \But we only need that the sets are equal.\ lemma mset_sublist_incl: assumes Perm: \mset xs' = mset xs\ and Fix: \\ i. i xs'!i = xs!i\ \\ j. \hi \ xs'!j = xs!j\ and bounds: \lo \ hi\ \hi < length xs\ shows \set (sublist xs' lo hi) \ set (sublist xs lo hi)\ proof fix x assume \x \ set (sublist xs' lo hi)\ then have \\i. lo\i\i\hi \ xs'!i=x\ by (metis assms(1) bounds(1) bounds(2) size_mset sublist_el') then obtain i where I: \lo\i\ \i\hi\ \xs'!i=x\ by blast have \\j. lo\j\j\hi \ xs!j=x\ using Perm I bounds(2) Fix by (rule mathias, auto) then show \x \ set (sublist xs lo hi)\ by (simp add: bounds(1) bounds(2) sublist_el') qed lemma mset_sublist_eq: assumes \mset xs' = mset xs\ and \\ i. i xs'!i = xs!i\ and \\ j. \hi \ xs'!j = xs!j\ and bounds: \lo \ hi\ \hi < length xs\ shows \set (sublist xs' lo hi) = set (sublist xs lo hi)\ proof show \set (sublist xs' lo hi) \ set (sublist xs lo hi)\ apply (rule mset_sublist_incl) using assms by auto show \set (sublist xs lo hi) \ set (sublist xs' lo hi)\ by (rule mset_sublist_incl) (metis assms size_mset)+ qed text \Our abstract recursive quicksort procedure. We abstract over a partition procedure.\ definition quicksort :: \('b \ 'b \ bool) \ ('a \ 'b) \ nat \ nat \ 'a list \ 'a list nres\ where \quicksort R h = (\(lo,hi,xs0). do { RECT (\f (lo,hi,xs). do { ASSERT(lo \ hi \ hi < length xs \ mset xs = mset xs0); \ \Premise for a partition function\ (xs, p) \ SPEC(uncurry (partition_spec R h xs lo hi)); \ \Abstract partition function\ ASSERT(mset xs = mset xs0); xs \ (if p-1\lo then RETURN xs else f (lo, p-1, xs)); ASSERT(mset xs = mset xs0); if hi\p+1 then RETURN xs else f (p+1, hi, xs) }) (lo,hi,xs0) })\ text \As premise for quicksor, we only need that the indices are ok.\ definition quicksort_pre :: \('b \ 'b \ bool) \ ('a \ 'b) \ 'a list \ nat \ nat \ 'a list \ bool\ where \quicksort_pre R h xs0 lo hi xs \ lo \ hi \ hi < length xs \ mset xs = mset xs0\ definition quicksort_post :: \('b \ 'b \ bool) \ ('a \ 'b) \ nat \ nat \ 'a list \ 'a list \ bool\ where \quicksort_post R h lo hi xs xs' \ mset xs' = mset xs \ sorted_sublist_map R h xs' lo hi \ (\ i. i xs'!i = xs!i) \ (\ j. hij xs'!j = xs!j)\ text \Convert Pure to HOL\ lemma quicksort_postI: \\mset xs' = mset xs; sorted_sublist_map R h xs' lo hi; (\ i. \i \ xs'!i = xs!i); (\ j. \hi \ xs'!j = xs!j)\ \ quicksort_post R h lo hi xs xs'\ by (auto simp add: quicksort_post_def) text \The first case for the correctness proof of (abstract) quicksort: We assume that we called the partition function, and we have \<^term>\p-1\lo\ and \<^term>\hi\p+1\.\ lemma quicksort_correct_case1: assumes trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. x \ y \ R (h x) (h y) \ R (h y) (h x)\ and pre: \quicksort_pre R h xs0 lo hi xs\ and part: \partition_spec R h xs lo hi xs' p\ and ifs: \p-1 \ lo\ \hi \ p+1\ shows \quicksort_post R h lo hi xs xs'\ proof - text \First boilerplate code step: 'unfold' the HOL definitions in the assumptions and convert them to Pure\ have pre: \lo \ hi\ \hi < length xs\ using pre by (auto simp add: quicksort_pre_def) (* have part_perm: \set (sublist xs' lo hi) = set (sublist xs lo hi)\ using part partition_spec_set_sublist pre(1) pre(2) by blast *) have part: \mset xs' = mset xs\ True \isPartition_map R h xs' lo hi p\ \lo \ p\ \p \ hi\ \\ i. i xs'!i=xs!i\ \\ i. \hi \ xs'!i=xs!i\ using part by (auto simp add: partition_spec_def) have sorted_lower: \sorted_sublist_map R h xs' lo (p - Suc 0)\ proof - show ?thesis apply (rule sorted_sublist_wrt_le) subgoal using ifs(1) by auto subgoal using ifs(1) mset_eq_length part(1) pre(1) pre(2) by fastforce done qed have sorted_upper: \sorted_sublist_map R h xs' (Suc p) hi\ proof - show ?thesis apply (rule sorted_sublist_wrt_le) subgoal using ifs(2) by auto subgoal using ifs(1) mset_eq_length part(1) pre(1) pre(2) by fastforce done qed have sorted_middle: \sorted_sublist_map R h xs' lo hi\ proof - show ?thesis apply (rule merge_sorted_map_partitions[where p=p]) subgoal by (rule trans) subgoal by (rule part) subgoal by (rule sorted_lower) subgoal by (rule sorted_upper) subgoal using pre(1) by auto subgoal by (simp add: part(4)) subgoal by (simp add: part(5)) subgoal by (metis part(1) pre(2) size_mset) done qed show ?thesis proof (intro quicksort_postI) show \mset xs' = mset xs\ by (simp add: part(1)) next show \sorted_sublist_map R h xs' lo hi\ by (rule sorted_middle) next show \\i. i < lo \ xs' ! i = xs ! i\ using part(6) by blast next show \\j. \hi < j; j < length xs\ \ xs' ! j = xs ! j\ by (metis part(1) part(7) size_mset) qed qed text \In the second case, we have to show that the precondition still holds for (p+1, hi, x') after the partition.\ lemma quicksort_correct_case2: assumes pre: \quicksort_pre R h xs0 lo hi xs\ and part: \partition_spec R h xs lo hi xs' p\ and ifs: \\ hi \ p + 1\ shows \quicksort_pre R h xs0 (Suc p) hi xs'\ proof - text \First boilerplate code step: 'unfold' the HOL definitions in the assumptions and convert them to Pure\ have pre: \lo \ hi\ \hi < length xs\ \mset xs = mset xs0\ using pre by (auto simp add: quicksort_pre_def) have part: \mset xs' = mset xs\ True \isPartition_map R h xs' lo hi p\ \lo \ p\ \p \ hi\ \\ i. i xs'!i=xs!i\ \\ i. \hi \ xs'!i=xs!i\ using part by (auto simp add: partition_spec_def) show ?thesis unfolding quicksort_pre_def proof (intro conjI) show \Suc p \ hi\ using ifs by linarith show \hi < length xs'\ by (metis part(1) pre(2) size_mset) show \mset xs' = mset xs0\ using pre(3) part(1) by (auto dest: mset_eq_setD) qed qed lemma quicksort_post_set: assumes \quicksort_post R h lo hi xs xs'\ and bounds: \lo \ hi\ \hi < length xs\ shows \set (sublist xs' lo hi) = set (sublist xs lo hi)\ proof - have \mset xs' = mset xs\ \\ i. i xs'!i = xs!i\ \\ j. \hi \ xs'!j = xs!j\ using assms by (auto simp add: quicksort_post_def) then show ?thesis using bounds by (rule mset_sublist_eq, auto) qed text \In the third case, we have run quicksort recursively on (p+1, hi, xs') after the partition, with hi<=p+1 and p-1<=lo.\ lemma quicksort_correct_case3: assumes trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. x \ y \ R (h x) (h y) \ R (h y) (h x)\ and pre: \quicksort_pre R h xs0 lo hi xs\ and part: \partition_spec R h xs lo hi xs' p\ and ifs: \p - Suc 0 \ lo\ \\ hi \ Suc p\ and IH1': \quicksort_post R h (Suc p) hi xs' xs''\ shows \quicksort_post R h lo hi xs xs''\ proof - text \First boilerplate code step: 'unfold' the HOL definitions in the assumptions and convert them to Pure\ have pre: \lo \ hi\ \hi < length xs\ \mset xs = mset xs0\ using pre by (auto simp add: quicksort_pre_def) have part: \mset xs' = mset xs\ True \isPartition_map R h xs' lo hi p\ \lo \ p\ \p \ hi\ \\ i. i xs'!i=xs!i\ \\ i. \hi \ xs'!i=xs!i\ using part by (auto simp add: partition_spec_def) have IH1: \mset xs'' = mset xs'\ \sorted_sublist_map R h xs'' (Suc p) hi\ \\ i. i xs'' ! i = xs' ! i\ \\ j. \hi < j; j < length xs'\ \ xs'' ! j = xs' ! j\ using IH1' by (auto simp add: quicksort_post_def) note IH1_perm = quicksort_post_set[OF IH1'] have still_partition: \isPartition_map R h xs'' lo hi p\ proof(intro isPartition_wrtI) fix i assume \lo \ i\ \i < p\ show \R (h (xs'' ! i)) (h (xs'' ! p))\ text \This holds because this part hasn't changed\ using IH1(3) \i < p\ \lo \ i\ isPartition_wrt_def part(3) by fastforce next fix j assume \p < j\ \j \ hi\ text \Obtain the position \<^term>\posJ\ where \<^term>\xs''!j\ was stored in \<^term>\xs'\.\ have \xs''!j \ set (sublist xs'' (Suc p) hi)\ by (metis IH1(1) Suc_leI \j \ hi\ \p < j\ less_le_trans mset_eq_length part(1) pre(2) sublist_el') then have \xs''!j \ set (sublist xs' (Suc p) hi)\ by (metis IH1_perm ifs(2) nat_le_linear part(1) pre(2) size_mset) then have \\ posJ. Suc p\posJ\posJ\hi \ xs''!j = xs'!posJ\ by (metis Suc_leI \j \ hi\ \p < j\ less_le_trans part(1) pre(2) size_mset sublist_el') then obtain posJ :: nat where PosJ: \Suc p\posJ\ \posJ\hi\ \xs''!j = xs'!posJ\ by blast then show \R (h (xs'' ! p)) (h (xs'' ! j))\ by (metis IH1(3) Suc_le_lessD isPartition_wrt_def lessI part(3)) qed have sorted_lower: \sorted_sublist_map R h xs'' lo (p - Suc 0)\ proof - show ?thesis apply (rule sorted_sublist_wrt_le) subgoal by (simp add: ifs(1)) subgoal using IH1(1) mset_eq_length part(1) part(5) pre(2) by fastforce done qed note sorted_upper = IH1(2) have sorted_middle: \sorted_sublist_map R h xs'' lo hi\ proof - show ?thesis apply (rule merge_sorted_map_partitions[where p=p]) subgoal by (rule trans) subgoal by (rule still_partition) subgoal by (rule sorted_lower) subgoal by (rule sorted_upper) subgoal using pre(1) by auto subgoal by (simp add: part(4)) subgoal by (simp add: part(5)) subgoal by (metis IH1(1) part(1) pre(2) size_mset) done qed show ?thesis proof (intro quicksort_postI) show \mset xs'' = mset xs\ using part(1) IH1(1) by auto \ \I was faster than sledgehammer :-)\ next show \sorted_sublist_map R h xs'' lo hi\ by (rule sorted_middle) next show \\i. i < lo \ xs'' ! i = xs ! i\ using IH1(3) le_SucI part(4) part(6) by auto next show \\j. hi < j \ j < length xs \ xs'' ! j = xs ! j\ by (metis IH1(4) part(1) part(7) size_mset) qed qed text \In the 4th case, we have to show that the premise holds for \<^term>\(lo,p-1,xs')\, in case \<^term>\\p-1\lo\\ text \Analogous to case 2.\ lemma quicksort_correct_case4: assumes pre: \quicksort_pre R h xs0 lo hi xs\ and part: \partition_spec R h xs lo hi xs' p\ and ifs: \\ p - Suc 0 \ lo \ shows \quicksort_pre R h xs0 lo (p-Suc 0) xs'\ proof - text \First boilerplate code step: 'unfold' the HOL definitions in the assumptions and convert them to Pure\ have pre: \lo \ hi\ \hi < length xs\ \mset xs0 = mset xs\ using pre by (auto simp add: quicksort_pre_def) have part: \mset xs' = mset xs\ True \isPartition_map R h xs' lo hi p\ \lo \ p\ \p \ hi\ \\ i. i xs'!i=xs!i\ \\ i. \hi \ xs'!i=xs!i\ using part by (auto simp add: partition_spec_def) show ?thesis unfolding quicksort_pre_def proof (intro conjI) show \lo \ p - Suc 0\ using ifs by linarith show \p - Suc 0 < length xs'\ using mset_eq_length part(1) part(5) pre(2) by fastforce show \mset xs' = mset xs0\ using pre(3) part(1) by (auto dest: mset_eq_setD) qed qed text \In the 5th case, we have run quicksort recursively on (lo, p-1, xs').\ lemma quicksort_correct_case5: assumes trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. x \ y \ R (h x) (h y) \ R (h y) (h x)\ and pre: \quicksort_pre R h xs0 lo hi xs\ and part: \partition_spec R h xs lo hi xs' p\ and ifs: \\ p - Suc 0 \ lo\ \hi \ Suc p\ and IH1': \quicksort_post R h lo (p - Suc 0) xs' xs''\ shows \quicksort_post R h lo hi xs xs''\ proof - text \First boilerplate code step: 'unfold' the HOL definitions in the assumptions and convert them to Pure\ have pre: \lo \ hi\ \hi < length xs\ using pre by (auto simp add: quicksort_pre_def) have part: \mset xs' = mset xs\ True \isPartition_map R h xs' lo hi p\ \lo \ p\ \p \ hi\ \\ i. i xs'!i=xs!i\ \\ i. \hi \ xs'!i=xs!i\ using part by (auto simp add: partition_spec_def) have IH1: \mset xs'' = mset xs'\ \sorted_sublist_map R h xs'' lo (p - Suc 0)\ \\ i. i xs''!i = xs'!i\ \\ j. \p-Suc 0 \ xs''!j = xs'!j\ using IH1' by (auto simp add: quicksort_post_def) note IH1_perm = quicksort_post_set[OF IH1'] have still_partition: \isPartition_map R h xs'' lo hi p\ proof(intro isPartition_wrtI) fix i assume \lo \ i\ \i < p\ text \Obtain the position \<^term>\posI\ where \<^term>\xs''!i\ was stored in \<^term>\xs'\.\ have \xs''!i \ set (sublist xs'' lo (p-Suc 0))\ by (metis (no_types, lifting) IH1(1) Suc_leI Suc_pred \i < p\ \lo \ i\ le_less_trans less_imp_diff_less mset_eq_length not_le not_less_zero part(1) part(5) pre(2) sublist_el') then have \xs''!i \ set (sublist xs' lo (p-Suc 0))\ by (metis IH1_perm ifs(1) le_less_trans less_imp_diff_less mset_eq_length nat_le_linear part(1) part(5) pre(2)) then have \\ posI. lo\posI\posI\p-Suc 0 \ xs''!i = xs'!posI\ proof - \ \sledgehammer\ have "p - Suc 0 < length xs" by (meson diff_le_self le_less_trans part(5) pre(2)) then show ?thesis by (metis (no_types) \xs'' ! i \ set (sublist xs' lo (p - Suc 0))\ ifs(1) mset_eq_length nat_le_linear part(1) sublist_el') qed then obtain posI :: nat where PosI: \lo\posI\ \posI\p-Suc 0\ \xs''!i = xs'!posI\ by blast then show \R (h (xs'' ! i)) (h (xs'' ! p))\ by (metis (no_types, lifting) IH1(4) \i < p\ diff_Suc_less isPartition_wrt_def le_less_trans mset_eq_length not_le not_less_eq part(1) part(3) part(5) pre(2) zero_less_Suc) next fix j assume \p < j\ \j \ hi\ then show \R (h (xs'' ! p)) (h (xs'' ! j))\ text \This holds because this part hasn't changed\ by (smt IH1(4) add_diff_cancel_left' add_diff_inverse_nat diff_Suc_eq_diff_pred diff_le_self ifs(1) isPartition_wrt_def le_less_Suc_eq less_le_trans mset_eq_length nat_less_le part(1) part(3) part(4) plus_1_eq_Suc pre(2)) qed note sorted_lower = IH1(2) have sorted_upper: \sorted_sublist_map R h xs'' (Suc p) hi\ proof - show ?thesis apply (rule sorted_sublist_wrt_le) subgoal by (simp add: ifs(2)) subgoal using IH1(1) mset_eq_length part(1) part(5) pre(2) by fastforce done qed have sorted_middle: \sorted_sublist_map R h xs'' lo hi\ proof - show ?thesis apply (rule merge_sorted_map_partitions[where p=p]) subgoal by (rule trans) subgoal by (rule still_partition) subgoal by (rule sorted_lower) subgoal by (rule sorted_upper) subgoal using pre(1) by auto subgoal by (simp add: part(4)) subgoal by (simp add: part(5)) subgoal by (metis IH1(1) part(1) pre(2) size_mset) done qed show ?thesis proof (intro quicksort_postI) show \mset xs'' = mset xs\ by (simp add: IH1(1) part(1)) next show \sorted_sublist_map R h xs'' lo hi\ by (rule sorted_middle) next show \\i. i < lo \ xs'' ! i = xs ! i\ by (simp add: IH1(3) part(6)) next show \\j. hi < j \ j < length xs \ xs'' ! j = xs ! j\ by (metis IH1(4) diff_le_self dual_order.strict_trans2 mset_eq_length part(1) part(5) part(7)) qed qed text \In the 6th case, we have run quicksort recursively on (lo, p-1, xs'). We show the precondition on the second call on (p+1, hi, xs'')\ lemma quicksort_correct_case6: assumes pre: \quicksort_pre R h xs0 lo hi xs\ and part: \partition_spec R h xs lo hi xs' p\ and ifs: \\ p - Suc 0 \ lo\ \\ hi \ Suc p\ and IH1: \quicksort_post R h lo (p - Suc 0) xs' xs''\ shows \quicksort_pre R h xs0 (Suc p) hi xs''\ proof - text \First boilerplate code step: 'unfold' the HOL definitions in the assumptions and convert them to Pure\ have pre: \lo \ hi\ \hi < length xs\ \mset xs0 = mset xs\ using pre by (auto simp add: quicksort_pre_def) have part: \mset xs' = mset xs\ True \isPartition_map R h xs' lo hi p\ \lo \ p\ \p \ hi\ \\ i. i xs'!i=xs!i\ \\ i. \hi \ xs'!i=xs!i\ using part by (auto simp add: partition_spec_def) have IH1: \mset xs'' = mset xs'\ \sorted_sublist_map R h xs'' lo (p - Suc 0)\ \\ i. i xs''!i = xs'!i\ \\ j. \p-Suc 0 \ xs''!j = xs'!j\ using IH1 by (auto simp add: quicksort_post_def) show ?thesis unfolding quicksort_pre_def proof (intro conjI) show \Suc p \ hi\ using ifs(2) by linarith show \hi < length xs''\ using IH1(1) mset_eq_length part(1) pre(2) by fastforce show \mset xs'' = mset xs0\ using pre(3) part(1) IH1(1) by (auto dest: mset_eq_setD) qed qed text \In the 7th (and last) case, we have run quicksort recursively on (lo, p-1, xs'). We show the postcondition on the second call on (p+1, hi, xs'')\ lemma quicksort_correct_case7: assumes trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. x \ y \ R (h x) (h y) \ R (h y) (h x)\ and pre: \quicksort_pre R h xs0 lo hi xs\ and part: \partition_spec R h xs lo hi xs' p\ and ifs: \\ p - Suc 0 \ lo\ \\ hi \ Suc p\ and IH1': \quicksort_post R h lo (p - Suc 0) xs' xs''\ and IH2': \quicksort_post R h (Suc p) hi xs'' xs'''\ shows \quicksort_post R h lo hi xs xs'''\ proof - text \First boilerplate code step: 'unfold' the HOL definitions in the assumptions and convert them to Pure\ have pre: \lo \ hi\ \hi < length xs\ using pre by (auto simp add: quicksort_pre_def) have part: \mset xs' = mset xs\ True \isPartition_map R h xs' lo hi p\ \lo \ p\ \p \ hi\ \\ i. i xs'!i=xs!i\ \\ i. \hi \ xs'!i=xs!i\ using part by (auto simp add: partition_spec_def) have IH1: \mset xs'' = mset xs'\ \sorted_sublist_map R h xs'' lo (p - Suc 0)\ \\ i. i xs''!i = xs'!i\ \\ j. \p-Suc 0 \ xs''!j = xs'!j\ using IH1' by (auto simp add: quicksort_post_def) note IH1_perm = quicksort_post_set[OF IH1'] have IH2: \mset xs''' = mset xs''\ \sorted_sublist_map R h xs''' (Suc p) hi\ \\ i. i xs'''!i = xs''!i\ \\ j. \hi \ xs'''!j = xs''!j\ using IH2' by (auto simp add: quicksort_post_def) note IH2_perm = quicksort_post_set[OF IH2'] text \We still have a partition after the first call (same as in case 5)\ have still_partition1: \isPartition_map R h xs'' lo hi p\ proof(intro isPartition_wrtI) fix i assume \lo \ i\ \i < p\ text \Obtain the position \<^term>\posI\ where \<^term>\xs''!i\ was stored in \<^term>\xs'\.\ have \xs''!i \ set (sublist xs'' lo (p-Suc 0))\ by (metis (no_types, lifting) IH1(1) Suc_leI Suc_pred \i < p\ \lo \ i\ le_less_trans less_imp_diff_less mset_eq_length not_le not_less_zero part(1) part(5) pre(2) sublist_el') then have \xs''!i \ set (sublist xs' lo (p-Suc 0))\ by (metis IH1_perm ifs(1) le_less_trans less_imp_diff_less mset_eq_length nat_le_linear part(1) part(5) pre(2)) then have \\ posI. lo\posI\posI\p-Suc 0 \ xs''!i = xs'!posI\ proof - \ \sledgehammer\ have "p - Suc 0 < length xs" by (meson diff_le_self le_less_trans part(5) pre(2)) then show ?thesis by (metis (no_types) \xs'' ! i \ set (sublist xs' lo (p - Suc 0))\ ifs(1) mset_eq_length nat_le_linear part(1) sublist_el') qed then obtain posI :: nat where PosI: \lo\posI\ \posI\p-Suc 0\ \xs''!i = xs'!posI\ by blast then show \R (h (xs'' ! i)) (h (xs'' ! p))\ by (metis (no_types, lifting) IH1(4) \i < p\ diff_Suc_less isPartition_wrt_def le_less_trans mset_eq_length not_le not_less_eq part(1) part(3) part(5) pre(2) zero_less_Suc) next fix j assume \p < j\ \j \ hi\ then show \R (h (xs'' ! p)) (h (xs'' ! j))\ text \This holds because this part hasn't changed\ by (smt IH1(4) add_diff_cancel_left' add_diff_inverse_nat diff_Suc_eq_diff_pred diff_le_self ifs(1) isPartition_wrt_def le_less_Suc_eq less_le_trans mset_eq_length nat_less_le part(1) part(3) part(4) plus_1_eq_Suc pre(2)) qed text \We still have a partition after the second call (similar as in case 3)\ have still_partition2: \isPartition_map R h xs''' lo hi p\ proof(intro isPartition_wrtI) fix i assume \lo \ i\ \i < p\ show \R (h (xs''' ! i)) (h (xs''' ! p))\ text \This holds because this part hasn't changed\ using IH2(3) \i < p\ \lo \ i\ isPartition_wrt_def still_partition1 by fastforce next fix j assume \p < j\ \j \ hi\ text \Obtain the position \<^term>\posJ\ where \<^term>\xs'''!j\ was stored in \<^term>\xs'''\.\ have \xs'''!j \ set (sublist xs''' (Suc p) hi)\ by (metis IH1(1) IH2(1) Suc_leI \j \ hi\ \p < j\ ifs(2) nat_le_linear part(1) pre(2) size_mset sublist_el') then have \xs'''!j \ set (sublist xs'' (Suc p) hi)\ by (metis IH1(1) IH2_perm ifs(2) mset_eq_length nat_le_linear part(1) pre(2)) then have \\ posJ. Suc p\posJ\posJ\hi \ xs'''!j = xs''!posJ\ by (metis IH1(1) ifs(2) mset_eq_length nat_le_linear part(1) pre(2) sublist_el') then obtain posJ :: nat where PosJ: \Suc p\posJ\ \posJ\hi\ \xs'''!j = xs''!posJ\ by blast then show \R (h (xs''' ! p)) (h (xs''' ! j))\ proof - \ \sledgehammer\ have "\n na as p. (p (as ! na::'a) (as ! posJ) \ posJ \ na) \ \ isPartition_wrt p as n hi na" by (metis (no_types) PosJ(2) isPartition_wrt_def not_less) then show ?thesis by (metis IH2(3) PosJ(1) PosJ(3) lessI not_less_eq_eq still_partition1) qed qed text \We have that the lower part is sorted after the first recursive call\ note sorted_lower1 = IH1(2) text \We show that it is still sorted after the second call.\ have sorted_lower2: \sorted_sublist_map R h xs''' lo (p-Suc 0)\ proof - show ?thesis using sorted_lower1 apply (rule sorted_wrt_lower_sublist_still_sorted) subgoal by (rule part) subgoal using IH1(1) mset_eq_length part(1) part(5) pre(2) by fastforce subgoal by (simp add: IH2(3)) subgoal by (metis IH2(1) size_mset) done qed text \The second IH gives us the the upper list is sorted after the second recursive call\ note sorted_upper2 = IH2(2) text \Finally, we have to show that the entire list is sorted after the second recursive call.\ have sorted_middle: \sorted_sublist_map R h xs''' lo hi\ proof - show ?thesis apply (rule merge_sorted_map_partitions[where p=p]) subgoal by (rule trans) subgoal by (rule still_partition2) subgoal by (rule sorted_lower2) subgoal by (rule sorted_upper2) subgoal using pre(1) by auto subgoal by (simp add: part(4)) subgoal by (simp add: part(5)) subgoal by (metis IH1(1) IH2(1) part(1) pre(2) size_mset) done qed show ?thesis proof (intro quicksort_postI) show \mset xs''' = mset xs\ by (simp add: IH1(1) IH2(1) part(1)) next show \sorted_sublist_map R h xs''' lo hi\ by (rule sorted_middle) next show \\i. i < lo \ xs''' ! i = xs ! i\ using IH1(3) IH2(3) part(4) part(6) by auto next show \\j. hi < j \ j < length xs \ xs''' ! j = xs ! j\ by (metis IH1(1) IH1(4) IH2(4) diff_le_self ifs(2) le_SucI less_le_trans nat_le_eq_or_lt not_less part(1) part(7) size_mset) qed qed text \We can now show the correctness of the abstract quicksort procedure, using the refinement framework and the above case lemmas.\ lemma quicksort_correct: assumes trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. x \ y \ R (h x) (h y) \ R (h y) (h x)\ and Pre: \lo0 \ hi0\ \hi0 < length xs0\ shows \quicksort R h (lo0,hi0,xs0) \ \ Id (SPEC(\xs. quicksort_post R h lo0 hi0 xs0 xs))\ proof - have wf: \wf (measure (\(lo, hi, xs). Suc hi - lo))\ by auto define pre where \pre = (\(lo,hi,xs). quicksort_pre R h xs0 lo hi xs)\ define post where \post = (\(lo,hi,xs). quicksort_post R h lo hi xs)\ have pre: \pre (lo0,hi0,xs0)\ unfolding quicksort_pre_def pre_def by (simp add: Pre) text \We first generalize the goal a over all states.\ have \WB_Sort.quicksort R h (lo0,hi0,xs0) \ \ Id (SPEC (post (lo0,hi0,xs0)))\ unfolding quicksort_def prod.case apply (rule RECT_rule) apply (refine_mono) apply (rule wf) apply (rule pre) subgoal premises IH for f x apply (refine_vcg ASSERT_leI) unfolding pre_def post_def subgoal \ \First premise (assertion) for partition\ using IH(2) by (simp add: quicksort_pre_def pre_def) subgoal \ \Second premise (assertion) for partition\ using IH(2) by (simp add: quicksort_pre_def pre_def) subgoal using IH(2) by (auto simp add: quicksort_pre_def pre_def dest: mset_eq_setD) text \Termination case: \<^term>\(p-1 \ lo')\ and \<^term>\(hi' \ p+1)\; directly show postcondition\ subgoal unfolding partition_spec_def by (auto dest: mset_eq_setD) subgoal \ \Postcondition (after partition)\ apply - using IH(2) unfolding pre_def apply (simp, elim conjE, split prod.splits) using trans lin apply (rule quicksort_correct_case1) by auto text \Case \<^term>\(p-1 \ lo')\ and \<^term>\(hi' < p+1)\ (Only second recursive call)\ subgoal apply (rule IH(1)[THEN order_trans]) text \Show that the invariant holds for the second recursive call\ subgoal using IH(2) unfolding pre_def apply (simp, elim conjE, split prod.splits) apply (rule quicksort_correct_case2) by auto text \Wellfoundness (easy)\ subgoal by (auto simp add: quicksort_pre_def partition_spec_def) text \Show that the postcondition holds\ subgoal apply (simp add: Misc.subset_Collect_conv post_def, intro allI impI, elim conjE) using trans lin apply (rule quicksort_correct_case3) using IH(2) unfolding pre_def by auto done text \Case: At least the first recursive call\ subgoal apply (rule IH(1)[THEN order_trans]) text \Show that the precondition holds for the first recursive call\ subgoal using IH(2) unfolding pre_def post_def apply (simp, elim conjE, split prod.splits) apply auto apply (rule quicksort_correct_case4) by auto text \Wellfoundness for first recursive call (easy)\ subgoal by (auto simp add: quicksort_pre_def partition_spec_def) text \Simplify some refinement suff...\ apply (simp add: Misc.subset_Collect_conv ASSERT_leI, intro allI impI conjI, elim conjE) apply (rule ASSERT_leI) apply (simp_all add: Misc.subset_Collect_conv ASSERT_leI) subgoal unfolding quicksort_post_def pre_def post_def by (auto dest: mset_eq_setD) text \Only the first recursive call: show postcondition\ subgoal using trans lin apply (rule quicksort_correct_case5) using IH(2) unfolding pre_def post_def by auto apply (rule ASSERT_leI) subgoal unfolding quicksort_post_def pre_def post_def by (auto dest: mset_eq_setD) text \Both recursive calls.\ subgoal apply (rule IH(1)[THEN order_trans]) text \Show precondition for second recursive call (after the first call)\ subgoal unfolding pre_def post_def apply auto apply (rule quicksort_correct_case6) using IH(2) unfolding pre_def post_def by auto text \Wellfoundedness for second recursive call (easy)\ subgoal by (auto simp add: quicksort_pre_def partition_spec_def) text \Show that the postcondition holds (after both recursive calls)\ subgoal apply (simp add: Misc.subset_Collect_conv, intro allI impI, elim conjE) using trans lin apply (rule quicksort_correct_case7) using IH(2) unfolding pre_def post_def by auto done done done done text \Finally, apply the generalized lemma to show the thesis.\ then show ?thesis unfolding post_def by auto qed (* TODO: Show that our (abstract) partition satisifies the specification *) definition partition_main_inv :: \('b \ 'b \ bool) \ ('a \ 'b) \ nat \ nat \ 'a list \ (nat\nat\'a list) \ bool\ where \partition_main_inv R h lo hi xs0 p \ case p of (i,j,xs) \ j < length xs \ j \ hi \ i < length xs \ lo \ i \ i \ j \ mset xs = mset xs0 \ (\k. k \ lo \ k < i \ R (h (xs!k)) (h (xs!hi))) \ \ \All elements from \<^term>\lo\ to \<^term>\i-1\ are smaller than the pivot\ (\k. k \ i \ k < j \ R (h (xs!hi)) (h (xs!k))) \ \ \All elements from \<^term>\i\ to \<^term>\j-1\ are greater than the pivot\ (\k. k < lo \ xs!k = xs0!k) \ \ \Everything below \<^term>\lo\ is unchanged\ (\k. k \ j \ k < length xs \ xs!k = xs0!k) \ \All elements from \<^term>\j\ are unchanged (including everyting above \<^term>\hi\)\ \ text \The main part of the partition function. The pivot is assumed to be the last element. This is exactly the "Lomuto partition scheme" partition function from Wikipedia.\ definition partition_main :: \('b \ 'b \ bool) \ ('a \ 'b) \ nat \ nat \ 'a list \ ('a list \ nat) nres\ where \partition_main R h lo hi xs0 = do { ASSERT(hi < length xs0); pivot \ RETURN (h (xs0 ! hi)); (i,j,xs) \ WHILE\<^sub>T\<^bsup>partition_main_inv R h lo hi xs0\<^esup> \ \We loop from \<^term>\j=lo\ to \<^term>\j=hi-1\.\ (\(i,j,xs). j < hi) (\(i,j,xs). do { ASSERT(i < length xs \ j < length xs); if R (h (xs!j)) pivot then RETURN (i+1, j+1, swap xs i j) else RETURN (i, j+1, xs) }) (lo, lo, xs0); \ \i and j are both initialized to lo\ ASSERT(i < length xs \ j = hi \ lo \ i \ hi < length xs \ mset xs = mset xs0); RETURN (swap xs i hi, i) }\ (* definition partition_spec :: \('b \ 'b \ bool) \ ('a \ 'b) \ 'a list \ nat \ nat \ 'a list \ nat \ bool\ where \partition_spec R h xs lo hi xs' p \ mset xs' = mset xs \ \ \The list is a permutation\ isPartition_map R h xs' lo hi p \ \ \We have a valid partition on the resulting list\ lo \ p \ p \ hi \ \ \The partition index is in bounds\ (\ i. i xs'!i=xs!i) \ (\ i. hii xs'!i=xs!i)\ \ \Everything else is unchanged.\ *) lemma partition_main_correct: assumes bounds: \hi < length xs\ \lo \ hi\ and trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. R (h x) (h y) \ R (h y) (h x)\ shows \partition_main R h lo hi xs \ SPEC(\(xs', p). mset xs = mset xs' \ lo \ p \ p \ hi \ isPartition_map R h xs' lo hi p \ (\ i. i xs'!i=xs!i) \ (\ i. hii xs'!i=xs!i))\ proof - have K: \b \ hi - Suc n \ n > 0 \ Suc n \ hi \ Suc b \ hi - n\ for b hi n by auto have L: \~ R (h x) (h y) \ R (h y) (h x)\ for x y \ \Corollary of linearity\ using assms by blast have M: \a < Suc b \ a = b \ a < b\ for a b by linarith have N: \(a::nat) \ b \ a = b \ a < b\ for a b by arith show ?thesis unfolding partition_main_def choose_pivot_def apply (refine_vcg WHILEIT_rule[where R = \measure(\(i,j,xs). hi-j)\]) subgoal using assms by blast \ \We feed our assumption to the assertion\ subgoal by auto \ \WF\ subgoal \ \Invariant holds before the first iteration\ unfolding partition_main_inv_def using assms apply simp by linarith subgoal unfolding partition_main_inv_def by simp subgoal unfolding partition_main_inv_def by simp subgoal unfolding partition_main_inv_def apply (auto dest: mset_eq_length) done subgoal unfolding partition_main_inv_def by (auto dest: mset_eq_length) subgoal unfolding partition_main_inv_def apply (auto dest: mset_eq_length) by (metis L M mset_eq_length nat_le_eq_or_lt) subgoal unfolding partition_main_inv_def by simp \ \assertions, etc\ subgoal unfolding partition_main_inv_def by simp subgoal unfolding partition_main_inv_def by (auto dest: mset_eq_length) subgoal unfolding partition_main_inv_def by simp subgoal unfolding partition_main_inv_def by (auto dest: mset_eq_length) subgoal unfolding partition_main_inv_def by (auto dest: mset_eq_length) subgoal unfolding partition_main_inv_def by (auto dest: mset_eq_length) subgoal unfolding partition_main_inv_def by simp subgoal unfolding partition_main_inv_def by simp subgoal \ \After the last iteration, we have a partitioning! :-)\ unfolding partition_main_inv_def by (auto simp add: isPartition_wrt_def) subgoal \ \And the lower out-of-bounds parts of the list haven't been changed\ unfolding partition_main_inv_def by auto subgoal \ \And the upper out-of-bounds parts of the list haven't been changed\ unfolding partition_main_inv_def by auto done qed definition partition_between :: \('b \ 'b \ bool) \ ('a \ 'b) \ nat \ nat \ 'a list \ ('a list \ nat) nres\ where \partition_between R h lo hi xs0 = do { ASSERT(hi < length xs0 \ lo \ hi); k \ choose_pivot R h xs0 lo hi; \ \choice of pivot\ ASSERT(k < length xs0); xs \ RETURN (swap xs0 k hi); \ \move the pivot to the last position, before we start the actual loop\ ASSERT(length xs = length xs0); partition_main R h lo hi xs }\ lemma partition_between_correct: assumes \hi < length xs\ and \lo \ hi\ and \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and \\x y. R (h x) (h y) \ R (h y) (h x)\ shows \partition_between R h lo hi xs \ SPEC(uncurry (partition_spec R h xs lo hi))\ proof - have K: \b \ hi - Suc n \ n > 0 \ Suc n \ hi \ Suc b \ hi - n\ for b hi n by auto show ?thesis unfolding partition_between_def choose_pivot_def apply (refine_vcg partition_main_correct) using assms apply (auto dest: mset_eq_length simp add: partition_spec_def) by (metis dual_order.strict_trans2 less_imp_not_eq2 mset_eq_length swap_nth) qed text \We use the median of the first, the middle, and the last element.\ definition choose_pivot3 where \choose_pivot3 R h xs lo (hi::nat) = do { ASSERT(lo < length xs); ASSERT(hi < length xs); let k' = (hi - lo) div 2; let k = lo + k'; ASSERT(k < length xs); let start = h (xs ! lo); let mid = h (xs ! k); let end = h (xs ! hi); if (R start mid \ R mid end) \ (R end mid \ R mid start) then RETURN k else if (R start end \ R end mid) \ (R mid end \ R end start) then RETURN hi else RETURN lo }\ \ \We only have to show that this procedure yields a valid index between \lo\ and \hi\.\ lemma choose_pivot3_choose_pivot: assumes \lo < length xs\ \hi < length xs\ \hi \ lo\ shows \choose_pivot3 R h xs lo hi \ \ Id (choose_pivot R h xs lo hi)\ unfolding choose_pivot3_def choose_pivot_def using assms by (auto intro!: ASSERT_leI simp: Let_def) text \The refined partion function: We use the above pivot function and fold instead of non-deterministic iteration.\ definition partition_between_ref :: \('b \ 'b \ bool) \ ('a \ 'b) \ nat \ nat \ 'a list \ ('a list \ nat) nres\ where \partition_between_ref R h lo hi xs0 = do { ASSERT(hi < length xs0 \ hi < length xs0 \ lo \ hi); k \ choose_pivot3 R h xs0 lo hi; \ \choice of pivot\ ASSERT(k < length xs0); xs \ RETURN (swap xs0 k hi); \ \move the pivot to the last position, before we start the actual loop\ ASSERT(length xs = length xs0); partition_main R h lo hi xs }\ lemma partition_main_ref': \partition_main R h lo hi xs \ \ ((\ a b c d. Id) a b c d) (partition_main R h lo hi xs)\ by auto (*TODO already exists somewhere*) lemma Down_id_eq: \\Id x = x\ by auto lemma partition_between_ref_partition_between: \partition_between_ref R h lo hi xs \ (partition_between R h lo hi xs)\ proof - have swap: \(swap xs k hi, swap xs ka hi) \ Id\ if \k = ka\ for k ka using that by auto have [refine0]: \(h (xsa ! hi), h (xsaa ! hi)) \ Id\ if \(xsa, xsaa) \ Id\ for xsa xsaa using that by auto show ?thesis apply (subst (2) Down_id_eq[symmetric]) unfolding partition_between_ref_def partition_between_def OP_def apply (refine_vcg choose_pivot3_choose_pivot swap partition_main_correct) subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto by (auto intro: Refine_Basic.Id_refine dest: mset_eq_length) qed text \Technical lemma for sepref\ lemma partition_between_ref_partition_between': \(uncurry2 (partition_between_ref R h), uncurry2 (partition_between R h)) \ (nat_rel \\<^sub>r nat_rel) \\<^sub>r \Id\list_rel \\<^sub>f \\Id\list_rel \\<^sub>r nat_rel\nres_rel\ by (intro frefI nres_relI) (auto intro: partition_between_ref_partition_between) text \Example instantiation for pivot\ definition choose_pivot3_impl where \choose_pivot3_impl = choose_pivot3 (\) id\ lemma partition_between_ref_correct: assumes trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. R (h x) (h y) \ R (h y) (h x)\ and bounds: \hi < length xs\ \lo \ hi\ shows \partition_between_ref R h lo hi xs \ SPEC (uncurry (partition_spec R h xs lo hi))\ proof - show ?thesis apply (rule partition_between_ref_partition_between[THEN order_trans]) using bounds apply (rule partition_between_correct[where h=h]) subgoal by (rule trans) subgoal by (rule lin) done qed text \Refined quicksort algorithm: We use the refined partition function.\ definition quicksort_ref :: \_ \ _ \ nat \ nat \ 'a list \ 'a list nres\ where \quicksort_ref R h = (\(lo,hi,xs0). do { RECT (\f (lo,hi,xs). do { ASSERT(lo \ hi \ hi < length xs0 \ mset xs = mset xs0); (xs, p) \ partition_between_ref R h lo hi xs; \ \This is the refined partition function. Note that we need the premises (trans,lin,bounds) here.\ ASSERT(mset xs = mset xs0 \ p \ lo \ p < length xs0); xs \ (if p-1\lo then RETURN xs else f (lo, p-1, xs)); ASSERT(mset xs = mset xs0); if hi\p+1 then RETURN xs else f (p+1, hi, xs) }) (lo,hi,xs0) })\ (*TODO share*) lemma fref_to_Down_curry2: \(uncurry2 f, uncurry2 g) \ [P]\<^sub>f A \ \B\nres_rel \ (\x x' y y' z z'. P ((x', y'), z') \ (((x, y), z), ((x', y'), z')) \ A\ f x y z \ \ B (g x' y' z'))\ unfolding fref_def uncurry_def nres_rel_def by auto lemma fref_to_Down_curry: \(f, g) \ [P]\<^sub>f A \ \B\nres_rel \ (\x x' . P x' \ (x, x') \ A\ f x \ \ B (g x'))\ unfolding fref_def uncurry_def nres_rel_def by auto lemma quicksort_ref_quicksort: assumes bounds: \hi < length xs\ \lo \ hi\ and trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. R (h x) (h y) \ R (h y) (h x)\ shows \quicksort_ref R h x0 \ \ Id (quicksort R h x0)\ proof - have wf: \wf (measure (\(lo, hi, xs). Suc hi - lo))\ by auto have pre: \x0 = x0' \ (x0, x0') \ Id \\<^sub>r Id \\<^sub>r \Id\list_rel\ for x0 x0' :: \nat \ nat \ 'b list\ by auto have [refine0]: \(x1e = x1d) \ (x1e,x1d) \ Id\ for x1e x1d :: \'b list\ by auto show ?thesis unfolding quicksort_def quicksort_ref_def apply (refine_vcg pre partition_between_ref_partition_between'[THEN fref_to_Down_curry2]) text \First assertion (premise for partition)\ subgoal by auto text \First assertion (premise for partition)\ subgoal by auto subgoal by (auto dest: mset_eq_length) subgoal by (auto dest: mset_eq_length mset_eq_setD) text \Correctness of the concrete partition function\ subgoal apply (simp, rule partition_between_ref_correct) subgoal by (rule trans) subgoal by (rule lin) subgoal by auto \ \first premise\ subgoal by auto \ \second premise\ done subgoal by (auto dest: mset_eq_length mset_eq_setD) subgoal by (auto simp: partition_spec_def isPartition_wrt_def) subgoal by (auto simp: partition_spec_def isPartition_wrt_def dest: mset_eq_length) subgoal by (auto dest: mset_eq_length mset_eq_setD) subgoal by (auto dest: mset_eq_length mset_eq_setD) subgoal by (auto dest: mset_eq_length mset_eq_setD) subgoal by (auto dest: mset_eq_length mset_eq_setD) by simp+ qed \ \Sort the entire list\ definition full_quicksort where \full_quicksort R h xs \ if xs = [] then RETURN xs else quicksort R h (0, length xs - 1, xs)\ definition full_quicksort_ref where \full_quicksort_ref R h xs \ if List.null xs then RETURN xs else quicksort_ref R h (0, length xs - 1, xs)\ definition full_quicksort_impl :: \nat list \ nat list nres\ where \full_quicksort_impl xs = full_quicksort_ref (\) id xs\ lemma full_quicksort_ref_full_quicksort: assumes trans: \\ x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. R (h x) (h y) \ R (h y) (h x)\ shows \(full_quicksort_ref R h, full_quicksort R h) \ \Id\list_rel \\<^sub>f \ \Id\list_rel\nres_rel\ proof - show ?thesis unfolding full_quicksort_ref_def full_quicksort_def apply (intro frefI nres_relI) apply (auto intro!: quicksort_ref_quicksort[unfolded Down_id_eq] simp: List.null_def) subgoal by (rule trans) subgoal using lin by blast done qed lemma sublist_entire: \sublist xs 0 (length xs - 1) = xs\ by (simp add: sublist_def) lemma sorted_sublist_wrt_entire: assumes \sorted_sublist_wrt R xs 0 (length xs - 1)\ shows \sorted_wrt R xs\ proof - have \sorted_wrt R (sublist xs 0 (length xs - 1))\ using assms by (simp add: sorted_sublist_wrt_def ) then show ?thesis by (metis sublist_entire) qed lemma sorted_sublist_map_entire: assumes \sorted_sublist_map R h xs 0 (length xs - 1)\ shows \sorted_wrt (\ x y. R (h x) (h y)) xs\ proof - show ?thesis using assms by (rule sorted_sublist_wrt_entire) qed text \Final correctness lemma\ theorem full_quicksort_correct_sorted: assumes trans: \\x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. x \ y \ R (h x) (h y) \ R (h y) (h x)\ shows \full_quicksort R h xs \ \ Id (SPEC(\xs'. mset xs' = mset xs \ sorted_wrt (\ x y. R (h x) (h y)) xs'))\ proof - show ?thesis unfolding full_quicksort_def apply (refine_vcg) subgoal by simp \ \case xs=[]\ subgoal by simp \ \case xs=[]\ apply (rule quicksort_correct[THEN order_trans]) subgoal by (rule trans) subgoal by (rule lin) subgoal by linarith subgoal by simp apply (simp add: Misc.subset_Collect_conv, intro allI impI conjI) subgoal by (auto simp add: quicksort_post_def) subgoal apply (rule sorted_sublist_map_entire) by (auto simp add: quicksort_post_def dest: mset_eq_length) done qed lemma full_quicksort_correct: assumes trans: \\x y z. \R (h x) (h y); R (h y) (h z)\ \ R (h x) (h z)\ and lin: \\x y. R (h x) (h y) \ R (h y) (h x)\ shows \full_quicksort R h xs \ \ Id (SPEC(\xs'. mset xs' = mset xs))\ by (rule order_trans[OF full_quicksort_correct_sorted]) (use assms in auto) end diff --git a/thys/Refine_Imperative_HOL/IICF/Intf/IICF_Multiset.thy b/thys/Refine_Imperative_HOL/IICF/Intf/IICF_Multiset.thy --- a/thys/Refine_Imperative_HOL/IICF/Intf/IICF_Multiset.thy +++ b/thys/Refine_Imperative_HOL/IICF/Intf/IICF_Multiset.thy @@ -1,227 +1,225 @@ section \Multiset Interface\ theory IICF_Multiset imports "../../Sepref" begin subsection \Additions to Multiset Theory\ lemma rel_mset_Plus_gen: assumes "rel_mset A m m'" assumes "rel_mset A n n'" shows "rel_mset A (m+n) (m'+n')" using assms by induction (auto simp: algebra_simps dest: rel_mset_Plus) lemma rel_mset_single: assumes "A x y" shows "rel_mset A {#x#} {#y#}" unfolding rel_mset_def apply (rule exI[where x="[x]"]) apply (rule exI[where x="[y]"]) using assms by auto lemma rel_mset_Minus: assumes BIU: "bi_unique A" shows "\ rel_mset A m n; A x y \ \ rel_mset A (m-{#x#}) (n-{#y#})" unfolding rel_mset_def proof clarsimp fix ml nl assume A: "A x y" assume R: "list_all2 A ml nl" show "\ml'. mset ml' = mset ml - {#x#} \ (\nl'. mset nl' = mset nl - {#y#} \ list_all2 A ml' nl')" proof (cases "x\set ml") case False have "y \ set nl" using A R apply (auto simp: in_set_conv_decomp list_all2_append2 list_all2_Cons2) using False BIU[unfolded bi_unique_alt_def] apply (auto dest: left_uniqueD) done with False R show ?thesis by (auto simp: diff_single_trivial in_multiset_in_set) next case True then obtain ml1 ml2 where [simp]: "ml=ml1@x#ml2" by (auto simp: in_set_conv_decomp) then obtain nl1 nl2 where [simp]: "nl=nl1@y#nl2" and LA: "list_all2 A ml1 nl1" "list_all2 A ml2 nl2" using A R apply (auto simp: in_set_conv_decomp list_all2_append1 list_all2_Cons1) using BIU[unfolded bi_unique_alt_def] apply (auto dest: right_uniqueD) done have "mset (ml1@ml2) = mset ml - {#x#}" "mset (nl1@nl2) = mset nl - {#y#}" using R by (auto simp: algebra_simps add_implies_diff union_assoc) moreover have "list_all2 A (ml1@ml2) (nl1@nl2)" by (rule list_all2_appendI) fact+ ultimately show ?thesis by blast qed qed lemma rel_mset_Minus_gen: assumes BIU: "bi_unique A" assumes "rel_mset A m m'" assumes "rel_mset A n n'" shows "rel_mset A (m-n) (m'-n')" using assms(3,2) apply (induction R\A _ _ rule: rel_mset_induct) apply (auto dest: rel_mset_Minus[OF BIU] simp: algebra_simps) done lemma pcr_count: assumes "bi_unique A" shows "rel_fun (rel_mset A) (rel_fun A (=)) count count" apply (intro rel_funI) unfolding rel_mset_def apply clarsimp subgoal for x y xs ys apply (rotate_tac,induction xs ys rule: list_all2_induct) using assms by (auto simp: bi_unique_alt_def left_uniqueD right_uniqueD) done subsection \Parametricity Setup\ definition [to_relAPP]: "mset_rel A \ p2rel (rel_mset (rel2p A))" lemma rel2p_mset[rel2p]: "rel2p (\A\mset_rel) = rel_mset (rel2p A)" by (simp add: mset_rel_def) lemma p2re_mset[p2rel]: "p2rel (rel_mset A) = \p2rel A\mset_rel" by (simp add: mset_rel_def) lemma mset_rel_empty[simp]: "(a,{#})\\A\mset_rel \ a={#}" "({#},b)\\A\mset_rel \ b={#}" by (auto simp: mset_rel_def p2rel_def rel_mset_def) lemma param_mset_empty[param]: "({#},{#}) \ \A\mset_rel" - unfolding mset_rel_def - apply (simp add: p2rel_def) - by (rule rel_mset_Zero) + by simp lemma param_mset_Plus[param]: "((+),(+))\\A\mset_rel \ \A\mset_rel \ \A\mset_rel" apply (rule rel2pD) apply (simp add: rel2p) apply (intro rel_funI) by (rule rel_mset_Plus_gen) (*lemma param_mset_single[param]: "(Multiset.single,Multiset.single) \ A \ \A\mset_rel" apply (rule rel2pD) apply (simp add: rel2p) apply (intro rel_funI) by (rule rel_mset_single)*) lemma param_mset_add[param]: "(add_mset, add_mset) \ A \ \A\mset_rel \ \A\mset_rel" apply (rule rel2pD) apply (simp add: rel2p) apply (intro rel_funI) by (rule rel_mset_Plus) lemma param_mset_minus[param]: "\single_valued A; single_valued (A\)\ \ ((-), (-)) \ \A\mset_rel \ \A\mset_rel \ \A\mset_rel" apply (rule rel2pD) apply (simp add: rel2p) apply (intro rel_funI) apply (rule rel_mset_Minus_gen) subgoal apply (unfold IS_LEFT_UNIQUE_def[symmetric]) by (simp add: prop2p bi_unique_alt_def) apply (simp; fail) apply (simp; fail) done lemma param_count[param]: "\single_valued A; single_valued (A\)\ \ (count,count)\\A\mset_rel \ A \ nat_rel" apply (rule rel2pD) apply (simp add: prop2p rel2p) apply (rule pcr_count) apply (simp add: bi_unique_alt_def) done lemma param_set_mset[param]: shows "(set_mset, set_mset) \ \A\mset_rel \ \A\set_rel" apply (rule rel2pD; simp add: rel2p) by (rule multiset.set_transfer) definition [simp]: "mset_is_empty m \ m = {#}" lemma mset_is_empty_param[param]: "(mset_is_empty,mset_is_empty) \ \A\mset_rel \ bool_rel" unfolding mset_rel_def mset_is_empty_def[abs_def] by (auto simp: p2rel_def rel_mset_def intro: nres_relI) subsection \Operations\ sepref_decl_op mset_empty: "{#}" :: "\A\mset_rel" . sepref_decl_op mset_is_empty: "\m. m={#}" :: "\A\mset_rel \ bool_rel" unfolding mset_is_empty_def[symmetric] apply (rule frefI) by parametricity (*sepref_decl_op mset_single: "\m. {#m#}" :: "A \ \A\mset_rel" .*) sepref_decl_op mset_insert: "add_mset" :: "A \ \A\mset_rel \ \A\mset_rel" . sepref_decl_op mset_delete: "\x m. m - {#x#}" :: "A \ \A\mset_rel \ \A\mset_rel" where "single_valued A" "single_valued (A\)" . sepref_decl_op mset_plus: "(+)::_ multiset \ _" :: "\A\mset_rel \ \A\mset_rel \ \A\mset_rel" . sepref_decl_op mset_minus: "(-)::_ multiset \ _" :: "\A\mset_rel \ \A\mset_rel \ \A\mset_rel" where "single_valued A" "single_valued (A\)" . sepref_decl_op mset_contains: "(\#)" :: "A \ \A\mset_rel \ bool_rel" where "single_valued A" "single_valued (A\)" . sepref_decl_op mset_count: "\x y. count y x" :: "A \ \A\mset_rel \ nat_rel" where "single_valued A" "single_valued (A\)" . sepref_decl_op mset_pick: "\m. SPEC (\(x,m'). m = {#x#} + m')" :: "[\m. m\{#}]\<^sub>f \A\mset_rel \ A \\<^sub>r \A\mset_rel" unfolding mset_is_empty_def[symmetric] apply (intro frefI nres_relI) apply (refine_vcg SPEC_refine) apply1 (rule ccontr; clarsimp) applyS (metis msed_rel_invL rel2p_def rel2p_mset union_ac(2)) applyS parametricity done subsection \Patterns\ lemma [def_pat_rules]: "{#} \ op_mset_empty" "add_mset \ op_mset_insert" "(=) $b${#} \ op_mset_is_empty$b" "(=) ${#}$b \ op_mset_is_empty$b" "(+) $a$b \ op_mset_plus$a$b" "(-) $a$b \ op_mset_minus$a$b" by (auto intro!: eq_reflection simp: algebra_simps) lemma [def_pat_rules]: "(+) $b$(add_mset$x${#}) \ op_mset_insert$x$b" "(+) $(add_mset$x${#})$b \ op_mset_insert$x$b" "(-) $b$(add_mset$x${#}) \ op_mset_delete$x$b" "(<) $0$(count$a$x) \ op_mset_contains$x$a" "(\) $x$(set_mset$a) \ op_mset_contains$x$a" by (auto intro!: eq_reflection simp: algebra_simps) locale mset_custom_empty = fixes rel empty and op_custom_empty :: "'a multiset" assumes customize_hnr_aux: "(uncurry0 empty,uncurry0 (RETURN (op_mset_empty::'a multiset))) \ unit_assn\<^sup>k \\<^sub>a rel" assumes op_custom_empty_def: "op_custom_empty = op_mset_empty" begin sepref_register op_custom_empty :: "'ax multiset" lemma fold_custom_empty: "{#} = op_custom_empty" "op_mset_empty = op_custom_empty" "mop_mset_empty = RETURN op_custom_empty" unfolding op_custom_empty_def by simp_all lemmas custom_hnr[sepref_fr_rules] = customize_hnr_aux[folded op_custom_empty_def] end 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,10508 +1,10501 @@ (*********************************************************************************** * Copyright (c) 2016-2020 The University of Sheffield, UK * 2019-2020 University of Exeter, UK * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are met: * * * Redistributions of source code must retain the above copyright notice, this * list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above copyright notice, * this list of conditions and the following disclaimer in the documentation * and/or other materials provided with the distribution. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS" * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. * * SPDX-License-Identifier: BSD-2-Clause ***********************************************************************************) section\The Shadow DOM\ theory Shadow_DOM imports "monads/ShadowRootMonad" Core_DOM.Core_DOM begin abbreviation "safe_shadow_root_element_types \ {''article'', ''aside'', ''blockquote'', ''body'', ''div'', ''footer'', ''h1'', ''h2'', ''h3'', ''h4'', ''h5'', ''h6'', ''header'', ''main'', ''nav'', ''p'', ''section'', ''span''}" subsection \Function Definitions\ subsubsection \get\_child\_nodes\ locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) node_ptr list) dom_prog" where "get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr _ = get_M shadow_root_ptr RShadowRoot.child_nodes" definition a_get_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) node_ptr list) dom_prog)) list" where "a_get_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" where "a_get_child_nodes ptr = invoke (CD.a_get_child_nodes_tups @ a_get_child_nodes_tups) ptr ()" definition a_get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then {preserved (get_M (the (cast ptr)) RShadowRoot.child_nodes)} else {}) \ CD.a_get_child_nodes_locs ptr" definition first_child :: "(_) object_ptr \ (_, (_) node_ptr option) dom_prog" where "first_child ptr = do { children \ a_get_child_nodes ptr; return (case children of [] \ None | child#_ \ Some child)}" end global_interpretation l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_child_nodes = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes and get_child_nodes_locs = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes_locs . locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_child_nodes_impl: "get_child_nodes = a_get_child_nodes" assumes get_child_nodes_locs_impl: "get_child_nodes_locs = a_get_child_nodes_locs" begin lemmas get_child_nodes_def = get_child_nodes_impl[unfolded a_get_child_nodes_def get_child_nodes_def] lemmas get_child_nodes_locs_def = get_child_nodes_locs_impl[unfolded a_get_child_nodes_locs_def get_child_nodes_locs_def, folded CD.get_child_nodes_locs_impl] lemma get_child_nodes_ok: assumes "known_ptr ptr" assumes "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_child_nodes ptr)" using assms[unfolded known_ptr_impl type_wf_impl] apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t CD.get_child_nodes_ok CD.known_ptr_impl CD.type_wf_impl apply blast apply(auto simp add: CD.known_ptr_impl a_get_child_nodes_tups_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok dest!: known_ptr_new_shadow_root_ptr intro!: bind_is_OK_I2)[1] by (metis is_shadow_root_ptr_kind_none l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas.get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas_axioms option.case_eq_if shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes) lemma get_child_nodes_ptr_in_heap: assumes "h \ get_child_nodes ptr \\<^sub>r children" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_child_nodes_def invoke_ptr_in_heap dest: is_OK_returns_result_I) lemma get_child_nodes_pure [simp]: "pure (get_child_nodes ptr) h" unfolding get_child_nodes_def a_get_child_nodes_tups_def proof(split CD.get_child_nodes_splits, rule conjI; clarify) assume "known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr" then show "pure (get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr) h" by simp next assume "\ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr" then show "pure (invoke [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r)] ptr ()) h" by(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro: bind_pure_I split: invoke_splits) qed lemma get_child_nodes_reads: "reads (get_child_nodes_locs ptr) (get_child_nodes ptr) h h'" apply (simp add: get_child_nodes_def a_get_child_nodes_tups_def get_child_nodes_locs_def CD.get_child_nodes_locs_def) apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto intro!: reads_subset[OF CD.get_child_nodes_reads[unfolded CD.get_child_nodes_locs_def]] split: if_splits)[1] apply(split invoke_splits, rule conjI)+ apply(auto)[1] apply(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro: reads_subset[OF reads_singleton] reads_subset[OF check_in_heap_reads] intro!: reads_bind_pure reads_subset[OF return_reads] split: option.splits)[1] done end interpretation i_get_child_nodes?: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(simp add: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_child_nodes_is_l_get_child_nodes [instances]: "l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_get_child_nodes_def instances)[1] using get_child_nodes_reads get_child_nodes_ok get_child_nodes_ptr_in_heap get_child_nodes_pure by blast+ paragraph \new\_document\ locale l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_document: "ptr' \ cast new_document_ptr \ h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] using CD.get_child_nodes_new_document apply (metis document_ptr_casts_commute3 empty_iff is_document_ptr_kind_none new_document_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) by (simp add: CD.get_child_nodes_new_document) lemma new_document_no_child_nodes: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using CD.new_document_no_child_nodes apply auto[1] by(auto simp add: DocumentClass.a_known_ptr_def CD.known_ptr_impl known_ptr_def dest!: new_document_is_document_ptr) end interpretation i_new_document_get_child_nodes?: l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_document_get_child_nodes_is_l_new_document_get_child_nodes [instances]: "l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_document_is_l_new_document get_child_nodes_is_l_get_child_nodes apply(simp add: l_new_document_get_child_nodes_def l_new_document_get_child_nodes_axioms_def) using get_child_nodes_new_document new_document_no_child_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] apply (metis document_ptr_casts_commute3 insert_absorb insert_not_empty is_document_ptr_kind_none new_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) apply(auto simp add: CD.get_child_nodes_locs_def)[1] using new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t apply blast apply (smt insertCI new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t singleton_iff) apply (metis document_ptr_casts_commute3 empty_iff new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t singletonD) done lemma new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def )[1] apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto simp add: CD.get_child_nodes_def CD.a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_element_ptr local.CD.known_ptr_impl apply blast apply(auto simp add: is_document_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_character_data_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_element_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply(auto simp add: is_shadow_root_ptr_def split: shadow_root_ptr.splits dest!: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_is_shadow_root_ptr)[1] apply(auto intro!: bind_pure_returns_result_I)[1] apply(drule(1) new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap) apply(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)[1] apply (metis check_in_heap_ptr_in_heap is_OK_returns_result_E old.unit.exhaust) using new_shadow_root_children by (simp add: new_shadow_root_children get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def) end interpretation i_new_shadow_root_get_child_nodes?: l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] locale l_new_shadow_root_get_child_nodes = l_get_child_nodes + assumes get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" assumes new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" lemma new_shadow_root_get_child_nodes_is_l_new_shadow_root_get_child_nodes [instances]: "l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(simp add: l_new_shadow_root_get_child_nodes_def l_new_shadow_root_get_child_nodes_axioms_def instances) using get_child_nodes_new_shadow_root new_shadow_root_no_child_nodes by fast paragraph \new\_element\ locale l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_element_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_new_element: "ptr' \ cast new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by (auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_child_nodes: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using local.new_element_no_child_nodes apply auto[1] apply(auto simp add: invoke_def)[1] apply(auto simp add: new_element_ptr_in_heap get_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def check_in_heap_def new_element_child_nodes intro!: bind_pure_returns_result_I intro: new_element_is_element_ptr elim!: new_element_ptr_in_heap)[1] proof - assume " h \ new_element \\<^sub>r new_element_ptr" assume "h \ new_element \\<^sub>h h'" assume "\ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr)" moreover have "known_ptr (cast new_element_ptr)" using new_element_is_element_ptr \h \ new_element \\<^sub>r new_element_ptr\ by(auto simp add: known_ptr_impl ShadowRootClass.a_known_ptr_def DocumentClass.a_known_ptr_def CharacterDataClass.a_known_ptr_def ElementClass.a_known_ptr_def) ultimately show "False" by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def is_document_ptr_kind_none) qed end interpretation i_new_element_get_child_nodes?: l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_element_get_child_nodes_is_l_new_element_get_child_nodes [instances]: "l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_element_is_l_new_element get_child_nodes_is_l_get_child_nodes apply(auto simp add: l_new_element_get_child_nodes_def l_new_element_get_child_nodes_axioms_def)[1] using get_child_nodes_new_element new_element_no_child_nodes by fast+ subsubsection \delete\_shadow\_root\ locale l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: if_splits option.splits intro: is_shadow_root_ptr_kind_obtains) end locale l_delete_shadow_root_get_child_nodes = l_get_child_nodes_defs + assumes get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_child_nodes_get_child_nodes_locs [instances]: "l_delete_shadow_root_get_child_nodes get_child_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_child_nodes_def)[1] using get_child_nodes_delete_shadow_root apply fast done subsubsection \set\_child\_nodes\ locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = put_M shadow_root_ptr RShadowRoot.child_nodes_update" definition a_set_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog)) list" where "a_set_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "a_set_child_nodes ptr children = invoke (CD.a_set_child_nodes_tups @ a_set_child_nodes_tups) ptr children" definition a_set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" where "a_set_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then all_args (put_M (the (cast ptr)) RShadowRoot.child_nodes_update) else {}) \ CD.a_set_child_nodes_locs ptr" end global_interpretation l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_child_nodes = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes and set_child_nodes_locs = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes_locs . locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_set_child_nodes_defs set_child_nodes set_child_nodes_locs + CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, unit) dom_prog set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_child_nodes_impl: "set_child_nodes = a_set_child_nodes" assumes set_child_nodes_locs_impl: "set_child_nodes_locs = a_set_child_nodes_locs" begin lemmas set_child_nodes_def = set_child_nodes_impl[unfolded a_set_child_nodes_def set_child_nodes_def] lemmas set_child_nodes_locs_def = set_child_nodes_locs_impl[unfolded a_set_child_nodes_locs_def set_child_nodes_locs_def, folded CD.set_child_nodes_locs_impl] lemma set_child_nodes_writes: "writes (set_child_nodes_locs ptr) (set_child_nodes ptr children) h h'" apply (simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes_locs_def) apply(split CD.set_child_nodes_splits, rule conjI)+ apply (simp add: CD.set_child_nodes_writes writes_union_right_I) apply(split invoke_splits, rule conjI)+ apply(auto simp add: a_set_child_nodes_def)[1] apply(auto simp add: set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: writes_bind_pure intro: writes_union_right_I writes_union_left_I split: list.splits)[1] by (simp add: is_shadow_root_ptr_kind_none) lemma set_child_nodes_pointers_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits) lemma set_child_nodes_types_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def type_wf_impl a_set_child_nodes_tups_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits option.splits) end interpretation i_set_child_nodes?: l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs apply(unfold_locales) by (auto simp add: set_child_nodes_def set_child_nodes_locs_def) declare l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_is_l_set_child_nodes [instances]: "l_set_child_nodes type_wf set_child_nodes set_child_nodes_locs" - using instances - apply(auto simp add: l_set_child_nodes_def)[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 - + 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 - 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 + unfolding l_get_shadow_root_def + by (metis (no_types, lifting) ElementClass.l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms get_shadow_root_ok get_shadow_root_pure get_shadow_root_reads l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.intro l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_shadow_root_def) paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_shadow_root_locs_def all_args_def) end locale l_set_disconnected_nodes_get_shadow_root = l_set_disconnected_nodes_defs + l_get_shadow_root_defs + assumes set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_disconnected_nodes_get_shadow_root?: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_shadow_root_is_l_set_disconnected_nodes_get_shadow_root [instances]: "l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_disconnected_nodes_get_shadow_root_def)[1] using set_disconnected_nodes_get_shadow_root apply fast done paragraph \set\_tag\_type\ locale l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_tag_name_locs_def get_shadow_root_locs_def all_args_def intro: element_put_get_preserved[where setter=tag_name_update and getter=shadow_root_opt]) end locale l_set_tag_name_get_shadow_root = l_set_tag_name + l_get_shadow_root + assumes set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_tag_name_get_shadow_root?: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_shadow_root_is_l_set_tag_name_get_shadow_root [instances]: "l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs" using set_tag_name_is_l_set_tag_name get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_tag_name_get_shadow_root_def l_set_tag_name_get_shadow_root_axioms_def) using set_tag_name_get_shadow_root by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" apply(auto simp add: set_child_nodes_locs_def get_shadow_root_locs_def CD.set_child_nodes_locs_def all_args_def)[1] by(auto intro!: element_put_get_preserved[where getter=shadow_root_opt and setter=RElement.child_nodes_update]) end locale l_set_child_nodes_get_shadow_root = l_set_child_nodes_defs + l_get_shadow_root_defs + assumes set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_child_nodes_get_shadow_root?: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_shadow_root_is_l_set_child_nodes_get_shadow_root [instances]: "l_set_child_nodes_get_shadow_root set_child_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_child_nodes_get_shadow_root_def)[1] using set_child_nodes_get_shadow_root apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by(auto simp add: get_shadow_root_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_shadow_root = l_get_shadow_root_defs + assumes get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(auto simp add: l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_shadow_root_get_shadow_root_locs [instances]: "l_delete_shadow_root_get_shadow_root get_shadow_root_locs" apply(auto simp add: l_delete_shadow_root_get_shadow_root_def)[1] using get_shadow_root_delete_shadow_root apply fast done paragraph \new\_character\_data\ locale l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_character_data_get_shadow_root = l_new_character_data + l_get_shadow_root + assumes get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_character_data_get_shadow_root?: l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_character_data_get_shadow_root_is_l_new_character_data_get_shadow_root [instances]: "l_new_character_data_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_character_data_is_l_new_character_data get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_character_data_get_shadow_root_def l_new_character_data_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_character_data by fast paragraph \new\_document\ locale l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_document_get_shadow_root = l_new_document + l_get_shadow_root + assumes get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_document_get_shadow_root?: l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_document_get_shadow_root_is_l_new_document_get_shadow_root [instances]: "l_new_document_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_document_is_l_new_document get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_document_get_shadow_root_def l_new_document_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_document by fast paragraph \new\_element\ locale l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" by(simp add: get_shadow_root_def new_element_shadow_root_opt) end locale l_new_element_get_shadow_root = l_new_element + l_get_shadow_root + assumes get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" assumes new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" interpretation i_new_element_get_shadow_root?: l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_shadow_root_is_l_new_element_get_shadow_root [instances]: "l_new_element_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_element_is_l_new_element get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_element_get_shadow_root_def l_new_element_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_element new_element_no_shadow_root by fast+ paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_shadow_root = l_get_shadow_root + assumes get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_shadow_root?: l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_shadow_root_is_l_new_shadow_root_get_shadow_root [instances]: "l_new_shadow_root_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_shadow_root_get_shadow_root_def l_new_shadow_root_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_shadow_root by fast subsubsection \set\_shadow\_root\ locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" where "a_set_shadow_root element_ptr = put_M element_ptr shadow_root_opt_update" definition a_set_shadow_root_locs :: "(_) element_ptr \ ((_, unit) dom_prog) set" where "a_set_shadow_root_locs element_ptr \ all_args (put_M element_ptr shadow_root_opt_update)" end global_interpretation l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_shadow_root = a_set_shadow_root and set_shadow_root_locs = a_set_shadow_root_locs . locale l_set_shadow_root_defs = fixes set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" fixes set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_shadow_root_impl: "set_shadow_root = a_set_shadow_root" assumes set_shadow_root_locs_impl: "set_shadow_root_locs = a_set_shadow_root_locs" begin lemmas set_shadow_root_def = set_shadow_root_impl[unfolded set_shadow_root_def a_set_shadow_root_def] lemmas set_shadow_root_locs_def = set_shadow_root_locs_impl[unfolded set_shadow_root_locs_def a_set_shadow_root_locs_def] lemma set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr tag)" apply(unfold type_wf_impl) unfolding set_shadow_root_def using get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok) lemma set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" by(simp add: set_shadow_root_def ElementMonad.put_M_ptr_in_heap) lemma set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr tag) h h'" by(auto simp add: set_shadow_root_def set_shadow_root_locs_def intro: writes_bind_pure) lemma set_shadow_root_pointers_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) lemma set_shadow_root_types_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) end interpretation i_set_shadow_root?: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs by (auto simp add: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_shadow_root = l_type_wf +l_set_shadow_root_defs + assumes set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr disc_nodes) h h'" assumes set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr shadow_root)" assumes set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" assumes set_shadow_root_pointers_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_shadow_root_types_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_shadow_root_is_l_set_shadow_root [instances]: "l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs" apply(auto simp add: l_set_shadow_root_def instances)[1] using set_shadow_root_writes apply blast using set_shadow_root_ok apply (blast) using set_shadow_root_ptr_in_heap apply blast using set_shadow_root_pointers_preserved apply(blast, blast) using set_shadow_root_types_preserved apply(blast, blast) done paragraph \get\_shadow\_root\ locale l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" by(auto simp add: set_shadow_root_def get_shadow_root_def) lemma set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ \w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_shadow_root_get_shadow_root?: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_shadow_root = l_type_wf + l_set_shadow_root_defs + l_get_shadow_root_defs + assumes set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" assumes set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ w \ set_shadow_root_locs ptr \ h \ w \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" lemma set_shadow_root_get_shadow_root_is_l_set_shadow_root_get_shadow_root [instances]: "l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs" apply(auto simp add: l_set_shadow_root_get_shadow_root_def instances)[1] using set_shadow_root_get_shadow_root apply fast using set_shadow_root_get_shadow_root_different_pointers apply fast done subsubsection \set\_mode\ locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" where "a_set_mode shadow_root_ptr = put_M shadow_root_ptr mode_update" definition a_set_mode_locs :: "(_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_set_mode_locs shadow_root_ptr \ all_args (put_M shadow_root_ptr mode_update)" end global_interpretation l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_mode = a_set_mode and set_mode_locs = a_set_mode_locs . locale l_set_mode_defs = fixes set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" fixes set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_mode_defs set_mode set_mode_locs + l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_mode_impl: "set_mode = a_set_mode" assumes set_mode_locs_impl: "set_mode_locs = a_set_mode_locs" begin lemmas set_mode_def = set_mode_impl[unfolded set_mode_def a_set_mode_def] lemmas set_mode_locs_def = set_mode_locs_impl[unfolded set_mode_locs_def a_set_mode_locs_def] lemma set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" apply(unfold type_wf_impl) unfolding set_mode_def using get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok) lemma set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" by(simp add: set_mode_def put_M_ptr_in_heap) lemma set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" by(auto simp add: set_mode_def set_mode_locs_def intro: writes_bind_pure) lemma set_mode_pointers_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) lemma set_mode_types_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) end interpretation i_set_mode?: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs by (auto simp add: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_mode = l_type_wf +l_set_mode_defs + assumes set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" assumes set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" assumes set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes set_mode_pointers_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_mode_types_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_mode_is_l_set_mode [instances]: "l_set_mode type_wf set_mode set_mode_locs" apply(auto simp add: l_set_mode_def instances)[1] using set_mode_writes apply blast using set_mode_ok apply (blast) using set_mode_ptr_in_heap apply blast using set_mode_pointers_preserved apply(blast, blast) using set_mode_types_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def set_shadow_root_locs_def CD.get_child_nodes_locs_def all_args_def intro: element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_child_nodes?: l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs set_shadow_root set_shadow_root_locs by(unfold_locales) declare l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_child_nodes = l_set_shadow_root + l_get_child_nodes + assumes set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_child_nodes_is_l_set_shadow_root_get_child_nodes [instances]: "l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_set_shadow_root_get_child_nodes_def l_set_shadow_root_get_child_nodes_axioms_def instances)[1] using set_shadow_root_get_child_nodes apply blast done paragraph \get\_shadow\_root\ locale l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_mode_get_shadow_root?: l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs by unfold_locales declare l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_shadow_root = l_set_mode + l_get_shadow_root + assumes set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" lemma set_mode_get_shadow_root_is_l_set_mode_get_shadow_root [instances]: "l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs" using set_mode_is_l_set_mode get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_mode_get_shadow_root_def l_set_mode_get_shadow_root_axioms_def) using set_mode_get_shadow_root by fast paragraph \get\_child\_nodes\ locale l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def set_mode_locs_def all_args_def)[1] end interpretation i_set_mode_get_child_nodes?: l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by unfold_locales declare l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_child_nodes = l_set_mode + l_get_child_nodes + assumes set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_mode_get_child_nodes_is_l_set_mode_get_child_nodes [instances]: "l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs" using set_mode_is_l_set_mode get_child_nodes_is_l_get_child_nodes apply(simp add: l_set_mode_get_child_nodes_def l_set_mode_get_child_nodes_axioms_def) using set_mode_get_child_nodes by fast subsubsection \get\_host\ locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for get_shadow_root :: "(_::linorder) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" where "a_get_host shadow_root_ptr = do { host_ptrs \ element_ptr_kinds_M \ filter_M (\element_ptr. do { shadow_root_opt \ get_shadow_root element_ptr; return (shadow_root_opt = Some shadow_root_ptr) }); (case host_ptrs of host_ptr#[] \ return host_ptr | _ \ error HierarchyRequestError) }" definition "a_get_host_locs \ (\element_ptr. (get_shadow_root_locs element_ptr)) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end global_interpretation l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs defines get_host = "a_get_host" and get_host_locs = "a_get_host_locs" . locale l_get_host_defs = fixes get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" fixes get_host_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_host_defs + l_get_shadow_root + assumes get_host_impl: "get_host = a_get_host" assumes get_host_locs_impl: "get_host_locs = a_get_host_locs" begin lemmas get_host_def = get_host_impl[unfolded a_get_host_def] lemmas get_host_locs_def = get_host_locs_impl[unfolded a_get_host_locs_def] lemma get_host_pure [simp]: "pure (get_host element_ptr) h" by(auto simp add: get_host_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_host_reads: "reads get_host_locs (get_host element_ptr) h h'" using get_shadow_root_reads[unfolded reads_def] by(auto simp add: get_host_def get_host_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_shadow_root_reads] reads_subset[OF return_reads] reads_subset[OF element_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_host = l_get_host_defs + assumes get_host_pure [simp]: "pure (get_host element_ptr) h" assumes get_host_reads: "reads get_host_locs (get_host node_ptr) h h'" interpretation i_get_host?: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf using instances by (simp add: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_host_def get_host_locs_def) declare l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_host_is_l_get_host [instances]: "l_get_host get_host get_host_locs" apply(auto simp add: l_get_host_def)[1] using get_host_reads apply fast done subsubsection \get\_mode\ locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" where "a_get_mode shadow_root_ptr = get_M shadow_root_ptr mode" definition a_get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_mode_locs shadow_root_ptr \ {preserved (get_M shadow_root_ptr mode)}" end global_interpretation l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_mode = a_get_mode and get_mode_locs = a_get_mode_locs . locale l_get_mode_defs = fixes get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" fixes get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_mode_defs get_mode get_mode_locs + l_type_wf type_wf for get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf :: "(_) heap \ bool" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_mode_impl: "get_mode = a_get_mode" assumes get_mode_locs_impl: "get_mode_locs = a_get_mode_locs" begin lemmas get_mode_def = get_mode_impl[unfolded get_mode_def a_get_mode_def] lemmas get_mode_locs_def = get_mode_locs_impl[unfolded get_mode_locs_def a_get_mode_locs_def] lemma get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" unfolding get_mode_def type_wf_impl using ShadowRootMonad.get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by blast lemma get_mode_pure [simp]: "pure (get_mode element_ptr) h" unfolding get_mode_def by simp lemma get_mode_ptr_in_heap: assumes "h \ get_mode shadow_root_ptr \\<^sub>r children" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: get_mode_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap dest: is_OK_returns_result_I) lemma get_mode_reads: "reads (get_mode_locs element_ptr) (get_mode element_ptr) h h'" by(simp add: get_mode_def get_mode_locs_def reads_bind_pure reads_insert_writes_set_right) end interpretation i_get_mode?: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_mode get_mode_locs type_wf using instances by (auto simp add: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_mode = l_type_wf + l_get_mode_defs + assumes get_mode_reads: "reads (get_mode_locs shadow_root_ptr) (get_mode shadow_root_ptr) h h'" assumes get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" assumes get_mode_ptr_in_heap: "h \ ok (get_mode shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes get_mode_pure [simp]: "pure (get_mode shadow_root_ptr) h" lemma get_mode_is_l_get_mode [instances]: "l_get_mode type_wf get_mode get_mode_locs" apply(auto simp add: l_get_mode_def instances)[1] using get_mode_reads apply blast using get_mode_ok apply blast using get_mode_ptr_in_heap apply blast done subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs for get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" where "a_get_shadow_root_safe element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { mode \ get_mode shadow_root_ptr; (if mode = Open then return (Some shadow_root_ptr) else return None ) } | None \ return None) }" definition a_get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_shadow_root_safe_locs element_ptr shadow_root_ptr \ (get_shadow_root_locs element_ptr) \ (get_mode_locs shadow_root_ptr)" end global_interpretation l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs defines get_shadow_root_safe = a_get_shadow_root_safe and get_shadow_root_safe_locs = a_get_shadow_root_safe_locs . locale l_get_shadow_root_safe_defs = fixes get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" fixes get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs + l_get_shadow_root_safe_defs get_shadow_root_safe get_shadow_root_safe_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_get_mode type_wf get_mode get_mode_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_shadow_root_safe_impl: "get_shadow_root_safe = a_get_shadow_root_safe" assumes get_shadow_root_safe_locs_impl: "get_shadow_root_safe_locs = a_get_shadow_root_safe_locs" begin lemmas get_shadow_root_safe_def = get_shadow_root_safe_impl[unfolded get_shadow_root_safe_def a_get_shadow_root_safe_def] lemmas get_shadow_root_safe_locs_def = get_shadow_root_safe_locs_impl[unfolded get_shadow_root_safe_locs_def a_get_shadow_root_safe_locs_def] lemma get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" apply(auto simp add: get_shadow_root_safe_def)[1] by (smt bind_returns_heap_E is_OK_returns_heap_E local.get_mode_pure local.get_shadow_root_pure option.case_eq_if pure_def pure_returns_heap_eq return_pure) end interpretation i_get_shadow_root_safe?: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs using instances by (auto simp add: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_shadow_root_safe_def get_shadow_root_safe_locs_def) declare l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_shadow_root_safe = l_get_shadow_root_safe_defs + assumes get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" lemma get_shadow_root_safe_is_l_get_shadow_root_safe [instances]: "l_get_shadow_root_safe get_shadow_root_safe" using instances apply(auto simp add: l_get_shadow_root_safe_def)[1] done subsubsection \set\_disconnected\_nodes\ locale l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (set_disconnected_nodes document_ptr node_ptrs)" using CD.set_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf_defs local.type_wf_impl by blast lemma set_disconnected_nodes_typess_preserved: assumes "w \ set_disconnected_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] apply(unfold type_wf_impl) by(auto simp add: all_args_def CD.set_disconnected_nodes_locs_def split: if_splits) end interpretation i_set_disconnected_nodes?: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs by(auto simp add: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_is_l_set_disconnected_nodes [instances]: "l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_def)[1] apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_writes) using set_disconnected_nodes_ok apply blast apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_ptr_in_heap) using i_set_disconnected_nodes.set_disconnected_nodes_pointers_preserved apply (blast, blast) using set_disconnected_nodes_typess_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_child_nodes: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_child_nodes_locs_def CD.get_child_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_child_nodes?: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_child_nodes_is_l_set_disconnected_nodes_get_child_nodes [instances]: "l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes_locs get_child_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_child_nodes_def)[1] using set_disconnected_nodes_get_child_nodes apply fast done paragraph \get\_host\ locale l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def get_shadow_root_locs_def get_host_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_host?: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_host get_host_locs by(auto simp add: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_disconnected_nodes_get_host = l_set_disconnected_nodes_defs + l_get_host_defs + assumes set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" lemma set_disconnected_nodes_get_host_is_l_set_disconnected_nodes_get_host [instances]: "l_set_disconnected_nodes_get_host set_disconnected_nodes_locs get_host_locs" apply(auto simp add: l_set_disconnected_nodes_get_host_def instances)[1] using set_disconnected_nodes_get_host by fast subsubsection \get\_tag\_name\ locale l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_tag_name_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_tag_name element_ptr)" apply(unfold type_wf_impl get_tag_name_impl[unfolded a_get_tag_name_def]) using CD.get_tag_name_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_tag_name?: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_tag_name_is_l_get_tag_name [instances]: "l_get_tag_name type_wf get_tag_name get_tag_name_locs" apply(auto simp add: l_get_tag_name_def)[1] using get_tag_name_reads apply fast using get_tag_name_ok apply fast done paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_tag_name: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_tag_name?: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_get_tag_name_is_l_set_disconnected_nodes_get_tag_name [instances]: "l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_disconnected_nodes_get_tag_name_def l_set_disconnected_nodes_get_tag_name_axioms_def instances)[1] using set_disconnected_nodes_get_tag_name by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_tag_name: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_child_nodes_locs_def set_child_nodes_locs_def CD.get_tag_name_locs_def all_args_def intro: element_put_get_preserved[where getter=tag_name and setter=RElement.child_nodes_update]) end interpretation i_set_child_nodes_get_tag_name?: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_child_nodes_get_tag_name_is_l_set_child_nodes_get_tag_name [instances]: "l_set_child_nodes_get_tag_name type_wf set_child_nodes set_child_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_child_nodes_get_tag_name_def l_set_child_nodes_get_tag_name_axioms_def instances)[1] using set_child_nodes_get_tag_name by fast paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_tag_name_get_tag_name_locs [instances]: "l_delete_shadow_root_get_tag_name get_tag_name_locs" apply(auto simp add: l_delete_shadow_root_get_tag_name_def)[1] using get_tag_name_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_tag_name_locs_def all_args_def element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_tag_name?: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_tag_name get_tag_name_locs apply(auto simp add: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_tag_name = l_set_shadow_root_defs + l_get_tag_name_defs + assumes set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_shadow_root_get_tag_name_is_l_set_shadow_root_get_tag_name [instances]: "l_set_shadow_root_get_tag_name set_shadow_root_locs get_tag_name_locs" using set_shadow_root_is_l_set_shadow_root get_tag_name_is_l_get_tag_name apply(simp add: l_set_shadow_root_get_tag_name_def ) using set_shadow_root_get_tag_name by fast paragraph \new\_element\ locale l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" by(simp add: CD.get_tag_name_def new_element_tag_name) end locale l_new_element_get_tag_name = l_new_element + l_get_tag_name + assumes get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" assumes new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" interpretation i_new_element_get_tag_name?: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_tag_name_is_l_new_element_get_tag_name [instances]: "l_new_element_get_tag_name type_wf get_tag_name get_tag_name_locs" using new_element_is_l_new_element get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_element_get_tag_name_def l_new_element_get_tag_name_axioms_def instances)[1] using get_tag_name_new_element new_element_empty_tag_name by fast+ paragraph \get\_shadow\_root\ locale l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_mode_get_tag_name?: l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_tag_name = l_set_mode + l_get_tag_name + assumes set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_mode_get_tag_name_is_l_set_mode_get_tag_name [instances]: "l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs" using set_mode_is_l_set_mode get_tag_name_is_l_get_tag_name apply(simp add: l_set_mode_get_tag_name_def l_set_mode_get_tag_name_axioms_def) using set_mode_get_tag_name by fast paragraph \new\_document\ locale l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_document_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_document_get_tag_name?: l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_document_get_tag_name_is_l_new_document_get_tag_name [instances]: "l_new_document_get_tag_name get_tag_name_locs" unfolding l_new_document_get_tag_name_def unfolding get_tag_name_locs_def using new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_tag_name = l_get_tag_name + assumes get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_tag_name?: l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(unfold_locales) declare l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_tag_name_is_l_new_shadow_root_get_tag_name [instances]: "l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs" using get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_shadow_root_get_tag_name_def l_new_shadow_root_get_tag_name_axioms_def instances)[1] using get_tag_name_new_shadow_root by fast paragraph \new\_character\_data\ locale l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_character_data_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_character_data_get_tag_name?: l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_character_data_get_tag_name_is_l_new_character_data_get_tag_name [instances]: "l_new_character_data_get_tag_name get_tag_name_locs" unfolding l_new_character_data_get_tag_name_def unfolding get_tag_name_locs_def using new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \get\_tag\_type\ locale l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_tag_name: assumes "h \ CD.a_set_tag_name element_ptr tag \\<^sub>h h'" shows "h' \ CD.a_get_tag_name element_ptr \\<^sub>r tag" using assms by(auto simp add: CD.a_get_tag_name_def CD.a_set_tag_name_def) lemma set_tag_name_get_tag_name_different_pointers: assumes "ptr \ ptr'" assumes "w \ CD.a_set_tag_name_locs ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ CD.a_get_tag_name_locs ptr'" shows "r h h'" using assms by(auto simp add: all_args_def CD.a_set_tag_name_locs_def CD.a_get_tag_name_locs_def split: if_splits option.splits ) end interpretation i_set_tag_name_get_tag_name?: l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs by unfold_locales declare l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_tag_name_is_l_set_tag_name_get_tag_name [instances]: "l_set_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs" using set_tag_name_is_l_set_tag_name get_tag_name_is_l_get_tag_name apply(simp add: l_set_tag_name_get_tag_name_def l_set_tag_name_get_tag_name_axioms_def) using set_tag_name_get_tag_name set_tag_name_get_tag_name_different_pointers by fast+ subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_mode_defs set_mode set_mode_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" where "a_attach_shadow_root element_ptr shadow_root_mode = do { tag \ get_tag_name element_ptr; (if tag \ safe_shadow_root_element_types then error NotSupportedError else return ()); prev_shadow_root \ get_shadow_root element_ptr; (if prev_shadow_root \ None then error NotSupportedError else return ()); new_shadow_root_ptr \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M; set_mode new_shadow_root_ptr shadow_root_mode; set_shadow_root element_ptr (Some new_shadow_root_ptr); return new_shadow_root_ptr }" end locale l_attach_shadow_root_defs = fixes attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" global_interpretation l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs defines attach_shadow_root = a_attach_shadow_root . locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_attach_shadow_root_defs attach_shadow_root + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs + l_set_mode type_wf set_mode set_mode_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ ((_) heap, exception, (_) shadow_root_ptr) prog" and type_wf :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = a_known_ptr" assumes attach_shadow_root_impl: "attach_shadow_root = a_attach_shadow_root" begin lemmas attach_shadow_root_def = a_attach_shadow_root_def[folded attach_shadow_root_impl] lemma attach_shadow_root_element_ptr_in_heap: assumes "h \ ok (attach_shadow_root element_ptr shadow_root_mode)" shows "element_ptr |\| element_ptr_kinds h" proof - obtain h' where "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>h h'" using assms by auto then obtain h2 h3 new_shadow_root_ptr where h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr shadow_root_mode \\<^sub>h h3" and "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) then have "element_ptr |\| element_ptr_kinds h3" using set_shadow_root_ptr_in_heap by blast moreover have "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr new_shadow_root_ptr by auto moreover have "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_mode_writes h3]) using set_mode_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) ultimately show ?thesis by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) qed lemma create_shadow_root_known_ptr: assumes "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>r new_shadow_root_ptr" shows "known_ptr (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr)" using assms by(auto simp add: attach_shadow_root_def known_ptr_impl ShadowRootClass.a_known_ptr_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def elim!: bind_returns_result_E) end locale l_attach_shadow_root = l_attach_shadow_root_defs interpretation i_attach_shadow_root?: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def attach_shadow_root_def instances) declare l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_parent\ global_interpretation l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines get_parent = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent get_child_nodes" and get_parent_locs = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent_locs get_child_nodes_locs" . interpretation i_get_parent?: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs by(simp add: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_parent_def get_parent_locs_def instances) declare l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_parent_is_l_get_parent [instances]: "l_get_parent type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs" apply(simp add: l_get_parent_def l_get_parent_axioms_def instances) using get_parent_reads get_parent_ok get_parent_ptr_in_heap get_parent_pure get_parent_parent_in_heap get_parent_child_dual get_parent_reads_pointers by blast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_child_nodes + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_parent [simp]: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_parent_locs. r h h'))" by(auto simp add: get_parent_locs_def CD.set_disconnected_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_parent?: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs type_wf DocumentClass.type_wf known_ptr known_ptrs get_parent get_parent_locs by (simp add: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_parent_is_l_set_disconnected_nodes_get_parent [instances]: "l_set_disconnected_nodes_get_parent set_disconnected_nodes_locs get_parent_locs" by(simp add: l_set_disconnected_nodes_get_parent_def) subsubsection \get\_root\_node\ global_interpretation l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs defines get_root_node = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node get_parent" and get_root_node_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node_locs get_parent_locs" and get_ancestors = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors get_parent" and get_ancestors_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors_locs get_parent_locs" . declare a_get_ancestors.simps [code] interpretation i_get_root_node?: l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_root_node_def get_root_node_locs_def get_ancestors_def get_ancestors_locs_def instances) declare l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_ancestors_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors" apply(auto simp add: l_get_ancestors_def)[1] using get_ancestors_ptr_in_heap apply fast using get_ancestors_ptr apply fast done lemma get_root_node_is_l_get_root_node [instances]: "l_get_root_node get_root_node get_parent" by (simp add: l_get_root_node_def Shadow_DOM.i_get_root_node.get_root_node_no_parent) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_host_defs get_host get_host_locs for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_get_ancestors_si ptr = do { check_in_heap ptr; ancestors \ (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some node_ptr \ do { parent_ptr_opt \ get_parent node_ptr; (case parent_ptr_opt of Some parent_ptr \ a_get_ancestors_si parent_ptr | None \ return []) } | None \ (case cast ptr of Some shadow_root_ptr \ do { host \ get_host shadow_root_ptr; a_get_ancestors_si (cast host) } | None \ return [])); return (ptr # ancestors) }" definition "a_get_ancestors_si_locs = get_parent_locs \ get_host_locs" definition a_get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" where "a_get_root_node_si ptr = do { ancestors \ a_get_ancestors_si ptr; return (last ancestors) }" definition "a_get_root_node_si_locs = a_get_ancestors_si_locs" end locale l_get_ancestors_si_defs = fixes get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" fixes get_ancestors_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si_defs = fixes get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" fixes get_root_node_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent + l_get_host + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_ancestors_si_defs + l_get_root_node_si_defs + assumes get_ancestors_si_impl: "get_ancestors_si = a_get_ancestors_si" assumes get_ancestors_si_locs_impl: "get_ancestors_si_locs = a_get_ancestors_si_locs" assumes get_root_node_si_impl: "get_root_node_si = a_get_root_node_si" assumes get_root_node_si_locs_impl: "get_root_node_si_locs = a_get_root_node_si_locs" begin lemmas get_ancestors_si_def = a_get_ancestors_si.simps[folded get_ancestors_si_impl] lemmas get_ancestors_si_locs_def = a_get_ancestors_si_locs_def[folded get_ancestors_si_locs_impl] lemmas get_root_node_si_def = a_get_root_node_si_def[folded get_root_node_si_impl get_ancestors_si_impl] lemmas get_root_node_si_locs_def = a_get_root_node_si_locs_def[folded get_root_node_si_locs_impl get_ancestors_si_locs_impl] lemma get_ancestors_si_pure [simp]: "pure (get_ancestors_si ptr) h" proof - have "\ptr h h' x. h \ get_ancestors_si ptr \\<^sub>r x \ h \ get_ancestors_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_get_ancestors_si.fixp_induct[folded get_ancestors_si_impl]) case 1 then show ?case by(rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then show ?case using get_parent_pure get_host_pure apply(auto simp add: pure_returns_heap_eq pure_def split: option.splits elim!: bind_returns_heap_E bind_returns_result_E dest!: pure_returns_heap_eq[rotated, OF check_in_heap_pure])[1] apply (meson option.simps(3) returns_result_eq) apply(metis get_parent_pure pure_returns_heap_eq) apply(metis get_host_pure pure_returns_heap_eq) done qed then show ?thesis by (meson pure_eq_iff) qed lemma get_root_node_si_pure [simp]: "pure (get_root_node_si ptr) h" by(auto simp add: get_root_node_si_def bind_pure_I) lemma get_ancestors_si_ptr_in_heap: assumes "h \ ok (get_ancestors_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_ancestors_si_def check_in_heap_ptr_in_heap elim!: bind_is_OK_E dest: is_OK_returns_result_I) lemma get_ancestors_si_ptr: assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" shows "ptr \ set ancestors" using assms by(simp add: get_ancestors_si_def) (auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I) lemma get_ancestors_si_never_empty: assumes "h \ get_ancestors_si child \\<^sub>r ancestors" shows "ancestors \ []" using assms apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits) lemma get_root_node_si_no_parent: "h \ get_parent node_ptr \\<^sub>r None \ h \ get_root_node_si (cast node_ptr) \\<^sub>r cast node_ptr" apply(auto simp add: check_in_heap_def get_root_node_si_def get_ancestors_si_def intro!: bind_pure_returns_result_I )[1] using get_parent_ptr_in_heap by blast lemma get_root_node_si_root_not_shadow_root: assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root" using assms proof(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2) fix y assume "h \ get_ancestors_si ptr \\<^sub>r y" and "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)" and "root = last y" then show False proof(induct y arbitrary: ptr) case Nil then show ?case using assms(1) get_ancestors_si_never_empty by blast next case (Cons a x) then show ?case apply(auto simp add: get_ancestors_si_def[of ptr] elim!: bind_returns_result_E2 split: option.splits if_splits)[1] using get_ancestors_si_never_empty apply blast using Cons.prems(2) apply auto[1] using \is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)\ \root = last y\ by auto qed qed end global_interpretation l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_host get_host_locs defines get_root_node_si = a_get_root_node_si and get_root_node_si_locs = a_get_root_node_si_locs and get_ancestors_si = a_get_ancestors_si and get_ancestors_si_locs = a_get_ancestors_si_locs . declare a_get_ancestors_si.simps [code] interpretation i_get_root_node_si?: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs apply(auto simp add: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)[1] by(auto simp add: get_root_node_si_def get_root_node_si_locs_def get_ancestors_si_def get_ancestors_si_locs_def) declare l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_si_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors_si" unfolding l_get_ancestors_def using get_ancestors_si_pure get_ancestors_si_ptr get_ancestors_si_ptr_in_heap by blast lemma get_root_node_si_is_l_get_root_node [instances]: "l_get_root_node get_root_node_si get_parent" apply(simp add: l_get_root_node_def) using get_root_node_si_no_parent by fast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_parent + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_host begin lemma set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" by(auto simp add: get_parent_locs_def set_disconnected_nodes_locs_def set_disconnected_nodes_get_host get_ancestors_si_locs_def all_args_def) end locale l_set_disconnected_nodes_get_ancestors_si = l_set_disconnected_nodes_defs + l_get_ancestors_si_defs + assumes set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" interpretation i_set_disconnected_nodes_get_ancestors_si?: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_parent get_parent_locs type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs DocumentClass.type_wf by (auto simp add: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_ancestors_si_is_l_set_disconnected_nodes_get_ancestors_si [instances]: "l_set_disconnected_nodes_get_ancestors_si set_disconnected_nodes_locs get_ancestors_si_locs" using instances apply(simp add: l_set_disconnected_nodes_get_ancestors_si_def) using set_disconnected_nodes_get_ancestors_si by fast subsubsection \get\_attribute\ lemma get_attribute_is_l_get_attribute [instances]: "l_get_attribute type_wf get_attribute get_attribute_locs" apply(auto simp add: l_get_attribute_def)[1] using i_get_attribute.get_attribute_reads apply fast using type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t i_get_attribute.get_attribute_ok apply blast using i_get_attribute.get_attribute_ptr_in_heap apply fast done subsubsection \to\_tree\_order\ global_interpretation l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines to_tree_order = "l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_to_tree_order get_child_nodes" . declare a_to_tree_order.simps [code] interpretation i_to_tree_order?: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ShadowRootClass.known_ptr ShadowRootClass.type_wf Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs to_tree_order by(auto simp add: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_def instances) declare l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_to_tree_order_si ptr = (do { children \ get_child_nodes ptr; shadow_root_part \ (case cast ptr of Some element_ptr \ do { shadow_root_opt \ get_shadow_root element_ptr; (case shadow_root_opt of Some shadow_root_ptr \ return [cast shadow_root_ptr] | None \ return []) } | None \ return []); treeorders \ map_M a_to_tree_order_si ((map (cast) children) @ shadow_root_part); return (ptr # concat treeorders) })" end locale l_to_tree_order_si_defs = fixes to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" global_interpretation l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs defines to_tree_order_si = "a_to_tree_order_si" . declare a_to_tree_order_si.simps [code] locale l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order_si_defs + l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes + l_get_shadow_root + assumes to_tree_order_si_impl: "to_tree_order_si = a_to_tree_order_si" begin lemmas to_tree_order_si_def = a_to_tree_order_si.simps[folded to_tree_order_si_impl] lemma to_tree_order_si_pure [simp]: "pure (to_tree_order_si ptr) h" proof - have "\ptr h h' x. h \ to_tree_order_si ptr \\<^sub>r x \ h \ to_tree_order_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_to_tree_order_si.fixp_induct[folded to_tree_order_si_impl]) case 1 then show ?case by (rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then have "\x h. pure (f x) h" by (metis is_OK_returns_heap_E is_OK_returns_result_E pure_def) then have "\xs h. pure (map_M f xs) h" by(rule map_M_pure_I) then show ?case by(auto elim!: bind_returns_heap_E2 split: option.splits) qed then show ?thesis unfolding pure_def by (metis is_OK_returns_heap_E is_OK_returns_result_E) qed end interpretation i_to_tree_order_si?: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order_si get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs type_wf known_ptr by(auto simp add: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_si_def instances) declare l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \first\_in\_tree\_order\ global_interpretation l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order defines first_in_tree_order = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order" . interpretation i_first_in_tree_order?: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order by(auto simp add: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def first_in_tree_order_def) declare l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_is_l_to_tree_order [instances]: "l_to_tree_order to_tree_order" by(auto simp add: l_to_tree_order_def) subsubsection \first\_in\_tree\_order\ global_interpretation l_dummy defines first_in_tree_order_si = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order_si" . subsubsection \get\_element\_by\ global_interpretation l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order first_in_tree_order get_attribute get_attribute_locs defines get_element_by_id = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order get_attribute" and get_elements_by_class_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order get_attribute" and get_elements_by_tag_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order" . interpretation i_get_element_by?: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order get_attribute get_attribute_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name type_wf by(auto simp add: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_element_by_id_def get_elements_by_class_name_def get_elements_by_tag_name_def instances) declare l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_element_by_is_l_get_element_by [instances]: "l_get_element_by get_element_by_id get_elements_by_tag_name to_tree_order" apply(auto simp add: l_get_element_by_def)[1] using get_element_by_id_result_in_tree_order apply fast done subsubsection \get\_element\_by\_si\ global_interpretation l_dummy defines get_element_by_id_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order_si get_attribute" and get_elements_by_class_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order_si get_attribute" and get_elements_by_tag_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order_si" . subsubsection \find\_slot\ locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs + l_get_attribute_defs get_attribute get_attribute_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_first_in_tree_order_defs first_in_tree_order for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_attribute :: "(_) element_ptr \ char list \ ((_) heap, exception, char list option) prog" and get_attribute_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and first_in_tree_order :: "(_) object_ptr \ ((_) object_ptr \ ((_) heap, exception, (_) element_ptr option) prog) \ ((_) heap, exception, (_) element_ptr option) prog" begin definition a_find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_find_slot open_flag slotable = do { parent_opt \ get_parent slotable; (case parent_opt of Some parent \ if is_element_ptr_kind parent then do { shadow_root_ptr_opt \ get_shadow_root (the (cast parent)); (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { shadow_root_mode \ get_mode shadow_root_ptr; if open_flag \ shadow_root_mode \ Open then return None else first_in_tree_order (cast shadow_root_ptr) (\ptr. if is_element_ptr_kind ptr then do { tag \ get_tag_name (the (cast ptr)); name_attr \ get_attribute (the (cast ptr)) ''name''; slotable_name_attr \ (if is_element_ptr_kind slotable then get_attribute (the (cast slotable)) ''slot'' else return None); (if (tag = ''slot'' \ (name_attr = slotable_name_attr \ (name_attr = None \ slotable_name_attr = Some '''') \ (name_attr = Some '''' \ slotable_name_attr = None))) then return (Some (the (cast ptr))) else return None)} else return None)} | None \ return None)} else return None | _ \ return None)}" definition a_assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_assigned_slot = a_find_slot True" end global_interpretation l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order defines find_slot = a_find_slot and assigned_slot = a_assigned_slot . locale l_find_slot_defs = fixes find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" and assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_find_slot_defs + l_get_parent + l_get_shadow_root + l_get_mode + l_get_attribute + l_get_tag_name + l_to_tree_order + l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + assumes find_slot_impl: "find_slot = a_find_slot" assumes assigned_slot_impl: "assigned_slot = a_assigned_slot" begin lemmas find_slot_def = find_slot_impl[unfolded a_find_slot_def] lemmas assigned_slot_def = assigned_slot_impl[unfolded a_assigned_slot_def] lemma find_slot_ptr_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r slot_opt" shows "slotable |\| node_ptr_kinds h" using assms apply(auto simp add: find_slot_def elim!: bind_returns_result_E2)[1] using get_parent_ptr_in_heap by blast lemma find_slot_slot_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r Some slot" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: find_slot_def first_in_tree_order_def elim!: bind_returns_result_E2 map_filter_M_pure_E[where y=slot] split: option.splits if_splits list.splits intro!: map_filter_M_pure bind_pure_I)[1] using get_tag_name_ptr_in_heap by blast+ lemma find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" by(auto simp add: find_slot_def first_in_tree_order_def intro!: bind_pure_I map_filter_M_pure split: option.splits list.splits) end interpretation i_find_slot?: l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order find_slot assigned_slot type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs to_tree_order by (auto simp add: find_slot_def assigned_slot_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_find_slot = l_find_slot_defs + assumes find_slot_ptr_in_heap: "h \ find_slot open_flag slotable \\<^sub>r slot_opt \ slotable |\| node_ptr_kinds h" assumes find_slot_slot_in_heap: "h \ find_slot open_flag slotable \\<^sub>r Some slot \ slot |\| element_ptr_kinds h" assumes find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" lemma find_slot_is_l_find_slot [instances]: "l_find_slot find_slot" apply(auto simp add: l_find_slot_def)[1] using find_slot_ptr_in_heap apply fast using find_slot_slot_in_heap apply fast done subsubsection \get\_disconnected\_nodes\ locale l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (get_disconnected_nodes document_ptr)" apply(unfold type_wf_impl get_disconnected_nodes_impl[unfolded a_get_disconnected_nodes_def]) using CD.get_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_disconnected_nodes?: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_nodes_is_l_get_disconnected_nodes [instances]: "l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_get_disconnected_nodes_def)[1] using i_get_disconnected_nodes.get_disconnected_nodes_reads apply fast using get_disconnected_nodes_ok apply fast using i_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap apply fast done paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_disconnected_nodes: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def CD.get_disconnected_nodes_locs_def all_args_def) end interpretation i_set_child_nodes_get_disconnected_nodes?: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_disconnected_nodes_is_l_set_child_nodes_get_disconnected_nodes [instances]: "l_set_child_nodes_get_disconnected_nodes type_wf set_child_nodes set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_child_nodes_is_l_set_child_nodes get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_child_nodes_get_disconnected_nodes_def l_set_child_nodes_get_disconnected_nodes_axioms_def) using set_child_nodes_get_disconnected_nodes by fast paragraph \set\_disconnected\_nodes\ lemma set_disconnected_nodes_get_disconnected_nodes_l_set_disconnected_nodes_get_disconnected_nodes [instances]: "l_set_disconnected_nodes_get_disconnected_nodes ShadowRootClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_def l_set_disconnected_nodes_get_disconnected_nodes_axioms_def instances)[1] using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes apply fast using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes_different_pointers apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_disconnected_nodes_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_disconnected_nodes_get_disconnected_nodes_locs [instances]: "l_delete_shadow_root_get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_disconnected_nodes_locs_def all_args_def) end interpretation i_set_shadow_root_get_disconnected_nodes?: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_disconnected_nodes = l_set_shadow_root_defs + l_get_disconnected_nodes_defs + assumes set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_disconnected_nodes_is_l_set_shadow_root_get_disconnected_nodes [instances]: "l_set_shadow_root_get_disconnected_nodes set_shadow_root_locs get_disconnected_nodes_locs" using set_shadow_root_is_l_set_shadow_root get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_shadow_root_get_disconnected_nodes_def ) using set_shadow_root_get_disconnected_nodes by fast paragraph \set\_mode\ locale l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_mode_get_disconnected_nodes?: l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_disconnected_nodes = l_set_mode + l_get_disconnected_nodes + assumes set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_mode_get_disconnected_nodes_is_l_set_mode_get_disconnected_nodes [instances]: "l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_mode_is_l_set_mode get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_mode_get_disconnected_nodes_def l_set_mode_get_disconnected_nodes_axioms_def) using set_mode_get_disconnected_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_disconnected_nodes_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end interpretation i_new_shadow_root_get_disconnected_nodes?: l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_new_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" lemma new_shadow_root_get_disconnected_nodes_is_l_new_shadow_root_get_disconnected_nodes [instances]: "l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes_locs" apply (auto simp add: l_new_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_new_shadow_root apply fast done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" where "a_remove_shadow_root element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { children \ get_child_nodes (cast shadow_root_ptr); (if children = [] then do { set_shadow_root element_ptr None; delete_M shadow_root_ptr } else do { error HierarchyRequestError }) } | None \ error HierarchyRequestError) }" definition a_remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_remove_shadow_root_locs element_ptr shadow_root_ptr \ set_shadow_root_locs element_ptr \ {delete_M shadow_root_ptr}" end global_interpretation l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs defines remove_shadow_root = "a_remove_shadow_root" and remove_shadow_root_locs = a_remove_shadow_root_locs . locale l_remove_shadow_root_defs = fixes remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" fixes remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_remove_shadow_root_defs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + assumes remove_shadow_root_impl: "remove_shadow_root = a_remove_shadow_root" assumes remove_shadow_root_locs_impl: "remove_shadow_root_locs = a_remove_shadow_root_locs" begin lemmas remove_shadow_root_def = remove_shadow_root_impl[unfolded remove_shadow_root_def a_remove_shadow_root_def] lemmas remove_shadow_root_locs_def = remove_shadow_root_locs_impl[unfolded remove_shadow_root_locs_def a_remove_shadow_root_locs_def] lemma remove_shadow_root_writes: "writes (remove_shadow_root_locs element_ptr (the |h \ get_shadow_root element_ptr|\<^sub>r)) (remove_shadow_root element_ptr) h h'" apply(auto simp add: remove_shadow_root_locs_def remove_shadow_root_def all_args_def writes_union_right_I writes_union_left_I set_shadow_root_writes intro!: writes_bind writes_bind_pure[OF get_shadow_root_pure] writes_bind_pure[OF get_child_nodes_pure] intro: writes_subset[OF set_shadow_root_writes] writes_subset[OF writes_singleton2] split: option.splits)[1] using writes_union_left_I[OF set_shadow_root_writes] apply (metis inf_sup_aci(5) insert_is_Un) using writes_union_right_I[OF writes_singleton[of delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M]] by (smt insert_is_Un writes_singleton2 writes_union_left_I) end interpretation i_remove_shadow_root?: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs type_wf known_ptr by(auto simp add: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_shadow_root_def remove_shadow_root_locs_def instances) declare l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] paragraph \get\_child\_nodes\ locale l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_child_nodes_different_pointers: assumes "ptr \ cast shadow_root_ptr" assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_child_nodes_locs ptr" shows "r h h'" using assms apply(auto simp add: all_args_def get_child_nodes_locs_def CD.get_child_nodes_locs_def remove_shadow_root_locs_def set_shadow_root_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t[rotated] element_put_get_preserved[where setter=shadow_root_opt_update] intro: is_shadow_root_ptr_kind_obtains elim: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains split: if_splits option.splits)[1] done end locale l_remove_shadow_root_get_child_nodes = l_get_child_nodes_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_child_nodes_different_pointers: "ptr \ cast shadow_root_ptr \ w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_child_nodes_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_child_nodes?: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs by(auto simp add: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_child_nodes_is_l_remove_shadow_root_get_child_nodes [instances]: "l_remove_shadow_root_get_child_nodes get_child_nodes_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_child_nodes_def instances )[1] using remove_shadow_root_get_child_nodes_different_pointers apply fast done paragraph \get\_tag\_name\ locale l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_tag_name: assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_tag_name_locs ptr" shows "r h h'" using assms by(auto simp add: all_args_def remove_shadow_root_locs_def set_shadow_root_locs_def CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_put_get_preserved[where setter=shadow_root_opt_update] split: if_splits option.splits) end locale l_remove_shadow_root_get_tag_name = l_get_tag_name_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_tag_name: "w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_tag_name_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_tag_name?: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs known_ptr by(auto simp add: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_tag_name_is_l_remove_shadow_root_get_tag_name [instances]: "l_remove_shadow_root_get_tag_name get_tag_name_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_tag_name_def instances )[1] using remove_shadow_root_get_tag_name apply fast done subsubsection \get\_owner\_document\ locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_host_defs get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs get_disconnected_nodes get_disconnected_nodes_locs for get_root_node :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin definition a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr _ = do { host \ get_host shadow_root_ptr; CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast host) () }" definition a_get_owner_document_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) document_ptr) dom_prog)) list" where "a_get_owner_document_tups = [(is_shadow_root_ptr, a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_owner_document :: "(_::linorder) object_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document ptr = invoke (CD.a_get_owner_document_tups @ a_get_owner_document_tups) ptr ()" end global_interpretation l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs defines get_owner_document_tups = a_get_owner_document_tups and get_owner_document = a_get_owner_document and get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r = a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r and get_owner_document_tups\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = "l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document_tups get_root_node_si get_disconnected_nodes" and get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r = "l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r get_root_node_si get_disconnected_nodes" . locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_known_ptr known_ptr + l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs + l_get_owner_document_defs get_owner_document + l_get_host get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node_si get_root_node_si_locs get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and type_wf :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node_si :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_si_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" + assumes known_ptr_impl: "known_ptr = a_known_ptr" assumes get_owner_document_impl: "get_owner_document = a_get_owner_document" begin lemmas known_ptr_def = known_ptr_impl[unfolded a_known_ptr_def] lemmas get_owner_document_def = a_get_owner_document_def[folded get_owner_document_impl] lemma get_owner_document_pure [simp]: "pure (get_owner_document ptr) h" proof - have "\shadow_root_ptr. pure (a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr ()) h" apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits)[1] by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits) then show ?thesis apply(auto simp add: get_owner_document_def)[1] apply(split CD.get_owner_document_splits, rule conjI)+ apply(simp) apply(auto simp add: a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply simp by(auto intro!: bind_pure_I) qed lemma get_owner_document_ptr_in_heap: assumes "h \ ok (get_owner_document ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_owner_document_def invoke_ptr_in_heap dest: is_OK_returns_heap_I) end interpretation i_get_owner_document?: l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr DocumentClass.known_ptr get_parent get_parent_locs type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_host get_host_locs get_owner_document by(auto simp add: instances l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_owner_document_def Core_DOM_Functions.get_owner_document_def) declare l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_is_l_get_owner_document [instances]: "l_get_owner_document get_owner_document" apply(auto simp add: l_get_owner_document_def)[1] using get_owner_document_ptr_in_heap apply fast done subsubsection \remove\_child\ global_interpretation l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs defines remove = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove get_child_nodes set_child_nodes get_parent get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child get_child_nodes set_child_nodes get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child_locs = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child_locs set_child_nodes_locs set_disconnected_nodes_locs" . interpretation i_remove_child?: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs Shadow_DOM.set_child_nodes Shadow_DOM.set_child_nodes_locs Shadow_DOM.get_parent Shadow_DOM.get_parent_locs Shadow_DOM.get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs by(auto simp add: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_child_def remove_child_locs_def remove_def instances) declare l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_disconnected_nodes :: "(_::linorder) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_disconnected_document :: "(_) node_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_disconnected_document node_ptr = do { check_in_heap (cast node_ptr); ptrs \ document_ptr_kinds_M; candidates \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (node_ptr \ set disconnected_nodes) }) ptrs; (case candidates of Cons document_ptr [] \ return document_ptr | _ \ error HierarchyRequestError ) }" definition "a_get_disconnected_document_locs = (\document_ptr. get_disconnected_nodes_locs document_ptr) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end locale l_get_disconnected_document_defs = fixes get_disconnected_document :: "(_) node_ptr \ (_, (_::linorder) document_ptr) dom_prog" fixes get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_disconnected_document_defs + l_get_disconnected_nodes + assumes get_disconnected_document_impl: "get_disconnected_document = a_get_disconnected_document" assumes get_disconnected_document_locs_impl: "get_disconnected_document_locs = a_get_disconnected_document_locs" begin lemmas get_disconnected_document_def = get_disconnected_document_impl[unfolded a_get_disconnected_document_def] lemmas get_disconnected_document_locs_def = get_disconnected_document_locs_impl[unfolded a_get_disconnected_document_locs_def] lemma get_disconnected_document_pure [simp]: "pure (get_disconnected_document ptr) h" using get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_disconnected_document_ptr_in_heap [simp]: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" using get_disconnected_document_def is_OK_returns_result_I check_in_heap_ptr_in_heap by (metis (no_types, lifting) bind_returns_heap_E get_disconnected_document_pure node_ptr_kinds_commutes pure_pure) lemma get_disconnected_document_disconnected_document_in_heap: assumes "h \ get_disconnected_document child_node \\<^sub>r disconnected_document" shows "disconnected_document |\| document_ptr_kinds h" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def elim!: bind_returns_result_E2 dest!: filter_M_not_more_elements[where x=disconnected_document] intro!: filter_M_pure_I bind_pure_I split: if_splits list.splits) lemma get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" using get_disconnected_nodes_reads[unfolded reads_def] by(auto simp add: get_disconnected_document_def get_disconnected_document_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_disconnected_nodes_reads] reads_subset[OF return_reads] reads_subset[OF document_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_disconnected_document = l_get_disconnected_document_defs + assumes get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" assumes get_disconnected_document_ptr_in_heap: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" assumes get_disconnected_document_pure [simp]: "pure (get_disconnected_document node_ptr) h" assumes get_disconnected_document_disconnected_document_in_heap: "h \ get_disconnected_document child_node \\<^sub>r disconnected_document \ disconnected_document |\| document_ptr_kinds h" global_interpretation l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs defines get_disconnected_document = a_get_disconnected_document and get_disconnected_document_locs = a_get_disconnected_document_locs . interpretation i_get_disconnected_document?: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf by(auto simp add: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_disconnected_document_def get_disconnected_document_locs_def instances) declare l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_document_is_l_get_disconnected_document [instances]: "l_get_disconnected_document get_disconnected_document get_disconnected_document_locs" apply(auto simp add: l_get_disconnected_document_def instances)[1] using get_disconnected_document_ptr_in_heap get_disconnected_document_pure get_disconnected_document_disconnected_document_in_heap get_disconnected_document_reads by blast+ paragraph \get\_disconnected\_nodes\ locale l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_disconnected_nodes: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: CD.set_tag_name_locs_impl[unfolded CD.a_set_tag_name_locs_def] CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_tag_name_get_disconnected_nodes?: l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_tag_name_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_disconnected_nodes_is_l_set_tag_name_get_disconnected_nodes [instances]: "l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_tag_name_is_l_set_tag_name get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_tag_name_get_disconnected_nodes_def l_set_tag_name_get_disconnected_nodes_axioms_def) using set_tag_name_get_disconnected_nodes by fast subsubsection \adopt\_node\ global_interpretation l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs defines adopt_node = a_adopt_node and adopt_node_locs = a_adopt_node_locs . interpretation i_adopt_node?: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove by(auto simp add: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def adopt_node_def adopt_node_locs_def instances) declare l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma adopt_node_is_l_adopt_node [instances]: "l_adopt_node type_wf known_ptr known_ptrs get_parent adopt_node adopt_node_locs get_child_nodes get_owner_document" apply(auto simp add: l_adopt_node_def l_adopt_node_axioms_def instances)[1] using adopt_node_writes apply fast using adopt_node_pointers_preserved apply (fast, fast) using adopt_node_types_preserved apply (fast, fast) using adopt_node_child_in_heap apply fast using adopt_node_children_subset apply fast done paragraph \get\_shadow\_root\ locale l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: adopt_node_locs_def remove_child_locs_def all_args_def set_disconnected_nodes_get_shadow_root set_child_nodes_get_shadow_root) end locale l_adopt_node_get_shadow_root = l_adopt_node_defs + l_get_shadow_root_defs + assumes adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_adopt_node_get_shadow_root?: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs adopt_node adopt_node_locs get_child_nodes get_child_nodes_locs known_ptrs remove by(auto simp add: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma adopt_node_get_shadow_root_is_l_adopt_node_get_shadow_root [instances]: "l_adopt_node_get_shadow_root adopt_node_locs get_shadow_root_locs" apply(auto simp add: l_adopt_node_get_shadow_root_def)[1] using adopt_node_get_shadow_root apply fast done subsubsection \insert\_before\ global_interpretation l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_si get_ancestors_si_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document defines next_sibling = a_next_sibling and insert_node = a_insert_node and ensure_pre_insertion_validity = a_ensure_pre_insertion_validity and insert_before = a_insert_before and insert_before_locs = a_insert_before_locs . global_interpretation l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs insert_before defines append_child = "l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_append_child insert_before" . interpretation i_insert_before?: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_si get_ancestors_si_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs by(auto simp add: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def insert_before_def insert_before_locs_def instances) declare l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_append_child?: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M append_child insert_before insert_before_locs by(simp add: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances append_child_def) declare l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsubsection \get\_assigned\_nodes\ fun map_filter_M2 :: "('x \ ('heap, 'e, 'y option) prog) \ 'x list \ ('heap, 'e, 'y list) prog" where "map_filter_M2 f [] = return []" | "map_filter_M2 f (x # xs) = do { res \ f x; remainder \ map_filter_M2 f xs; return ((case res of Some r \ [r] | None \ []) @ remainder) }" lemma map_filter_M2_pure [simp]: assumes "\x. x \ set xs \ pure (f x) h" shows "pure (map_filter_M2 f xs) h" using assms apply(induct xs arbitrary: h) by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I) lemma map_filter_pure_no_monad: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" shows "ys = map the (filter (\x. x \ None) (map (\x. |h \ f x|\<^sub>r) xs))" and "\x. x \ set xs \ h \ ok (f x)" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_pure_foo: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" assumes "y \ set ys" obtains x where "h \ f x \\<^sub>r Some y" and "x \ set xs" using assms apply(induct xs arbitrary: ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_M2_in_result: assumes "h \ map_filter_M2 P xs \\<^sub>r ys" assumes "a \ set xs" assumes "\x. x \ set xs \ pure (P x) h" assumes "h \ P a \\<^sub>r Some b" shows "b \ set ys" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2 ) locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_root_node_defs get_root_node get_root_node_locs + l_get_host_defs get_host get_host_locs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_find_slot_defs find_slot assigned_slot + l_remove_defs remove + l_insert_before_defs insert_before insert_before_locs + l_append_child_defs append_child + l_remove_shadow_root_defs remove_shadow_root remove_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" begin definition a_assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); if is_shadow_root_ptr_kind root then do { host \ get_host (the (cast root)); children \ get_child_nodes (cast host); filter_M (\slotable. do { found_slot \ find_slot False slotable; return (found_slot = Some slot)}) children} else return []}" partial_function (dom_prog) a_assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes_flatten slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); (if is_shadow_root_ptr_kind root then do { slotables \ a_assigned_nodes slot; slotables_or_child_nodes \ (if slotables = [] then do { get_child_nodes (cast slot) } else do { return slotables }); list_of_lists \ map_M (\node_ptr. do { (case cast node_ptr of Some element_ptr \ do { tag \ get_tag_name element_ptr; (if tag = ''slot'' then do { root \ get_root_node (cast element_ptr); (if is_shadow_root_ptr_kind root then do { a_assigned_nodes_flatten element_ptr } else do { return [node_ptr] }) } else do { return [node_ptr] }) } | None \ return [node_ptr]) }) slotables_or_child_nodes; return (concat list_of_lists) } else return []) }" definition a_flatten_dom :: "(_, unit) dom_prog" where "a_flatten_dom = do { tups \ element_ptr_kinds_M \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ a_assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}); forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups; shadow_root_ptr_kinds_M \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }); return () }" end global_interpretation l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs defines assigned_nodes = a_assigned_nodes and assigned_nodes_flatten = a_assigned_nodes_flatten and flatten_dom = a_flatten_dom . declare a_assigned_nodes_flatten.simps [code] locale l_assigned_nodes_defs = fixes assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes flatten_dom :: "(_, unit) dom_prog" locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes_defs assigned_nodes assigned_nodes_flatten flatten_dom + l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_remove + l_insert_before insert_before insert_before_locs + l_find_slot find_slot assigned_slot + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_root_node get_root_node get_root_node_locs get_parent get_parent_locs + l_get_host get_host get_host_locs + l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_to_tree_order to_tree_order for known_ptr :: "(_::linorder) object_ptr \ bool" and assigned_nodes :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and assigned_nodes_flatten :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and flatten_dom :: "((_) heap, exception, unit) prog" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" + assumes assigned_nodes_impl: "assigned_nodes = a_assigned_nodes" assumes flatten_dom_impl: "flatten_dom = a_flatten_dom" begin lemmas assigned_nodes_def = assigned_nodes_impl[unfolded a_assigned_nodes_def] lemmas flatten_dom_def = flatten_dom_impl[unfolded a_flatten_dom_def, folded assigned_nodes_impl] lemma assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" by(auto simp add: assigned_nodes_def intro!: bind_pure_I filter_M_pure_I) lemma assigned_nodes_ptr_in_heap: assumes "h \ ok (assigned_nodes slot)" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: assigned_nodes_def)[1] by (meson bind_is_OK_E is_OK_returns_result_I local.get_tag_name_ptr_in_heap) lemma assigned_nodes_slot_is_slot: assumes "h \ ok (assigned_nodes slot)" shows "h \ get_tag_name slot \\<^sub>r ''slot''" using assms by(auto simp add: assigned_nodes_def elim!: bind_is_OK_E split: if_splits) lemma assigned_nodes_different_ptr: assumes "h \ assigned_nodes slot \\<^sub>r nodes" assumes "h \ assigned_nodes slot' \\<^sub>r nodes'" assumes "slot \ slot'" shows "set nodes \ set nodes' = {}" proof (rule ccontr) assume "set nodes \ set nodes' \ {} " then obtain common_ptr where "common_ptr \ set nodes" and "common_ptr \ set nodes'" by auto have "h \ find_slot False common_ptr \\<^sub>r Some slot" using \common_ptr \ set nodes\ using assms(1) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) moreover have "h \ find_slot False common_ptr \\<^sub>r Some slot'" using \common_ptr \ set nodes'\ using assms(2) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) ultimately show False using assms(3) by (meson option.inject returns_result_eq) qed end interpretation i_assigned_nodes?: l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order by(auto simp add: instances l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def assigned_nodes_def flatten_dom_def) declare l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_assigned_nodes = l_assigned_nodes_defs + assumes assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" assumes assigned_nodes_ptr_in_heap: "h \ ok (assigned_nodes slot) \ slot |\| element_ptr_kinds h" assumes assigned_nodes_slot_is_slot: "h \ ok (assigned_nodes slot) \ h \ get_tag_name slot \\<^sub>r ''slot''" assumes assigned_nodes_different_ptr: "h \ assigned_nodes slot \\<^sub>r nodes \ h \ assigned_nodes slot' \\<^sub>r nodes' \ slot \ slot' \ set nodes \ set nodes' = {}" lemma assigned_nodes_is_l_assigned_nodes [instances]: "l_assigned_nodes assigned_nodes" apply(auto simp add: l_assigned_nodes_def)[1] using assigned_nodes_ptr_in_heap apply fast using assigned_nodes_slot_is_slot apply fast using assigned_nodes_different_ptr apply fast done subsubsection \set\_val\ locale l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_val set_val_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_val :: "(_) character_data_ptr \ char list \ (_, unit) dom_prog" and set_val_locs :: "(_) character_data_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_val_ok: "type_wf h \ character_data_ptr |\| character_data_ptr_kinds h \ h \ ok (set_val character_data_ptr tag)" using CD.set_val_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t local.type_wf_impl by blast lemma set_val_writes: "writes (set_val_locs character_data_ptr) (set_val character_data_ptr tag) h h'" using CD.set_val_writes . lemma set_val_pointers_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms CD.set_val_pointers_preserved by simp lemma set_val_typess_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] split: if_splits) end interpretation i_set_val?: l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs apply(unfold_locales) by (auto simp add: set_val_def set_val_locs_def) declare l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_is_l_set_val [instances]: "l_set_val type_wf set_val set_val_locs" apply(simp add: l_set_val_def) using set_val_ok set_val_writes set_val_pointers_preserved set_val_typess_preserved by blast paragraph \get\_shadow\_root\ locale l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] get_shadow_root_locs_def all_args_def) end locale l_set_val_get_shadow_root = l_set_val + l_get_shadow_root + assumes set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_val_get_shadow_root?: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_shadow_root_is_l_set_val_get_shadow_root [instances]: "l_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs" using set_val_is_l_set_val get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_val_get_shadow_root_def l_set_val_get_shadow_root_axioms_def) using set_val_get_shadow_root by fast paragraph \get\_tag\_type\ locale l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] CD.get_tag_name_locs_impl[unfolded CD.a_get_tag_name_locs_def] all_args_def) end locale l_set_val_get_tag_name = l_set_val + l_get_tag_name + assumes set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" interpretation i_set_val_get_tag_name?: l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_tag_name get_tag_name_locs by unfold_locales declare l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_tag_name_is_l_set_val_get_tag_name [instances]: "l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs" using set_val_is_l_set_val get_tag_name_is_l_get_tag_name apply(simp add: l_set_val_get_tag_name_def l_set_val_get_tag_name_axioms_def) using set_val_get_tag_name by fast subsubsection \create\_character\_data\ locale l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemma create_character_data_document_in_heap: assumes "h \ ok (create_character_data document_ptr text)" shows "document_ptr |\| document_ptr_kinds h" using assms CD.create_character_data_document_in_heap by simp lemma create_character_data_known_ptr: assumes "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" shows "known_ptr (cast new_character_data_ptr)" using assms CD.create_character_data_known_ptr by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def) end locale l_create_character_data = l_create_character_data_defs interpretation i_create_character_data?: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs create_character_data known_ptr DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_element\ locale l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_element known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemmas create_element_def = CD.create_element_def lemma create_element_document_in_heap: assumes "h \ ok (create_element document_ptr tag)" shows "document_ptr |\| document_ptr_kinds h" using CD.create_element_document_in_heap assms . lemma create_element_known_ptr: assumes "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" shows "known_ptr (cast new_element_ptr)" proof - have "is_element_ptr new_element_ptr" using assms apply(auto simp add: create_element_def elim!: bind_returns_result_E)[1] using new_element_is_element_ptr by blast then show ?thesis by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs) qed end interpretation i_create_element?: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsection \A wellformed heap (Core DOM)\ subsubsection \wellformed\_heap\ locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_host_shadow_root_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" where "a_host_shadow_root_rel h = (\(x, y). (cast x, cast y)) ` {(host, shadow_root). host |\| element_ptr_kinds h \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root}" lemma a_host_shadow_root_rel_code [code]: "a_host_shadow_root_rel h = set (concat (map (\host. (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ [(cast host, cast shadow_root)] | None \ [])) (sorted_list_of_fset (element_ptr_kinds h))) )" by(auto simp add: a_host_shadow_root_rel_def) definition a_all_ptrs_in_heap :: "(_) heap \ bool" where "a_all_ptrs_in_heap h = ((\host shadow_root_ptr. (h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h))" definition a_distinct_lists :: "(_) heap \ bool" where "a_distinct_lists h = distinct (concat ( map (\element_ptr. (case |h \ get_shadow_root element_ptr|\<^sub>r of Some shadow_root_ptr \ [shadow_root_ptr] | None \ [])) |h \ element_ptr_kinds_M|\<^sub>r ))" definition a_shadow_root_valid :: "(_) heap \ bool" where "a_shadow_root_valid h = (\shadow_root_ptr \ fset (shadow_root_ptr_kinds h). (\host \ fset(element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr))" definition a_heap_is_wellformed :: "(_) heap \ bool" where "a_heap_is_wellformed h \ CD.a_heap_is_wellformed h \ acyclic (CD.a_parent_child_rel h \ a_host_shadow_root_rel h) \ a_all_ptrs_in_heap h \ a_distinct_lists h \ a_shadow_root_valid h" end global_interpretation l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs defines heap_is_wellformed = a_heap_is_wellformed and parent_child_rel = CD.a_parent_child_rel and host_shadow_root_rel = a_host_shadow_root_rel and all_ptrs_in_heap = a_all_ptrs_in_heap and distinct_lists = a_distinct_lists and shadow_root_valid = a_shadow_root_valid and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_heap_is_wellformed and parent_child_rel\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_parent_child_rel and acyclic_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_acyclic_heap and all_ptrs_in_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_all_ptrs_in_heap and distinct_lists\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_distinct_lists and owner_document_valid\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_owner_document_valid . interpretation i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by (auto simp add: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def instances) declare i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def)[1] using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disc_nodes_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_disc_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes_different apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disconnected_nodes_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply (blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child_in_heap apply blast done locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs + CD: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel + l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" + assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed" begin lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def, folded CD.heap_is_wellformed_impl CD.parent_child_rel_impl] lemma a_distinct_lists_code [code]: "a_all_ptrs_in_heap h = ((\host \ fset (element_ptr_kinds h). h \ ok (get_shadow_root host) \ (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root_ptr \ shadow_root_ptr |\| shadow_root_ptr_kinds h | None \ True)))" apply(auto simp add: a_all_ptrs_in_heap_def split: option.splits)[1] by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap notin_fset select_result_I2) lemma get_shadow_root_shadow_root_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def) lemma get_host_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: get_host_def elim!: bind_returns_result_E2 dest!: filter_M_holds_for_result intro!: bind_pure_I split: list.splits) lemma shadow_root_same_host: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" assumes "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" shows "host = host'" proof (rule ccontr) assume " host \ host'" have "host |\| element_ptr_kinds h" using assms(3) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) moreover have "host' |\| element_ptr_kinds h" using assms(4) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) ultimately show False using assms apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1] apply(drule distinct_concat_map_E(1)[where x=host and y=host']) apply(simp) apply(simp) using \host \ host'\ apply(simp) apply(auto)[1] done qed lemma shadow_root_host_dual: assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms by(auto simp add: get_host_def dest: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: bind_pure_I split: list.splits) lemma disc_doc_disc_node_dual: assumes "h \ get_disconnected_document disc_node \\<^sub>r disc_doc" obtains disc_nodes where "h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes" and "disc_node \ set disc_nodes" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def bind_pure_I dest!: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: filter_M_pure_I split: if_splits list.splits) lemma get_host_valid_tag_name: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" assumes "h \ get_tag_name host \\<^sub>r tag" shows "tag \ safe_shadow_root_element_types" proof - obtain host' where "host' |\| element_ptr_kinds h" and "|h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types" and "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" using assms apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] by (smt assms(1) finite_set_in get_host_ptr_in_heap local.get_shadow_root_ok returns_result_select_result) then have "host = host'" by (meson assms(1) assms(2) assms(3) shadow_root_host_dual shadow_root_same_host) then show ?thesis by (smt\\thesis. (\host'. \host' |\| element_ptr_kinds h; |h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types; h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ \ thesis) \ thesis\ \h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) assms(4) select_result_I2 shadow_root_same_host) qed lemma a_host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)" proof - have "a_host_shadow_root_rel h = (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast host, cast shadow_root)} | None \ {}))" by(auto simp add: a_host_shadow_root_rel_def split: option.splits) moreover have "finite (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root)} | None \ {}))" by(auto split: option.splits) ultimately show ?thesis by auto qed lemma heap_is_wellformed_children_in_heap: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ child |\| node_ptr_kinds h" using CD.heap_is_wellformed_children_in_heap local.heap_is_wellformed_def by blast lemma heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" using CD.heap_is_wellformed_disc_nodes_in_heap local.heap_is_wellformed_def by blast lemma heap_is_wellformed_one_parent: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>r children' \ set children \ set children' \ {} \ ptr = ptr'" using CD.heap_is_wellformed_one_parent local.heap_is_wellformed_def by blast lemma heap_is_wellformed_one_disc_parent: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes' \ set disc_nodes \ set disc_nodes' \ {} \ document_ptr = document_ptr'" using CD.heap_is_wellformed_one_disc_parent local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_disc_nodes_different: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ set children \ set disc_nodes = {}" using CD.heap_is_wellformed_children_disc_nodes_different local.heap_is_wellformed_def by blast lemma heap_is_wellformed_disconnected_nodes_distinct: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ distinct disc_nodes" using CD.heap_is_wellformed_disconnected_nodes_distinct local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using CD.heap_is_wellformed_children_distinct local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_disc_nodes: "heap_is_wellformed h \ node_ptr |\| node_ptr_kinds h \ \(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using CD.heap_is_wellformed_children_disc_nodes local.heap_is_wellformed_def by blast lemma parent_child_rel_finite: "heap_is_wellformed h \ finite (parent_child_rel h)" using CD.parent_child_rel_finite by blast lemma parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" using CD.parent_child_rel_acyclic heap_is_wellformed_def by blast lemma parent_child_rel_child_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptr parent \ (parent, child_ptr) \ parent_child_rel h \ child_ptr |\| object_ptr_kinds h" using CD.parent_child_rel_child_in_heap local.heap_is_wellformed_def by blast end interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs by(auto simp add: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def heap_is_wellformed_def instances) declare l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def instances)[1] using heap_is_wellformed_children_in_heap apply metis using heap_is_wellformed_disc_nodes_in_heap apply metis using heap_is_wellformed_one_parent apply blast using heap_is_wellformed_one_disc_parent apply blast using heap_is_wellformed_children_disc_nodes_different apply blast using heap_is_wellformed_disconnected_nodes_distinct apply metis using heap_is_wellformed_children_distinct apply metis using heap_is_wellformed_children_disc_nodes apply metis using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply(blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using parent_child_rel_child_in_heap apply metis done subsubsection \get\_parent\ interpretation i_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] interpretation i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_parent_wf [instances]: "l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual apply fast using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct_rev apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent apply fast done subsubsection \get\_disconnected\_nodes\ subsubsection \set\_disconnected\_nodes\ paragraph \get\_disconnected\_nodes\ interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] using i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_from_disconnected_nodes_removes apply fast done paragraph \get\_root\_node\ interpretation i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M:l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_never_empty apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_reads apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ptrs_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_remains_not_in_ancestors apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_also_parent apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_obtains_children apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ptr_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_root_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_same_root_node apply(blast, blast) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_same_no_parent apply blast (* using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_not_node_same apply blast *) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_parent_same apply (blast, blast) done subsubsection \to\_tree\_order\ interpretation i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) done declare i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ok apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptrs_in_heap apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent_child_rel apply(blast, blast) using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child2 apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_node_ptrs apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptr_in_result apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_subset apply blast done paragraph \get\_root\_node\ interpretation i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order by(auto simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order get_root_node heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M" apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def instances)[1] using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_get_root_node apply blast using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_same_root apply blast done subsubsection \remove\_child\ interpretation i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by unfold_locales declare i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_first_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_for_all_empty_children apply fast done subsection \A wellformed heap\ subsubsection \get\_parent\ interpretation i_get_parent_wf?: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes using instances by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_parent_wf_is_l_get_parent_wf [instances]: "l_get_parent_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs heap_is_wellformed parent_child_rel Shadow_DOM.get_child_nodes Shadow_DOM.get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using child_parent_dual apply blast using heap_wellformed_induct apply metis using heap_wellformed_induct_rev apply metis using parent_child_rel_parent apply metis done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name + l_get_disconnected_nodes + l_set_shadow_root_get_tag_name + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_delete_shadow_root_get_disconnected_nodes + l_delete_shadow_root_get_child_nodes + l_set_shadow_root_get_disconnected_nodes + l_set_shadow_root_get_child_nodes + l_delete_shadow_root_get_tag_name + l_set_shadow_root_get_shadow_root + l_delete_shadow_root_get_shadow_root + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_shadow_root ptr \\<^sub>h h'" shows "known_ptrs h'" and "type_wf h'" "heap_is_wellformed h'" proof - obtain shadow_root_ptr h2 where "h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr" and "h \ get_child_nodes (cast shadow_root_ptr) \\<^sub>r []" and h2: "h \ set_shadow_root ptr None \\<^sub>h h2" and h': "h2 \ delete_M shadow_root_ptr \\<^sub>h h'" using assms(4) by(auto simp add: remove_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] split: option.splits if_splits) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_shadow_root_writes h2] using \type_wf h\ set_shadow_root_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using h' delete_shadow_root_type_wf_preserved local.type_wf_impl by blast have object_ptr_kinds_eq_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_shadow_root_writes h2]) using set_shadow_root_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) have node_ptr_kinds_eq_h: "node_ptr_kinds h = node_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: node_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h = element_ptr_kinds h2" using node_ptr_kinds_eq_h by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h = document_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: shadow_root_ptr_kinds_def) have "known_ptrs h2" using \known_ptrs h\ object_ptr_kinds_eq_h known_ptrs_subset by blast have object_ptr_kinds_eq_h2: "object_ptr_kinds h' |\| object_ptr_kinds h2" using h' delete_shadow_root_pointers by auto have object_ptr_kinds_eq2_h2: "object_ptr_kinds h2 = object_ptr_kinds h' |\| {|cast shadow_root_ptr|}" using h' delete_shadow_root_pointers by auto have node_ptr_kinds_eq_h2: "node_ptr_kinds h2 = node_ptr_kinds h'" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def delete_shadow_root_pointers[OF h']) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h'" using node_ptr_kinds_eq_h2 by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h'" using object_ptr_kinds_eq_h2 by(auto simp add: document_ptr_kinds_def delete_shadow_root_pointers[OF h']) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h' |\| shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by (auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq2_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h' |\| {|shadow_root_ptr|}" using object_ptr_kinds_eq2_h2 apply (auto simp add: shadow_root_ptr_kinds_def)[1] by (metis \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) fset.map_comp local.get_shadow_root_shadow_root_ptr_in_heap object_ptr_kinds_eq_h shadow_root_ptr_kinds_def) show "known_ptrs h'" using object_ptr_kinds_eq_h2 \known_ptrs h2\ known_ptrs_subset by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_shadow_root_writes h2 set_shadow_root_get_disconnected_nodes by(rule reads_writes_preserved) then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads get_disconnected_nodes_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\doc_ptr disc_nodes. h \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h2 set_shadow_root_get_tag_name by(rule reads_writes_preserved) then have tag_name_eq2_h: "\doc_ptr. |h \ get_tag_name doc_ptr|\<^sub>r = |h2 \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\doc_ptr disc_nodes. h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h' \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads get_tag_name_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h2: "\doc_ptr. |h2 \ get_tag_name doc_ptr|\<^sub>r = |h' \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h: "\ptr' children. h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_shadow_root_writes h2 set_shadow_root_get_child_nodes by(rule reads_writes_preserved) then have children_eq2_h: "\ptr'. |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. ptr' \ cast shadow_root_ptr \ h2 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h' get_child_nodes_delete_shadow_root apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h2: "\ptr'. ptr' \ cast shadow_root_ptr \ |h2 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "cast shadow_root_ptr |\| object_ptr_kinds h'" using h' delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by auto have get_shadow_root_eq_h: "\shadow_root_opt ptr'. ptr \ ptr' \ h \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads set_shadow_root_writes h2 apply(rule reads_writes_preserved) using set_shadow_root_get_shadow_root_different_pointers by fast have get_shadow_root_eq_h2: "\shadow_root_opt ptr'. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads get_shadow_root_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have get_shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "parent_child_rel h = parent_child_rel h2" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h children_eq2_h) moreover have "parent_child_rel h' \ parent_child_rel h2" using object_ptr_kinds_eq_h2 apply(auto simp add: CD.parent_child_rel_def)[1] by (metis \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ children_eq2_h2) ultimately have "CD.a_acyclic_heap h'" using acyclic_subset by (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" by(auto simp add: children_eq2_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h CD.a_all_ptrs_in_heap_def object_ptr_kinds_eq_h node_ptr_kinds_def children_eq_h disconnected_nodes_eq_h) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 children_eq_h2 disconnected_nodes_eq_h2)[1] apply(case_tac "ptr = cast shadow_root_ptr") using object_ptr_kinds_eq_h2 children_eq_h2 apply (meson \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ is_OK_returns_result_I local.get_child_nodes_ptr_in_heap) apply (metis (no_types, lifting) children_eq2_h2 fin_mono finite_set_in object_ptr_kinds_eq_h2 subsetD) by (metis (full_types) assms(1) assms(2) disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 returns_result_select_result) moreover have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" by(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] apply(auto simp add: intro!: distinct_concat_map_I)[1] apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 concat_map_all_distinct[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] apply (metis (no_types, lifting) children_eq2_h2 finite_fset fmember.rep_eq fset_mp object_ptr_kinds_eq_h2 set_sorted_list_of_set) apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp apply(case_tac "y = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 distinct_concat_map_E(1)[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] apply (smt IntI children_eq2_h2 empty_iff finite_fset fmember.rep_eq fset_mp object_ptr_kinds_eq_h2 set_sorted_list_of_set) apply(case_tac "xa = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp by (smt \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 disconnected_nodes_eq_h2 fset_mp is_OK_returns_result_E local.CD.distinct_lists_no_parent local.get_disconnected_nodes_ok object_ptr_kinds_eq_h2 select_result_I2) moreover have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h2" by(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h node_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def document_ptr_kinds_eq_h2 node_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] proof - fix node_ptr assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h2 \ node_ptr \ set |h2 \ get_child_nodes parent_ptr|\<^sub>r)" and 1: "node_ptr |\| node_ptr_kinds h'" and 2: "\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h' \ get_child_nodes parent_ptr|\<^sub>r" then have "\parent_ptr. parent_ptr |\| object_ptr_kinds h2 \ node_ptr \ set |h2 \ get_child_nodes parent_ptr|\<^sub>r" apply(auto)[1] apply(case_tac "parent_ptr = cast shadow_root_ptr") using \h \ get_child_nodes (cast shadow_root_ptr) \\<^sub>r []\ children_eq_h apply auto[1] using children_eq2_h2 object_ptr_kinds_eq_h2 h' delete_shadow_root_pointers by (metis fempty_iff finsert_iff funionE) then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using 0 1 by auto qed ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by(simp add: CD.heap_is_wellformed_def) moreover have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2)" proof - have "a_host_shadow_root_rel h2 \ a_host_shadow_root_rel h" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h)[1] apply(case_tac "aa = ptr") apply(simp) apply (metis (no_types, lifting) \type_wf h2\ assms(2) h2 local.get_shadow_root_ok local.type_wf_impl option.distinct(1) returns_result_eq returns_result_select_result set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by (metis (mono_tags, lifting) \type_wf h2\ image_eqI is_OK_returns_result_E local.get_shadow_root_ok mem_Collect_eq prod.simps(2) select_result_I2) then show ?thesis using \parent_child_rel h = parent_child_rel h2\ by (metis (no_types, hide_lams) \acyclic (parent_child_rel h \ a_host_shadow_root_rel h)\ acyclic_subset order_refl sup_mono) qed then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" proof - have "a_host_shadow_root_rel h' \ a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) then show ?thesis using \parent_child_rel h' \ parent_child_rel h2\ \acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2)\ using acyclic_subset sup_mono by (metis (no_types, hide_lams)) qed moreover have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] apply(case_tac "host = ptr") apply(simp) apply (metis assms(2) h2 local.type_wf_impl option.distinct(1) returns_result_eq set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def get_shadow_root_eq_h2)[1] apply(auto simp add: shadow_root_ptr_kinds_eq2_h2)[1] by (metis (no_types, lifting) \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) get_shadow_root_eq_h get_shadow_root_eq_h2 h2 local.shadow_root_same_host local.type_wf_impl option.distinct(1) select_result_I2 set_shadow_root_get_shadow_root) moreover have "a_distinct_lists h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = ptr") apply(simp) apply (metis (no_types, hide_lams) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) apply(case_tac "y = ptr") apply(simp) apply (metis (no_types, hide_lams) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) by (metis \type_wf h2\ assms(1) assms(2) get_shadow_root_eq_h local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) moreover have "a_shadow_root_valid h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" apply(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h tag_name_eq2_h)[1] apply(simp add: shadow_root_ptr_kinds_eq2_h2 element_ptr_kinds_eq_h2 tag_name_eq2_h2) using get_shadow_root_eq_h get_shadow_root_eq_h2 by (smt \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(2) element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok option.inject returns_result_select_result select_result_I2 shadow_root_ptr_kinds_commutes) ultimately show "heap_is_wellformed h'" by(simp add: heap_is_wellformed_def) qed end interpretation i_remove_shadow_root_wf?: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_tag_name get_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove_shadow_root remove_shadow_root_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_root\_node\ interpretation i_get_root_node_wf?: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[1] using get_ancestors_never_empty apply blast using get_ancestors_ok apply blast using get_ancestors_reads apply blast using get_ancestors_ptrs_in_heap apply blast using get_ancestors_remains_not_in_ancestors apply blast using get_ancestors_also_parent apply blast using get_ancestors_obtains_children apply blast using get_ancestors_parent_child_rel apply blast using get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1] using get_root_node_ok apply blast using get_root_node_ptr_in_heap apply blast using get_root_node_root_in_heap apply blast using get_ancestors_same_root_node apply(blast, blast) using get_root_node_same_no_parent apply blast (* using get_root_node_not_node_same apply blast *) using get_root_node_parent_same apply (blast, blast) done subsubsection \get\_parent\_get\_host\ locale l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_shadow_root + l_get_host + l_get_child_nodes begin lemma host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)" proof - have "a_host_shadow_root_rel h = (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast host, cast shadow_root)} | None \ {}))" by(auto simp add: a_host_shadow_root_rel_def split: option.splits) moreover have "finite (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root)} | None \ {}))" by(auto split: option.splits) ultimately show ?thesis by auto qed lemma host_shadow_root_rel_shadow_root: "h \ get_shadow_root host \\<^sub>r shadow_root_option \ shadow_root_option = Some shadow_root \ ((cast host, cast shadow_root) \ a_host_shadow_root_rel h)" apply(auto simp add: a_host_shadow_root_rel_def)[1] by(metis (mono_tags, lifting) case_prodI is_OK_returns_result_I l_get_shadow_root.get_shadow_root_ptr_in_heap local.l_get_shadow_root_axioms mem_Collect_eq pair_imageI select_result_I2) lemma host_shadow_root_rel_host: "heap_is_wellformed h \ h \ get_host shadow_root \\<^sub>r host \ (cast host, cast shadow_root) \ a_host_shadow_root_rel h" apply(auto simp add: a_host_shadow_root_rel_def)[1] using shadow_root_host_dual by (metis (no_types, lifting) Collect_cong host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def split_cong) lemma heap_wellformed_induct_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ P parent" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h)" using host_shadow_root_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf ((parent_child_rel h \ a_host_shadow_root_rel h)\)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf_converse local.CD.parent_child_rel_impl) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using assms host_shadow_root_rel_shadow_root local.CD.parent_child_rel_child by blast qed qed lemma heap_wellformed_induct_rev_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ P child" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h)" using host_shadow_root_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf (parent_child_rel h \ a_host_shadow_root_rel h)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using parent_child_rel_parent host_shadow_root_rel_host using assms(1) assms(2) by auto qed qed end interpretation i_get_parent_get_host_wf?: l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_parent_get_host_wf = l_heap_is_wellformed_defs + l_get_parent_defs + l_get_shadow_root_defs + l_get_host_defs + l_get_child_nodes_defs + assumes heap_wellformed_induct_si [consumes 1, case_names step]: "heap_is_wellformed h \ (\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ P parent) \ P ptr" assumes heap_wellformed_induct_rev_si [consumes 1, case_names step]: "heap_is_wellformed h \ (\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ P child) \ P ptr" lemma l_get_parent_get_host_wf_is_get_parent_get_host_wf [instances]: "l_get_parent_get_host_wf heap_is_wellformed get_parent get_shadow_root get_host get_child_nodes" apply(auto simp add: l_get_parent_get_host_wf_def instances)[1] using heap_wellformed_induct_si apply metis using heap_wellformed_induct_rev_si apply blast done subsubsection \get\_host\ locale l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma get_host_ok [simp]: assumes "heap_is_wellformed h" assumes "type_wf h" assumes "known_ptrs h" assumes "shadow_root_ptr |\| shadow_root_ptr_kinds h" shows "h \ ok (get_host shadow_root_ptr)" proof - obtain host where host: "host |\| element_ptr_kinds h" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms(1) assms(4) get_shadow_root_ok assms(2) apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] by (smt finite_set_in returns_result_select_result) obtain host_candidates where host_candidates: "h \ filter_M (\element_ptr. Heap_Error_Monad.bind (get_shadow_root element_ptr) (\shadow_root_opt. return (shadow_root_opt = Some shadow_root_ptr))) (sorted_list_of_set (fset (element_ptr_kinds h))) \\<^sub>r host_candidates" apply(drule is_OK_returns_result_E[rotated]) using get_shadow_root_ok assms(2) by(auto intro!: filter_M_is_OK_I bind_pure_I bind_is_OK_I2) then have "host_candidates = [host]" apply(rule filter_M_ex1) apply (simp add: host) apply (smt assms(1) assms(2) bind_pure_returns_result_I2 bind_returns_result_E finite_set_in host local.get_shadow_root_ok local.get_shadow_root_pure local.shadow_root_same_host return_returns_result returns_result_eq shadow_root sorted_list_of_fset.rep_eq sorted_list_of_fset_simps(1)) by (simp_all add: assms(2) bind_pure_I bind_pure_returns_result_I2 host local.get_shadow_root_ok returns_result_eq shadow_root) then show ?thesis using host_candidates host assms(1) get_shadow_root_ok apply(auto simp add: get_host_def known_ptrs_known_ptr intro!: bind_is_OK_pure_I filter_M_pure_I filter_M_is_OK_I bind_pure_I split: list.splits)[1] using assms(2) apply blast apply (meson list.distinct(1) returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_host_wf?: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptr known_ptrs type_wf get_host get_host_locs get_shadow_root get_shadow_root_locs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_host_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_host_defs + assumes get_host_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_host shadow_root_ptr)" lemma get_host_wf_is_l_get_host_wf [instances]: "l_get_host_wf heap_is_wellformed known_ptr known_ptrs type_wf get_host" by(auto simp add: l_get_host_wf_def l_get_host_wf_axioms_def instances) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent_get_host_wf + l_get_host_wf begin lemma get_root_node_si_ptr_in_heap: assumes "h \ ok (get_root_node_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms unfolding get_root_node_si_def using get_ancestors_si_ptr_in_heap by auto lemma get_ancestors_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_ancestors_si ptr)" proof (insert assms(1) assms(4), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using assms(2) assms(3) apply(auto simp add: get_ancestors_si_def[of child] assms(1) get_parent_parent_in_heap intro!: bind_is_OK_pure_I split: option.splits)[1] using local.get_parent_ok apply blast using get_host_ok assms(1) apply blast by (meson assms(1) is_OK_returns_result_I local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual) qed lemma get_ancestors_si_remains_not_in_ancestors: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "h' \ get_ancestors_si ptr \\<^sub>r ancestors'" and "\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>r children' \ set children' \ set children" and "\p shadow_root_option shadow_root_option'. h \ get_shadow_root p \\<^sub>r shadow_root_option \ h' \ get_shadow_root p \\<^sub>r shadow_root_option' \ (if shadow_root_option = None then shadow_root_option' = None else shadow_root_option' = None \ shadow_root_option' = shadow_root_option)" and "node \ set ancestors" and object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "node \ set ancestors'" proof - have object_ptr_kinds_M_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" using object_ptr_kinds_eq3 by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) show ?thesis proof (insert assms(1) assms(3) assms(4) assms(7), induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev_si) case (step child) obtain ancestors_remains where ancestors_remains: "ancestors = child # ancestors_remains" using \h \ get_ancestors_si child \\<^sub>r ancestors\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) obtain ancestors_remains' where ancestors_remains': "ancestors' = child # ancestors_remains'" using \h' \ get_ancestors_si child \\<^sub>r ancestors'\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) have "child |\| object_ptr_kinds h" using local.get_ancestors_si_ptr_in_heap object_ptr_kinds_eq3 step.prems(2) by fastforce have "node \ child" using ancestors_remains step.prems(3) by auto have 1: "\p parent. h' \ get_parent p \\<^sub>r Some parent \ h \ get_parent p \\<^sub>r Some parent" proof - fix p parent assume "h' \ get_parent p \\<^sub>r Some parent" then obtain children' where children': "h' \ get_child_nodes parent \\<^sub>r children'" and p_in_children': "p \ set children'" using get_parent_child_dual by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using get_child_nodes_ok assms(1) get_child_nodes_ptr_in_heap object_ptr_kinds_eq children' known_ptrs using type_wf type_wf' by (metis \h' \ get_parent p \\<^sub>r Some parent\ get_parent_parent_in_heap is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) have "p \ set children" using assms(5) children children' p_in_children' by blast then show "h \ get_parent p \\<^sub>r Some parent" using child_parent_dual assms(1) children known_ptrs type_wf by blast qed have 2: "\p host. h' \ get_host p \\<^sub>r host \ h \ get_host p \\<^sub>r host" proof - fix p host assume "h' \ get_host p \\<^sub>r host" then have "h' \ get_shadow_root host \\<^sub>r Some p" using local.shadow_root_host_dual by blast then have "h \ get_shadow_root host \\<^sub>r Some p" by (metis assms(6) element_ptr_kinds_commutes is_OK_returns_result_I local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq3 option.distinct(1) returns_result_select_result type_wf) then show "h \ get_host p \\<^sub>r host" by (metis assms(1) is_OK_returns_result_E known_ptrs local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host type_wf) qed show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?thesis using step(3) step(4) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] by (metis "2" assms(1) l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.shadow_root_same_host list.set_intros(2) local.l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.shadow_root_host_dual step.hyps(2) step.prems(3) type_wf) next case (Some node_child) then show ?thesis using step(3) step(4) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] apply (meson "1" option.distinct(1) returns_result_eq) by (metis "1" list.set_intros(2) option.inject returns_result_eq step.hyps(1) step.prems(3)) qed qed qed lemma get_ancestors_si_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" shows "ptr' |\| object_ptr_kinds h" proof (insert assms(4) assms(5), induct ancestors arbitrary: ptr) case Nil then show ?case by(auto) next case (Cons a ancestors) then obtain x where x: "h \ get_ancestors_si x \\<^sub>r a # ancestors" by(auto simp add: get_ancestors_si_def[of a] elim!: bind_returns_result_E2 split: option.splits) then have "x = a" by(auto simp add: get_ancestors_si_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?case proof (cases "ptr' = a") case True then show ?thesis using Cons.hyps Cons.prems(2) get_ancestors_si_ptr_in_heap x using \x = a\ by blast next case False obtain ptr'' where ptr'': "h \ get_ancestors_si ptr'' \\<^sub>r ancestors" using \ h \ get_ancestors_si x \\<^sub>r a # ancestors\ Cons.prems(2) False by(auto simp add: get_ancestors_si_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?thesis using Cons.hyps Cons.prems(2) False by auto qed qed lemma get_ancestors_si_reads: assumes "heap_is_wellformed h" shows "reads get_ancestors_si_locs (get_ancestors_si node_ptr) h h'" proof (insert assms(1), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def] get_host_reads[unfolded reads_def] apply(simp (no_asm) add: get_ancestors_si_def) by(auto simp add: get_ancestors_si_locs_def get_parent_reads_pointers intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF return_reads] reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads] reads_subset[OF get_host_reads] split: option.splits) qed lemma get_ancestors_si_subset: assumes "heap_is_wellformed h" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors_si ancestor \\<^sub>r ancestor_ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "set ancestor_ancestors \ set ancestors" proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev_si) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_si_ptr_in_heap step(3) by auto (* then have "h \ check_in_heap child \\<^sub>r ()" using returns_result_select_result by force *) obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using step(3) by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_I elim!: bind_returns_result_E2 split: option.splits) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?case using step(3) \None = cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2)[1] by (metis (no_types, lifting) assms(4) empty_iff empty_set select_result_I2 set_ConsD step.prems(1) step.prems(2)) next case (Some shadow_root_child) then have "shadow_root_child |\| shadow_root_ptr_kinds h" using \child |\| object_ptr_kinds h\ by (metis (no_types, lifting) shadow_root_ptr_casts_commute shadow_root_ptr_kinds_commutes) obtain host where host: "h \ get_host shadow_root_child \\<^sub>r host" using get_host_ok assms by (meson \shadow_root_child |\| shadow_root_ptr_kinds h\ is_OK_returns_result_E) then have "h \ get_ancestors_si (cast host) \\<^sub>r tl_ancestors" using Some step(3) tl_ancestors None by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_returns_result_I elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) then show ?case using step(2) Some host step(4) tl_ancestors by (metis (no_types, lifting) assms(4) dual_order.trans eq_iff returns_result_eq set_ConsD set_subset_Cons shadow_root_ptr_casts_commute step.prems(1)) qed next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs] by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then have "ancestors = [child]" using step(3) s1 apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(3) step(4) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some parent) then have "h \ get_ancestors_si parent \\<^sub>r tl_ancestors" using s1 tl_ancestors step(3) by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case by (metis (no_types, lifting) Some.prems \h \ get_ancestors_si parent \\<^sub>r tl_ancestors\ assms(4) eq_iff node_ptr_casts_commute order_trans s1 select_result_I2 set_ConsD set_subset_Cons step.hyps(1) step.prems(1) step.prems(2) tl_ancestors) qed qed qed lemma get_ancestors_si_also_parent: assumes "heap_is_wellformed h" and "h \ get_ancestors_si some_ptr \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>r Some parent" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "parent \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_si (cast child) \\<^sub>r child_ancestors" by (meson assms(1) assms(4) get_ancestors_si_ok is_OK_returns_result_I known_ptrs local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result type_wf) then have "parent \ set child_ancestors" apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_also_host: assumes "heap_is_wellformed h" and "h \ get_ancestors_si some_ptr \\<^sub>r ancestors" and "cast shadow_root \ set ancestors" and "h \ get_host shadow_root \\<^sub>r host" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "cast host \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_si (cast shadow_root) \\<^sub>r child_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_si_ok get_ancestors_si_ptrs_in_heap is_OK_returns_result_E known_ptrs type_wf) then have "cast host \ set child_ancestors" apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_parent_child_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_si child \\<^sub>r ancestors" assumes "((ptr, child) \ (parent_child_rel h)\<^sup>*)" shows "ptr \ set ancestors" proof (insert assms(5), induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) local.get_ancestors_si_ptr by blast next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h) \ (ptr_child, child) \ (parent_child_rel h)\<^sup>*" using converse_rtranclE[OF 1(3)] \ptr \ child\ by metis then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 CD.parent_child_rel_node_ptr by (metis ) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using CD.parent_child_rel_parent_in_heap ptr_child by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis calculation \known_ptrs h\ local.get_child_nodes_ok local.known_ptrs_known_ptr CD.parent_child_rel_child ptr_child ptr_child_ptr_child_node returns_result_select_result \type_wf h\) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using ptr_child ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using \known_ptrs h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_si_also_parent assms \type_wf h\ by blast qed qed lemma get_ancestors_si_parent_child_host_shadow_root_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_si child \\<^sub>r ancestors" assumes "((ptr, child) \ (parent_child_rel h \ a_host_shadow_root_rel h)\<^sup>*)" shows "ptr \ set ancestors" proof (insert assms(5), induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) local.get_ancestors_si_ptr by blast next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h) \ (ptr_child, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using converse_rtranclE[OF 1(3)] \ptr \ child\ by metis then show ?thesis proof(cases "(ptr, ptr_child) \ parent_child_rel h") case True then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 CD.parent_child_rel_node_ptr by (metis) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using CD.parent_child_rel_parent_in_heap True by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis True assms(2) assms(3) calculation local.CD.parent_child_rel_child local.get_child_nodes_ok local.known_ptrs_known_ptr ptr_child_ptr_child_node returns_result_select_result) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child True ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using \known_ptrs h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_si_also_parent assms \type_wf h\ by blast next case False then obtain host where host: "ptr = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host" using ptr_child by(auto simp add: a_host_shadow_root_rel_def) then obtain shadow_root where shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root" and ptr_child_shadow_root: "ptr_child = cast shadow_root" using ptr_child False apply(auto simp add: a_host_shadow_root_rel_def)[1] by (metis (no_types, lifting) assms(3) local.get_shadow_root_ok select_result_I) moreover have "(cast shadow_root, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child ptr_child_shadow_root by blast ultimately have "cast shadow_root \ set ancestors" using "1.hyps"(2) host by blast moreover have "h \ get_host shadow_root \\<^sub>r host" by (metis assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host shadow_root) ultimately show ?thesis using get_ancestors_si_also_host assms(1) assms(2) assms(3) assms(4) host by blast qed qed qed lemma get_root_node_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_root_node_si ptr)" using assms get_ancestors_si_ok by(auto simp add: get_root_node_si_def) lemma get_root_node_si_root_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "root |\| object_ptr_kinds h" using assms apply(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2)[1] by (simp add: get_ancestors_si_never_empty get_ancestors_si_ptrs_in_heap) lemma get_root_node_si_same_no_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si ptr \\<^sub>r cast child" shows "h \ get_parent child \\<^sub>r None" proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev_si) case (step c) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c") case None then show ?thesis using step(3) by(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: if_splits option.splits intro!: step(2) bind_pure_returns_result_I) next case (Some child_node) note s = this then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using step(3) apply(auto simp add: get_root_node_si_def get_ancestors_si_def intro!: bind_pure_I elim!: bind_returns_result_E2)[1] by(auto split: option.splits) then show ?thesis proof(induct parent_opt) case None then show ?case using Some get_root_node_si_no_parent returns_result_eq step.prems by fastforce next case (Some parent) then show ?case using step(3) s apply(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: option.splits list.splits if_splits)[1] using assms(1) get_ancestors_si_never_empty apply blast by(auto simp add: get_root_node_si_def dest: returns_result_eq intro!: step(1) bind_pure_returns_result_I) qed qed qed end interpretation i_get_root_node_si_wf?: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: instances l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent begin lemma get_disconnected_document_ok: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_parent node_ptr \\<^sub>r None" shows "h \ ok (get_disconnected_document node_ptr)" proof - have "node_ptr |\| node_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap) have "\(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r)" apply(auto)[1] using assms(4) child_parent_dual[OF assms(1)] assms(1) assms(2) assms(3) known_ptrs_known_ptr option.simps(3) returns_result_eq returns_result_select_result by (metis (no_types, lifting) CD.get_child_nodes_ok) then have "(\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using heap_is_wellformed_children_disc_nodes using \node_ptr |\| node_ptr_kinds h\ assms(1) by blast then obtain some_owner_document where "some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" and "node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto have h5: "\!x. x \ set (sorted_list_of_set (fset (document_ptr_kinds h))) \ h \ Heap_Error_Monad.bind (get_disconnected_nodes x) (\children. return (node_ptr \ set children)) \\<^sub>r True" apply(auto intro!: bind_pure_returns_result_I)[1] apply (smt CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure \\document_ptr\fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r\ assms(2) bind_pure_returns_result_I2 notin_fset return_returns_result select_result_I2) apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1] using heap_is_wellformed_one_disc_parent assms(1) by blast let ?filter_M = "filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h)))" have "h \ ok (?filter_M)" using CD.get_disconnected_nodes_ok by (smt CD.get_disconnected_nodes_pure DocumentMonad.ptr_kinds_M_ptr_kinds DocumentMonad.ptr_kinds_ptr_kinds_M assms(2) bind_is_OK_pure_I bind_pure_I document_ptr_kinds_M_def filter_M_is_OK_I l_ptr_kinds_M.ptr_kinds_M_ok return_ok return_pure returns_result_select_result) then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by auto have "candidates = [some_owner_document]" apply(rule filter_M_ex1[OF candidates \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ h5]) using \node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ by(auto simp add: CD.get_disconnected_nodes_ok assms(2) intro!: bind_pure_I intro!: bind_pure_returns_result_I) then show ?thesis using candidates \node_ptr |\| node_ptr_kinds h\ apply(auto simp add: get_disconnected_document_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I split: list.splits)[1] apply (meson not_Cons_self2 returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_disconnected_document_wf?: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_owner\_document\ locale l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes + l_get_child_nodes + l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_known_ptrs + l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" begin lemma get_owner_document_disconnected_nodes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" assumes known_ptrs: "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" proof - have 2: "node_ptr |\| node_ptr_kinds h" using assms apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.a_all_ptrs_in_heap_def)[1] using assms(1) local.heap_is_wellformed_disc_nodes_in_heap by blast have 3: "document_ptr |\| document_ptr_kinds h" using assms(2) get_disconnected_nodes_ptr_in_heap by blast then have 4: "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using CD.distinct_lists_no_parent assms unfolding heap_is_wellformed_def CD.heap_is_wellformed_def by simp moreover have "(\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using assms(1) 2 "3" assms(2) assms(3) by auto ultimately have 0: "\!document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using concat_map_distinct assms(1) known_ptrs_implies by (smt CD.heap_is_wellformed_one_disc_parent DocumentMonad.ptr_kinds_ptr_kinds_M disjoint_iff_not_equal local.get_disconnected_nodes_ok local.heap_is_wellformed_def returns_result_select_result type_wf) have "h \ get_parent node_ptr \\<^sub>r None" using 4 2 apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I )[1] apply(auto intro!: filter_M_empty_I bind_pure_I bind_pure_returns_result_I)[1] using get_child_nodes_ok assms(4) type_wf by (metis get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have 4: "h \ get_root_node_si (cast node_ptr) \\<^sub>r cast node_ptr" using get_root_node_si_no_parent by simp obtain document_ptrs where document_ptrs: "h \ document_ptr_kinds_M \\<^sub>r document_ptrs" by simp then have "h \ ok (filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs)" using assms(1) get_disconnected_nodes_ok type_wf by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I) then obtain candidates where candidates: "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r candidates" by auto have filter: "filter (\document_ptr. |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \ cast ` set disconnected_nodes) }|\<^sub>r) document_ptrs = [document_ptr]" apply(rule filter_ex1) using 0 document_ptrs apply(simp)[1] apply (smt "0" "3" assms bind_is_OK_pure_I bind_pure_returns_result_I bind_pure_returns_result_I2 bind_returns_result_E2 bind_returns_result_E3 document_ptr_kinds_M_def get_disconnected_nodes_ok get_disconnected_nodes_pure image_eqI is_OK_returns_result_E l_ptr_kinds_M.ptr_kinds_ptr_kinds_M return_ok return_returns_result returns_result_eq select_result_E select_result_I select_result_I2 select_result_I2) using assms(2) assms(3) apply (metis (no_types, lifting) bind_pure_returns_result_I2 is_OK_returns_result_I local.get_disconnected_nodes_pure node_ptr_inclusion return_returns_result select_result_I2) using document_ptrs 3 apply(simp) using document_ptrs by simp have "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r [document_ptr]" apply(rule filter_M_filter2) using get_disconnected_nodes_ok document_ptrs 3 assms(1) type_wf filter by(auto intro: bind_pure_I bind_is_OK_I2) with 4 document_ptrs have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r document_ptr" by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits) moreover have "known_ptr (cast node_ptr)" using known_ptrs_known_ptr[OF known_ptrs, where ptr="cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"] 2 known_ptrs_implies by(simp) ultimately show ?thesis using 2 apply(auto simp add: CD.a_get_owner_document_tups_def get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_shadow_root_ptr) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto split: option.splits intro!: bind_pure_returns_result_I) qed lemma in_disconnected_nodes_no_parent: assumes "heap_is_wellformed h" assumes "h \ get_parent node_ptr \\<^sub>r None" assumes "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" assumes "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assumes "known_ptrs h" assumes "type_wf h" shows "node_ptr \ set disc_nodes" proof - have "\parent. parent |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent|\<^sub>r" using assms(2) by (meson get_child_nodes_ok assms(1) assms(5) assms(6) local.child_parent_dual local.known_ptrs_known_ptr option.distinct(1) returns_result_eq returns_result_select_result) then show ?thesis by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok local.get_owner_document_disconnected_nodes local.get_parent_ptr_in_heap local.heap_is_wellformed_children_disc_nodes returns_result_select_result select_result_I2) qed lemma get_owner_document_owner_document_in_heap_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" proof - obtain root where root: "h \ get_root_node_si (cast node_ptr) \\<^sub>r root" using assms(4) by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) then show ?thesis proof (cases "is_document_ptr root") case True then show ?thesis using assms(4) root apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using assms document_ptr_kinds_commutes get_root_node_si_root_in_heap by blast next case False have "known_ptr root" using assms local.get_root_node_si_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using assms local.get_root_node_si_root_in_heap by blast have "\is_shadow_root_ptr root" using root using local.get_root_node_si_root_not_shadow_root by blast then have "is_node_ptr_kind root" using False \known_ptr root\ \root |\| object_ptr_kinds h\ apply(simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs) using is_node_ptr_kind_none by force then have "(\document_ptr \ fset (document_ptr_kinds h). root \ cast ` set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using local.child_parent_dual local.get_child_nodes_ok local.get_root_node_si_same_no_parent local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq returns_result_select_result root by (metis (no_types, lifting) assms \root |\| object_ptr_kinds h\) then obtain some_owner_document where "some_owner_document |\| document_ptr_kinds h" and "root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by (metis (no_types, lifting) assms bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset return_ok return_pure sorted_list_of_set(1)) then have "some_owner_document \ set candidates" apply(rule filter_M_in_result_if_ok) using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto simp add: assms local.get_disconnected_nodes_ok intro!: bind_pure_I bind_pure_returns_result_I)[1] done then have "candidates \ []" by auto then have "owner_document \ set candidates" using assms(4) root apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis candidates list.set_sel(1) returns_result_eq) by (metis \is_node_ptr_kind root\ node_ptr_no_document_ptr_cast returns_result_eq) then show ?thesis using candidates by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure) qed qed lemma get_owner_document_owner_document_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" using assms apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_split_asm)+ proof - assume "h \ invoke [] ptr () \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by (meson invoke_empty is_OK_returns_result_I) next assume "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 6: "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 7: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" apply(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: filter_M_pure_I bind_pure_I elim!: bind_returns_result_E2 split: if_splits option.splits)[1] using get_owner_document_owner_document_in_heap_node by blast qed lemma get_owner_document_ok: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_owner_document ptr)" proof - have "known_ptr ptr" using assms(2) assms(4) local.known_ptrs_known_ptr by blast then show ?thesis apply(simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def) apply(split invoke_splits, (rule conjI | rule impI)+)+ proof - assume 0: "known_ptr ptr" and 1: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 2: "\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 3: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" then show "h \ ok invoke [] ptr ()" using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_shadow_root_ptr known_ptr_not_element_ptr known_ptr_impl by blast next assume 0: "known_ptr ptr" and 1: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 2: "\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 3: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" then show "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) next show "is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(4) by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits) next show "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) next show "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) qed qed end interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs DocumentClass.known_ptr get_parent get_parent_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_host get_host_locs get_owner_document get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptrs get_ancestors_si get_ancestors_si_locs by(auto simp add: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]: "l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes get_owner_document get_parent" apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def instances)[1] using get_owner_document_disconnected_nodes apply fast using in_disconnected_nodes_no_parent apply fast using get_owner_document_owner_document_in_heap apply fast using get_owner_document_ok apply fast done subsubsection \remove\_child\ locale l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_disconnected_nodes + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + CD: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "type_wf h'" using CD.remove_child_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_child_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_child_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_child_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "heap_is_wellformed h'" proof - have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using CD.remove_child_heap_is_wellformed_preserved(3) assms unfolding heap_is_wellformed_def by auto have shadow_root_eq: "\ptr' shadow_root_ptr_opt. h \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2: "\ptr'. |h \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq: "\ptr' tag. h \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(4)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq: "element_ptr_kinds h = element_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "parent_child_rel h' \ parent_child_rel h" using \heap_is_wellformed h\ heap_is_wellformed_def using CD.remove_child_parent_child_rel_subset using \known_ptrs h\ \type_wf h\ assms(4) by simp show ?thesis using \heap_is_wellformed h\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ \parent_child_rel h' \ parent_child_rel h\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def object_ptr_kinds_eq element_ptr_kinds_eq shadow_root_ptr_kinds_eq shadow_root_eq shadow_root_eq2 tag_name_eq tag_name_eq2)[1] by (meson acyclic_subset order_refl sup_mono) qed lemma remove_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "type_wf h'" using CD.remove_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "heap_is_wellformed h'" using assms by(auto simp add: remove_def elim!: bind_returns_heap_E2 intro: remove_child_heap_is_wellformed_preserved split: option.splits) lemma remove_child_removes_child: "heap_is_wellformed h \ h \ remove_child ptr' child \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children \ known_ptrs h \ type_wf h \ child \ set children" using CD.remove_child_removes_child local.heap_is_wellformed_def by blast lemma remove_child_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove_child ptr node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" using CD.remove_child_removes_first_child local.heap_is_wellformed_def by blast lemma remove_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" using CD.remove_removes_child local.heap_is_wellformed_def by blast lemma remove_for_all_empty_children: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r []" using CD.remove_for_all_empty_children local.heap_is_wellformed_def by blast end interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr get_parent get_parent_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_owner_document known_ptrs get_ancestors_si get_ancestors_si_locs set_child_nodes set_child_nodes_locs remove_child remove_child_locs remove by(auto simp add: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma remove_child_wf2_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using remove_child_preserves_type_wf apply fast using remove_child_preserves_known_ptrs apply fast using remove_child_heap_is_wellformed_preserved apply (fast) using remove_preserves_type_wf apply fast using remove_preserves_known_ptrs apply fast using remove_heap_is_wellformed_preserved apply (fast) using remove_child_removes_child apply fast using remove_child_removes_first_child apply fast using remove_removes_child apply fast using remove_for_all_empty_children apply fast done subsubsection \adopt\_node\ locale l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_set_disconnected_nodes_get_child_nodes + l_get_owner_document_wf + l_remove_child_wf2 + l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_child: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h2" and children: "h2 \ get_child_nodes ptr \\<^sub>r children" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set children" proof - obtain old_document parent_opt h' where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h': "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return () ) \\<^sub>h h'" using adopt_node by(auto simp add: adopt_node_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: if_splits) then have "h' \ get_child_nodes ptr \\<^sub>r children" using adopt_node apply(auto simp add: adopt_node_def dest!: bind_returns_heap_E3[rotated, OF old_document, rotated] bind_returns_heap_E3[rotated, OF parent_opt, rotated] elim!: bind_returns_heap_E4[rotated, OF h', rotated])[1] apply(auto split: if_splits elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] apply (simp add: set_disconnected_nodes_get_child_nodes children reads_writes_preserved[OF get_child_nodes_reads set_disconnected_nodes_writes]) using children by blast show ?thesis proof(insert parent_opt h', induct parent_opt) case None then show ?case using child_parent_dual wellformed known_ptrs type_wf \h' \ get_child_nodes ptr \\<^sub>r children\ returns_result_eq by fastforce next case (Some option) then show ?case using remove_child_removes_child \h' \ get_child_nodes ptr \\<^sub>r children\ known_ptrs type_wf wellformed by auto qed qed lemma adopt_node_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ adopt_node document_ptr child \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if document_ptr \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 child old_disc_nodes); disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (child # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have object_ptr_kinds_h_eq3: "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "type_wf h2" using h2 remove_child_preserves_type_wf assms by(auto split: option.splits) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs assms by(auto split: option.splits) then have "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" proof(cases "document_ptr = old_document") case True then show "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" using h' wellformed_h2 \known_ptrs h2\ \type_wf h2\ by auto next case False then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where docs_neq: "document_ptr \ old_document" and old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>h h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h2: "node_ptr_kinds h2 = node_ptr_kinds h3" by auto have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto have children_eq_h2: "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr. |h2 \ get_child_nodes ptr|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have object_ptr_kinds_h3_eq3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h3: "node_ptr_kinds h3 = node_ptr_kinds h'" by auto have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h3: "document_ptr_kinds h3 = document_ptr_kinds h'" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto have children_eq_h3: "\ptr children. h3 \ get_child_nodes ptr \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr. |h3 \ get_child_nodes ptr|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. old_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force obtain disc_nodes_old_document_h2 where disc_nodes_old_document_h2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" using old_disc_nodes by blast then have disc_nodes_old_document_h3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" using h3 old_disc_nodes returns_result_eq set_disconnected_nodes_get_disconnected_nodes by fastforce have "distinct disc_nodes_old_document_h2" using disc_nodes_old_document_h2 local.heap_is_wellformed_disconnected_nodes_distinct wellformed_h2 by blast have "type_wf h2" proof (insert h2, induct parent_opt) case None then show ?case using type_wf by simp next case (Some option) then show ?case using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes] type_wf remove_child_types_preserved by (simp add: reflp_def transp_def) qed then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" using disconnected_nodes_eq_h2 docs_neq disc_nodes_document_ptr_h3 by auto have disc_nodes_document_ptr_h': "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" using h' disc_nodes_document_ptr_h3 using set_disconnected_nodes_get_disconnected_nodes by blast have document_ptr_in_heap: "document_ptr |\| document_ptr_kinds h2" using disc_nodes_document_ptr_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using disc_nodes_document_ptr_h2 get_disconnected_nodes_ptr_in_heap by blast have old_document_in_heap: "old_document |\| document_ptr_kinds h2" using disc_nodes_old_document_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using get_disconnected_nodes_ptr_in_heap old_disc_nodes by blast have "child \ set disc_nodes_old_document_h2" proof (insert parent_opt h2, induct parent_opt) case None then have "h = h2" by(auto) moreover have "CD.a_owner_document_valid h" using assms(1) by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) ultimately show ?case using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)] in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast next case (Some option) then show ?case apply(simp split: option.splits) using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs by blast qed have "child \ set (remove1 child disc_nodes_old_document_h2)" using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \distinct disc_nodes_old_document_h2\ by auto have "child \ set disc_nodes_document_ptr_h3" proof - have "CD.a_distinct_lists h2" using heap_is_wellformed_def CD.heap_is_wellformed_def wellformed_h2 by blast then have 0: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>r))" by(simp add: CD.a_distinct_lists_def) show ?thesis using distinct_concat_map_E(1)[OF 0] \child \ set disc_nodes_old_document_h2\ disc_nodes_old_document_h2 disc_nodes_document_ptr_h2 by (meson \type_wf h2\ docs_neq known_ptrs local.get_owner_document_disconnected_nodes local.known_ptrs_preserved object_ptr_kinds_h_eq3 returns_result_eq wellformed_h2) qed have child_in_heap: "child |\| node_ptr_kinds h" using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]] node_ptr_kinds_commutes by blast have "CD.a_acyclic_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h2" proof fix x assume "x \ parent_child_rel h'" then show "x \ parent_child_rel h2" using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3 mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong unfolding CD.parent_child_rel_def by(simp) qed then have " CD.a_acyclic_heap h'" using \ CD.a_acyclic_heap h2\ CD.acyclic_heap_def acyclic_subset by blast moreover have " CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h3" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1] apply (metis \type_wf h'\ children_eq2_h3 children_eq_h2 children_eq_h3 known_ptrs l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms node_ptr_kinds_eq3_h2 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result wellformed_h2) by (metis (no_types, hide_lams) disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 finite_set_in select_result_I2 set_remove1_subset subsetD) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1] apply (metis (no_types, hide_lams) children_eq2_h3 finite_set_in object_ptr_kinds_h3_eq3 subsetD) by (metis (no_types, hide_lams) \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3 finite_set_in local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 select_result_I2 set_ConsD subsetD wellformed_h2) moreover have "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(simp add: CD.a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 ) by (metis (no_types) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2) have a_distinct_lists_h2: "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2 children_eq2_h2 children_eq2_h3)[1] proof - assume 1: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" proof(rule distinct_concat_map_I) show "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))" by(auto simp add: document_ptr_kinds_M_def ) next fix x assume a1: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" have 4: "distinct |h2 \ get_disconnected_nodes x|\<^sub>r" using a_distinct_lists_h2 "2" a1 concat_map_all_distinct document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 by fastforce then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "old_document \ x") case True then show ?thesis proof (cases "document_ptr \ x") case True then show ?thesis using disconnected_nodes_eq2_h2[OF \old_document \ x\] disconnected_nodes_eq2_h3[OF \document_ptr \ x\] 4 by(auto) next case False then show ?thesis using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4 \child \ set disc_nodes_document_ptr_h3\ by(auto simp add: disconnected_nodes_eq2_h2[OF \old_document \ x\] ) qed next case False then show ?thesis by (metis (no_types, hide_lams) \distinct disc_nodes_old_document_h2\ disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 distinct_remove1 docs_neq select_result_I2) qed next fix x y assume a0: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a1: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a2: "x \ y" moreover have 5: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>r = {}" using 2 calculation by (auto simp add: document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 dest: distinct_concat_map_E(1)) ultimately show "set |h' \ get_disconnected_nodes x|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>r = {}" proof(cases "old_document = x") case True have "old_document \ y" using \x \ y\ \old_document = x\ by simp have "document_ptr \ x" using docs_neq \old_document = x\ by auto show ?thesis proof(cases "document_ptr = y") case True then show ?thesis using 5 True select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document = x\ by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ \document_ptr \ x\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1 set_ConsD) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \old_document = x\ docs_neq \old_document \ y\ by (metis (no_types, lifting) disjoint_iff_not_equal notin_set_remove1) qed next case False then show ?thesis proof(cases "old_document = y") case True then show ?thesis proof(cases "document_ptr = x") case True show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr = x\ apply(simp) by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr \ x\ by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal docs_neq notin_set_remove1) qed next case False have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False \type_wf h2\ a1 disc_nodes_old_document_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result wellformed_h2) then show ?thesis proof(cases "document_ptr = x") case True then have "document_ptr \ y" using \x \ y\ by auto have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by blast then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document \ y\ \document_ptr = x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by(auto) next case False then show ?thesis proof(cases "document_ptr = y") case True have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set disc_nodes_document_ptr_h3 = {}" using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \document_ptr \ x\ select_result_I2[OF disc_nodes_document_ptr_h3, symmetric] disconnected_nodes_eq2_h2[OF docs_neq[symmetric], symmetric] by (simp add: "5" True) moreover have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>r = {}" using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \old_document \ x\ by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 finite_fset fmember.rep_eq set_sorted_list_of_set) ultimately show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr = y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by auto next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by (metis \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ empty_iff inf.idem) qed qed qed qed qed next fix x xa xb assume 0: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" then show False using \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 old_document_in_heap apply(auto)[1] apply(cases "xb = old_document") proof - assume a1: "xb = old_document" assume a2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f6: "old_document |\| document_ptr_kinds h'" using a1 \xb |\| document_ptr_kinds h'\ by blast have f7: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a2 by simp have "x \ set disc_nodes_old_document_h2" using f6 a3 a1 by (metis (no_types) \type_wf h'\ \x \ set |h' \ get_disconnected_nodes xb|\<^sub>r\ disconnected_nodes_eq_h3 docs_neq get_disconnected_nodes_ok returns_result_eq returns_result_select_result set_remove1_subset subsetCE) then have "set |h' \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using f7 f6 a5 a4 \xa |\| object_ptr_kinds h'\ by fastforce then show ?thesis using \x \ set disc_nodes_old_document_h2\ a1 a4 f7 by blast next assume a1: "xb \ old_document" assume a2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'" assume a10: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>r \ h2 \ ok get_disconnected_nodes d" using a9 \type_wf h2\ get_disconnected_nodes_ok by simp then have f14: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) then have "x = child" using f13 a11 a10 a7 a5 a2 a1 by (metis (no_types, lifting) select_result_I2 set_ConsD) then have "child \ set disc_nodes_old_document_h2" using f14 a12 a8 a6 a4 by (metis \type_wf h'\ adopt_node_removes_child assms(1) assms(2) type_wf get_child_nodes_ok known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result) then show ?thesis using \child \ set disc_nodes_old_document_h2\ by fastforce qed qed ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using \CD.a_owner_document_valid h'\ CD.heap_is_wellformed_def by simp have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "known_ptrs h3" using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 by blast then have "known_ptrs h'" using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast show "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ \known_ptrs h'\ \type_wf h'\ using \parent_child_rel h' \ parent_child_rel h2\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3 CD.parent_child_rel_def children_eq2_h2 children_eq2_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3)[1] done qed then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" by auto qed lemma adopt_node_node_in_disconnected_nodes: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h'" and "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set disc_nodes" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node_ptr # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) show ?thesis proof (cases "owner_document = old_document") case True then show ?thesis proof (insert parent_opt h2, induct parent_opt) case None then have "h = h'" using h2 h' by(auto) then show ?case using in_disconnected_nodes_no_parent assms None old_document by blast next case (Some parent) then show ?case using remove_child_in_disconnected_nodes known_ptrs True h' assms(3) old_document by auto qed next case False then show ?thesis using assms(3) h' list.set_intros(1) select_result_I2 set_disconnected_nodes_get_disconnected_nodes apply(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] proof - fix x and h'a and xb assume a1: "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assume a2: "\h document_ptr disc_nodes h'. h \ set_disconnected_nodes document_ptr disc_nodes \\<^sub>h h' \ h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume "h'a \ set_disconnected_nodes owner_document (node_ptr # xb) \\<^sub>h h'" then have "node_ptr # xb = disc_nodes" using a2 a1 by (meson returns_result_eq) then show ?thesis by (meson list.set_intros(1)) qed qed qed end interpretation l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M Shadow_DOM.get_owner_document Shadow_DOM.get_parent Shadow_DOM.get_parent_locs Shadow_DOM.remove_child Shadow_DOM.remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs Shadow_DOM.adopt_node Shadow_DOM.adopt_node_locs ShadowRootClass.known_ptr ShadowRootClass.type_wf Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs ShadowRootClass.known_ptrs Shadow_DOM.set_child_nodes Shadow_DOM.set_child_nodes_locs Shadow_DOM.remove Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel by(auto simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_root_node get_root_node_locs get_parent get_parent_locs known_ptrs get_owner_document remove_child remove_child_locs remove adopt_node adopt_node_locs by(auto simp add: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma adopt_node_wf_is_l_adopt_node_wf [instances]: "l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes known_ptrs adopt_node" apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def instances)[1] using adopt_node_preserves_wellformedness apply blast using adopt_node_removes_child apply blast using adopt_node_node_in_disconnected_nodes apply blast using adopt_node_removes_first_child apply blast using adopt_node_document_in_heap apply blast using adopt_node_preserves_wellformedness apply blast using adopt_node_preserves_wellformedness apply blast done subsubsection \insert\_before\ locale l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_set_disconnected_nodes_get_disconnected_nodes + l_set_child_nodes_get_disconnected_nodes + l_set_disconnected_nodes_get_disconnected_nodes_wf + l_set_disconnected_nodes_get_ancestors_si + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ get_ancestors_si get_ancestors_si_locs + l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document + l_adopt_node + l_adopt_node_wf + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_get_shadow_root begin lemma insert_before_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ insert_before ptr node child \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors_si ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and reference_child: "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" (* children: "h3 \ get_child_nodes ptr \\<^sub>r children" and *) (* h': "h3 \ set_child_nodes ptr (insert_before_list node reference_child children) \\<^sub>h h'" *) using assms(4) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "type_wf h2" using \type_wf h\ using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using adopt_node_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF insert_node_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have "object_ptr_kinds h = object_ptr_kinds h2" using adopt_node_writes h2 apply(rule writes_small_big) using adopt_node_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h3" using set_disconnected_nodes_writes h3 apply(rule writes_small_big) using set_disconnected_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h'" using insert_node_writes h' apply(rule writes_small_big) using set_child_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) ultimately show "known_ptrs h'" using \known_ptrs h\ known_ptrs_preserved by blast have "known_ptrs h2" using \known_ptrs h\ known_ptrs_preserved \object_ptr_kinds h = object_ptr_kinds h2\ by blast then have "known_ptrs h3" using known_ptrs_preserved \object_ptr_kinds h2 = object_ptr_kinds h3\ by blast have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I \known_ptrs h\ l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF \heap_is_wellformed h\ h2] \known_ptrs h\ \type_wf h\ . have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) unfolding a_remove_child_locs_def using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have shadow_root_eq_h2: "\ptr' shadow_root. h \ get_shadow_root ptr' \\<^sub>r shadow_root = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root" using get_shadow_root_reads adopt_node_writes h2 apply(rule reads_writes_preserved) using local.adopt_node_get_shadow_root by blast have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. owner_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ owner_document \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>r remove1 node disconnected_nodes_h2" using h3 set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) using set_child_nodes_get_disconnected_nodes by fast then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) by (auto simp add: set_child_nodes_get_child_nodes_different_pointers) then have children_eq2_h3: "\ptr'. ptr \ ptr' \ |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>r insert_before_list node reference_child children_h3" using h' \type_wf h3\ \known_ptr ptr\ by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2 dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3]) have ptr_in_heap: "ptr |\| object_ptr_kinds h3" using children_h3 get_child_nodes_ptr_in_heap by blast have node_in_heap: "node |\| node_ptr_kinds h" using h2 adopt_node_child_in_heap by fast have child_not_in_any_children: "\p children. h2 \ get_child_nodes p \\<^sub>r children \ node \ set children" using \heap_is_wellformed h\ h2 adopt_node_removes_child \type_wf h\ \known_ptrs h\ by auto have "node \ set disconnected_nodes_h2" using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1) \type_wf h\ \known_ptrs h\ by blast have node_not_in_disconnected_nodes: "\d. d |\| document_ptr_kinds h3 \ node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof (cases "d = owner_document") case True then show ?thesis using disconnected_nodes_h2 wellformed_h2 h3 remove_from_disconnected_nodes_removes wellformed_h2 \d |\| document_ptr_kinds h3\ disconnected_nodes_h3 by fastforce next case False then have "set |h2 \ get_disconnected_nodes d|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>r = {}" using distinct_concat_map_E(1) wellformed_h2 by (metis (no_types, lifting) \d |\| document_ptr_kinds h3\ \type_wf h2\ disconnected_nodes_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result select_result_I2) then show ?thesis using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ disconnected_nodes_h2 by fastforce qed qed have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast obtain ancestors_h2 where ancestors_h2: "h2 \ get_ancestors_si ptr \\<^sub>r ancestors_h2" using get_ancestors_si_ok by (metis \known_ptrs h2\ \type_wf h2\ is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2) have ancestors_h3: "h3 \ get_ancestors_si ptr \\<^sub>r ancestors_h2" using get_ancestors_si_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_separate_forwards) using \heap_is_wellformed h2\ apply simp using ancestors_h2 apply simp apply(auto simp add: get_ancestors_si_locs_def get_parent_locs_def)[1] apply (simp add: local.get_ancestors_si_locs_def local.get_parent_reads_pointers local.set_disconnected_nodes_get_ancestors_si) using local.get_ancestors_si_locs_def local.set_disconnected_nodes_get_ancestors_si by blast have node_not_in_ancestors_h2: "cast node \ set ancestors_h2" using \heap_is_wellformed h\ \heap_is_wellformed h2\ ancestors ancestors_h2 apply(rule get_ancestors_si_remains_not_in_ancestors) using assms(2) assms(3) h2 local.adopt_node_children_subset apply blast using shadow_root_eq_h2 node_not_in_ancestors object_ptr_kinds_M_eq2_h assms(2) assms(3) \type_wf h2\ by(auto dest: returns_result_eq) moreover have "parent_child_rel h2 = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))" using children_h3 children_h' ptr_in_heap apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3 insert_before_list_node_in_set)[1] apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2) by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2) have "CD.a_acyclic_heap h'" proof - have "acyclic (parent_child_rel h2)" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) then have "acyclic (parent_child_rel h3)" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "cast node \ {x. (x, ptr) \ (parent_child_rel h2)\<^sup>*}" using get_ancestors_si_parent_child_rel using \known_ptrs h2\ \type_wf h2\ ancestors_h2 node_not_in_ancestors_h2 wellformed_h2 by blast then have "cast node \ {x. (x, ptr) \ (parent_child_rel h3)\<^sup>*}" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) ultimately show ?thesis using \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\ by(auto simp add: CD.acyclic_heap_def) qed moreover have "CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) have "CD.a_all_ptrs_in_heap h'" proof - have "CD.a_all_ptrs_in_heap h3" using \CD.a_all_ptrs_in_heap h2\ apply(auto simp add: CD.a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2 children_eq_h2)[1] using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 using node_ptr_kinds_eq2_h2 apply auto[1] apply (metis (no_types, lifting) children_eq2_h2 in_mono notin_fset object_ptr_kinds_M_eq3_h2) by (metis (no_types, hide_lams) NodeMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in node_ptr_kinds_eq2_h2 object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD) have "set children_h3 \ set |h' \ node_ptr_kinds_M|\<^sub>r" using children_h3 \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1] using CD.parent_child_rel_child_nodes2 \known_ptr ptr\ \parent_child_rel h2 = parent_child_rel h3\ \type_wf h2\ local.parent_child_rel_child_in_heap node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 wellformed_h2 by blast then have "set (insert_before_list node reference_child children_h3) \ set |h' \ node_ptr_kinds_M|\<^sub>r" using node_in_heap apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1] by (metis (no_types, hide_lams) contra_subsetD finite_set_in insert_before_list_in_set node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2) then show ?thesis using \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: object_ptr_kinds_M_eq3_h' CD.a_all_ptrs_in_heap_def node_ptr_kinds_def node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1] using children_eq_h3 children_h' apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD) by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h3 finite_set_in subsetD) qed moreover have "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_distinct_lists h3" proof(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2 children_eq2_h2 intro!: distinct_concat_map_I) fix x assume 1: "x |\| document_ptr_kinds h3" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" show "distinct |h3 \ get_disconnected_nodes x|\<^sub>r" using distinct_concat_map_E(2)[OF 2] select_result_I2[OF disconnected_nodes_h3] disconnected_nodes_eq2_h2 select_result_I2[OF disconnected_nodes_h2] 1 by (metis (full_types) distinct_remove1 finite_fset fmember.rep_eq set_sorted_list_of_set) next fix x y xa assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and 2: "x |\| document_ptr_kinds h3" and 3: "y |\| document_ptr_kinds h3" and 4: "x \ y" and 5: "xa \ set |h3 \ get_disconnected_nodes x|\<^sub>r" and 6: "xa \ set |h3 \ get_disconnected_nodes y|\<^sub>r" show False proof (cases "x = owner_document") case True then have "y \ owner_document" using 4 by simp show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \y \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis proof (cases "y = owner_document") case True then show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \x \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6 using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 disjoint_iff_not_equal finite_fset fmember.rep_eq notin_set_remove1 select_result_I2 set_sorted_list_of_set by (metis (no_types, lifting)) qed qed next fix x xa xb assume 1: "(\x\fset (object_ptr_kinds h3). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 2: "xa |\| object_ptr_kinds h3" and 3: "x \ set |h3 \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h3" and 5: "x \ set |h3 \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 4 by (metis \type_wf h2\ children_eq2_h2 document_ptr_kinds_commutes \known_ptrs h\ local.get_child_nodes_ok local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result wellformed_h2) show False proof (cases "xb = owner_document") case True then show ?thesis using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]] by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1) next case False show ?thesis using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto qed qed then have "CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3 disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I) fix x assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" have 3: "\p. p |\| object_ptr_kinds h' \ distinct |h3 \ get_child_nodes p|\<^sub>r" using 1 by (auto elim: distinct_concat_map_E) show "distinct |h' \ get_child_nodes x|\<^sub>r" proof(cases "ptr = x") case True show ?thesis using 3[OF 2] children_h3 children_h' by(auto simp add: True insert_before_list_distinct dest: child_not_in_any_children[unfolded children_eq_h2]) next case False show ?thesis using children_eq2_h3[OF False] 3[OF 2] by auto qed next fix x y xa assume 1:"distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" and 3: "y |\| object_ptr_kinds h'" and 4: "x \ y" and 5: "xa \ set |h' \ get_child_nodes x|\<^sub>r" and 6: "xa \ set |h' \ get_child_nodes y|\<^sub>r" have 7:"set |h3 \ get_child_nodes x|\<^sub>r \ set |h3 \ get_child_nodes y|\<^sub>r = {}" using distinct_concat_map_E(1)[OF 1] 2 3 4 by auto show False proof (cases "ptr = x") case True then have "ptr \ y" using 4 by simp then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ y\])[1] by (metis (no_types, hide_lams) "3" "7" \type_wf h3\ children_eq2_h3 disjoint_iff_not_equal get_child_nodes_ok insert_before_list_in_set \known_ptrs h\ local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2) next case False then show ?thesis proof (cases "ptr = y") case True then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ x\])[1] by (metis (no_types, hide_lams) "2" "4" "7" IntI \known_ptrs h3\ \type_wf h'\ children_eq_h3 empty_iff insert_before_list_in_set local.get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h' returns_result_select_result select_result_I2) next case False then show ?thesis using children_eq2_h3[OF \ptr \ x\] children_eq2_h3[OF \ptr \ y\] 5 6 7 by auto qed qed next fix x xa xb assume 1: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h' \ get_disconnected_nodes x|\<^sub>r) = {} " and 2: "xa |\| object_ptr_kinds h'" and 3: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h'" and 5: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h' \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 3 4 5 proof - have "\h d. \ type_wf h \ d |\| document_ptr_kinds h \ h \ ok get_disconnected_nodes d" using local.get_disconnected_nodes_ok by satx then have "h' \ ok get_disconnected_nodes xb" using "4" \type_wf h'\ by fastforce then have f1: "h3 \ get_disconnected_nodes xb \\<^sub>r |h' \ get_disconnected_nodes xb|\<^sub>r" by (simp add: disconnected_nodes_eq_h3) have "xa |\| object_ptr_kinds h3" using "2" object_ptr_kinds_M_eq3_h' by blast then show ?thesis using f1 \local.CD.a_distinct_lists h3\ CD.distinct_lists_no_parent by fastforce qed show False proof (cases "ptr = xa") case True show ?thesis using 6 node_not_in_disconnected_nodes 3 4 5 select_result_I2[OF children_h'] select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3 by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M \CD.a_distinct_lists h3\ \type_wf h'\ disconnected_nodes_eq_h3 CD.distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result) next case False then show ?thesis using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce qed qed moreover have "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_M_eq2_h2 object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 )[1] thm children_eq2_h3 apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified] object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified] node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1] apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1] by (metis (no_types, lifting) Core_DOM_Functions.i_insert_before.insert_before_list_in_set children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1 is_OK_returns_result_I object_ptr_kinds_M_eq3_h' ptr_in_heap returns_result_eq returns_result_select_result) ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by (simp add: CD.heap_is_wellformed_def) have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq2_h2) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 shadow_root_eq2_h3) have "cast node \ {x. (x, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}" using get_ancestors_si_parent_child_host_shadow_root_rel using \known_ptrs h2\ \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ \type_wf h2\ ancestors_h2 node_not_in_ancestors_h2 wellformed_h2 by auto have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \heap_is_wellformed h2\ by(auto simp add: heap_is_wellformed_def \parent_child_rel h2 = parent_child_rel h3\ \a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\) then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" apply(auto simp add: \a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'\ \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\)[1] using \cast node \ {x. (x, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}\ by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\) then show "heap_is_wellformed h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1] by(auto simp add: object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3) qed end interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_si get_ancestors_si_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf_is_l_insert_before_wf [instances]: "l_insert_before_wf Shadow_DOM.heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.get_child_nodes" apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1] using insert_before_removes_child apply fast done lemma l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] by (metis Diff_iff Shadow_DOM.i_heap_is_wellformed.heap_is_wellformed_disconnected_nodes_distinct Shadow_DOM.i_remove_child.set_disconnected_nodes_get_disconnected_nodes insert_iff returns_result_eq set_remove1_eq) interpretation l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel get_ancestors_si get_ancestors_si_locs get_parent get_parent_locs adopt_node adopt_node_locs get_owner_document insert_before insert_before_locs append_child known_ptrs get_host get_host_locs get_root_node_si get_root_node_si_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf2_is_l_insert_before_wf2 [instances]: "l_insert_before_wf2 ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.heap_is_wellformed" apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1] using insert_before_child_preserves apply(fast, fast, fast) done subsubsection \append\_child\ interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove get_ancestors_si get_ancestors_si_locs insert_before insert_before_locs append_child heap_is_wellformed parent_child_rel by(auto simp add: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr known_ptrs append_child heap_is_wellformed" apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1] using append_child_heap_is_wellformed_preserved by fast+ subsubsection \to\_tree\_order\ interpretation i_to_tree_order_wf?: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using to_tree_order_ok apply fast using to_tree_order_ptrs_in_heap apply fast using to_tree_order_parent_child_rel apply(fast, fast) using to_tree_order_child2 apply blast using to_tree_order_node_ptrs apply fast using to_tree_order_child apply fast using to_tree_order_ptr_in_result apply fast using to_tree_order_parent apply fast using to_tree_order_subset apply fast done paragraph \get\_root\_node\ interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order by(auto simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs to_tree_order Shadow_DOM.get_root_node Shadow_DOM.heap_is_wellformed" apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def instances)[1] using to_tree_order_get_root_node apply fast using to_tree_order_same_root apply fast done subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma to_tree_order_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (to_tree_order_si ptr)" proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct_si) case (step parent) have "known_ptr parent" using assms(2) local.known_ptrs_known_ptr step.prems by blast then show ?case using step using assms(1) assms(2) assms(3) using local.heap_is_wellformed_children_in_heap local.get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: to_tree_order_si_def[of parent] intro: get_child_nodes_ok get_shadow_root_ok intro!: bind_is_OK_pure_I map_M_pure_I bind_pure_I map_M_ok_I split: option.splits) qed end interpretation i_to_tree_order_si_wf?: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs to_tree_order_si by(auto simp add: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_assigned\_nodes\ lemma forall_M_small_big: "h \ forall_M f xs \\<^sub>h h' \ P h \ (\h h' x. x \ set xs \ h \ f x \\<^sub>h h' \ P h \ P h') \ P h'" by(induct xs arbitrary: h) (auto elim!: bind_returns_heap_E) locale l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed + l_remove_child_wf2 + l_append_child_wf + l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma assigned_nodes_distinct: assumes "heap_is_wellformed h" assumes "h \ assigned_nodes slot \\<^sub>r nodes" shows "distinct nodes" proof - have "\ptr children. h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using assms(1) local.heap_is_wellformed_children_distinct by blast then show ?thesis using assms apply(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits)[1] by (simp add: filter_M_distinct) qed lemma flatten_dom_preserves: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ flatten_dom \\<^sub>h h'" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain tups h2 element_ptrs shadow_root_ptrs where "h \ element_ptr_kinds_M \\<^sub>r element_ptrs" and tups: "h \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}) element_ptrs \\<^sub>r tups" (is "h \ map_filter_M2 ?f element_ptrs \\<^sub>r tups") and h2: "h \ forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups \\<^sub>h h2" and "h2 \ shadow_root_ptr_kinds_M \\<^sub>r shadow_root_ptrs" and h': "h2 \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }) shadow_root_ptrs \\<^sub>h h'" using \h \ flatten_dom \\<^sub>h h'\ apply(auto simp add: flatten_dom_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF ElementMonad.ptr_kinds_M_pure, rotated] bind_returns_heap_E2[rotated, OF ShadowRootMonad.ptr_kinds_M_pure, rotated])[1] apply(drule pure_returns_heap_eq) by(auto intro!: map_filter_M2_pure bind_pure_I) have "heap_is_wellformed h2 \ known_ptrs h2 \ type_wf h2" using h2 \heap_is_wellformed h\ \known_ptrs h\ \type_wf h\ by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] elim!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf) then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_host_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] dest!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf remove_shadow_root_preserves ) qed end interpretation i_assigned_nodes_wf?: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs known_ptrs remove_child remove_child_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" begin end subsubsection \create\_element\ locale l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name type_wf set_tag_name set_tag_name_locs + l_create_element_defs create_element + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs + l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs + l_new_element_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_new_element_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs + l_new_element type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma create_element_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_element document_ptr tag \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_element_ptr h2 h3 disc_nodes_h3 where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF CD.get_disconnected_nodes_pure, rotated] ) then have "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I CD.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_element_ptr \ set |h \ element_ptr_kinds_M|\<^sub>r" using new_element_ptr ElementMonad.ptr_kinds_ptr_kinds_M h2 using new_element_ptr_not_in_heap by blast then have "cast new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_element_ptr|}" using new_element_new_ptr h2 new_element_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_element_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h |\| {|new_element_ptr|}" apply(simp add: element_ptr_kinds_def) by force have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def character_data_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) using set_tag_name_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" by(simp add: element_ptr_kinds_def) have "known_ptr (cast new_element_ptr)" using \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ local.create_element_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_element_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_element_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes by blast have tag_name_eq_h: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by(blast)+ then have tag_name_eq2_h: "\ptr'. ptr' \ new_element_ptr \ |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name new_element_ptr \\<^sub>r ''''" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_empty_tag_name by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name new_element_ptr \\<^sub>r ''''" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_empty_tag_name by blast have "type_wf h2" using \type_wf h\ new_element_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_tag_name_writes h3] using set_tag_name_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_element_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_disconnected_nodes_writes h']) using set_disconnected_nodes_get_tag_name by blast then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] apply (metis \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(3) funion_iff CD.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap CD.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h returns_result_select_result) by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subsetD) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (smt children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 element_ptr_kinds_commutes h' h2 local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_element_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_element_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_element_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: CD.heap_is_wellformed_def heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) apply(case_tac "x=cast new_element_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \ CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) then have " CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2) then have " CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h3\ \ CD.a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) CD.distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply(smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \CD.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' set_disconnected_nodes_get_disconnected_nodes CD.a_all_ptrs_in_heap_def select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes CD.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply - apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \ CD.a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 CD.distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \ CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: CD.a_owner_document_valid_def)[1] apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1] apply(auto simp add: object_ptr_kinds_eq_h2)[1] apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1] apply(auto simp add: document_ptr_kinds_eq_h2)[1] apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1] apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1] apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes select_result_I2) have "CD.a_heap_is_wellformed h'" using \CD.a_acyclic_heap h'\ \CD.a_all_ptrs_in_heap h'\ \CD.a_distinct_lists h'\ \CD.a_owner_document_valid h'\ by(simp add: CD.a_heap_is_wellformed_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\element_ptr shadow_root_opt. element_ptr \ new_element_ptr \ h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" proof - fix element_ptr shadow_root_opt assume "element_ptr \ new_element_ptr " have "\P \ get_shadow_root_locs element_ptr. P h h2" using get_shadow_root_new_element new_element_ptr h2 using \element_ptr \ new_element_ptr\ by blast then have "preserved (get_shadow_root element_ptr) h h2" using get_shadow_root_new_element[rotated, OF new_element_ptr h2] using get_shadow_root_reads by(simp add: reads_def) then show "h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" by (simp add: preserved_def) qed have shadow_root_none: "h2 \ get_shadow_root (new_element_ptr) \\<^sub>r None" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_shadow_root by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h shadow_root_none by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = new_element_ptr") using shadow_root_none apply auto[1] using shadow_root_eq_h by (smt Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) finite_set_in h2 insort_split local.get_shadow_root_ok local.shadow_root_same_host new_element_ptr new_element_ptr_not_in_heap option.distinct(1) returns_result_select_result select_result_I2 shadow_root_none) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h) moreover have "previous_host \ new_element_ptr" using calculation(1) h2 new_element_ptr new_element_ptr_not_in_heap by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h apply (simp add: tag_name_eq2_h) by (metis \previous_host \ new_element_ptr\ \|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" by (meson \previous_host \ fset (element_ptr_kinds h)\ \previous_host \ new_element_ptr\ assms(3) local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap notin_fset returns_result_select_result shadow_root_eq_h) qed then have "a_shadow_root_valid h3" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h2). \host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h2)" and "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h2\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2) moreover have "previous_host \ new_element_ptr" using calculation(1) h3 new_element_ptr new_element_ptr_not_in_heap using calculation(3) shadow_root_none by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h2 apply (simp add: tag_name_eq2_h2) by (metis \previous_host \ new_element_ptr\ \|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h3). |h3 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h3 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" by (smt \previous_host \ fset (element_ptr_kinds h2)\ \previous_host \ new_element_ptr\ \type_wf h2\ \type_wf h3\ element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok returns_result_eq returns_result_select_result shadow_root_eq_h2 tag_name_eq2_h2) qed then have "a_shadow_root_valid h'" apply(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h3 shadow_root_eq_h3 shadow_root_ptr_kinds_eq_h3 tag_name_eq2_h3)[1] by (smt \type_wf h3\ finite_set_in local.get_shadow_root_ok returns_result_select_result select_result_I2 shadow_root_eq_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq_h)[1] apply (smt assms(3) case_prod_conv h2 image_iff local.get_shadow_root_ok mem_Collect_eq new_element_ptr new_element_ptr_not_in_heap returns_result_select_result select_result_I2 shadow_root_eq_h) using shadow_root_none apply auto[1] by (metis (no_types, lifting) Collect_cong assms(3) case_prodE case_prodI h2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok local.new_element_no_shadow_root new_element_ptr option.distinct(1) returns_result_select_result select_result_I2 shadow_root_eq_h) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] apply (smt Collect_cong Shadow_DOM.a_host_shadow_root_rel_def assms(3) h2 host_shadow_root_rel_shadow_root is_OK_returns_result_E local.get_shadow_root_impl local.get_shadow_root_ok local.new_element_types_preserved select_result_I2 shadow_root_eq_h2 split_cong) apply (metis (no_types, lifting) Collect_cong \type_wf h3\ element_ptr_kinds_eq_h2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h2 split_cong) done have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] apply (metis (no_types, lifting) Collect_cong \type_wf h3\ case_prodE case_prodI element_ptr_kinds_eq_h2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h3) apply (smt Collect_cong \type_wf h'\ \type_wf h2\ case_prodD case_prodI2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root is_OK_returns_result_E local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok select_result_I2 shadow_root_eq_h2 shadow_root_eq_h3) done have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \parent_child_rel h3 = parent_child_rel h'\) have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3\) show " heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation i_create_element_wf?: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr DocumentClass.type_wf by(auto simp add: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs create_character_data known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_character_data_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_val_get_disconnected_nodes type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs + l_new_character_data_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_val_get_child_nodes type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs + l_new_character_data type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma create_character_data_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_character_data document_ptr text \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes_h3 where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: CD.create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF CD.get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" apply(auto simp add: CD.create_character_data_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.CD.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>r" using new_character_data_ptr CharacterDataMonad.ptr_kinds_ptr_kinds_M h2 using new_character_data_ptr_not_in_heap by blast then have "cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF CD.set_val_writes h3]) using CD.set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_character_data_ptr)" using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ local.create_character_data_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF CD.set_val_writes h3]) using CD.set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []" using new_character_data_ptr h2 new_character_data_ptr_in_heap[OF h2 new_character_data_ptr] new_character_data_is_character_data_ptr[OF new_character_data_ptr] new_character_data_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h2 get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_character_data_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF CD.set_val_writes h3] using set_val_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_character_data_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h finite_set_in finsert_iff funion_finsert_right CD.parent_child_rel_child CD.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \h2 \ get_child_nodes (cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr) \\<^sub>r []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subset_code(1)) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (smt character_data_ptr_kinds_commutes character_data_ptr_kinds_eq_h2 children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 h' h2 local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr new_character_data_ptr_in_heap node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_character_data_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_character_data_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_character_data_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_character_data_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) apply(case_tac "x=cast new_character_data_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] thm children_eq2_h using \CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result by metis then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, hide_lams) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set disc_nodes_h3\ \type_wf h2\ assms(1) disc_nodes_document_ptr_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disconnected_nodes_eq_h distinct.simps(2) document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result select_result_I2) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" using NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ by (smt local.CD.a_all_ptrs_in_heap_def \CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply(cases "document_ptr = xb") apply (metis (no_types, lifting) "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 children_eq2_h3 disc_nodes_document_ptr_h2 document_ptr_kinds_eq_h3 h' local.get_disconnected_nodes_ok local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr_not_in_any_children object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_eq returns_result_select_result set_ConsD) by (metis "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 local.get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(simp add: CD.a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by (metis (mono_tags, lifting) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_ptr_kinds_M.ptr_kinds_ptr_kinds_M l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms object_ptr_kinds_M_def select_result_I2) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_character_data[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_character_data_ptr apply blast using local.get_shadow_root_locs_impl new_character_data_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h] tag_name_eq2_h) then have "a_shadow_root_valid h3" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2 element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2] tag_name_eq2_h2) then have "a_shadow_root_valid h'" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h3 element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3] tag_name_eq2_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \parent_child_rel h3 = parent_child_rel h'\) have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end subsubsection \create\_document\ locale l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_new_document_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document + l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_new_document_get_tag_name get_tag_name get_tag_name_locs + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs + l_new_document type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_document :: "((_) heap, exception, (_) document_ptr) prog" and known_ptrs :: "(_) heap \ bool" begin lemma create_document_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_document \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" proof - obtain new_document_ptr where new_document_ptr: "h \ new_document \\<^sub>r new_document_ptr" and h': "h \ new_document \\<^sub>h h'" using assms(2) apply(simp add: create_document_def) using new_document_ok by blast have "new_document_ptr \ set |h \ document_ptr_kinds_M|\<^sub>r" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have "new_document_ptr |\| document_ptr_kinds h" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr |\| object_ptr_kinds h" by simp have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_document_ptr|}" using new_document_new_ptr h' new_document_ptr by blast then have node_ptr_kinds_eq: "node_ptr_kinds h' = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h' = character_data_ptr_kinds h" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h' = element_ptr_kinds h" using object_ptr_kinds_eq by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h' = document_ptr_kinds h |\| {|new_document_ptr|}" using object_ptr_kinds_eq apply(auto simp add: document_ptr_kinds_def)[1] by (metis (no_types, lifting) document_ptr_kinds_commutes document_ptr_kinds_def finsertI1 fset.map_comp) have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h" using object_ptr_kinds_eq apply(simp add: shadow_root_ptr_kinds_def) by force have children_eq: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_document_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h' get_child_nodes_new_document[rotated, OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2: "\ptr'. ptr' \ cast new_document_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" using new_document_ptr h' new_document_ptr_in_heap[OF h' new_document_ptr] new_document_is_document_ptr[OF new_document_ptr] new_document_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. doc_ptr \ new_document_ptr \ h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis(full_types) \\thesis. (\new_document_ptr. \h \ new_document \\<^sub>r new_document_ptr; h \ new_document \\<^sub>h h'\ \ thesis) \ thesis\ local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ new_document_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" using h' local.new_document_no_disconnected_nodes new_document_ptr by blast have "type_wf h'" using \type_wf h\ new_document_types_preserved h' by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h'" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h'" by (simp add: object_ptr_kinds_eq) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h' \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ children_eq2 empty_iff empty_set image_eqI select_result_I2) qed finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \parent_child_rel h = parent_child_rel h'\ assms(1) children_eq fset_of_list_elem local.heap_is_wellformed_children_in_heap CD.parent_child_rel_child CD.parent_child_rel_parent_in_heap node_ptr_kinds_eq apply (metis (no_types, lifting) \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \type_wf h'\ assms(1) disconnected_nodes_eq_h empty_iff empty_set local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq returns_result_select_result select_result_I2) have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h'" using \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ apply(auto simp add: children_eq2[symmetric] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(auto simp add: dest: distinct_concat_map_E)[1] using \new_document_ptr |\| document_ptr_kinds h\ apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1] apply (metis assms(1) assms(3) disconnected_nodes_eq2_h get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result) proof - fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr" assume a1: "x \ y" assume a2: "x |\| document_ptr_kinds h" assume a3: "x \ new_document_ptr" assume a4: "y |\| document_ptr_kinds h" assume a5: "y \ new_document_ptr" assume a6: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by presburger have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>r" using a8 a5 disconnected_nodes_eq2_h by presburger have f11: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a4 by simp have "x \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a2 by simp then show False using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1)) next fix x xa xb assume 0: "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" and 1: "h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr" and 9: "xb |\| document_ptr_kinds h" and 10: "xb \ new_document_ptr" then show "False" by (metis \CD.a_distinct_lists h\ assms(3) disconnected_nodes_eq2_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def)[1] by (metis \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff node_ptr_kinds_eq object_ptr_kinds_eq) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h' \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads assms(2) get_shadow_root_new_document[rotated, OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_document_ptr apply blast using local.get_shadow_root_locs_impl new_document_ptr by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq document_ptr_kinds_eq_h)[1] using shadow_root_eq_h by fastforce have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h'\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h' get_tag_name_new_document[OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" using new_document_is_document_ptr[OF new_document_ptr] by(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h document_ptr_kinds_eq_h shadow_root_ptr_kinds_eq select_result_eq[OF shadow_root_eq_h] select_result_eq[OF tag_name_eq_h]) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) moreover have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h' \ a_host_shadow_root_rel h'" by (simp add: \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h'\ \parent_child_rel h = parent_child_rel h'\) ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by simp have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h'" using CD.heap_is_wellformed_impl \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\ local.heap_is_wellformed_def by auto qed end interpretation l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf DocumentClass.type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_document known_ptrs by(auto simp add: l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs + l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_set_shadow_root_get_disconnected_nodes set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs + l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_shadow_root_get_tag_name set_shadow_root set_shadow_root_locs get_tag_name get_tag_name_locs + l_new_shadow_root type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" begin lemma attach_shadow_root_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ attach_shadow_root element_ptr new_mode \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain h2 h3 new_shadow_root_ptr element_tag_name where element_tag_name: "h \ get_tag_name element_ptr \\<^sub>r element_tag_name" and "element_tag_name \ safe_shadow_root_element_types" and prev_shadow_root: "h \ get_shadow_root element_ptr \\<^sub>r None" and h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3" and h': "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" using assms(4) by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) have "h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr" thm bind_pure_returns_result_I[OF get_tag_name_pure] apply(unfold attach_shadow_root_def)[1] using element_tag_name apply(rule bind_pure_returns_result_I[OF get_tag_name_pure]) apply(rule bind_pure_returns_result_I) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using prev_shadow_root apply(rule bind_pure_returns_result_I[OF get_shadow_root_pure]) apply(rule bind_pure_returns_result_I) apply(simp) apply(simp) using h2 new_shadow_root_ptr h3 h' by(auto intro!: bind_returns_result_I intro: is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h3]] is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h']]) have "new_shadow_root_ptr \ set |h \ shadow_root_ptr_kinds_M|\<^sub>r" using new_shadow_root_ptr ShadowRootMonad.ptr_kinds_ptr_kinds_M h2 using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by blast then have "cast new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr h2 new_shadow_root_ptr by blast then have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" apply(simp add: document_ptr_kinds_def) by force have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h |\| {|new_shadow_root_ptr|}" using object_ptr_kinds_eq_h apply(simp add: shadow_root_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_mode_writes h3]) using set_mode_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by (auto simp add: shadow_root_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_shadow_root_writes h']) using set_shadow_root_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by (auto simp add: shadow_root_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_shadow_root_ptr)" using \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ create_shadow_root_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "element_ptr |\| element_ptr_kinds h" by (meson \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ is_OK_returns_result_I local.attach_shadow_root_element_ptr_in_heap) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_shadow_root_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr h2 new_shadow_root_ptr object_ptr_kinds_eq_h by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" apply(simp add: character_data_ptr_kinds_def) done have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_mode_writes h3]) using set_mode_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_shadow_root_writes h']) using set_shadow_root_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_shadow_root_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" using h2 local.new_shadow_root_no_child_nodes new_shadow_root_ptr by auto have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_shadow_root[rotated, OF h2,rotated,OF new_shadow_root_ptr] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis (no_types, lifting))+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_shadow_root[OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_shadow_root_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_mode_writes h3] using set_mode_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_shadow_root_writes h'] using set_shadow_root_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_disconnected_nodes) then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) CD.get_child_nodes_ok CD.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(2) l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap node_ptr_kinds_commutes returns_result_select_result) by (metis assms(1) assms(2) disconnected_nodes_eq2_h document_ptr_kinds_eq_h local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (simp add: children_eq2_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3) have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply(auto simp add: select_result_eq[OF disconnected_nodes_eq_h] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I dest: distinct_concat_map_E)[1] apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(case_tac "x = cast new_shadow_root_ptr") using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply blast apply(case_tac "y = cast new_shadow_root_ptr") using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply blast proof - fix x y :: "(_) object_ptr" fix xa :: "(_) node_ptr" assume a1: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" assume "x \ y" assume "xa \ set |h2 \ get_child_nodes x|\<^sub>r" assume "xa \ set |h2 \ get_child_nodes y|\<^sub>r" assume "x |\| object_ptr_kinds h" assume "x \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr" assume "y |\| object_ptr_kinds h" assume "y \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr" show False using distinct_concat_map_E(1)[OF a1, of x y] using \x |\| object_ptr_kinds h\ \y |\| object_ptr_kinds h\ using \xa \ set |h2 \ get_child_nodes x|\<^sub>r\ \xa \ set |h2 \ get_child_nodes y|\<^sub>r\ using \x \ y\ by(auto simp add: children_eq2_h[OF \x \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr\] children_eq2_h[OF \y \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr\]) next fix x :: "(_) node_ptr" fix xa :: "(_) object_ptr" fix xb :: "(_) document_ptr" assume "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" assume "x \ set |h2 \ get_child_nodes xa|\<^sub>r" assume "xb |\| document_ptr_kinds h" assume "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" assume "xa |\| object_ptr_kinds h" assume "xa \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr" have "set |h \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" by (metis (no_types, lifting) CD.get_child_nodes_ok \xa |\| object_ptr_kinds h\ \xb |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) disconnected_nodes_eq2_h is_OK_returns_result_E local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr select_result_I2) then show "False" using \x \ set |h2 \ get_child_nodes xa|\<^sub>r\ \x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r\ \xa \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr\ children_eq2_h by auto qed then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I) have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" (* using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ *) apply(simp add: CD.a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] by (metis CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ assms(2) assms(3) children_eq2_h children_eq_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.known_ptrs_known_ptr object_ptr_kinds_M_def returns_result_select_result) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_shadow_root[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_shadow_root_ptr apply blast using local.get_shadow_root_locs_impl new_shadow_root_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. element_ptr \ ptr' \ h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_shadow_root_different_pointers) have shadow_root_h3: "h' \ get_shadow_root element_ptr \\<^sub>r Some new_shadow_root_ptr" using \type_wf h3\ h' local.set_shadow_root_get_shadow_root by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] apply(case_tac "shadow_root_ptr = new_shadow_root_ptr") using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_eq_h2 apply blast using \type_wf h3\ h' local.set_shadow_root_get_shadow_root returns_result_eq shadow_root_eq_h3 apply fastforce done have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(2) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (smt \type_wf h3\ assms(1) assms(2) h' h2 local.get_shadow_root_ok local.get_shadow_root_shadow_root_ptr_in_heap local.set_shadow_root_get_shadow_root local.shadow_root_same_host new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" proof(unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "a_shadow_root_valid h" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h')" show "\host\fset (element_ptr_kinds h'). |h' \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h' \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" proof (cases "shadow_root_ptr = new_shadow_root_ptr") case True have "element_ptr \ fset (element_ptr_kinds h')" by (simp add: \element_ptr |\| element_ptr_kinds h\ element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3) moreover have "|h' \ get_tag_name element_ptr|\<^sub>r \ safe_shadow_root_element_types" by (smt \\thesis. (\h2 h3 new_shadow_root_ptr element_tag_name. \h \ get_tag_name element_ptr \\<^sub>r element_tag_name; element_tag_name \ safe_shadow_root_element_types; h \ get_shadow_root element_ptr \\<^sub>r None; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr; h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3; h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'\ \ thesis) \ thesis\ select_result_I2 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) moreover have "|h' \ get_shadow_root element_ptr|\<^sub>r = Some shadow_root_ptr" using shadow_root_h3 by (simp add: True) ultimately show ?thesis by meson next case False then obtain host where host: "host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" using \shadow_root_ptr \ fset (shadow_root_ptr_kinds h')\ using \\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr\ apply(simp add: shadow_root_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h) by (meson finite_set_in) moreover have "host \ element_ptr" using calculation(3) prev_shadow_root by auto ultimately show ?thesis using element_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h by (smt \type_wf h'\ assms(2) finite_set_in local.get_shadow_root_ok returns_result_eq returns_result_select_result shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) qed qed have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 )[1] apply(case_tac "element_ptr \ aa") using select_result_eq[OF shadow_root_eq_h3] apply (simp add: image_iff) using select_result_eq[OF shadow_root_eq_h3] apply (metis (no_types, lifting) \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \type_wf h3\ host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok option.distinct(1) prev_shadow_root returns_result_select_result) apply (metis (mono_tags, lifting) \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ case_prod_conv image_iff is_OK_returns_result_I mem_Collect_eq option.inject returns_result_eq returns_result_select_result shadow_root_h3) using element_ptr_kinds_eq_h3 local.get_shadow_root_ptr_in_heap shadow_root_h3 apply fastforce using Shadow_DOM.a_host_shadow_root_rel_def \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ \type_wf h3\ case_prodE case_prodI host_shadow_root_rel_shadow_root image_iff local.get_shadow_root_impl local.get_shadow_root_ok mem_Collect_eq option.discI prev_shadow_root returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 apply(auto)[1] by (smt case_prodI mem_Collect_eq option.distinct(1) pair_imageI returns_result_eq returns_result_select_result) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h' = {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr)} \ local.a_host_shadow_root_rel h3\ \parent_child_rel h3 = parent_child_rel h'\) have "\a b. (a, b) \ parent_child_rel h3 \ a \ cast new_shadow_root_ptr" using CD.parent_child_rel_parent_in_heap \parent_child_rel h = parent_child_rel h2\ \parent_child_rel h2 = parent_child_rel h3\ document_ptr_kinds_commutes by (metis h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_commutes) moreover have "\a b. (a, b) \ a_host_shadow_root_rel h3 \ a \ cast new_shadow_root_ptr" using shadow_root_eq_h2 by(auto simp add: a_host_shadow_root_rel_def) moreover have "cast new_shadow_root_ptr \ {x. (x, cast element_ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}" by (metis (no_types, lifting) UnE calculation(1) calculation(2) cast_shadow_root_ptr_not_node_ptr(1) converse_rtranclE mem_Collect_eq) moreover have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation l_attach_shadow_root_wf?: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs set_val set_val_locs create_character_data DocumentClass.known_ptr DocumentClass.type_wf set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root by(auto simp add: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end