Page MenuHomeIsabelle/Phabricator

No OneTemporary

This file is larger than 256 KB, so syntax highlighting was skipped.
diff --git a/thys/BNF_CC/DDS.thy b/thys/BNF_CC/DDS.thy
--- a/thys/BNF_CC/DDS.thy
+++ b/thys/BNF_CC/DDS.thy
@@ -1,396 +1,396 @@
(* Author: Andreas Lochbihler, ETH Zurich
Author: Joshua Schneider, ETH Zurich *)
section \<open>Example: deterministic discrete system\<close>
theory DDS imports
Concrete_Examples
"HOL-Library.Rewrite"
"HOL-Library.FSet"
begin
unbundle lifting_syntax
subsection \<open>Definition and generalised mapper and relator\<close>
codatatype ('a, 'b) dds = DDS (run: "'a \<Rightarrow> 'b \<times> ('a, 'b) dds")
for map: map_dds'
rel: rel_dds'
primcorec map_dds :: "('a' \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'b') \<Rightarrow> ('a, 'b) dds \<Rightarrow> ('a', 'b') dds" where
"run (map_dds f g S) = (\<lambda>a. map_prod g (map_dds f g) (run S (f a)))"
lemma map_dds_id: "map_dds id id S = S"
by(coinduction arbitrary: S)(auto simp add: rel_fun_def prod.rel_map intro: prod.rel_refl_strong)
lemma map_dds_comp: "map_dds f g (map_dds f' g' S) = map_dds (f' \<circ> f) (g \<circ> g') S"
by(coinduction arbitrary: S)(auto simp add: rel_fun_def prod.rel_map intro: prod.rel_refl_strong)
coinductive rel_dds :: "('a \<Rightarrow> 'a' \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b' \<Rightarrow> bool) \<Rightarrow> ('a, 'b) dds \<Rightarrow> ('a', 'b') dds \<Rightarrow> bool"
for A B where
"rel_dds A B S S'" if "rel_fun A (rel_prod B (rel_dds A B)) (run S) (run S')"
lemma rel_dds'_rel_dds: "rel_dds' B = rel_dds (=) B"
apply (intro ext iffI)
apply (erule rel_dds.coinduct)
apply (erule dds.rel_cases)
apply (simp)
apply (erule rel_fun_mono[THEN predicate2D, OF order_refl, rotated -1])
apply (rule prod.rel_mono[OF order_refl])
apply (blast)
apply (erule dds.rel_coinduct)
apply (erule rel_dds.cases)
apply (simp)
done
lemma rel_dds_eq [relator_eq]: "rel_dds (=) (=) = (=)"
apply(rule ext iffI)+
subgoal by(erule dds.coinduct)(erule rel_dds.cases; simp)
subgoal by(erule rel_dds.coinduct)(auto simp add: rel_fun_def intro!: prod.rel_refl_strong)
done
lemma rel_dds_mono [relator_mono]: "rel_dds A B \<le> rel_dds A' B'" if "A' \<le> A" "B \<le> B'"
apply(rule predicate2I)
apply(erule rel_dds.coinduct)
apply(erule rel_dds.cases)
apply simp
apply(erule BNF_Def.rel_fun_mono)
apply(auto intro: that[THEN predicate2D])
done
lemma rel_dds_conversep: "rel_dds A\<inverse>\<inverse> B\<inverse>\<inverse> = (rel_dds A B)\<inverse>\<inverse>"
apply(intro ext iffI; simp)
subgoal
apply(erule rel_dds.coinduct; erule rel_dds.cases; simp del: conversep_iff)
apply(rewrite conversep_iff[symmetric])
apply(fold rel_fun_conversep prod.rel_conversep)
apply(erule BNF_Def.rel_fun_mono)
apply(auto simp del: conversep_iff)
done
subgoal
apply(erule rel_dds.coinduct; erule rel_dds.cases; clarsimp simp del: conversep_iff)
apply(rewrite in asm conversep_iff[symmetric])
apply(fold rel_fun_conversep prod.rel_conversep)
apply(erule BNF_Def.rel_fun_mono)
apply(auto simp del: conversep_iff)
done
done
lemma DDS_parametric [transfer_rule]:
"((A ===> rel_prod B (rel_dds A B)) ===> rel_dds A B) DDS DDS"
by(auto intro!: rel_dds.intros)
lemma run_parametric [transfer_rule]:
"(rel_dds A B ===> A ===> rel_prod B (rel_dds A B)) run run"
by(auto elim: rel_dds.cases)
lemma corec_dds_parametric [transfer_rule]:
"((S ===> A ===> rel_prod B (rel_sum (rel_dds A B) S)) ===> S ===> rel_dds A B) corec_dds corec_dds"
apply(rule rel_funI)+
subgoal premises prems for f g s s' using prems(2)
apply(coinduction arbitrary: s s')
apply simp
apply(rule comp_transfer[THEN rel_funD, THEN rel_funD, rotated])
apply(erule prems(1)[THEN rel_funD])
apply(rule prod.map_transfer[THEN rel_funD, THEN rel_funD, OF id_transfer])
apply(fastforce elim!: rel_sum.cases)
done
done
lemma map_dds_parametric [transfer_rule]:
"((A' ===> A) ===> (B ===> B') ===> rel_dds A B ===> rel_dds A' B') map_dds map_dds"
unfolding map_dds_def by transfer_prover
lemmas map_dds_rel_cong = map_dds_parametric[unfolded rel_fun_def, rule_format, rotated -1]
lemma rel_dds_Grp:
"rel_dds (Grp UNIV f)\<inverse>\<inverse> (Grp UNIV g) = Grp UNIV (map_dds f g)"
apply(rule ext iffI)+
apply(simp add: Grp_apply)
apply(rule sym)
apply(fold rel_dds_eq)
apply(rewrite in "rel_dds _ _ _ \<hole>" map_dds_id[symmetric])
apply(erule map_dds_parametric[THEN rel_funD, THEN rel_funD, THEN rel_funD, rotated -1];
simp add: Grp_apply rel_fun_def)
apply(erule GrpE; clarsimp)
apply(rewrite in "rel_dds _ _ \<hole>" map_dds_id[symmetric])
apply(rule map_dds_parametric[THEN rel_funD, THEN rel_funD, THEN rel_funD, rotated -1])
apply(subst rel_dds_eq; simp)
apply(simp_all add: Grp_apply rel_fun_def)
done
lemma rel_dds_pos_distr [relator_distr]:
"rel_dds A B OO rel_dds C D \<le> rel_dds (A OO C) (B OO D)"
apply (rule predicate2I)
apply (erule relcomppE)
subgoal for x y z
apply (coinduction arbitrary: x y z)
apply (simp)
apply (rule rel_fun_mono[THEN predicate2D, OF order_refl,
of "rel_prod B (rel_dds A B) OO rel_prod D (rel_dds C D)"])
apply (rule order_trans)
apply (rule prod.rel_compp[symmetric, THEN eq_refl])
apply (rule prod.rel_mono[OF order_refl])
apply (blast)
apply (rule rel_fun_pos_distr[THEN predicate2D])
apply (simp)
apply (rule relcomppI)
apply (auto elim: rel_dds.cases)
done
done
lemma Quotient_dds [quot_map]:
assumes "Quotient R1 Abs1 Rep1 T1" and "Quotient R2 Abs2 Rep2 T2"
shows "Quotient (rel_dds R1 R2) (map_dds Rep1 Abs2) (map_dds Abs1 Rep2) (rel_dds T1 T2)"
unfolding Quotient_alt_def5
proof (intro conjI, goal_cases)
case 1
have "rel_dds T1 T2 \<le> rel_dds (Grp UNIV Rep1)\<inverse>\<inverse> (Grp UNIV Abs2)"
apply (rule rel_dds_mono)
apply (rule assms(1)[unfolded Quotient_alt_def5, THEN conjunct2, THEN conjunct1,
unfolded conversep_le_swap])
apply (rule assms(2)[unfolded Quotient_alt_def5, THEN conjunct1])
done
also have "... = Grp UNIV (map_dds Rep1 Abs2)" using rel_dds_Grp .
finally show ?case .
next
case 2
have "Grp UNIV (map_dds Abs1 Rep2) = rel_dds (Grp UNIV Abs1)\<inverse>\<inverse> (Grp UNIV Rep2)"
using rel_dds_Grp[symmetric] .
also have "... \<le> rel_dds T1\<inverse>\<inverse> T2\<inverse>\<inverse>"
apply (rule rel_dds_mono)
apply (simp)
apply (rule assms(1)[unfolded Quotient_alt_def5, THEN conjunct1])
apply (rule assms(2)[unfolded Quotient_alt_def5, THEN conjunct2, THEN conjunct1])
done
finally show ?case by (simp add: rel_dds_conversep)
next
case 3
show ?case
apply (rule antisym)
apply (rule predicate2I)
apply (rule relcomppI)
apply (subst map_dds_id[symmetric])
apply (erule map_dds_rel_cong)
apply (simp_all)[2]
apply (erule assms(1)[THEN Quotient_rep_equiv1])
apply (erule assms(2)[THEN Quotient_equiv_abs1])
apply (subst rel_dds_conversep[symmetric])
subgoal for x y
apply (subgoal_tac "map_dds Rep1 Abs2 y = map_dds Rep1 Abs2 x")
apply (simp)
apply (subst (3) map_dds_id[symmetric])
apply (erule map_dds_rel_cong)
apply (simp_all)
apply (erule assms(1)[THEN Quotient_rep_equiv2])
apply (erule assms(2)[THEN Quotient_equiv_abs2])
apply (rule sym)
apply (subst rel_dds_eq[symmetric])
apply (erule map_dds_rel_cong)
apply (simp, rule assms(1)[THEN Quotient_rep_reflp])
apply (erule assms(2)[THEN Quotient_rel_abs])
done
apply (unfold rel_dds_conversep[symmetric]
assms[unfolded Quotient_alt_def5, THEN conjunct2, THEN conjunct2])
apply (rule rel_dds_pos_distr)
done
qed
text \<open>This is just the co-iterator.\<close>
primcorec dds_of :: "('s \<Rightarrow> 'a \<Rightarrow> ('b \<times> 's)) \<Rightarrow> 's \<Rightarrow> ('a, 'b) dds" where
"run (dds_of f s) = map_prod id (dds_of f) \<circ> f s"
lemma dds_of_parametric [transfer_rule]:
"((S ===> A ===> rel_prod B S) ===> S ===> rel_dds A B) dds_of dds_of"
unfolding dds_of_def by transfer_prover
subsection \<open>Evenness of partial sums\<close>
definition even_psum :: "(int, bool) dds" where
"even_psum = dds_of (\<lambda>psum n. (even (psum + n), psum + n)) 0"
definition even_psum_nat :: "(nat, bool) dds" where
"even_psum_nat = map_dds int id even_psum"
subsection \<open>Composition\<close>
primcorec compose :: "('a, 'b) dds \<Rightarrow> ('b, 'c) dds \<Rightarrow> ('a, 'c) dds" (infixl "\<bullet>" 120) where
"run (S1 \<bullet> S2) = (\<lambda>a. let (b, S1') = run S1 a; (c, S2') = run S2 b in (c, S1' \<bullet> S2'))"
lemma compose_parametric [transfer_rule]:
"(rel_dds A B ===> rel_dds B C ===> rel_dds A C) (\<bullet>) (\<bullet>)"
unfolding compose_def by transfer_prover
text \<open>
For the following lemma, a direct proof by induction is easy as the inner functor of
the @{type dds} codatatype is fairly simple.
\<close>
lemma "map_dds f g S1 \<bullet> S2 = map_dds f id (S1 \<bullet> map_dds g id S2)"
apply(coinduction arbitrary: S1 S2)
apply(auto simp add: case_prod_map_prod rel_fun_def split: prod.split)
done
text \<open>However, we can also follow the systematic route via parametricity:\<close>
lemma compose_map1: "map_dds f g S1 \<bullet> S2 = map_dds f id (S1 \<bullet> map_dds g id S2)"
for S1 :: "('a, 'b) dds" and S2 :: "('b, 'c) dds"
using compose_parametric[of "(Grp UNIV f)\<inverse>\<inverse>" "Grp UNIV g" "Grp UNIV id :: 'c \<Rightarrow> 'c \<Rightarrow> bool"]
apply(rewrite in "_ ===> rel_dds \<hole> _ ===> _" in asm conversep_conversep[symmetric])
apply(rewrite in "_ ===> rel_dds _ \<hole> ===> _" in asm conversep_Grp_id[symmetric])
apply(simp only: rel_dds_conversep rel_dds_Grp)
apply(simp add: rel_fun_def Grp_apply)
done
lemma compose_map2: "S1 \<bullet> map_dds f g S2 = map_dds id g (map_dds id f S1 \<bullet> S2)"
for S1 :: "('a, 'b) dds" and S2 :: "('b, 'c) dds"
using compose_parametric[of "Grp UNIV id :: 'a \<Rightarrow> 'a \<Rightarrow> bool" "(Grp UNIV f)\<inverse>\<inverse>" "Grp UNIV g"]
apply(rewrite in "rel_dds \<hole> _ ===> _" in asm conversep_conversep[symmetric])
apply(rewrite in "_ ===> _ ===> rel_dds \<hole> _" in asm conversep_Grp_id[symmetric])
apply(simp only: rel_dds_conversep rel_dds_Grp)
apply(simp add: rel_fun_def Grp_apply)
done
primcorec parallel :: "('a, 'b) dds \<Rightarrow> ('c, 'd) dds \<Rightarrow> ('a + 'c, 'b + 'd) dds" (infixr "\<parallel>" 130) where
"run (S1 \<parallel> S2) = (\<lambda>x. case x of
Inl a \<Rightarrow> let (b, S1') = run S1 a in (Inl b, S1' \<parallel> S2)
| Inr c \<Rightarrow> let (d, S2') = run S2 c in (Inr d, S1 \<parallel> S2'))"
lemma parallel_parametric [transfer_rule]:
"(rel_dds A B ===> rel_dds C D ===> rel_dds (rel_sum A C) (rel_sum B D)) (\<parallel>) (\<parallel>)"
unfolding parallel_def by transfer_prover
lemma map_parallel:
"map_dds f h S1 \<parallel> map_dds g k S2 = map_dds (map_sum f g) (map_sum h k) (S1 \<parallel> S2)"
using parallel_parametric[where A="(Grp UNIV f)\<inverse>\<inverse>" and B="Grp UNIV h" and
C="(Grp UNIV g)\<inverse>\<inverse>" and D="Grp UNIV k"]
by(simp add: sum.rel_conversep sum.rel_Grp rel_dds_Grp)(simp add: rel_fun_def Grp_apply)
subsection \<open>Graph traversal: refinement and quotients\<close>
lemma finite_Image:
"finite A \<Longrightarrow> finite (R `` A) \<longleftrightarrow> (\<forall>x\<in>A. finite {y. (x, y) \<in> R})"
by(simp add: Image_def)
context includes fset.lifting begin
lift_definition fImage :: "('a \<times> 'b) fset \<Rightarrow> 'a fset \<Rightarrow> 'b fset" is Image parametric Image_parametric
by(auto simp add: finite_Image intro: finite_subset[OF _ finite_imageI])
lemmas fImage_iff = Image_iff[Transfer.transferred]
lemmas fImageI [intro] = ImageI[Transfer.transferred]
lemmas fImageE [elim!] = ImageE[Transfer.transferred]
lemmas rev_fImageI = rev_ImageI[Transfer.transferred]
lemmas fImage_mono = Image_mono[Transfer.transferred]
lifting_update fset.lifting
lifting_forget fset.lifting
end
type_synonym 'a graph = "('a \<times> 'a) fset"
definition traverse :: "'a graph \<Rightarrow> ('a fset, 'a fset) dds" where
"traverse E = dds_of (\<lambda>visited A. ((fImage E A) |-| visited, visited |\<union>| A)) {||}"
type_synonym 'a graph' = "('a \<times> 'a) list"
definition traverse_impl :: "'a graph' \<Rightarrow> ('a list, 'a list) dds" where
"traverse_impl E =
dds_of (\<lambda>visited A. (map snd [(x, y)\<leftarrow>E . x \<in> set A \<and> y |\<notin>| visited],
visited |\<union>| fset_of_list A)) {||}"
definition list_fset_rel :: "'a list \<Rightarrow> 'a fset \<Rightarrow> bool" where
"list_fset_rel xs A \<longleftrightarrow> fset_of_list xs = A"
lemma traverse_refinement: \<comment> \<open>This is the refinement lemma.\<close>
"(list_fset_rel ===> rel_dds list_fset_rel list_fset_rel) traverse_impl traverse"
unfolding traverse_impl_def traverse_def
apply(rule rel_funI)
apply(rule dds_of_parametric[where S="(=)", THEN rel_funD, THEN rel_funD])
- apply(auto simp add: rel_fun_def list_fset_rel_def fset_of_list_elem intro: rev_fimage_eqI)
+ apply(auto simp add: rel_fun_def list_fset_rel_def fset_of_list_elem intro!: rev_image_eqI)
done
lemma fset_of_list_parametric [transfer_rule]:
"(list_all2 A ===> rel_fset A) fset_of_list fset_of_list"
including fset.lifting unfolding rel_fun_def
by transfer(rule list.set_transfer[unfolded rel_fun_def])
lemma traverse_impl_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
shows "(list_all2 (rel_prod A A) ===> rel_dds (list_all2 A) (list_all2 A)) traverse_impl traverse_impl"
unfolding traverse_impl_def by transfer_prover
text \<open>
By constructing finite sets as a quotient of lists, we can synthesise an abstract version
of @{const traverse_impl} automatically, together with a polymorphic refinement lemma.
\<close>
quotient_type 'a fset' = "'a list" / "vimage2p set set (=)"
by (auto intro: equivpI reflpI sympI transpI simp add: vimage2p_def)
lift_definition traverse'' :: "('a \<times> 'a) fset' \<Rightarrow> ('a fset', 'a fset') dds"
is "traverse_impl :: 'a graph' \<Rightarrow> _" parametric traverse_impl_parametric
unfolding traverse_impl_def
apply (rule dds_of_parametric[where S="(=)", THEN rel_funD, THEN rel_funD])
apply (auto simp add: rel_fun_def vimage2p_def fset_of_list_elem)
done
subsection \<open>Generalised rewriting\<close>
definition accumulate :: "('a fset, 'a fset) dds" where
"accumulate = dds_of (\<lambda>A X. (A |\<union>| X, A |\<union>| X)) {||}"
lemma accumulate_mono: "rel_dds (|\<subseteq>|) (|\<subseteq>|) accumulate accumulate"
unfolding accumulate_def
apply (rule dds_of_parametric[THEN rel_funD, THEN rel_funD, of "(|\<subseteq>|)"])
apply (intro rel_funI rel_prod.intros)
apply (erule (1) funion_mono)+
apply (simp)
done
lemma traverse_mono: "((|\<subseteq>|) ===> rel_dds (=) (|\<subseteq>|)) traverse traverse"
unfolding traverse_def
apply (rule rel_funI)
apply (rule dds_of_parametric[THEN rel_funD, THEN rel_funD, of "(=)"])
apply (intro rel_funI rel_prod.intros)
apply (simp)
apply (rule fminus_mono)
apply (erule fImage_mono)
apply (simp_all)
done
lemma
assumes "G |\<subseteq>| H"
shows "rel_dds (=) (|\<subseteq>|) (traverse G \<bullet> accumulate) (traverse H \<bullet> accumulate)"
apply (rule compose_parametric[THEN rel_funD, THEN rel_funD])
apply (rule traverse_mono[THEN rel_funD])
apply (rule assms)
apply (rule accumulate_mono)
done
definition seen :: "('a fset, 'a fset) dds" where
"seen = dds_of (\<lambda>S X. (S |\<inter>| X, S |\<union>| X)) {||}"
lemma seen_mono: "rel_dds (|\<subseteq>|) (|\<subseteq>|) seen seen"
unfolding seen_def
apply (rule dds_of_parametric[THEN rel_funD, THEN rel_funD, of "(|\<subseteq>|)"])
apply (intro rel_funI rel_prod.intros)
apply (erule (1) finter_mono funion_mono)+
apply (simp)
done
lemma
assumes "G |\<subseteq>| H"
shows "rel_dds (=) (|\<subseteq>|) (traverse G \<bullet> seen) (traverse H \<bullet> seen)"
apply (rule compose_parametric[THEN rel_funD, THEN rel_funD])
apply (rule traverse_mono[THEN rel_funD])
apply (rule assms)
apply (rule seen_mono)
done
end
\ No newline at end of file
diff --git a/thys/CakeML_Codegen/Backend/CakeML_Correctness.thy b/thys/CakeML_Codegen/Backend/CakeML_Correctness.thy
--- a/thys/CakeML_Codegen/Backend/CakeML_Correctness.thy
+++ b/thys/CakeML_Codegen/Backend/CakeML_Correctness.thy
@@ -1,1215 +1,1218 @@
subsection \<open>Correctness of compilation\<close>
theory CakeML_Correctness
imports
CakeML_Backend
"../Rewriting/Big_Step_Value_ML"
begin
context cakeml' begin
lemma mk_rec_env_related:
assumes "fmrel (\<lambda>cs (n, e). related_fun cs n e) css (fmap_of_list (map (map_prod Name (map_prod Name id)) funs))"
assumes "fmrel_on_fset (fbind (fmran css) (ids \<circ> Sabs)) related_v \<Gamma>\<^sub>\<Lambda> (fmap_of_ns (sem_env.v env\<^sub>\<Lambda>))"
shows "fmrel related_v (mk_rec_env css \<Gamma>\<^sub>\<Lambda>) (cake_mk_rec_env funs env\<^sub>\<Lambda>)"
proof (rule fmrelI)
fix name
have "rel_option (\<lambda>cs (n, e). related_fun cs n e) (fmlookup css name) (map_of (map (map_prod Name (map_prod Name id)) funs) name)"
using assms by (auto simp: fmap_of_list.rep_eq)
then have "rel_option (\<lambda>cs (n, e). related_fun cs (Name n) e) (fmlookup css name) (map_of funs (as_string name))"
unfolding name.map_of_rekey'
by cases auto
have *: "related_v (Vrecabs css name \<Gamma>\<^sub>\<Lambda>) (Recclosure env\<^sub>\<Lambda> funs (as_string name))"
using assms by (auto intro: related_v.rec_closure)
show "rel_option related_v (fmlookup (mk_rec_env css \<Gamma>\<^sub>\<Lambda>) name) (fmlookup (cake_mk_rec_env funs env\<^sub>\<Lambda>) name)"
unfolding mk_rec_env_def cake_mk_rec_env_def fmap_of_list.rep_eq
apply (simp add: map_of_map_keyed name.map_of_rekey option.rel_map)
apply (rule option.rel_mono_strong)
apply fact
apply (rule *)
done
qed
lemma mk_exp_correctness:
"ids t |\<subseteq>| S \<Longrightarrow> all_consts |\<subseteq>| S \<Longrightarrow> \<not> shadows_consts t \<Longrightarrow> related_exp t (mk_exp S t)"
"ids (Sabs cs) |\<subseteq>| S \<Longrightarrow> all_consts |\<subseteq>| S \<Longrightarrow> \<not> shadows_consts (Sabs cs) \<Longrightarrow> list_all2 (rel_prod related_pat related_exp) cs (mk_clauses S cs)"
"ids t |\<subseteq>| S \<Longrightarrow> all_consts |\<subseteq>| S \<Longrightarrow> \<not> shadows_consts t \<Longrightarrow> related_exp t (mk_con S t)"
proof (induction rule: mk_exp_mk_clauses_mk_con.induct)
case (2 S name)
show ?case
proof (cases "name |\<in>| C")
case True
hence "related_exp (name $$ []) (mk_exp S (Sconst name))"
by (auto intro: related_exp.intros simp del: list_comb.simps)
thus ?thesis
by (simp add: const_sterm_def)
qed (auto intro: related_exp.intros)
next
case (4 S cs)
have "fresh_fNext (S |\<union>| all_consts) |\<notin>| S |\<union>| all_consts"
by (rule fNext_not_member)
hence "fresh_fNext S |\<notin>| S |\<union>| all_consts"
using \<open>all_consts |\<subseteq>| S\<close> by (simp add: sup_absorb1)
hence "fresh_fNext S |\<notin>| ids (Sabs cs) |\<union>| all_consts"
using 4 by auto
show ?case
apply (simp add: Let_def)
apply (rule related_exp.fun)
apply (rule "4.IH"[unfolded mk_clauses.simps])
apply (rule refl)
apply fact+
using \<open>fresh_fNext S |\<notin>| ids (Sabs cs) |\<union>| all_consts\<close> by auto
next
case (5 S t)
show ?case
apply (simp add: mk_con.simps split!: prod.splits sterm.splits if_splits)
subgoal premises prems for args c
proof -
from prems have "t = c $$ args"
apply (fold const_sterm_def)
by (metis fst_conv list_strip_comb snd_conv)
show ?thesis
unfolding \<open>t = _\<close>
apply (rule related_exp.constr)
apply fact
apply (simp add: list.rel_map)
apply (rule list.rel_refl_strong)
apply (rule 5(1))
apply (rule prems(1)[symmetric])
apply (rule refl)
subgoal by (rule prems)
subgoal by assumption
subgoal
using \<open>ids t |\<subseteq>| S\<close> unfolding \<open>t = _\<close>
apply (auto simp: ids_list_comb)
by (meson ffUnion_subset_elem fimage_eqI fset_of_list_elem fset_rev_mp)
subgoal by (rule 5)
subgoal
using \<open>\<not> shadows_consts t\<close> unfolding \<open>t = _\<close>
unfolding shadows.list_comb
by (auto simp: list_ex_iff)
done
qed
using 5 by (auto split: prod.splits sterm.splits)
next
case (6 S cs)
have "list_all2 (\<lambda>x y. rel_prod related_pat related_exp x (case y of (pat, t) \<Rightarrow> (mk_ml_pat (mk_pat pat), mk_con (frees pat |\<union>| S) t))) cs cs"
proof (rule list.rel_refl_strong, safe intro!: rel_prod.intros)
fix pat rhs
assume "(pat, rhs) \<in> set cs"
hence "consts rhs |\<subseteq>| S"
using \<open>ids (Sabs cs) |\<subseteq>| S\<close>
unfolding ids_def
apply auto
apply (drule ffUnion_least_rev)+
apply (auto simp: fset_of_list_elem elim!: fBallE)
done
have "frees rhs |\<subseteq>| frees pat |\<union>| S"
using \<open>ids (Sabs cs) |\<subseteq>| S\<close> \<open>(pat, rhs) \<in> set cs\<close>
unfolding ids_def
apply auto
apply (drule ffUnion_least_rev)+
apply (auto simp: fset_of_list_elem elim!: fBallE)
done
have "\<not> shadows_consts rhs"
using \<open>(pat, rhs) \<in> set cs\<close> 6
by (auto simp: list_ex_iff)
show "related_exp rhs (mk_con (frees pat |\<union>| S) rhs)"
apply (rule 6)
apply fact
subgoal by simp
subgoal
unfolding ids_def
using \<open>consts rhs |\<subseteq>| S\<close> \<open>frees rhs |\<subseteq>| frees pat |\<union>| S\<close>
by auto
- subgoal using 6(3) by fast
+ subgoal using 6(3) by auto
subgoal by fact
done
qed
thus ?case
by (simp add: list.rel_map)
qed (auto intro: related_exp.intros simp: ids_def fdisjnt_alt_def)
context begin
private lemma semantic_correctness0:
fixes exp
assumes "cupcake_evaluate_single env exp r" "is_cupcake_all_env env"
assumes "fmrel_on_fset (ids t) related_v \<Gamma> (fmap_of_ns (sem_env.v env))"
assumes "related_exp t exp"
assumes "wellformed t" "wellformed_venv \<Gamma>"
assumes "closed_venv \<Gamma>" "closed_except t (fmdom \<Gamma>)"
assumes "fmpred (\<lambda>_. vwelldefined') \<Gamma>" "consts t |\<subseteq>| fmdom \<Gamma> |\<union>| C"
assumes "fdisjnt C (fmdom \<Gamma>)"
assumes "\<not> shadows_consts t" "not_shadows_vconsts_env \<Gamma>"
shows "if_rval (\<lambda>ml_v. \<exists>v. \<Gamma> \<turnstile>\<^sub>v t \<down> v \<and> related_v v ml_v) r"
using assms proof (induction arbitrary: \<Gamma> t)
case (con1 env cn es ress ml_vs ml_v')
obtain name ts where "cn = Some (Short (as_string name))" "name |\<in>| C" "t = name $$ ts" "list_all2 related_exp ts es"
using \<open>related_exp t (Con cn es)\<close>
by cases auto
with con1 obtain tid where "ml_v' = Conv (Some (id_to_n (Short (as_string name)), tid)) (rev ml_vs)"
by (auto split: option.splits)
have "ress = map Rval ml_vs"
using con1 by auto
define ml_vs' where "ml_vs' = rev ml_vs"
note IH = \<open>list_all2_shortcircuit _ _ _\<close>[
unfolded \<open>ress = _\<close> list_all2_shortcircuit_rval list_all2_rev1,
folded ml_vs'_def]
moreover have
"list_all wellformed ts" "list_all (\<lambda>t. \<not> shadows_consts t) ts"
"list_all (\<lambda>t. consts t |\<subseteq>| fmdom \<Gamma> |\<union>| C) ts" "list_all (\<lambda>t. closed_except t (fmdom \<Gamma>)) ts"
subgoal
using \<open>wellformed t\<close> unfolding \<open>t = _\<close> wellformed.list_comb by simp
subgoal
using \<open>\<not> shadows_consts t\<close> unfolding \<open>t = _\<close> shadows.list_comb
by (simp add: list_all_iff list_ex_iff)
subgoal
using \<open>consts t |\<subseteq>| fmdom \<Gamma> |\<union>| C\<close>
unfolding list_all_iff
by (metis Ball_set \<open>t = name $$ ts\<close> con1.prems(9) special_constants.sconsts_list_comb)
subgoal
using \<open>closed_except t (fmdom \<Gamma>)\<close> unfolding \<open>t = _\<close> closed.list_comb by simp
done
moreover have
"list_all (\<lambda>t'. fmrel_on_fset (ids t') related_v \<Gamma> (fmap_of_ns (sem_env.v env))) ts"
proof (standard, rule fmrel_on_fsubset)
fix t'
assume "t' \<in> set ts"
thus "ids t' |\<subseteq>| ids t"
unfolding \<open>t = _\<close>
apply (simp add: ids_list_comb)
apply (subst (2) ids_def)
apply simp
apply (rule fsubset_finsertI2)
apply (auto simp: fset_of_list_elem intro!: ffUnion_subset_elem)
done
show "fmrel_on_fset (ids t) related_v \<Gamma> (fmap_of_ns (sem_env.v env))"
by fact
qed
ultimately obtain us where "list_all2 (veval' \<Gamma>) ts us" "list_all2 related_v us ml_vs'"
using \<open>list_all2 related_exp ts es\<close>
proof (induction es ml_vs' arbitrary: ts thesis rule: list.rel_induct)
case (Cons e es ml_v ml_vs ts0)
then obtain t ts where "ts0 = t # ts" "related_exp t e" by (cases ts0) auto
with Cons have "list_all2 related_exp ts es" by simp
with Cons obtain us where "list_all2 (veval' \<Gamma>) ts us" "list_all2 related_v us ml_vs"
unfolding \<open>ts0 = _\<close>
by auto
from Cons.hyps[simplified, THEN conjunct2, rule_format, of t \<Gamma>]
obtain u where "\<Gamma> \<turnstile>\<^sub>v t \<down> u " "related_v u ml_v"
proof
show
"is_cupcake_all_env env" "related_exp t e" "wellformed_venv \<Gamma>" "closed_venv \<Gamma>"
"fmpred (\<lambda>_. vwelldefined') \<Gamma>" "fdisjnt C (fmdom \<Gamma>)"
"not_shadows_vconsts_env \<Gamma>"
by fact+
next
show
"wellformed t" "\<not> shadows_consts t" "closed_except t (fmdom \<Gamma>)"
"consts t |\<subseteq>| fmdom \<Gamma> |\<union>| C" "fmrel_on_fset (ids t) related_v \<Gamma> (fmap_of_ns (sem_env.v env))"
using Cons unfolding \<open>ts0 = _\<close>
by auto
qed blast
show ?case
apply (rule Cons(3)[of "u # us"])
unfolding \<open>ts0 = _\<close>
apply auto
apply fact+
done
qed auto
show ?case
apply simp
apply (intro exI conjI)
unfolding \<open>t = _\<close>
apply (rule veval'.constr)
apply fact+
unfolding \<open>ml_v' = _\<close>
apply (subst ml_vs'_def[symmetric])
apply simp
apply (rule related_v.conv)
apply fact
done
next
case (var1 env id ml_v)
from \<open>related_exp t (Var id)\<close> obtain name where "id = Short (as_string name)"
by cases auto
with var1 have "cupcake_nsLookup (sem_env.v env) (as_string name) = Some ml_v"
by auto
from \<open>related_exp t (Var id)\<close> consider
(var) "t = Svar name"
| (const) "t = Sconst name" "name |\<notin>| C"
unfolding \<open>id = _\<close>
apply (cases t)
using name.expand by blast+
thus ?case
proof cases
case var
hence "name |\<in>| ids t"
unfolding ids_def by simp
have "rel_option related_v (fmlookup \<Gamma> name) (cupcake_nsLookup (sem_env.v env) (as_string name))"
using \<open>fmrel_on_fset (ids t) _ _ _\<close>
apply -
apply (drule fmrel_on_fsetD[OF \<open>name |\<in>| ids t\<close>])
apply simp
done
then obtain v where "related_v v ml_v" "fmlookup \<Gamma> name = Some v"
using \<open>cupcake_nsLookup (sem_env.v env) _ = _\<close>
by cases auto
show ?thesis
unfolding \<open>t = _\<close>
apply simp
apply (rule exI)
apply (rule conjI)
apply (rule veval'.var)
apply fact+
done
next
case const
hence "name |\<in>| ids t"
unfolding ids_def by simp
have "rel_option related_v (fmlookup \<Gamma> name) (cupcake_nsLookup (sem_env.v env) (as_string name))"
using \<open>fmrel_on_fset (ids t) _ _ _\<close>
apply -
apply (drule fmrel_on_fsetD[OF \<open>name |\<in>| ids t\<close>])
apply simp
done
then obtain v where "related_v v ml_v" "fmlookup \<Gamma> name = Some v"
using \<open>cupcake_nsLookup (sem_env.v env) _ = _\<close>
by cases auto
show ?thesis
unfolding \<open>t = _\<close>
apply simp
apply (rule exI)
apply (rule conjI)
apply (rule veval'.const)
apply fact+
done
qed
next
case (fn env n u)
obtain n' where "n = as_string n'"
by (metis name.sel)
obtain cs ml_cs
where "t = Sabs cs" "u = Mat (Var (Short (as_string n'))) ml_cs" "n' |\<notin>| ids (Sabs cs)" "n' |\<notin>| all_consts"
and "list_all2 (rel_prod related_pat related_exp) cs ml_cs"
using \<open>related_exp t (Fun n u)\<close> unfolding \<open>n = _\<close>
by cases (auto dest: name.expand)
obtain ns where "fmap_of_ns (sem_env.v env) = fmap_of_list ns"
apply (cases env)
apply simp
subgoal for v by (cases v) simp
done
show ?case
apply simp
apply (rule exI)
apply (rule conjI)
unfolding \<open>t = _\<close>
apply (rule veval'.abs)
unfolding \<open>n = _\<close>
apply (rule related_v.closure)
unfolding \<open>u = _\<close>
apply (subst related_fun_alt_def; rule conjI)
apply fact
apply (rule conjI; fact)
using \<open>fmrel_on_fset (ids t) _ _ _\<close>
unfolding \<open>t = _\<close> \<open>fmap_of_ns (sem_env.v env) = _\<close>
by simp
next
case (app1 env exps ress ml_vs env' exp' bv)
from \<open>related_exp t _\<close> obtain exp\<^sub>1 exp\<^sub>2 t\<^sub>1 t\<^sub>2
where "rev exps = [exp\<^sub>2, exp\<^sub>1]" "exps = [exp\<^sub>1, exp\<^sub>2]" "t = t\<^sub>1 $\<^sub>s t\<^sub>2"
and "related_exp t\<^sub>1 exp\<^sub>1" "related_exp t\<^sub>2 exp\<^sub>2"
by cases auto
moreover from app1 have "ress = map Rval ml_vs"
by auto
ultimately obtain ml_v\<^sub>1 ml_v\<^sub>2 where "ml_vs = [ml_v\<^sub>2, ml_v\<^sub>1]"
using app1(1)
by (smt list_all2_shortcircuit_rval list_all2_Cons1 list_all2_Nil)
have "is_cupcake_exp exp\<^sub>1" "is_cupcake_exp exp\<^sub>2"
using app1 unfolding \<open>exps = _\<close> by (auto dest: related_exp_is_cupcake)
moreover have "fmrel_on_fset (ids t\<^sub>1) related_v \<Gamma> (fmap_of_ns (sem_env.v env))"
using app1 unfolding ids_def \<open>t = _\<close>
by (auto intro: fmrel_on_fsubset)
moreover have "fmrel_on_fset (ids t\<^sub>2) related_v \<Gamma> (fmap_of_ns (sem_env.v env))"
using app1 unfolding ids_def \<open>t = _\<close>
by (auto intro: fmrel_on_fsubset)
ultimately have
"cupcake_evaluate_single env exp\<^sub>1 (Rval ml_v\<^sub>1)" "cupcake_evaluate_single env exp\<^sub>2 (Rval ml_v\<^sub>2)" and
"\<exists>t\<^sub>1'. \<Gamma> \<turnstile>\<^sub>v t\<^sub>1 \<down> t\<^sub>1' \<and> related_v t\<^sub>1' ml_v\<^sub>1" "\<exists>t\<^sub>2'. \<Gamma> \<turnstile>\<^sub>v t\<^sub>2 \<down> t\<^sub>2' \<and> related_v t\<^sub>2' ml_v\<^sub>2"
using app1 \<open>related_exp t\<^sub>1 exp\<^sub>1\<close> \<open>related_exp t\<^sub>2 exp\<^sub>2\<close>
unfolding \<open>ress = _\<close> \<open>exps = _\<close> \<open>ml_vs = _\<close> \<open>t = _\<close>
by (auto simp: closed_except_def)
then obtain v\<^sub>1 v\<^sub>2
where "\<Gamma> \<turnstile>\<^sub>v t\<^sub>1 \<down> v\<^sub>1" "related_v v\<^sub>1 ml_v\<^sub>1"
and "\<Gamma> \<turnstile>\<^sub>v t\<^sub>2 \<down> v\<^sub>2" "related_v v\<^sub>2 ml_v\<^sub>2"
by blast
have "is_cupcake_value ml_v\<^sub>1"
by (rule cupcake_single_preserve) fact+
moreover have "is_cupcake_value ml_v\<^sub>2"
by (rule cupcake_single_preserve) fact+
ultimately have "list_all is_cupcake_value (rev ml_vs)"
unfolding \<open>ml_vs = _\<close> by simp
hence "is_cupcake_exp exp'" "is_cupcake_all_env env'"
using \<open>do_opapp _ = _\<close> by (metis cupcake_opapp_preserve)+
have "vclosed v\<^sub>1"
proof (rule veval'_closed)
show "closed_except t\<^sub>1 (fmdom \<Gamma>)"
using \<open>closed_except _ (fmdom \<Gamma>)\<close>
unfolding \<open>t = _\<close>
by (simp add: closed_except_def)
next
show "wellformed t\<^sub>1"
using \<open>wellformed t\<close> unfolding \<open>t = _\<close>
by simp
qed fact+
have "vclosed v\<^sub>2"
apply (rule veval'_closed)
apply fact
using app1 unfolding \<open>t = _\<close> by (auto simp: closed_except_def)
have "vwellformed v\<^sub>1"
apply (rule veval'_wellformed)
apply fact
using app1 unfolding \<open>t = _\<close> by auto
have "vwellformed v\<^sub>2"
apply (rule veval'_wellformed)
apply fact
using app1 unfolding \<open>t = _\<close> by auto
have "vwelldefined' v\<^sub>1"
apply (rule veval'_welldefined')
apply fact
using app1 unfolding \<open>t = _\<close> by auto
have "vwelldefined' v\<^sub>2"
apply (rule veval'_welldefined')
apply fact
using app1 unfolding \<open>t = _\<close> by auto
have "not_shadows_vconsts v\<^sub>1"
apply (rule veval'_shadows)
apply fact
using app1 unfolding \<open>t = _\<close> by auto
have "not_shadows_vconsts v\<^sub>2"
apply (rule veval'_shadows)
apply fact
using app1 unfolding \<open>t = _\<close> by auto
show ?case
proof (rule if_rvalI)
fix ml_v
assume "bv = Rval ml_v"
show "\<exists>v. \<Gamma> \<turnstile>\<^sub>v t \<down> v \<and> related_v v ml_v"
using \<open>do_opapp _ = _\<close>
proof (cases rule: do_opapp_cases)
case (closure env\<^sub>\<Lambda> n)
then have closure':
"ml_v\<^sub>1 = Closure env\<^sub>\<Lambda> (as_string (Name n)) exp'"
"env' = update_v (\<lambda>_. nsBind (as_string (Name n)) ml_v\<^sub>2 (sem_env.v env\<^sub>\<Lambda>)) env\<^sub>\<Lambda>"
unfolding \<open>ml_vs = _\<close> by auto
obtain \<Gamma>\<^sub>\<Lambda> cs
where "v\<^sub>1 = Vabs cs \<Gamma>\<^sub>\<Lambda>" "related_fun cs (Name n) exp'"
and "fmrel_on_fset (ids (Sabs cs)) related_v \<Gamma>\<^sub>\<Lambda> (fmap_of_ns (sem_env.v env\<^sub>\<Lambda>))"
using \<open>related_v v\<^sub>1 ml_v\<^sub>1\<close> unfolding \<open>ml_v\<^sub>1 = _\<close>
by cases auto
then obtain ml_cs
where "exp' = Mat (Var (Short (as_string (Name n)))) ml_cs" "Name n |\<notin>| ids (Sabs cs)" "Name n |\<notin>| all_consts"
and "list_all2 (rel_prod related_pat related_exp) cs ml_cs"
by (auto elim: related_funE)
hence "cupcake_evaluate_single env' (Mat (Var (Short (as_string (Name n)))) ml_cs) (Rval ml_v)"
using \<open>cupcake_evaluate_single env' exp' bv\<close>
unfolding \<open>bv = _\<close>
by simp
then obtain m_env v' ml_rhs ml_pat
where "cupcake_evaluate_single env' (Var (Short (as_string (Name n)))) (Rval v')"
and "cupcake_match_result (sem_env.c env') v' ml_cs Bindv = Rval (ml_rhs, ml_pat, m_env)"
and "cupcake_evaluate_single (env' \<lparr> sem_env.v := nsAppend (alist_to_ns m_env) (sem_env.v env') \<rparr>) ml_rhs (Rval ml_v)"
by cases auto
have
"closed_venv (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>)" "wellformed_venv (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>)"
"not_shadows_vconsts_env (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>)" "fmpred (\<lambda>_. vwelldefined') (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>)"
using \<open>vclosed v\<^sub>1\<close> \<open>vclosed v\<^sub>2\<close>
using \<open>vwellformed v\<^sub>1\<close> \<open>vwellformed v\<^sub>2\<close>
using \<open>not_shadows_vconsts v\<^sub>1\<close> \<open>not_shadows_vconsts v\<^sub>2\<close>
using \<open>vwelldefined' v\<^sub>1\<close> \<open>vwelldefined' v\<^sub>2\<close>
unfolding \<open>v\<^sub>1 = _\<close>
by auto
have "closed_except (Sabs cs) (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))"
using \<open>vclosed v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
apply (auto simp: Sterm.closed_except_simps list_all_iff)
apply (auto simp: closed_except_def)
done
have "consts (Sabs cs) |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>) |\<union>| C"
using \<open>vwelldefined' v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
unfolding sconsts_sabs
by (auto simp: list_all_iff)
have "\<not> shadows_consts (Sabs cs)"
using \<open>not_shadows_vconsts v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
by (auto simp: list_all_iff list_ex_iff)
have "fdisjnt C (fmdom \<Gamma>\<^sub>\<Lambda>)"
using \<open>vwelldefined' v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
by simp
have "if_rval (\<lambda>ml_v. \<exists>v. fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> \<turnstile>\<^sub>v Sabs cs $\<^sub>s Svar (Name n) \<down> v \<and> related_v v ml_v) bv"
proof (rule app1(2))
show "fmrel_on_fset (ids (Sabs cs $\<^sub>s Svar (Name n))) related_v (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>) (fmap_of_ns (sem_env.v env'))"
unfolding closure'
apply (simp del: frees_sterm.simps(3) consts_sterm.simps(3) name.sel add: ids_def split!: sem_env.splits)
apply (rule fmrel_on_fset_updateI)
apply (fold ids_def)
using \<open>fmrel_on_fset (ids (Sabs cs)) related_v \<Gamma>\<^sub>\<Lambda> _\<close> apply simp
apply (rule \<open>related_v v\<^sub>2 ml_v\<^sub>2\<close>)
done
next
show "wellformed (Sabs cs $\<^sub>s Svar (Name n))"
using \<open>vwellformed v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
by simp
next
show "related_exp (Sabs cs $\<^sub>s Svar (Name n)) exp'"
unfolding \<open>exp' = _\<close>
using \<open>list_all2 (rel_prod related_pat related_exp) cs ml_cs\<close>
by (auto intro:related_exp.intros simp del: name.sel)
next
show "closed_except (Sabs cs $\<^sub>s Svar (Name n)) (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))"
using \<open>closed_except (Sabs cs) (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))\<close> by (simp add: closed_except_def)
next
show "\<not> shadows_consts (Sabs cs $\<^sub>s Svar (Name n))"
using \<open>\<not> shadows_consts (Sabs cs)\<close> \<open>Name n |\<notin>| all_consts\<close> by simp
next
show "consts (Sabs cs $\<^sub>s Svar (Name n)) |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>) |\<union>| C"
using \<open>consts (Sabs cs) |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>) |\<union>| C\<close> by simp
next
show "fdisjnt C (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))"
using \<open>Name n |\<notin>| all_consts\<close> \<open>fdisjnt C (fmdom \<Gamma>\<^sub>\<Lambda>)\<close>
unfolding fdisjnt_alt_def all_consts_def by auto
qed fact+
then obtain v where "fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> \<turnstile>\<^sub>v Sabs cs $\<^sub>s Svar (Name n) \<down> v" "related_v v ml_v"
unfolding \<open>bv = _\<close>
by auto
then obtain env pat rhs
where "vfind_match cs v\<^sub>2 = Some (env, pat, rhs)"
and "fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env \<turnstile>\<^sub>v rhs \<down> v"
by (auto elim: veval'_sabs_svarE)
hence "(pat, rhs) \<in> set cs" "vmatch (mk_pat pat) v\<^sub>2 = Some env"
by (metis vfind_match_elem)+
hence "linear pat" "wellformed rhs"
using \<open>vwellformed v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
by (auto simp: list_all_iff)
hence "frees pat = patvars (mk_pat pat)"
by (simp add: mk_pat_frees)
hence "fmdom env = frees pat"
apply simp
apply (rule vmatch_dom)
apply (rule \<open>vmatch (mk_pat pat) v\<^sub>2 = Some env\<close>)
done
obtain v' where "\<Gamma>\<^sub>\<Lambda> ++\<^sub>f env \<turnstile>\<^sub>v rhs \<down> v'" "v' \<approx>\<^sub>e v"
proof (rule veval'_agree_eq)
show "fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env \<turnstile>\<^sub>v rhs \<down> v" by fact
next
have *: "Name n |\<notin>| ids rhs" if "Name n |\<notin>| fmdom env"
proof
assume "Name n |\<in>| ids rhs"
thus False
unfolding ids_def
proof (cases rule: funion_strictE)
case A
moreover have "Name n |\<notin>| frees pat"
using that unfolding \<open>fmdom env = frees pat\<close> .
ultimately have "Name n |\<in>| frees (Sabs cs)"
apply auto
unfolding ffUnion_alt_def
apply simp
apply (rule fBexI[where x = "(pat, rhs)"])
apply (auto simp: fset_of_list_elem)
apply (rule \<open>(pat, rhs) \<in> set cs\<close>)
done
thus ?thesis
using \<open>Name n |\<notin>| ids (Sabs cs)\<close> unfolding ids_def
by blast
next
case B
hence "Name n |\<in>| consts (Sabs cs)"
apply auto
unfolding ffUnion_alt_def
apply simp
apply (rule fBexI[where x = "(pat, rhs)"])
apply (auto simp: fset_of_list_elem)
apply (rule \<open>(pat, rhs) \<in> set cs\<close>)
done
thus ?thesis
using \<open>Name n |\<notin>| ids (Sabs cs)\<close> unfolding ids_def
by blast
qed
qed
show "fmrel_on_fset (ids rhs) erelated (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f env) (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env)"
apply rule
apply auto
apply (rule option.rel_refl; rule erelated_refl)
using * apply auto[]
apply (rule option.rel_refl; rule erelated_refl)+
done
next
show "closed_venv (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env)"
apply rule
apply fact
apply (rule vclosed.vmatch_env)
apply fact
apply fact
done
next
show "wellformed_venv (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env)"
apply rule
apply fact
apply (rule vwellformed.vmatch_env)
apply fact
apply fact
done
next
show "closed_except rhs (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env))"
using \<open>fmdom env = frees pat\<close> \<open>(pat, rhs) \<in> set cs\<close>
using \<open>closed_except (Sabs cs) (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))\<close>
by (auto simp: Sterm.closed_except_simps list_all_iff)
next
show "wellformed rhs" by fact
next
show "consts rhs |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env) |\<union>| C"
using \<open>consts (Sabs cs) |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>) |\<union>| C\<close> \<open>(pat, rhs) \<in> set cs\<close>
unfolding sconsts_sabs
by (auto simp: list_all_iff)
next
have "fdisjnt C (fmdom env)"
using \<open>(pat, rhs) \<in> set cs\<close> \<open>\<not> shadows_consts (Sabs cs)\<close>
unfolding \<open>fmdom env = frees pat\<close>
by (auto simp: list_ex_iff fdisjnt_alt_def all_consts_def)
thus "fdisjnt C (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env))"
using \<open>Name n |\<notin>| all_consts\<close> \<open>fdisjnt C (fmdom \<Gamma>\<^sub>\<Lambda>)\<close>
unfolding fdisjnt_alt_def
by (auto simp: all_consts_def)
next
show "\<not> shadows_consts rhs"
using \<open>(pat, rhs) \<in> set cs\<close> \<open>\<not> shadows_consts (Sabs cs)\<close>
by (auto simp: list_ex_iff)
next
have "not_shadows_vconsts_env env"
by (rule not_shadows_vconsts.vmatch_env) fact+
thus "not_shadows_vconsts_env (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env)"
using \<open>not_shadows_vconsts_env (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>)\<close> by blast
next
have "fmpred (\<lambda>_. vwelldefined') env"
by (rule vmatch_welldefined') fact+
thus "fmpred (\<lambda>_. vwelldefined') (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env)"
using \<open>fmpred (\<lambda>_. vwelldefined') (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>)\<close> by blast
qed blast
show ?thesis
apply (intro exI conjI)
unfolding \<open>t = _\<close>
apply (rule veval'.comb)
using \<open>\<Gamma> \<turnstile>\<^sub>v t\<^sub>1 \<down> v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
apply blast
apply fact
apply fact+
apply (rule related_v_ext)
apply fact+
done
next
case (recclosure env\<^sub>\<Lambda> funs name n)
with recclosure have recclosure':
"ml_v\<^sub>1 = Recclosure env\<^sub>\<Lambda> funs name"
"env' = update_v (\<lambda>_. nsBind (as_string (Name n)) ml_v\<^sub>2 (build_rec_env funs env\<^sub>\<Lambda> (sem_env.v env\<^sub>\<Lambda>))) env\<^sub>\<Lambda>"
unfolding \<open>ml_vs = _\<close> by auto
obtain \<Gamma>\<^sub>\<Lambda> css
where "v\<^sub>1 = Vrecabs css (Name name) \<Gamma>\<^sub>\<Lambda>"
and "fmrel_on_fset (fbind (fmran css) (ids \<circ> Sabs)) related_v \<Gamma>\<^sub>\<Lambda> (fmap_of_ns (sem_env.v env\<^sub>\<Lambda>))"
and "fmrel (\<lambda>cs (n, e). related_fun cs n e) css (fmap_of_list (map (map_prod Name (map_prod Name id)) funs))"
using \<open>related_v v\<^sub>1 ml_v\<^sub>1\<close> unfolding \<open>ml_v\<^sub>1 = _\<close>
by cases auto
from \<open>fmrel _ _ _\<close> have "rel_option (\<lambda>cs (n, e). related_fun cs (Name n) e) (fmlookup css (Name name)) (find_recfun name funs)"
apply -
apply (subst option.rel_sel)
apply auto
apply (drule fmrel_fmdom_eq)
apply (drule fmdom_notI)
using \<open>v\<^sub>1 = Vrecabs css (Name name) \<Gamma>\<^sub>\<Lambda>\<close> \<open>vwellformed v\<^sub>1\<close> apply auto[1]
using recclosure(3) apply auto[1]
apply (erule fmrel_cases[where x = "Name name"])
apply simp
apply (subst (asm) fmlookup_of_list)
apply (simp add: name.map_of_rekey')
by blast
then obtain cs where "fmlookup css (Name name) = Some cs" "related_fun cs (Name n) exp'"
using \<open>find_recfun _ _ = _\<close>
by cases auto
then obtain ml_cs
where "exp' = Mat (Var (Short (as_string (Name n)))) ml_cs" "Name n |\<notin>| ids (Sabs cs)" "Name n |\<notin>| all_consts"
and "list_all2 (rel_prod related_pat related_exp) cs ml_cs"
by (auto elim: related_funE)
hence "cupcake_evaluate_single env' (Mat (Var (Short n)) ml_cs) (Rval ml_v)"
using \<open>cupcake_evaluate_single env' exp' bv\<close>
unfolding \<open>bv = _\<close>
by simp
then obtain m_env v' ml_rhs ml_pat
where "cupcake_evaluate_single env' (Var (Short n)) (Rval v')"
and "cupcake_match_result (sem_env.c env') v' ml_cs Bindv = Rval (ml_rhs, ml_pat, m_env)"
and "cupcake_evaluate_single (env' \<lparr> sem_env.v := nsAppend (alist_to_ns m_env) (sem_env.v env') \<rparr>) ml_rhs (Rval ml_v)"
by cases auto
have "closed_venv (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>))"
using \<open>vclosed v\<^sub>1\<close> \<open>vclosed v\<^sub>2\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
unfolding \<open>v\<^sub>1 = _\<close> mk_rec_env_def
apply auto
apply rule
apply rule
apply (auto intro: fmdomI)
done
have "wellformed_venv (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>))"
using \<open>vwellformed v\<^sub>1\<close> \<open>vwellformed v\<^sub>2\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
unfolding \<open>v\<^sub>1 = _\<close> mk_rec_env_def
apply auto
apply rule
apply rule
apply (auto intro: fmdomI)
done
have "not_shadows_vconsts_env (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>))"
using \<open>not_shadows_vconsts v\<^sub>1\<close> \<open>not_shadows_vconsts v\<^sub>2\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
unfolding \<open>v\<^sub>1 = _\<close> mk_rec_env_def
apply auto
apply rule
apply rule
apply (auto intro: fmdomI)
done
have "fmpred (\<lambda>_. vwelldefined') (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>))"
using \<open>vwelldefined' v\<^sub>1\<close> \<open>vwelldefined' v\<^sub>2\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
unfolding \<open>v\<^sub>1 = _\<close> mk_rec_env_def
apply auto
apply rule
apply rule
apply (auto intro: fmdomI)
done
have "closed_except (Sabs cs) (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))"
using \<open>vclosed v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
apply (auto simp: Sterm.closed_except_simps list_all_iff)
using \<open>fmlookup css (Name name) = Some cs\<close>
apply (auto simp: closed_except_def dest!: fmpredD[where m = css])
done
have "consts (Sabs cs) |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>) |\<union>| (C |\<union>| fmdom css)"
using \<open>vwelldefined' v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
unfolding sconsts_sabs
using \<open>fmlookup css (Name name) = Some cs\<close>
by (auto simp: list_all_iff dest!: fmpredD[where m = css])
have "\<not> shadows_consts (Sabs cs)"
using \<open>not_shadows_vconsts v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
by (auto simp: list_all_iff list_ex_iff)
have "fdisjnt C (fmdom \<Gamma>\<^sub>\<Lambda>)"
using \<open>vwelldefined' v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
by auto
have "if_rval (\<lambda>ml_v. \<exists>v. fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) \<turnstile>\<^sub>v Sabs cs $\<^sub>s Svar (Name n) \<down> v \<and> related_v v ml_v) bv"
proof (rule app1(2))
have "fmrel_on_fset (ids (Sabs cs)) related_v \<Gamma>\<^sub>\<Lambda> (fmap_of_ns (sem_env.v env\<^sub>\<Lambda>))"
apply (rule fmrel_on_fsubset)
apply fact
apply (subst funion_image_bind_eq[symmetric])
apply (rule ffUnion_subset_elem)
apply (subst fimage_iff)
apply (rule fBexI)
apply simp
apply (rule fmranI)
apply fact
done
have "fmrel_on_fset (ids (Sabs cs)) related_v (mk_rec_env css \<Gamma>\<^sub>\<Lambda>) (cake_mk_rec_env funs env\<^sub>\<Lambda>)"
apply rule
apply (rule mk_rec_env_related[THEN fmrelD])
apply (rule \<open>fmrel _ css _\<close>)
apply (rule \<open>fmrel_on_fset (fbind _ _) related_v \<Gamma>\<^sub>\<Lambda> _\<close>)
done
show "fmrel_on_fset (ids (Sabs cs $\<^sub>s Svar (Name n))) related_v (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>)) (fmap_of_ns (sem_env.v env'))"
unfolding recclosure'
apply (simp del: frees_sterm.simps(3) consts_sterm.simps(3) name.sel add: ids_def split!: sem_env.splits)
apply (rule fmrel_on_fset_updateI)
unfolding build_rec_env_fmap
apply (rule fmrel_on_fset_addI)
apply (fold ids_def)
subgoal
using \<open>fmrel_on_fset (ids (Sabs cs)) related_v \<Gamma>\<^sub>\<Lambda> _\<close> by simp
subgoal
using \<open>fmrel_on_fset (ids (Sabs cs)) related_v (mk_rec_env css \<Gamma>\<^sub>\<Lambda>) _\<close> by simp
apply (rule \<open>related_v v\<^sub>2 ml_v\<^sub>2\<close>)
done
next
show "wellformed (Sabs cs $\<^sub>s Svar (Name n))"
using \<open>vwellformed v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
by auto
next
show "related_exp (Sabs cs $\<^sub>s Svar (Name n)) exp'"
unfolding \<open>exp' = _\<close>
apply (rule related_exp.intros)
apply fact
apply (rule related_exp.intros)
done
next
show "closed_except (Sabs cs $\<^sub>s Svar (Name n)) (fmdom (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>)))"
using \<open>closed_except (Sabs cs) (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))\<close>
by (auto simp: list_all_iff closed_except_def)
next
show "\<not> shadows_consts (Sabs cs $\<^sub>s Svar (Name n))"
using \<open>\<not> shadows_consts (Sabs cs)\<close> \<open>Name n |\<notin>| all_consts\<close> by simp
next
show "consts (Sabs cs $\<^sub>s Svar (Name n)) |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>)) |\<union>| C"
using \<open>consts (Sabs cs) |\<subseteq>| _\<close> unfolding mk_rec_env_def
by auto
next
show "fdisjnt C (fmdom (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>)))"
using \<open>Name n |\<notin>| all_consts\<close> \<open>fdisjnt C (fmdom \<Gamma>\<^sub>\<Lambda>)\<close> \<open>vwelldefined' v\<^sub>1\<close>
unfolding mk_rec_env_def \<open>v\<^sub>1 = _\<close>
by (auto simp: fdisjnt_alt_def all_consts_def)
qed fact+
then obtain v
where "fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) \<turnstile>\<^sub>v Sabs cs $\<^sub>s Svar (Name n) \<down> v" "related_v v ml_v"
unfolding \<open>bv = _\<close>
by auto
then obtain env pat rhs
where "vfind_match cs v\<^sub>2 = Some (env, pat, rhs)"
and "fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env \<turnstile>\<^sub>v rhs \<down> v"
by (auto elim: veval'_sabs_svarE)
hence "(pat, rhs) \<in> set cs" "vmatch (mk_pat pat) v\<^sub>2 = Some env"
by (metis vfind_match_elem)+
hence "linear pat" "wellformed rhs"
using \<open>vwellformed v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
using \<open>fmlookup css (Name name) = Some cs\<close>
by (auto simp: list_all_iff)
hence "frees pat = patvars (mk_pat pat)"
by (simp add: mk_pat_frees)
hence "fmdom env = frees pat"
apply simp
apply (rule vmatch_dom)
apply (rule \<open>vmatch (mk_pat pat) v\<^sub>2 = Some env\<close>)
done
obtain v' where "\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env \<turnstile>\<^sub>v rhs \<down> v'" "v' \<approx>\<^sub>e v"
proof (rule veval'_agree_eq)
show "fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env \<turnstile>\<^sub>v rhs \<down> v" by fact
next
have *: "Name n |\<notin>| ids rhs" if "Name n |\<notin>| fmdom env"
proof
assume "Name n |\<in>| ids rhs"
thus False
unfolding ids_def
proof (cases rule: funion_strictE)
case A
moreover have "Name n |\<notin>| frees pat"
using that unfolding \<open>fmdom env = frees pat\<close> .
ultimately have "Name n |\<in>| frees (Sabs cs)"
apply auto
unfolding ffUnion_alt_def
apply simp
apply (rule fBexI[where x = "(pat, rhs)"])
apply (auto simp: fset_of_list_elem)
apply (rule \<open>(pat, rhs) \<in> set cs\<close>)
done
thus ?thesis
using \<open>Name n |\<notin>| ids (Sabs cs)\<close> unfolding ids_def
by blast
next
case B
hence "Name n |\<in>| consts (Sabs cs)"
apply auto
unfolding ffUnion_alt_def
apply simp
apply (rule fBexI[where x = "(pat, rhs)"])
apply (auto simp: fset_of_list_elem)
apply (rule \<open>(pat, rhs) \<in> set cs\<close>)
done
thus ?thesis
using \<open>Name n |\<notin>| ids (Sabs cs)\<close> unfolding ids_def
by blast
qed
qed
show "fmrel_on_fset (ids rhs) erelated (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda> ++\<^sub>f env) (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env)"
apply rule
apply auto
apply (rule option.rel_refl; rule erelated_refl)
using * apply auto[]
apply (rule option.rel_refl; rule erelated_refl)+
using * apply auto[]
apply (rule option.rel_refl; rule erelated_refl)+
done
next
show "closed_venv (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env)"
apply rule
apply fact
apply (rule vclosed.vmatch_env)
apply fact
apply fact
done
next
show "wellformed_venv (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env)"
apply rule
apply fact
apply (rule vwellformed.vmatch_env)
apply fact
apply fact
done
next
show "closed_except rhs (fmdom (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env))"
using \<open>fmdom env = frees pat\<close> \<open>(pat, rhs) \<in> set cs\<close>
using \<open>closed_except (Sabs cs) (fmdom (fmupd (Name n) v\<^sub>2 \<Gamma>\<^sub>\<Lambda>))\<close>
apply (auto simp: Sterm.closed_except_simps list_all_iff)
apply (erule ballE[where x = "(pat, rhs)"])
apply (auto simp: closed_except_def)
done
next
show "wellformed rhs" by fact
next
show "consts rhs |\<subseteq>| fmdom (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env) |\<union>| C"
using \<open>consts (Sabs cs) |\<subseteq>| _\<close> \<open>(pat, rhs) \<in> set cs\<close>
unfolding sconsts_sabs mk_rec_env_def
by (auto simp: list_all_iff)
next
have "fdisjnt C (fmdom env)"
using \<open>(pat, rhs) \<in> set cs\<close> \<open>\<not> shadows_consts (Sabs cs)\<close>
unfolding \<open>fmdom env = frees pat\<close>
by (auto simp: list_ex_iff all_consts_def fdisjnt_alt_def)
moreover have "fdisjnt C (fmdom css)"
using \<open>vwelldefined' v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close>
by simp
ultimately show "fdisjnt C (fmdom (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env))"
using \<open>Name n |\<notin>| all_consts\<close> \<open>fdisjnt C (fmdom \<Gamma>\<^sub>\<Lambda>)\<close>
unfolding fdisjnt_alt_def mk_rec_env_def
by (auto simp: all_consts_def)
next
show "\<not> shadows_consts rhs"
using \<open>(pat, rhs) \<in> set cs\<close> \<open>\<not> shadows_consts (Sabs cs)\<close>
by (auto simp: list_ex_iff)
next
have "not_shadows_vconsts_env env"
by (rule not_shadows_vconsts.vmatch_env) fact+
thus "not_shadows_vconsts_env (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env)"
using \<open>not_shadows_vconsts_env (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>))\<close> by blast
next
have "fmpred (\<lambda>_. vwelldefined') env"
by (rule vmatch_welldefined') fact+
thus "fmpred (\<lambda>_. vwelldefined') (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>) ++\<^sub>f env)"
using \<open>fmpred (\<lambda>_. vwelldefined') (fmupd (Name n) v\<^sub>2 (\<Gamma>\<^sub>\<Lambda> ++\<^sub>f mk_rec_env css \<Gamma>\<^sub>\<Lambda>))\<close> by blast
qed simp
show ?thesis
apply (intro exI conjI)
unfolding \<open>t = _\<close>
apply (rule veval'.rec_comb)
using \<open>\<Gamma> \<turnstile>\<^sub>v t\<^sub>1 \<down> v\<^sub>1\<close> unfolding \<open>v\<^sub>1 = _\<close> apply blast
apply fact+
apply (rule related_v_ext)
apply fact+
done
qed
qed
next
case (mat1 env ml_scr ml_scr' ml_cs ml_rhs ml_pat env' ml_res)
obtain scr cs
where "t = Sabs cs $\<^sub>s scr" "related_exp scr ml_scr"
and "list_all2 (rel_prod related_pat related_exp) cs ml_cs"
using \<open>related_exp t (Mat ml_scr ml_cs)\<close>
by cases
have "sem_env.c env = as_static_cenv"
using \<open>is_cupcake_all_env env\<close>
by (auto elim: is_cupcake_all_envE)
obtain scr' where "\<Gamma> \<turnstile>\<^sub>v scr \<down> scr'" "related_v scr' ml_scr'"
using mat1(4) unfolding if_rval.simps
proof
show
"is_cupcake_all_env env" "related_exp scr ml_scr" "wellformed_venv \<Gamma>"
"closed_venv \<Gamma>" "fmpred (\<lambda>_. vwelldefined') \<Gamma>" "fdisjnt C (fmdom \<Gamma>)"
"not_shadows_vconsts_env \<Gamma>"
by fact+
next
show "fmrel_on_fset (ids scr) related_v \<Gamma> (fmap_of_ns (sem_env.v env))"
apply (rule fmrel_on_fsubset)
apply fact
unfolding \<open>t = _\<close> ids_def
apply auto
done
next
show
"wellformed scr" "consts scr |\<subseteq>| fmdom \<Gamma> |\<union>| C"
"closed_except scr (fmdom \<Gamma>)" "\<not> shadows_consts scr"
using mat1 unfolding \<open>t = _\<close> by (auto simp: closed_except_def)
qed blast
have "list_all (\<lambda>(pat, _). linear pat) cs"
using mat1 unfolding \<open>t = _\<close> by (auto simp: list_all_iff)
obtain rhs pat \<Gamma>'
where "ml_pat = mk_ml_pat (mk_pat pat)" "related_exp rhs ml_rhs"
and "vfind_match cs scr' = Some (\<Gamma>', pat, rhs)"
and "var_env \<Gamma>' env'"
using \<open>list_all2 _ cs ml_cs\<close> \<open>list_all _ cs\<close> \<open>related_v scr' ml_scr'\<close> mat1(2)
unfolding \<open>sem_env.c env = as_static_cenv\<close>
by (auto elim: match_all_related)
hence "vmatch (mk_pat pat) scr' = Some \<Gamma>'"
by (auto dest: vfind_match_elem)
hence "patvars (mk_pat pat) = fmdom \<Gamma>'"
by (auto simp: vmatch_dom)
have "(pat, rhs) \<in> set cs"
by (rule vfind_match_elem) fact
have "linear pat"
using \<open>(pat, rhs) \<in> set cs\<close> \<open>wellformed t\<close> unfolding \<open>t = _\<close>
by (auto simp: list_all_iff)
hence "fmdom \<Gamma>' = frees pat"
using \<open>patvars (mk_pat pat) = fmdom \<Gamma>'\<close>
by (simp add: mk_pat_frees)
show ?case
proof (rule if_rvalI)
fix ml_rhs'
assume "ml_res = Rval ml_rhs'"
obtain rhs' where "\<Gamma> ++\<^sub>f \<Gamma>' \<turnstile>\<^sub>v rhs \<down> rhs'" "related_v rhs' ml_rhs'"
using mat1(5) unfolding \<open>ml_res = _\<close> if_rval.simps
proof
show "is_cupcake_all_env (env \<lparr> sem_env.v := nsAppend (alist_to_ns env') (sem_env.v env) \<rparr>)"
proof (rule cupcake_v_update_preserve)
have "list_all (is_cupcake_value \<circ> snd) env'"
proof (rule match_all_preserve)
show "cupcake_c_ns (sem_env.c env)"
unfolding \<open>sem_env.c env = _\<close> by (rule static_cenv)
next
have "is_cupcake_exp (Mat ml_scr ml_cs)"
apply (rule related_exp_is_cupcake)
using mat1 by auto
thus "cupcake_clauses ml_cs"
by simp
show "is_cupcake_value ml_scr'"
apply (rule cupcake_single_preserve)
apply (rule mat1)
apply (rule mat1)
using \<open>is_cupcake_exp (Mat ml_scr ml_cs)\<close> by simp
qed fact+
hence "is_cupcake_ns (alist_to_ns env')"
by (rule cupcake_alist_to_ns_preserve)
moreover have "is_cupcake_ns (sem_env.v env)"
by (rule is_cupcake_all_envD) fact
ultimately show "is_cupcake_ns (nsAppend (alist_to_ns env') (sem_env.v env))"
by (rule cupcake_nsAppend_preserve)
qed fact
next
show "related_exp rhs ml_rhs"
by fact
next
have *: "fmdom (fmap_of_list (map (map_prod Name id) env')) = fmdom \<Gamma>'"
using \<open>var_env \<Gamma>' env'\<close>
by (metis fmrel_fmdom_eq)
have **: "id |\<in>| ids t" if "id |\<in>| ids rhs" "id |\<notin>| fmdom \<Gamma>'" for id
using \<open>id |\<in>| ids rhs\<close> unfolding ids_def
proof (cases rule: funion_strictE)
case A
from that have "id |\<notin>| frees pat"
unfolding \<open>fmdom \<Gamma>' = frees pat\<close> by simp
hence "id |\<in>| frees t"
using \<open>(pat, rhs) \<in> set cs\<close>
unfolding \<open>t = _\<close>
apply auto
apply (subst ffUnion_alt_def)
apply simp
apply (rule fBexI[where x = "(pat, rhs)"])
using A apply (auto simp: fset_of_list_elem)
done
thus "id |\<in>| frees t |\<union>| consts t" by simp
next
case B
hence "id |\<in>| consts t"
using \<open>(pat, rhs) \<in> set cs\<close>
unfolding \<open>t = _\<close>
apply auto
apply (subst ffUnion_alt_def)
apply simp
apply (rule fBexI[where x = "(pat, rhs)"])
apply (auto simp: fset_of_list_elem)
done
thus "id |\<in>| frees t |\<union>| consts t" by simp
qed
have "fmrel_on_fset (ids rhs) related_v (\<Gamma> ++\<^sub>f \<Gamma>') (fmap_of_ns (sem_env.v env) ++\<^sub>f fmap_of_list (map (map_prod Name id) env'))"
apply rule
apply simp
apply safe
subgoal
apply (rule fmrelD)
apply (rule \<open>var_env \<Gamma>' env'\<close>)
done
- subgoal using * by simp
+ subgoal
+ unfolding *[symmetric]
+ using fmdom_of_list fset_of_list_map fset.set_map image_image fst_map_prod
+ by simp
subgoal using *
by (metis (no_types, opaque_lifting) comp_def fimageI fmdom_fmap_of_list fset_of_list_map fst_comp_map_prod)
subgoal using **
by (metis fmlookup_ns fmrel_on_fsetD mat1.prems(2))
done
thus "fmrel_on_fset (ids rhs) related_v (\<Gamma> ++\<^sub>f \<Gamma>') (fmap_of_ns (sem_env.v (env \<lparr> sem_env.v := nsAppend (alist_to_ns env') (sem_env.v env) \<rparr>)))"
by (auto split: sem_env.splits)
next
show "wellformed_venv (\<Gamma> ++\<^sub>f \<Gamma>')"
apply rule
apply fact
apply (rule vwellformed.vmatch_env)
apply fact
apply (rule veval'_wellformed)
apply fact
using mat1 unfolding \<open>t = _\<close> by auto
next
show "closed_venv (\<Gamma> ++\<^sub>f \<Gamma>')"
apply rule
apply fact
apply (rule vclosed.vmatch_env)
apply fact
apply (rule veval'_closed)
apply fact
using mat1 unfolding \<open>t = _\<close> by (auto simp: closed_except_def)
next
show "fmpred (\<lambda>_. vwelldefined') (\<Gamma> ++\<^sub>f \<Gamma>')"
apply rule
apply fact
apply (rule vmatch_welldefined')
apply fact
apply (rule veval'_welldefined')
apply fact
using mat1 unfolding \<open>t = _\<close> by auto
next
show "not_shadows_vconsts_env (\<Gamma> ++\<^sub>f \<Gamma>')"
apply rule
apply fact
apply (rule not_shadows_vconsts.vmatch_env)
apply fact
apply (rule veval'_shadows)
apply fact
using mat1 unfolding \<open>t = _\<close> by auto
next
show "wellformed rhs"
using \<open>(pat, rhs) \<in> set cs\<close> \<open>wellformed t\<close> unfolding \<open>t = _\<close>
by (auto simp: list_all_iff)
next
show "closed_except rhs (fmdom (\<Gamma> ++\<^sub>f \<Gamma>'))"
apply simp
unfolding \<open>fmdom \<Gamma>' = frees pat\<close>
using \<open>(pat, rhs) \<in> set cs\<close> \<open>closed_except t (fmdom \<Gamma>)\<close> unfolding \<open>t = _\<close>
by (auto simp: Sterm.closed_except_simps list_all_iff)
next
have "consts (Sabs cs) |\<subseteq>| fmdom \<Gamma> |\<union>| C"
using mat1 unfolding \<open>t = _\<close> by auto
show "consts rhs |\<subseteq>| fmdom (\<Gamma> ++\<^sub>f \<Gamma>') |\<union>| C"
apply simp
unfolding \<open>fmdom \<Gamma>' = frees pat\<close>
using \<open>(pat, rhs) \<in> set cs\<close> \<open>consts (Sabs cs) |\<subseteq>| _\<close>
unfolding sconsts_sabs
by (auto simp: list_all_iff)
next
have "fdisjnt C (fmdom \<Gamma>')"
unfolding \<open>fmdom \<Gamma>' = frees pat\<close>
using \<open>\<not> shadows_consts t\<close> \<open>(pat, rhs) \<in> set cs\<close>
unfolding \<open>t = _\<close>
by (auto simp: list_ex_iff fdisjnt_alt_def all_consts_def)
thus "fdisjnt C (fmdom (\<Gamma> ++\<^sub>f \<Gamma>'))"
using \<open>fdisjnt C (fmdom \<Gamma>)\<close>
unfolding fdisjnt_alt_def by auto
next
show "\<not> shadows_consts rhs"
using \<open>(pat, rhs) \<in> set cs\<close> \<open>\<not> shadows_consts t\<close> unfolding \<open>t = _\<close>
by (auto simp: list_ex_iff)
qed blast
show "\<exists>t'. \<Gamma> \<turnstile>\<^sub>v t \<down> t' \<and> related_v t' ml_rhs'"
unfolding \<open>t = _\<close>
apply (intro exI conjI seval.intros)
apply (rule veval'.intros)
apply (rule veval'.intros)
apply fact+
done
qed
qed auto
theorem semantic_correctness:
fixes exp
assumes "cupcake_evaluate_single env exp (Rval ml_v)" "is_cupcake_all_env env"
assumes "fmrel_on_fset (ids t) related_v \<Gamma> (fmap_of_ns (sem_env.v env))"
assumes "related_exp t exp"
assumes "wellformed t" "wellformed_venv \<Gamma>"
assumes "closed_venv \<Gamma>" "closed_except t (fmdom \<Gamma>)"
assumes "fmpred (\<lambda>_. vwelldefined') \<Gamma>" "consts t |\<subseteq>| fmdom \<Gamma> |\<union>| C"
assumes "fdisjnt C (fmdom \<Gamma>)"
assumes "\<not> shadows_consts t" "not_shadows_vconsts_env \<Gamma>"
obtains v where "\<Gamma> \<turnstile>\<^sub>v t \<down> v" "related_v v ml_v"
using semantic_correctness0[OF assms]
by auto
end end
end
\ No newline at end of file
diff --git a/thys/CakeML_Codegen/Compiler/Compiler.thy b/thys/CakeML_Codegen/Compiler/Compiler.thy
--- a/thys/CakeML_Codegen/Compiler/Compiler.thy
+++ b/thys/CakeML_Codegen/Compiler/Compiler.thy
@@ -1,102 +1,102 @@
section \<open>Executable compilation chain\<close>
theory Compiler
imports Composition
begin
definition term_to_exp :: "C_info \<Rightarrow> rule fset \<Rightarrow> term \<Rightarrow> exp" where
"term_to_exp C_info rs t =
cakeml.mk_con C_info (heads_of rs |\<union>| constructors.C C_info)
(pterm_to_sterm (nterm_to_pterm (fresh_frun (term_to_nterm [] t) (heads_of rs |\<union>| constructors.C C_info))))"
lemma (in rules) "Compiler.term_to_exp C_info rs = term_to_cake"
unfolding term_to_exp_def by (simp add: all_consts_def)
primrec compress_pterm :: "pterm \<Rightarrow> pterm" where
"compress_pterm (Pabs cs) = Pabs (fcompress (map_prod id compress_pterm |`| cs))" |
"compress_pterm (Pconst name) = Pconst name" |
"compress_pterm (Pvar name) = Pvar name" |
"compress_pterm (t $\<^sub>p u) = compress_pterm t $\<^sub>p compress_pterm u"
lemma compress_pterm_eq[simp]: "compress_pterm t = t"
-by (induction t) (auto simp: subst_pabs_id fset_map_snd_id map_prod_def fmember_iff_member_fset)
+by (induction t) (auto simp: subst_pabs_id fset_map_snd_id map_prod_def)
definition compress_crule_set :: "crule_set \<Rightarrow> crule_set" where
"compress_crule_set = fcompress \<circ> fimage (map_prod id fcompress)"
definition compress_irule_set :: "irule_set \<Rightarrow> irule_set" where
"compress_irule_set = fcompress \<circ> fimage (map_prod id (fcompress \<circ> fimage (map_prod id compress_pterm)))"
definition compress_prule_set :: "prule fset \<Rightarrow> prule fset" where
"compress_prule_set = fcompress \<circ> fimage (map_prod id compress_pterm)"
lemma compress_crule_set_eq[simp]: "compress_crule_set rs = rs"
unfolding compress_crule_set_def by force
lemma compress_irule_set_eq[simp]: "compress_irule_set rs = rs"
unfolding compress_irule_set_def map_prod_def by simp
lemma compress_prule_set[simp]: "compress_prule_set rs = rs"
unfolding compress_prule_set_def by force
definition transform_irule_set_iter :: "irule_set \<Rightarrow> irule_set" where
"transform_irule_set_iter rs = ((transform_irule_set \<circ> compress_irule_set) ^^ max_arity rs) rs"
definition as_sem_env :: "C_info \<Rightarrow> srule list \<Rightarrow> v sem_env \<Rightarrow> v sem_env" where
"as_sem_env C_info rs env =
\<lparr> sem_env.v =
build_rec_env (cakeml.mk_letrec_body C_info (fset_of_list (map fst rs) |\<union>| constructors.C C_info) rs) env nsEmpty,
sem_env.c =
nsEmpty \<rparr>"
definition empty_sem_env :: "C_info \<Rightarrow> v sem_env" where
"empty_sem_env C_info = \<lparr> sem_env.v = nsEmpty, sem_env.c = constructors.as_static_cenv C_info \<rparr>"
definition sem_env :: "C_info \<Rightarrow> srule list \<Rightarrow> v sem_env" where
"sem_env C_info rs = extend_dec_env (as_sem_env C_info rs (empty_sem_env C_info)) (empty_sem_env C_info)"
definition compile :: "C_info \<Rightarrow> rule fset \<Rightarrow> Ast.prog" where
"compile C_info =
CakeML_Backend.compile' C_info \<circ>
Rewriting_Sterm.compile \<circ>
compress_prule_set \<circ>
Rewriting_Pterm.compile \<circ>
transform_irule_set_iter \<circ>
compress_irule_set \<circ>
Rewriting_Pterm_Elim.compile \<circ>
compress_crule_set \<circ>
Rewriting_Nterm.consts_of \<circ>
fcompress \<circ>
Rewriting_Nterm.compile' C_info \<circ>
fcompress"
definition compile_to_env :: "C_info \<Rightarrow> rule fset \<Rightarrow> v sem_env" where
"compile_to_env C_info =
sem_env C_info \<circ>
Rewriting_Sterm.compile \<circ>
compress_prule_set \<circ>
Rewriting_Pterm.compile \<circ>
transform_irule_set_iter \<circ>
compress_irule_set \<circ>
Rewriting_Pterm_Elim.compile \<circ>
compress_crule_set \<circ>
Rewriting_Nterm.consts_of \<circ>
fcompress \<circ>
Rewriting_Nterm.compile' C_info \<circ>
fcompress"
lemma (in rules) "Compiler.compile_to_env C_info rs = rules.cake_sem_env C_info rs"
unfolding Compiler.compile_to_env_def Compiler.sem_env_def Compiler.as_sem_env_def Compiler.empty_sem_env_def
unfolding rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.sem_env_def
unfolding rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_sem_env_def
unfolding empty_sem_env_def
by (auto simp:
Compiler.compress_irule_set_eq[abs_def]
Composition.transform_irule_set_iter_def[abs_def]
Compiler.transform_irule_set_iter_def[abs_def] comp_def pre_constants.all_consts_def)
export_code
term_to_exp compile compile_to_env
checking Scala
end
\ No newline at end of file
diff --git a/thys/CakeML_Codegen/Compiler/Composition.thy b/thys/CakeML_Codegen/Compiler/Composition.thy
--- a/thys/CakeML_Codegen/Compiler/Composition.thy
+++ b/thys/CakeML_Codegen/Compiler/Composition.thy
@@ -1,1081 +1,1081 @@
section \<open>Composition of correctness results\<close>
theory Composition
imports "../Backend/CakeML_Correctness"
begin
hide_const (open) sem_env.v
text \<open>@{typ term} \<open>\<longrightarrow>\<close> @{typ nterm} \<open>\<longrightarrow>\<close> @{typ pterm} \<open>\<longrightarrow>\<close> @{typ sterm}\<close>
subsection \<open>Reflexive-transitive closure of @{thm [source=true] irules.compile_correct}.\<close>
lemma (in prules) prewrite_closed:
assumes "rs \<turnstile>\<^sub>p t \<longrightarrow> t'" "closed t"
shows "closed t'"
using assms proof induction
case (step name rhs)
thus ?case
using all_rules by force
next
case (beta c)
obtain pat rhs where "c = (pat, rhs)" by (cases c) auto
with beta have "closed_except rhs (frees pat)"
by (auto simp: closed_except_simps)
show ?case
apply (rule rewrite_step_closed[OF _ beta(2)[unfolded \<open>c = _\<close>]])
using \<open>closed_except rhs (frees pat)\<close> beta by (auto simp: closed_except_def)
qed (auto simp: closed_except_def)
corollary (in prules) prewrite_rt_closed:
assumes "rs \<turnstile>\<^sub>p t \<longrightarrow>* t'" "closed t"
shows "closed t'"
using assms
by induction (auto intro: prewrite_closed)
corollary (in irules) compile_correct_rt:
assumes "Rewriting_Pterm.compile rs \<turnstile>\<^sub>p t \<longrightarrow>* t'" "finished rs"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'"
using assms proof (induction rule: rtranclp_induct)
case step
thus ?case
by (meson compile_correct rtranclp.simps)
qed auto
subsection \<open>Reflexive-transitive closure of @{thm [source=true] prules.compile_correct}.\<close>
lemma (in prules) compile_correct_rt:
assumes "Rewriting_Sterm.compile rs \<turnstile>\<^sub>s u \<longrightarrow>* u'"
shows "rs \<turnstile>\<^sub>p sterm_to_pterm u \<longrightarrow>* sterm_to_pterm u'"
using assms proof induction
case step
thus ?case
by (meson compile_correct rtranclp.simps)
qed auto
lemma srewrite_stepD:
assumes "srewrite_step rs name t"
shows "(name, t) \<in> set rs"
using assms by induct auto
lemma (in srules) srewrite_wellformed:
assumes "rs \<turnstile>\<^sub>s t \<longrightarrow> t'" "wellformed t"
shows "wellformed t'"
using assms proof induction
case (step name rhs)
hence "(name, rhs) \<in> set rs"
by (auto dest: srewrite_stepD)
thus ?case
using all_rules by (auto simp: list_all_iff)
next
case (beta cs t t')
then obtain pat rhs env where "(pat, rhs) \<in> set cs" "match pat t = Some env" "t' = subst rhs env"
by (elim rewrite_firstE)
show ?case
unfolding \<open>t' = _\<close>
proof (rule subst_wellformed)
show "wellformed rhs"
using \<open>(pat, rhs) \<in> set cs\<close> beta by (auto simp: list_all_iff)
next
show "wellformed_env env"
using \<open>match pat t = Some env\<close> beta
by (auto intro: wellformed.match)
qed
qed auto
lemma (in srules) srewrite_wellformed_rt:
assumes "rs \<turnstile>\<^sub>s t \<longrightarrow>* t'" "wellformed t"
shows "wellformed t'"
using assms
by induction (auto intro: srewrite_wellformed)
lemma vno_abs_value_to_sterm: "no_abs (value_to_sterm v) \<longleftrightarrow> vno_abs v" for v
by (induction v) (auto simp: no_abs.list_comb list_all_iff)
subsection \<open>Reflexive-transitive closure of @{thm [source=true] rules.compile_correct}.\<close>
corollary (in rules) compile_correct_rt:
assumes "compile \<turnstile>\<^sub>n u \<longrightarrow>* u'" "closed u"
shows "rs \<turnstile> nterm_to_term' u \<longrightarrow>* nterm_to_term' u'"
using assms
proof induction
case (step u' u'')
hence "rs \<turnstile> nterm_to_term' u \<longrightarrow>* nterm_to_term' u'"
by auto
also have "rs \<turnstile> nterm_to_term' u' \<longrightarrow> nterm_to_term' u''"
using step by (auto dest: rewrite_rt_closed intro!: compile_correct simp: closed_except_def)
finally show ?case .
qed auto
subsection \<open>Reflexive-transitive closure of @{thm [source=true] irules.transform_correct}.\<close>
corollary (in irules) transform_correct_rt:
assumes "transform_irule_set rs \<turnstile>\<^sub>i u \<longrightarrow>* u''" "t \<approx>\<^sub>p u" "closed t"
obtains t'' where "rs \<turnstile>\<^sub>i t \<longrightarrow>* t''" "t'' \<approx>\<^sub>p u''"
using assms proof (induction arbitrary: thesis t)
case (step u' u'')
obtain t' where "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'" "t' \<approx>\<^sub>p u'"
using step by blast
obtain t'' where "rs \<turnstile>\<^sub>i t' \<longrightarrow>* t''" "t'' \<approx>\<^sub>p u''"
apply (rule transform_correct)
apply (rule \<open>transform_irule_set rs \<turnstile>\<^sub>i u' \<longrightarrow> u''\<close>)
apply (rule \<open>t' \<approx>\<^sub>p u'\<close>)
apply (rule irewrite_rt_closed)
apply (rule \<open>rs \<turnstile>\<^sub>i t \<longrightarrow>* t'\<close>)
apply (rule \<open>closed t\<close>)
apply blast
done
show ?case
apply (rule step.prems)
apply (rule rtranclp_trans)
apply fact+
done
qed blast
corollary (in irules) transform_correct_rt_no_abs:
assumes "transform_irule_set rs \<turnstile>\<^sub>i t \<longrightarrow>* u" "closed t" "no_abs u"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* u"
proof -
have "t \<approx>\<^sub>p t" by (rule prelated_refl)
obtain t' where "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'" "t' \<approx>\<^sub>p u"
apply (rule transform_correct_rt)
apply (rule assms)
apply (rule \<open>t \<approx>\<^sub>p t\<close>)
apply (rule assms)
apply blast
done
thus ?thesis
using assms by (metis prelated_no_abs_right)
qed
corollary transform_correct_rt_n_no_abs0:
assumes "irules C rs" "(transform_irule_set ^^ n) rs \<turnstile>\<^sub>i t \<longrightarrow>* u" "closed t" "no_abs u"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* u"
using assms(1,2) proof (induction n arbitrary: rs)
case (Suc n)
interpret irules C rs by fact
show ?case
apply (rule transform_correct_rt_no_abs)
apply (rule Suc.IH)
apply (rule rules_transform)
using Suc(3) apply (simp add: funpow_swap1)
apply fact+
done
qed auto
corollary (in irules) transform_correct_rt_n_no_abs:
assumes "(transform_irule_set ^^ n) rs \<turnstile>\<^sub>i t \<longrightarrow>* u" "closed t" "no_abs u"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* u"
by (rule transform_correct_rt_n_no_abs0) (rule irules_axioms assms)+
hide_fact transform_correct_rt_n_no_abs0
subsection \<open>Iterated application of @{const transform_irule_set}.\<close>
definition max_arity :: "irule_set \<Rightarrow> nat" where
"max_arity rs = fMax ((arity \<circ> snd) |`| rs)"
lemma rules_transform_iter0:
assumes "irules C_info rs"
shows "irules C_info ((transform_irule_set ^^ n) rs)"
using assms
by (induction n) (auto intro: irules.rules_transform del: irulesI)
lemma (in irules) rules_transform_iter: "irules C_info ((transform_irule_set ^^ n) rs)"
by (rule rules_transform_iter0) (rule irules_axioms)
lemma transform_irule_set_n_heads: "fst |`| ((transform_irule_set ^^ n) rs) = fst |`| rs"
by (induction n) (auto simp: transform_irule_set_heads)
hide_fact rules_transform_iter0
definition transform_irule_set_iter :: "irule_set \<Rightarrow> irule_set" where
"transform_irule_set_iter rs = (transform_irule_set ^^ max_arity rs) rs"
lemma transform_irule_set_iter_heads: "fst |`| transform_irule_set_iter rs = fst |`| rs"
unfolding transform_irule_set_iter_def by (simp add: transform_irule_set_n_heads)
lemma (in irules) finished_alt_def: "finished rs \<longleftrightarrow> max_arity rs = 0"
proof
assume "max_arity rs = 0"
hence "\<not> fBex ((arity \<circ> snd) |`| rs) (\<lambda>x. 0 < x)"
using nonempty
unfolding max_arity_def
by (metis fBex_fempty fmax_ex_gr not_less0)
thus "finished rs"
unfolding finished_def
by force
next
assume "finished rs"
have "fMax ((arity \<circ> snd) |`| rs) \<le> 0"
proof (rule fMax_le)
show "fBall ((arity \<circ> snd) |`| rs) (\<lambda>x. x \<le> 0)"
using \<open>finished rs\<close> unfolding finished_def by force
next
show "(arity \<circ> snd) |`| rs \<noteq> {||}"
using nonempty by force
qed
thus "max_arity rs = 0"
unfolding max_arity_def by simp
qed
lemma (in irules) transform_finished_id: "finished rs \<Longrightarrow> transform_irule_set rs = rs"
unfolding transform_irule_set_def finished_def transform_irules_def map_prod_def id_apply
-by (rule fset_map_snd_id) (auto simp: fmember_iff_member_fset elim!: fBallE)
+by (rule fset_map_snd_id) (auto elim!: fBallE)
lemma (in irules) max_arity_decr: "max_arity (transform_irule_set rs) = max_arity rs - 1"
proof (cases "finished rs")
case True
thus ?thesis
by (auto simp: transform_finished_id finished_alt_def)
next
case False
have "(arity \<circ> snd) |`| transform_irule_set rs = (\<lambda>x. x - 1) |`| (arity \<circ> snd) |`| rs"
unfolding transform_irule_set_def fset.map_comp
proof (rule fset.map_cong0, safe, unfold o_apply map_prod_simp id_apply snd_conv)
fix name irs
- assume "(name, irs) \<in> fset rs"
- hence "(name, irs) |\<in>| rs"
- by (simp add: fmember_iff_member_fset)
+ assume "(name, irs) |\<in>| rs"
hence "arity_compatibles irs" "irs \<noteq> {||}"
- using nonempty inner by (blast dest: fpairwiseD)+
+ using nonempty inner
+ unfolding atomize_conj
+ by (smt (verit, del_insts) \<open>(name, irs) |\<in>| rs\<close> case_prodD fBallE inner)
thus "arity (transform_irules irs) = arity irs - 1"
by (simp add: arity_transform_irules)
qed
hence "max_arity (transform_irule_set rs) = fMax ((\<lambda>x. x - 1) |`| (arity \<circ> snd) |`| rs)"
unfolding max_arity_def by simp
also have "\<dots> = fMax ((arity \<circ> snd) |`| rs) - 1"
proof (rule fmax_decr)
show "fBex ((arity \<circ> snd) |`| rs) ((\<le>) 1)"
using False unfolding finished_def by force
qed
finally show ?thesis
unfolding max_arity_def
by simp
qed
lemma max_arity_decr'0:
assumes "irules C rs"
shows "max_arity ((transform_irule_set ^^ n) rs) = max_arity rs - n"
proof (induction n)
case (Suc n)
show ?case
apply simp
apply (subst irules.max_arity_decr)
using Suc assms by (auto intro: irules.rules_transform_iter del: irulesI)
qed auto
lemma (in irules) max_arity_decr': "max_arity ((transform_irule_set ^^ n) rs) = max_arity rs - n"
by (rule max_arity_decr'0) (rule irules_axioms)
hide_fact max_arity_decr'0
lemma (in irules) transform_finished: "finished (transform_irule_set_iter rs)"
unfolding transform_irule_set_iter_def
by (subst irules.finished_alt_def)
(auto simp: max_arity_decr' intro: rules_transform_iter del: Rewriting_Pterm_Elim.irulesI)
text \<open>Trick as described in \<open>\<section>7.1\<close> in the locale manual.\<close>
locale irules' = irules
sublocale irules' \<subseteq> irules'_as_irules: irules C_info "transform_irule_set_iter rs"
unfolding transform_irule_set_iter_def by (rule rules_transform_iter)
sublocale crules \<subseteq> crules_as_irules': irules' C_info "Rewriting_Pterm_Elim.compile rs"
unfolding irules'_def by (fact compile_rules)
sublocale irules' \<subseteq> irules'_as_prules: prules C_info "Rewriting_Pterm.compile (transform_irule_set_iter rs)"
by (rule irules'_as_irules.compile_rules) (rule transform_finished)
subsection \<open>Big-step semantics\<close>
context srules begin
definition global_css :: "(name, sclauses) fmap" where
"global_css = fmap_of_list (map (map_prod id clauses) rs)"
lemma fmdom_global_css: "fmdom global_css = fst |`| fset_of_list rs"
unfolding global_css_def by simp
definition as_vrules :: "vrule list" where
"as_vrules = map (\<lambda>(name, _). (name, Vrecabs global_css name fmempty)) rs"
lemma as_vrules_fst[simp]: "fst |`| fset_of_list as_vrules = fst |`| fset_of_list rs"
unfolding as_vrules_def
apply simp
apply (rule fset.map_cong)
apply (rule refl)
by auto
lemma as_vrules_fst'[simp]: "map fst as_vrules = map fst rs"
unfolding as_vrules_def
by auto
lemma list_all_as_vrulesI:
assumes "list_all (\<lambda>(_, t). P fmempty (clauses t)) rs"
assumes "R (fst |`| fset_of_list rs)"
shows "list_all (\<lambda>(_, t). value_pred.pred P Q R t) as_vrules"
proof (rule list_allI, safe)
fix name rhs
assume "(name, rhs) \<in> set as_vrules"
hence "rhs = Vrecabs global_css name fmempty"
unfolding as_vrules_def by auto
moreover have "fmpred (\<lambda>_. P fmempty) global_css"
unfolding global_css_def list.pred_map
using assms by (auto simp: list_all_iff intro!: fmpred_of_list)
moreover have "name |\<in>| fmdom global_css"
unfolding global_css_def
apply auto
using \<open>(name, rhs) \<in> set as_vrules\<close> unfolding as_vrules_def
including fset.lifting apply transfer'
by force
moreover have "R (fmdom global_css)"
using assms unfolding global_css_def
by auto
ultimately show "value_pred.pred P Q R rhs"
by (simp add: value_pred.pred_alt_def)
qed
lemma srules_as_vrules: "vrules C_info as_vrules"
proof (standard, unfold as_vrules_fst)
have "list_all (\<lambda>(_, t). vwellformed t) as_vrules"
unfolding vwellformed_def
apply (rule list_all_as_vrulesI)
apply (rule list.pred_mono_strong)
apply (rule all_rules)
apply (auto elim: clausesE)
done
moreover have "list_all (\<lambda>(_, t). vclosed t) as_vrules"
unfolding vclosed_def
apply (rule list_all_as_vrulesI)
apply auto
apply (rule list.pred_mono_strong)
apply (rule all_rules)
apply (auto elim: clausesE simp: Sterm.closed_except_simps)
done
moreover have "list_all (\<lambda>(_, t). \<not> is_Vconstr t) as_vrules"
unfolding as_vrules_def
by (auto simp: list_all_iff)
ultimately show "list_all vrule as_vrules"
unfolding list_all_iff by fastforce
next
show "distinct (map fst as_vrules)"
using distinct by auto
next
show "fdisjnt (fst |`| fset_of_list rs) C"
using disjnt by simp
next
show "list_all (\<lambda>(_, rhs). not_shadows_vconsts rhs) as_vrules"
unfolding not_shadows_vconsts_def
apply (rule list_all_as_vrulesI)
apply auto
apply (rule list.pred_mono_strong)
apply (rule not_shadows)
by (auto simp: list_all_iff list_ex_iff all_consts_def elim!: clausesE)
next
show "vconstructor_value_rs as_vrules"
unfolding vconstructor_value_rs_def
apply (rule conjI)
unfolding vconstructor_value_def
apply (rule list_all_as_vrulesI)
apply (simp add: list_all_iff)
apply simp
apply simp
using disjnt by simp
next
show "list_all (\<lambda>(_, rhs). vwelldefined rhs) as_vrules"
unfolding vwelldefined_def
apply (rule list_all_as_vrulesI)
apply auto
apply (rule list.pred_mono_strong)
apply (rule swelldefined_rs)
apply auto
apply (erule clausesE)
apply hypsubst_thin
apply (subst (asm) welldefined_sabs)
by simp
next
show "distinct all_constructors"
by (fact distinct_ctr)
qed
sublocale srules_as_vrules: vrules C_info as_vrules
by (fact srules_as_vrules)
lemma rs'_rs_eq: "srules_as_vrules.rs' = rs"
unfolding srules_as_vrules.rs'_def
unfolding as_vrules_def
apply (subst map_prod_def)
apply simp
unfolding comp_def
apply (subst case_prod_twice)
apply (rule list_map_snd_id)
unfolding global_css_def
using all_rules map
apply (auto simp: list_all_iff map_of_is_map map_of_map map_prod_def fmap_of_list.rep_eq)
subgoal for a b
by (erule ballE[where x = "(a, b)"], cases b, auto simp: is_abs_def term_cases_def)
done
lemma veval_correct:
fixes v
assumes "as_vrules, fmempty \<turnstile>\<^sub>v t \<down> v" "wellformed t" "closed t"
shows "rs, fmempty \<turnstile>\<^sub>s t \<down> value_to_sterm v"
using assms
by (rule srules_as_vrules.veval_correct[unfolded rs'_rs_eq])
end
subsection \<open>ML-style semantics\<close>
context srules begin
lemma as_vrules_mk_rec_env: "fmap_of_list as_vrules = mk_rec_env global_css fmempty"
apply (subst global_css_def)
apply (subst as_vrules_def)
apply (subst mk_rec_env_def)
apply (rule fmap_ext)
apply (subst fmlookup_fmmap_keys)
apply (subst fmap_of_list.rep_eq)
apply (subst fmap_of_list.rep_eq)
apply (subst map_of_map_keyed)
apply (subst (2) map_prod_def)
apply (subst id_apply)
apply (subst map_of_map)
apply simp
apply (subst option.map_comp)
apply (rule option.map_cong)
apply (rule refl)
apply simp
apply (subst global_css_def)
apply (rule refl)
done
abbreviation (input) "vrelated \<equiv> srules_as_vrules.vrelated"
notation srules_as_vrules.vrelated ("\<turnstile>\<^sub>v/ _ \<approx> _" [0, 50] 50)
lemma vrecabs_global_css_refl:
assumes "name |\<in>| fmdom global_css"
shows "\<turnstile>\<^sub>v Vrecabs global_css name fmempty \<approx> Vrecabs global_css name fmempty"
using assms
proof (coinduction arbitrary: name)
case vrelated
have "rel_option (\<lambda>v\<^sub>1 v\<^sub>2. (\<exists>name. v\<^sub>1 = Vrecabs global_css name fmempty \<and> v\<^sub>2 = Vrecabs global_css name fmempty \<and> name |\<in>| fmdom global_css) \<or> \<turnstile>\<^sub>v v\<^sub>1 \<approx> v\<^sub>2) (fmlookup (fmap_of_list as_vrules) y) (fmlookup (mk_rec_env global_css fmempty) y)" for y
apply (subst as_vrules_mk_rec_env)
apply (rule option.rel_refl_strong)
apply (rule disjI1)
apply (simp add: mk_rec_env_def)
apply (elim conjE exE)
by (auto intro: fmdomI)
with vrelated show ?case
by fastforce
qed
lemma as_vrules_refl_rs: "fmrel_on_fset (fst |`| fset_of_list as_vrules) vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
apply rule
apply (subst (2) as_vrules_def)
apply (subst (2) as_vrules_def)
apply (simp add: fmap_of_list.rep_eq)
apply (rule rel_option_reflI)
apply simp
apply (drule map_of_SomeD)
apply auto
apply (rule vrecabs_global_css_refl)
unfolding global_css_def
-by (auto simp: fset_of_list_elem intro: rev_fimage_eqI)
+by (auto simp: fset_of_list_elem intro: rev_image_eqI)
lemma as_vrules_refl_C: "fmrel_on_fset C vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
proof
fix c
assume "c |\<in>| C"
hence "c |\<notin>| fset_of_list (map fst as_vrules)"
using srules_as_vrules.vconstructor_value_rs
unfolding vconstructor_value_rs_def fdisjnt_alt_def
by auto
hence "c |\<notin>| fmdom (fmap_of_list as_vrules)"
by simp
hence "fmlookup (fmap_of_list as_vrules) c = None"
by (metis fmdom_notD)
thus "rel_option vrelated (fmlookup (fmap_of_list as_vrules) c) (fmlookup (fmap_of_list as_vrules) c)"
by simp
qed
lemma veval'_correct'':
fixes t v
assumes "fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> v"
assumes "wellformed t"
assumes "\<not> shadows_consts t"
assumes "welldefined t"
assumes "closed t"
assumes "vno_abs v"
shows "as_vrules, fmempty \<turnstile>\<^sub>v t \<down> v"
proof -
obtain v\<^sub>1 where "as_vrules, fmempty \<turnstile>\<^sub>v t \<down> v\<^sub>1" "\<turnstile>\<^sub>v v\<^sub>1 \<approx> v"
using \<open>fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> v\<close>
proof (rule srules_as_vrules.veval'_correct', unfold as_vrules_fst)
show "wellformed t" "\<not> shadows_consts t" "closed t" "consts t |\<subseteq>| all_consts"
by fact+
next
show "wellformed_venv (fmap_of_list as_vrules)"
apply rule
using srules_as_vrules.all_rules
apply (auto simp: list_all_iff)
done
next
show "not_shadows_vconsts_env (fmap_of_list as_vrules) "
apply rule
using srules_as_vrules.not_shadows
apply (auto simp: list_all_iff)
done
next
have "fmrel_on_fset (fst |`| fset_of_list as_vrules |\<union>| C) vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
apply (rule fmrel_on_fset_unionI)
apply (rule as_vrules_refl_rs)
apply (rule as_vrules_refl_C)
done
show "fmrel_on_fset (consts t) vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
apply (rule fmrel_on_fsubset)
apply fact+
using assms by (auto simp: all_consts_def)
qed
thus ?thesis
using assms by (metis srules_as_vrules.vrelated.eq_right)
qed
end
subsection \<open>CakeML\<close>
context srules begin
definition as_sem_env :: "v sem_env \<Rightarrow> v sem_env" where
"as_sem_env env = \<lparr> sem_env.v = build_rec_env (mk_letrec_body all_consts rs) env nsEmpty, sem_env.c = nsEmpty \<rparr>"
lemma compile_sem_env:
"evaluate_dec ck mn env state (compile_group all_consts rs) (state, Rval (as_sem_env env))"
unfolding compile_group_def as_sem_env_def
apply (rule evaluate_dec.dletrec1)
unfolding mk_letrec_body_def Let_def
apply (simp add:comp_def case_prod_twice)
using name_as_string.fst_distinct[OF distinct]
by auto
lemma compile_sem_env':
"fun_evaluate_decs mn state env [(compile_group all_consts rs)] = (state, Rval (as_sem_env env))"
unfolding compile_group_def as_sem_env_def mk_letrec_body_def Let_def
apply (simp add: comp_def case_prod_twice)
using name_as_string.fst_distinct[OF distinct]
by auto
lemma compile_prog[unfolded combine_dec_result.simps, simplified]:
"evaluate_prog ck env state (compile rs) (state, combine_dec_result (as_sem_env env) (Rval \<lparr> sem_env.v = nsEmpty, sem_env.c = nsEmpty \<rparr>))"
unfolding compile_def
apply (rule evaluate_prog.cons1)
apply rule
apply (rule evaluate_top.tdec1)
apply (rule compile_sem_env)
apply (rule evaluate_prog.empty)
done
lemma compile_prog'[unfolded combine_dec_result.simps, simplified]:
"fun_evaluate_prog state env (compile rs) = (state, combine_dec_result (as_sem_env env) (Rval \<lparr> sem_env.v = nsEmpty, sem_env.c = nsEmpty \<rparr>))"
unfolding compile_def fun_evaluate_prog_def no_dup_mods_def no_dup_top_types_def prog_to_mods_def prog_to_top_types_def decs_to_types_def
using compile_sem_env' compile_group_def by simp
definition sem_env :: "v sem_env" where
"sem_env \<equiv> extend_dec_env (as_sem_env empty_sem_env) empty_sem_env"
(* FIXME introduce lemma: is_cupcake_all_env extend_dec_env *)
(* FIXME introduce lemma: is_cupcake_all_env empty_sem_env *)
lemma cupcake_sem_env: "is_cupcake_all_env sem_env"
unfolding as_sem_env_def sem_env_def
apply (rule is_cupcake_all_envI)
apply (simp add: extend_dec_env_def empty_sem_env_def nsEmpty_def)
apply (rule cupcake_nsAppend_preserve)
apply (rule cupcake_build_rec_preserve)
apply (simp add: empty_sem_env_def)
apply (simp add: nsEmpty_def)
apply (rule mk_letrec_cupcake)
apply simp
apply (simp add: empty_sem_env_def)
done
lemma sem_env_refl: "fmrel related_v (fmap_of_list as_vrules) (fmap_of_ns (sem_env.v sem_env))"
proof
fix name
show "rel_option related_v (fmlookup (fmap_of_list as_vrules) name) (fmlookup (fmap_of_ns (sem_env.v sem_env)) name)"
apply (simp add:
as_sem_env_def build_rec_env_fmap cake_mk_rec_env_def sem_env_def
fmap_of_list.rep_eq map_of_map_keyed option.rel_map
as_vrules_def mk_letrec_body_def comp_def case_prod_twice)
apply (rule option.rel_refl_strong)
apply (rule related_v.rec_closure)
apply auto[]
apply (simp add:
fmmap_of_list[symmetric, unfolded apsnd_def map_prod_def id_def] fmap.rel_map
global_css_def Let_def map_prod_def comp_def case_prod_twice)
apply (thin_tac "map_of rs name = _")
apply (rule fmap.rel_refl_strong)
apply simp
subgoal premises prems for rhs
proof -
obtain name where "(name, rhs) \<in> set rs"
using prems
including fmap.lifting
by transfer' (auto dest: map_of_SomeD)
hence "is_abs rhs" "closed rhs" "welldefined rhs"
using all_rules swelldefined_rs by (auto simp add: list_all_iff)
then obtain cs where "clauses rhs = cs" "rhs = Sabs cs" "wellformed_clauses cs"
using \<open>(name, rhs) \<in> set rs\<close> all_rules
by (cases rhs) (auto simp: list_all_iff is_abs_def term_cases_def)
show ?thesis
unfolding related_fun_alt_def \<open>clauses rhs = cs\<close>
proof (intro conjI)
show "list_all2 (rel_prod related_pat related_exp) cs (map (\<lambda>(pat, t). (mk_ml_pat (mk_pat pat), mk_con (frees pat |\<union>| all_consts) t)) cs)"
unfolding list.rel_map
apply (rule list.rel_refl_strong)
apply (rename_tac z, case_tac z, hypsubst_thin)
apply simp
subgoal premises prems for pat t
proof (rule mk_exp_correctness)
have "\<not> shadows_consts rhs"
using \<open>(name, rhs) \<in> set rs\<close> not_shadows
by (auto simp: list_all_iff all_consts_def)
thus "\<not> shadows_consts t"
unfolding \<open>rhs = Sabs cs\<close> using prems
by (auto simp: list_all_iff list_ex_iff)
next
have "frees t |\<subseteq>| frees pat"
using \<open>closed rhs\<close> prems unfolding \<open>rhs = _\<close>
apply (auto simp: list_all_iff Sterm.closed_except_simps)
apply (erule ballE[where x = "(pat, t)"])
apply (auto simp: closed_except_def)
done
moreover have "consts t |\<subseteq>| all_consts"
using \<open>welldefined rhs\<close> prems unfolding \<open>rhs = _\<close> welldefined_sabs
by (auto simp: list_all_iff all_consts_def)
ultimately show "ids t |\<subseteq>| frees pat |\<union>| all_consts"
unfolding ids_def by auto
qed (auto simp: all_consts_def)
done
next
have 1: "frees (Sabs cs) = {||}"
using \<open>closed rhs\<close> unfolding \<open>rhs = Sabs cs\<close>
by (auto simp: closed_except_def)
have 2: "welldefined rhs"
using swelldefined_rs \<open>(name, rhs) \<in> set rs\<close>
by (auto simp: list_all_iff)
show "fresh_fNext all_consts |\<notin>| ids (Sabs cs)"
apply (rule fNext_not_member_subset)
unfolding ids_def 1
using 2 \<open>rhs = _\<close> by (simp add: all_consts_def del: consts_sterm.simps)
next
show "fresh_fNext all_consts |\<notin>| all_consts"
by (rule fNext_not_member)
qed
qed
done
qed
lemma semantic_correctness':
assumes "cupcake_evaluate_single sem_env (mk_con all_consts t) (Rval ml_v)"
assumes "welldefined t" "closed t" "\<not> shadows_consts t" "wellformed t"
obtains v where "fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> v" "related_v v ml_v"
using assms(1) proof (rule semantic_correctness)
show "is_cupcake_all_env sem_env"
by (fact cupcake_sem_env)
next
show "related_exp t (mk_con all_consts t)"
apply (rule mk_exp_correctness)
using assms
unfolding ids_def closed_except_def by (auto simp: all_consts_def)
next
show "wellformed t" "\<not> shadows_consts t" by fact+
next
show "closed_except t (fmdom (fmap_of_list as_vrules))"
using \<open>closed t\<close> by (auto simp: closed_except_def)
next
show "closed_venv (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
using srules_as_vrules.all_rules
by (auto simp: list_all_iff)
show "wellformed_venv (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
using srules_as_vrules.all_rules
by (auto simp: list_all_iff)
next
have 1: "fmpred (\<lambda>_. list_all (\<lambda>(pat, t). consts t |\<subseteq>| C |\<union>| fmdom global_css)) global_css"
apply (subst (2) global_css_def)
apply (rule fmpred_of_list)
apply (auto simp: map_prod_def)
subgoal premises prems for pat t
proof -
from prems obtain cs where "t = Sabs cs"
by (elim clausesE)
have "welldefined t"
using swelldefined_rs prems
by (auto simp: list_all_iff fmdom_global_css)
show ?thesis
using \<open>welldefined t\<close>
unfolding \<open>t = _\<close> welldefined_sabs
by (auto simp: all_consts_def list_all_iff fmdom_global_css)
qed
done
show "fmpred (\<lambda>_. vwelldefined') (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
unfolding as_vrules_def
apply simp
apply (erule imageE)
apply (auto split: prod.splits)
apply (subst fdisjnt_alt_def)
apply simp
apply (rule 1)
apply (subst global_css_def)
apply simp
subgoal for x1 x2
- apply (rule fimage_eqI[where x = "(x1, x2)"])
+ apply (rule image_eqI[where x = "(x1, x2)"])
by (auto simp: fset_of_list_elem)
subgoal
using disjnt by (auto simp: fdisjnt_alt_def fmdom_global_css)
done
next
show "not_shadows_vconsts_env (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
using srules_as_vrules.not_shadows
unfolding list_all_iff
by auto
next
show "fdisjnt C (fmdom (fmap_of_list as_vrules))"
using disjnt by (auto simp: fdisjnt_alt_def)
next
show "fmrel_on_fset (ids t) related_v (fmap_of_list as_vrules) (fmap_of_ns (sem_env.v sem_env))"
unfolding fmrel_on_fset_fmrel_restrict
apply (rule fmrel_restrict_fset)
apply (rule sem_env_refl)
done
next
show "consts t |\<subseteq>| fmdom (fmap_of_list as_vrules) |\<union>| C"
apply (subst fmdom_fmap_of_list)
apply (subst as_vrules_fst')
apply simp
using assms by (auto simp: all_consts_def)
qed blast
end
fun cake_to_value :: "v \<Rightarrow> value" where
"cake_to_value (Conv (Some (name, _)) vs) = Vconstr (Name name) (map cake_to_value vs)"
context cakeml' begin
lemma cake_to_value_abs_free:
assumes "is_cupcake_value v" "cake_no_abs v"
shows "vno_abs (cake_to_value v)"
using assms by (induction v) (auto elim: is_cupcake_value.elims simp: list_all_iff)
lemma cake_to_value_related:
assumes "cake_no_abs v" "is_cupcake_value v"
shows "related_v (cake_to_value v) v"
using assms proof (induction v)
case (Conv c vs)
then obtain name tid where "c = Some ((as_string name), TypeId (Short tid))"
apply (elim is_cupcake_value.elims)
subgoal
by (metis name.sel v.simps(2))
by auto
show ?case
unfolding \<open>c = _\<close>
apply simp
apply (rule related_v.conv)
apply (simp add: list.rel_map)
apply (rule list.rel_refl_strong)
apply (rule Conv)
using Conv unfolding \<open>c = _\<close>
by (auto simp: list_all_iff)
qed auto
lemma related_v_abs_free_uniq:
assumes "related_v v\<^sub>1 ml_v" "related_v v\<^sub>2 ml_v" "cake_no_abs ml_v"
shows "v\<^sub>1 = v\<^sub>2"
using assms proof (induction arbitrary: v\<^sub>2)
case (conv vs\<^sub>1 ml_vs name)
then obtain vs\<^sub>2 where "v\<^sub>2 = Vconstr name vs\<^sub>2" "list_all2 related_v vs\<^sub>2 ml_vs"
by (auto elim: related_v.cases simp: name.expand)
moreover have "list_all cake_no_abs ml_vs"
using conv by simp
have "list_all2 (=) vs\<^sub>1 vs\<^sub>2"
using \<open>list_all2 _ vs\<^sub>1 _\<close> \<open>list_all2 _ vs\<^sub>2 _\<close> \<open>list_all cake_no_abs ml_vs\<close>
by (induction arbitrary: vs\<^sub>2 rule: list.rel_induct) (auto simp: list_all2_Cons2)
thus ?case
unfolding \<open>v\<^sub>2 = _\<close>
by (simp add: list.rel_eq)
qed auto
corollary related_v_abs_free_cake_to_value:
assumes "related_v v ml_v" "cake_no_abs ml_v" "is_cupcake_value ml_v"
shows "v = cake_to_value ml_v"
using assms by (metis cake_to_value_related related_v_abs_free_uniq)
end
context srules begin
lemma cupcake_sem_env_preserve:
assumes "cupcake_evaluate_single sem_env (mk_con S t) (Rval ml_v)" "wellformed t"
shows "is_cupcake_value ml_v"
apply (rule cupcake_single_preserve[OF assms(1)])
apply (rule cupcake_sem_env)
apply (rule mk_exp_cupcake)
apply fact
done
lemma semantic_correctness'':
assumes "cupcake_evaluate_single sem_env (mk_con all_consts t) (Rval ml_v)"
assumes "welldefined t" "closed t" "\<not> shadows_consts t" "wellformed t"
assumes "cake_no_abs ml_v"
shows "fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> cake_to_value ml_v"
using assms
by (metis cupcake_sem_env_preserve semantic_correctness' related_v_abs_free_cake_to_value)
end
subsection \<open>Composition\<close>
context rules begin
abbreviation term_to_nterm where
"term_to_nterm t \<equiv> fresh_frun (Term_to_Nterm.term_to_nterm [] t) all_consts"
abbreviation sterm_to_cake where
"sterm_to_cake \<equiv> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.mk_con all_consts"
abbreviation "term_to_cake t \<equiv> sterm_to_cake (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
abbreviation "cake_to_term t \<equiv> (convert_term (value_to_sterm (cake_to_value t)) :: term)"
abbreviation "cake_sem_env \<equiv> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.sem_env"
definition "compiled \<equiv> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules"
lemma fmdom_compiled: "fmdom (fmap_of_list compiled) = heads_of rs"
unfolding compiled_def
by (simp add:
rules_as_nrules.crules_as_irules'.irules'_as_prules.compile_heads
Rewriting_Pterm.compile_heads transform_irule_set_iter_heads
Rewriting_Pterm_Elim.compile_heads
compile_heads consts_of_heads)
lemma cake_semantic_correctness:
assumes "cupcake_evaluate_single cake_sem_env (sterm_to_cake t) (Rval ml_v)"
assumes "welldefined t" "closed t" "\<not> shadows_consts t" "wellformed t"
assumes "cake_no_abs ml_v"
shows "fmap_of_list compiled \<turnstile>\<^sub>v t \<down> cake_to_value ml_v"
unfolding compiled_def
apply (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.semantic_correctness'')
using assms
by (simp_all add:
rules_as_nrules.crules_as_irules'.irules'_as_prules.compile_heads
Rewriting_Pterm.compile_heads transform_irule_set_iter_heads
Rewriting_Pterm_Elim.compile_heads
compile_heads consts_of_heads all_consts_def)
text \<open>Lo and behold, this is the final correctness theorem!\<close>
theorem compiled_correct:
\<comment> \<open>If CakeML evaluation of a term succeeds ...\<close>
assumes "\<exists>k. Evaluate_Single.evaluate cake_sem_env (s \<lparr> clock := k \<rparr>) (term_to_cake t) = (s', Rval ml_v)"
\<comment> \<open>... producing a constructor term without closures ...\<close>
assumes "cake_no_abs ml_v"
\<comment> \<open>... and some syntactic properties of the involved terms hold ...\<close>
assumes "closed t" "\<not> shadows_consts t" "welldefined t" "wellformed t"
\<comment> \<open>... then this evaluation can be reproduced in the term-rewriting semantics\<close>
shows "rs \<turnstile> t \<longrightarrow>* cake_to_term ml_v"
proof -
let ?heads = "fst |`| fset_of_list rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules"
have "?heads = heads_of rs"
using fmdom_compiled unfolding compiled_def by simp
have "wellformed (nterm_to_pterm (term_to_nterm t))"
by auto
hence "wellformed (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
by (auto intro: pterm_to_sterm_wellformed)
have "is_cupcake_all_env cake_sem_env"
by (rule rules_as_nrules.nrules_as_crules.crules_as_irules'.irules'_as_prules.prules_as_srules.cupcake_sem_env)
have "is_cupcake_exp (term_to_cake t)"
by (rule rules_as_nrules.nrules_as_crules.crules_as_irules'.irules'_as_prules.prules_as_srules.srules_as_cake.mk_exp_cupcake) fact
obtain k where "Evaluate_Single.evaluate cake_sem_env (s \<lparr> clock := k \<rparr>) (term_to_cake t) = (s', Rval ml_v)"
using assms by blast
then have "Big_Step_Unclocked_Single.evaluate cake_sem_env (s \<lparr> clock := (clock s') \<rparr>) (term_to_cake t) (s', Rval ml_v)"
using unclocked_single_fun_eq by fastforce
have "cupcake_evaluate_single cake_sem_env (sterm_to_cake (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))) (Rval ml_v)"
apply (rule cupcake_single_complete)
apply fact+
done
hence "is_cupcake_value ml_v"
apply (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.cupcake_sem_env_preserve)
by (auto intro: pterm_to_sterm_wellformed)
hence "vno_abs (cake_to_value ml_v)"
using \<open>cake_no_abs _\<close>
by (metis rules_as_nrules.nrules_as_crules.crules_as_irules'.irules'_as_prules.prules_as_srules.srules_as_cake.cake_to_value_abs_free)
hence "no_abs (value_to_sterm (cake_to_value ml_v))"
by (metis vno_abs_value_to_sterm)
hence "no_abs (sterm_to_pterm (value_to_sterm (cake_to_value ml_v)))"
by (metis sterm_to_pterm convert_term_no_abs)
have "welldefined (term_to_nterm t)"
unfolding term_to_nterm'_def
apply (subst fresh_frun_def)
apply (subst pred_stateD[OF term_to_nterm_consts])
apply (subst surjective_pairing)
apply (rule refl)
apply fact
done
have "welldefined (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
apply (subst pterm_to_sterm_consts)
apply fact
apply (subst consts_nterm_to_pterm)
apply fact+
done
have "\<not> shadows_consts t"
using assms unfolding shadows_consts_def fdisjnt_alt_def
by auto
hence "\<not> shadows_consts (term_to_nterm t)"
unfolding shadows_consts_def shadows_consts_def
apply auto
using term_to_nterm_all_vars[folded wellformed_term_def]
by (metis assms(6) fdisjnt_swap sup_idem)
have "\<not> shadows_consts (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
apply (subst pterm_to_sterm_shadows[symmetric])
apply fact
apply (subst shadows_nterm_to_pterm)
unfolding shadows_consts_def
apply simp
apply (rule term_to_nterm_all_vars[where T = "fempty", simplified, THEN fdisjnt_swap])
apply (fold wellformed_term_def)
apply fact
using \<open>closed t\<close> unfolding closed_except_def by (auto simp: fdisjnt_alt_def)
have "closed (term_to_nterm t)"
using assms unfolding closed_except_def
using term_to_nterm_vars unfolding wellformed_term_def by blast
hence "closed (nterm_to_pterm (term_to_nterm t))"
using closed_nterm_to_pterm unfolding closed_except_def
by auto
have "closed (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
unfolding closed_except_def
apply (subst pterm_to_sterm_frees)
apply fact
using \<open>closed (term_to_nterm t)\<close> closed_nterm_to_pterm unfolding closed_except_def
by auto
have "fmap_of_list compiled \<turnstile>\<^sub>v pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> cake_to_value ml_v"
by (rule cake_semantic_correctness) fact+
hence "fmap_of_list rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules \<turnstile>\<^sub>v pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> cake_to_value ml_v"
using assms unfolding compiled_def by simp
hence "rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules, fmempty \<turnstile>\<^sub>v pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> cake_to_value ml_v"
proof (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.veval'_correct'')
show "\<not> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.shadows_consts (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
using \<open>\<not> shadows_consts (_::sterm)\<close> \<open>?heads = heads_of rs\<close> by auto
next
show "consts (pterm_to_sterm (nterm_to_pterm (term_to_nterm t))) |\<subseteq>| rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.all_consts"
using \<open>welldefined (pterm_to_sterm _)\<close> \<open>?heads = _\<close> by auto
qed fact+
hence "Rewriting_Sterm.compile (Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile)))), fmempty \<turnstile>\<^sub>s pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> value_to_sterm (cake_to_value ml_v)"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.veval_correct) fact+
hence "Rewriting_Sterm.compile (Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile)))) \<turnstile>\<^sub>s pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<longrightarrow>* value_to_sterm (cake_to_value ml_v)"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.seval_correct) fact
hence "Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile))) \<turnstile>\<^sub>p sterm_to_pterm (pterm_to_sterm (nterm_to_pterm (term_to_nterm t))) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.compile_correct_rt)
hence "Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile))) \<turnstile>\<^sub>p nterm_to_pterm (term_to_nterm t) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
by (subst (asm) pterm_to_sterm_sterm_to_pterm) fact
hence "transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile)) \<turnstile>\<^sub>i nterm_to_pterm (term_to_nterm t) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_irules.compile_correct_rt)
(rule rules_as_nrules.crules_as_irules.transform_finished)
have "Rewriting_Pterm_Elim.compile (consts_of compile) \<turnstile>\<^sub>i nterm_to_pterm (term_to_nterm t) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
apply (rule rules_as_nrules.crules_as_irules.transform_correct_rt_n_no_abs)
using \<open>transform_irule_set_iter _ \<turnstile>\<^sub>i _ \<longrightarrow>* _\<close> unfolding transform_irule_set_iter_def
apply simp
apply fact+
done
then obtain t' where "compile \<turnstile>\<^sub>n term_to_nterm t \<longrightarrow>* t'" "t' \<approx>\<^sub>i sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
using \<open>closed (term_to_nterm t)\<close>
by (metis rules_as_nrules.compile_correct_rt)
hence "no_abs t'"
using \<open>no_abs (sterm_to_pterm _)\<close>
by (metis irelated_no_abs)
have "rs \<turnstile> nterm_to_term' (term_to_nterm t) \<longrightarrow>* nterm_to_term' t'"
by (rule compile_correct_rt) fact+
hence "rs \<turnstile> t \<longrightarrow>* nterm_to_term' t'"
apply (subst (asm) fresh_frun_def)
apply (subst (asm) term_to_nterm_nterm_to_term[where S = "fempty" and t = t, simplified])
apply (fold wellformed_term_def)
apply fact
using assms unfolding closed_except_def by auto
have "nterm_to_pterm t' = sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
using \<open>t' \<approx>\<^sub>i _\<close>
by auto
hence "(convert_term t' :: pterm) = convert_term (value_to_sterm (cake_to_value ml_v))"
apply (subst (asm) nterm_to_pterm)
apply fact
apply (subst (asm) sterm_to_pterm)
apply fact
apply assumption
done
hence "nterm_to_term' t' = convert_term (value_to_sterm (cake_to_value ml_v))"
apply (subst nterm_to_term')
apply (rule \<open>no_abs t'\<close>)
apply (rule convert_term_inj)
subgoal premises
apply (rule convert_term_no_abs)
apply fact
done
subgoal premises
apply (rule convert_term_no_abs)
apply fact
done
apply (subst convert_term_idem)
apply (rule \<open>no_abs t'\<close>)
apply (subst convert_term_idem)
apply (rule \<open>no_abs (value_to_sterm (cake_to_value ml_v))\<close>)
apply assumption
done
thus ?thesis
using \<open>rs \<turnstile> t \<longrightarrow>* nterm_to_term' t'\<close> by simp
qed
end
end
\ No newline at end of file
diff --git a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy
--- a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy
+++ b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy
@@ -1,1733 +1,1733 @@
section \<open>Higher-order term rewriting with explicit pattern matching\<close>
theory Rewriting_Pterm_Elim
imports
Rewriting_Nterm
"../Terms/Pterm"
begin
subsection \<open>Intermediate rule sets\<close>
type_synonym irules = "(term list \<times> pterm) fset"
type_synonym irule_set = "(name \<times> irules) fset"
locale pre_irules = constants C_info "fst |`| rs" for C_info and rs :: "irule_set"
locale irules = pre_irules +
assumes fmap: "is_fmap rs"
assumes nonempty: "rs \<noteq> {||}"
assumes inner:
"fBall rs (\<lambda>(_, irs).
arity_compatibles irs \<and>
is_fmap irs \<and>
patterns_compatibles irs \<and>
irs \<noteq> {||} \<and>
fBall irs (\<lambda>(pats, rhs).
linears pats \<and>
abs_ish pats rhs \<and>
closed_except rhs (freess pats) \<and>
fdisjnt (freess pats) all_consts \<and>
wellformed rhs \<and>
\<not> shadows_consts rhs \<and>
welldefined rhs))"
lemma (in pre_irules) irulesI:
assumes "\<And>name irs. (name, irs) |\<in>| rs \<Longrightarrow> arity_compatibles irs"
assumes "\<And>name irs. (name, irs) |\<in>| rs \<Longrightarrow> is_fmap irs"
assumes "\<And>name irs. (name, irs) |\<in>| rs \<Longrightarrow> patterns_compatibles irs"
assumes "\<And>name irs. (name, irs) |\<in>| rs \<Longrightarrow> irs \<noteq> {||}"
assumes "\<And>name irs pats rhs. (name, irs) |\<in>| rs \<Longrightarrow> (pats, rhs) |\<in>| irs \<Longrightarrow> linears pats"
assumes "\<And>name irs pats rhs. (name, irs) |\<in>| rs \<Longrightarrow> (pats, rhs) |\<in>| irs \<Longrightarrow> abs_ish pats rhs"
assumes "\<And>name irs pats rhs. (name, irs) |\<in>| rs \<Longrightarrow> (pats, rhs) |\<in>| irs \<Longrightarrow> fdisjnt (freess pats) all_consts"
assumes "\<And>name irs pats rhs. (name, irs) |\<in>| rs \<Longrightarrow> (pats, rhs) |\<in>| irs \<Longrightarrow> closed_except rhs (freess pats)"
assumes "\<And>name irs pats rhs. (name, irs) |\<in>| rs \<Longrightarrow> (pats, rhs) |\<in>| irs \<Longrightarrow> wellformed rhs"
assumes "\<And>name irs pats rhs. (name, irs) |\<in>| rs \<Longrightarrow> (pats, rhs) |\<in>| irs \<Longrightarrow> \<not> shadows_consts rhs"
assumes "\<And>name irs pats rhs. (name, irs) |\<in>| rs \<Longrightarrow> (pats, rhs) |\<in>| irs \<Longrightarrow> welldefined rhs"
assumes "is_fmap rs" "rs \<noteq> {||}"
shows "irules C_info rs"
using assms unfolding irules_axioms_def irules_def
by (auto simp: prod_fBallI intro: pre_irules_axioms)
lemmas irulesI[intro!] = pre_irules.irulesI[unfolded pre_irules_def]
subsubsection \<open>Translation from @{typ nterm} to @{typ pterm}\<close>
fun nterm_to_pterm :: "nterm \<Rightarrow> pterm" where
"nterm_to_pterm (Nvar s) = Pvar s" |
"nterm_to_pterm (Nconst s) = Pconst s" |
"nterm_to_pterm (t\<^sub>1 $\<^sub>n t\<^sub>2) = nterm_to_pterm t\<^sub>1 $\<^sub>p nterm_to_pterm t\<^sub>2" |
"nterm_to_pterm (\<Lambda>\<^sub>n x. t) = (\<Lambda>\<^sub>p x. nterm_to_pterm t)"
lemma nterm_to_pterm_inj: "nterm_to_pterm x = nterm_to_pterm y \<Longrightarrow> x = y"
by (induction y arbitrary: x) (auto elim: nterm_to_pterm.elims)
lemma nterm_to_pterm:
assumes "no_abs t"
shows "nterm_to_pterm t = convert_term t"
using assms
apply induction
apply auto
by (auto simp: free_nterm_def free_pterm_def const_nterm_def const_pterm_def app_nterm_def app_pterm_def)
lemma nterm_to_pterm_frees[simp]: "frees (nterm_to_pterm t) = frees t"
by (induct t) auto
lemma closed_nterm_to_pterm[intro]: "closed_except (nterm_to_pterm t) (frees t)"
unfolding closed_except_def by simp
lemma (in constants) shadows_nterm_to_pterm[simp]: "shadows_consts (nterm_to_pterm t) = shadows_consts t"
by (induct t) (auto simp: shadows_consts_def fdisjnt_alt_def)
lemma wellformed_nterm_to_pterm[intro]: "wellformed (nterm_to_pterm t)"
by (induct t) auto
lemma consts_nterm_to_pterm[simp]: "consts (nterm_to_pterm t) = consts t"
by (induct t) auto
subsubsection \<open>Translation from @{typ crule_set} to @{typ irule_set}\<close>
definition translate_crules :: "crules \<Rightarrow> irules" where
"translate_crules = fimage (map_prod id nterm_to_pterm)"
definition compile :: "crule_set \<Rightarrow> irule_set" where
"compile = fimage (map_prod id translate_crules)"
lemma compile_heads: "fst |`| compile rs = fst |`| rs"
unfolding compile_def by simp
lemma (in crules) compile_rules: "irules C_info (compile rs)"
proof
have "is_fmap rs"
using fmap by simp
thus "is_fmap (compile rs)"
unfolding compile_def map_prod_def id_apply by (rule is_fmap_image)
show "compile rs \<noteq> {||}"
using nonempty unfolding compile_def by auto
show "constants C_info (fst |`| compile rs)"
proof
show "fdisjnt (fst |`| compile rs) C"
using disjnt unfolding compile_def
by force
next
show "distinct all_constructors"
by (fact distinct_ctr)
qed
fix name irs
assume irs: "(name, irs) |\<in>| compile rs"
then obtain irs' where "(name, irs') |\<in>| rs" "irs = translate_crules irs'"
unfolding compile_def by force
hence "arity_compatibles irs'"
using inner by (blast dest: fpairwiseD)
thus "arity_compatibles irs"
unfolding \<open>irs = translate_crules irs'\<close> translate_crules_def
by (force dest: fpairwiseD)
have "patterns_compatibles irs'"
using \<open>(name, irs') |\<in>| rs\<close> inner
by (blast dest: fpairwiseD)
thus "patterns_compatibles irs"
unfolding \<open>irs = _\<close> translate_crules_def
by (auto dest: fpairwiseD)
have "is_fmap irs'"
using \<open>(name, irs') |\<in>| rs\<close> inner by auto
thus "is_fmap irs"
unfolding \<open>irs = translate_crules irs'\<close> translate_crules_def map_prod_def id_apply
by (rule is_fmap_image)
have "irs' \<noteq> {||}"
using \<open>(name, irs') |\<in>| rs\<close> inner by auto
thus "irs \<noteq> {||}"
unfolding \<open>irs = translate_crules irs'\<close> translate_crules_def by simp
fix pats rhs
assume "(pats, rhs) |\<in>| irs"
then obtain rhs' where "(pats, rhs') |\<in>| irs'" "rhs = nterm_to_pterm rhs'"
unfolding \<open>irs = translate_crules irs'\<close> translate_crules_def by force
hence "linears pats" "pats \<noteq> []" "frees rhs' |\<subseteq>| freess pats" "\<not> shadows_consts rhs'"
using fbspec[OF inner \<open>(name, irs') |\<in>| rs\<close>]
by blast+
show "linears pats" by fact
show "closed_except rhs (freess pats)"
unfolding \<open>rhs = nterm_to_pterm rhs'\<close>
using \<open>frees rhs' |\<subseteq>| freess pats\<close>
by (metis dual_order.trans closed_nterm_to_pterm closed_except_def)
show "wellformed rhs"
unfolding \<open>rhs = nterm_to_pterm rhs'\<close> by auto
have "fdisjnt (freess pats) all_consts"
using \<open>(pats, rhs') |\<in>| irs'\<close> \<open>(name, irs') |\<in>| rs\<close> inner
by blast
thus "fdisjnt (freess pats) (pre_constants.all_consts C_info (fst |`| compile rs))"
unfolding compile_def by simp
have "\<not> shadows_consts rhs"
unfolding \<open>rhs = _\<close> using \<open>\<not> shadows_consts _\<close> by simp
thus "\<not> pre_constants.shadows_consts C_info (fst |`| compile rs) rhs"
unfolding compile_heads .
show "abs_ish pats rhs"
using \<open>pats \<noteq> []\<close> unfolding abs_ish_def by simp
have "welldefined rhs'"
using fbspec[OF inner \<open>(name, irs') |\<in>| rs\<close>, simplified]
using \<open>(pats, rhs') |\<in>| irs'\<close>
by blast
thus "pre_constants.welldefined C_info (fst |`| compile rs) rhs"
unfolding compile_def \<open>rhs = _\<close>
by simp
qed
sublocale crules \<subseteq> crules_as_irules: irules C_info "compile rs"
by (fact compile_rules)
subsubsection \<open>Transformation of @{typ irule_set}\<close>
definition transform_irules :: "irules \<Rightarrow> irules" where
"transform_irules rs = (
if arity rs = 0 then rs
else map_prod id Pabs |`| fgroup_by (\<lambda>(pats, rhs). (butlast pats, (last pats, rhs))) rs)"
lemma arity_compatibles_transform_irules:
assumes "arity_compatibles rs"
shows "arity_compatibles (transform_irules rs)"
proof (cases "arity rs = 0")
case True
thus ?thesis
unfolding transform_irules_def using assms by simp
next
case False
let ?rs' = "transform_irules rs"
let ?f = "\<lambda>(pats, rhs). (butlast pats, (last pats, rhs))"
let ?grp = "fgroup_by ?f rs"
have rs': "?rs' = map_prod id Pabs |`| ?grp"
using False unfolding transform_irules_def by simp
show ?thesis
proof safe
fix pats\<^sub>1 rhs\<^sub>1 pats\<^sub>2 rhs\<^sub>2
assume "(pats\<^sub>1, rhs\<^sub>1) |\<in>| ?rs'" "(pats\<^sub>2, rhs\<^sub>2) |\<in>| ?rs'"
then obtain rhs\<^sub>1' rhs\<^sub>2' where "(pats\<^sub>1, rhs\<^sub>1') |\<in>| ?grp" "(pats\<^sub>2, rhs\<^sub>2') |\<in>| ?grp"
unfolding rs' by auto
then obtain pats\<^sub>1' pats\<^sub>2' x y \<comment> \<open>dummies\<close>
where "fst (?f (pats\<^sub>1', x)) = pats\<^sub>1" "(pats\<^sub>1', x) |\<in>| rs"
and "fst (?f (pats\<^sub>2', y)) = pats\<^sub>2" "(pats\<^sub>2', y) |\<in>| rs"
by (fastforce simp: split_beta elim: fgroup_byE2)
hence "pats\<^sub>1 = butlast pats\<^sub>1'" "pats\<^sub>2 = butlast pats\<^sub>2'" "length pats\<^sub>1' = length pats\<^sub>2'"
using assms by (force dest: fpairwiseD)+
thus "length pats\<^sub>1 = length pats\<^sub>2"
by auto
qed
qed
lemma arity_transform_irules:
assumes "arity_compatibles rs" "rs \<noteq> {||}"
shows "arity (transform_irules rs) = (if arity rs = 0 then 0 else arity rs - 1)"
proof (cases "arity rs = 0")
case True
thus ?thesis
unfolding transform_irules_def by simp
next
case False
let ?f = "\<lambda>(pats, rhs). (butlast pats, (last pats, rhs))"
let ?grp = "fgroup_by ?f rs"
let ?rs' = "map_prod id Pabs |`| ?grp"
have "arity ?rs' = arity rs - 1"
proof (rule arityI)
show "fBall ?rs' (\<lambda>(pats, _). length pats = arity rs - 1)"
proof (rule prod_fBallI)
fix pats rhs
assume "(pats, rhs) |\<in>| ?rs'"
then obtain cs where "(pats, cs) |\<in>| ?grp" "rhs = Pabs cs"
by force
then obtain pats' x \<comment> \<open>dummy\<close>
where "pats = butlast pats'" "(pats', x) |\<in>| rs"
by (fastforce simp: split_beta elim: fgroup_byE2)
hence "length pats' = arity rs"
using assms by (metis arity_compatible_length)
thus "length pats = arity rs - 1"
unfolding \<open>pats = butlast pats'\<close> using False by simp
qed
next
show "?rs' \<noteq> {||}"
using assms by (simp add: fgroup_by_nonempty)
qed
with False show ?thesis
unfolding transform_irules_def by simp
qed
definition transform_irule_set :: "irule_set \<Rightarrow> irule_set" where
"transform_irule_set = fimage (map_prod id transform_irules)"
lemma transform_irule_set_heads: "fst |`| transform_irule_set rs = fst |`| rs"
unfolding transform_irule_set_def by simp
lemma (in irules) rules_transform: "irules C_info (transform_irule_set rs)"
proof
have "is_fmap rs"
using fmap by simp
thus "is_fmap (transform_irule_set rs)"
unfolding transform_irule_set_def map_prod_def id_apply by (rule is_fmap_image)
show "transform_irule_set rs \<noteq> {||}"
using nonempty unfolding transform_irule_set_def by auto
show "constants C_info (fst |`| transform_irule_set rs)"
proof
show "fdisjnt (fst |`| transform_irule_set rs) C"
using disjnt unfolding transform_irule_set_def
by force
next
show "distinct all_constructors"
by (fact distinct_ctr)
qed
fix name irs
assume irs: "(name, irs) |\<in>| transform_irule_set rs"
then obtain irs' where "(name, irs') |\<in>| rs" "irs = transform_irules irs'"
unfolding transform_irule_set_def by force
hence "arity_compatibles irs'"
using inner by (blast dest: fpairwiseD)
thus "arity_compatibles irs"
unfolding \<open>irs = transform_irules irs'\<close> by (rule arity_compatibles_transform_irules)
have "irs' \<noteq> {||}"
using \<open>(name, irs') |\<in>| rs\<close> inner by blast
thus "irs \<noteq> {||}"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (simp add: fgroup_by_nonempty)
let ?f = "\<lambda>(pats, rhs). (butlast pats, (last pats, rhs))"
let ?grp = "fgroup_by ?f irs'"
have "patterns_compatibles irs'"
using \<open>(name, irs') |\<in>| rs\<close> inner
by (blast dest: fpairwiseD)
show "patterns_compatibles irs"
proof (cases "arity irs' = 0")
case True
thus ?thesis
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
using \<open>patterns_compatibles irs'\<close> by simp
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
show ?thesis
proof safe
fix pats\<^sub>1 rhs\<^sub>1 pats\<^sub>2 rhs\<^sub>2
assume "(pats\<^sub>1, rhs\<^sub>1) |\<in>| irs" "(pats\<^sub>2, rhs\<^sub>2) |\<in>| irs"
with irs' obtain cs\<^sub>1 cs\<^sub>2 where "(pats\<^sub>1, cs\<^sub>1) |\<in>| ?grp" "(pats\<^sub>2, cs\<^sub>2) |\<in>| ?grp"
by force
then obtain pats\<^sub>1' pats\<^sub>2' and x y \<comment> \<open>dummies\<close>
where "(pats\<^sub>1', x) |\<in>| irs'" "(pats\<^sub>2', y) |\<in>| irs'"
and "pats\<^sub>1 = butlast pats\<^sub>1'" "pats\<^sub>2 = butlast pats\<^sub>2'"
unfolding irs'
by (fastforce elim: fgroup_byE2)
hence "patterns_compatible pats\<^sub>1' pats\<^sub>2'"
using \<open>patterns_compatibles irs'\<close> by (auto dest: fpairwiseD)
thus "patterns_compatible pats\<^sub>1 pats\<^sub>2"
unfolding \<open>pats\<^sub>1 = _\<close> \<open>pats\<^sub>2 = _\<close>
by auto
qed
qed
have "is_fmap irs'"
using \<open>(name, irs') |\<in>| rs\<close> inner by blast
show "is_fmap irs"
proof (cases "arity irs' = 0")
case True
thus ?thesis
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
using \<open>is_fmap irs'\<close> by simp
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
show ?thesis
proof
fix pats rhs\<^sub>1 rhs\<^sub>2
assume "(pats, rhs\<^sub>1) |\<in>| irs" "(pats, rhs\<^sub>2) |\<in>| irs"
with irs' obtain cs\<^sub>1 cs\<^sub>2
where "(pats, cs\<^sub>1) |\<in>| ?grp" "rhs\<^sub>1 = Pabs cs\<^sub>1"
and "(pats, cs\<^sub>2) |\<in>| ?grp" "rhs\<^sub>2 = Pabs cs\<^sub>2"
by force
moreover have "is_fmap ?grp"
by auto
ultimately show "rhs\<^sub>1 = rhs\<^sub>2"
by (auto dest: is_fmapD)
qed
qed
fix pats rhs
assume "(pats, rhs) |\<in>| irs"
show "linears pats"
proof (cases "arity irs' = 0")
case True
thus ?thesis
using \<open>(pats, rhs) |\<in>| irs\<close> \<open>(name, irs') |\<in>| rs\<close> inner
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (smt fBallE split_conv)
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
then obtain cs where "(pats, cs) |\<in>| ?grp"
using \<open>(pats, rhs) |\<in>| irs\<close> by force
then obtain pats' x \<comment> \<open>dummy\<close>
where "fst (?f (pats', x)) = pats" "(pats', x) |\<in>| irs'"
by (fastforce simp: split_beta elim: fgroup_byE2)
hence "pats = butlast pats'"
by simp
moreover have "linears pats'"
using \<open>(pats', x) |\<in>| irs'\<close> \<open>(name, irs') |\<in>| rs\<close> inner
by blast
ultimately show ?thesis
by auto
qed
have "fdisjnt (freess pats) all_consts"
proof (cases "arity irs' = 0")
case True
thus ?thesis
using \<open>(pats, rhs) |\<in>| irs\<close> \<open>(name, irs') |\<in>| rs\<close> inner
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (smt fBallE split_conv)
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
then obtain cs where "(pats, cs) |\<in>| ?grp"
using \<open>(pats, rhs) |\<in>| irs\<close> by force
then obtain pats' x \<comment> \<open>dummy\<close>
where "fst (?f (pats', x)) = pats" "(pats', x) |\<in>| irs'"
by (fastforce simp: split_beta elim: fgroup_byE2)
hence "pats = butlast pats'"
by simp
moreover have "fdisjnt (freess pats') all_consts"
using \<open>(pats', x) |\<in>| irs'\<close> \<open>(name, irs') |\<in>| rs\<close> inner by blast
ultimately show ?thesis
by (metis subsetI in_set_butlastD freess_subset fdisjnt_subset_left)
qed
thus "fdisjnt (freess pats) (pre_constants.all_consts C_info (fst |`| transform_irule_set rs))"
unfolding transform_irule_set_def by simp
show "closed_except rhs (freess pats)"
proof (cases "arity irs' = 0")
case True
thus ?thesis
using \<open>(pats, rhs) |\<in>| irs\<close> \<open>(name, irs') |\<in>| rs\<close> inner
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (smt fBallE split_conv)
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
then obtain cs where "(pats, cs) |\<in>| ?grp" "rhs = Pabs cs"
using \<open>(pats, rhs) |\<in>| irs\<close> by force
show ?thesis
unfolding \<open>rhs = Pabs cs\<close> closed_except_simps
proof safe
fix pat t
assume "(pat, t) |\<in>| cs"
then obtain pats' where "(pats', t) |\<in>| irs'" "?f (pats', t) = (pats, (pat, t))"
using \<open>(pats, cs) |\<in>| ?grp\<close> by auto
hence "closed_except t (freess pats')"
using \<open>(name, irs') |\<in>| rs\<close> inner by blast
have "pats' \<noteq> []"
using \<open>arity_compatibles irs'\<close> \<open>(pats', t) |\<in>| irs'\<close> False
by (metis list.size(3) arity_compatible_length)
hence "pats' = pats @ [pat]"
using \<open>?f (pats', t) = (pats, (pat, t))\<close>
by (fastforce simp: split_beta snoc_eq_iff_butlast)
hence "freess pats |\<union>| frees pat = freess pats'"
unfolding freess_def by auto
thus "closed_except t (freess pats |\<union>| frees pat)"
using \<open>closed_except t (freess pats')\<close> by simp
qed
qed
show "wellformed rhs"
proof (cases "arity irs' = 0")
case True
thus ?thesis
using \<open>(pats, rhs) |\<in>| irs\<close> \<open>(name, irs') |\<in>| rs\<close> inner
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (smt fBallE split_conv)
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
then obtain cs where "(pats, cs) |\<in>| ?grp" "rhs = Pabs cs"
using \<open>(pats, rhs) |\<in>| irs\<close> by force
show ?thesis
unfolding \<open>rhs = Pabs cs\<close>
proof (rule wellformed_PabsI)
show "cs \<noteq> {||}"
using \<open>(pats, cs) |\<in>| ?grp\<close> \<open>irs' \<noteq> {||}\<close>
by (meson femptyE fgroup_by_nonempty_inner)
next
show "is_fmap cs"
proof
fix pat t\<^sub>1 t\<^sub>2
assume "(pat, t\<^sub>1) |\<in>| cs" "(pat, t\<^sub>2) |\<in>| cs"
then obtain pats\<^sub>1' pats\<^sub>2'
where "(pats\<^sub>1', t\<^sub>1) |\<in>| irs'" "?f (pats\<^sub>1', t\<^sub>1) = (pats, (pat, t\<^sub>1))"
and "(pats\<^sub>2', t\<^sub>2) |\<in>| irs'" "?f (pats\<^sub>2', t\<^sub>2) = (pats, (pat, t\<^sub>2))"
using \<open>(pats, cs) |\<in>| ?grp\<close> by force
moreover hence "pats\<^sub>1' \<noteq> []" "pats\<^sub>2' \<noteq> []"
using \<open>arity_compatibles irs'\<close> False
unfolding prod.case
by (metis list.size(3) arity_compatible_length)+
ultimately have "pats\<^sub>1' = pats @ [pat]" "pats\<^sub>2' = pats @ [pat]"
unfolding split_beta fst_conv snd_conv
by (metis prod.inject snoc_eq_iff_butlast)+
with \<open>is_fmap irs'\<close> show "t\<^sub>1 = t\<^sub>2"
using \<open>(pats\<^sub>1', t\<^sub>1) |\<in>| irs'\<close> \<open>(pats\<^sub>2', t\<^sub>2) |\<in>| irs'\<close>
by (blast dest: is_fmapD)
qed
next
show "pattern_compatibles cs"
proof safe
fix pat\<^sub>1 rhs\<^sub>1 pat\<^sub>2 rhs\<^sub>2
assume "(pat\<^sub>1, rhs\<^sub>1) |\<in>| cs" "(pat\<^sub>2, rhs\<^sub>2) |\<in>| cs"
then obtain pats\<^sub>1' pats\<^sub>2'
where "(pats\<^sub>1', rhs\<^sub>1) |\<in>| irs'" "?f (pats\<^sub>1', rhs\<^sub>1) = (pats, (pat\<^sub>1, rhs\<^sub>1))"
and "(pats\<^sub>2', rhs\<^sub>2) |\<in>| irs'" "?f (pats\<^sub>2', rhs\<^sub>2) = (pats, (pat\<^sub>2, rhs\<^sub>2))"
using \<open>(pats, cs) |\<in>| ?grp\<close>
by force
moreover hence "pats\<^sub>1' \<noteq> []" "pats\<^sub>2' \<noteq> []"
using \<open>arity_compatibles irs'\<close> False
unfolding prod.case
by (metis list.size(3) arity_compatible_length)+
ultimately have "pats\<^sub>1' = pats @ [pat\<^sub>1]" "pats\<^sub>2' = pats @ [pat\<^sub>2]"
unfolding split_beta fst_conv snd_conv
by (metis prod.inject snoc_eq_iff_butlast)+
moreover have "patterns_compatible pats\<^sub>1' pats\<^sub>2'"
using \<open>(pats\<^sub>1', rhs\<^sub>1) |\<in>| irs'\<close> \<open>(pats\<^sub>2', rhs\<^sub>2) |\<in>| irs'\<close> \<open>patterns_compatibles irs'\<close>
by (auto dest: fpairwiseD)
ultimately show "pattern_compatible pat\<^sub>1 pat\<^sub>2"
by (auto elim: rev_accum_rel_snoc_eqE)
qed
next
fix pat t
assume "(pat, t) |\<in>| cs"
then obtain pats' where "(pats', t) |\<in>| irs'" "pat = last pats'"
using \<open>(pats, cs) |\<in>| ?grp\<close> by auto
moreover hence "pats' \<noteq> []"
using \<open>arity_compatibles irs'\<close> False
by (metis list.size(3) arity_compatible_length)
ultimately have "pat \<in> set pats'"
by auto
moreover have "linears pats'"
using \<open>(pats', t) |\<in>| irs'\<close> \<open>(name, irs') |\<in>| rs\<close> inner by blast
ultimately show "linear pat"
by (metis linears_linear)
show "wellformed t"
using \<open>(pats', t) |\<in>| irs'\<close> \<open>(name, irs') |\<in>| rs\<close> inner by blast
qed
qed
have "\<not> shadows_consts rhs"
proof (cases "arity irs' = 0")
case True
thus ?thesis
using \<open>(pats, rhs) |\<in>| irs\<close> \<open>(name, irs') |\<in>| rs\<close> inner
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (smt fBallE split_conv)
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
then obtain cs where "(pats, cs) |\<in>| ?grp" "rhs = Pabs cs"
using \<open>(pats, rhs) |\<in>| irs\<close> by force
show ?thesis
unfolding \<open>rhs = _\<close>
proof
assume "shadows_consts (Pabs cs)"
then obtain pat t where "(pat, t) |\<in>| cs" "shadows_consts t \<or> shadows_consts pat"
by force
then obtain pats' where "(pats', t) |\<in>| irs'" "pat = last pats'"
using \<open>(pats, cs) |\<in>| ?grp\<close> by auto
moreover hence "pats' \<noteq> []"
using \<open>arity_compatibles irs'\<close> False
by (metis list.size(3) arity_compatible_length)
ultimately have "pat \<in> set pats'"
by auto
show False
using \<open>shadows_consts t \<or> shadows_consts pat\<close>
proof
assume "shadows_consts t"
thus False
using \<open>(name, irs') |\<in>| rs\<close> \<open>(pats', t) |\<in>| irs'\<close> inner by blast
next
assume "shadows_consts pat"
have "fdisjnt (freess pats') all_consts"
using \<open>(name, irs') |\<in>| rs\<close> \<open>(pats', t) |\<in>| irs'\<close> inner by blast
have "fdisjnt (frees pat) all_consts"
apply (rule fdisjnt_subset_left)
apply (subst freess_single[symmetric])
apply (rule freess_subset)
apply simp
apply fact+
done
thus False
using \<open>shadows_consts pat\<close>
unfolding shadows_consts_def fdisjnt_alt_def by auto
qed
qed
qed
thus "\<not> pre_constants.shadows_consts C_info (fst |`| transform_irule_set rs) rhs"
by (simp add: transform_irule_set_heads)
show "abs_ish pats rhs"
proof (cases "arity irs' = 0")
case True
thus ?thesis
using \<open>(pats, rhs) |\<in>| irs\<close> \<open>(name, irs') |\<in>| rs\<close> inner
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (smt fBallE split_conv)
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
then obtain cs where "(pats, cs) |\<in>| ?grp" "rhs = Pabs cs"
using \<open>(pats, rhs) |\<in>| irs\<close> by force
thus ?thesis
unfolding abs_ish_def by (simp add: is_abs_def term_cases_def)
qed
have "welldefined rhs"
proof (cases "arity irs' = 0")
case True
hence \<open>(pats, rhs) |\<in>| irs'\<close>
using \<open>(pats, rhs) |\<in>| irs\<close> \<open>(name, irs') |\<in>| rs\<close> inner
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def
by (smt fBallE split_conv)
thus ?thesis
unfolding transform_irule_set_def
using fbspec[OF inner \<open>(name, irs') |\<in>| rs\<close>, simplified]
by force
next
case False
hence irs': "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = transform_irules irs'\<close> transform_irules_def by simp
then obtain cs where "(pats, cs) |\<in>| ?grp" "rhs = Pabs cs"
using \<open>(pats, rhs) |\<in>| irs\<close> by force
show ?thesis
unfolding \<open>rhs = _\<close>
apply simp
apply (rule ffUnion_least)
unfolding ball_simps
apply rule
apply (rename_tac x, case_tac x, hypsubst_thin)
apply simp
subgoal premises prems for pat t
proof -
from prems obtain pats' where "(pats', t) |\<in>| irs'"
using \<open>(pats, cs) |\<in>| ?grp\<close> by auto
hence "welldefined t"
using fbspec[OF inner \<open>(name, irs') |\<in>| rs\<close>, simplified]
by blast
thus ?thesis
unfolding transform_irule_set_def
by simp
qed
done
qed
thus "pre_constants.welldefined C_info (fst |`| transform_irule_set rs) rhs"
unfolding transform_irule_set_heads .
qed
subsubsection \<open>Matching and rewriting\<close>
definition irewrite_step :: "name \<Rightarrow> term list \<Rightarrow> pterm \<Rightarrow> pterm \<Rightarrow> pterm option" where
"irewrite_step name pats rhs t = map_option (subst rhs) (match (name $$ pats) t)"
abbreviation irewrite_step' :: "name \<Rightarrow> term list \<Rightarrow> pterm \<Rightarrow> pterm \<Rightarrow> pterm \<Rightarrow> bool" ("_, _, _ \<turnstile>\<^sub>i/ _ \<rightarrow>/ _" [50,0,50] 50) where
"name, pats, rhs \<turnstile>\<^sub>i t \<rightarrow> u \<equiv> irewrite_step name pats rhs t = Some u"
lemma irewrite_stepI:
assumes "match (name $$ pats) t = Some env" "subst rhs env = u"
shows "name, pats, rhs \<turnstile>\<^sub>i t \<rightarrow> u"
using assms unfolding irewrite_step_def by simp
inductive irewrite :: "irule_set \<Rightarrow> pterm \<Rightarrow> pterm \<Rightarrow> bool" ("_/ \<turnstile>\<^sub>i/ _ \<longrightarrow>/ _" [50,0,50] 50) for irs where
step: "\<lbrakk> (name, rs) |\<in>| irs; (pats, rhs) |\<in>| rs; name, pats, rhs \<turnstile>\<^sub>i t \<rightarrow> t' \<rbrakk> \<Longrightarrow> irs \<turnstile>\<^sub>i t \<longrightarrow> t'" |
beta: "\<lbrakk> c |\<in>| cs; c \<turnstile> t \<rightarrow> t' \<rbrakk> \<Longrightarrow> irs \<turnstile>\<^sub>i Pabs cs $\<^sub>p t \<longrightarrow> t'" |
"fun": "irs \<turnstile>\<^sub>i t \<longrightarrow> t' \<Longrightarrow> irs \<turnstile>\<^sub>i t $\<^sub>p u \<longrightarrow> t' $\<^sub>p u" |
arg: "irs \<turnstile>\<^sub>i u \<longrightarrow> u' \<Longrightarrow> irs \<turnstile>\<^sub>i t $\<^sub>p u \<longrightarrow> t $\<^sub>p u'"
global_interpretation irewrite: rewriting "irewrite rs" for rs
by standard (auto intro: irewrite.intros simp: app_pterm_def)+
abbreviation irewrite_rt :: "irule_set \<Rightarrow> pterm \<Rightarrow> pterm \<Rightarrow> bool" ("_/ \<turnstile>\<^sub>i/ _ \<longrightarrow>*/ _" [50,0,50] 50) where
"irewrite_rt rs \<equiv> (irewrite rs)\<^sup>*\<^sup>*"
lemma (in irules) irewrite_closed:
assumes "rs \<turnstile>\<^sub>i t \<longrightarrow> u" "closed t"
shows "closed u"
using assms proof induction
case (step name rs pats rhs t t')
then obtain env where "match (name $$ pats) t = Some env" "t' = subst rhs env"
unfolding irewrite_step_def by auto
hence "closed_env env"
using step by (auto intro: closed.match)
show ?case
unfolding \<open>t' = _\<close>
apply (subst closed_except_def)
apply (subst subst_frees)
apply fact
apply (subst match_dom)
apply fact
apply (subst frees_list_comb)
apply simp
apply (subst closed_except_def[symmetric])
using inner step by blast
next
case (beta c cs t t')
then obtain pat rhs where "c = (pat, rhs)"
by (cases c) auto
with beta obtain env where "match pat t = Some env" "t' = subst rhs env"
by auto
moreover have "closed t"
using beta unfolding closed_except_def by simp
ultimately have "closed_env env"
using beta by (auto intro: closed.match)
show ?case
unfolding \<open>t' = subst rhs env\<close>
apply (subst closed_except_def)
apply (subst subst_frees)
apply fact
apply (subst match_dom)
apply fact
apply simp
apply (subst closed_except_def[symmetric])
using inner beta \<open>c = _\<close> by (auto simp: closed_except_simps)
qed (auto simp: closed_except_def)
corollary (in irules) irewrite_rt_closed:
assumes "rs \<turnstile>\<^sub>i t \<longrightarrow>* u" "closed t"
shows "closed u"
using assms by induction (auto intro: irewrite_closed)
subsubsection \<open>Correctness of translation\<close>
abbreviation irelated :: "nterm \<Rightarrow> pterm \<Rightarrow> bool" ("_ \<approx>\<^sub>i _" [0,50] 50) where
"n \<approx>\<^sub>i p \<equiv> nterm_to_pterm n = p"
global_interpretation irelated: term_struct_rel_strong irelated
by standard
(auto simp: app_pterm_def app_nterm_def const_pterm_def const_nterm_def elim: nterm_to_pterm.elims)
lemma irelated_vars: "t \<approx>\<^sub>i u \<Longrightarrow> frees t = frees u"
by auto
lemma irelated_no_abs:
assumes "t \<approx>\<^sub>i u"
shows "no_abs t \<longleftrightarrow> no_abs u"
using assms
apply (induction arbitrary: t)
apply (auto elim!: nterm_to_pterm.elims)
apply (fold const_nterm_def const_pterm_def free_nterm_def free_pterm_def app_pterm_def app_nterm_def)
by auto
lemma irelated_subst:
assumes "t \<approx>\<^sub>i u" "irelated.P_env nenv penv"
shows "subst t nenv \<approx>\<^sub>i subst u penv"
using assms proof (induction arbitrary: nenv penv u rule: nterm_to_pterm.induct)
case (1 s)
then show ?case
by (auto elim!: fmrel_cases[where x = s])
next
case 4
from 4(2)[symmetric] show ?case
apply simp
apply (rule 4)
apply simp
using 4(3)
by (simp add: fmrel_drop)
qed auto
lemma related_irewrite_step:
assumes "name, pats, nterm_to_pterm rhs \<turnstile>\<^sub>i u \<rightarrow> u'" "t \<approx>\<^sub>i u"
obtains t' where "unsplit_rule (name, pats, rhs) \<turnstile> t \<rightarrow> t'" "t' \<approx>\<^sub>i u'"
proof -
let ?rhs' = "nterm_to_pterm rhs"
let ?x = "name $$ pats"
from assms obtain env where "match ?x u = Some env" "u' = subst ?rhs' env"
unfolding irewrite_step_def by blast
then obtain nenv where "match ?x t = Some nenv" "irelated.P_env nenv env"
using assms
by (metis Option.is_none_def not_None_eq option.rel_distinct(1) option.sel rel_option_unfold irelated.match_rel)
show thesis
proof
show "unsplit_rule (name, pats, rhs) \<turnstile> t \<rightarrow> subst rhs nenv"
using \<open>match ?x t = _\<close> by auto
next
show "subst rhs nenv \<approx>\<^sub>i u'"
unfolding \<open>u' = _\<close> using \<open>irelated.P_env nenv env\<close>
by (auto intro: irelated_subst)
qed
qed
theorem (in nrules) compile_correct:
assumes "compile (consts_of rs) \<turnstile>\<^sub>i u \<longrightarrow> u'" "t \<approx>\<^sub>i u" "closed t"
obtains t' where "rs \<turnstile>\<^sub>n t \<longrightarrow> t'" "t' \<approx>\<^sub>i u'"
using assms(1-3) proof (induction arbitrary: t thesis rule: irewrite.induct)
case (step name irs pats rhs u u')
then obtain crs where "irs = translate_crules crs" "(name, crs) |\<in>| consts_of rs"
unfolding compile_def by force
moreover with step obtain rhs' where "rhs = nterm_to_pterm rhs'" "(pats, rhs') |\<in>| crs"
unfolding translate_crules_def by force
ultimately obtain rule where "split_rule rule = (name, (pats, rhs'))" "rule |\<in>| rs"
unfolding consts_of_def by blast
hence "nrule rule"
using all_rules by blast
obtain t' where "unsplit_rule (name, pats, rhs') \<turnstile> t \<rightarrow> t'" "t' \<approx>\<^sub>i u'"
using \<open>name, pats, rhs \<turnstile>\<^sub>i u \<rightarrow> u'\<close> \<open>t \<approx>\<^sub>i u\<close> unfolding \<open>rhs = nterm_to_pterm rhs'\<close>
by (elim related_irewrite_step)
hence "rule \<turnstile> t \<rightarrow> t'"
using \<open>nrule rule\<close> \<open>split_rule rule = (name, (pats, rhs'))\<close>
by (metis unsplit_split)
show ?case
proof (rule step.prems)
show "rs \<turnstile>\<^sub>n t \<longrightarrow> t'"
apply (rule nrewrite.step)
apply fact
apply fact
done
next
show "t' \<approx>\<^sub>i u'"
by fact
qed
next
case (beta c cs u u')
then obtain pat rhs where "c = (pat, rhs)" "(pat, rhs) |\<in>| cs"
by (cases c) auto
obtain v w where "t = v $\<^sub>n w" "v \<approx>\<^sub>i Pabs cs" "w \<approx>\<^sub>i u"
using \<open>t \<approx>\<^sub>i Pabs cs $\<^sub>p u\<close> by (auto elim: nterm_to_pterm.elims)
obtain x nrhs irhs where "v = (\<Lambda>\<^sub>n x. nrhs)" "cs = {| (Free x, irhs) |}" "nrhs \<approx>\<^sub>i irhs"
using \<open>v \<approx>\<^sub>i Pabs cs\<close> by (auto elim: nterm_to_pterm.elims)
hence "t = (\<Lambda>\<^sub>n x. nrhs) $\<^sub>n w" "\<Lambda>\<^sub>n x. nrhs \<approx>\<^sub>i \<Lambda>\<^sub>p x. irhs"
unfolding \<open>t = v $\<^sub>n w\<close> using \<open>v \<approx>\<^sub>i Pabs cs\<close> by auto
have "pat = Free x" "rhs = irhs"
using \<open>cs = {| (Free x, irhs) |}\<close> \<open>(pat, rhs) |\<in>| cs\<close> by auto
hence "(Free x, irhs) \<turnstile> u \<rightarrow> u'"
using beta \<open>c = _\<close> by simp
hence "u' = subst irhs (fmap_of_list [(x, u)])"
by simp
show ?case
proof (rule beta.prems)
show "rs \<turnstile>\<^sub>n t \<longrightarrow> subst nrhs (fmap_of_list [(x, w)])"
unfolding \<open>t = (\<Lambda>\<^sub>n x. nrhs) $\<^sub>n w\<close>
by (rule nrewrite.beta)
next
show "subst nrhs (fmap_of_list [(x, w)]) \<approx>\<^sub>i u'"
unfolding \<open>u' = subst irhs _\<close>
apply (rule irelated_subst)
apply fact
apply simp
apply rule
apply rule
apply fact
done
qed
next
case ("fun" v v' u)
obtain w x where "t = w $\<^sub>n x" "w \<approx>\<^sub>i v" "x \<approx>\<^sub>i u"
using \<open>t \<approx>\<^sub>i v $\<^sub>p u\<close> by (auto elim: nterm_to_pterm.elims)
with "fun" obtain w' where "rs \<turnstile>\<^sub>n w \<longrightarrow> w'" "w' \<approx>\<^sub>i v'"
unfolding closed_except_def by auto
show ?case
proof (rule fun.prems)
show "rs \<turnstile>\<^sub>n t \<longrightarrow> w' $\<^sub>n x"
unfolding \<open>t = w $\<^sub>n x\<close>
by (rule nrewrite.fun) fact
next
show "w' $\<^sub>n x \<approx>\<^sub>i v' $\<^sub>p u"
by auto fact+
qed
next
case (arg u u' v)
obtain w x where "t = w $\<^sub>n x" "w \<approx>\<^sub>i v" "x \<approx>\<^sub>i u"
using \<open> t \<approx>\<^sub>i v $\<^sub>p u\<close> by (auto elim: nterm_to_pterm.elims)
with arg obtain x' where "rs \<turnstile>\<^sub>n x \<longrightarrow> x'" "x' \<approx>\<^sub>i u'"
unfolding closed_except_def by auto
show ?case
proof (rule arg.prems)
show "rs \<turnstile>\<^sub>n t \<longrightarrow> w $\<^sub>n x'"
unfolding \<open>t = w $\<^sub>n x\<close>
by (rule nrewrite.arg) fact
next
show "w $\<^sub>n x' \<approx>\<^sub>i v $\<^sub>p u'"
by auto fact+
qed
qed
corollary (in nrules) compile_correct_rt:
assumes "compile (consts_of rs) \<turnstile>\<^sub>i u \<longrightarrow>* u'" "t \<approx>\<^sub>i u" "closed t"
obtains t' where "rs \<turnstile>\<^sub>n t \<longrightarrow>* t'" "t' \<approx>\<^sub>i u'"
using assms proof (induction arbitrary: thesis t) (* FIXME clone of transform_correct_rt, maybe locale? *)
case (step u' u'')
obtain t' where "rs \<turnstile>\<^sub>n t \<longrightarrow>* t'" "t' \<approx>\<^sub>i u'"
using step by blast
obtain t'' where "rs \<turnstile>\<^sub>n t' \<longrightarrow>* t''" "t'' \<approx>\<^sub>i u''"
proof (rule compile_correct)
show "compile (consts_of rs) \<turnstile>\<^sub>i u' \<longrightarrow> u''" "t' \<approx>\<^sub>i u'"
by fact+
next
show "closed t'"
using \<open>rs \<turnstile>\<^sub>n t \<longrightarrow>* t'\<close> \<open>closed t\<close>
by (rule nrewrite_rt_closed)
qed blast
show ?case
proof (rule step.prems)
show "rs \<turnstile>\<^sub>n t \<longrightarrow>* t''"
using \<open>rs \<turnstile>\<^sub>n t \<longrightarrow>* t'\<close> \<open>rs \<turnstile>\<^sub>n t' \<longrightarrow>* t''\<close> by auto
qed fact
qed blast
subsubsection \<open>Completeness of translation\<close>
lemma (in nrules) compile_complete:
assumes "rs \<turnstile>\<^sub>n t \<longrightarrow> t'" "closed t"
shows "compile (consts_of rs) \<turnstile>\<^sub>i nterm_to_pterm t \<longrightarrow> nterm_to_pterm t'"
using assms proof induction
case (step r t t')
then obtain pat rhs' where "r = (pat, rhs')"
by force
then have "(pat, rhs') |\<in>| rs" "(pat, rhs') \<turnstile> t \<rightarrow> t'"
using step by blast+
then have "nrule (pat, rhs')"
using all_rules by blast
then obtain name pats where "(name, (pats, rhs')) = split_rule r" "pat = name $$ pats"
unfolding split_rule_def \<open>r = _\<close>
apply atomize_elim
by (auto simp: split_beta)
obtain crs where "(name, crs) |\<in>| consts_of rs" "(pats, rhs') |\<in>| crs"
using step \<open>_ = split_rule r\<close> \<open>r = _\<close>
by (metis consts_of_def fgroup_by_complete fst_conv snd_conv)
then obtain irs where "irs = translate_crules crs"
by blast
then have "(name, irs) |\<in>| compile (consts_of rs)"
unfolding compile_def
using \<open>(name, _) |\<in>| _\<close>
by (metis fimageI id_def map_prod_simp)
obtain rhs where "rhs = nterm_to_pterm rhs'" "(pats, rhs) |\<in>| irs"
using \<open>irs = _\<close> \<open>_ |\<in>| crs\<close>
unfolding translate_crules_def
by (metis fimageI id_def map_prod_simp)
from step obtain env' where "match pat t = Some env'" "t' = subst rhs' env'"
unfolding \<open>r = _\<close> using rewrite_step.simps
by force
then obtain env where "match pat (nterm_to_pterm t) = Some env" "irelated.P_env env' env"
by (metis irelated.match_rel option_rel_Some1)
then have "subst rhs env = nterm_to_pterm t'"
unfolding \<open>t' = _\<close>
apply -
apply (rule sym)
apply (rule irelated_subst)
unfolding \<open>rhs = _\<close>
by auto
have "name, pats, rhs \<turnstile>\<^sub>i nterm_to_pterm t \<rightarrow> nterm_to_pterm t'"
apply (rule irewrite_stepI)
using \<open>match _ _ = Some env\<close> unfolding \<open>pat = _\<close>
apply assumption
by fact
show ?case
by rule fact+
next
case (beta x t t')
obtain c where "c = (Free x, nterm_to_pterm t)"
by blast
from beta have "closed (nterm_to_pterm t')"
using closed_nterm_to_pterm[where t = t']
unfolding closed_except_def
by auto
show ?case
apply simp
apply rule
using \<open>c = _\<close>
by (fastforce intro: irelated_subst[THEN sym])+
next
case ("fun" t t' u)
show ?case
apply simp
apply rule
apply (rule "fun")
using "fun"
unfolding closed_except_def
apply simp
done
next
case (arg u u' t)
show ?case
apply simp
apply rule
apply (rule arg)
using arg
unfolding closed_except_def
by simp
qed
subsubsection \<open>Correctness of transformation\<close>
abbreviation irules_deferred_matches :: "pterm list \<Rightarrow> irules \<Rightarrow> (term \<times> pterm) fset" where
"irules_deferred_matches args \<equiv> fselect
(\<lambda>(pats, rhs). map_option (\<lambda>env. (last pats, subst rhs env)) (matchs (butlast pats) args))"
context irules begin
inductive prelated :: "pterm \<Rightarrow> pterm \<Rightarrow> bool" ("_ \<approx>\<^sub>p _" [0,50] 50) where
const: "Pconst x \<approx>\<^sub>p Pconst x" |
var: "Pvar x \<approx>\<^sub>p Pvar x" |
app: "t\<^sub>1 \<approx>\<^sub>p u\<^sub>1 \<Longrightarrow> t\<^sub>2 \<approx>\<^sub>p u\<^sub>2 \<Longrightarrow> t\<^sub>1 $\<^sub>p t\<^sub>2 \<approx>\<^sub>p u\<^sub>1 $\<^sub>p u\<^sub>2" |
pat: "rel_fset (rel_prod (=) prelated) cs\<^sub>1 cs\<^sub>2 \<Longrightarrow> Pabs cs\<^sub>1 \<approx>\<^sub>p Pabs cs\<^sub>2" |
"defer":
"(name, rsi) |\<in>| rs \<Longrightarrow> 0 < arity rsi \<Longrightarrow>
rel_fset (rel_prod (=) prelated) (irules_deferred_matches args rsi) cs \<Longrightarrow>
list_all closed args \<Longrightarrow>
name $$ args \<approx>\<^sub>p Pabs cs"
inductive_cases prelated_absE[consumes 1, case_names pat "defer"]: "t \<approx>\<^sub>p Pabs cs"
lemma prelated_refl[intro!]: "t \<approx>\<^sub>p t"
proof (induction t)
case Pabs
thus ?case
- by (auto simp: snds.simps fmember_iff_member_fset intro!: prelated.pat rel_fset_refl_strong rel_prod.intros)
+ by (auto simp: snds.simps intro!: prelated.pat rel_fset_refl_strong rel_prod.intros)
qed (auto intro: prelated.intros)
sublocale prelated: term_struct_rel prelated
by standard (auto simp: const_pterm_def app_pterm_def intro: prelated.intros elim: prelated.cases)
lemma prelated_pvars:
assumes "t \<approx>\<^sub>p u"
shows "frees t = frees u"
using assms proof (induction rule: prelated.induct)
case (pat cs\<^sub>1 cs\<^sub>2)
show ?case
apply simp
apply (rule arg_cong[where f = ffUnion])
apply (rule rel_fset_image_eq)
apply fact
apply auto
done
next
case ("defer" name rsi args cs)
{
fix pat t
assume "(pat, t) |\<in>| cs"
with "defer" obtain t'
where "(pat, t') |\<in>| irules_deferred_matches args rsi" "frees t = frees t'"
by (auto elim: rel_fsetE2)
then obtain pats rhs env
where "pat = last pats" "(pats, rhs) |\<in>| rsi"
and "matchs (butlast pats) args = Some env" "t' = subst rhs env"
by auto
have "closed_except rhs (freess pats)" "linears pats"
using \<open>(pats, rhs) |\<in>| rsi\<close> \<open>(name, rsi) |\<in>| rs\<close> inner by blast+
have "arity_compatibles rsi"
using "defer" inner by (blast dest: fpairwiseD)
have "length pats > 0"
by (subst arity_compatible_length) fact+
hence "pats = butlast pats @ [last pats]"
by simp
note \<open>frees t = frees t'\<close>
also have "frees t' = frees rhs - fmdom env"
unfolding \<open>t' = _\<close>
apply (rule subst_frees)
apply (rule closed.matchs)
apply fact+
done
also have "\<dots> = frees rhs - freess (butlast pats)"
using \<open>matchs _ _ = _\<close> by (metis matchs_dom)
also have "\<dots> |\<subseteq>| freess pats - freess (butlast pats)"
using \<open>closed_except _ _\<close>
by (auto simp: closed_except_def)
also have "\<dots> = frees (last pats) |-| freess (butlast pats)"
by (subst \<open>pats = _\<close>) (simp add: funion_fminus)
also have "\<dots> = frees (last pats)"
proof (rule fminus_triv)
have "fdisjnt (freess (butlast pats)) (freess [last pats])"
using \<open>linears pats\<close> \<open>pats = _\<close>
by (metis linears_appendD)
thus "frees (last pats) |\<inter>| freess (butlast pats) = {||}"
by (fastforce simp: fdisjnt_alt_def)
qed
also have "\<dots> = frees pat" unfolding \<open>pat = _\<close> ..
finally have "frees t |\<subseteq>| frees pat" .
}
hence "closed (Pabs cs)"
unfolding closed_except_simps
by (auto simp: closed_except_def)
moreover have "closed (name $$ args)"
unfolding closed_list_comb by fact
ultimately show ?case
unfolding closed_except_def by simp
qed auto
corollary prelated_closed: "t \<approx>\<^sub>p u \<Longrightarrow> closed t \<longleftrightarrow> closed u"
unfolding closed_except_def
by (auto simp: prelated_pvars)
lemma prelated_no_abs_right:
assumes "t \<approx>\<^sub>p u" "no_abs u"
shows "t = u"
using assms
apply (induction rule: prelated.induct)
apply auto
apply (fold app_pterm_def)
apply auto
done
corollary env_prelated_refl[intro!]: "prelated.P_env env env"
by (auto intro: fmap.rel_refl)
text \<open>
The following, more general statement does not hold:
@{prop "t \<approx>\<^sub>p u \<Longrightarrow> rel_option prelated.P_env (match x t) (match x u)"}
If @{text t} and @{text u} are related because of the @{thm [source=true] prelated.defer} rule,
they have completely different shapes.
Establishing @{prop "is_abs t \<longleftrightarrow> is_abs u"} as a precondition would rule out this case, but at
the same time be too restrictive.
Instead, we use @{thm prelated.related_match}.
\<close>
lemma prelated_subst:
assumes "t\<^sub>1 \<approx>\<^sub>p t\<^sub>2" "prelated.P_env env\<^sub>1 env\<^sub>2"
shows "subst t\<^sub>1 env\<^sub>1 \<approx>\<^sub>p subst t\<^sub>2 env\<^sub>2"
using assms proof (induction arbitrary: env\<^sub>1 env\<^sub>2 rule: prelated.induct)
case (var x)
thus ?case
proof (cases rule: fmrel_cases[where x = x])
case none
thus ?thesis
by (auto intro: prelated.var)
next
case (some t u)
thus ?thesis
by simp
qed
next
case (pat cs\<^sub>1 cs\<^sub>2)
let ?drop = "\<lambda>env. \<lambda>(pat::term). fmdrop_fset (frees pat) env"
from pat have "prelated.P_env (?drop env\<^sub>1 pat) (?drop env\<^sub>2 pat)" for pat
by blast
with pat show ?case
by (auto intro!: prelated.pat rel_fset_image)
next
case ("defer" name rsi args cs)
have "name $$ args \<approx>\<^sub>p Pabs cs"
apply (rule prelated.defer)
apply fact+
apply (rule fset.rel_mono_strong)
apply fact
apply force
apply fact
done
moreover have "closed (name $$ args)"
unfolding closed_list_comb by fact
ultimately have "closed (Pabs cs)"
by (metis prelated_closed)
let ?drop = "\<lambda>env. \<lambda>pat. fmdrop_fset (frees pat) env"
let ?f = "\<lambda>env. (\<lambda>(pat, rhs). (pat, subst rhs (?drop env pat)))"
have "name $$ args \<approx>\<^sub>p Pabs (?f env\<^sub>2 |`| cs)"
proof (rule prelated.defer)
show "(name, rsi) |\<in>| rs" "0 < arity rsi" "list_all closed args"
using "defer" by auto
next
{
fix pat\<^sub>1 rhs\<^sub>1
fix pat\<^sub>2 rhs\<^sub>2
assume "(pat\<^sub>2, rhs\<^sub>2) |\<in>| cs"
assume "pat\<^sub>1 = pat\<^sub>2" "rhs\<^sub>1 \<approx>\<^sub>p rhs\<^sub>2"
have "rhs\<^sub>1 \<approx>\<^sub>p subst rhs\<^sub>2 (fmdrop_fset (frees pat\<^sub>2) env\<^sub>2)"
by (subst subst_closed_pabs) fact+
}
hence "rel_fset (rel_prod (=) prelated) (id |`| irules_deferred_matches args rsi) (?f env\<^sub>2 |`| cs)"
by (force intro!: rel_fset_image[OF \<open>rel_fset _ _ _\<close>])
thus "rel_fset (rel_prod (=) prelated) (irules_deferred_matches args rsi) (?f env\<^sub>2 |`| cs)"
by simp
qed
moreover have "map (\<lambda>t. subst t env\<^sub>1) args = args"
apply (rule map_idI)
apply (rule subst_closed_id)
using "defer" by (simp add: list_all_iff)
ultimately show ?case
by (simp add: subst_list_comb)
qed (auto intro: prelated.intros)
lemma prelated_step:
assumes "name, pats, rhs \<turnstile>\<^sub>i u \<rightarrow> u'" "t \<approx>\<^sub>p u"
obtains t' where "name, pats, rhs \<turnstile>\<^sub>i t \<rightarrow> t'" "t' \<approx>\<^sub>p u'"
proof -
let ?lhs = "name $$ pats"
from assms obtain env where "match ?lhs u = Some env" "u' = subst rhs env"
unfolding irewrite_step_def by blast
then obtain env' where "match ?lhs t = Some env'" "prelated.P_env env' env"
using assms by (auto elim: prelated.related_match)
hence "subst rhs env' \<approx>\<^sub>p subst rhs env"
using assms by (auto intro: prelated_subst)
show thesis
proof
show "name, pats, rhs \<turnstile>\<^sub>i t \<rightarrow> subst rhs env'"
unfolding irewrite_step_def using \<open>match ?lhs t = Some env'\<close>
by simp
next
show "subst rhs env' \<approx>\<^sub>p u'"
unfolding \<open>u' = subst rhs env\<close>
by fact
qed
qed
(* FIXME write using relators *)
lemma prelated_beta: \<comment> \<open>same problem as @{thm [source=true] prelated.related_match}\<close>
assumes "(pat, rhs\<^sub>2) \<turnstile> t\<^sub>2 \<rightarrow> u\<^sub>2" "rhs\<^sub>1 \<approx>\<^sub>p rhs\<^sub>2" "t\<^sub>1 \<approx>\<^sub>p t\<^sub>2"
obtains u\<^sub>1 where "(pat, rhs\<^sub>1) \<turnstile> t\<^sub>1 \<rightarrow> u\<^sub>1" "u\<^sub>1 \<approx>\<^sub>p u\<^sub>2"
proof -
from assms obtain env\<^sub>2 where "match pat t\<^sub>2 = Some env\<^sub>2" "u\<^sub>2 = subst rhs\<^sub>2 env\<^sub>2"
by auto
with assms obtain env\<^sub>1 where "match pat t\<^sub>1 = Some env\<^sub>1" "prelated.P_env env\<^sub>1 env\<^sub>2"
by (auto elim: prelated.related_match)
with assms have "subst rhs\<^sub>1 env\<^sub>1 \<approx>\<^sub>p subst rhs\<^sub>2 env\<^sub>2"
by (auto intro: prelated_subst)
show thesis
proof
show "(pat, rhs\<^sub>1) \<turnstile> t\<^sub>1 \<rightarrow> subst rhs\<^sub>1 env\<^sub>1"
using \<open>match pat t\<^sub>1 = _\<close> by simp
next
show "subst rhs\<^sub>1 env\<^sub>1 \<approx>\<^sub>p u\<^sub>2"
unfolding \<open>u\<^sub>2 = _\<close> by fact
qed
qed
theorem transform_correct:
assumes "transform_irule_set rs \<turnstile>\<^sub>i u \<longrightarrow> u'" "t \<approx>\<^sub>p u" "closed t"
obtains t' where "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'" \<comment> \<open>zero or one step\<close> and "t' \<approx>\<^sub>p u'"
using assms(1-3) proof (induction arbitrary: t thesis rule: irewrite.induct)
case (beta c cs\<^sub>2 u\<^sub>2 x\<^sub>2')
obtain v u\<^sub>1 where "t = v $\<^sub>p u\<^sub>1" "v \<approx>\<^sub>p Pabs cs\<^sub>2" "u\<^sub>1 \<approx>\<^sub>p u\<^sub>2"
using \<open>t \<approx>\<^sub>p Pabs cs\<^sub>2 $\<^sub>p u\<^sub>2\<close> by cases
with beta have "closed u\<^sub>1"
by (simp add: closed_except_def)
obtain pat rhs\<^sub>2 where "c = (pat, rhs\<^sub>2)" by (cases c) auto
from \<open>v \<approx>\<^sub>p Pabs cs\<^sub>2\<close> show ?case
proof (cases rule: prelated_absE)
case (pat cs\<^sub>1)
with beta \<open>c = _\<close> obtain rhs\<^sub>1 where "(pat, rhs\<^sub>1) |\<in>| cs\<^sub>1" "rhs\<^sub>1 \<approx>\<^sub>p rhs\<^sub>2"
by (auto elim: rel_fsetE2)
with beta obtain x\<^sub>1' where "(pat, rhs\<^sub>1) \<turnstile> u\<^sub>1 \<rightarrow> x\<^sub>1'" "x\<^sub>1' \<approx>\<^sub>p x\<^sub>2'"
using \<open>u\<^sub>1 \<approx>\<^sub>p u\<^sub>2\<close> assms \<open>c = _\<close>
by (auto elim: prelated_beta simp del: rewrite_step.simps)
show ?thesis
proof (rule beta.prems)
show "rs \<turnstile>\<^sub>i t \<longrightarrow>* x\<^sub>1'"
unfolding \<open>t = _\<close> \<open>v = _\<close>
by (intro r_into_rtranclp irewrite.beta) fact+
next
show "x\<^sub>1' \<approx>\<^sub>p x\<^sub>2'"
by fact
qed
next
case ("defer" name rsi args)
with beta \<open>c = _\<close> obtain rhs\<^sub>1' where "(pat, rhs\<^sub>1') |\<in>| irules_deferred_matches args rsi" "rhs\<^sub>1' \<approx>\<^sub>p rhs\<^sub>2"
by (auto elim: rel_fsetE2)
then obtain env\<^sub>a rhs\<^sub>1 pats
where "matchs (butlast pats) args = Some env\<^sub>a" "pat = last pats" "rhs\<^sub>1' = subst rhs\<^sub>1 env\<^sub>a"
and "(pats, rhs\<^sub>1) |\<in>| rsi"
by auto
hence "linears pats"
using \<open>(name, rsi) |\<in>| rs\<close> inner unfolding irules_def by blast
have "arity_compatibles rsi"
using "defer" inner by (blast dest: fpairwiseD)
have "length pats > 0"
by (subst arity_compatible_length) fact+
hence "pats = butlast pats @ [pat]"
unfolding \<open>pat = _\<close> by simp
from beta \<open>c = _\<close> obtain env\<^sub>b where "match pat u\<^sub>2 = Some env\<^sub>b" "x\<^sub>2' = subst rhs\<^sub>2 env\<^sub>b"
by fastforce
with \<open>u\<^sub>1 \<approx>\<^sub>p u\<^sub>2\<close> obtain env\<^sub>b' where "match pat u\<^sub>1 = Some env\<^sub>b'" "prelated.P_env env\<^sub>b' env\<^sub>b"
by (metis prelated.related_match)
have "closed_env env\<^sub>a"
by (rule closed.matchs) fact+
have "closed_env env\<^sub>b'"
apply (rule closed.matchs[where pats = "[pat]" and ts = "[u\<^sub>1]"])
apply simp
apply fact
apply simp
apply fact
done
have "fmdom env\<^sub>a = freess (butlast pats)"
by (rule matchs_dom) fact
moreover have "fmdom env\<^sub>b' = frees pat"
by (rule match_dom) fact
moreover have "fdisjnt (freess (butlast pats)) (frees pat)"
using \<open>pats = _\<close> \<open>linears pats\<close>
by (metis freess_single linears_appendD(3))
ultimately have "fdisjnt (fmdom env\<^sub>a) (fmdom env\<^sub>b')"
by simp
show ?thesis
proof (rule beta.prems)
have "rs \<turnstile>\<^sub>i name $$ args $\<^sub>p u\<^sub>1 \<longrightarrow> subst rhs\<^sub>1' env\<^sub>b'"
proof (rule irewrite.step)
show "(name, rsi) |\<in>| rs" "(pats, rhs\<^sub>1) |\<in>| rsi"
by fact+
next
show "name, pats, rhs\<^sub>1 \<turnstile>\<^sub>i name $$ args $\<^sub>p u\<^sub>1 \<rightarrow> subst rhs\<^sub>1' env\<^sub>b'"
apply (rule irewrite_stepI)
apply (fold app_pterm_def)
apply (subst list_comb_snoc)
apply (subst matchs_match_list_comb)
apply (subst \<open>pats = _\<close>)
apply (rule matchs_appI)
apply fact
apply simp
apply fact
unfolding \<open>rhs\<^sub>1' = _\<close>
apply (rule subst_indep')
apply fact+
done
qed
thus "rs \<turnstile>\<^sub>i t \<longrightarrow>* subst rhs\<^sub>1' env\<^sub>b'"
unfolding \<open>t = _\<close> \<open>v = _\<close>
by (rule r_into_rtranclp)
next
show "subst rhs\<^sub>1' env\<^sub>b' \<approx>\<^sub>p x\<^sub>2'"
unfolding \<open>x\<^sub>2' = _\<close>
by (rule prelated_subst) fact+
qed
qed
next
case (step name rs\<^sub>2 pats rhs u u')
then obtain rs\<^sub>1 where "rs\<^sub>2 = transform_irules rs\<^sub>1" "(name, rs\<^sub>1) |\<in>| rs"
unfolding transform_irule_set_def by force
hence "arity_compatibles rs\<^sub>1"
using inner by (blast dest: fpairwiseD)
show ?case
proof (cases "arity rs\<^sub>1 = 0")
case True
hence "rs\<^sub>2 = rs\<^sub>1"
unfolding \<open>rs\<^sub>2 = _\<close> transform_irules_def by simp
with step have "(pats, rhs) |\<in>| rs\<^sub>1"
by simp
from step obtain t' where "name, pats, rhs \<turnstile>\<^sub>i t \<rightarrow> t'" "t' \<approx>\<^sub>p u'"
using assms
by (auto elim: prelated_step)
show ?thesis
proof (rule step.prems)
show "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'"
by (intro conjI exI r_into_rtranclp irewrite.step) fact+
qed fact
next
let ?f = "\<lambda>(pats, rhs). (butlast pats, last pats, rhs)"
let ?grp = "fgroup_by ?f rs\<^sub>1"
case False
hence "rs\<^sub>2 = map_prod id Pabs |`| ?grp"
unfolding \<open>rs\<^sub>2 = _\<close> transform_irules_def by simp
with step obtain cs where "rhs = Pabs cs" "(pats, cs) |\<in>| ?grp"
by force
from step obtain env\<^sub>2 where "match (name $$ pats) u = Some env\<^sub>2" "u' = subst rhs env\<^sub>2"
unfolding irewrite_step_def by auto
then obtain args\<^sub>2 where "u = name $$ args\<^sub>2" "matchs pats args\<^sub>2 = Some env\<^sub>2"
by (auto elim: match_list_combE)
with step obtain args\<^sub>1 where "t = name $$ args\<^sub>1" "list_all2 prelated args\<^sub>1 args\<^sub>2"
by (auto elim: prelated.list_combE)
then obtain env\<^sub>1 where "matchs pats args\<^sub>1 = Some env\<^sub>1" "prelated.P_env env\<^sub>1 env\<^sub>2"
using \<open>matchs pats args\<^sub>2 = _\<close> by (metis prelated.related_matchs)
hence "fmdom env\<^sub>1 = freess pats"
by (auto simp: matchs_dom)
obtain cs' where "u' = Pabs cs'"
unfolding \<open>u' = _\<close> \<open>rhs = _\<close> by auto
hence "cs' = (\<lambda>(pat, rhs). (pat, subst rhs (fmdrop_fset (frees pat) env\<^sub>2 ))) |`| cs"
using \<open>u' = subst rhs env\<^sub>2\<close> unfolding \<open>rhs = _\<close>
by simp
show ?thesis
proof (rule step.prems)
show "rs \<turnstile>\<^sub>i t \<longrightarrow>* t"
by (rule rtranclp.rtrancl_refl)
next
show "t \<approx>\<^sub>p u'"
unfolding \<open>u' = Pabs cs'\<close> \<open>t = _\<close>
proof (intro prelated.defer rel_fsetI; safe?)
show "(name, rs\<^sub>1) |\<in>| rs"
by fact
next
show "0 < arity rs\<^sub>1"
using False by simp
next
show "list_all closed args\<^sub>1"
using \<open>closed t\<close> unfolding \<open>t = _\<close> closed_list_comb .
next
fix pat rhs'
assume "(pat, rhs') |\<in>| irules_deferred_matches args\<^sub>1 rs\<^sub>1"
then obtain pats' rhs env
where "(pats', rhs) |\<in>| rs\<^sub>1"
and "matchs (butlast pats') args\<^sub>1 = Some env" "pat = last pats'" "rhs' = subst rhs env"
by auto
with False have "pats' \<noteq> []"
using \<open>arity_compatibles rs\<^sub>1\<close>
by (metis list.size(3) arity_compatible_length)
hence "butlast pats' @ [last pats'] = pats'"
by simp
from \<open>(pats, cs) |\<in>| ?grp\<close> obtain pats\<^sub>e rhs\<^sub>e
where "(pats\<^sub>e, rhs\<^sub>e) |\<in>| rs\<^sub>1" "pats = butlast pats\<^sub>e"
by (auto elim: fgroup_byE2)
have "patterns_compatible (butlast pats') pats"
unfolding \<open>pats = _\<close>
apply (rule rev_accum_rel_butlast)
using \<open>(pats', rhs) |\<in>| rs\<^sub>1\<close> \<open>(pats\<^sub>e, rhs\<^sub>e) |\<in>| rs\<^sub>1\<close> \<open>(name, rs\<^sub>1) |\<in>| rs\<close> inner
by (blast dest: fpairwiseD)
interpret irules': irules C_info "transform_irule_set rs" by (rule rules_transform)
have "butlast pats' = pats" "env = env\<^sub>1"
apply (rule matchs_compatible_eq)
subgoal by fact
subgoal
apply (rule linears_butlastI)
using \<open>(pats', rhs) |\<in>| rs\<^sub>1\<close> \<open>(name, rs\<^sub>1) |\<in>| rs\<close> inner by blast
subgoal
using \<open>(pats, _) |\<in>| rs\<^sub>2\<close> \<open>(name, rs\<^sub>2) |\<in>| transform_irule_set rs\<close>
using irules'.inner by blast
apply fact+
subgoal
apply (rule matchs_compatible_eq)
apply fact
apply (rule linears_butlastI)
using \<open>(pats', rhs) |\<in>| rs\<^sub>1\<close> \<open>(name, rs\<^sub>1) |\<in>| rs\<close> inner
apply blast
using \<open>(pats, _) |\<in>| rs\<^sub>2\<close> \<open>(name, rs\<^sub>2) |\<in>| transform_irule_set rs\<close>
using irules'.inner apply blast
by fact+
done
let ?rhs_subst = "\<lambda>env. subst rhs (fmdrop_fset (frees pat) env)"
have "fmdom env\<^sub>2 = freess pats"
using \<open>match (_ $$ _) _ = Some env\<^sub>2\<close>
by (simp add: match_dom)
show "fBex cs' (rel_prod (=) prelated (pat, rhs'))"
unfolding \<open>rhs' = _\<close>
proof (rule fBexI, rule rel_prod.intros)
have "fdisjnt (freess (butlast pats')) (frees (last pats'))"
apply (subst freess_single[symmetric])
apply (rule linears_appendD)
apply (subst \<open>butlast pats' @ [last pats'] = pats'\<close>)
using \<open>(pats', rhs) |\<in>| rs\<^sub>1\<close> \<open>(name, rs\<^sub>1) |\<in>| rs\<close> inner
by blast
show "subst rhs env \<approx>\<^sub>p ?rhs_subst env\<^sub>2"
apply (rule prelated_subst)
apply (rule prelated_refl)
unfolding fmfilter_alt_defs
apply (subst fmfilter_true)
subgoal premises prems for x y
using fmdomI[OF prems]
unfolding \<open>pat = _\<close> \<open>fmdom env\<^sub>2 = _\<close>
apply (subst (asm) \<open>butlast pats' = pats\<close>[symmetric])
using \<open>fdisjnt (freess (butlast pats')) (frees (last pats'))\<close>
by (auto simp: fdisjnt_alt_def)
subgoal
unfolding \<open>env = _\<close>
by fact
done
next
have "(pat, rhs) |\<in>| cs"
unfolding \<open>pat = _\<close>
apply (rule fgroup_byD[where a = "(x, y)" for x y])
apply fact
apply simp
apply (intro conjI)
apply fact
apply (rule refl)+
apply fact
done
thus "(pat, ?rhs_subst env\<^sub>2) |\<in>| cs'"
unfolding \<open>cs' = _\<close> by force
qed simp
next
fix pat rhs'
assume "(pat, rhs') |\<in>| cs'"
then obtain rhs
where "(pat, rhs) |\<in>| cs"
and "rhs' = subst rhs (fmdrop_fset (frees pat) env\<^sub>2 )"
unfolding \<open>cs' = _\<close> by auto
with \<open>(pats, cs) |\<in>| ?grp\<close> obtain pats'
where "(pats', rhs) |\<in>| rs\<^sub>1" "pats = butlast pats'" "pat = last pats'"
by auto
with False have "length pats' \<noteq> 0"
using \<open>arity_compatibles _\<close> by (metis arity_compatible_length)
hence "pats' = pats @ [pat]"
unfolding \<open>pats = _\<close> \<open>pat = _\<close> by auto
moreover have "linears pats'"
using \<open>(pats', rhs) |\<in>| rs\<^sub>1\<close> \<open>(name, rs\<^sub>1) |\<in>| _\<close> inner by blast
ultimately have "fdisjnt (fmdom env\<^sub>1) (frees pat)"
unfolding \<open>fmdom env\<^sub>1 = _\<close>
by (auto dest: linears_appendD)
let ?rhs_subst = "\<lambda>env. subst rhs (fmdrop_fset (frees pat) env)"
show "fBex (irules_deferred_matches args\<^sub>1 rs\<^sub>1) (\<lambda>e. rel_prod (=) prelated e (pat, rhs'))"
unfolding \<open>rhs' = _\<close>
proof (rule fBexI, rule rel_prod.intros)
show "?rhs_subst env\<^sub>1 \<approx>\<^sub>p ?rhs_subst env\<^sub>2"
using \<open>prelated.P_env env\<^sub>1 env\<^sub>2\<close> inner
by (auto intro: prelated_subst)
next
have "matchs (butlast pats') args\<^sub>1 = Some env\<^sub>1"
using \<open>matchs pats args\<^sub>1 = _\<close> \<open>pats = _\<close> by simp
moreover have "subst rhs env\<^sub>1 = ?rhs_subst env\<^sub>1"
apply (rule arg_cong[where f = "subst rhs"])
unfolding fmfilter_alt_defs
apply (rule fmfilter_true[symmetric])
using \<open>fdisjnt (fmdom env\<^sub>1) _\<close>
by (auto simp: fdisjnt_alt_def intro: fmdomI)
ultimately show "(pat, ?rhs_subst env\<^sub>1) |\<in>| irules_deferred_matches args\<^sub>1 rs\<^sub>1"
using \<open>(pats', rhs) |\<in>| rs\<^sub>1\<close> \<open>pat = last pats'\<close>
by auto
qed simp
qed
qed
qed
next
case ("fun" v v' u)
obtain w x where "t = w $\<^sub>p x" "w \<approx>\<^sub>p v" "x \<approx>\<^sub>p u" "closed w"
using \<open>t \<approx>\<^sub>p v $\<^sub>p u\<close> \<open>closed t\<close> by cases (auto simp: closed_except_def)
with "fun" obtain w' where "rs \<turnstile>\<^sub>i w \<longrightarrow>* w'" "w' \<approx>\<^sub>p v'"
by blast
show ?case
proof (rule fun.prems)
show "rs \<turnstile>\<^sub>i t \<longrightarrow>* w' $\<^sub>p x"
unfolding \<open>t = _\<close>
by (intro irewrite.rt_comb[unfolded app_pterm_def] rtranclp.rtrancl_refl) fact
next
show "w' $\<^sub>p x \<approx>\<^sub>p v' $\<^sub>p u"
by (rule prelated.app) fact+
qed
next
case (arg u u' v)
obtain w x where "t = w $\<^sub>p x" "w \<approx>\<^sub>p v" "x \<approx>\<^sub>p u" "closed x"
using \<open>t \<approx>\<^sub>p v $\<^sub>p u\<close> \<open>closed t\<close> by cases (auto simp: closed_except_def)
with arg obtain x' where "rs \<turnstile>\<^sub>i x \<longrightarrow>* x'" "x' \<approx>\<^sub>p u'"
by blast
show ?case
proof (rule arg.prems)
show "rs \<turnstile>\<^sub>i t \<longrightarrow>* w $\<^sub>p x'"
unfolding \<open>t = w $\<^sub>p x\<close>
by (intro irewrite.rt_comb[unfolded app_pterm_def] rtranclp.rtrancl_refl) fact
next
show "w $\<^sub>p x' \<approx>\<^sub>p v $\<^sub>p u'"
by (rule prelated.app) fact+
qed
qed
end
subsubsection \<open>Completeness of transformation\<close>
lemma (in irules) transform_completeness:
assumes "rs \<turnstile>\<^sub>i t \<longrightarrow> t'" "closed t"
shows "transform_irule_set rs \<turnstile>\<^sub>i t \<longrightarrow>* t'"
using assms proof induction
case (step name irs' pats' rhs' t t')
then obtain irs where "irs = transform_irules irs'" "(name, irs) |\<in>| transform_irule_set rs"
unfolding transform_irule_set_def
by (metis fimageI id_apply map_prod_simp)
show ?case
proof (cases "arity irs' = 0")
case True
hence "irs = irs'"
unfolding \<open>irs = _\<close>
unfolding transform_irules_def
by simp
with step have "(pats', rhs') |\<in>| irs" "name, pats', rhs' \<turnstile>\<^sub>i t \<rightarrow> t'"
by blast+
have "transform_irule_set rs \<turnstile>\<^sub>i t \<longrightarrow>* t'"
apply (rule r_into_rtranclp)
apply rule
by fact+
show ?thesis by fact
next
let ?f = "\<lambda>(pats, rhs). (butlast pats, last pats, rhs)"
let ?grp = "fgroup_by ?f irs'"
note closed_except_def [simp add]
case False
then have "irs = map_prod id Pabs |`| ?grp"
unfolding \<open>irs = _\<close>
unfolding transform_irules_def
by simp
with False have "irs = transform_irules irs'"
unfolding transform_irules_def
by simp
obtain pat pats where "pat = last pats'" "pats = butlast pats'"
by blast
from step False have "length pats' \<noteq> 0"
using arity_compatible_length inner
by (smt fBallE prod.simps(2))
then have "pats' = pats @ [pat]"
unfolding \<open>pat = _\<close> \<open>pats = _\<close>
by simp
from step have "linears pats'"
using inner fBallE
by (metis (mono_tags, lifting) old.prod.case)
then have "fdisjnt (freess pats) (frees pat)"
unfolding \<open>pats' = _\<close>
using linears_appendD(3) freess_single
by force
from step obtain cs where "(pats, cs) |\<in>| ?grp"
unfolding \<open>pats = _\<close>
by (metis (no_types, lifting) fgroup_by_complete fst_conv prod.simps(2))
with step have "(pat, rhs') |\<in>| cs"
unfolding \<open>pat = _\<close> \<open>pats = _\<close>
by (meson fgroup_byD old.prod.case)
have "(pats, Pabs cs) |\<in>| irs"
using \<open>irs = map_prod id Pabs |`| ?grp\<close> \<open>(pats, cs) |\<in>| _\<close>
by (metis (no_types, lifting) eq_snd_iff fst_conv fst_map_prod id_def rev_fimage_eqI snd_map_prod)
from step obtain env' where "match (name $$ pats') t = Some env'" "subst rhs' env' = t'"
using irewrite_step_def by auto
have "name $$ pats' = (name $$ pats) $ pat"
unfolding \<open>pats' = _\<close>
by (simp add: app_term_def)
then obtain t\<^sub>0 t\<^sub>1 env\<^sub>0 env\<^sub>1 where "t = t\<^sub>0 $\<^sub>p t\<^sub>1" "match (name $$ pats) t\<^sub>0 = Some env\<^sub>0" "match pat t\<^sub>1 = Some env\<^sub>1" "env' = env\<^sub>0 ++\<^sub>f env\<^sub>1"
using match_appE_split[OF \<open>match (name $$ pats') _ = _\<close>[unfolded \<open>name $$ pats' = _\<close>], unfolded app_pterm_def]
by blast
with step have "closed t\<^sub>0" "closed t\<^sub>1"
by auto
then have "closed_env env\<^sub>0" "closed_env env\<^sub>1"
using match_vars[OF \<open>match _ t\<^sub>0 = _\<close>] match_vars[OF \<open>match _ t\<^sub>1 = _\<close>]
unfolding closed_except_def
by auto
obtain t\<^sub>0' where "subst (Pabs cs) env\<^sub>0 = t\<^sub>0'"
by blast
then obtain cs' where "t\<^sub>0' = Pabs cs'" "cs' = ((\<lambda>(pat, rhs). (pat, subst rhs (fmdrop_fset (frees pat) env\<^sub>0))) |`| cs)"
using subst_pterm.simps(3) by blast
obtain rhs where "subst rhs' (fmdrop_fset (frees pat) env\<^sub>0) = rhs"
by blast
then have "(pat, rhs) |\<in>| cs'"
unfolding \<open>cs' = _\<close>
using \<open>_ |\<in>| cs\<close>
by (metis (mono_tags, lifting) old.prod.case rev_fimage_eqI)
have "env\<^sub>0 ++\<^sub>f env\<^sub>1 = (fmdrop_fset (frees pat) env\<^sub>0) ++\<^sub>f env\<^sub>1"
apply (subst fmadd_drop_left_dom[symmetric])
using \<open>match pat _ = _\<close> match_dom
by metis
have "fdisjnt (fmdom env\<^sub>0) (fmdom env\<^sub>1)"
using match_dom
using \<open>match pat _ = _\<close> \<open>match (name $$ pats) _ = _\<close>
using \<open>fdisjnt _ _\<close>
unfolding fdisjnt_alt_def
by (metis matchs_dom match_list_combE)
have "subst rhs env\<^sub>1 = t'"
unfolding \<open>_ = rhs\<close>[symmetric]
unfolding \<open>_ = t'\<close>[symmetric]
unfolding \<open>env' = _\<close>
unfolding \<open>env\<^sub>0 ++\<^sub>f _ = _\<close>
apply (subst subst_indep')
using \<open>closed_env env\<^sub>0\<close>
apply blast
using \<open>fdisjnt (fmdom _) _\<close>
unfolding fdisjnt_alt_def
by auto
show ?thesis
unfolding \<open>t = _\<close>
apply rule
apply (rule r_into_rtranclp)
apply (rule irewrite.intros(3))
apply rule
apply fact+
apply (rule irewrite_stepI)
apply fact+
unfolding \<open>t\<^sub>0' = _\<close>
apply rule
apply fact
using \<open>match pat t\<^sub>1 = _\<close> \<open>subst rhs _ = _\<close>
by force
qed
qed (auto intro: irewrite.rt_comb[unfolded app_pterm_def] intro!: irewrite.intros simp: closed_except_def)
subsubsection \<open>Computability\<close>
export_code
compile transform_irules
checking Scala SML
end
\ No newline at end of file
diff --git a/thys/CakeML_Codegen/Rewriting/Rewriting_Sterm.thy b/thys/CakeML_Codegen/Rewriting/Rewriting_Sterm.thy
--- a/thys/CakeML_Codegen/Rewriting/Rewriting_Sterm.thy
+++ b/thys/CakeML_Codegen/Rewriting/Rewriting_Sterm.thy
@@ -1,668 +1,666 @@
section \<open>Sequential pattern matching\<close>
theory Rewriting_Sterm
imports Rewriting_Pterm
begin
type_synonym srule = "name \<times> sterm"
abbreviation closed_srules :: "srule list \<Rightarrow> bool" where
"closed_srules \<equiv> list_all (closed \<circ> snd)"
primrec srule :: "srule \<Rightarrow> bool" where
"srule (_, rhs) \<longleftrightarrow> wellformed rhs \<and> closed rhs \<and> is_abs rhs"
lemma sruleI[intro!]: "wellformed rhs \<Longrightarrow> closed rhs \<Longrightarrow> is_abs rhs \<Longrightarrow> srule (name, rhs)"
by simp
locale srules = constants C_info "fst |`| fset_of_list rs" for C_info and rs :: "srule list" +
assumes all_rules: "list_all srule rs"
assumes distinct: "distinct (map fst rs)"
assumes not_shadows: "list_all (\<lambda>(_, rhs). \<not> shadows_consts rhs) rs"
assumes swelldefined_rs: "list_all (\<lambda>(_, rhs). welldefined rhs) rs"
begin
lemma map: "is_map (set rs)"
using distinct by (rule distinct_is_map)
lemma clausesE:
assumes "(name, rhs) \<in> set rs"
obtains cs where "rhs = Sabs cs"
proof -
from assms have "is_abs rhs"
using all_rules unfolding list_all_iff by auto
then obtain cs where "rhs = Sabs cs"
by (cases rhs) (auto simp: is_abs_def term_cases_def)
with that show thesis .
qed
end
subsubsection \<open>Rewriting\<close>
inductive srewrite_step where
cons_match: "srewrite_step ((name, rhs) # rest) name rhs" |
cons_nomatch: "name \<noteq> name' \<Longrightarrow> srewrite_step rs name rhs \<Longrightarrow> srewrite_step ((name', rhs') # rs) name rhs"
lemma srewrite_stepI0:
assumes "(name, rhs) \<in> set rs" "is_map (set rs)"
shows "srewrite_step rs name rhs"
using assms proof (induction rs)
case (Cons r rs)
then obtain name' rhs' where "r = (name', rhs')" by force
show ?case
proof (cases "name = name'")
case False
show ?thesis
unfolding \<open>r = _\<close>
apply (rule srewrite_step.cons_nomatch)
subgoal by fact
apply (rule Cons)
using False Cons(2) \<open>r = _\<close> apply force
using Cons(3) unfolding is_map_def by auto
next
case True
have "rhs = rhs'"
apply (rule is_mapD)
apply fact
unfolding \<open>r = _\<close>
using Cons(2) \<open>r = _\<close> apply simp
using True apply simp
done
show ?thesis
unfolding \<open>r = _\<close> \<open>name = _\<close> \<open>rhs = _\<close>
by (rule srewrite_step.cons_match)
qed
qed auto
lemma (in srules) srewrite_stepI: "(name, rhs) \<in> set rs \<Longrightarrow> srewrite_step rs name rhs"
using map
by (metis srewrite_stepI0)
hide_fact srewrite_stepI0
inductive srewrite :: "srule list \<Rightarrow> sterm \<Rightarrow> sterm \<Rightarrow> bool" ("_/ \<turnstile>\<^sub>s/ _ \<longrightarrow>/ _" [50,0,50] 50) for rs where
step: "srewrite_step rs name rhs \<Longrightarrow> rs \<turnstile>\<^sub>s Sconst name \<longrightarrow> rhs" |
beta: "rewrite_first cs t t' \<Longrightarrow> rs \<turnstile>\<^sub>s Sabs cs $\<^sub>s t \<longrightarrow> t'" |
"fun": "rs \<turnstile>\<^sub>s t \<longrightarrow> t' \<Longrightarrow> rs \<turnstile>\<^sub>s t $\<^sub>s u \<longrightarrow> t' $\<^sub>s u" |
arg: "rs \<turnstile>\<^sub>s u \<longrightarrow> u' \<Longrightarrow> rs \<turnstile>\<^sub>s t $\<^sub>s u \<longrightarrow> t $\<^sub>s u'"
code_pred srewrite .
abbreviation srewrite_rt :: "srule list \<Rightarrow> sterm \<Rightarrow> sterm \<Rightarrow> bool" ("_/ \<turnstile>\<^sub>s/ _ \<longrightarrow>*/ _" [50,0,50] 50) where
"srewrite_rt rs \<equiv> (srewrite rs)\<^sup>*\<^sup>*"
global_interpretation srewrite: rewriting "srewrite rs" for rs
by standard (auto intro: srewrite.intros simp: app_sterm_def)+
code_pred (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) srewrite_step .
code_pred (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) srewrite .
subsubsection \<open>Translation from @{typ pterm} to @{typ sterm}\<close>
text \<open>
In principle, any function of type @{typ \<open>('a \<times> 'b) fset \<Rightarrow> ('a \<times> 'b) list\<close>} that orders
by keys would do here. However, For simplicity's sake, we choose a fixed one
(@{const ordered_fmap}) here.
\<close>
primrec pterm_to_sterm :: "pterm \<Rightarrow> sterm" where
"pterm_to_sterm (Pconst name) = Sconst name" |
"pterm_to_sterm (Pvar name) = Svar name" |
"pterm_to_sterm (t $\<^sub>p u) = pterm_to_sterm t $\<^sub>s pterm_to_sterm u" |
"pterm_to_sterm (Pabs cs) = Sabs (ordered_fmap (map_prod id pterm_to_sterm |`| cs))"
lemma pterm_to_sterm:
assumes "no_abs t"
shows "pterm_to_sterm t = convert_term t"
using assms proof induction
case (free name)
show ?case
apply simp
apply (simp add: free_sterm_def free_pterm_def)
done
next
case (const name)
show ?case
apply simp
apply (simp add: const_sterm_def const_pterm_def)
done
next
case (app t\<^sub>1 t\<^sub>2)
then show ?case
apply simp
apply (simp add: app_sterm_def app_pterm_def)
done
qed
text \<open>
@{const sterm_to_pterm} has to be defined, for technical reasons, in
@{theory CakeML_Codegen.Pterm}.
\<close>
lemma pterm_to_sterm_wellformed:
assumes "wellformed t"
shows "wellformed (pterm_to_sterm t)"
using assms proof (induction t rule: pterm_induct)
case (Pabs cs)
show ?case
apply simp
unfolding map_prod_def id_apply
apply (intro conjI)
subgoal
apply (subst list_all_iff_fset)
apply (subst ordered_fmap_set_eq)
apply (rule is_fmap_image)
using Pabs apply simp
apply (rule fBallI)
apply (erule fimageE)
apply auto[]
using Pabs(2) apply auto[]
apply (rule Pabs)
using Pabs(2) by auto
subgoal
apply (rule ordered_fmap_distinct)
apply (rule is_fmap_image)
using Pabs(2) by simp
subgoal
apply (subgoal_tac "cs \<noteq> {||}")
including fset.lifting apply transfer
unfolding ordered_map_def
using Pabs(2) by auto
done
qed auto
lemma pterm_to_sterm_sterm_to_pterm:
assumes "wellformed t"
shows "sterm_to_pterm (pterm_to_sterm t) = t"
using assms proof (induction t)
case (Pabs cs)
note fset_of_list_map[simp del]
show ?case
apply simp
unfolding map_prod_def id_apply
apply (subst ordered_fmap_image)
subgoal
apply (rule is_fmap_image)
using Pabs by simp
apply (subst ordered_fmap_set_eq)
subgoal
apply (rule is_fmap_image)
apply (rule is_fmap_image)
using Pabs by simp
subgoal
apply (subst fset.map_comp)
apply (subst map_prod_def[symmetric])+
unfolding o_def
apply (subst prod.map_comp)
apply (subst id_def[symmetric])+
apply simp
apply (subst map_prod_def)
unfolding id_def
apply (rule fset_map_snd_id)
apply simp
apply (rule Pabs)
- using Pabs(2) by (auto simp: fmember_iff_member_fset snds.simps)
+ using Pabs(2) by (auto simp: snds.simps)
done
qed auto
corollary pterm_to_sterm_frees: "wellformed t \<Longrightarrow> frees (pterm_to_sterm t) = frees t"
by (metis pterm_to_sterm_sterm_to_pterm sterm_to_pterm_frees)
corollary pterm_to_sterm_closed:
"closed_except t S \<Longrightarrow> wellformed t \<Longrightarrow> closed_except (pterm_to_sterm t) S"
unfolding closed_except_def
by (simp add: pterm_to_sterm_frees)
corollary pterm_to_sterm_consts: "wellformed t \<Longrightarrow> consts (pterm_to_sterm t) = consts t"
by (metis pterm_to_sterm_sterm_to_pterm sterm_to_pterm_consts)
corollary (in constants) pterm_to_sterm_shadows:
"wellformed t \<Longrightarrow> shadows_consts t \<longleftrightarrow> shadows_consts (pterm_to_sterm t)"
unfolding shadows_consts_def
by (metis pterm_to_sterm_sterm_to_pterm sterm_to_pterm_all_frees)
definition compile :: "prule fset \<Rightarrow> srule list" where
"compile rs = ordered_fmap (map_prod id pterm_to_sterm |`| rs)"
subsubsection \<open>Correctness of translation\<close>
context prules begin
lemma compile_heads: "fst |`| fset_of_list (compile rs) = fst |`| rs"
unfolding compile_def
apply (subst ordered_fmap_set_eq)
apply (subst map_prod_def, subst id_apply)
apply (rule is_fmap_image)
apply (rule fmap)
apply simp
done
lemma compile_rules: "srules C_info (compile rs)"
proof
show "list_all srule (compile rs)"
using fmap all_rules
unfolding compile_def list_all_iff
including fset.lifting
apply transfer
apply (subst ordered_map_set_eq)
subgoal by simp
subgoal
unfolding map_prod_def id_def
by (erule is_map_image)
subgoal
apply (rule ballI)
apply safe
subgoal
apply (rule pterm_to_sterm_wellformed)
apply fastforce
done
subgoal
apply (rule pterm_to_sterm_closed)
apply fastforce
apply fastforce
done
subgoal for _ _ a b
apply (erule ballE[where x = "(a, b)"])
apply (cases b; auto)
apply (auto simp: is_abs_def term_cases_def)
done
done
done
next
show "distinct (map fst (compile rs))"
unfolding compile_def
apply (rule ordered_fmap_distinct)
unfolding map_prod_def id_def
apply (rule is_fmap_image)
apply (rule fmap)
done
next
have "list_all (\<lambda>(_, rhs). welldefined rhs) (compile rs)"
unfolding compile_def
apply (subst ordered_fmap_list_all)
subgoal
apply (subst map_prod_def)
apply (subst id_apply)
apply (rule is_fmap_image)
by (fact fmap)
apply simp
apply (rule fBallI)
subgoal for x
apply (cases x, simp)
apply (subst pterm_to_sterm_consts)
using all_rules apply force
using welldefined_rs by force
done
thus "list_all (\<lambda>(_, rhs). consts rhs |\<subseteq>| pre_constants.all_consts C_info (fst |`| fset_of_list (compile rs))) (compile rs)"
by (simp add: compile_heads)
next
interpret c: constants _ "fset_of_list (map fst (compile rs))"
by (simp add: constants_axioms compile_heads)
have all_consts: "c.all_consts = all_consts"
by (simp add: compile_heads)
note fset_of_list_map[simp del]
have "list_all (\<lambda>(_, rhs). \<not> shadows_consts rhs) (compile rs)"
unfolding compile_def
apply (subst list_all_iff_fset)
apply (subst ordered_fmap_set_eq)
apply (subst map_prod_def)
unfolding id_apply
apply (rule is_fmap_image)
apply (fact fmap)
apply simp
apply (rule fBall_pred_weaken[where P = "\<lambda>(_, rhs). \<not> shadows_consts rhs"])
subgoal for x
apply (cases x, simp)
apply (subst (asm) pterm_to_sterm_shadows)
using all_rules apply force
by simp
subgoal
using not_shadows by force
done
thus "list_all (\<lambda>(_, rhs). \<not> pre_constants.shadows_consts C_info (fst |`| fset_of_list (compile rs)) rhs) (compile rs)"
unfolding compile_heads all_consts .
next
show "fdisjnt (fst |`| fset_of_list (compile rs)) C"
unfolding compile_def
apply (subst fset_of_list_map[symmetric])
apply (subst ordered_fmap_keys)
apply (subst map_prod_def)
apply (subst id_apply)
apply (rule is_fmap_image)
using fmap disjnt by auto
next
show "distinct all_constructors"
by (fact distinct_ctr)
qed
sublocale prules_as_srules: srules C_info "compile rs"
by (fact compile_rules)
end
global_interpretation srelated: term_struct_rel_strong "(\<lambda>p s. p = sterm_to_pterm s)"
proof (standard, goal_cases)
case (5 name t)
then show ?case by (cases t) (auto simp: const_sterm_def const_pterm_def split: option.splits)
next
case (6 u\<^sub>1 u\<^sub>2 t)
then show ?case by (cases t) (auto simp: app_sterm_def app_pterm_def split: option.splits)
qed (auto simp: const_sterm_def const_pterm_def app_sterm_def app_pterm_def)
lemma srelated_subst:
assumes "srelated.P_env penv senv"
shows "subst (sterm_to_pterm t) penv = sterm_to_pterm (subst t senv)"
using assms
proof (induction t arbitrary: penv senv)
case (Svar name)
thus ?case
by (cases rule: fmrel_cases[where x = name]) auto
next
case (Sabs cs)
show ?case
apply simp
including fset.lifting
apply (transfer' fixing: cs penv senv)
unfolding set_map image_comp
apply (rule image_cong[OF refl])
unfolding comp_apply
apply (case_tac x)
apply hypsubst_thin
apply simp
apply (rule Sabs)
apply assumption
apply (simp add: snds.simps)
apply rule
apply (rule Sabs)
done
qed auto
context begin
private lemma srewrite_step_non_empty: "srewrite_step rs' name rhs \<Longrightarrow> rs' \<noteq> []"
by (induct rule: srewrite_step.induct) auto
private lemma compile_consE:
assumes "(name, rhs') # rest = compile rs" "is_fmap rs"
obtains rhs where "rhs' = pterm_to_sterm rhs" "(name, rhs) |\<in>| rs" "rest = compile (rs - {| (name, rhs) |})"
proof -
from assms have "ordered_fmap (map_prod id pterm_to_sterm |`| rs) = (name, rhs') # rest"
unfolding compile_def
by simp
hence "(name, rhs') \<in> set (ordered_fmap (map_prod id pterm_to_sterm |`| rs))"
by simp
have "(name, rhs') |\<in>| map_prod id pterm_to_sterm |`| rs"
apply (rule ordered_fmap_sound)
subgoal
unfolding map_prod_def id_apply
apply (rule is_fmap_image)
apply fact
done
subgoal by fact
done
then obtain rhs where "rhs' = pterm_to_sterm rhs" "(name, rhs) |\<in>| rs"
by auto
have "rest = compile (rs - {| (name, rhs) |})"
unfolding compile_def
apply (subst inj_on_fimage_set_diff[where C = rs])
subgoal
apply (rule inj_onI)
apply safe
apply auto
- apply (subst (asm) fmember_iff_member_fset[symmetric])+
using \<open>is_fmap rs\<close> by (blast dest: is_fmapD)
subgoal by simp
subgoal using \<open>(name, rhs) |\<in>| rs\<close> by simp
subgoal
apply simp
apply (subst ordered_fmap_remove)
apply (subst map_prod_def)
unfolding id_apply
apply (rule is_fmap_image)
apply fact
using \<open>(name, rhs) |\<in>| rs\<close> apply force
apply (subst \<open>rhs' = pterm_to_sterm rhs\<close>[symmetric])
apply (subst \<open>ordered_fmap _ = _\<close>[unfolded id_def])
by simp
done
show thesis
by (rule that) fact+
qed
private lemma compile_correct_step:
assumes "srewrite_step (compile rs) name rhs" "is_fmap rs" "fBall rs prule"
shows "(name, sterm_to_pterm rhs) |\<in>| rs"
using assms proof (induction "compile rs" name rhs arbitrary: rs)
case (cons_match name rhs' rest)
then obtain rhs where "rhs' = pterm_to_sterm rhs" "(name, rhs) |\<in>| rs"
by (auto elim: compile_consE)
show ?case
unfolding \<open>rhs' = _\<close>
apply (subst pterm_to_sterm_sterm_to_pterm)
using fbspec[OF \<open>fBall rs prule\<close> \<open>(name, rhs) |\<in>| rs\<close>] apply force
by fact
next
case (cons_nomatch name name\<^sub>1 rest rhs rhs\<^sub>1')
then obtain rhs\<^sub>1 where "rhs\<^sub>1' = pterm_to_sterm rhs\<^sub>1" "(name\<^sub>1, rhs\<^sub>1) |\<in>| rs" "rest = compile (rs - {| (name\<^sub>1, rhs\<^sub>1) |})"
by (auto elim: compile_consE)
let ?rs' = "rs - {| (name\<^sub>1, rhs\<^sub>1) |}"
have "(name, sterm_to_pterm rhs) |\<in>| ?rs'"
proof (intro cons_nomatch)
show "rest = compile ?rs'"
by fact
show "is_fmap (rs |-| {|(name\<^sub>1, rhs\<^sub>1)|})"
using \<open>is_fmap rs\<close>
by (rule is_fmap_subset) auto
show "fBall ?rs' prule"
using cons_nomatch by blast
qed
thus ?case
by simp
qed
lemma compile_correct0:
assumes "compile rs \<turnstile>\<^sub>s u \<longrightarrow> u'" "prules C rs"
shows "rs \<turnstile>\<^sub>p sterm_to_pterm u \<longrightarrow> sterm_to_pterm u'"
using assms proof induction
case (beta cs t t')
then obtain pat rhs env where "(pat, rhs) \<in> set cs" "match pat t = Some env" "t' = subst rhs env"
by (auto elim: rewrite_firstE)
then obtain env' where "match pat (sterm_to_pterm t) = Some env'" "srelated.P_env env' env"
by (metis option.distinct(1) option.inject option.rel_cases srelated.match_rel)
hence "subst (sterm_to_pterm rhs) env' = sterm_to_pterm (subst rhs env)"
by (simp add: srelated_subst)
let ?rhs' = "sterm_to_pterm rhs"
have "(pat, ?rhs') |\<in>| fset_of_list (map (map_prod id sterm_to_pterm) cs)"
using \<open>(pat, rhs) \<in> set cs\<close>
including fset.lifting
by transfer' force
note fset_of_list_map[simp del]
show ?case
apply simp
apply (rule prewrite.intros)
apply fact
unfolding rewrite_step.simps
apply (subst map_option_eq_Some)
apply (intro exI conjI)
apply fact
unfolding \<open>t' = _\<close>
by fact
next
case (step name rhs)
hence "(name, sterm_to_pterm rhs) |\<in>| rs"
unfolding prules_def prules_axioms_def
by (metis compile_correct_step)
thus ?case
by (auto intro: prewrite.intros)
qed (auto intro: prewrite.intros)
end
lemma (in prules) compile_correct:
assumes "compile rs \<turnstile>\<^sub>s u \<longrightarrow> u'"
shows "rs \<turnstile>\<^sub>p sterm_to_pterm u \<longrightarrow> sterm_to_pterm u'"
by (rule compile_correct0) (fact | standard)+
hide_fact compile_correct0
subsubsection \<open>Completeness of translation\<close>
global_interpretation srelated': term_struct_rel_strong "(\<lambda>p s. pterm_to_sterm p = s)"
proof (standard, goal_cases)
case (1 t name)
then show ?case by (cases t) (auto simp: const_sterm_def const_pterm_def split: option.splits)
next
case (3 t u\<^sub>1 u\<^sub>2)
then show ?case by (cases t) (auto simp: app_sterm_def app_pterm_def split: option.splits)
qed (auto simp: const_sterm_def const_pterm_def app_sterm_def app_pterm_def)
corollary srelated_env_unique:
"srelated'.P_env penv senv \<Longrightarrow> srelated'.P_env penv senv' \<Longrightarrow> senv = senv'"
apply (subst (asm) fmrel_iff)+
apply (subst (asm) option.rel_sel)+
apply (rule fmap_ext)
by (metis option.exhaust_sel)
lemma srelated_subst':
assumes "srelated'.P_env penv senv" "wellformed t"
shows "pterm_to_sterm (subst t penv) = subst (pterm_to_sterm t) senv"
using assms proof (induction t arbitrary: penv senv)
case (Pvar name)
thus ?case
by (cases rule: fmrel_cases[where x = name]) auto
next
case (Pabs cs)
hence "is_fmap cs"
by force
show ?case
apply simp
unfolding map_prod_def id_apply
apply (subst ordered_fmap_image[symmetric])
apply fact
apply (subst fset.map_comp[symmetric])
apply (subst ordered_fmap_image[symmetric])
subgoal by (rule is_fmap_image) fact
apply (subst ordered_fmap_image[symmetric])
apply fact
apply auto
apply (drule ordered_fmap_sound[OF \<open>is_fmap cs\<close>])
subgoal for pat rhs
apply (rule Pabs)
- apply (subst (asm) fmember_iff_member_fset)
apply assumption
apply auto
using Pabs by force+
done
qed auto
lemma srelated_find_match:
assumes "find_match cs t = Some (penv, pat, rhs)" "srelated'.P_env penv senv"
shows "find_match (map (map_prod id pterm_to_sterm) cs) (pterm_to_sterm t) = Some (senv, pat, pterm_to_sterm rhs)"
proof -
let ?cs' = "map (map_prod id pterm_to_sterm) cs"
let ?t' = "pterm_to_sterm t"
have *: "list_all2 (rel_prod (=) (\<lambda>p s. pterm_to_sterm p = s)) cs ?cs'"
unfolding list.rel_map
by (auto intro: list.rel_refl)
obtain senv0
where "find_match ?cs' ?t' = Some (senv0, pat, pterm_to_sterm rhs)" "srelated'.P_env penv senv0"
using srelated'.find_match_rel[OF * refl, where t = t, unfolded assms]
unfolding option_rel_Some1 rel_prod_conv
by auto
with assms have "senv = senv0"
by (metis srelated_env_unique)
show ?thesis
unfolding \<open>senv = _\<close> by fact
qed
lemma (in prules) compile_complete:
assumes "rs \<turnstile>\<^sub>p t \<longrightarrow> t'" "wellformed t"
shows "compile rs \<turnstile>\<^sub>s pterm_to_sterm t \<longrightarrow> pterm_to_sterm t'"
using assms proof induction
case (step name rhs)
then show ?case
apply simp
apply rule
apply (rule prules_as_srules.srewrite_stepI)
unfolding compile_def
apply (subst fset_of_list_elem[symmetric])
apply (subst ordered_fmap_set_eq)
apply (insert fmap)
apply (rule is_fmapI)
apply (force dest: is_fmapD)
- by (simp add: rev_fimage_eqI)
+ by (metis fimage_eqI id_def map_prod_simp)
next
case (beta c cs t t')
from beta obtain pat rhs penv where "c = (pat, rhs)" "match pat t = Some penv" "subst rhs penv = t'"
by (metis (no_types, lifting) map_option_eq_Some rewrite_step.simps surj_pair)
then obtain senv where "match pat (pterm_to_sterm t) = Some senv" "srelated'.P_env penv senv"
by (metis option_rel_Some1 srelated'.match_rel)
have "wellformed rhs"
using beta \<open>c = _\<close> prules.all_rules prule.simps
by force
then have "subst (pterm_to_sterm rhs) senv = pterm_to_sterm t'"
using srelated_subst' \<open>_ = t'\<close> \<open>srelated'.P_env _ _\<close>
by metis
have "(pat, pterm_to_sterm rhs) |\<in>| map_prod id pterm_to_sterm |`| cs"
using beta \<open>c = _\<close>
by (metis fimage_eqI id_def map_prod_simp)
have "is_fmap cs"
using beta
by auto
have "find_match (ordered_fmap cs) t = Some (penv, pat, rhs)"
apply (rule compatible_find_match)
subgoal
apply (subst ordered_fmap_set_eq[OF \<open>is_fmap cs\<close>])+
using beta by simp
subgoal
unfolding list_all_iff
apply rule
apply (rename_tac x, case_tac x)
apply simp
apply (drule ordered_fmap_sound[OF \<open>is_fmap cs\<close>])
using beta by auto
subgoal
apply (subst ordered_fmap_set_eq)
by fact
subgoal
by fact
subgoal
using beta(1) \<open>c = _\<close> \<open>is_fmap cs\<close>
using fset_of_list_elem ordered_fmap_set_eq by fast
done
show ?case
apply simp
apply rule
apply (subst \<open>_ = pterm_to_sterm t'\<close>[symmetric])
apply (rule find_match_rewrite_first)
unfolding map_prod_def id_apply
apply (subst ordered_fmap_image[symmetric])
apply fact
apply (subst map_prod_def[symmetric])
apply (subst id_def[symmetric])
apply (rule srelated_find_match)
by fact+
qed (auto intro: srewrite.intros)
subsubsection \<open>Computability\<close>
export_code compile
checking Scala
end
\ No newline at end of file
diff --git a/thys/CakeML_Codegen/Terms/Pterm.thy b/thys/CakeML_Codegen/Terms/Pterm.thy
--- a/thys/CakeML_Codegen/Terms/Pterm.thy
+++ b/thys/CakeML_Codegen/Terms/Pterm.thy
@@ -1,474 +1,472 @@
section \<open>Terms with explicit pattern matching\<close>
theory Pterm
imports
"../Utils/Compiler_Utils"
Consts
Sterm \<comment> \<open>Inclusion of this theory might seem a bit strange. Indeed, it is only for technical
reasons: to allow for a \<^theory_text>\<open>quickcheck\<close> setup.\<close>
begin
datatype pterm =
Pconst name |
Pvar name |
Pabs "(term \<times> pterm) fset" |
Papp pterm pterm (infixl "$\<^sub>p" 70)
primrec sterm_to_pterm :: "sterm \<Rightarrow> pterm" where
"sterm_to_pterm (Sconst name) = Pconst name" |
"sterm_to_pterm (Svar name) = Pvar name" |
"sterm_to_pterm (t $\<^sub>s u) = sterm_to_pterm t $\<^sub>p sterm_to_pterm u" |
"sterm_to_pterm (Sabs cs) = Pabs (fset_of_list (map (map_prod id sterm_to_pterm) cs))"
quickcheck_generator pterm
\<comment> \<open>will print some fishy ``constructor'' names, but at least it works\<close>
constructors: sterm_to_pterm
lemma sterm_to_pterm_total:
obtains t' where "t = sterm_to_pterm t'"
proof (induction t arbitrary: thesis)
case (Pconst x)
then show ?case
by (metis sterm_to_pterm.simps)
next
case (Pvar x)
then show ?case
by (metis sterm_to_pterm.simps)
next
case (Pabs cs)
from Pabs.IH obtain cs' where "cs = fset_of_list (map (map_prod id sterm_to_pterm) cs')"
apply atomize_elim
proof (induction cs)
case empty
show ?case
apply (rule exI[where x = "[]"])
by simp
next
case (insert c cs)
obtain pat rhs where "c = (pat, rhs)" by (cases c) auto
have "\<exists>cs'. cs = fset_of_list (map (map_prod id sterm_to_pterm) cs')"
apply (rule insert)
using insert.prems unfolding finsert.rep_eq
by blast
then obtain cs' where "cs = fset_of_list (map (map_prod id sterm_to_pterm) cs')"
by blast
obtain rhs' where "rhs = sterm_to_pterm rhs'"
apply (rule insert.prems[of "(pat, rhs)" rhs])
unfolding \<open>c = _\<close> by simp+
show ?case
apply (rule exI[where x = "(pat, rhs') # cs'"])
unfolding \<open>c = _\<close> \<open>cs = _\<close> \<open>rhs = _\<close>
by (simp add: id_def)
qed
hence "Pabs cs = sterm_to_pterm (Sabs cs')"
by simp
then show ?case
using Pabs by metis
next
case (Papp t1 t2)
then obtain t1' t2' where "t1 = sterm_to_pterm t1'" "t2 = sterm_to_pterm t2'"
by metis
then have "t1 $\<^sub>p t2 = sterm_to_pterm (t1' $\<^sub>s t2')"
by simp
with Papp show ?case
by metis
qed
lemma pterm_induct[case_names Pconst Pvar Pabs Papp]:
assumes "\<And>x. P (Pconst x)"
assumes "\<And>x. P (Pvar x)"
assumes "\<And>cs. (\<And>pat t. (pat, t) |\<in>| cs \<Longrightarrow> P t) \<Longrightarrow> P (Pabs cs)"
assumes "\<And>t u. P t \<Longrightarrow> P u \<Longrightarrow> P (t $\<^sub>p u)"
shows "P t"
proof (rule pterm.induct, goal_cases)
case (3 cs)
show ?case
apply (rule assms)
using 3
- apply (subst (asm) fmember_iff_member_fset[symmetric])
by auto
qed fact+
instantiation pterm :: pre_term begin
definition app_pterm where
"app_pterm t u = t $\<^sub>p u"
fun unapp_pterm where
"unapp_pterm (t $\<^sub>p u) = Some (t, u)" |
"unapp_pterm _ = None"
definition const_pterm where
"const_pterm = Pconst"
fun unconst_pterm where
"unconst_pterm (Pconst name) = Some name" |
"unconst_pterm _ = None"
definition free_pterm where
"free_pterm = Pvar"
fun unfree_pterm where
"unfree_pterm (Pvar name) = Some name" |
"unfree_pterm _ = None"
function (sequential) subst_pterm where
"subst_pterm (Pvar s) env = (case fmlookup env s of Some t \<Rightarrow> t | None \<Rightarrow> Pvar s)" |
"subst_pterm (t\<^sub>1 $\<^sub>p t\<^sub>2) env = subst_pterm t\<^sub>1 env $\<^sub>p subst_pterm t\<^sub>2 env" |
"subst_pterm (Pabs cs) env = Pabs ((\<lambda>(pat, rhs). (pat, subst_pterm rhs (fmdrop_fset (frees pat) env))) |`| cs)" |
"subst_pterm t _ = t"
by pat_completeness auto
termination
proof (relation "measure (size \<circ> fst)", goal_cases)
case 4
then show ?case
apply auto
including fset.lifting apply transfer
apply (rule le_imp_less_Suc)
apply (rule sum_nat_le_single[where y = "(a, (b, size b))" for a b])
by auto
qed auto
primrec consts_pterm :: "pterm \<Rightarrow> name fset" where
"consts_pterm (Pconst x) = {|x|}" |
"consts_pterm (t\<^sub>1 $\<^sub>p t\<^sub>2) = consts_pterm t\<^sub>1 |\<union>| consts_pterm t\<^sub>2" |
"consts_pterm (Pabs cs) = ffUnion (snd |`| map_prod id consts_pterm |`| cs)" |
"consts_pterm (Pvar _) = {||}"
primrec frees_pterm :: "pterm \<Rightarrow> name fset" where
"frees_pterm (Pvar x) = {|x|}" |
"frees_pterm (t\<^sub>1 $\<^sub>p t\<^sub>2) = frees_pterm t\<^sub>1 |\<union>| frees_pterm t\<^sub>2" |
"frees_pterm (Pabs cs) = ffUnion ((\<lambda>(pv, tv). tv - frees pv) |`| map_prod id frees_pterm |`| cs)" |
"frees_pterm (Pconst _) = {||}"
instance
by standard
(auto
simp: app_pterm_def const_pterm_def free_pterm_def
elim: unapp_pterm.elims unconst_pterm.elims unfree_pterm.elims
split: option.splits)
end
corollary subst_pabs_id:
assumes "\<And>pat rhs. (pat, rhs) |\<in>| cs \<Longrightarrow> subst rhs (fmdrop_fset (frees pat) env) = rhs"
shows "subst (Pabs cs) env = Pabs cs"
apply (subst subst_pterm.simps)
apply (rule arg_cong[where f = Pabs])
apply (rule fset_map_snd_id)
apply (rule assms)
-apply (subst (asm) fmember_iff_member_fset[symmetric])
apply assumption
done
corollary frees_pabs_alt_def:
"frees (Pabs cs) = ffUnion ((\<lambda>(pat, rhs). frees rhs - frees pat) |`| cs)"
apply simp
apply (rule arg_cong[where f = ffUnion])
apply (rule fset.map_cong[OF refl])
by auto
lemma sterm_to_pterm_frees[simp]: "frees (sterm_to_pterm t) = frees t"
proof (induction t)
case (Sabs cs)
show ?case
apply simp
apply (rule arg_cong[where f = ffUnion])
apply (rule fimage_cong[OF refl])
apply clarsimp
apply (subst Sabs)
by (auto simp: fset_of_list_elem snds.simps)
qed auto
lemma sterm_to_pterm_consts[simp]: "consts (sterm_to_pterm t) = consts t"
proof (induction t)
case (Sabs cs)
show ?case
apply simp
apply (rule arg_cong[where f = ffUnion])
apply (rule fimage_cong[OF refl])
apply clarsimp
apply (subst Sabs)
by (auto simp: fset_of_list_elem snds.simps)
qed auto
lemma subst_sterm_to_pterm:
"subst (sterm_to_pterm t) (fmmap sterm_to_pterm env) = sterm_to_pterm (subst t env)"
proof (induction t arbitrary: env rule: sterm_induct)
case (Sabs cs)
show ?case
apply simp
apply (rule fset.map_cong0)
apply (auto split: prod.splits)
apply (rule Sabs)
by (auto simp: fset_of_list.rep_eq)
qed (auto split: option.splits)
instantiation pterm :: "term" begin
definition abs_pred_pterm :: "(pterm \<Rightarrow> bool) \<Rightarrow> pterm \<Rightarrow> bool" where
[code del]: "abs_pred P t \<longleftrightarrow> (\<forall>cs. t = Pabs cs \<longrightarrow> (\<forall>pat t. (pat, t) |\<in>| cs \<longrightarrow> P t) \<longrightarrow> P t)"
context begin
private lemma abs_pred_trivI0: "P t \<Longrightarrow> abs_pred P (t::pterm)"
unfolding abs_pred_pterm_def by auto
instance proof (standard, goal_cases)
case (1 P t)
then show ?case
by (induction t rule: pterm_induct)
(auto simp: const_pterm_def free_pterm_def app_pterm_def abs_pred_pterm_def)
next
(* FIXME proving 2, 3 and 4 via sterm probably requires lifting setup *)
(* lifting setup requires a consistent ordering without assumptions! *)
(* but: other parts (in Eq_Logic_PM_Seq) require a key-ordering that only works with assumptions *)
(* solution: don't try to abstract over the ordering *)
case (2 t)
show ?case
unfolding abs_pred_pterm_def
apply clarify
apply (rule subst_pabs_id)
by blast
next
case (3 x t)
show ?case
unfolding abs_pred_pterm_def
apply clarsimp
apply (rule fset.map_cong0)
apply (rename_tac c)
apply (case_tac c; hypsubst_thin)
apply simp
subgoal for cs env pat rhs
apply (cases "x |\<in>| frees pat")
subgoal
apply (rule arg_cong[where f = "subst rhs"])
by (auto intro: fmap_ext)
subgoal premises prems[rule_format]
apply (subst (2) prems(1)[symmetric, where pat = pat])
subgoal
- by (subst fmember_iff_member_fset) fact
+ by fact
subgoal
using prems unfolding ffUnion_alt_def
- by (auto simp: fmember_iff_member_fset fset_of_list.rep_eq elim!: fBallE)
+ by (auto simp: fset_of_list.rep_eq elim!: fBallE)
subgoal
apply (rule arg_cong[where f = "subst rhs"])
by (auto intro: fmap_ext)
done
done
done
next
case (4 t)
show ?case
unfolding abs_pred_pterm_def
apply clarsimp
apply (rule fset.map_cong0)
apply clarsimp
subgoal premises prems[rule_format] for cs env\<^sub>1 env\<^sub>2 a b
- apply (rule prems(2)[unfolded fmember_iff_member_fset, OF prems(5)])
+ apply (rule prems(2)[OF prems(5)])
using prems unfolding fdisjnt_alt_def by auto
done
next
case 5
show ?case
proof (rule abs_pred_trivI0, clarify)
fix t :: pterm
fix env :: "(name, pterm) fmap"
obtain t' where "t = sterm_to_pterm t'"
by (rule sterm_to_pterm_total)
obtain env' where "env = fmmap sterm_to_pterm env'"
by (metis fmmap_total sterm_to_pterm_total)
show "frees (subst t env) = frees t - fmdom env" if "fmpred (\<lambda>_. closed) env"
unfolding \<open>t = _\<close> \<open>env = _\<close>
apply simp
apply (subst subst_sterm_to_pterm)
apply simp
apply (rule subst_frees)
using that unfolding \<open>env = _\<close>
apply simp
apply (rule fmpred_mono_strong; assumption?)
unfolding closed_except_def by simp
qed
next
case 6
show ?case
proof (rule abs_pred_trivI0, clarify)
fix t :: pterm
fix env :: "(name, pterm) fmap"
obtain t' where "t = sterm_to_pterm t'"
by (rule sterm_to_pterm_total)
obtain env' where "env = fmmap sterm_to_pterm env'"
by (metis fmmap_total sterm_to_pterm_total)
show "consts (subst t env) = consts t |\<union>| ffUnion (consts |`| fmimage env (frees t))"
unfolding \<open>t = _\<close> \<open>env = _\<close>
apply simp
apply (subst comp_def)
apply simp
apply (subst subst_sterm_to_pterm)
apply simp
apply (rule subst_consts')
done
qed
qed (rule abs_pred_trivI0)
end
end
lemma no_abs_abs[simp]: "\<not> no_abs (Pabs cs)"
by (subst no_abs.simps) (auto simp: term_cases_def)
lemma sterm_to_pterm:
assumes "no_abs t"
shows "sterm_to_pterm t = convert_term t"
using assms proof induction
case (free name)
show ?case
apply simp
apply (simp add: free_sterm_def free_pterm_def)
done
next
case (const name)
show ?case
apply simp
apply (simp add: const_sterm_def const_pterm_def)
done
next
case (app t\<^sub>1 t\<^sub>2)
then show ?case
apply simp
apply (simp add: app_sterm_def app_pterm_def)
done
qed
abbreviation Pabs_single ("\<Lambda>\<^sub>p _. _" [0, 50] 50) where
"Pabs_single x rhs \<equiv> Pabs {| (Free x, rhs) |}"
lemma closed_except_simps:
"closed_except (Pvar x) S \<longleftrightarrow> x |\<in>| S"
"closed_except (t\<^sub>1 $\<^sub>p t\<^sub>2) S \<longleftrightarrow> closed_except t\<^sub>1 S \<and> closed_except t\<^sub>2 S"
"closed_except (Pabs cs) S \<longleftrightarrow> fBall cs (\<lambda>(pat, t). closed_except t (S |\<union>| frees pat))"
"closed_except (Pconst name) S \<longleftrightarrow> True"
proof goal_cases
case 3
show ?case
proof (standard, goal_cases)
case 1
then show ?case
apply (auto simp: ffUnion_alt_def closed_except_def)
apply (drule ffUnion_least_rev)
apply auto
by (smt case_prod_conv fBall_alt_def fminus_iff fset_rev_mp id_apply map_prod_simp)
next
case 2
then show ?case
by (fastforce simp: ffUnion_alt_def closed_except_def)
qed
qed (auto simp: ffUnion_alt_def closed_except_def)
instantiation pterm :: pre_strong_term begin
function (sequential) wellformed_pterm :: "pterm \<Rightarrow> bool" where
"wellformed_pterm (t\<^sub>1 $\<^sub>p t\<^sub>2) \<longleftrightarrow> wellformed_pterm t\<^sub>1 \<and> wellformed_pterm t\<^sub>2" |
"wellformed_pterm (Pabs cs) \<longleftrightarrow> fBall cs (\<lambda>(pat, t). linear pat \<and> wellformed_pterm t) \<and> is_fmap cs \<and> pattern_compatibles cs \<and> cs \<noteq> {||}" |
"wellformed_pterm _ \<longleftrightarrow> True"
by pat_completeness auto
termination
proof (relation "measure size", goal_cases)
case 4
then show ?case
apply auto
including fset.lifting apply transfer
apply (rule le_imp_less_Suc)
apply (rule sum_nat_le_single[where y = "(a, (b, size b))" for a b])
by auto
qed auto
primrec all_frees_pterm :: "pterm \<Rightarrow> name fset" where
"all_frees_pterm (Pvar x) = {|x|}" |
"all_frees_pterm (t\<^sub>1 $\<^sub>p t\<^sub>2) = all_frees_pterm t\<^sub>1 |\<union>| all_frees_pterm t\<^sub>2" |
"all_frees_pterm (Pabs cs) = ffUnion ((\<lambda>(P, T). P |\<union>| T) |`| map_prod frees all_frees_pterm |`| cs)" |
"all_frees_pterm (Pconst _) = {||}"
instance
by standard (auto simp: const_pterm_def free_pterm_def app_pterm_def)
end
lemma sterm_to_pterm_all_frees[simp]: "all_frees (sterm_to_pterm t) = all_frees t"
proof (induction t)
case (Sabs cs)
show ?case
apply simp
apply (rule arg_cong[where f = ffUnion])
apply (rule fimage_cong[OF refl])
apply clarsimp
apply (subst Sabs)
by (auto simp: fset_of_list_elem snds.simps)
qed auto
instance pterm :: strong_term proof (standard, goal_cases)
case (1 t)
obtain t' where "t = sterm_to_pterm t'"
by (metis sterm_to_pterm_total)
show ?case
apply (rule abs_pred_trivI)
unfolding \<open>t = _\<close> sterm_to_pterm_all_frees sterm_to_pterm_frees
by (rule frees_all_frees)
next
case (2 t)
show ?case
unfolding abs_pred_pterm_def
apply (intro allI impI)
apply (simp add: case_prod_twice, intro conjI)
subgoal by blast
subgoal by (auto intro: is_fmap_image)
subgoal
unfolding fpairwise_image fpairwise_alt_def
by (auto elim!: fBallE)
done
qed
lemma wellformed_PabsI:
assumes "is_fmap cs" "pattern_compatibles cs" "cs \<noteq> {||}"
assumes "\<And>pat t. (pat, t) |\<in>| cs \<Longrightarrow> linear pat"
assumes "\<And>pat t. (pat, t) |\<in>| cs \<Longrightarrow> wellformed t"
shows "wellformed (Pabs cs)"
using assms by auto
corollary subst_closed_pabs:
assumes "(pat, rhs) |\<in>| cs" "closed (Pabs cs)"
shows "subst rhs (fmdrop_fset (frees pat) env) = rhs"
using assms by (subst subst_closed_except_id) (auto simp: fdisjnt_alt_def closed_except_simps)
lemma (in constants) shadows_consts_pterm_simps[simp]:
"shadows_consts (t\<^sub>1 $\<^sub>p t\<^sub>2) \<longleftrightarrow> shadows_consts t\<^sub>1 \<or> shadows_consts t\<^sub>2"
"shadows_consts (Pvar name) \<longleftrightarrow> name |\<in>| all_consts"
"shadows_consts (Pabs cs) \<longleftrightarrow> fBex cs (\<lambda>(pat, t). shadows_consts pat \<or> shadows_consts t)"
"shadows_consts (Pconst name) \<longleftrightarrow> False"
proof goal_cases
case 3
(* FIXME duplicated from Sterm *)
show ?case
unfolding shadows_consts_def
apply rule
subgoal
by (force simp: ffUnion_alt_def fset_of_list_elem fdisjnt_alt_def elim!: ballE)
subgoal
apply (auto simp: fset_of_list_elem fdisjnt_alt_def)
by (auto simp: fset_eq_empty_iff ffUnion_alt_def fset_of_list_elem elim!: allE fBallE)
done
qed (auto simp: shadows_consts_def fdisjnt_alt_def)
end
\ No newline at end of file
diff --git a/thys/CakeML_Codegen/Terms/Sterm.thy b/thys/CakeML_Codegen/Terms/Sterm.thy
--- a/thys/CakeML_Codegen/Terms/Sterm.thy
+++ b/thys/CakeML_Codegen/Terms/Sterm.thy
@@ -1,381 +1,381 @@
section \<open>Terms with sequential pattern matching\<close>
theory Sterm
imports Strong_Term
begin
datatype sterm =
Sconst name |
Svar name |
Sabs (clauses: "(term \<times> sterm) list") |
Sapp sterm sterm (infixl "$\<^sub>s" 70)
datatype_compat sterm
derive linorder sterm
abbreviation Sabs_single ("\<Lambda>\<^sub>s _. _" [0, 50] 50) where
"Sabs_single x rhs \<equiv> Sabs [(Free x, rhs)]"
type_synonym sclauses = "(term \<times> sterm) list"
lemma sterm_induct[case_names Sconst Svar Sabs Sapp]:
assumes "\<And>x. P (Sconst x)"
assumes "\<And>x. P (Svar x)"
assumes "\<And>cs. (\<And>pat t. (pat, t) \<in> set cs \<Longrightarrow> P t) \<Longrightarrow> P (Sabs cs)"
assumes "\<And>t u. P t \<Longrightarrow> P u \<Longrightarrow> P (t $\<^sub>s u)"
shows "P t"
using assms
apply induction_schema
apply pat_completeness
apply lexicographic_order
done
instantiation sterm :: pre_term begin
definition app_sterm where
"app_sterm t u = t $\<^sub>s u"
fun unapp_sterm where
"unapp_sterm (t $\<^sub>s u) = Some (t, u)" |
"unapp_sterm _ = None"
definition const_sterm where
"const_sterm = Sconst"
fun unconst_sterm where
"unconst_sterm (Sconst name) = Some name" |
"unconst_sterm _ = None"
fun unfree_sterm where
"unfree_sterm (Svar name) = Some name" |
"unfree_sterm _ = None"
definition free_sterm where
"free_sterm = Svar"
fun frees_sterm where
"frees_sterm (Svar name) = {|name|}" |
"frees_sterm (Sconst _) = {||}" |
"frees_sterm (Sabs cs) = ffUnion (fset_of_list (map (\<lambda>(pat, rhs). frees_sterm rhs - frees pat) cs))" |
"frees_sterm (t $\<^sub>s u) = frees_sterm t |\<union>| frees_sterm u"
fun subst_sterm where
"subst_sterm (Svar s) env = (case fmlookup env s of Some t \<Rightarrow> t | None \<Rightarrow> Svar s)" |
"subst_sterm (t\<^sub>1 $\<^sub>s t\<^sub>2) env = subst_sterm t\<^sub>1 env $\<^sub>s subst_sterm t\<^sub>2 env" |
"subst_sterm (Sabs cs) env = Sabs (map (\<lambda>(pat, rhs). (pat, subst_sterm rhs (fmdrop_fset (frees pat) env))) cs)" |
"subst_sterm t env = t"
fun consts_sterm :: "sterm \<Rightarrow> name fset" where
"consts_sterm (Svar _) = {||}" |
"consts_sterm (Sconst name) = {|name|}" |
"consts_sterm (Sabs cs) = ffUnion (fset_of_list (map (\<lambda>(_, rhs). consts_sterm rhs) cs))" |
"consts_sterm (t $\<^sub>s u) = consts_sterm t |\<union>| consts_sterm u"
instance
by standard
(auto
simp: app_sterm_def const_sterm_def free_sterm_def
elim: unapp_sterm.elims unconst_sterm.elims unfree_sterm.elims
split: option.splits)
end
instantiation sterm :: "term" begin
definition abs_pred_sterm :: "(sterm \<Rightarrow> bool) \<Rightarrow> sterm \<Rightarrow> bool" where
[code del]: "abs_pred P t \<longleftrightarrow> (\<forall>cs. t = Sabs cs \<longrightarrow> (\<forall>pat t. (pat, t) \<in> set cs \<longrightarrow> P t) \<longrightarrow> P t)"
lemma abs_pred_stermI[intro]:
assumes "\<And>cs. (\<And>pat t. (pat, t) \<in> set cs \<Longrightarrow> P t) \<Longrightarrow> P (Sabs cs)"
shows "abs_pred P t"
using assms unfolding abs_pred_sterm_def by auto
instance proof (standard, goal_cases)
case (1 P t)
then show ?case
by (induction t) (auto simp: const_sterm_def free_sterm_def app_sterm_def abs_pred_sterm_def)
next
case (2 t)
show ?case
apply rule
apply auto
apply (subst (3) list.map_id[symmetric])
apply (rule list.map_cong0)
apply auto
by blast
next
case (3 x t)
show ?case
apply rule
apply clarsimp
subgoal for cs env pat rhs
apply (cases "x |\<in>| frees pat")
subgoal
apply (rule arg_cong[where f = "subst rhs"])
by (auto intro: fmap_ext)
subgoal premises prems[rule_format]
apply (subst (2) prems(1)[symmetric, where pat = pat])
subgoal by fact
subgoal
using prems
unfolding ffUnion_alt_def
- by (auto simp add: fmember_iff_member_fset fset_of_list.rep_eq elim!: fBallE)
+ by (auto simp add: fset_of_list.rep_eq elim!: fBallE)
subgoal
apply (rule arg_cong[where f = "subst rhs"])
by (auto intro: fmap_ext)
done
done
done
next
case (4 t)
show ?case
apply rule
apply clarsimp
subgoal premises prems[rule_format]
apply (rule prems(1)[OF prems(4)])
subgoal using prems by auto
subgoal using prems unfolding fdisjnt_alt_def by auto
done
done
next
case 5
show ?case
proof (intro abs_pred_stermI allI impI, goal_cases)
case (1 cs env)
show ?case
proof safe
fix name
assume "name |\<in>| frees (subst (Sabs cs) env)"
then obtain pat rhs
where "(pat, rhs) \<in> set cs"
and "name |\<in>| frees (subst rhs (fmdrop_fset (frees pat) env))"
and "name |\<notin>| frees pat"
by (auto simp: fset_of_list_elem case_prod_twice comp_def ffUnion_alt_def)
hence "name |\<in>| frees rhs |-| fmdom (fmdrop_fset (frees pat) env)"
using 1 by (simp add: fmpred_drop_fset)
hence "name |\<in>| frees rhs |-| frees pat"
using \<open>name |\<notin>| frees pat\<close> by blast
show "name |\<in>| frees (Sabs cs)"
apply (simp add: ffUnion_alt_def)
apply (rule fBexI[where x = "(pat, rhs)"])
unfolding prod.case
apply (fact \<open>name |\<in>| frees rhs |-| frees pat\<close>)
unfolding fset_of_list_elem
by fact
assume "name |\<in>| fmdom env"
thus False
using \<open>name |\<in>| frees rhs |-| fmdom (fmdrop_fset (frees pat) env)\<close> \<open>name |\<notin>| frees pat\<close>
by fastforce
next
fix name
assume "name |\<in>| frees (Sabs cs)" "name |\<notin>| fmdom env"
then obtain pat rhs
where "(pat, rhs) \<in> set cs" "name |\<in>| frees rhs" "name |\<notin>| frees pat"
by (auto simp: fset_of_list_elem ffUnion_alt_def)
moreover hence "name |\<in>| frees rhs |-| fmdom (fmdrop_fset (frees pat) env) |-| frees pat"
using \<open>name |\<notin>| fmdom env\<close> by fastforce
ultimately have "name |\<in>| frees (subst rhs (fmdrop_fset (frees pat) env)) |-| frees pat"
using 1 by (simp add: fmpred_drop_fset)
show "name |\<in>| frees (subst (Sabs cs) env)"
apply (simp add: case_prod_twice comp_def)
unfolding ffUnion_alt_def
apply (rule fBexI)
apply (fact \<open>name |\<in>| frees (subst rhs (fmdrop_fset (frees pat) env)) |-| frees pat\<close>)
apply (subst fimage_iff)
apply (rule fBexI[where x = "(pat, rhs)"])
apply simp
using \<open>(pat, rhs) \<in> set cs\<close>
by (auto simp: fset_of_list_elem)
qed
qed
next
case 6
show ?case
proof (intro abs_pred_stermI allI impI, goal_cases)
case (1 cs env)
\<comment> \<open>some property on various operations that is only useful in here\<close>
have *: "fbind (fmimage m (fbind A g)) f = fbind A (\<lambda>x. fbind (fmimage m (g x)) f)"
for m A f g
including fset.lifting fmap.lifting
by transfer' force
have "consts (subst (Sabs cs) env) = fbind (fset_of_list cs) (\<lambda>(pat, rhs). consts rhs |\<union>| ffUnion (consts |`| fmimage (fmdrop_fset (frees pat) env) (frees rhs)))"
apply (simp add: funion_image_bind_eq)
apply (rule fbind_cong[OF refl])
apply (clarsimp split: prod.splits)
apply (subst 1)
apply (subst (asm) fset_of_list_elem, assumption)
apply simp
by (simp add: funion_image_bind_eq)
also have "\<dots> = fbind (fset_of_list cs) (consts \<circ> snd) |\<union>| fbind (fset_of_list cs) (\<lambda>(pat, rhs). ffUnion (consts |`| fmimage (fmdrop_fset (frees pat) env) (frees rhs)))"
apply (subst fbind_fun_funion[symmetric])
apply (rule fbind_cong[OF refl])
by auto
also have "\<dots> = consts (Sabs cs) |\<union>| fbind (fset_of_list cs) (\<lambda>(pat, rhs). ffUnion (consts |`| fmimage (fmdrop_fset (frees pat) env) (frees rhs)))"
apply (rule cong[OF cong, OF refl _ refl, where f1 = "funion"])
apply (subst funion_image_bind_eq[symmetric])
unfolding consts_sterm.simps
apply (rule arg_cong[where f = ffUnion])
apply (subst fset_of_list_map)
apply (rule fset.map_cong[OF refl])
by auto
also have "\<dots> = consts (Sabs cs) |\<union>| fbind (fmimage env (fbind (fset_of_list cs) (\<lambda>(pat, rhs). frees rhs |-| frees pat))) consts"
apply (subst funion_image_bind_eq)
apply (subst fmimage_drop_fset)
apply (rule cong[OF cong, OF refl refl, where f1 = "funion"])
apply (subst *)
apply (rule fbind_cong[OF refl])
by auto
also have "\<dots> = consts (Sabs cs) |\<union>| ffUnion (consts |`| fmimage env (frees (Sabs cs)))"
by (simp only: frees_sterm.simps fset_of_list_map fmimage_Union funion_image_bind_eq)
finally show ?case .
qed
qed (auto simp: abs_pred_sterm_def)
end
lemma no_abs_abs[simp]: "\<not> no_abs (Sabs cs)"
by (subst no_abs.simps) (auto simp: term_cases_def)
lemma closed_except_simps:
"closed_except (Svar x) S \<longleftrightarrow> x |\<in>| S"
"closed_except (t\<^sub>1 $\<^sub>s t\<^sub>2) S \<longleftrightarrow> closed_except t\<^sub>1 S \<and> closed_except t\<^sub>2 S"
"closed_except (Sabs cs) S \<longleftrightarrow> list_all (\<lambda>(pat, t). closed_except t (S |\<union>| frees pat)) cs"
"closed_except (Sconst name) S \<longleftrightarrow> True"
proof goal_cases
case 3
show ?case
proof (standard, goal_cases)
case 1
then show ?case
apply (auto simp: list_all_iff ffUnion_alt_def fset_of_list_elem closed_except_def)
apply (drule ffUnion_least_rev)
apply auto
by (smt case_prod_conv fbspec fimageI fminusI fset_of_list_elem fset_rev_mp)
next
case 2
then show ?case
by (fastforce simp: list_all_iff ffUnion_alt_def fset_of_list_elem closed_except_def)
qed
qed (auto simp: ffUnion_alt_def closed_except_def)
lemma closed_except_sabs:
assumes "closed (Sabs cs)" "(pat, rhs) \<in> set cs"
shows "closed_except rhs (frees pat)"
using assms unfolding closed_except_def
apply auto
by (metis bot.extremum_uniqueI fempty_iff ffUnion_subset_elem fimageI fminusI fset_of_list_elem old.prod.case)
instantiation sterm :: strong_term begin
fun wellformed_sterm :: "sterm \<Rightarrow> bool" where
"wellformed_sterm (t\<^sub>1 $\<^sub>s t\<^sub>2) \<longleftrightarrow> wellformed_sterm t\<^sub>1 \<and> wellformed_sterm t\<^sub>2" |
"wellformed_sterm (Sabs cs) \<longleftrightarrow> list_all (\<lambda>(pat, t). linear pat \<and> wellformed_sterm t) cs \<and> distinct (map fst cs) \<and> cs \<noteq> []" |
"wellformed_sterm _ \<longleftrightarrow> True"
primrec all_frees_sterm :: "sterm \<Rightarrow> name fset" where
"all_frees_sterm (Svar x) = {|x|}" |
"all_frees_sterm (t\<^sub>1 $\<^sub>s t\<^sub>2) = all_frees_sterm t\<^sub>1 |\<union>| all_frees_sterm t\<^sub>2" |
"all_frees_sterm (Sabs cs) = ffUnion (fset_of_list (map (\<lambda>(P, T). P |\<union>| T) (map (map_prod frees all_frees_sterm) cs)))" |
"all_frees_sterm (Sconst _) = {||}"
instance proof (standard, goal_cases)
case (7 t)
show ?case
apply (intro abs_pred_stermI allI impI)
apply simp
apply (rule ffUnion_least)
apply (rule fBallI)
apply auto
apply (subst ffUnion_alt_def)
apply simp
apply (rule_tac x = "(a, b)" in fBexI)
by (auto simp: fset_of_list_elem)
next
case (8 t)
show ?case
apply (intro abs_pred_stermI allI impI)
apply (simp add: list.pred_map comp_def case_prod_twice, safe)
subgoal
apply (subst list_all_iff)
apply (rule ballI)
apply safe[1]
apply (fastforce simp: list_all_iff)
subgoal premises prems[rule_format]
apply (rule prems)
apply (fact prems)
using prems apply (fastforce simp: list_all_iff)
using prems by force
done
subgoal
apply (subst map_cong[OF refl])
by auto
done
qed (auto simp: const_sterm_def free_sterm_def app_sterm_def)
end
lemma match_sabs[simp]: "\<not> is_free t \<Longrightarrow> match t (Sabs cs) = None"
by (cases t) auto
context pre_constants begin
lemma welldefined_sabs: "welldefined (Sabs cs) \<longleftrightarrow> list_all (\<lambda>(_, t). welldefined t) cs"
apply (auto simp: list_all_iff ffUnion_alt_def dest!: ffUnion_least_rev)
apply (subst (asm) list_all_iff_fset[symmetric])
apply (auto simp: list_all_iff fset_of_list_elem)
done
lemma shadows_consts_sterm_simps[simp]:
"shadows_consts (t\<^sub>1 $\<^sub>s t\<^sub>2) \<longleftrightarrow> shadows_consts t\<^sub>1 \<or> shadows_consts t\<^sub>2"
"shadows_consts (Svar name) \<longleftrightarrow> name |\<in>| all_consts"
"shadows_consts (Sabs cs) \<longleftrightarrow> list_ex (\<lambda>(pat, t). \<not> fdisjnt all_consts (frees pat) \<or> shadows_consts t) cs"
"shadows_consts (Sconst name) \<longleftrightarrow> False"
proof (goal_cases)
case 3
show ?case
unfolding shadows_consts_def list_ex_iff
apply rule
subgoal
by (force simp: ffUnion_alt_def fset_of_list_elem fdisjnt_alt_def elim!: ballE)
subgoal
apply (auto simp: fset_of_list_elem fdisjnt_alt_def)
by (auto simp: fset_eq_empty_iff ffUnion_alt_def fset_of_list_elem elim!: allE fBallE)
done
qed (auto simp: shadows_consts_def fdisjnt_alt_def)
(* FIXME derive from axioms? *)
lemma subst_shadows:
assumes "\<not> shadows_consts (t::sterm)" "not_shadows_consts_env \<Gamma>"
shows "\<not> shadows_consts (subst t \<Gamma>)"
using assms proof (induction t arbitrary: \<Gamma> rule: sterm_induct)
case (Sabs cs)
show ?case
apply (simp add: list_ex_iff case_prod_twice)
apply (rule ballI)
subgoal for c
apply (cases c, hypsubst_thin, simp)
apply (rule conjI)
subgoal using Sabs(2) by (fastforce simp: list_ex_iff)
apply (rule Sabs(1))
apply assumption
subgoal using Sabs(2) by (fastforce simp: list_ex_iff)
subgoal using Sabs(3) by force
done
done
qed (auto split: option.splits)
end
end
\ No newline at end of file
diff --git a/thys/Coinductive_Languages/Context_Free_Grammar.thy b/thys/Coinductive_Languages/Context_Free_Grammar.thy
--- a/thys/Coinductive_Languages/Context_Free_Grammar.thy
+++ b/thys/Coinductive_Languages/Context_Free_Grammar.thy
@@ -1,177 +1,177 @@
section \<open>Word Problem for Context-Free Grammars\<close>
(*<*)
theory Context_Free_Grammar
imports Coinductive_Language "HOL-Library.FSet"
begin
(*>*)
section \<open>Context Free Languages\<close>
text \<open>
A context-free grammar consists of a list of productions for every nonterminal
and an initial nonterminal. The productions are required to be in weak Greibach
normal form, i.e. each right hand side of a production must either be empty or
start with a terminal.
\<close>
abbreviation "wgreibach \<alpha> \<equiv> (case \<alpha> of (Inr N # _) \<Rightarrow> False | _ \<Rightarrow> True)"
record ('t, 'n) cfg =
init :: "'n :: finite"
prod :: "'n \<Rightarrow> ('t + 'n) list fset"
context
fixes G :: "('t, 'n :: finite) cfg"
begin
inductive in_cfl where
"in_cfl [] []"
| "in_cfl \<alpha> w \<Longrightarrow> in_cfl (Inl a # \<alpha>) (a # w)"
| "fBex (prod G N) (\<lambda>\<beta>. in_cfl (\<beta> @ \<alpha>) w) \<Longrightarrow> in_cfl (Inr N # \<alpha>) w"
abbreviation lang_trad where
"lang_trad \<equiv> {w. in_cfl [Inr (init G)] w}"
fun \<oo>\<^sub>P where
"\<oo>\<^sub>P [] = True"
| "\<oo>\<^sub>P (Inl _ # _) = False"
| "\<oo>\<^sub>P (Inr N # \<alpha>) = ([] |\<in>| prod G N \<and> \<oo>\<^sub>P \<alpha>)"
fun \<dd>\<^sub>P where
"\<dd>\<^sub>P [] a = {||}"
| "\<dd>\<^sub>P (Inl b # \<alpha>) a = (if a = b then {|\<alpha>|} else {||})"
| "\<dd>\<^sub>P (Inr N # \<alpha>) a =
(\<lambda>\<beta>. tl \<beta> @ \<alpha>) |`| ffilter (\<lambda>\<beta>. \<beta> \<noteq> [] \<and> hd \<beta> = Inl a) (prod G N) |\<union>|
(if [] |\<in>| prod G N then \<dd>\<^sub>P \<alpha> a else {||})"
primcorec subst :: "('t + 'n) list fset \<Rightarrow> 't language" where
"subst P = Lang (fBex P \<oo>\<^sub>P) (\<lambda>a. subst (ffUnion ((\<lambda>r. \<dd>\<^sub>P r a) |`| P)))"
inductive in_cfls where
"fBex P \<oo>\<^sub>P \<Longrightarrow> in_cfls P []"
| "in_cfls (ffUnion ((\<lambda>\<alpha>. \<dd>\<^sub>P \<alpha> a) |`| P)) w \<Longrightarrow> in_cfls P (a # w)"
inductive_cases [elim!]: "in_cfls P []"
inductive_cases [elim!]: "in_cfls P (a # w)"
declare inj_eq[OF bij_is_inj[OF to_language_bij], simp]
lemma subst_in_cfls: "subst P = to_language {w. in_cfls P w}"
by (coinduction arbitrary: P) (auto intro: in_cfls.intros)
lemma \<oo>\<^sub>P_in_cfl: "\<oo>\<^sub>P \<alpha> \<Longrightarrow> in_cfl \<alpha> []"
by (induct \<alpha> rule: \<oo>\<^sub>P.induct) (auto intro!: in_cfl.intros elim: fBexI[rotated])
lemma \<dd>\<^sub>P_in_cfl: "\<beta> |\<in>| \<dd>\<^sub>P \<alpha> a \<Longrightarrow> in_cfl \<beta> w \<Longrightarrow> in_cfl \<alpha> (a # w)"
proof (induct \<alpha> a arbitrary: \<beta> w rule: \<dd>\<^sub>P.induct)
case (3 N \<alpha> a)
then show ?case
by (auto simp: rev_fBexI neq_Nil_conv split: if_splits
intro!: in_cfl.intros elim!: rev_fBexI[of "_ # _"])
qed (auto split: if_splits intro: in_cfl.intros)
lemma in_cfls_in_cfl: "in_cfls P w \<Longrightarrow> fBex P (\<lambda>\<alpha>. in_cfl \<alpha> w)"
by (induct P w rule: in_cfls.induct)
- (auto simp: \<oo>\<^sub>P_in_cfl \<dd>\<^sub>P_in_cfl ffUnion.rep_eq fmember_iff_member_fset fBex.rep_eq fBall.rep_eq
+ (auto simp: \<oo>\<^sub>P_in_cfl \<dd>\<^sub>P_in_cfl ffUnion.rep_eq fBex.rep_eq fBall.rep_eq
intro: in_cfl.intros elim: rev_bexI)
lemma in_cfls_mono: "in_cfls P w \<Longrightarrow> P |\<subseteq>| Q \<Longrightarrow> in_cfls Q w"
proof (induct P w arbitrary: Q rule: in_cfls.induct)
case (2 a P w)
from 2(3) 2(2)[of "ffUnion ((\<lambda>\<alpha>. local.\<dd>\<^sub>P \<alpha> a) |`| Q)"] show ?case
by (auto intro!: ffunion_mono in_cfls.intros)
qed (auto intro!: in_cfls.intros)
end
locale cfg_wgreibach =
fixes G :: "('t, 'n :: finite) cfg"
assumes weakGreibach: "\<And>N \<alpha>. \<alpha> |\<in>| prod G N \<Longrightarrow> wgreibach \<alpha>"
begin
lemma in_cfl_in_cfls: "in_cfl G \<alpha> w \<Longrightarrow> in_cfls G {|\<alpha>|} w"
proof (induct \<alpha> w rule: in_cfl.induct)
case (3 N \<alpha> w)
then obtain \<beta> where
\<beta>: "\<beta> |\<in>| prod G N" and
in_cfl: "in_cfl G (\<beta> @ \<alpha>) w" and
in_cfls: "in_cfls G {|\<beta> @ \<alpha>|} w" by blast
then show ?case
proof (cases \<beta>)
case [simp]: Nil
from \<beta> in_cfls show ?thesis
by (cases w) (auto intro!: in_cfls.intros elim: in_cfls_mono)
next
case [simp]: (Cons x \<gamma>)
from weakGreibach[OF \<beta>] obtain a where [simp]: "x = Inl a" by (cases x) auto
with \<beta> in_cfls show ?thesis
apply -
apply (rule in_cfl.cases[OF in_cfl]; auto)
apply (force intro: in_cfls.intros(2) elim!: in_cfls_mono)
done
qed
qed (auto intro!: in_cfls.intros)
abbreviation lang where
"lang \<equiv> subst G {|[Inr (init G)]|}"
lemma lang_lang_trad: "lang = to_language (lang_trad G)"
proof -
have "in_cfls G {|[Inr (init G)]|} w \<longleftrightarrow> in_cfl G [Inr (init G)] w" for w
by (auto dest: in_cfls_in_cfl in_cfl_in_cfls)
then show ?thesis
by (auto simp: subst_in_cfls)
qed
end
text \<open>
The function @{term in_language} decides the word problem for a given language.
Since we can construct the language of a CFG using @{const cfg_wgreibach.lang} we obtain an
executable (but not very efficient) decision procedure for CFGs for free.
\<close>
abbreviation "\<aa> \<equiv> Inl True"
abbreviation "\<bb> \<equiv> Inl False"
abbreviation "S \<equiv> Inr ()"
interpretation palindromes: cfg_wgreibach "\<lparr>init = (), prod = \<lambda>_. {|[], [\<aa>], [\<bb>], [\<aa>, S, \<aa>], [\<bb>, S, \<bb>]|}\<rparr>"
by unfold_locales auto
lemma "in_language palindromes.lang []" by normalization
lemma "in_language palindromes.lang [True]" by normalization
lemma "in_language palindromes.lang [False]" by normalization
lemma "in_language palindromes.lang [True, True]" by normalization
lemma "in_language palindromes.lang [True, False, True]" by normalization
lemma "\<not> in_language palindromes.lang [True, False]" by normalization
lemma "\<not> in_language palindromes.lang [True, False, True, False]" by normalization
lemma "in_language palindromes.lang [True, False, True, True, False, True]" by normalization
lemma "\<not> in_language palindromes.lang [True, False, True, False, False, True]" by normalization
interpretation Dyck: cfg_wgreibach "\<lparr>init = (), prod = \<lambda>_. {|[], [\<aa>, S, \<bb>, S]|}\<rparr>"
by unfold_locales auto
lemma "in_language Dyck.lang []" by normalization
lemma "\<not> in_language Dyck.lang [True]" by normalization
lemma "\<not> in_language Dyck.lang [False]" by normalization
lemma "in_language Dyck.lang [True, False, True, False]" by normalization
lemma "in_language Dyck.lang [True, True, False, False]" by normalization
lemma "in_language Dyck.lang [True, False, True, False]" by normalization
lemma "in_language Dyck.lang [True, False, True, False, True, True, False, False]" by normalization
lemma "\<not> in_language Dyck.lang [True, False, True, True, False]" by normalization
lemma "\<not> in_language Dyck.lang [True, True, False, False, False, True]" by normalization
interpretation abSSa: cfg_wgreibach "\<lparr>init = (), prod = \<lambda>_. {|[], [\<aa>, \<bb>, S, S, \<aa>]|}\<rparr>"
by unfold_locales auto
lemma "in_language abSSa.lang []" by normalization
lemma "\<not> in_language abSSa.lang [True]" by normalization
lemma "\<not> in_language abSSa.lang [False]" by normalization
lemma "in_language abSSa.lang [True, False, True]" by normalization
lemma "in_language abSSa.lang [True, False, True, False, True, True, False, True, True]" by normalization
lemma "in_language abSSa.lang [True, False, True, False, True, True]" by normalization
lemma "\<not> in_language abSSa.lang [True, False, True, True, False]" by normalization
lemma "\<not> in_language abSSa.lang [True, True, False, False, False, True]" by normalization
(*<*)
end
(*>*)
\ No newline at end of file
diff --git a/thys/Core_DOM/common/classes/CharacterDataClass.thy b/thys/Core_DOM/common/classes/CharacterDataClass.thy
--- a/thys/Core_DOM/common/classes/CharacterDataClass.thy
+++ b/thys/Core_DOM/common/classes/CharacterDataClass.thy
@@ -1,355 +1,360 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>CharacterData\<close>
text\<open>In this theory, we introduce the types for the CharacterData class.\<close>
theory CharacterDataClass
imports
ElementClass
begin
subsubsection\<open>CharacterData\<close>
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, defined
\autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
record RCharacterData = RNode +
nothing :: unit
val :: DOMString
register_default_tvars "'CharacterData RCharacterData_ext"
type_synonym 'CharacterData CharacterData = "'CharacterData option RCharacterData_scheme"
register_default_tvars "'CharacterData CharacterData"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
'Element, 'CharacterData) Node
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
'CharacterData option RCharacterData_ext + 'Node, 'Element) Node"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
'Element, 'CharacterData) Node"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'CharacterData option RCharacterData_ext + 'Node,
'Element) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'CharacterData option RCharacterData_ext + 'Node, 'Element) heap"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition character_data_ptr_kinds :: "(_) heap \<Rightarrow> (_) character_data_ptr fset"
where
"character_data_ptr_kinds heap = the |`| (cast |`| (ffilter is_character_data_ptr_kind
(node_ptr_kinds heap)))"
lemma character_data_ptr_kinds_simp [simp]:
"character_data_ptr_kinds (Heap (fmupd (cast character_data_ptr) character_data (the_heap h)))
= {|character_data_ptr|} |\<union>| character_data_ptr_kinds h"
- apply(auto simp add: character_data_ptr_kinds_def)[1]
- by force
+ by (auto simp add: character_data_ptr_kinds_def)
definition character_data_ptrs :: "(_) heap \<Rightarrow> _ character_data_ptr fset"
where
"character_data_ptrs heap = ffilter is_character_data_ptr (character_data_ptr_kinds heap)"
abbreviation "character_data_ptr_exts heap \<equiv> character_data_ptr_kinds heap - character_data_ptrs heap"
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) Node \<Rightarrow> (_) CharacterData option"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node = (case RNode.more node of
Inr (Inl character_data) \<Rightarrow> Some (RNode.extend (RNode.truncate node) character_data)
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) Object \<Rightarrow> (_) CharacterData option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a obj \<equiv> (case cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj of Some node \<Rightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node
| None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
definition cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) CharacterData \<Rightarrow> (_) Node"
where
"cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = RNode.extend (RNode.truncate character_data)
(Inr (Inl (RNode.more character_data)))"
adhoc_overloading cast cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e
abbreviation cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) CharacterData \<Rightarrow> (_) Object"
where
"cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr)"
adhoc_overloading cast cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
consts is_character_data_kind :: 'a
definition is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Node \<Rightarrow> bool"
where
"is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr \<longleftrightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr \<noteq> None"
adhoc_overloading is_character_data_kind is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e
lemmas is_character_data_kind_def = is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
abbreviation is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) Object \<Rightarrow> bool"
where
"is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr \<noteq> None"
adhoc_overloading is_character_data_kind is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma character_data_ptr_kinds_commutes [simp]:
"cast character_data_ptr |\<in>| node_ptr_kinds h
\<longleftrightarrow> character_data_ptr |\<in>| character_data_ptr_kinds h"
- apply(auto simp add: character_data_ptr_kinds_def)[1]
- by (metis character_data_ptr_casts_commute2 comp_eq_dest_lhs ffmember_filter fimage_eqI
- is_character_data_ptr_kind_none
- option.distinct(1) option.sel)
+proof (rule iffI)
+ show "cast character_data_ptr |\<in>| node_ptr_kinds h \<Longrightarrow>
+ character_data_ptr |\<in>| character_data_ptr_kinds h"
+ by (metis (no_types, opaque_lifting) character_data_ptr_casts_commute2
+ character_data_ptr_kinds_def ffmember_filter fimage_eqI is_character_data_ptr_kind\<^sub>_cast
+ option.sel)
+next
+ show "character_data_ptr |\<in>| character_data_ptr_kinds h \<Longrightarrow>
+ cast character_data_ptr |\<in>| node_ptr_kinds h"
+ by (auto simp add: character_data_ptr_kinds_def)
+qed
definition get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) character_data_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) CharacterData option"
where
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h = Option.bind (get\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast character_data_ptr) h) cast"
adhoc_overloading get get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
locale l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (ElementClass.type_wf h
\<and> (\<forall>character_data_ptr \<in> fset (character_data_ptr_kinds h).
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a: "type_wf h \<Longrightarrow> CharacterDataClass.type_wf h"
sublocale l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a \<subseteq> l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
using ElementClass.a_type_wf_def
by (meson CharacterDataClass.a_type_wf_def l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
locale l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
sublocale l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
lemma get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf:
assumes "type_wf h"
shows "character_data_ptr |\<in>| character_data_ptr_kinds h
\<longleftrightarrow> get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h \<noteq> None"
using l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms assms
apply(simp add: type_wf_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
- by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes fmember_iff_member_fset
+ by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes
local.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf option.exhaust option.simps(3))
end
global_interpretation l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf
by unfold_locales
definition put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) character_data_ptr \<Rightarrow> (_) CharacterData \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast character_data_ptr)
(cast character_data)"
adhoc_overloading put put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
lemma put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap:
assumes "put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data h = h'"
shows "character_data_ptr |\<in>| character_data_ptr_kinds h'"
using assms put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap
unfolding put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def character_data_ptr_kinds_def
by (metis character_data_ptr_kinds_commutes character_data_ptr_kinds_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap)
lemma put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_put_ptrs:
assumes "put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast character_data_ptr|}"
using assms
by (simp add: put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs)
lemma cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]: "cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e x = cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def)
by (metis (full_types) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_none [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node = None \<longleftrightarrow> \<not> (\<exists>character_data. cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = node)"
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)[1]
by (metis (full_types) RNode.select_convs(2) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_some [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node = Some character_data \<longleftrightarrow> cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = node"
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_inv [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a (cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data) = Some character_data"
by simp
lemma cast_element_not_character_data [simp]:
"(cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element \<noteq> cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data)"
"(cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data \<noteq> cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element)"
by(auto simp add: cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RNode.extend_def)
lemma get_CharacterData_simp1 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data h)
= Some character_data"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma get_CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr' \<Longrightarrow> get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr
(put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' character_data h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma get_CharacterData_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma get_CharacterData_simp4 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t character_data_ptr f h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a element_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
abbreviation "create_character_data_obj val_arg
\<equiv> \<lparr> RObject.nothing = (), RNode.nothing = (), RCharacterData.nothing = (), val = val_arg, \<dots> = None \<rparr>"
definition new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) heap \<Rightarrow> ((_) character_data_ptr \<times> (_) heap)"
where
"new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h =
(let new_character_data_ptr = character_data_ptr.Ref (Suc (fMax (character_data_ptr.the_ref
|`| (character_data_ptrs h)))) in
(new_character_data_ptr, put new_character_data_ptr (create_character_data_obj '''') h))"
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "new_character_data_ptr |\<in>| character_data_ptr_kinds h'"
using assms
unfolding new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def
using put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap by blast
lemma new_character_data_ptr_new:
"character_data_ptr.Ref (Suc (fMax (finsert 0 (character_data_ptr.the_ref |`| character_data_ptrs h))))
|\<notin>| character_data_ptrs h"
by (metis Suc_n_not_le_n character_data_ptr.sel(1) fMax_ge fimage_finsert finsertI1
finsertI2 set_finsert)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "new_character_data_ptr |\<notin>| character_data_ptr_kinds h"
using assms
unfolding new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
by (metis Pair_inject character_data_ptrs_def fMax_finsert fempty_iff ffmember_filter
fimage_is_fempty is_character_data_ptr_ref max_0L new_character_data_ptr_new)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_new_ptr:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using assms
by (metis Pair_inject new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_put_ptrs)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_is_character_data_ptr:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "is_character_data_ptr new_character_data_ptr"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
assumes "ptr \<noteq> cast new_character_data_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
assumes "ptr \<noteq> cast new_character_data_ptr"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
assumes "ptr \<noteq> new_character_data_ptr"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
locale l_known_ptr\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_character_data_ptr ptr)"
lemma known_ptr_not_character_data_ptr:
"\<not>is_character_data_ptr ptr \<Longrightarrow> a_known_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def
by blast
end
diff --git a/thys/Core_DOM/common/classes/DocumentClass.thy b/thys/Core_DOM/common/classes/DocumentClass.thy
--- a/thys/Core_DOM/common/classes/DocumentClass.thy
+++ b/thys/Core_DOM/common/classes/DocumentClass.thy
@@ -1,345 +1,348 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Document\<close>
text\<open>In this theory, we introduce the types for the Document class.\<close>
theory DocumentClass
imports
CharacterDataClass
begin
text\<open>The type @{type "doctype"} is a type synonym for @{type "string"}, defined
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
record ('node_ptr, 'element_ptr, 'character_data_ptr) RDocument = RObject +
nothing :: unit
doctype :: doctype
document_element :: "(_) element_ptr option"
disconnected_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document) Document
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option) RDocument_scheme"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'Document) Document"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData, 'Document) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option)
RDocument_ext + 'Object, 'Node, 'Element, 'CharacterData) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'Document) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr,
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option) RDocument_ext + 'Object, 'Node,
'Element, 'CharacterData) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition document_ptr_kinds :: "(_) heap \<Rightarrow> (_) document_ptr fset"
where
"document_ptr_kinds heap = 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 |`|
(ffilter is_document_ptr_kind (object_ptr_kinds heap)))"
definition document_ptrs :: "(_) heap \<Rightarrow> (_) document_ptr fset"
where
"document_ptrs heap = ffilter is_document_ptr (document_ptr_kinds heap)"
definition cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Object \<Rightarrow> (_) Document option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = (case RObject.more obj of
Inr (Inl document) \<Rightarrow> Some (RObject.extend (RObject.truncate obj) document)
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
definition cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:: "(_) Document \<Rightarrow> (_) Object"
where
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document = (RObject.extend (RObject.truncate document)
(Inr (Inl (RObject.more document))))"
adhoc_overloading cast cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
definition is_document_kind :: "(_) Object \<Rightarrow> bool"
where
"is_document_kind ptr \<longleftrightarrow> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
lemma document_ptr_kinds_simp [simp]:
"document_ptr_kinds (Heap (fmupd (cast document_ptr) document (the_heap h)))
= {|document_ptr|} |\<union>| document_ptr_kinds h"
- apply(auto simp add: document_ptr_kinds_def)[1]
- by force
+ by (auto simp add: document_ptr_kinds_def)
lemma document_ptr_kinds_commutes [simp]:
"cast document_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> document_ptr |\<in>| document_ptr_kinds h"
- apply(auto simp add: object_ptr_kinds_def document_ptr_kinds_def)[1]
- by (metis (no_types, lifting) document_ptr_casts_commute2 document_ptr_document_ptr_cast
- ffmember_filter fimage_eqI fset.map_comp option.sel)
+proof (rule iffI)
+ show "cast document_ptr |\<in>| object_ptr_kinds h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h"
+ by (metis (no_types, lifting) document_ptr_casts_commute2 document_ptr_document_ptr_cast
+ document_ptr_kinds_def ffmember_filter fimage_eqI option.sel)
+next
+ show "document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow> cast document_ptr |\<in>| object_ptr_kinds h"
+ by (auto simp add: object_ptr_kinds_def document_ptr_kinds_def)
+qed
definition get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) document_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Document option"
where
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h = Option.bind (get (cast document_ptr) h) cast"
adhoc_overloading get get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_type_wf_def\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (CharacterDataClass.type_wf h \<and>
(\<forall>document_ptr \<in> fset (document_ptr_kinds h). get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t: "type_wf h \<Longrightarrow> DocumentClass.type_wf h"
sublocale l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t \<subseteq> l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
apply(unfold_locales)
by (metis (full_types) type_wf_defs l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
locale l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
lemma get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf:
assumes "type_wf h"
shows "document_ptr |\<in>| document_ptr_kinds h \<longleftrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None"
using l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- by (metis document_ptr_kinds_commutes fmember_iff_member_fset is_none_bind is_none_simps(1)
+ by (metis document_ptr_kinds_commutes is_none_bind is_none_simps(1)
is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
end
global_interpretation l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
definition put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) document_ptr \<Rightarrow> (_) Document \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document = put (cast document_ptr) (cast document)"
adhoc_overloading put put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
lemma put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h = h'"
shows "document_ptr |\<in>| document_ptr_kinds h'"
using assms
unfolding put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis document_ptr_kinds_commutes put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap)
lemma put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs:
assumes "put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast document_ptr|}"
using assms
by (simp add: put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs)
lemma cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_inject [simp]: "cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x = cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
by (metis (full_types) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = None \<longleftrightarrow> \<not> (\<exists>document. cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document = obj)"
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
split: sum.splits)[1]
by (metis (full_types) RObject.select_convs(2) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = Some document \<longleftrightarrow> cast document = obj"
by(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
split: sum.splits)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document) = Some document"
by simp
lemma cast_document_not_node [simp]:
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document \<noteq> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node"
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node \<noteq> cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document"
by(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
lemma get_document_ptr_simp1 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h) = Some document"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp2 [simp]:
"document_ptr \<noteq> document_ptr'
\<Longrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' document h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp4 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_document_ptr_simp5 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp6 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
abbreviation
create_document_obj :: "char list \<Rightarrow> (_) element_ptr option \<Rightarrow> (_) node_ptr list \<Rightarrow> (_) Document"
where
"create_document_obj doctype_arg document_element_arg disconnected_nodes_arg
\<equiv> \<lparr> RObject.nothing = (), RDocument.nothing = (), doctype = doctype_arg,
document_element = document_element_arg,
disconnected_nodes = disconnected_nodes_arg, \<dots> = None \<rparr>"
definition new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_)heap \<Rightarrow> ((_) document_ptr \<times> (_) heap)"
where
"new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_document_ptr = document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| (document_ptrs h)))))
in
(new_document_ptr, put new_document_ptr (create_document_obj '''' None []) h))"
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "new_document_ptr |\<in>| document_ptr_kinds h'"
using assms
unfolding new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
using put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap by blast
lemma new_document_ptr_new:
"document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| document_ptrs h))))
|\<notin>| document_ptrs h"
by (metis Suc_n_not_le_n document_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "new_document_ptr |\<notin>| document_ptr_kinds h"
using assms
unfolding new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis Pair_inject document_ptrs_def fMax_finsert fempty_iff ffmember_filter
fimage_is_fempty is_document_ptr_ref max_0L new_document_ptr_new)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_document_ptr|}"
using assms
by (metis Pair_inject new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_document_ptr:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "is_document_ptr new_document_ptr"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
assumes "ptr \<noteq> cast new_document_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
apply(simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
assumes "ptr \<noteq> new_document_ptr"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
locale l_known_ptr\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_document_ptr ptr)"
lemma known_ptr_not_document_ptr: "\<not>is_document_ptr ptr \<Longrightarrow> a_known_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs [instances]: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
end
diff --git a/thys/Core_DOM/common/classes/NodeClass.thy b/thys/Core_DOM/common/classes/NodeClass.thy
--- a/thys/Core_DOM/common/classes/NodeClass.thy
+++ b/thys/Core_DOM/common/classes/NodeClass.thy
@@ -1,209 +1,212 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Node\<close>
text\<open>In this theory, we introduce the types for the Node class.\<close>
theory NodeClass
imports
ObjectClass
"../pointers/NodePointer"
begin
subsubsection\<open>Node\<close>
record RNode = RObject
+ nothing :: unit
register_default_tvars "'Node RNode_ext"
type_synonym 'Node Node = "'Node RNode_scheme"
register_default_tvars "'Node Node"
type_synonym ('Object, 'Node) Object = "('Node RNode_ext + 'Object) Object"
register_default_tvars "('Object, 'Node) Object"
type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node) heap
= "('node_ptr node_ptr + 'object_ptr, 'Node RNode_ext + 'Object) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'Object, 'Node) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit) heap"
definition node_ptr_kinds :: "(_) heap \<Rightarrow> (_) node_ptr fset"
where
"node_ptr_kinds heap =
(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 |`| (ffilter is_node_ptr_kind (object_ptr_kinds heap))))"
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (Heap (fmupd (cast node_ptr) node (the_heap h)))
= {|node_ptr|} |\<union>| node_ptr_kinds h"
- apply(auto simp add: node_ptr_kinds_def)[1]
- by force
+ by (auto simp add: node_ptr_kinds_def)
definition cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Object \<Rightarrow> (_) Node option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = (case RObject.more obj of Inl node
\<Rightarrow> Some (RObject.extend (RObject.truncate obj) node) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:: "(_) Node \<Rightarrow> (_) Object"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node = (RObject.extend (RObject.truncate node) (Inl (RObject.more node)))"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
definition is_node_kind :: "(_) Object \<Rightarrow> bool"
where
"is_node_kind ptr \<longleftrightarrow> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr \<noteq> None"
definition get\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) node_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Node option"
where
"get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h = Option.bind (get (cast node_ptr) h) cast"
adhoc_overloading get get\<^sub>N\<^sub>o\<^sub>d\<^sub>e
locale l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (ObjectClass.type_wf h
\<and> (\<forall>node_ptr \<in> fset( node_ptr_kinds h). get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e: "type_wf h \<Longrightarrow> NodeClass.type_wf h"
sublocale l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e \<subseteq> l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
apply(unfold_locales)
using ObjectClass.a_type_wf_def by auto
locale l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
sublocale l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales
lemma get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf:
assumes "type_wf h"
shows "node_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None"
using l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_axioms assms
apply(simp add: type_wf_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
- by (metis bind_eq_None_conv ffmember_filter fimage_eqI fmember_iff_member_fset is_node_ptr_kind_cast
+ by (metis bind_eq_None_conv ffmember_filter fimage_eqI is_node_ptr_kind_cast
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf node_ptr_casts_commute2 node_ptr_kinds_def option.sel option.simps(3))
end
global_interpretation l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf
by unfold_locales
definition put\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) node_ptr \<Rightarrow> (_) Node \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node = put (cast node_ptr) (cast node)"
adhoc_overloading put put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
lemma put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap:
assumes "put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h = h'"
shows "node_ptr |\<in>| node_ptr_kinds h'"
using assms
unfolding put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def node_ptr_kinds_def
by (metis ffmember_filter fimage_eqI is_node_ptr_kind_cast node_ptr_casts_commute2
option.sel put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap)
lemma put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs:
assumes "put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast node_ptr|}"
using assms
by (simp add: put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs)
lemma node_ptr_kinds_commutes [simp]:
"cast node_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> node_ptr |\<in>| node_ptr_kinds h"
- apply(auto simp add: node_ptr_kinds_def split: option.splits)[1]
- by (metis (no_types, lifting) ffmember_filter fimage_eqI fset.map_comp
- is_node_ptr_kind_none node_ptr_casts_commute2
- option.distinct(1) option.sel)
+proof (rule iffI)
+ show "cast node_ptr |\<in>| object_ptr_kinds h \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h"
+ by (metis ffmember_filter fimage_eqI is_node_ptr_kind_cast node_ptr_casts_commute2
+ node_ptr_kinds_def option.sel)
+next
+ show "node_ptr |\<in>| node_ptr_kinds h \<Longrightarrow> cast node_ptr |\<in>| object_ptr_kinds h"
+ by (auto simp add: node_ptr_kinds_def)
+qed
lemma node_empty [simp]:
"\<lparr>RObject.nothing = (), RNode.nothing = (), \<dots> = RNode.more node\<rparr> = node"
by simp
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_inject [simp]: "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x = cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
by (metis (full_types) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = None \<longleftrightarrow> \<not> (\<exists>node. cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node = obj)"
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def split: sum.splits)[1]
by (metis (full_types) RObject.select_convs(2) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_some [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = Some node \<longleftrightarrow> cast node = obj"
by(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def split: sum.splits)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node) = Some node"
by simp
locale l_known_ptr\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = False"
end
global_interpretation l_known_ptr\<^sub>N\<^sub>o\<^sub>d\<^sub>e defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
+
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset
known_ptrs_new_ptr
by blast
lemma get_node_ptr_simp1 [simp]: "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h) = Some node"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_node_ptr_simp2 [simp]:
"node_ptr \<noteq> node_ptr' \<Longrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr' node h) = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
end
diff --git a/thys/Core_DOM/common/classes/ObjectClass.thy b/thys/Core_DOM/common/classes/ObjectClass.thy
--- a/thys/Core_DOM/common/classes/ObjectClass.thy
+++ b/thys/Core_DOM/common/classes/ObjectClass.thy
@@ -1,225 +1,224 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Object\<close>
text\<open>In this theory, we introduce the definition of the class Object. This class is the
common superclass of our class model.\<close>
theory ObjectClass
imports
BaseClass
"../pointers/ObjectPointer"
begin
record RObject =
nothing :: unit
register_default_tvars "'Object RObject_ext"
type_synonym 'Object Object = "'Object RObject_scheme"
register_default_tvars "'Object Object"
datatype ('object_ptr, 'Object) heap = Heap (the_heap: "((_) object_ptr, (_) Object) fmap")
register_default_tvars "('object_ptr, 'Object) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit) heap"
definition object_ptr_kinds :: "(_) heap \<Rightarrow> (_) object_ptr fset"
where
"object_ptr_kinds = fmdom \<circ> the_heap"
lemma object_ptr_kinds_simp [simp]:
"object_ptr_kinds (Heap (fmupd object_ptr object (the_heap h)))
= {|object_ptr|} |\<union>| object_ptr_kinds h"
by(auto simp add: object_ptr_kinds_def)
definition get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Object option"
where
"get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = fmlookup (the_heap h) ptr"
adhoc_overloading get get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
locale l_type_wf_def\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = True"
end
global_interpretation l_type_wf_def\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t: "type_wf h \<Longrightarrow> ObjectClass.type_wf h"
locale l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas = l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
lemma get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf:
assumes "type_wf h"
shows "object_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h \<noteq> None"
using l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms assms
apply(simp add: type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
by (simp add: fmlookup_dom_iff object_ptr_kinds_def)
end
global_interpretation l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas type_wf
by (simp add: l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas.intro l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t.intro)
definition put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) Object \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h = Heap (fmupd ptr obj (the_heap h))"
adhoc_overloading put put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap:
assumes "put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h = h'"
shows "object_ptr |\<in>| object_ptr_kinds h'"
using assms
unfolding put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
by auto
lemma put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs:
assumes "put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|object_ptr|}"
using assms
by (metis comp_apply fmdom_fmupd funion_finsert_right heap.sel object_ptr_kinds_def
sup_bot.right_neutral put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemma object_more_extend_id [simp]: "more (extend x y) = y"
by(simp add: extend_def)
lemma object_empty [simp]: "\<lparr>nothing = (), \<dots> = more x\<rparr> = x"
by simp
locale l_known_ptr\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = False"
lemma known_ptr_not_object_ptr:
"a_known_ptr ptr \<Longrightarrow> \<not>is_object_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool" +
fixes known_ptrs :: "(_) heap \<Rightarrow> bool"
assumes known_ptrs_known_ptr: "known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
assumes known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h = known_ptrs h'"
assumes known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'"
assumes known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
known_ptrs h \<Longrightarrow> known_ptrs h'"
locale l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
"a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
lemma get_object_ptr_simp1 [simp]: "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h) = Some object"
by(simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemma get_object_ptr_simp2 [simp]:
"object_ptr \<noteq> object_ptr'
\<Longrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr' object h) = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h"
by(simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
subsection\<open>Limited Heap Modifications\<close>
definition heap_unchanged_except :: "(_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"heap_unchanged_except S h h' = (\<forall>ptr \<in> (fset (object_ptr_kinds h)
\<union> (fset (object_ptr_kinds h'))) - S. get ptr h = get ptr h')"
definition delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) heap option" where
"delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = (if ptr |\<in>| object_ptr_kinds h then Some (Heap (fmdrop ptr (the_heap h)))
else None)"
lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_removed:
assumes "delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = Some h'"
shows "ptr |\<notin>| object_ptr_kinds h'"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def object_ptr_kinds_def split: if_splits)
lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_ptr_in_heap:
assumes "delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = Some h'"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def object_ptr_kinds_def split: if_splits)
lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok:
assumes "ptr |\<in>| object_ptr_kinds h"
shows "delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h \<noteq> None"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def object_ptr_kinds_def split: if_splits)
subsection \<open>Code Generator Setup\<close>
definition "create_heap xs = Heap (fmap_of_list xs)"
code_datatype ObjectClass.heap.Heap create_heap
lemma object_ptr_kinds_code3 [code]:
"fmlookup (the_heap (create_heap xs)) x = map_of xs x"
by(auto simp add: create_heap_def fmlookup_of_list)
lemma object_ptr_kinds_code4 [code]:
"the_heap (create_heap xs) = fmap_of_list xs"
by(simp add: create_heap_def)
lemma object_ptr_kinds_code5 [code]:
"the_heap (Heap x) = x"
by simp
end
diff --git a/thys/Core_DOM/common/monads/BaseMonad.thy b/thys/Core_DOM/common/monads/BaseMonad.thy
--- a/thys/Core_DOM/common/monads/BaseMonad.thy
+++ b/thys/Core_DOM/common/monads/BaseMonad.thy
@@ -1,384 +1,381 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>The Monad Infrastructure\<close>
text\<open>In this theory, we introduce the basic infrastructure for our monadic class encoding.\<close>
theory BaseMonad
imports
"../classes/BaseClass"
"../preliminaries/Heap_Error_Monad"
begin
subsection\<open>Datatypes\<close>
datatype exception = NotFoundError | HierarchyRequestError | NotSupportedError | SegmentationFault
| AssertException | NonTerminationException | InvokeError | TypeError
-lemma finite_set_in [simp]: "x \<in> fset FS \<longleftrightarrow> x |\<in>| FS"
- by (meson notin_fset)
-
consts put_M :: 'a
consts get_M :: 'a
consts delete_M :: 'a
lemma sorted_list_of_set_eq [dest]:
"sorted_list_of_set (fset x) = sorted_list_of_set (fset y) \<Longrightarrow> x = y"
by (metis finite_fset fset_inject sorted_list_of_set(1))
locale l_ptr_kinds_M =
fixes ptr_kinds :: "'heap \<Rightarrow> 'ptr::linorder fset"
begin
definition a_ptr_kinds_M :: "('heap, exception, 'ptr list) prog"
where
"a_ptr_kinds_M = do {
h \<leftarrow> get_heap;
return (sorted_list_of_set (fset (ptr_kinds h)))
}"
lemma ptr_kinds_M_ok [simp]: "h \<turnstile> ok a_ptr_kinds_M"
by(simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_pure [simp]: "pure a_ptr_kinds_M h"
by (auto simp add: a_ptr_kinds_M_def intro: bind_pure_I)
lemma ptr_kinds_ptr_kinds_M [simp]: "ptr \<in> set |h \<turnstile> a_ptr_kinds_M|\<^sub>r \<longleftrightarrow> ptr |\<in>| ptr_kinds h"
by(simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds [simp]:
"h \<turnstile> a_ptr_kinds_M \<rightarrow>\<^sub>r xa \<longleftrightarrow> xa = sorted_list_of_set (fset (ptr_kinds h))"
by(auto simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds_returns_result [simp]:
"h \<turnstile> a_ptr_kinds_M \<bind> f \<rightarrow>\<^sub>r x \<longleftrightarrow> h \<turnstile> f (sorted_list_of_set (fset (ptr_kinds h))) \<rightarrow>\<^sub>r x"
by(auto simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds_returns_heap [simp]:
"h \<turnstile> a_ptr_kinds_M \<bind> f \<rightarrow>\<^sub>h h' \<longleftrightarrow> h \<turnstile> f (sorted_list_of_set (fset (ptr_kinds h))) \<rightarrow>\<^sub>h h'"
by(auto simp add: a_ptr_kinds_M_def)
end
locale l_get_M =
fixes get :: "'ptr \<Rightarrow> 'heap \<Rightarrow> 'obj option"
fixes type_wf :: "'heap \<Rightarrow> bool"
fixes ptr_kinds :: "'heap \<Rightarrow> 'ptr fset"
assumes "type_wf h \<Longrightarrow> ptr |\<in>| ptr_kinds h \<Longrightarrow> get ptr h \<noteq> None"
assumes "get ptr h \<noteq> None \<Longrightarrow> ptr |\<in>| ptr_kinds h"
begin
definition a_get_M :: "'ptr \<Rightarrow> ('obj \<Rightarrow> 'result) \<Rightarrow> ('heap, exception, 'result) prog"
where
"a_get_M ptr getter = (do {
h \<leftarrow> get_heap;
(case get ptr h of
Some res \<Rightarrow> return (getter res)
| None \<Rightarrow> error SegmentationFault)
})"
lemma get_M_pure [simp]: "pure (a_get_M ptr getter) h"
by(auto simp add: a_get_M_def bind_pure_I split: option.splits)
lemma get_M_ok:
"type_wf h \<Longrightarrow> ptr |\<in>| ptr_kinds h \<Longrightarrow> h \<turnstile> ok (a_get_M ptr getter)"
apply(simp add: a_get_M_def)
by (metis l_get_M_axioms l_get_M_def option.case_eq_if return_ok)
lemma get_M_ptr_in_heap:
"h \<turnstile> ok (a_get_M ptr getter) \<Longrightarrow> ptr |\<in>| ptr_kinds h"
apply(simp add: a_get_M_def)
by (metis error_returns_result is_OK_returns_result_E l_get_M_axioms l_get_M_def option.simps(4))
end
locale l_put_M = l_get_M get for get :: "'ptr \<Rightarrow> 'heap \<Rightarrow> 'obj option" +
fixes put :: "'ptr \<Rightarrow> 'obj \<Rightarrow> 'heap \<Rightarrow> 'heap"
begin
definition a_put_M :: "'ptr \<Rightarrow> (('v \<Rightarrow> 'v) \<Rightarrow> 'obj \<Rightarrow> 'obj) \<Rightarrow> 'v \<Rightarrow> ('heap, exception, unit) prog"
where
"a_put_M ptr setter v = (do {
obj \<leftarrow> a_get_M ptr id;
h \<leftarrow> get_heap;
return_heap (put ptr (setter (\<lambda>_. v) obj) h)
})"
lemma put_M_ok:
"type_wf h \<Longrightarrow> ptr |\<in>| ptr_kinds h \<Longrightarrow> h \<turnstile> ok (a_put_M ptr setter v)"
by(auto simp add: a_put_M_def intro!: bind_is_OK_I2 dest: get_M_ok elim!: bind_is_OK_E)
lemma put_M_ptr_in_heap:
"h \<turnstile> ok (a_put_M ptr setter v) \<Longrightarrow> ptr |\<in>| ptr_kinds h"
by(auto simp add: a_put_M_def intro!: bind_is_OK_I2 elim: get_M_ptr_in_heap
dest: is_OK_returns_result_I elim!: bind_is_OK_E)
end
subsection \<open>Setup for Defining Partial Functions\<close>
lemma execute_admissible:
"ccpo.admissible (fun_lub (flat_lub (Inl (e::'e)))) (fun_ord (flat_ord (Inl e)))
((\<lambda>a. \<forall>(h::'heap) h2 (r::'result). h \<turnstile> a = Inr (r, h2) \<longrightarrow> P h h2 r) \<circ> Prog)"
proof (unfold comp_def, rule ccpo.admissibleI, clarify)
fix A :: "('heap \<Rightarrow> 'e + 'result \<times> 'heap) set"
let ?lub = "Prog (fun_lub (flat_lub (Inl e)) A)"
fix h h2 r
assume 1: "Complete_Partial_Order.chain (fun_ord (flat_ord (Inl e))) A"
and 2: "\<forall>xa\<in>A. \<forall>h h2 r. h \<turnstile> Prog xa = Inr (r, h2) \<longrightarrow> P h h2 r"
and 4: "h \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r, h2)"
have h1:"\<And>a. Complete_Partial_Order.chain (flat_ord (Inl e)) {y. \<exists>f\<in>A. y = f a}"
by (rule chain_fun[OF 1])
show "P h h2 r"
using CollectD Inl_Inr_False prog.sel chain_fun[OF 1] flat_lub_in_chain[OF chain_fun[OF 1]] 2 4
unfolding execute_def fun_lub_def
proof -
assume a1: "the_prog (Prog (\<lambda>x. flat_lub (Inl e) {y. \<exists>f\<in>A. y = f x})) h = Inr (r, h2)"
assume a2: "\<forall>xa\<in>A. \<forall>h h2 r. the_prog (Prog xa) h = Inr (r, h2) \<longrightarrow> P h h2 r"
have "Inr (r, h2) \<in> {s. \<exists>f. f \<in> A \<and> s = f h} \<or> Inr (r, h2) = Inl e"
using a1 by (metis (lifting) \<open>\<And>aa a. flat_lub (Inl e) {y. \<exists>f\<in>A. y = f aa} = a \<Longrightarrow> a = Inl e \<or> a \<in> {y. \<exists>f\<in>A. y = f aa}\<close> prog.sel)
then show ?thesis
using a2 by fastforce
qed
qed
lemma execute_admissible2:
"ccpo.admissible (fun_lub (flat_lub (Inl (e::'e)))) (fun_ord (flat_ord (Inl e)))
((\<lambda>a. \<forall>(h::'heap) h' h2 h2' (r::'result) r'.
h \<turnstile> a = Inr (r, h2) \<longrightarrow> h' \<turnstile> a = Inr (r', h2') \<longrightarrow> P h h' h2 h2' r r') \<circ> Prog)"
proof (unfold comp_def, rule ccpo.admissibleI, clarify)
fix A :: "('heap \<Rightarrow> 'e + 'result \<times> 'heap) set"
let ?lub = "Prog (fun_lub (flat_lub (Inl e)) A)"
fix h h' h2 h2' r r'
assume 1: "Complete_Partial_Order.chain (fun_ord (flat_ord (Inl e))) A"
and 2 [rule_format]: "\<forall>xa\<in>A. \<forall>h h' h2 h2' r r'. h \<turnstile> Prog xa = Inr (r, h2)
\<longrightarrow> h' \<turnstile> Prog xa = Inr (r', h2') \<longrightarrow> P h h' h2 h2' r r'"
and 4: "h \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r, h2)"
and 5: "h' \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r', h2')"
have h1:"\<And>a. Complete_Partial_Order.chain (flat_ord (Inl e)) {y. \<exists>f\<in>A. y = f a}"
by (rule chain_fun[OF 1])
have "h \<turnstile> ?lub \<in> {y. \<exists>f\<in>A. y = f h}"
using flat_lub_in_chain[OF h1] 4
unfolding execute_def fun_lub_def
by (metis (mono_tags, lifting) Collect_cong Inl_Inr_False prog.sel)
moreover have "h' \<turnstile> ?lub \<in> {y. \<exists>f\<in>A. y = f h'}"
using flat_lub_in_chain[OF h1] 5
unfolding execute_def fun_lub_def
by (metis (no_types, lifting) Collect_cong Inl_Inr_False prog.sel)
ultimately obtain f where
"f \<in> A" and
"h \<turnstile> Prog f = Inr (r, h2)" and
"h' \<turnstile> Prog f = Inr (r', h2')"
using 1 4 5
apply(auto simp add: chain_def fun_ord_def flat_ord_def execute_def)[1]
by (metis Inl_Inr_False)
then show "P h h' h2 h2' r r'"
by(fact 2)
qed
definition dom_prog_ord ::
"('heap, exception, 'result) prog \<Rightarrow> ('heap, exception, 'result) prog \<Rightarrow> bool" where
"dom_prog_ord = img_ord (\<lambda>a b. execute b a) (fun_ord (flat_ord (Inl NonTerminationException)))"
definition dom_prog_lub ::
"('heap, exception, 'result) prog set \<Rightarrow> ('heap, exception, 'result) prog" where
"dom_prog_lub = img_lub (\<lambda>a b. execute b a) Prog (fun_lub (flat_lub (Inl NonTerminationException)))"
lemma dom_prog_lub_empty: "dom_prog_lub {} = error NonTerminationException"
by(simp add: dom_prog_lub_def img_lub_def fun_lub_def flat_lub_def error_def)
lemma dom_prog_interpretation: "partial_function_definitions dom_prog_ord dom_prog_lub"
proof -
have "partial_function_definitions (fun_ord (flat_ord (Inl NonTerminationException)))
(fun_lub (flat_lub (Inl NonTerminationException)))"
by (rule partial_function_lift) (rule flat_interpretation)
then show ?thesis
apply (simp add: dom_prog_lub_def dom_prog_ord_def flat_interpretation execute_def)
using partial_function_image prog.expand prog.sel by blast
qed
interpretation dom_prog: partial_function_definitions dom_prog_ord dom_prog_lub
rewrites "dom_prog_lub {} \<equiv> error NonTerminationException"
by (fact dom_prog_interpretation)(simp add: dom_prog_lub_empty)
lemma admissible_dom_prog:
"dom_prog.admissible (\<lambda>f. \<forall>x h h' r. h \<turnstile> f x \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h' \<longrightarrow> P x h h' r)"
proof (rule admissible_fun[OF dom_prog_interpretation])
fix x
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h' r. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h'
\<longrightarrow> P x h h' r)"
unfolding dom_prog_ord_def dom_prog_lub_def
proof (intro admissible_image partial_function_lift flat_interpretation)
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
(fun_ord (flat_ord (Inl NonTerminationException)))
((\<lambda>a. \<forall>h h' r. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> P x h h' r) \<circ> Prog)"
by(auto simp add: execute_admissible returns_result_def returns_heap_def split: sum.splits)
next
show "\<And>x y. (\<lambda>b. b \<turnstile> x) = (\<lambda>b. b \<turnstile> y) \<Longrightarrow> x = y"
by(simp add: execute_def prog.expand)
next
show "\<And>x. (\<lambda>b. b \<turnstile> Prog x) = x"
by(simp add: execute_def)
qed
qed
lemma admissible_dom_prog2:
"dom_prog.admissible (\<lambda>f. \<forall>x h h2 h' h2' r r2. h \<turnstile> f x \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h'
\<longrightarrow> h2 \<turnstile> f x \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> f x \<rightarrow>\<^sub>h h2' \<longrightarrow> P x h h2 h' h2' r r2)"
proof (rule admissible_fun[OF dom_prog_interpretation])
fix x
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r
\<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>h h2' \<longrightarrow> P x h h2 h' h2' r r2)"
unfolding dom_prog_ord_def dom_prog_lub_def
proof (intro admissible_image partial_function_lift flat_interpretation)
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
(fun_ord (flat_ord (Inl NonTerminationException)))
((\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>h h2'
\<longrightarrow> P x h h2 h' h2' r r2) \<circ> Prog)"
by(auto simp add: returns_result_def returns_heap_def intro!: ccpo.admissibleI
dest!: ccpo.admissibleD[OF execute_admissible2[where P="P x"]]
split: sum.splits)
next
show "\<And>x y. (\<lambda>b. b \<turnstile> x) = (\<lambda>b. b \<turnstile> y) \<Longrightarrow> x = y"
by(simp add: execute_def prog.expand)
next
show "\<And>x. (\<lambda>b. b \<turnstile> Prog x) = x"
by(simp add: execute_def)
qed
qed
lemma fixp_induct_dom_prog:
fixes F :: "'c \<Rightarrow> 'c" and
U :: "'c \<Rightarrow> 'b \<Rightarrow> ('heap, exception, 'result) prog" and
C :: "('b \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> 'c" and
P :: "'b \<Rightarrow> 'heap \<Rightarrow> 'heap \<Rightarrow> 'result \<Rightarrow> bool"
assumes mono: "\<And>x. monotone (fun_ord dom_prog_ord) dom_prog_ord (\<lambda>f. U (F (C f)) x)"
assumes eq: "f \<equiv> C (ccpo.fixp (fun_lub dom_prog_lub) (fun_ord dom_prog_ord) (\<lambda>f. U (F (C f))))"
assumes inverse2: "\<And>f. U (C f) = f"
assumes step: "\<And>f x h h' r. (\<And>x h h' r. h \<turnstile> (U f x) \<rightarrow>\<^sub>r r \<Longrightarrow> h \<turnstile> (U f x) \<rightarrow>\<^sub>h h' \<Longrightarrow> P x h h' r)
\<Longrightarrow> h \<turnstile> (U (F f) x) \<rightarrow>\<^sub>r r \<Longrightarrow> h \<turnstile> (U (F f) x) \<rightarrow>\<^sub>h h' \<Longrightarrow> P x h h' r"
assumes defined: "h \<turnstile> (U f x) \<rightarrow>\<^sub>r r" and "h \<turnstile> (U f x) \<rightarrow>\<^sub>h h'"
shows "P x h h' r"
using step defined dom_prog.fixp_induct_uc[of U F C, OF mono eq inverse2 admissible_dom_prog, of P]
by (metis assms(6) error_returns_heap)
declaration \<open>Partial_Function.init "dom_prog" @{term dom_prog.fixp_fun}
@{term dom_prog.mono_body} @{thm dom_prog.fixp_rule_uc} @{thm dom_prog.fixp_induct_uc}
(SOME @{thm fixp_induct_dom_prog})\<close>
abbreviation "mono_dom_prog \<equiv> monotone (fun_ord dom_prog_ord) dom_prog_ord"
lemma dom_prog_ordI:
assumes "\<And>h. h \<turnstile> f \<rightarrow>\<^sub>e NonTerminationException \<or> h \<turnstile> f = h \<turnstile> g"
shows "dom_prog_ord f g"
proof(auto simp add: dom_prog_ord_def img_ord_def fun_ord_def flat_ord_def)[1]
fix x
assume "x \<turnstile> f \<noteq> x \<turnstile> g"
then show "x \<turnstile> f = Inl NonTerminationException"
using assms[where h=x]
by(auto simp add: returns_error_def split: sum.splits)
qed
lemma dom_prog_ordE:
assumes "dom_prog_ord x y"
obtains "h \<turnstile> x \<rightarrow>\<^sub>e NonTerminationException" | " h \<turnstile> x = h \<turnstile> y"
using assms unfolding dom_prog_ord_def img_ord_def fun_ord_def flat_ord_def
using returns_error_def by force
lemma bind_mono [partial_function_mono]:
fixes B :: "('a \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> ('heap, exception, 'result2) prog"
assumes mf: "mono_dom_prog B" and mg: "\<And>y. mono_dom_prog (\<lambda>f. C y f)"
shows "mono_dom_prog (\<lambda>f. B f \<bind> (\<lambda>y. C y f))"
proof (rule monotoneI)
fix f g :: "'a \<Rightarrow> ('heap, exception, 'result) prog"
assume fg: "dom_prog.le_fun f g"
from mf
have 1: "dom_prog_ord (B f) (B g)" by (rule monotoneD) (rule fg)
from mg
have 2: "\<And>y'. dom_prog_ord (C y' f) (C y' g)" by (rule monotoneD) (rule fg)
have "dom_prog_ord (B f \<bind> (\<lambda>y. C y f)) (B g \<bind> (\<lambda>y. C y f))"
(is "dom_prog_ord ?L ?R")
proof (rule dom_prog_ordI)
fix h
from 1 show "h \<turnstile> ?L \<rightarrow>\<^sub>e NonTerminationException \<or> h \<turnstile> ?L = h \<turnstile> ?R"
apply(rule dom_prog_ordE)
apply(auto)[1]
using bind_cong by fastforce
qed
also
have h1: "dom_prog_ord (B g \<bind> (\<lambda>y'. C y' f)) (B g \<bind> (\<lambda>y'. C y' g))"
(is "dom_prog_ord ?L ?R")
proof (rule dom_prog_ordI)
(* { *)
fix h
show "h \<turnstile> ?L \<rightarrow>\<^sub>e NonTerminationException \<or> h \<turnstile> ?L = h \<turnstile> ?R"
proof (cases "h \<turnstile> ok (B g)")
case True
then obtain x h' where x: "h \<turnstile> B g \<rightarrow>\<^sub>r x" and h': "h \<turnstile> B g \<rightarrow>\<^sub>h h'"
by blast
then have "dom_prog_ord (C x f) (C x g)"
using 2 by simp
then show ?thesis
using x h'
apply(auto intro!: bind_returns_error_I3 dest: returns_result_eq dest!: dom_prog_ordE)[1]
apply(auto simp add: execute_bind_simp)[1]
using "2" dom_prog_ordE by metis
next
case False
then obtain e where e: "h \<turnstile> B g \<rightarrow>\<^sub>e e"
by(simp add: is_OK_def returns_error_def split: sum.splits)
have "h \<turnstile> B g \<bind> (\<lambda>y'. C y' f) \<rightarrow>\<^sub>e e"
using e by(auto)
moreover have "h \<turnstile> B g \<bind> (\<lambda>y'. C y' g) \<rightarrow>\<^sub>e e"
using e by auto
ultimately show ?thesis
using bind_returns_error_eq by metis
qed
qed
finally (dom_prog.leq_trans)
show "dom_prog_ord (B f \<bind> (\<lambda>y. C y f)) (B g \<bind> (\<lambda>y'. C y' g))" .
qed
lemma mono_dom_prog1 [partial_function_mono]:
fixes g :: "('a \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> 'b \<Rightarrow> ('heap, exception, 'result) prog"
assumes "\<And>x. (mono_dom_prog (\<lambda>f. g f x))"
shows "mono_dom_prog (\<lambda>f. map_M (g f) xs)"
using assms
apply (induct xs)
by(auto simp add: call_mono dom_prog.const_mono intro!: bind_mono)
lemma mono_dom_prog2 [partial_function_mono]:
fixes g :: "('a \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> 'b \<Rightarrow> ('heap, exception, 'result) prog"
assumes "\<And>x. (mono_dom_prog (\<lambda>f. g f x))"
shows "mono_dom_prog (\<lambda>f. forall_M (g f) xs)"
using assms
apply (induct xs)
by(auto simp add: call_mono dom_prog.const_mono intro!: bind_mono)
lemma sorted_list_set_cong [simp]:
"sorted_list_of_set (fset FS) = sorted_list_of_set (fset FS') \<longleftrightarrow> FS = FS'"
by auto
end
diff --git a/thys/Core_DOM/common/monads/CharacterDataMonad.thy b/thys/Core_DOM/common/monads/CharacterDataMonad.thy
--- a/thys/Core_DOM/common/monads/CharacterDataMonad.thy
+++ b/thys/Core_DOM/common/monads/CharacterDataMonad.thy
@@ -1,534 +1,534 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>CharacterData\<close>
text\<open>In this theory, we introduce the monadic method setup for the CharacterData class.\<close>
theory CharacterDataMonad
imports
ElementMonad
"../classes/CharacterDataClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog"
global_interpretation l_ptr_kinds_M character_data_ptr_kinds
defines character_data_ptr_kinds_M = a_ptr_kinds_M .
lemmas character_data_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma character_data_ptr_kinds_M_eq:
assumes "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> character_data_ptr_kinds_M|\<^sub>r = |h' \<turnstile> character_data_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: character_data_ptr_kinds_M_defs node_ptr_kinds_M_defs
character_data_ptr_kinds_def)
lemma character_data_ptr_kinds_M_reads:
"reads (\<Union>node_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node_ptr RObject.nothing)}) character_data_ptr_kinds_M h h'"
using node_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs character_data_ptr_kinds_M_defs
character_data_ptr_kinds_def preserved_def)
by (metis (mono_tags, lifting) node_ptr_kinds_small old.unit.exhaust preserved_def)
global_interpretation l_dummy defines get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = "l_get_M.a_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf character_data_ptr_kinds"
apply(simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf l_get_M_def)
by (metis (no_types, opaque_lifting) NodeMonad.get_M_is_l_get_M bind_eq_Some_conv
character_data_ptr_kinds_commutes get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def l_get_M_def option.distinct(1))
lemmas get_M_defs = get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
locale l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
sublocale l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
interpretation l_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf character_data_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf local.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a)
by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ok = get_M_ok[folded get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def]
end
global_interpretation l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf character_data_ptr_kinds get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
rewrites "a_get_M = get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a" defines put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
locale l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
sublocale l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
interpretation l_put_M type_wf character_data_ptr_kinds get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
apply(unfold_locales)
using get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a local.l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms
apply blast
by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ok = put_M_ok[folded put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def]
end
global_interpretation l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf by unfold_locales
lemma CharacterData_simp1 [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = v) \<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp3 [simp]: "
(\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' getter) h h'"
apply(cases "character_data_ptr = character_data_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp4 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp5 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp6 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply (cases "cast character_data_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv split: option.splits)
lemma CharacterData_simp7 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast character_data_ptr = node_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv split: option.splits)
lemma CharacterData_simp8 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def NodeMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp9 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast character_data_ptr \<noteq> node_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
NodeMonad.get_M_defs preserved_def split: option.splits bind_splits
dest: get_heap_E)
lemma CharacterData_simp10 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: NodeMonad.put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def NodeMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp11 [simp]:
"cast character_data_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp12 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast character_data_ptr \<noteq> object_ptr")
apply(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)[1]
lemma CharacterData_simp13 [simp]:
"cast character_data_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma new_element_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def split: prod.splits option.splits
elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Creating CharacterData\<close>
definition new_character_data :: "(_, (_) character_data_ptr) dom_prog"
where
"new_character_data = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h);
return_heap h';
return new_ptr
}"
lemma new_character_data_ok [simp]:
"h \<turnstile> ok new_character_data"
by(auto simp add: new_character_data_def split: prod.splits)
lemma new_character_data_ptr_in_heap:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "new_character_data_ptr |\<in>| character_data_ptr_kinds h'"
using assms
unfolding new_character_data_def
by(auto simp add: new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap
is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_ptr_not_in_heap:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "new_character_data_ptr |\<notin>| character_data_ptr_kinds h"
using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap
by(auto simp add: new_character_data_def split: prod.splits
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_new_ptr:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_new_ptr
by(auto simp add: new_character_data_def split: prod.splits
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_is_character_data_ptr:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "is_character_data_ptr new_character_data_ptr"
using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_is_character_data_ptr
by(auto simp add: new_character_data_def elim!: bind_returns_result_E split: prod.splits)
lemma new_character_data_child_nodes:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "h' \<turnstile> get_M new_character_data_ptr val \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> cast new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> cast new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_character_data_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_CharacterData_ptr_simp [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast character_data_ptr then cast obj else get character_data_ptr h)"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def split: option.splits Option.bind_splits)
lemma Character_data_ptr_kinds_simp [simp]:
"character_data_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = character_data_ptr_kinds h |\<union>|
(if is_character_data_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: character_data_ptr_kinds_def is_node_ptr_kind_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "ElementClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_character_data_ptr_kind ptr \<Longrightarrow> is_character_data_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: ElementMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)[1]
using assms(2) node_ptr_kinds_commutes by blast
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "ElementClass.type_wf h"
assumes "is_character_data_ptr_kind ptr \<Longrightarrow> is_character_data_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
by (metis (no_types, lifting) ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf assms(2) bind.bind_lunit
- cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def notin_fset option.collapse)
+ cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def option.collapse)
subsection\<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I split: if_splits)[1]
using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast
- using element_ptrs_def apply fastforce
+ using element_ptrs_def apply force
using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_element_ptr_ref)
lemma new_element_is_l_new_element: "l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma new_character_data_type_wf_preserved [simp]:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(simp_all add: type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs is_node_kind_def)
by (meson new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap)
locale l_new_character_data = l_type_wf +
assumes new_character_data_types_preserved: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_character_data_is_l_new_character_data: "l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_val_type_wf_preserved [simp]:
"h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
CharacterDataClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e CharacterDataClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs
ObjectClass.a_type_wf_def
split: option.splits)[1]
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
+ apply metis
done
lemma character_data_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "character_data_ptr_kinds h = character_data_ptr_kinds h'"
by(simp add: character_data_ptr_kinds_def node_ptr_kinds_def preserved_def
object_ptr_kinds_preserved_small[OF assms])
lemma character_data_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "character_data_ptr_kinds h = character_data_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def character_data_ptr_kinds_def)
by (metis assms node_ptr_kinds_preserved)
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr
RCharacterData.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3)]
allI[OF assms(4), of id, simplified] character_data_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs character_data_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply(force)
by force
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr
RCharacterData.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_def ElementMonad.type_wf_drop
l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a.a_type_wf_def)[1]
using type_wf_drop
by (metis (no_types, lifting) ElementClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
- character_data_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
+ character_data_ptr_kinds_commutes fmlookup_drop get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def node_ptr_kinds_commutes object_ptr_kinds_code5)
end
diff --git a/thys/Core_DOM/common/monads/DocumentMonad.thy b/thys/Core_DOM/common/monads/DocumentMonad.thy
--- a/thys/Core_DOM/common/monads/DocumentMonad.thy
+++ b/thys/Core_DOM/common/monads/DocumentMonad.thy
@@ -1,614 +1,606 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Document\<close>
text\<open>In this theory, we introduce the monadic method setup for the Document class.\<close>
theory DocumentMonad
imports
CharacterDataMonad
"../classes/DocumentClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'result) dom_prog"
global_interpretation l_ptr_kinds_M document_ptr_kinds defines document_ptr_kinds_M = a_ptr_kinds_M .
lemmas document_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma document_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: document_ptr_kinds_M_defs object_ptr_kinds_M_defs document_ptr_kinds_def)
lemma document_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) document_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def object_ptr_kinds_M_defs document_ptr_kinds_M_defs
document_ptr_kinds_def preserved_def cong del: image_cong_simp)
apply (metis (mono_tags, opaque_lifting) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def)
done
global_interpretation l_dummy defines get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = "l_get_M.a_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf document_ptr_kinds"
apply(simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf l_get_M_def)
by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs bind_eq_None_conv
document_ptr_kinds_commutes get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.simps(3))
lemmas get_M_defs = get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf document_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson DocumentMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = get_M_ok[folded get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf document_ptr_kinds get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
rewrites "a_get_M = get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t" defines put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_put_M type_wf document_ptr_kinds get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
apply (simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson DocumentMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = put_M_ok[folded put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
lemma document_put_get [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Mdocument_preserved1 [simp]:
"document_ptr \<noteq> document_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma document_put_get_preserved [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' getter) h h'"
apply(cases "document_ptr = document_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved2 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved3 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved4 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast document_ptr \<noteq> object_ptr")[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Mdocument_preserved5 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved7 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved8 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved10 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast document_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection \<open>Creating Documents\<close>
definition new_document :: "(_, (_) document_ptr) dom_prog"
where
"new_document = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h);
return_heap h';
return new_ptr
}"
lemma new_document_ok [simp]:
"h \<turnstile> ok new_document"
by(auto simp add: new_document_def split: prod.splits)
lemma new_document_ptr_in_heap:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "new_document_ptr |\<in>| document_ptr_kinds h'"
using assms
unfolding new_document_def
by(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_ptr_not_in_heap:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "new_document_ptr |\<notin>| document_ptr_kinds h"
using assms new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap
by(auto simp add: new_document_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_new_ptr:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_document_ptr|}"
using assms new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr
by(auto simp add: new_document_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_is_document_ptr:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "is_document_ptr new_document_ptr"
using assms new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_document_ptr
by(auto simp add: new_document_def elim!: bind_returns_result_E split: prod.splits)
lemma new_document_doctype:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr doctype \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_document_element:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr document_element \<rightarrow>\<^sub>r None"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_disconnected_nodes:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr disconnected_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> ptr \<noteq> cast new_document_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_document_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new_document_def CharacterDataMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> ptr \<noteq> new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection \<open>Modified Heaps\<close>
lemma get_document_ptr_simp [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast document_ptr then cast obj else get document_ptr h)"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: option.splits Option.bind_splits)
lemma document_ptr_kidns_simp [simp]:
"document_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= document_ptr_kinds h |\<union>| (if is_document_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: document_ptr_kinds_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "CharacterDataClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_document_ptr_kind ptr \<Longrightarrow> is_document_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs is_document_kind_def split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "CharacterDataClass.type_wf h"
assumes "is_document_ptr_kind ptr \<Longrightarrow> is_document_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1]
by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
- is_document_kind_def notin_fset option.exhaust_sel)
+ is_document_kind_def option.exhaust_sel)
subsection \<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def element_ptrs_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
- apply fastforce
+ apply force
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_element_ptr_ref)
lemma new_element_is_l_new_element [instances]:
"l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma new_character_data_type_wf_preserved [simp]:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2 bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
by (meson new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap)
lemma new_character_data_is_l_new_character_data [instances]:
"l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_val_type_wf_preserved [simp]:
"h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t is_node_kind_def
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis bind.bind_lzero finite_set_in get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def option.distinct(1) option.exhaust_sel)
- by (metis finite_set_in)
+ apply (metis bind.bind_lzero get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def option.distinct(1) option.exhaust_sel)
+ by metis
lemma new_document_type_wf_preserved [simp]: "h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(auto simp add: type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_document_kind_def
split: option.splits)[1]
- using document_ptrs_def apply fastforce
+ using document_ptrs_def apply force
apply (simp add: is_document_kind_def)
- apply (metis Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge ffmember_filter
- fimage_eqI is_document_ptr_ref)
- done
+ by (metis (no_types, opaque_lifting) Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge
+ ffmember_filter fimageI is_document_ptr_ref)
locale l_new_document = l_type_wf +
assumes new_document_types_preserved: "h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_document_is_l_new_document [instances]: "l_new_document type_wf"
using l_new_document.intro new_document_type_wf_preserved
by blast
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doctype_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr doctype_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: get_M_defs)[1]
- by (metis (mono_tags) error_returns_result finite_set_in option.exhaust_sel option.simps(4))
+ by (metis (mono_tags) error_returns_result option.exhaust_sel option.simps(4))
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr document_element_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t is_node_ptr_kind_none
cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none is_document_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: get_M_defs is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs
split: option.splits)[1]
- by (metis finite_set_in)
+ by metis
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr disconnected_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none
cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none is_document_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def get_M_defs type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- by (metis finite_set_in)
+ by metis
lemma document_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "document_ptr_kinds h = document_ptr_kinds h'"
by(simp add: document_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms])
lemma document_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "document_ptr_kinds h = document_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def document_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved
(get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
shows "DocumentClass.type_wf h = DocumentClass.type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3) assms(4)]
allI[OF assms(5), of id, simplified] document_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs )[1]
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply force
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
by force
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved
(get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
shows "DocumentClass.type_wf h = DocumentClass.type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> DocumentClass.type_wf h = DocumentClass.type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs)[1]
using type_wf_drop
apply blast
- by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf CharacterDataMonad.type_wf_drop
- document_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel)
+ by (metis (no_types, opaque_lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
+ CharacterDataMonad.type_wf_drop document_ptr_kinds_commutes fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
+ get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel)
end
diff --git a/thys/Core_DOM/common/monads/ElementMonad.thy b/thys/Core_DOM/common/monads/ElementMonad.thy
--- a/thys/Core_DOM/common/monads/ElementMonad.thy
+++ b/thys/Core_DOM/common/monads/ElementMonad.thy
@@ -1,445 +1,445 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Element\<close>
text\<open>In this theory, we introduce the monadic method setup for the Element class.\<close>
theory ElementMonad
imports
NodeMonad
"ElementClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element,'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element,'result) dom_prog"
global_interpretation l_ptr_kinds_M element_ptr_kinds defines element_ptr_kinds_M = a_ptr_kinds_M .
lemmas element_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma element_ptr_kinds_M_eq:
assumes "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> element_ptr_kinds_M|\<^sub>r = |h' \<turnstile> element_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: element_ptr_kinds_M_defs node_ptr_kinds_M_defs element_ptr_kinds_def)
lemma element_ptr_kinds_M_reads:
"reads (\<Union>element_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t element_ptr RObject.nothing)}) element_ptr_kinds_M h h'"
apply (simp add: reads_def node_ptr_kinds_M_defs element_ptr_kinds_M_defs element_ptr_kinds_def
node_ptr_kinds_M_reads preserved_def cong del: image_cong_simp)
apply (metis (mono_tags, opaque_lifting) node_ptr_kinds_small old.unit.exhaust preserved_def)
done
global_interpretation l_dummy defines get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = "l_get_M.a_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf element_ptr_kinds"
apply(simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf l_get_M_def)
by (metis (no_types, lifting) ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs
bind_eq_Some_conv bind_eq_Some_conv element_ptr_kinds_commutes get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def node_ptr_kinds_commutes option.simps(3))
lemmas get_M_defs = get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas by unfold_locales
interpretation l_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf element_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson ElementMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = get_M_ok[folded get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
lemmas get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap = get_M_ptr_in_heap[folded get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf element_ptr_kinds get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
rewrites "a_get_M = get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t"
defines put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas by unfold_locales
interpretation l_put_M type_wf element_ptr_kinds get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
apply (simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson ElementMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = put_M_ok[folded put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
lemma element_put_get [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Element_preserved1 [simp]:
"element_ptr \<noteq> element_ptr' \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma element_put_get_preserved [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = getter x) \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' getter) h h'"
apply(cases "element_ptr = element_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved3 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast element_ptr = object_ptr")
by (auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma get_M_Element_preserved4 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast element_ptr = node_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma get_M_Element_preserved5 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast element_ptr \<noteq> node_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Element_preserved7 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: NodeMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved8 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast element_ptr \<noteq> object_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Element_preserved10 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
subsection\<open>Creating Elements\<close>
definition new_element :: "(_, (_) element_ptr) dom_prog"
where
"new_element = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h);
return_heap h';
return new_ptr
}"
lemma new_element_ok [simp]:
"h \<turnstile> ok new_element"
by(auto simp add: new_element_def split: prod.splits)
lemma new_element_ptr_in_heap:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "new_element_ptr |\<in>| element_ptr_kinds h'"
using assms
unfolding new_element_def
by(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_ptr_not_in_heap:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "new_element_ptr |\<notin>| element_ptr_kinds h"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
bind_returns_heap_E)
lemma new_element_new_ptr:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
bind_returns_heap_E)
lemma new_element_is_element_ptr:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "is_element_ptr new_element_ptr"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_element_ptr
by(auto simp add: new_element_def elim!: bind_returns_result_E split: prod.splits)
lemma new_element_child_nodes:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr child_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_tag_name:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr tag_name \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_attrs:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr attrs \<rightarrow>\<^sub>r fmempty"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_shadow_root_opt:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr shadow_root_opt \<rightarrow>\<^sub>r None"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_element_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_Element_ptr_simp [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast element_ptr then cast obj else get element_ptr h)"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: option.splits Option.bind_splits)
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= element_ptr_kinds h |\<union>| (if is_element_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: element_ptr_kinds_def is_node_ptr_kind_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "NodeClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_element_ptr_kind ptr \<Longrightarrow> is_element_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: NodeMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)[1]
using assms(2) node_ptr_kinds_commutes by blast
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "NodeClass.type_wf h"
assumes "is_element_ptr_kind ptr \<Longrightarrow> is_element_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
by (metis (no_types, lifting) NodeClass.l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_axioms assms(2) bind.bind_lunit
- cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
+ cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf option.collapse)
subsection\<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]: "h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
new_element_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits if_splits elim!: bind_returns_heap_E)[1]
- apply (metis element_ptr_kinds_commutes element_ptrs_def fempty_iff ffmember_filter finite_set_in
+ apply (metis element_ptr_kinds_commutes element_ptrs_def fempty_iff ffmember_filter
is_element_ptr_ref)
- apply (metis element_ptrs_def fempty_iff ffmember_filter finite_set_in is_element_ptr_ref)
+ apply (metis element_ptrs_def fempty_iff ffmember_filter is_element_ptr_ref)
apply (metis (no_types, lifting) Suc_n_not_le_n element_ptr.sel(1) element_ptr_kinds_commutes
- element_ptrs_def fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref notin_fset)
+ element_ptrs_def fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref)
apply (metis (no_types, lifting) Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def
- fMax_ge ffmember_filter fimage_eqI finite_set_in is_element_ptr_ref)
+ fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref)
done
locale l_new_element = l_type_wf +
assumes new_element_types_preserved: "h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_element_is_l_new_element: "l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs Let_def
put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M_pointers_preserved:
assumes "h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
using assms
apply(auto simp add: put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
elim!: bind_returns_heap_E2 dest!: get_heap_E)[1]
by (meson get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap is_OK_returns_result_I)
lemma element_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "element_ptr_kinds h = element_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def element_ptr_kinds_def)
by (metis assms node_ptr_kinds_preserved)
lemma element_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "element_ptr_kinds h = element_ptr_kinds h'"
by(simp add: element_ptr_kinds_def node_ptr_kinds_def preserved_def
object_ptr_kinds_preserved_small[OF assms])
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2)] allI[OF assms(3), of id, simplified]
apply(auto simp add: type_wf_defs )[1]
apply(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
split: option.splits,force)[1]
by(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
split: option.splits,force)
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
node_ptr_kinds_def object_ptr_kinds_def is_node_ptr_kind_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)[1]
- apply (metis (no_types, lifting) element_ptr_kinds_commutes finite_set_in fmdom_notD fmdom_notI
+ apply (metis (no_types, lifting) element_ptr_kinds_commutes fmdom_notD fmdom_notI
fmlookup_drop heap.sel node_ptr_kinds_commutes o_apply object_ptr_kinds_def)
by (metis element_ptr_kinds_commutes fmdom_notI fmdrop_lookup heap.sel node_ptr_kinds_commutes
o_apply object_ptr_kinds_def)
end
diff --git a/thys/Core_DOM/common/monads/NodeMonad.thy b/thys/Core_DOM/common/monads/NodeMonad.thy
--- a/thys/Core_DOM/common/monads/NodeMonad.thy
+++ b/thys/Core_DOM/common/monads/NodeMonad.thy
@@ -1,218 +1,218 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Node\<close>
text\<open>In this theory, we introduce the monadic method setup for the Node class.\<close>
theory NodeMonad
imports
ObjectMonad
"../classes/NodeClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog"
global_interpretation l_ptr_kinds_M node_ptr_kinds defines node_ptr_kinds_M = a_ptr_kinds_M .
lemmas node_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma node_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: node_ptr_kinds_M_defs object_ptr_kinds_M_defs node_ptr_kinds_def)
global_interpretation l_dummy defines get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = "l_get_M.a_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf node_ptr_kinds"
apply(simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf l_get_M_def)
by (metis ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind_eq_None_conv get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
node_ptr_kinds_commutes option.simps(3))
lemmas get_M_defs = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e
locale l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
sublocale l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales
interpretation l_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf node_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf local.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e)
by (meson NodeMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ok = get_M_ok[folded get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def]
end
global_interpretation l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales
lemma node_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) node_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_def
object_ptr_kinds_M_reads preserved_def)
by (metis (mono_tags, lifting) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def)
global_interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
rewrites "a_get_M = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e"
defines put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e
locale l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
sublocale l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales
interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
apply(unfold_locales)
apply (simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf local.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e)
by (meson NodeMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ok = put_M_ok[folded put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def]
end
global_interpretation l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales
lemma get_M_Object_preserved1 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x)) \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast node_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv
split: option.splits)
lemma get_M_Object_preserved2 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Object_preserved3 [simp]:
"h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast node_ptr \<noteq> object_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Object_preserved4 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_node_ptr_simp [simp]:
"get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = cast node_ptr then cast obj else get node_ptr h)"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= node_ptr_kinds h |\<union>| (if is_node_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: node_ptr_kinds_def)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "ObjectClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_node_ptr_kind ptr \<Longrightarrow> is_node_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
apply(auto simp add: type_wf_defs split: option.splits)[1]
using cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none is_node_kind_def apply blast
using cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none is_node_kind_def apply blast
done
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: ObjectMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "ObjectClass.type_wf h"
assumes "is_node_ptr_kind ptr \<Longrightarrow> is_node_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
- by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit finite_set_in get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def is_node_kind_def option.exhaust_sel)
+ by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def is_node_kind_def option.exhaust_sel)
subsection\<open>Preserving Types\<close>
lemma node_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "node_ptr_kinds h = node_ptr_kinds h'"
by(simp add: node_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms])
lemma node_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "node_ptr_kinds h = node_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def node_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved allI[OF assms(2), of id, simplified]
apply(auto simp add: type_wf_defs)[1]
apply(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
- apply (metis notin_fset option.simps(3))
+ apply (metis option.simps(3))
by(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits, force)[1]
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
end
diff --git a/thys/Core_DOM/common/monads/ObjectMonad.thy b/thys/Core_DOM/common/monads/ObjectMonad.thy
--- a/thys/Core_DOM/common/monads/ObjectMonad.thy
+++ b/thys/Core_DOM/common/monads/ObjectMonad.thy
@@ -1,258 +1,258 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Object\<close>
text\<open>In this theory, we introduce the monadic method setup for the Object class.\<close>
theory ObjectMonad
imports
BaseMonad
"../classes/ObjectClass"
begin
type_synonym ('object_ptr, 'Object, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'Object, 'result) dom_prog"
global_interpretation l_ptr_kinds_M object_ptr_kinds defines object_ptr_kinds_M = a_ptr_kinds_M .
lemmas object_ptr_kinds_M_defs = a_ptr_kinds_M_def
global_interpretation l_dummy defines get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = "l_get_M.a_get_M get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t type_wf object_ptr_kinds"
by (simp add: a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf l_get_M_def)
lemmas get_M_defs = get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
locale l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas = l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
interpretation l_get_M get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t type_wf object_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf local.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
by (simp add: a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
lemmas get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok = get_M_ok[folded get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
lemmas get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap = get_M_ptr_in_heap[folded get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas type_wf
by (simp add: l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_def l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms)
lemma object_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) object_ptr_kinds_M h h'"
apply(auto simp add: object_ptr_kinds_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf type_wf_defs reads_def
preserved_def get_M_defs
split: option.splits)[1]
using a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf by blast+
global_interpretation l_put_M type_wf object_ptr_kinds get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
rewrites "a_get_M = get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"
defines put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
locale l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas = l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
interpretation l_put_M type_wf object_ptr_kinds get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
apply(unfold_locales)
using get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t local.l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms apply blast
by (simp add: a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
lemmas put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok = put_M_ok[folded put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
lemmas put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap = put_M_ptr_in_heap[folded put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas type_wf
by (simp add: l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_def l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms)
definition check_in_heap :: "(_) object_ptr \<Rightarrow> (_, unit) dom_prog"
where
"check_in_heap ptr = do {
h \<leftarrow> get_heap;
(if ptr |\<in>| object_ptr_kinds h then
return ()
else
error SegmentationFault
)}"
lemma check_in_heap_ptr_in_heap: "ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> h \<turnstile> ok (check_in_heap ptr)"
by(auto simp add: check_in_heap_def)
lemma check_in_heap_pure [simp]: "pure (check_in_heap ptr) h"
by(auto simp add: check_in_heap_def intro!: bind_pure_I)
lemma check_in_heap_is_OK [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (check_in_heap ptr \<bind> f) = h \<turnstile> ok (f ())"
by(simp add: check_in_heap_def)
lemma check_in_heap_returns_result [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> (check_in_heap ptr \<bind> f) \<rightarrow>\<^sub>r x = h \<turnstile> f () \<rightarrow>\<^sub>r x"
by(simp add: check_in_heap_def)
lemma check_in_heap_returns_heap [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> (check_in_heap ptr \<bind> f) \<rightarrow>\<^sub>h h' = h \<turnstile> f () \<rightarrow>\<^sub>h h'"
by(simp add: check_in_heap_def)
lemma check_in_heap_reads:
"reads {preserved (get_M object_ptr nothing)} (check_in_heap object_ptr) h h'"
apply(simp add: check_in_heap_def reads_def preserved_def)
by (metis a_type_wf_def get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap is_OK_returns_result_E
is_OK_returns_result_I unit_all_impI)
subsection\<open>Invoke\<close>
fun invoke_rec :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list \<Rightarrow> (_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog"
where
"invoke_rec ((P, f)#xs) ptr args = (if P ptr then f ptr args else invoke_rec xs ptr args)"
| "invoke_rec [] ptr args = error InvokeError"
definition invoke :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list
\<Rightarrow> (_) object_ptr \<Rightarrow> 'args \<Rightarrow> (_, 'result) dom_prog"
where
"invoke xs ptr args = do { check_in_heap ptr; invoke_rec xs ptr args}"
lemma invoke_split: "P (invoke ((Pred, f) # xs) ptr args) =
((\<not>(Pred ptr) \<longrightarrow> P (invoke xs ptr args))
\<and> (Pred ptr \<longrightarrow> P (do {check_in_heap ptr; f ptr args})))"
by(simp add: invoke_def)
lemma invoke_split_asm: "P (invoke ((Pred, f) # xs) ptr args) =
(\<not>((\<not>(Pred ptr) \<and> (\<not> P (invoke xs ptr args)))
\<or> (Pred ptr \<and> (\<not> P (do {check_in_heap ptr; f ptr args})))))"
by(simp add: invoke_def)
lemmas invoke_splits = invoke_split invoke_split_asm
lemma invoke_ptr_in_heap: "h \<turnstile> ok (invoke xs ptr args) \<Longrightarrow> ptr |\<in>| object_ptr_kinds h"
by (metis bind_is_OK_E check_in_heap_ptr_in_heap invoke_def is_OK_returns_heap_I)
lemma invoke_pure [simp]: "pure (invoke [] ptr args) h"
by(auto simp add: invoke_def intro!: bind_pure_I)
lemma invoke_is_OK [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> ok (invoke ((Pred, f) # xs) ptr args) = h \<turnstile> ok (f ptr args)"
by(simp add: invoke_def)
lemma invoke_returns_result [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> (invoke ((Pred, f) # xs) ptr args) \<rightarrow>\<^sub>r x = h \<turnstile> f ptr args \<rightarrow>\<^sub>r x"
by(simp add: invoke_def)
lemma invoke_returns_heap [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> (invoke ((Pred, f) # xs) ptr args) \<rightarrow>\<^sub>h h' = h \<turnstile> f ptr args \<rightarrow>\<^sub>h h'"
by(simp add: invoke_def)
lemma invoke_not [simp]: "\<not>Pred ptr \<Longrightarrow> invoke ((Pred, f) # xs) ptr args = invoke xs ptr args"
by(auto simp add: invoke_def)
lemma invoke_empty [simp]: "\<not>h \<turnstile> ok (invoke [] ptr args)"
by(auto simp add: invoke_def check_in_heap_def)
lemma invoke_empty_reads [simp]: "\<forall>P \<in> S. reflp P \<and> transp P \<Longrightarrow> reads S (invoke [] ptr args) h h'"
apply(simp add: invoke_def reads_def preserved_def)
by (meson bind_returns_result_E error_returns_result)
subsection\<open>Modified Heaps\<close>
lemma get_object_ptr_simp [simp]:
"get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = object_ptr then Some obj else get object_ptr h)"
by(auto simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: option.splits Option.bind_splits)
lemma object_ptr_kinds_simp [simp]: "object_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = object_ptr_kinds h |\<union>| {|ptr|}"
by(auto simp add: object_ptr_kinds_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs split: option.splits if_splits)
subsection\<open>Preserving Types\<close>
lemma type_wf_preserved: "type_wf h = type_wf h'"
by(auto simp add: type_wf_defs)
lemma object_ptr_kinds_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
using assms
apply(auto simp add: object_ptr_kinds_def preserved_def get_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: option.splits)[1]
- apply (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember_iff_member_fset
+ apply (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq
old.unit.exhaust option.case_eq_if return_returns_result)
- by (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember_iff_member_fset
+ by (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq
old.unit.exhaust option.case_eq_if return_returns_result)
lemma object_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w object_ptr. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
proof -
{
fix object_ptr w
have "preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
apply(rule writes_small_big[OF assms])
by auto
}
then show ?thesis
using object_ptr_kinds_preserved_small by blast
qed
lemma reads_writes_preserved2:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' x. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
shows "preserved (get_M ptr getter) h h'"
apply(clarsimp simp add: preserved_def)
using reads_singleton assms(1) assms(2)
apply(rule reads_writes_preserved)
using assms(3)
by(auto simp add: preserved_def)
end
diff --git a/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy b/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy
--- a/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy
+++ b/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy
@@ -1,8044 +1,8042 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Wellformedness\<close>
text\<open>In this theory, we discuss the wellformedness of the DOM. First, we define
wellformedness and, second, we show for all functions for querying and modifying the
DOM to what extend they preserve wellformendess.\<close>
theory Core_DOM_Heap_WF
imports
"Core_DOM_Functions"
begin
locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs =
l_get_child_nodes_defs get_child_nodes get_child_nodes_locs +
l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs
for get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_owner_document_valid :: "(_) heap \<Rightarrow> bool"
where
"a_owner_document_valid h \<longleftrightarrow> (\<forall>node_ptr \<in> fset (node_ptr_kinds h).
((\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h
\<and> node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
\<or> (\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h
\<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)))"
lemma a_owner_document_valid_code [code]: "a_owner_document_valid h \<longleftrightarrow> node_ptr_kinds h |\<subseteq>|
fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)) @ map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))
"
apply(auto simp add: a_owner_document_valid_def
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_owner_document_valid_def)[1]
proof -
fix x
assume 1: " \<forall>node_ptr\<in>fset (node_ptr_kinds h).
(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
assume 2: "x |\<in>| node_ptr_kinds h"
assume 3: "x |\<notin>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
have "\<not>(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
x \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
using 1 2 3
- by (smt (verit) UN_I fset_of_list_elem image_eqI notin_fset set_concat set_map sorted_list_of_fset_simps(1))
+ by (smt (verit) UN_I fset_of_list_elem image_eqI set_concat set_map sorted_list_of_fset_simps(1))
then
have "(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
using 1 2
by auto
then obtain parent_ptr where parent_ptr:
"parent_ptr |\<in>| object_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
by auto
moreover have "parent_ptr \<in> set (sorted_list_of_fset (object_ptr_kinds h))"
using parent_ptr by auto
moreover have "|h \<turnstile> get_child_nodes parent_ptr|\<^sub>r \<in> set (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))"
using calculation(2) by auto
ultimately
show "x |\<in>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h))))"
using fset_of_list_elem by fastforce
next
fix node_ptr
assume 1: "node_ptr_kinds h |\<subseteq>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))) |\<union>|
fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
assume 2: "node_ptr |\<in>| node_ptr_kinds h"
assume 3: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<longrightarrow>
node_ptr \<notin> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
have "node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))) \<or>
node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
using 1 2
by (meson fin_mono fset_of_list_elem funion_iff)
then
show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using 3
by auto
qed
definition a_parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
where
"a_parent_child_rel h = {(parent, child). parent |\<in>| object_ptr_kinds h
\<and> child \<in> cast ` set |h \<turnstile> get_child_nodes parent|\<^sub>r}"
lemma a_parent_child_rel_code [code]: "a_parent_child_rel h = set (concat (map
(\<lambda>parent. map
(\<lambda>child. (parent, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child))
|h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))
)"
by(auto simp add: a_parent_child_rel_def)
definition a_acyclic_heap :: "(_) heap \<Rightarrow> bool"
where
"a_acyclic_heap h = acyclic (a_parent_child_rel h)"
definition a_all_ptrs_in_heap :: "(_) heap \<Rightarrow> bool"
where
"a_all_ptrs_in_heap h \<longleftrightarrow>
(\<forall>ptr \<in> fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h)) \<and>
(\<forall>document_ptr \<in> fset (document_ptr_kinds h).
set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h))"
definition a_distinct_lists :: "(_) heap \<Rightarrow> bool"
where
"a_distinct_lists h = distinct (concat (
(map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r)
@ (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) |h \<turnstile> document_ptr_kinds_M|\<^sub>r)
))"
definition a_heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
where
"a_heap_is_wellformed h \<longleftrightarrow>
a_acyclic_heap h \<and> a_all_ptrs_in_heap h \<and> a_distinct_lists h \<and> a_owner_document_valid h"
end
locale l_heap_is_wellformed_defs =
fixes heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
fixes parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
global_interpretation l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
defines heap_is_wellformed = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_heap_is_wellformed get_child_nodes
get_disconnected_nodes"
and parent_child_rel = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_parent_child_rel get_child_nodes"
and acyclic_heap = a_acyclic_heap
and all_ptrs_in_heap = a_all_ptrs_in_heap
and distinct_lists = a_distinct_lists
and owner_document_valid = a_owner_document_valid
.
locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs
+ l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs
+ l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel
+ l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set" +
assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed"
assumes parent_child_rel_impl: "parent_child_rel = a_parent_child_rel"
begin
lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def]
lemmas parent_child_rel_def = parent_child_rel_impl[unfolded a_parent_child_rel_def]
lemmas acyclic_heap_def = a_acyclic_heap_def[folded parent_child_rel_impl]
lemma parent_child_rel_node_ptr:
"(parent, child) \<in> parent_child_rel h \<Longrightarrow> is_node_ptr_kind child"
by(auto simp add: parent_child_rel_def)
lemma parent_child_rel_child_nodes:
assumes "known_ptr parent"
and "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
and "child \<in> set children"
shows "(parent, cast child) \<in> parent_child_rel h"
using assms
apply(auto simp add: parent_child_rel_def is_OK_returns_result_I )[1]
using get_child_nodes_ptr_in_heap by blast
lemma parent_child_rel_child_nodes2:
assumes "known_ptr parent"
and "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
and "child \<in> set children"
and "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = child_obj"
shows "(parent, child_obj) \<in> parent_child_rel h"
using assms parent_child_rel_child_nodes by blast
lemma parent_child_rel_finite: "finite (parent_child_rel h)"
proof -
have "parent_child_rel h = (\<Union>ptr \<in> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r.
(\<Union>child \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r. {(ptr, cast child)}))"
by(auto simp add: parent_child_rel_def)
moreover have "finite (\<Union>ptr \<in> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r.
(\<Union>child \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r. {(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)}))"
by simp
ultimately show ?thesis
by simp
qed
lemma distinct_lists_no_parent:
assumes "a_distinct_lists h"
assumes "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "node_ptr \<in> set disc_nodes"
shows "\<not>(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h
\<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
using assms
apply(auto simp add: a_distinct_lists_def)[1]
proof -
fix parent_ptr :: "(_) object_ptr"
assume a1: "parent_ptr |\<in>| object_ptr_kinds h"
assume a2: "(\<Union>x\<in>fset (object_ptr_kinds h).
set |h \<turnstile> get_child_nodes x|\<^sub>r) \<inter> (\<Union>x\<in>fset (document_ptr_kinds h).
set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
assume a3: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume a4: "node_ptr \<in> set disc_nodes"
assume a5: "node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
have f6: "parent_ptr \<in> fset (object_ptr_kinds h)"
using a1 by auto
have f7: "document_ptr \<in> fset (document_ptr_kinds h)"
- using a3 by (meson fmember_iff_member_fset get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I)
+ using a3 by (meson get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I)
have "|h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes"
using a3 by simp
then show False
using f7 f6 a5 a4 a2 by blast
qed
lemma distinct_lists_disconnected_nodes:
assumes "a_distinct_lists h"
and "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
shows "distinct disc_nodes"
proof -
have h1: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
|h \<turnstile> document_ptr_kinds_M|\<^sub>r))"
using assms(1)
by(simp add: a_distinct_lists_def)
then show ?thesis
using concat_map_all_distinct[OF h1] assms(2) is_OK_returns_result_I get_disconnected_nodes_ok
by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M
l_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap
l_get_disconnected_nodes_axioms select_result_I2)
qed
lemma distinct_lists_children:
assumes "a_distinct_lists h"
and "known_ptr ptr"
and "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
shows "distinct children"
proof (cases "children = []", simp)
assume "children \<noteq> []"
have h1: "distinct (concat ((map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r)))"
using assms(1)
by(simp add: a_distinct_lists_def)
show ?thesis
using concat_map_all_distinct[OF h1] assms(2) assms(3)
by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M get_child_nodes_ptr_in_heap
is_OK_returns_result_I select_result_I2)
qed
lemma heap_is_wellformed_children_in_heap:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "child |\<in>| node_ptr_kinds h"
using assms
apply(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def)[1]
- by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I
+ by (metis (no_types, opaque_lifting) is_OK_returns_result_I
local.get_child_nodes_ptr_in_heap select_result_I2 subsetD)
lemma heap_is_wellformed_one_parent:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'"
assumes "set children \<inter> set children' \<noteq> {}"
shows "ptr = ptr'"
using assms
proof (auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1]
fix x :: "(_) node_ptr"
assume a1: "ptr \<noteq> ptr'"
assume a2: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assume a3: "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'"
assume a4: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
have f5: "|h \<turnstile> get_child_nodes ptr|\<^sub>r = children"
using a2 by simp
have "|h \<turnstile> get_child_nodes ptr'|\<^sub>r = children'"
using a3 by (meson select_result_I2)
then have "ptr \<notin> set (sorted_list_of_set (fset (object_ptr_kinds h)))
\<or> ptr' \<notin> set (sorted_list_of_set (fset (object_ptr_kinds h)))
\<or> set children \<inter> set children' = {}"
using f5 a4 a1 by (meson distinct_concat_map_E(1))
then show False
- using a3 a2 by (metis (no_types) assms(4) finite_fset fmember_iff_member_fset is_OK_returns_result_I
+ using a3 a2 by (metis (no_types) assms(4) finite_fset is_OK_returns_result_I
local.get_child_nodes_ptr_in_heap set_sorted_list_of_set)
qed
lemma parent_child_rel_child:
"h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<longleftrightarrow> (ptr, cast child) \<in> parent_child_rel h"
by (simp add: is_OK_returns_result_I get_child_nodes_ptr_in_heap parent_child_rel_def)
lemma parent_child_rel_acyclic: "heap_is_wellformed h \<Longrightarrow> acyclic (parent_child_rel h)"
by (simp add: acyclic_heap_def local.heap_is_wellformed_def)
lemma heap_is_wellformed_disconnected_nodes_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
distinct disc_nodes"
using distinct_lists_disconnected_nodes local.heap_is_wellformed_def by blast
lemma parent_child_rel_parent_in_heap:
"(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> parent |\<in>| object_ptr_kinds h"
using local.parent_child_rel_def by blast
lemma parent_child_rel_child_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptr parent
\<Longrightarrow> (parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> child_ptr |\<in>| object_ptr_kinds h"
apply(auto simp add: heap_is_wellformed_def parent_child_rel_def a_all_ptrs_in_heap_def)[1]
using get_child_nodes_ok
- by (meson finite_set_in subsetD)
+ by (meson subsetD)
lemma heap_is_wellformed_disc_nodes_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> node \<in> set disc_nodes \<Longrightarrow> node |\<in>| node_ptr_kinds h"
- by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I local.a_all_ptrs_in_heap_def
+ by (metis (no_types, opaque_lifting) is_OK_returns_result_I local.a_all_ptrs_in_heap_def
local.get_disconnected_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2 subsetD)
lemma heap_is_wellformed_one_disc_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes'
\<Longrightarrow> set disc_nodes \<inter> set disc_nodes' \<noteq> {} \<Longrightarrow> document_ptr = document_ptr'"
using DocumentMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append distinct_concat_map_E(1)
is_OK_returns_result_I local.a_distinct_lists_def local.get_disconnected_nodes_ptr_in_heap
local.heap_is_wellformed_def select_result_I2
proof -
assume a1: "heap_is_wellformed h"
assume a2: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume a3: "h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes'"
assume a4: "set disc_nodes \<inter> set disc_nodes' \<noteq> {}"
have f5: "|h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes"
using a2 by (meson select_result_I2)
have f6: "|h \<turnstile> get_disconnected_nodes document_ptr'|\<^sub>r = disc_nodes'"
using a3 by (meson select_result_I2)
have "\<And>nss nssa. \<not> distinct (concat (nss @ nssa)) \<or> distinct (concat nssa::(_) node_ptr list)"
by (metis (no_types) concat_append distinct_append)
then have "distinct (concat (map (\<lambda>d. |h \<turnstile> get_disconnected_nodes d|\<^sub>r) |h \<turnstile> document_ptr_kinds_M|\<^sub>r))"
using a1 local.a_distinct_lists_def local.heap_is_wellformed_def by blast
then show ?thesis
using f6 f5 a4 a3 a2 by (meson DocumentMonad.ptr_kinds_ptr_kinds_M distinct_concat_map_E(1)
is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
qed
lemma heap_is_wellformed_children_disc_nodes_different:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> set children \<inter> set disc_nodes = {}"
by (metis (no_types, opaque_lifting) disjoint_iff_not_equal distinct_lists_no_parent
is_OK_returns_result_I local.get_child_nodes_ptr_in_heap
local.heap_is_wellformed_def select_result_I2)
lemma heap_is_wellformed_children_disc_nodes:
"heap_is_wellformed h \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h
\<Longrightarrow> \<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)
\<Longrightarrow> (\<exists>document_ptr \<in> fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
- apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)[1]
- by (meson fmember_iff_member_fset)
+ by (auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)
+
lemma heap_is_wellformed_children_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> distinct children"
by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append
distinct_concat_map_E(2) is_OK_returns_result_I local.a_distinct_lists_def
local.get_child_nodes_ptr_in_heap local.heap_is_wellformed_def
select_result_I2)
end
locale l_heap_is_wellformed = l_type_wf + l_known_ptr + l_heap_is_wellformed_defs
+ l_get_child_nodes_defs + l_get_disconnected_nodes_defs +
assumes heap_is_wellformed_children_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children
\<Longrightarrow> child |\<in>| node_ptr_kinds h"
assumes heap_is_wellformed_disc_nodes_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> node \<in> set disc_nodes \<Longrightarrow> node |\<in>| node_ptr_kinds h"
assumes heap_is_wellformed_one_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'
\<Longrightarrow> set children \<inter> set children' \<noteq> {} \<Longrightarrow> ptr = ptr'"
assumes heap_is_wellformed_one_disc_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes'
\<Longrightarrow> set disc_nodes \<inter> set disc_nodes' \<noteq> {} \<Longrightarrow> document_ptr = document_ptr'"
assumes heap_is_wellformed_children_disc_nodes_different:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> set children \<inter> set disc_nodes = {}"
assumes heap_is_wellformed_disconnected_nodes_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> distinct disc_nodes"
assumes heap_is_wellformed_children_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> distinct children"
assumes heap_is_wellformed_children_disc_nodes:
"heap_is_wellformed h \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h
\<Longrightarrow> \<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)
\<Longrightarrow> (\<exists>document_ptr \<in> fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
assumes parent_child_rel_child:
"h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<longleftrightarrow> (ptr, cast child) \<in> parent_child_rel h"
assumes parent_child_rel_finite:
"heap_is_wellformed h \<Longrightarrow> finite (parent_child_rel h)"
assumes parent_child_rel_acyclic:
"heap_is_wellformed h \<Longrightarrow> acyclic (parent_child_rel h)"
assumes parent_child_rel_node_ptr:
"(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> is_node_ptr_kind child_ptr"
assumes parent_child_rel_parent_in_heap:
"(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> parent |\<in>| object_ptr_kinds h"
assumes parent_child_rel_child_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptr parent
\<Longrightarrow> (parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> child_ptr |\<in>| object_ptr_kinds h"
interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel
apply(unfold_locales)
by(auto simp add: heap_is_wellformed_def parent_child_rel_def)
declare l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]:
"l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes
get_disconnected_nodes"
apply(auto simp add: l_heap_is_wellformed_def)[1]
using heap_is_wellformed_children_in_heap
apply blast
using heap_is_wellformed_disc_nodes_in_heap
apply blast
using heap_is_wellformed_one_parent
apply blast
using heap_is_wellformed_one_disc_parent
apply blast
using heap_is_wellformed_children_disc_nodes_different
apply blast
using heap_is_wellformed_disconnected_nodes_distinct
apply blast
using heap_is_wellformed_children_distinct
apply blast
using heap_is_wellformed_children_disc_nodes
apply blast
using parent_child_rel_child
apply (blast)
using parent_child_rel_child
apply(blast)
using parent_child_rel_finite
apply blast
using parent_child_rel_acyclic
apply blast
using parent_child_rel_node_ptr
apply blast
using parent_child_rel_parent_in_heap
apply blast
using parent_child_rel_child_in_heap
apply blast
done
subsection \<open>get\_parent\<close>
locale l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
+ l_heap_is_wellformed
type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma child_parent_dual:
assumes heap_is_wellformed: "heap_is_wellformed h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
assumes "known_ptrs h"
assumes type_wf: "type_wf h"
shows "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr"
proof -
obtain ptrs where ptrs: "h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have h1: "ptr \<in> set ptrs"
using get_child_nodes_ok assms(2) is_OK_returns_result_I
by (metis (no_types, opaque_lifting) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>\<And>thesis. (\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
get_child_nodes_ptr_in_heap returns_result_eq select_result_I2)
let ?P = "(\<lambda>ptr. get_child_nodes ptr \<bind> (\<lambda>children. return (child \<in> set children)))"
let ?filter = "filter_M ?P ptrs"
have "h \<turnstile> ok ?filter"
using ptrs type_wf
using get_child_nodes_ok
apply(auto intro!: filter_M_is_OK_I bind_is_OK_pure_I get_child_nodes_ok simp add: bind_pure_I)[1]
using assms(4) local.known_ptrs_known_ptr by blast
then obtain parent_ptrs where parent_ptrs: "h \<turnstile> ?filter \<rightarrow>\<^sub>r parent_ptrs"
by auto
have h5: "\<exists>!x. x \<in> set ptrs \<and> h \<turnstile> Heap_Error_Monad.bind (get_child_nodes x)
(\<lambda>children. return (child \<in> set children)) \<rightarrow>\<^sub>r True"
apply(auto intro!: bind_pure_returns_result_I)[1]
using heap_is_wellformed_one_parent
proof -
have "h \<turnstile> (return (child \<in> set children)::((_) heap, exception, bool) prog) \<rightarrow>\<^sub>r True"
by (simp add: assms(3))
then show
"\<exists>z. z \<in> set ptrs \<and> h \<turnstile> Heap_Error_Monad.bind (get_child_nodes z)
(\<lambda>ns. return (child \<in> set ns)) \<rightarrow>\<^sub>r True"
by (metis (no_types) assms(2) bind_pure_returns_result_I2 h1 is_OK_returns_result_I
local.get_child_nodes_pure select_result_I2)
next
fix x y
assume 0: "x \<in> set ptrs"
and 1: "h \<turnstile> Heap_Error_Monad.bind (get_child_nodes x)
(\<lambda>children. return (child \<in> set children)) \<rightarrow>\<^sub>r True"
and 2: "y \<in> set ptrs"
and 3: "h \<turnstile> Heap_Error_Monad.bind (get_child_nodes y)
(\<lambda>children. return (child \<in> set children)) \<rightarrow>\<^sub>r True"
and 4: "(\<And>h ptr children ptr' children'. heap_is_wellformed h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'
\<Longrightarrow> set children \<inter> set children' \<noteq> {} \<Longrightarrow> ptr = ptr')"
then show "x = y"
by (metis (no_types, lifting) bind_returns_result_E disjoint_iff_not_equal heap_is_wellformed
return_returns_result)
qed
have "child |\<in>| node_ptr_kinds h"
using heap_is_wellformed_children_in_heap heap_is_wellformed assms(2) assms(3)
by fast
moreover have "parent_ptrs = [ptr]"
apply(rule filter_M_ex1[OF parent_ptrs h1 h5])
using ptrs assms(2) assms(3)
by(auto simp add: object_ptr_kinds_M_defs bind_pure_I intro!: bind_pure_returns_result_I)
ultimately show ?thesis
using ptrs parent_ptrs
by(auto simp add: bind_pure_I get_parent_def
elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I) (*slow, ca 1min *)
qed
lemma parent_child_rel_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent"
shows "(parent, cast child_node) \<in> parent_child_rel h"
using assms parent_child_rel_child get_parent_child_dual by auto
lemma heap_wellformed_induct [consumes 1, case_names step]:
assumes "heap_is_wellformed h"
and step: "\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<Longrightarrow> P (cast child)) \<Longrightarrow> P parent"
shows "P ptr"
proof -
fix ptr
have "wf ((parent_child_rel h)\<inverse>)"
by (simp add: assms(1) finite_acyclic_wf_converse parent_child_rel_acyclic parent_child_rel_finite)
then show "?thesis"
proof (induct rule: wf_induct_rule)
case (less parent)
then show ?case
using assms parent_child_rel_child
by (meson converse_iff)
qed
qed
lemma heap_wellformed_induct2 [consumes 3, case_names not_in_heap empty_children step]:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
and not_in_heap: "\<And>parent. parent |\<notin>| object_ptr_kinds h \<Longrightarrow> P parent"
and empty_children: "\<And>parent. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r [] \<Longrightarrow> P parent"
and step: "\<And>parent children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<Longrightarrow> P (cast child) \<Longrightarrow> P parent"
shows "P ptr"
proof(insert assms(1), induct rule: heap_wellformed_induct)
case (step parent)
then show ?case
proof(cases "parent |\<in>| object_ptr_kinds h")
case True
then obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
using get_child_nodes_ok assms(2) assms(3)
by (meson is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?thesis
proof (cases "children = []")
case True
then show ?thesis
using children empty_children
by simp
next
case False
then show ?thesis
using assms(6) children last_in_set step.hyps by blast
qed
next
case False
then show ?thesis
by (simp add: not_in_heap)
qed
qed
lemma heap_wellformed_induct_rev [consumes 1, case_names step]:
assumes "heap_is_wellformed h"
and step: "\<And>child. (\<And>parent child_node. cast child_node = child
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent) \<Longrightarrow> P child"
shows "P ptr"
proof -
fix ptr
have "wf ((parent_child_rel h))"
by (simp add: assms(1) local.parent_child_rel_acyclic local.parent_child_rel_finite
wf_iff_acyclic_if_finite)
then show "?thesis"
proof (induct rule: wf_induct_rule)
case (less child)
show ?case
using assms get_parent_child_dual
by (metis less.hyps parent_child_rel_parent)
qed
qed
end
interpretation i_get_parent_wf?: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed
parent_child_rel get_disconnected_nodes
using instances
by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
locale l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs
+ l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma preserves_wellformedness_writes_needed:
assumes heap_is_wellformed: "heap_is_wellformed h"
and "h \<turnstile> f \<rightarrow>\<^sub>h h'"
and "writes SW f h h'"
and preserved_get_child_nodes:
"\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. \<forall>r \<in> get_child_nodes_locs object_ptr. r h h'"
and preserved_get_disconnected_nodes:
"\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>document_ptr. \<forall>r \<in> get_disconnected_nodes_locs document_ptr. r h h'"
and preserved_object_pointers:
"\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "heap_is_wellformed h'"
proof -
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
using assms(2) assms(3) object_ptr_kinds_preserved preserved_object_pointers by blast
then have object_ptr_kinds_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq2: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
using select_result_eq by force
then have node_ptr_kinds_eq2: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by auto
then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
by auto
have document_ptr_kinds_eq2: "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
by auto
have children_eq:
"\<And>ptr children. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads assms(3) assms(2)])
using preserved_get_child_nodes by fast
then have children_eq2: "\<And>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r = |h' \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq:
"\<And>document_ptr disconnected_nodes.
h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes"
apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads assms(3) assms(2)])
using preserved_get_disconnected_nodes by fast
then have disconnected_nodes_eq2:
"\<And>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r
= |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using select_result_eq by force
have get_parent_eq: "\<And>ptr parent. h \<turnstile> get_parent ptr \<rightarrow>\<^sub>r parent = h' \<turnstile> get_parent ptr \<rightarrow>\<^sub>r parent"
apply(rule reads_writes_preserved[OF get_parent_reads assms(3) assms(2)])
using preserved_get_child_nodes preserved_object_pointers unfolding get_parent_locs_def by fast
have "a_acyclic_heap h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h"
proof
fix x
assume "x \<in> parent_child_rel h'"
then show "x \<in> parent_child_rel h"
by(simp add: parent_child_rel_def children_eq2 object_ptr_kinds_eq3)
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'"
by (simp add: children_eq2 disconnected_nodes_eq2 document_ptr_kinds_eq3
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_all_ptrs_in_heap_def node_ptr_kinds_eq3 object_ptr_kinds_eq3)
moreover have h0: "a_distinct_lists h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
have h1: "map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))
= map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))"
by (simp add: children_eq2 object_ptr_kinds_eq3)
have h2: "map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))
= map (\<lambda>document_ptr. |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))"
using disconnected_nodes_eq document_ptr_kinds_eq2 select_result_eq by force
have "a_distinct_lists h'"
using h0
by(simp add: a_distinct_lists_def h1 h2)
moreover have "a_owner_document_valid h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
by(auto simp add: a_owner_document_valid_def children_eq2 disconnected_nodes_eq2
object_ptr_kinds_eq3 node_ptr_kinds_eq3 document_ptr_kinds_eq3)
ultimately show ?thesis
by (simp add: heap_is_wellformed_def)
qed
end
interpretation i_get_parent_wf2?: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs known_ptrs get_parent get_parent_locs
heap_is_wellformed parent_child_rel get_disconnected_nodes
get_disconnected_nodes_locs
using l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms
by (simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
locale l_get_parent_wf = l_type_wf + l_known_ptrs + l_heap_is_wellformed_defs
+ l_get_child_nodes_defs + l_get_parent_defs +
assumes child_parent_dual:
"heap_is_wellformed h
\<Longrightarrow> type_wf h
\<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children
\<Longrightarrow> h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr"
assumes heap_wellformed_induct [consumes 1, case_names step]:
"heap_is_wellformed h
\<Longrightarrow> (\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<Longrightarrow> P (cast child)) \<Longrightarrow> P parent)
\<Longrightarrow> P ptr"
assumes heap_wellformed_induct_rev [consumes 1, case_names step]:
"heap_is_wellformed h
\<Longrightarrow> (\<And>child. (\<And>parent child_node. cast child_node = child
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent) \<Longrightarrow> P child)
\<Longrightarrow> P ptr"
assumes parent_child_rel_parent: "heap_is_wellformed h
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent
\<Longrightarrow> (parent, cast child_node) \<in> parent_child_rel h"
lemma get_parent_wf_is_l_get_parent_wf [instances]:
"l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel
get_child_nodes get_parent"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def)[1]
using child_parent_dual heap_wellformed_induct heap_wellformed_induct_rev parent_child_rel_parent
by metis+
subsection \<open>get\_disconnected\_nodes\<close>
subsection \<open>set\_disconnected\_nodes\<close>
subsubsection \<open>get\_disconnected\_nodes\<close>
locale l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_set_disconnected_nodes_get_disconnected_nodes
type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs
+ l_heap_is_wellformed
type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma remove_from_disconnected_nodes_removes:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "h \<turnstile> set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \<rightarrow>\<^sub>h h'"
assumes "h' \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes'"
shows "node_ptr \<notin> set disc_nodes'"
using assms
by (metis distinct_remove1_removeAll heap_is_wellformed_disconnected_nodes_distinct
set_disconnected_nodes_get_disconnected_nodes member_remove remove_code(1)
returns_result_eq)
end
locale l_set_disconnected_nodes_get_disconnected_nodes_wf = l_heap_is_wellformed
+ l_set_disconnected_nodes_get_disconnected_nodes +
assumes remove_from_disconnected_nodes_removes:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> h \<turnstile> set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes'
\<Longrightarrow> node_ptr \<notin> set disc_nodes'"
interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M?:
l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed
parent_child_rel get_child_nodes
using instances
by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma set_disconnected_nodes_get_disconnected_nodes_wf_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]:
"l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed parent_child_rel
get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs"
apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def
l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1]
using remove_from_disconnected_nodes_removes apply fast
done
subsection \<open>get\_root\_node\<close>
locale l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed
type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
+ l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
+ l_get_parent_wf
type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes
get_child_nodes_locs get_parent get_parent_locs
+ l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_ancestors get_ancestors_locs get_root_node get_root_node_locs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_ancestors_reads:
assumes "heap_is_wellformed h"
shows "reads get_ancestors_locs (get_ancestors node_ptr) h h'"
proof (insert assms(1), induct rule: heap_wellformed_induct_rev)
case (step child)
then show ?case
using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def]
apply(simp (no_asm) add: get_ancestors_def)
by(auto simp add: get_ancestors_locs_def reads_subset[OF return_reads] get_parent_reads_pointers
intro!: reads_bind_pure reads_subset[OF check_in_heap_reads]
reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads]
split: option.splits)
qed
lemma get_ancestors_ok:
assumes "heap_is_wellformed h"
and "ptr |\<in>| object_ptr_kinds h"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "h \<turnstile> ok (get_ancestors ptr)"
proof (insert assms(1) assms(2), induct rule: heap_wellformed_induct_rev)
case (step child)
then show ?case
using assms(3) assms(4)
apply(simp (no_asm) add: get_ancestors_def)
apply(simp add: assms(1) get_parent_parent_in_heap)
by(auto intro!: bind_is_OK_pure_I bind_pure_I get_parent_ok split: option.splits)
qed
lemma get_root_node_ptr_in_heap:
assumes "h \<turnstile> ok (get_root_node ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
unfolding get_root_node_def
using get_ancestors_ptr_in_heap
by auto
lemma get_root_node_ok:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
and "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_root_node ptr)"
unfolding get_root_node_def
using assms get_ancestors_ok
by auto
lemma get_ancestors_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
shows "h \<turnstile> get_ancestors (cast child) \<rightarrow>\<^sub>r (cast child) # parent # ancestors
\<longleftrightarrow> h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r parent # ancestors"
proof
assume a1: "h \<turnstile> get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r
cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors"
then have "h \<turnstile> Heap_Error_Monad.bind (check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child))
(\<lambda>_. Heap_Error_Monad.bind (get_parent child)
(\<lambda>x. Heap_Error_Monad.bind (case x of None \<Rightarrow> return [] | Some x \<Rightarrow> get_ancestors x)
(\<lambda>ancestors. return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # ancestors))))
\<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors"
by(simp add: get_ancestors_def)
then show "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r parent # ancestors"
using assms(2) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1]
using returns_result_eq by fastforce
next
assume "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r parent # ancestors"
then show "h \<turnstile> get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors"
using assms(2)
apply(simp (no_asm) add: get_ancestors_def)
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
by (metis (full_types) assms(2) check_in_heap_ptr_in_heap is_OK_returns_result_I
local.get_parent_ptr_in_heap node_ptr_kinds_commutes old.unit.exhaust
select_result_I)
qed
lemma get_ancestors_never_empty:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors"
shows "ancestors \<noteq> []"
proof(insert assms(2), induct arbitrary: ancestors rule: heap_wellformed_induct_rev[OF assms(1)])
case (1 child)
then show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then show ?case
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
next
case (Some child_node)
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
with Some show ?case
proof(induct parent_opt)
case None
then show ?case
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
next
case (Some option)
then show ?case
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
qed
qed
qed
lemma get_ancestors_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and "ancestor \<in> set ancestors"
and "h \<turnstile> get_ancestors ancestor \<rightarrow>\<^sub>r ancestor_ancestors"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "set ancestor_ancestors \<subseteq> set ancestors"
proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors
rule: heap_wellformed_induct_rev)
case (step child)
have "child |\<in>| object_ptr_kinds h"
using get_ancestors_ptr_in_heap step(2) by auto
(* then have "h \<turnstile> check_in_heap child \<rightarrow>\<^sub>r ()"
using returns_result_select_result by force *)
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then have "ancestors = [child]"
using step(2) step(3)
by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2)
show ?case
using step(2) step(3)
apply(auto simp add: \<open>ancestors = [child]\<close>)[1]
using assms(4) returns_result_eq by fastforce
next
case (Some child_node)
note s1 = Some
obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using \<open>child |\<in>| object_ptr_kinds h\<close> assms(1) Some[symmetric]
get_parent_ok[OF type_wf known_ptrs]
by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes)
then show ?case
proof (induct parent_opt)
case None
then have "ancestors = [child]"
using step(2) step(3) s1
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
show ?case
using step(2) step(3)
apply(auto simp add: \<open>ancestors = [child]\<close>)[1]
using assms(4) returns_result_eq by fastforce
next
case (Some parent)
have "h \<turnstile> Heap_Error_Monad.bind (check_in_heap child)
(\<lambda>_. Heap_Error_Monad.bind
(case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \<Rightarrow> return []
| Some node_ptr \<Rightarrow> Heap_Error_Monad.bind (get_parent node_ptr)
(\<lambda>parent_ptr_opt. case parent_ptr_opt of None \<Rightarrow> return []
| Some x \<Rightarrow> get_ancestors x))
(\<lambda>ancestors. return (child # ancestors)))
\<rightarrow>\<^sub>r ancestors"
using step(2)
by(simp add: get_ancestors_def)
moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
using calculation
by(auto elim!: bind_returns_result_E2 split: option.splits)
ultimately have "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors"
using s1 Some
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
show ?case
using step(1)[OF s1[symmetric, simplified] Some \<open>h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors\<close>]
step(3)
apply(auto simp add: tl_ancestors)[1]
by (metis assms(4) insert_iff list.simps(15) local.step(2) returns_result_eq tl_ancestors)
qed
qed
qed
lemma get_ancestors_also_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast child \<in> set ancestors"
and "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "parent \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors: "h \<turnstile> get_ancestors (cast child) \<rightarrow>\<^sub>r child_ancestors"
by (meson assms(1) assms(4) get_ancestors_ok is_OK_returns_result_I known_ptrs
local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result
type_wf)
then have "parent \<in> set child_ancestors"
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)]
get_ancestors_ptr)
then show ?thesis
using assms child_ancestors get_ancestors_subset by blast
qed
lemma get_ancestors_obtains_children:
assumes "heap_is_wellformed h"
and "ancestor \<noteq> ptr"
and "ancestor \<in> set ancestors"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
obtains children ancestor_child where "h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children"
and "ancestor_child \<in> set children" and "cast ancestor_child \<in> set ancestors"
proof -
assume 0: "(\<And>children ancestor_child.
h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<Longrightarrow>
ancestor_child \<in> set children \<Longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set ancestors
\<Longrightarrow> thesis)"
have "\<exists>child. h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ancestor \<and> cast child \<in> set ancestors"
proof (insert assms(1) assms(2) assms(3) assms(4), induct ptr arbitrary: ancestors
rule: heap_wellformed_induct_rev)
case (step child)
have "child |\<in>| object_ptr_kinds h"
using get_ancestors_ptr_in_heap step(4) by auto
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then have "ancestors = [child]"
using step(3) step(4)
by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2)
show ?case
using step(2) step(3) step(4)
by(auto simp add: \<open>ancestors = [child]\<close>)
next
case (Some child_node)
note s1 = Some
obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using \<open>child |\<in>| object_ptr_kinds h\<close> assms(1) Some[symmetric]
using get_parent_ok known_ptrs type_wf
by (metis (no_types, lifting) is_OK_returns_result_E node_ptr_casts_commute
node_ptr_kinds_commutes)
then show ?case
proof (induct parent_opt)
case None
then have "ancestors = [child]"
using step(2) step(3) step(4) s1
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
show ?case
using step(2) step(3) step(4)
by(auto simp add: \<open>ancestors = [child]\<close>)
next
case (Some parent)
have "h \<turnstile> Heap_Error_Monad.bind (check_in_heap child)
(\<lambda>_. Heap_Error_Monad.bind
(case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \<Rightarrow> return []
| Some node_ptr \<Rightarrow> Heap_Error_Monad.bind (get_parent node_ptr)
(\<lambda>parent_ptr_opt. case parent_ptr_opt of None \<Rightarrow> return []
| Some x \<Rightarrow> get_ancestors x))
(\<lambda>ancestors. return (child # ancestors)))
\<rightarrow>\<^sub>r ancestors"
using step(4)
by(simp add: get_ancestors_def)
moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
using calculation
by(auto elim!: bind_returns_result_E2 split: option.splits)
ultimately have "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors"
using s1 Some
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
(* have "ancestor \<noteq> parent" *)
have "ancestor \<in> set tl_ancestors"
using tl_ancestors step(2) step(3) by auto
show ?case
proof (cases "ancestor \<noteq> parent")
case True
show ?thesis
using step(1)[OF s1[symmetric, simplified] Some True
\<open>ancestor \<in> set tl_ancestors\<close> \<open>h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors\<close>]
using tl_ancestors by auto
next
case False
have "child \<in> set ancestors"
using step(4) get_ancestors_ptr by simp
then show ?thesis
using Some False s1[symmetric] by(auto)
qed
qed
qed
qed
then obtain child where child: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ancestor"
and in_ancestors: "cast child \<in> set ancestors"
by auto
then obtain children where
children: "h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children" and
child_in_children: "child \<in> set children"
using get_parent_child_dual by blast
show thesis
using 0[OF children child_in_children] child assms(3) in_ancestors by blast
qed
lemma get_ancestors_parent_child_rel:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "(ptr, child) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> ptr \<in> set ancestors"
proof (safe)
assume 3: "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
show "ptr \<in> set ancestors"
proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by (metis (no_types, lifting) assms(2) bind_returns_result_E get_ancestors_def
in_set_member member_rec(1) return_returns_result)
next
case False
obtain ptr_child where
ptr_child: "(ptr, ptr_child) \<in> (parent_child_rel h) \<and> (ptr_child, child) \<in> (parent_child_rel h)\<^sup>*"
using converse_rtranclE[OF 1(2)] \<open>ptr \<noteq> child\<close>
by metis
then obtain ptr_child_node
where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node"
using ptr_child node_ptr_casts_commute3 parent_child_rel_node_ptr
by (metis )
then obtain children where
children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children"
proof -
assume a1: "\<And>children. \<lbrakk>h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children; ptr_child_node \<in> set children\<rbrakk>
\<Longrightarrow> thesis"
have "ptr |\<in>| object_ptr_kinds h"
using local.parent_child_rel_parent_in_heap ptr_child by blast
moreover have "ptr_child_node \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r"
by (metis calculation known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr
local.parent_child_rel_child ptr_child ptr_child_ptr_child_node
returns_result_select_result type_wf)
ultimately show ?thesis
using a1 get_child_nodes_ok type_wf known_ptrs
by (meson local.known_ptrs_known_ptr returns_result_select_result)
qed
moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \<in> (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 \<in> set ancestors"
using 1 by auto
moreover have "h \<turnstile> get_parent ptr_child_node \<rightarrow>\<^sub>r Some ptr"
using assms(1) children ptr_child_node child_parent_dual
using known_ptrs type_wf by blast
ultimately show ?thesis
using get_ancestors_also_parent assms type_wf by blast
qed
qed
next
assume 3: "ptr \<in> set ancestors"
show "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by simp
next
case False
then obtain children ptr_child_node where
children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children" and
ptr_child_node_in_ancestors: "cast ptr_child_node \<in> set ancestors"
using 1(2) assms(2) get_ancestors_obtains_children assms(1)
using known_ptrs type_wf by blast
then have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \<in> (parent_child_rel h)\<^sup>*"
using 1(1) by blast
moreover have "(ptr, cast ptr_child_node) \<in> parent_child_rel h"
using children ptr_child_node assms(1) parent_child_rel_child_nodes2
using child_parent_dual known_ptrs parent_child_rel_parent type_wf
by blast
ultimately show ?thesis
by auto
qed
qed
qed
lemma get_root_node_parent_child_rel:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r root"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "(root, child) \<in> (parent_child_rel h)\<^sup>*"
using assms get_ancestors_parent_child_rel
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
using get_ancestors_never_empty last_in_set by blast
lemma get_ancestors_eq:
assumes "heap_is_wellformed h"
and "heap_is_wellformed h'"
and "\<And>object_ptr w. object_ptr \<noteq> ptr \<Longrightarrow> w \<in> get_child_nodes_locs object_ptr \<Longrightarrow> w h h'"
and pointers_preserved: "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
and known_ptrs: "known_ptrs h"
and known_ptrs': "known_ptrs h'"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and type_wf: "type_wf h"
and type_wf': "type_wf h'"
shows "h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
proof -
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
using pointers_preserved object_ptr_kinds_preserved_small by blast
then have object_ptr_kinds_M_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
have "h' \<turnstile> ok (get_ancestors ptr)"
using get_ancestors_ok get_ancestors_ptr_in_heap object_ptr_kinds_eq3 assms(1) known_ptrs
known_ptrs' assms(2) assms(7) type_wf'
by blast
then obtain ancestors' where ancestors': "h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors'"
by auto
obtain root where root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
proof -
assume 0: "(\<And>root. h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow> thesis)"
show thesis
apply(rule 0)
using assms(7)
by(auto simp add: get_root_node_def elim!: bind_returns_result_E2 split: option.splits)
qed
have children_eq:
"\<And>p children. p \<noteq> ptr \<Longrightarrow> h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
using get_child_nodes_reads assms(3)
apply(simp add: reads_def reflp_def transp_def preserved_def)
by blast
have "acyclic (parent_child_rel h)"
using assms(1) local.parent_child_rel_acyclic by auto
have "acyclic (parent_child_rel h')"
using assms(2) local.parent_child_rel_acyclic by blast
have 2: "\<And>c parent_opt. cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \<in> set ancestors \<inter> set ancestors'
\<Longrightarrow> h \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt = h' \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt"
proof -
fix c parent_opt
assume 1: " cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \<in> set ancestors \<inter> set ancestors'"
obtain ptrs where ptrs: "h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by simp
let ?P = "(\<lambda>ptr. Heap_Error_Monad.bind (get_child_nodes ptr) (\<lambda>children. return (c \<in> set children)))"
have children_eq_True: "\<And>p. p \<in> set ptrs \<Longrightarrow> h \<turnstile> ?P p \<rightarrow>\<^sub>r True \<longleftrightarrow> h' \<turnstile> ?P p \<rightarrow>\<^sub>r True"
proof -
fix p
assume "p \<in> set ptrs"
then show "h \<turnstile> ?P p \<rightarrow>\<^sub>r True \<longleftrightarrow> h' \<turnstile> ?P p \<rightarrow>\<^sub>r True"
proof (cases "p = ptr")
case True
have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h)\<^sup>*"
using get_ancestors_parent_child_rel 1 assms by blast
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h)"
proof (cases "cast c = ptr")
case True
then show ?thesis
using \<open>acyclic (parent_child_rel h)\<close> by(auto simp add: acyclic_def)
next
case False
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h)\<^sup>*"
using \<open>acyclic (parent_child_rel h)\<close> False rtrancl_eq_or_trancl rtrancl_trancl_trancl
\<open>(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h)\<^sup>*\<close>
by (metis acyclic_def)
then show ?thesis
using r_into_rtrancl by auto
qed
obtain children where children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using type_wf
by (metis \<open>h' \<turnstile> ok get_ancestors ptr\<close> assms(1) get_ancestors_ptr_in_heap get_child_nodes_ok
heap_is_wellformed_def is_OK_returns_result_E known_ptrs local.known_ptrs_known_ptr
object_ptr_kinds_eq3)
then have "c \<notin> set children"
using \<open>(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h)\<close> assms(1)
using parent_child_rel_child_nodes2
using child_parent_dual known_ptrs parent_child_rel_parent
type_wf by blast
with children have "h \<turnstile> ?P p \<rightarrow>\<^sub>r False"
by(auto simp add: True)
moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h')\<^sup>*"
using get_ancestors_parent_child_rel assms(2) ancestors' 1 known_ptrs' type_wf
type_wf' by blast
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')"
proof (cases "cast c = ptr")
case True
then show ?thesis
using \<open>acyclic (parent_child_rel h')\<close> by(auto simp add: acyclic_def)
next
case False
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')\<^sup>*"
using \<open>acyclic (parent_child_rel h')\<close> False rtrancl_eq_or_trancl rtrancl_trancl_trancl
\<open>(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h')\<^sup>*\<close>
by (metis acyclic_def)
then show ?thesis
using r_into_rtrancl by auto
qed
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')"
using r_into_rtrancl by auto
obtain children' where children': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'"
using type_wf type_wf'
by (meson \<open>h' \<turnstile> ok (get_ancestors ptr)\<close> assms(2) get_ancestors_ptr_in_heap
get_child_nodes_ok is_OK_returns_result_E known_ptrs'
local.known_ptrs_known_ptr)
then have "c \<notin> set children'"
using \<open>(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')\<close> assms(2) type_wf type_wf'
using parent_child_rel_child_nodes2 child_parent_dual known_ptrs' parent_child_rel_parent
by auto
with children' have "h' \<turnstile> ?P p \<rightarrow>\<^sub>r False"
by(auto simp add: True)
ultimately show ?thesis
by (metis returns_result_eq)
next
case False
then show ?thesis
using children_eq ptrs
by (metis (no_types, lifting) bind_pure_returns_result_I bind_returns_result_E
get_child_nodes_pure return_returns_result)
qed
qed
have "\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> ok (get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children))) = h' \<turnstile> ok ( get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)))"
using assms(1) assms(2) object_ptr_kinds_eq ptrs type_wf type_wf'
by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M bind_is_OK_pure_I
get_child_nodes_ok get_child_nodes_pure known_ptrs'
local.known_ptrs_known_ptr return_ok select_result_I2)
have children_eq_False:
"\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False = h' \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
proof
fix pa
assume "pa \<in> set ptrs"
and "h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
have "h \<turnstile> ok (get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))
\<Longrightarrow> h' \<turnstile> ok ( get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))"
using \<open>pa \<in> set ptrs\<close> \<open>\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> ok (get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children))) = h' \<turnstile> ok ( get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)))\<close>
by auto
moreover have "h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False
\<Longrightarrow> h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
by (metis (mono_tags, lifting) \<open>\<And>pa. pa \<in> set ptrs
\<Longrightarrow> h \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True = h' \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True\<close> \<open>pa \<in> set ptrs\<close>
calculation is_OK_returns_result_I returns_result_eq returns_result_select_result)
ultimately show "h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
using \<open>h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False\<close>
by auto
next
fix pa
assume "pa \<in> set ptrs"
and "h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
have "h' \<turnstile> ok (get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))
\<Longrightarrow> h \<turnstile> ok ( get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))"
using \<open>pa \<in> set ptrs\<close> \<open>\<And>pa. pa \<in> set ptrs
\<Longrightarrow> h \<turnstile> ok (get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children))) = h' \<turnstile> ok ( get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)))\<close>
by auto
moreover have "h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False
\<Longrightarrow> h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
by (metis (mono_tags, lifting)
\<open>\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True = h' \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True\<close> \<open>pa \<in> set ptrs\<close>
calculation is_OK_returns_result_I returns_result_eq returns_result_select_result)
ultimately show "h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
using \<open>h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False\<close> by blast
qed
have filter_eq: "\<And>xs. h \<turnstile> filter_M ?P ptrs \<rightarrow>\<^sub>r xs = h' \<turnstile> filter_M ?P ptrs \<rightarrow>\<^sub>r xs"
proof (rule filter_M_eq)
show
"\<And>xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children))) h"
by(auto intro!: bind_pure_I)
next
show
"\<And>xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children))) h'"
by(auto intro!: bind_pure_I)
next
fix xs b x
assume 0: "x \<in> set ptrs"
then show "h \<turnstile> Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r b
= h' \<turnstile> Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r b"
apply(induct b)
using children_eq_True apply blast
using children_eq_False apply blast
done
qed
show "h \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt = h' \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt"
apply(simp add: get_parent_def)
apply(rule bind_cong_2)
apply(simp)
apply(simp)
apply(simp add: check_in_heap_def node_ptr_kinds_def object_ptr_kinds_eq3)
apply(rule bind_cong_2)
apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
apply(rule bind_cong_2)
apply(auto intro!: filter_M_pure_I bind_pure_I)[1]
apply(auto intro!: filter_M_pure_I bind_pure_I)[1]
apply(auto simp add: filter_eq (* dest!: returns_result_eq[OF ptrs] *))[1]
using filter_eq ptrs apply auto[1]
using filter_eq ptrs by auto
qed
have "ancestors = ancestors'"
proof(insert assms(1) assms(7) ancestors' 2, induct ptr arbitrary: ancestors ancestors'
rule: heap_wellformed_induct_rev)
case (step child)
show ?case
using step(2) step(3) step(4)
apply(simp add: get_ancestors_def)
apply(auto intro!: elim!: bind_returns_result_E2 split: option.splits)[1]
using returns_result_eq apply fastforce
apply (meson option.simps(3) returns_result_eq)
by (metis IntD1 IntD2 option.inject returns_result_eq step.hyps)
qed
then show ?thesis
using assms(5) ancestors'
by simp
qed
lemma get_ancestors_remains_not_in_ancestors:
assumes "heap_is_wellformed h"
and "heap_is_wellformed h'"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and "h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors'"
and "\<And>p children children'. h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children
\<Longrightarrow> h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children' \<Longrightarrow> set children' \<subseteq> set children"
and "node \<notin> 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 \<notin> set ancestors'"
proof -
have object_ptr_kinds_M_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
using object_ptr_kinds_eq3
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
show ?thesis
proof (insert assms(1) assms(3) assms(4) assms(6), induct ptr arbitrary: ancestors ancestors'
rule: heap_wellformed_induct_rev)
case (step child)
have 1: "\<And>p parent. h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent \<Longrightarrow> h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
proof -
fix p parent
assume "h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
then obtain children' where
children': "h' \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children'" and
p_in_children': "p \<in> set children'"
using get_parent_child_dual by blast
obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^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 \<open>h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent\<close> get_parent_parent_in_heap is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
have "p \<in> set children"
using assms(5) children children' p_in_children'
by blast
then show "h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
using child_parent_dual assms(1) children known_ptrs type_wf by blast
qed
have "node \<noteq> child"
using assms(1) get_ancestors_parent_child_rel step.prems(1) step.prems(3) known_ptrs
using type_wf type_wf'
by blast
then show ?case
using step(2) step(3)
apply(simp add: get_ancestors_def)
using step(4)
apply(auto elim!: bind_returns_result_E2 split: option.splits)[1]
using 1
apply (meson option.distinct(1) returns_result_eq)
by (metis "1" option.inject returns_result_eq step.hyps)
qed
qed
lemma get_ancestors_ptrs_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
shows "ptr' |\<in>| 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 \<turnstile> get_ancestors x \<rightarrow>\<^sub>r a # ancestors"
by(auto simp add: get_ancestors_def[of a] elim!: bind_returns_result_E2 split: option.splits)
then have "x = a"
by(auto simp add: get_ancestors_def[of x] elim!: bind_returns_result_E2 split: option.splits)
then show ?case
using Cons.hyps Cons.prems(2) get_ancestors_ptr_in_heap x
by (metis assms(1) assms(2) assms(3) get_ancestors_obtains_children get_child_nodes_ptr_in_heap
is_OK_returns_result_I)
qed
lemma get_ancestors_prefix:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
assumes "h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'"
shows "\<exists>pre. ancestors = pre @ ancestors'"
proof (insert assms(1) assms(5) assms(6), induct ptr' arbitrary: ancestors'
rule: heap_wellformed_induct)
case (step parent)
then show ?case
proof (cases "parent \<noteq> ptr" )
case True
then obtain children ancestor_child where "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
and "ancestor_child \<in> set children" and "cast ancestor_child \<in> set ancestors"
using assms(1) assms(2) assms(3) assms(4) get_ancestors_obtains_children step.prems(1) by blast
then have "h \<turnstile> get_parent ancestor_child \<rightarrow>\<^sub>r Some parent"
using assms(1) assms(2) assms(3) child_parent_dual by blast
then have "h \<turnstile> get_ancestors (cast ancestor_child) \<rightarrow>\<^sub>r cast ancestor_child # ancestors'"
apply(simp add: get_ancestors_def)
using \<open>h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r ancestors'\<close> get_parent_ptr_in_heap
by(auto simp add: check_in_heap_def is_OK_returns_result_I intro!: bind_pure_returns_result_I)
then show ?thesis
using step(1) \<open>h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children\<close> \<open>ancestor_child \<in> set children\<close>
\<open>cast ancestor_child \<in> set ancestors\<close>
\<open>h \<turnstile> get_ancestors (cast ancestor_child) \<rightarrow>\<^sub>r cast ancestor_child # ancestors'\<close>
by fastforce
next
case False
then show ?thesis
by (metis append_Nil assms(4) returns_result_eq step.prems(2))
qed
qed
lemma get_ancestors_same_root_node:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
assumes "ptr'' \<in> set ancestors"
shows "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr \<longleftrightarrow> h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
proof -
have "ptr' |\<in>| object_ptr_kinds h"
by (metis assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_obtains_children
get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I)
then obtain ancestors' where ancestors': "h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'"
by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E)
then have "\<exists>pre. ancestors = pre @ ancestors'"
using get_ancestors_prefix assms by blast
moreover have "ptr'' |\<in>| object_ptr_kinds h"
by (metis assms(1) assms(2) assms(3) assms(4) assms(6) get_ancestors_obtains_children
get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I)
then obtain ancestors'' where ancestors'': "h \<turnstile> get_ancestors ptr'' \<rightarrow>\<^sub>r ancestors''"
by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E)
then have "\<exists>pre. ancestors = pre @ ancestors''"
using get_ancestors_prefix assms by blast
ultimately show ?thesis
using ancestors' ancestors''
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I)[1]
apply (metis (no_types, lifting) assms(1) get_ancestors_never_empty last_appendR
returns_result_eq)
by (metis assms(1) get_ancestors_never_empty last_appendR returns_result_eq)
qed
lemma get_root_node_parent_same:
assumes "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr"
shows "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root \<longleftrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
proof
assume 1: " h \<turnstile> get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r root"
show "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
using 1[unfolded get_root_node_def] assms
apply(simp add: get_ancestors_def)
apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I split: option.splits)[1]
using returns_result_eq apply fastforce
using get_ancestors_ptr by fastforce
next
assume 1: " h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
show "h \<turnstile> get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r root"
apply(simp add: get_root_node_def)
using assms 1
apply(simp add: get_ancestors_def)
apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I split: option.splits)[1]
apply (simp add: check_in_heap_def is_OK_returns_result_I)
using get_ancestors_ptr get_parent_ptr_in_heap
apply (simp add: is_OK_returns_result_I)
by (meson list.distinct(1) list.set_cases local.get_ancestors_ptr)
qed
lemma get_root_node_same_no_parent:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r cast child"
shows "h \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev)
case (step c)
then show ?case
proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c")
case None
then have "c = cast child"
using step(2)
by(auto simp add: get_root_node_def get_ancestors_def[of c] elim!: bind_returns_result_E2)
then show ?thesis
using None by auto
next
case (Some child_node)
note s = this
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
by (metis (no_types, lifting) assms(2) assms(3) get_root_node_ptr_in_heap
is_OK_returns_result_I local.get_parent_ok node_ptr_casts_commute
node_ptr_kinds_commutes returns_result_select_result step.prems)
then show ?thesis
proof(induct parent_opt)
case None
then show ?case
using Some get_root_node_no_parent returns_result_eq step.prems by fastforce
next
case (Some parent)
then show ?case
using step s
apply(auto simp add: get_root_node_def get_ancestors_def[of c]
elim!: bind_returns_result_E2 split: option.splits list.splits)[1]
using get_root_node_parent_same step.hyps step.prems by auto
qed
qed
qed
lemma get_root_node_not_node_same:
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "\<not>is_node_ptr_kind ptr"
shows "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r ptr"
using assms
apply(simp add: get_root_node_def get_ancestors_def)
by(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I split: option.splits)
lemma get_root_node_root_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
shows "root |\<in>| object_ptr_kinds h"
using assms
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
by (simp add: get_ancestors_never_empty get_ancestors_ptrs_in_heap)
lemma get_root_node_same_no_parent_parent_child_rel:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r ptr'"
shows "\<not>(\<exists>p. (p, ptr') \<in> (parent_child_rel h))"
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) get_root_node_same_no_parent
l_heap_is_wellformed.parent_child_rel_child local.child_parent_dual local.get_child_nodes_ok
local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms local.parent_child_rel_node_ptr
local.parent_child_rel_parent_in_heap node_ptr_casts_commute3 option.simps(3) returns_result_eq
returns_result_select_result)
end
locale l_get_ancestors_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_ancestors_defs
+ l_get_child_nodes_defs + l_get_parent_defs +
assumes get_ancestors_never_empty:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors \<Longrightarrow> ancestors \<noteq> []"
assumes get_ancestors_ok:
"heap_is_wellformed h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> h \<turnstile> ok (get_ancestors ptr)"
assumes get_ancestors_reads:
"heap_is_wellformed h \<Longrightarrow> reads get_ancestors_locs (get_ancestors node_ptr) h h'"
assumes get_ancestors_ptrs_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors \<Longrightarrow> ptr' \<in> set ancestors
\<Longrightarrow> ptr' |\<in>| object_ptr_kinds h"
assumes get_ancestors_remains_not_in_ancestors:
"heap_is_wellformed h \<Longrightarrow> heap_is_wellformed h' \<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors
\<Longrightarrow> h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors'
\<Longrightarrow> (\<And>p children children'. h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children
\<Longrightarrow> h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children'
\<Longrightarrow> set children' \<subseteq> set children)
\<Longrightarrow> node \<notin> set ancestors
\<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> type_wf h' \<Longrightarrow> node \<notin> set ancestors'"
assumes get_ancestors_also_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_ancestors some_ptr \<rightarrow>\<^sub>r ancestors
\<Longrightarrow> cast child_node \<in> set ancestors
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> type_wf h
\<Longrightarrow> known_ptrs h \<Longrightarrow> parent \<in> set ancestors"
assumes get_ancestors_obtains_children:
"heap_is_wellformed h \<Longrightarrow> ancestor \<noteq> ptr \<Longrightarrow> ancestor \<in> set ancestors
\<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> (\<And>children ancestor_child . h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children
\<Longrightarrow> ancestor_child \<in> set children
\<Longrightarrow> cast ancestor_child \<in> set ancestors
\<Longrightarrow> thesis)
\<Longrightarrow> thesis"
assumes get_ancestors_parent_child_rel:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> (ptr, child) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> ptr \<in> set ancestors"
locale l_get_root_node_wf = l_heap_is_wellformed_defs + l_get_root_node_defs + l_type_wf
+ l_known_ptrs + l_get_ancestors_defs + l_get_parent_defs +
assumes get_root_node_ok:
"heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h
\<Longrightarrow> h \<turnstile> ok (get_root_node ptr)"
assumes get_root_node_ptr_in_heap:
"h \<turnstile> ok (get_root_node ptr) \<Longrightarrow> ptr |\<in>| object_ptr_kinds h"
assumes get_root_node_root_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow> root |\<in>| object_ptr_kinds h"
assumes get_ancestors_same_root_node:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors \<Longrightarrow> ptr' \<in> set ancestors
\<Longrightarrow> ptr'' \<in> set ancestors
\<Longrightarrow> h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr \<longleftrightarrow> h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
assumes get_root_node_same_no_parent:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r cast child \<Longrightarrow> h \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
assumes get_root_node_parent_same:
"h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr
\<Longrightarrow> h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root \<longleftrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
interpretation i_get_root_node_wf?:
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs
using instances
by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]:
"l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors
get_ancestors_locs get_child_nodes get_parent"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def)[1]
using get_ancestors_never_empty apply blast
using get_ancestors_ok apply blast
using get_ancestors_reads apply blast
using get_ancestors_ptrs_in_heap apply blast
using get_ancestors_remains_not_in_ancestors apply blast
using get_ancestors_also_parent apply blast
using get_ancestors_obtains_children apply blast
using get_ancestors_parent_child_rel apply blast
using get_ancestors_parent_child_rel apply blast
done
lemma get_root_node_wf_is_l_get_root_node_wf [instances]:
"l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs
get_ancestors get_parent"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1]
using get_root_node_ok apply blast
using get_root_node_ptr_in_heap apply blast
using get_root_node_root_in_heap apply blast
using get_ancestors_same_root_node apply(blast, blast)
using get_root_node_same_no_parent apply blast
using get_root_node_parent_same apply (blast, blast)
done
subsection \<open>to\_tree\_order\<close>
locale l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent +
l_get_parent_wf +
l_heap_is_wellformed
(* l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *)
begin
lemma to_tree_order_ptr_in_heap:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> ok (to_tree_order ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct)
case (step parent)
then show ?case
apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_is_OK_E3)[1]
using get_child_nodes_ptr_in_heap by blast
qed
lemma to_tree_order_either_ptr_or_in_children:
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
and "node \<in> set nodes"
and "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and "node \<noteq> ptr"
obtains child child_to where "child \<in> set children"
and "h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r child_to" and "node \<in> set child_to"
proof -
obtain treeorders where treeorders: "h \<turnstile> map_M to_tree_order (map cast children) \<rightarrow>\<^sub>r treeorders"
using assms
apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1]
using pure_returns_heap_eq returns_result_eq by fastforce
then have "node \<in> set (concat treeorders)"
using assms[simplified to_tree_order_def]
by(auto elim!: bind_returns_result_E4 dest: pure_returns_heap_eq)
then obtain treeorder where "treeorder \<in> set treeorders"
and node_in_treeorder: "node \<in> set treeorder"
by auto
then obtain child where "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r treeorder"
and "child \<in> set children"
using assms[simplified to_tree_order_def] treeorders
by(auto elim!: map_M_pure_E2)
then show ?thesis
using node_in_treeorder returns_result_eq that by auto
qed
lemma to_tree_order_ptrs_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "ptr' \<in> set to"
shows "ptr' |\<in>| object_ptr_kinds h"
proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wellformed_induct)
case (step parent)
have "parent |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) step.prems(1) to_tree_order_ptr_in_heap by blast
then obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then have "to = [parent]"
using step(2) children
apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_returns_result_E2)[1]
by (metis list.distinct(1) list.map_disc_iff list.set_cases map_M_pure_E2 returns_result_eq)
then show ?thesis
using \<open>parent |\<in>| object_ptr_kinds h\<close> step.prems(2) by auto
next
case False
note f = this
then show ?thesis
using children step to_tree_order_either_ptr_or_in_children
proof (cases "ptr' = parent")
case True
then show ?thesis
using \<open>parent |\<in>| object_ptr_kinds h\<close> by blast
next
case False
then show ?thesis
using children step.hyps to_tree_order_either_ptr_or_in_children
by (metis step.prems(1) step.prems(2))
qed
qed
qed
lemma to_tree_order_ok:
assumes wellformed: "heap_is_wellformed h"
and "ptr |\<in>| object_ptr_kinds h"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "h \<turnstile> ok (to_tree_order ptr)"
proof(insert assms(1) assms(2), induct rule: heap_wellformed_induct)
case (step parent)
then show ?case
using assms(3) type_wf
apply(simp add: to_tree_order_def)
apply(auto simp add: heap_is_wellformed_def intro!: map_M_ok_I bind_is_OK_pure_I map_M_pure_I)[1]
using get_child_nodes_ok known_ptrs_known_ptr apply blast
by (simp add: local.heap_is_wellformed_children_in_heap local.to_tree_order_def wellformed)
qed
lemma to_tree_order_child_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
and "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and "node \<in> set children"
and "h \<turnstile> to_tree_order (cast node) \<rightarrow>\<^sub>r nodes'"
shows "set nodes' \<subseteq> set nodes"
proof
fix x
assume a1: "x \<in> set nodes'"
moreover obtain treeorders
where treeorders: "h \<turnstile> map_M to_tree_order (map cast children) \<rightarrow>\<^sub>r treeorders"
using assms(2) assms(3)
apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1]
using pure_returns_heap_eq returns_result_eq by fastforce
then have "nodes' \<in> set treeorders"
using assms(4) assms(5)
by(auto elim!: map_M_pure_E dest: returns_result_eq)
moreover have "set (concat treeorders) \<subseteq> set nodes"
using treeorders assms(2) assms(3)
by(auto simp add: to_tree_order_def elim!: bind_returns_result_E4 dest: pure_returns_heap_eq)
ultimately show "x \<in> set nodes"
by auto
qed
lemma to_tree_order_ptr_in_result:
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
shows "ptr \<in> set nodes"
using assms
apply(simp add: to_tree_order_def)
by(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I bind_pure_I)
lemma to_tree_order_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
and "node \<in> set nodes"
and "h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "set nodes' \<subseteq> set nodes"
proof -
have "\<forall>nodes. h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes \<longrightarrow> (\<forall>node. node \<in> set nodes
\<longrightarrow> (\<forall>nodes'. h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes' \<longrightarrow> set nodes' \<subseteq> set nodes))"
proof(insert assms(1), induct ptr rule: heap_wellformed_induct)
case (step parent)
then show ?case
proof safe
fix nodes node nodes' x
assume 1: "(\<And>children child.
h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> \<forall>nodes. h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes
\<longrightarrow> (\<forall>node. node \<in> set nodes \<longrightarrow> (\<forall>nodes'. h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'
\<longrightarrow> set nodes' \<subseteq> set nodes)))"
and 2: "h \<turnstile> to_tree_order parent \<rightarrow>\<^sub>r nodes"
and 3: "node \<in> set nodes"
and "h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'"
and "x \<in> set nodes'"
have h1: "(\<And>children child nodes node nodes'.
h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes
\<longrightarrow> (node \<in> set nodes \<longrightarrow> (h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes' \<longrightarrow> set nodes' \<subseteq> set nodes)))"
using 1
by blast
obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
using 2
by(auto simp add: to_tree_order_def elim!: bind_returns_result_E)
then have "set nodes' \<subseteq> set nodes"
proof (cases "children = []")
case True
then show ?thesis
by (metis "2" "3" \<open>h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'\<close> children empty_iff list.set(1)
subsetI to_tree_order_either_ptr_or_in_children)
next
case False
then show ?thesis
proof (cases "node = parent")
case True
then show ?thesis
using "2" \<open>h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'\<close> returns_result_eq by fastforce
next
case False
then obtain child nodes_of_child where
"child \<in> set children" and
"h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes_of_child" and
"node \<in> set nodes_of_child"
using 2[simplified to_tree_order_def] 3
to_tree_order_either_ptr_or_in_children[where node=node and ptr=parent] children
apply(auto elim!: bind_returns_result_E2 intro: map_M_pure_I)[1]
using is_OK_returns_result_E 2 a_all_ptrs_in_heap_def assms(1) heap_is_wellformed_def
using "3" by blast
then have "set nodes' \<subseteq> set nodes_of_child"
using h1
using \<open>h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'\<close> children by blast
moreover have "set nodes_of_child \<subseteq> set nodes"
using "2" \<open>child \<in> set children\<close> \<open>h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes_of_child\<close>
assms children to_tree_order_child_subset by auto
ultimately show ?thesis
by blast
qed
qed
then show "x \<in> set nodes"
using \<open>x \<in> set nodes'\<close> by blast
qed
qed
then show ?thesis
using assms by blast
qed
lemma to_tree_order_parent:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
assumes "parent \<in> set nodes"
shows "cast child \<in> set nodes"
proof -
obtain nodes' where nodes': "h \<turnstile> to_tree_order parent \<rightarrow>\<^sub>r nodes'"
using assms to_tree_order_ok get_parent_parent_in_heap
by (meson get_parent_parent_in_heap is_OK_returns_result_E)
then have "set nodes' \<subseteq> set nodes"
using to_tree_order_subset assms
by blast
moreover obtain children where
children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children" and
child: "child \<in> set children"
using assms get_parent_child_dual by blast
then obtain child_to where child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r child_to"
by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E is_OK_returns_result_I
get_parent_ptr_in_heap node_ptr_kinds_commutes to_tree_order_ok)
then have "cast child \<in> set child_to"
apply(simp add: to_tree_order_def)
by(auto elim!: bind_returns_result_E2 map_M_pure_E
dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I)
have "cast child \<in> set nodes'"
using nodes' child
apply(simp add: to_tree_order_def)
apply(auto elim!: bind_returns_result_E2 map_M_pure_E
dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I)[1]
using child_to \<open>cast child \<in> set child_to\<close> returns_result_eq by fastforce
ultimately show ?thesis
by auto
qed
lemma to_tree_order_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
assumes "cast child \<noteq> ptr"
assumes "child \<in> set children"
assumes "cast child \<in> set nodes"
shows "parent \<in> set nodes"
proof(insert assms(1) assms(4) assms(6) assms(8), induct ptr arbitrary: nodes
rule: heap_wellformed_induct)
case (step p)
have "p |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> to_tree_order p \<rightarrow>\<^sub>r nodes\<close> to_tree_order_ptr_in_heap
using assms(1) assms(2) assms(3) by blast
then obtain children where children: "h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then show ?thesis
using step(2) step(3) step(4) children
by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])
next
case False
then obtain c child_to where
child: "c \<in> set children" and
child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<rightarrow>\<^sub>r child_to" and
"cast child \<in> set child_to"
using step(2) children
apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
by (metis (full_types) assms(1) assms(2) assms(3) get_parent_ptr_in_heap
is_OK_returns_result_I l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_kinds_commutes
returns_result_select_result step.prems(1) step.prems(2) step.prems(3)
to_tree_order_either_ptr_or_in_children to_tree_order_ok)
then have "set child_to \<subseteq> set nodes"
using assms(1) child children step.prems(1) to_tree_order_child_subset by auto
show ?thesis
proof (cases "c = child")
case True
then have "parent = p"
using step(3) children child assms(5) assms(7)
by (meson assms(1) assms(2) assms(3) child_parent_dual option.inject returns_result_eq)
then show ?thesis
using step.prems(1) to_tree_order_ptr_in_result by blast
next
case False
then show ?thesis
using step(1)[OF children child child_to] step(3) step(4)
using \<open>set child_to \<subseteq> set nodes\<close>
using \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \<in> set child_to\<close> by auto
qed
qed
qed
lemma to_tree_order_node_ptrs:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "ptr' \<noteq> ptr"
assumes "ptr' \<in> set nodes"
shows "is_node_ptr_kind ptr'"
proof(insert assms(1) assms(4) assms(5) assms(6), induct ptr arbitrary: nodes
rule: heap_wellformed_induct)
case (step p)
have "p |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> to_tree_order p \<rightarrow>\<^sub>r nodes\<close> to_tree_order_ptr_in_heap
using assms(1) assms(2) assms(3) by blast
then obtain children where children: "h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then show ?thesis
using step(2) step(3) step(4) children
by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
next
case False
then obtain c child_to where
child: "c \<in> set children" and
child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<rightarrow>\<^sub>r child_to" and
"ptr' \<in> set child_to"
using step(2) children
apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children by blast
then have "set child_to \<subseteq> set nodes"
using assms(1) child children step.prems(1) to_tree_order_child_subset by auto
show ?thesis
proof (cases "cast c = ptr")
case True
then show ?thesis
using step \<open>ptr' \<in> set child_to\<close> assms(5) child child_to children by blast
next
case False
then show ?thesis
using \<open>ptr' \<in> set child_to\<close> child child_to children is_node_ptr_kind_cast step.hyps by blast
qed
qed
qed
lemma to_tree_order_child2:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "cast child \<noteq> ptr"
assumes "cast child \<in> set nodes"
obtains parent where "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent" and "parent \<in> set nodes"
proof -
assume 1: "(\<And>parent. h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent \<Longrightarrow> parent \<in> set nodes \<Longrightarrow> thesis)"
show thesis
proof(insert assms(1) assms(4) assms(5) assms(6) 1, induct ptr arbitrary: nodes
rule: heap_wellformed_induct)
case (step p)
have "p |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> to_tree_order p \<rightarrow>\<^sub>r nodes\<close> to_tree_order_ptr_in_heap
using assms(1) assms(2) assms(3) by blast
then obtain children where children: "h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then show ?thesis
using step(2) step(3) step(4) children
by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])
next
case False
then obtain c child_to where
child: "c \<in> set children" and
child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<rightarrow>\<^sub>r child_to" and
"cast child \<in> set child_to"
using step(2) children
apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children
by blast
then have "set child_to \<subseteq> set nodes"
using assms(1) child children step.prems(1) to_tree_order_child_subset by auto
have "cast child |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) assms(4) assms(6) to_tree_order_ptrs_in_heap by blast
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt"
by (meson assms(2) assms(3) is_OK_returns_result_E get_parent_ok node_ptr_kinds_commutes)
then show ?thesis
proof (induct parent_opt)
case None
then show ?case
by (metis \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \<in> set child_to\<close> assms(1) assms(2) assms(3)
cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject child child_parent_dual child_to children
option.distinct(1) returns_result_eq step.hyps)
next
case (Some option)
then show ?case
by (meson assms(1) assms(2) assms(3) get_parent_child_dual step.prems(1) step.prems(2)
step.prems(3) step.prems(4) to_tree_order_child)
qed
qed
qed
qed
lemma to_tree_order_parent_child_rel:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
shows "(ptr, child) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> child \<in> set to"
proof
assume 3: "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
show "child \<in> set to"
proof (insert 3, induct child rule: heap_wellformed_induct_rev[OF assms(1)])
case (1 child)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
using assms(4)
apply(simp add: to_tree_order_def)
by(auto simp add: map_M_pure_I elim!: bind_returns_result_E2)
next
case False
obtain child_parent where
"(ptr, child_parent) \<in> (parent_child_rel h)\<^sup>*" and
"(child_parent, child) \<in> (parent_child_rel h)"
using \<open>ptr \<noteq> child\<close>
by (metis "1.prems" rtranclE)
obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child"
using \<open>(child_parent, child) \<in> parent_child_rel h\<close> node_ptr_casts_commute3
parent_child_rel_node_ptr
by blast
then have "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some child_parent"
using \<open>(child_parent, child) \<in> (parent_child_rel h)\<close>
by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E l_get_parent_wf.child_parent_dual
l_heap_is_wellformed.parent_child_rel_child local.get_child_nodes_ok
local.known_ptrs_known_ptr local.l_get_parent_wf_axioms
local.l_heap_is_wellformed_axioms local.parent_child_rel_parent_in_heap)
then show ?thesis
using 1(1) child_node \<open>(ptr, child_parent) \<in> (parent_child_rel h)\<^sup>*\<close>
using assms(1) assms(2) assms(3) assms(4) to_tree_order_parent by blast
qed
qed
next
assume "child \<in> set to"
then show "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
proof (induct child rule: heap_wellformed_induct_rev[OF assms(1)])
case (1 child)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by simp
next
case False
then have "\<exists>parent. (parent, child) \<in> (parent_child_rel h)"
using 1(2) assms(4) to_tree_order_child2[OF assms(1) assms(2) assms(3) assms(4)]
to_tree_order_node_ptrs
by (metis assms(1) assms(2) assms(3) node_ptr_casts_commute3 parent_child_rel_parent)
then obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child"
using node_ptr_casts_commute3 parent_child_rel_node_ptr by blast
then obtain child_parent where child_parent: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some child_parent"
using \<open>\<exists>parent. (parent, child) \<in> (parent_child_rel h)\<close>
by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) to_tree_order_child2)
then have "(child_parent, child) \<in> (parent_child_rel h)"
using assms(1) child_node parent_child_rel_parent by blast
moreover have "child_parent \<in> set to"
by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) child_node child_parent
get_parent_child_dual to_tree_order_child)
then have "(ptr, child_parent) \<in> (parent_child_rel h)\<^sup>*"
using 1 child_node child_parent by blast
ultimately show ?thesis
by auto
qed
qed
qed
end
interpretation i_to_tree_order_wf?: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs to_tree_order known_ptrs get_parent
get_parent_locs heap_is_wellformed parent_child_rel
get_disconnected_nodes get_disconnected_nodes_locs
using instances
apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
done
declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_to_tree_order_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
+ l_to_tree_order_defs
+ l_get_parent_defs + l_get_child_nodes_defs +
assumes to_tree_order_ok:
"heap_is_wellformed h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> h \<turnstile> ok (to_tree_order ptr)"
assumes to_tree_order_ptrs_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to
\<Longrightarrow> ptr' \<in> set to \<Longrightarrow> ptr' |\<in>| object_ptr_kinds h"
assumes to_tree_order_parent_child_rel:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to
\<Longrightarrow> (ptr, child_ptr) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> child_ptr \<in> set to"
assumes to_tree_order_child2:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> cast child \<noteq> ptr \<Longrightarrow> cast child \<in> set nodes
\<Longrightarrow> (\<And>parent. h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent
\<Longrightarrow> parent \<in> set nodes \<Longrightarrow> thesis)
\<Longrightarrow> thesis"
assumes to_tree_order_node_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> ptr' \<noteq> ptr \<Longrightarrow> ptr' \<in> set nodes \<Longrightarrow> is_node_ptr_kind ptr'"
assumes to_tree_order_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow> cast child \<noteq> ptr
\<Longrightarrow> child \<in> set children \<Longrightarrow> cast child \<in> set nodes
\<Longrightarrow> parent \<in> set nodes"
assumes to_tree_order_ptr_in_result:
"h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes \<Longrightarrow> ptr \<in> set nodes"
assumes to_tree_order_parent:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent \<Longrightarrow> parent \<in> set nodes
\<Longrightarrow> cast child \<in> set nodes"
assumes to_tree_order_subset:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes \<Longrightarrow> node \<in> set nodes
\<Longrightarrow> h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> set nodes' \<subseteq> set nodes"
lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]:
"l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
to_tree_order get_parent get_child_nodes"
using instances
apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def)[1]
using to_tree_order_ok
apply blast
using to_tree_order_ptrs_in_heap
apply blast
using to_tree_order_parent_child_rel
apply(blast, blast)
using to_tree_order_child2
apply blast
using to_tree_order_node_ptrs
apply blast
using to_tree_order_child
apply blast
using to_tree_order_ptr_in_result
apply blast
using to_tree_order_parent
apply blast
using to_tree_order_subset
apply blast
done
subsubsection \<open>get\_root\_node\<close>
locale l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
+ l_to_tree_order_wf
begin
lemma to_tree_order_get_root_node:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "ptr' \<in> set to"
assumes "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
assumes "ptr'' \<in> set to"
shows "h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
proof -
obtain ancestors' where ancestors': "h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'"
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_ok is_OK_returns_result_E
to_tree_order_ptrs_in_heap )
moreover have "ptr \<in> set ancestors'"
using \<open>h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'\<close>
using assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_parent_child_rel
to_tree_order_parent_child_rel by blast
ultimately have "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
using \<open>h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr\<close>
using assms(1) assms(2) assms(3) get_ancestors_ptr get_ancestors_same_root_node by blast
obtain ancestors'' where ancestors'': "h \<turnstile> get_ancestors ptr'' \<rightarrow>\<^sub>r ancestors''"
by (meson assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_ok is_OK_returns_result_E
to_tree_order_ptrs_in_heap)
moreover have "ptr \<in> set ancestors''"
using \<open>h \<turnstile> get_ancestors ptr'' \<rightarrow>\<^sub>r ancestors''\<close>
using assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_parent_child_rel
to_tree_order_parent_child_rel by blast
ultimately show ?thesis
using \<open>h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr\<close> assms(1) assms(2) assms(3) get_ancestors_ptr
get_ancestors_same_root_node by blast
qed
lemma to_tree_order_same_root:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
assumes "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r to"
assumes "ptr' \<in> set to"
shows "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
proof (insert assms(1)(* assms(4) assms(5) *) assms(6), induct ptr' rule: heap_wellformed_induct_rev)
case (step child)
then show ?case
proof (cases "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child")
case True
then have "child = root_ptr"
using assms(1) assms(2) assms(3) assms(5) step.prems
by (metis (no_types, lifting) get_root_node_same_no_parent node_ptr_casts_commute3
option.simps(3) returns_result_eq to_tree_order_child2 to_tree_order_node_ptrs)
then show ?thesis
using True by blast
next
case False
then obtain child_node parent where "cast child_node = child"
and "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent"
by (metis assms(1) assms(2) assms(3) assms(4) assms(5) local.get_root_node_no_parent
local.get_root_node_not_node_same local.get_root_node_same_no_parent
local.to_tree_order_child2 local.to_tree_order_ptrs_in_heap node_ptr_casts_commute3
step.prems)
then show ?thesis
proof (cases "child = root_ptr")
case True
then have "h \<turnstile> get_root_node root_ptr \<rightarrow>\<^sub>r root_ptr"
using assms(4)
using \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\<close> assms(1) assms(2) assms(3)
get_root_node_no_parent get_root_node_same_no_parent
by blast
then show ?thesis
using step assms(4)
using True by blast
next
case False
then have "parent \<in> set to"
using assms(5) step(2) to_tree_order_child \<open>h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent\<close>
\<open>cast child_node = child\<close>
by (metis False assms(1) assms(2) assms(3) get_parent_child_dual)
then show ?thesis
using \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\<close> \<open>h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent\<close>
get_root_node_parent_same
using step.hyps by blast
qed
qed
qed
end
interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes
get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs
get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order
using instances
by(simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
locale l_to_tree_order_wf_get_root_node_wf = l_type_wf + l_known_ptrs + l_to_tree_order_defs
+ l_get_root_node_defs + l_heap_is_wellformed_defs +
assumes to_tree_order_get_root_node:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to
\<Longrightarrow> ptr' \<in> set to \<Longrightarrow> h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr
\<Longrightarrow> ptr'' \<in> set to \<Longrightarrow> h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
assumes to_tree_order_same_root:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr
\<Longrightarrow> h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r to \<Longrightarrow> ptr' \<in> set to
\<Longrightarrow> h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]:
"l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order
get_root_node heap_is_wellformed"
using instances
apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def
l_to_tree_order_wf_get_root_node_wf_axioms_def)[1]
using to_tree_order_get_root_node apply blast
using to_tree_order_same_root apply blast
done
subsection \<open>get\_owner\_document\<close>
locale l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_known_ptrs
+ l_heap_is_wellformed
+ l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
+ l_get_ancestors
+ l_get_ancestors_wf
+ l_get_parent
+ l_get_parent_wf
+ l_get_root_node_wf
+ l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_owner_document_disconnected_nodes:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "node_ptr \<in> set disc_nodes"
assumes known_ptrs: "known_ptrs h"
assumes type_wf: "type_wf h"
shows "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r document_ptr"
proof -
have 2: "node_ptr |\<in>| node_ptr_kinds h"
using assms heap_is_wellformed_disc_nodes_in_heap
by blast
have 3: "document_ptr |\<in>| document_ptr_kinds h"
using assms(2) get_disconnected_nodes_ptr_in_heap by blast
have 0:
"\<exists>!document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r. node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
by (metis (no_types, lifting) "3" DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(2) assms(3)
disjoint_iff_not_equal l_heap_is_wellformed.heap_is_wellformed_one_disc_parent
local.get_disconnected_nodes_ok local.l_heap_is_wellformed_axioms
returns_result_select_result select_result_I2 type_wf)
have "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
using heap_is_wellformed_children_disc_nodes_different child_parent_dual assms
using "2" disjoint_iff_not_equal local.get_parent_child_dual local.get_parent_ok
returns_result_select_result split_option_ex
by (metis (no_types, lifting))
then have 4: "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
using 2 get_root_node_no_parent
by blast
obtain document_ptrs where document_ptrs: "h \<turnstile> document_ptr_kinds_M \<rightarrow>\<^sub>r document_ptrs"
by simp
then
have "h \<turnstile> ok (filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs)"
using assms(1) get_disconnected_nodes_ok type_wf unfolding heap_is_wellformed_def
by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I)
then obtain candidates where
candidates: "h \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^sub>r candidates"
by auto
have eq: "\<And>document_ptr. document_ptr |\<in>| document_ptr_kinds h
\<Longrightarrow> node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r \<longleftrightarrow> |h \<turnstile> do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}|\<^sub>r"
apply(auto dest!: get_disconnected_nodes_ok[OF type_wf]
intro!: select_result_I[where P=id, simplified] elim!: bind_returns_result_E2)[1]
apply(drule select_result_E[where P=id, simplified])
by(auto elim!: bind_returns_result_E2)
have filter: "filter (\<lambda>document_ptr. |h \<turnstile> do {
disconnected_nodes \<leftarrow> 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 \<in> cast ` set disconnected_nodes)
}|\<^sub>r) document_ptrs = [document_ptr]"
apply(rule filter_ex1)
using 0 document_ptrs apply(simp)[1]
using eq
using local.get_disconnected_nodes_ok apply auto[1]
using assms(2) assms(3)
apply(auto intro!: intro!: select_result_I[where P=id, simplified]
elim!: bind_returns_result_E2)[1]
using returns_result_eq apply fastforce
using document_ptrs 3 apply(simp)
using document_ptrs
by simp
have "h \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^sub>r [document_ptr]"
apply(rule filter_M_filter2)
using get_disconnected_nodes_ok document_ptrs 3 assms(1) type_wf filter
unfolding heap_is_wellformed_def
by(auto intro: bind_pure_I bind_is_OK_I2)
with 4 document_ptrs have "h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r document_ptr"
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I
split: option.splits)[1]
moreover have "known_ptr (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)"
using "4" assms(1) known_ptrs type_wf known_ptrs_known_ptr "2" node_ptr_kinds_commutes by blast
ultimately show ?thesis
using 2
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
by(auto split: option.splits intro!: bind_pure_returns_result_I)
qed
lemma in_disconnected_nodes_no_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
and "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document"
and "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
have 2: "cast node_ptr |\<in>| object_ptr_kinds h"
using assms(3) get_owner_document_ptr_in_heap by fast
then have 3: "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
using assms(2) local.get_root_node_no_parent by blast
have "\<not>(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
apply(auto)[1]
using assms(2) child_parent_dual[OF assms(1)] type_wf
assms(1) assms(5) get_child_nodes_ok known_ptrs_known_ptr option.simps(3)
returns_result_eq returns_result_select_result
by (metis (no_types, opaque_lifting))
moreover have "node_ptr |\<in>| node_ptr_kinds h"
using assms(2) get_parent_ptr_in_heap by blast
ultimately
have 0: "\<exists>document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r. node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
- by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) finite_set_in heap_is_wellformed_children_disc_nodes)
+ by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) heap_is_wellformed_children_disc_nodes)
then obtain document_ptr where
document_ptr: "document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r" and
node_ptr_in_disc_nodes: "node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
by auto
then show ?thesis
using get_owner_document_disconnected_nodes known_ptrs type_wf assms
using DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) assms(4) get_disconnected_nodes_ok
returns_result_select_result select_result_I2
by (metis (no_types, opaque_lifting) )
qed
lemma get_owner_document_owner_document_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
shows "owner_document |\<in>| document_ptr_kinds h"
using assms
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_split_asm)+
proof -
assume "h \<turnstile> invoke [] ptr () \<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| document_ptr_kinds h"
by (meson invoke_empty is_OK_returns_result_I)
next
assume "h \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ())
\<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| document_ptr_kinds h"
by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits)
next
assume 0: "heap_is_wellformed h"
and 1: "type_wf h"
and 2: "known_ptrs h"
and 3: "\<not> 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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^sub>r owner_document"
then obtain root where
root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
split: option.splits)
then show ?thesis
proof (cases "is_document_ptr root")
case True
then show ?thesis
using 4 5 root
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply(drule(1) returns_result_eq) apply(auto)[1]
using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
next
case False
have "known_ptr root"
using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast
have "root |\<in>| object_ptr_kinds h"
using root
using "0" "1" "2" local.get_root_node_root_in_heap
by blast
then have "is_node_ptr_kind root"
using False \<open>known_ptr root\<close>
apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs)
using is_node_ptr_kind_none by force
then
have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h).
root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
- by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close> local.child_parent_dual
+ by (metis (no_types, opaque_lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close> local.child_parent_dual
local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes
- local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset
+ local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes
option.distinct(1) returns_result_eq returns_result_select_result root)
then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
then
obtain candidates where
candidates: "h \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (root \<in> 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)))
\<rightarrow>\<^sub>r candidates"
by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset
- is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset
+ is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure
return_ok return_pure sorted_list_of_set(1))
then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
apply (simp add: \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>)
using "1" \<open>root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
\<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
local.get_disconnected_nodes_ok by auto
then have "candidates \<noteq> []"
by auto
then have "owner_document \<in> set candidates"
using 5 root 4
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply (metis candidates list.set_sel(1) returns_result_eq)
by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq)
then show ?thesis
using candidates
by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I
local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure)
qed
next
assume 0: "heap_is_wellformed h"
and 1: "type_wf h"
and 2: "known_ptrs h"
and 3: "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
and 4: "h \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^sub>r owner_document"
then obtain root where
root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
split: option.splits)
then show ?thesis
proof (cases "is_document_ptr root")
case True
then show ?thesis
using 3 4 root
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply(drule(1) returns_result_eq) apply(auto)[1]
using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
next
case False
have "known_ptr root"
using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast
have "root |\<in>| object_ptr_kinds h"
using root
using "0" "1" "2" local.get_root_node_root_in_heap
by blast
then have "is_node_ptr_kind root"
using False \<open>known_ptr root\<close>
apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs)
using is_node_ptr_kind_none by force
then
have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in>
cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
- by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close>
+ by (metis (no_types, opaque_lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close>
local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent
local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3
- node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq
+ node_ptr_inclusion node_ptr_kinds_commutes option.distinct(1) returns_result_eq
returns_result_select_result root)
then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
then
obtain candidates where
candidates: "h \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (root \<in> 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)))
\<rightarrow>\<^sub>r candidates"
by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset
- is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset
+ is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure
return_ok return_pure sorted_list_of_set(1))
then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
by (simp add: "1" local.get_disconnected_nodes_ok)
then have "candidates \<noteq> []"
by auto
then have "owner_document \<in> set candidates"
using 4 root 3
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply (metis candidates list.set_sel(1) returns_result_eq)
by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq)
then show ?thesis
using candidates
by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I
local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure)
qed
qed
lemma get_owner_document_ok:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_owner_document ptr)"
proof -
have "known_ptr ptr"
using assms(2) assms(4) local.known_ptrs_known_ptr
by blast
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(auto simp add: known_ptr_impl)[1]
using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr
known_ptr_not_element_ptr
apply blast
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply (metis (no_types, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes
is_document_ptr_kind_none option.case_eq_if)
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply (metis (no_types, lifting) assms(1) assms(2) assms(3) is_node_ptr_kind_none
local.get_root_node_ok node_ptr_casts_commute3 option.case_eq_if)
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply(auto split: option.splits
intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1]
using assms(3) local.get_disconnected_nodes_ok
apply blast
apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply(auto split: option.splits
intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1]
apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)[1]
apply(auto split: option.splits
intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1]
using assms(3) local.get_disconnected_nodes_ok by blast
qed
lemma get_owner_document_child_same:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document"
proof -
have "ptr |\<in>| object_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_child_nodes_ptr_in_heap)
then have "known_ptr ptr"
using assms(2) local.known_ptrs_known_ptr by blast
have "cast child |\<in>| object_ptr_kinds h"
using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes
by blast
then
have "known_ptr (cast child)"
using assms(2) local.known_ptrs_known_ptr by blast
obtain root where root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.get_root_node_ok)
then have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root"
using assms(1) assms(2) assms(3) assms(4) assms(5) local.child_parent_dual
local.get_root_node_parent_same
by blast
have "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr ptr")
case True
then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = ptr"
using case_optionE document_ptr_casts_commute by blast
then have "root = cast document_ptr"
using root
by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
split: option.splits)
then have "h \<turnstile> a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using document_ptr
\<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> document_ptr]
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>
[simplified \<open>root = cast document_ptr\<close> document_ptr], rotated]
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I
split: if_splits option.splits)[1]
using \<open>ptr |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes
by blast
then show ?thesis
using \<open>known_ptr ptr\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>ptr |\<in>| object_ptr_kinds h\<close> True
by(auto simp add: document_ptr[symmetric]
intro!: bind_pure_returns_result_I
split: option.splits)
next
case False
then obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = ptr"
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
then have "h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using root \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>
unfolding a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure)
then show ?thesis
using \<open>known_ptr ptr\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
apply (meson invoke_empty is_OK_returns_result_I)
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
by(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
qed
then show ?thesis
using \<open>known_ptr (cast child)\<close>
apply(auto simp add: get_owner_document_def[of "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child"]
a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
by (smt (verit) \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\<in>| object_ptr_kinds h\<close> cast_document_ptr_not_node_ptr(1)
comp_apply invoke_empty invoke_not invoke_returns_result is_OK_returns_result_I
node_ptr_casts_commute2 option.sel)
qed
end
locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
+ l_get_disconnected_nodes_defs + l_get_owner_document_defs
+ l_get_parent_defs +
assumes get_owner_document_disconnected_nodes:
"heap_is_wellformed h \<Longrightarrow>
known_ptrs h \<Longrightarrow>
type_wf h \<Longrightarrow>
h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
node_ptr \<in> set disc_nodes \<Longrightarrow>
h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r document_ptr"
assumes in_disconnected_nodes_no_parent:
"heap_is_wellformed h \<Longrightarrow>
h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None\<Longrightarrow>
h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document \<Longrightarrow>
h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
known_ptrs h \<Longrightarrow>
type_wf h\<Longrightarrow>
node_ptr \<in> set disc_nodes"
assumes get_owner_document_owner_document_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow>
h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<Longrightarrow>
owner_document |\<in>| document_ptr_kinds h"
assumes get_owner_document_ok:
"heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h
\<Longrightarrow> h \<turnstile> ok (get_owner_document ptr)"
interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr known_ptrs type_wf heap_is_wellformed parent_child_rel get_child_nodes
get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs
get_ancestors get_ancestors_locs get_root_node get_root_node_locs get_owner_document
by(auto simp add: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]:
"l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes
get_owner_document get_parent"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def)[1]
using get_owner_document_disconnected_nodes apply fast
using in_disconnected_nodes_no_parent apply fast
using get_owner_document_owner_document_in_heap apply fast
using get_owner_document_ok apply fast
done
subsubsection \<open>get\_root\_node\<close>
locale l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node_wf +
l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document_wf
begin
lemma get_root_node_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
assumes "is_document_ptr_kind root"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r the (cast root)"
proof -
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
by (meson is_OK_returns_result_I local.get_root_node_ptr_in_heap)
then have "known_ptr ptr"
using assms(3) local.known_ptrs_known_ptr by blast
{
assume "is_document_ptr_kind ptr"
then have "ptr = root"
using assms(4)
by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
split: option.splits)
then have ?thesis
using \<open>is_document_ptr_kind ptr\<close> \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I
split: option.splits)
}
moreover
{
assume "is_node_ptr_kind ptr"
then have ?thesis
using \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
apply(auto split: option.splits)[1]
using \<open>h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root\<close> assms(5)
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def
intro!: bind_pure_returns_result_I
split: option.splits)[2]
}
ultimately
show ?thesis
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
qed
lemma get_root_node_same_owner_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
proof -
have "ptr |\<in>| object_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_root_node_ptr_in_heap)
have "root |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) assms(4) local.get_root_node_root_in_heap by blast
have "known_ptr ptr"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(3) local.known_ptrs_known_ptr by blast
have "known_ptr root"
using \<open>root |\<in>| object_ptr_kinds h\<close> assms(3) local.known_ptrs_known_ptr by blast
show ?thesis
proof (cases "is_document_ptr_kind ptr")
case True
then
have "ptr = root"
using assms(4)
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
by (metis document_ptr_casts_commute3 last_ConsL local.get_ancestors_not_node
node_ptr_no_document_ptr_cast)
then show ?thesis
by auto
next
case False
then have "is_node_ptr_kind ptr"
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
by (metis node_ptr_casts_commute3)
show ?thesis
proof
assume "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
using node_ptr
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
by(auto elim!: bind_returns_result_E2 split: option.splits)
show "h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr_kind root")
case True
have "is_document_ptr root"
using True \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
have "root = cast owner_document"
using True
by (metis \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3) assms(4)
document_ptr_casts_commute3 get_root_node_document returns_result_eq)
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\<close> apply blast
using \<open>root |\<in>| object_ptr_kinds h\<close>
by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none)
next
case False
then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> assms(4)
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
apply (metis assms(1) assms(2) assms(3) local.get_root_node_no_parent
local.get_root_node_same_no_parent node_ptr returns_result_eq)
using \<open>is_node_ptr_kind root\<close> node_ptr returns_result_eq by fastforce
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_node_ptr_kind root\<close> \<open>known_ptr root\<close>
apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)[1]
apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)[1]
using \<open>root |\<in>| object_ptr_kinds h\<close>
by(auto simp add: root_node_ptr)
qed
next
assume "h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
show "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr_kind root")
case True
have "root = cast owner_document"
using \<open>h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
apply(auto simp add: True a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
split: if_splits)[1]
apply (metis True cast_document_ptr_not_node_ptr(2) is_document_ptr_kind_obtains
is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
by (metis True cast_document_ptr_not_node_ptr(1) document_ptr_casts_commute3
is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) get_root_node_document
by fastforce
next
case False
then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
using assms(1) assms(2) assms(3) assms(4) local.get_root_node_no_parent
local.get_root_node_same_no_parent node_ptr returns_result_eq root_node_ptr
by fastforce+
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using node_ptr \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
intro!: bind_pure_returns_result_I split: option.splits)
qed
qed
qed
qed
end
interpretation get_owner_document_wf_get_root_node_wf?: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_ancestors get_ancestors_locs get_root_node get_root_node_locs heap_is_wellformed parent_child_rel
get_disconnected_nodes get_disconnected_nodes_locs get_owner_document
by(auto simp add: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_get_owner_document_wf_get_root_node_wf = l_heap_is_wellformed_defs + l_type_wf +
l_known_ptrs + l_get_root_node_defs + l_get_owner_document_defs +
assumes get_root_node_document:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow>
is_document_ptr_kind root \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r the (cast root)"
assumes get_root_node_same_owner_document:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow>
h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]:
"l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs
get_root_node get_owner_document"
apply(auto simp add: l_get_owner_document_wf_get_root_node_wf_def
l_get_owner_document_wf_get_root_node_wf_axioms_def instances)[1]
using get_root_node_document apply blast
using get_root_node_same_owner_document apply (blast, blast)
done
subsection \<open>Preserving heap-wellformedness\<close>
subsection \<open>set\_attribute\<close>
locale l_set_attribute_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_attribute\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_attribute_get_disconnected_nodes +
l_set_attribute_get_child_nodes
begin
lemma set_attribute_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> set_attribute element_ptr k v \<rightarrow>\<^sub>h h'"
shows "heap_is_wellformed h'"
thm preserves_wellformedness_writes_needed
apply(rule preserves_wellformedness_writes_needed[OF assms set_attribute_writes])
using set_attribute_get_child_nodes
apply(fast)
using set_attribute_get_disconnected_nodes apply(fast)
by(auto simp add: all_args_def set_attribute_locs_def)
end
subsection \<open>remove\_child\<close>
locale l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed +
l_set_disconnected_nodes_get_child_nodes
begin
lemma remove_child_removes_parent:
assumes wellformed: "heap_is_wellformed h"
and remove_child: "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h2"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "h2 \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
proof -
obtain children where children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using remove_child remove_child_def by auto
then have "child \<in> set children"
using remove_child remove_child_def
by(auto elim!: bind_returns_heap_E dest: returns_result_eq split: if_splits)
then have h1: "\<And>other_ptr other_children. other_ptr \<noteq> ptr
\<Longrightarrow> h \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children \<Longrightarrow> child \<notin> set other_children"
using assms(1) known_ptrs type_wf child_parent_dual
by (meson child_parent_dual children option.inject returns_result_eq)
have known_ptr: "known_ptr ptr"
using known_ptrs
by (meson is_OK_returns_heap_I l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms
remove_child remove_child_ptr_in_heap)
obtain owner_document disc_nodes h' where
owner_document: "h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document" and
disc_nodes: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h': "h \<turnstile> set_disconnected_nodes owner_document (child # disc_nodes) \<rightarrow>\<^sub>h h'" and
h2: "h' \<turnstile> set_child_nodes ptr (remove1 child children) \<rightarrow>\<^sub>h h2"
using assms children unfolding remove_child_def
apply(auto split: if_splits elim!: bind_returns_heap_E)[1]
by (metis (full_types) get_child_nodes_pure get_disconnected_nodes_pure
get_owner_document_pure pure_returns_heap_eq returns_result_eq)
have "object_ptr_kinds h = object_ptr_kinds h2"
using remove_child_writes remove_child unfolding remove_child_locs_def
apply(rule writes_small_big)
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by(auto simp add: reflp_def transp_def)
then have "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
unfolding object_ptr_kinds_M_defs by simp
have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF set_disconnected_nodes_writes h']
using set_disconnected_nodes_types_preserved type_wf
by(auto simp add: reflp_def transp_def)
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF remove_child_writes remove_child] unfolding remove_child_locs_def
using set_disconnected_nodes_types_preserved set_child_nodes_types_preserved type_wf
apply(auto simp add: reflp_def transp_def)[1]
by blast
then obtain children' where children': "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'"
using h2 set_child_nodes_get_child_nodes known_ptr
by (metis \<open>object_ptr_kinds h = object_ptr_kinds h2\<close> children get_child_nodes_ok
get_child_nodes_ptr_in_heap is_OK_returns_result_E is_OK_returns_result_I)
have "child \<notin> set children'"
by (metis (mono_tags, lifting) \<open>type_wf h'\<close> children children' distinct_remove1_removeAll h2
known_ptr local.heap_is_wellformed_children_distinct
local.set_child_nodes_get_child_nodes member_remove remove_code(1) select_result_I2
wellformed)
moreover have "\<And>other_ptr other_children. other_ptr \<noteq> ptr
\<Longrightarrow> h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children \<Longrightarrow> child \<notin> set other_children"
proof -
fix other_ptr other_children
assume a1: "other_ptr \<noteq> ptr" and a3: "h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
have "h \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
using get_child_nodes_reads set_disconnected_nodes_writes h' a3
apply(rule reads_writes_separate_backwards)
using set_disconnected_nodes_get_child_nodes by fast
show "child \<notin> set other_children"
using \<open>h \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children\<close> a1 h1 by blast
qed
then have "\<And>other_ptr other_children. other_ptr \<noteq> ptr
\<Longrightarrow> h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children \<Longrightarrow> child \<notin> set other_children"
proof -
fix other_ptr other_children
assume a1: "other_ptr \<noteq> ptr" and a3: "h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
have "h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
using get_child_nodes_reads set_child_nodes_writes h2 a3
apply(rule reads_writes_separate_backwards)
using set_disconnected_nodes_get_child_nodes a1 set_child_nodes_get_child_nodes_different_pointers
by metis
then show "child \<notin> set other_children"
using \<open>\<And>other_ptr other_children. \<lbrakk>other_ptr \<noteq> ptr; h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children\<rbrakk>
\<Longrightarrow> child \<notin> set other_children\<close> a1 by blast
qed
ultimately have ha: "\<And>other_ptr other_children. h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children
\<Longrightarrow> child \<notin> set other_children"
by (metis (full_types) children' returns_result_eq)
moreover obtain ptrs where ptrs: "h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by (simp add: object_ptr_kinds_M_defs)
moreover have "\<And>ptr. ptr \<in> set ptrs \<Longrightarrow> h2 \<turnstile> ok (get_child_nodes ptr)"
using \<open>type_wf h2\<close> ptrs get_child_nodes_ok known_ptr
using \<open>object_ptr_kinds h = object_ptr_kinds h2\<close> known_ptrs local.known_ptrs_known_ptr by auto
ultimately show "h2 \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I)[1]
proof -
have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\<in>| object_ptr_kinds h"
using get_owner_document_ptr_in_heap owner_document by blast
then show "h2 \<turnstile> check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r ()"
by (simp add: \<open>object_ptr_kinds h = object_ptr_kinds h2\<close> check_in_heap_def)
next
show "(\<And>other_ptr other_children. h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children
\<Longrightarrow> child \<notin> set other_children) \<Longrightarrow>
ptrs = sorted_list_of_set (fset (object_ptr_kinds h2)) \<Longrightarrow>
(\<And>ptr. ptr |\<in>| object_ptr_kinds h2 \<Longrightarrow> h2 \<turnstile> ok get_child_nodes ptr) \<Longrightarrow>
h2 \<turnstile> filter_M (\<lambda>ptr. Heap_Error_Monad.bind (get_child_nodes ptr)
(\<lambda>children. return (child \<in> set children))) (sorted_list_of_set (fset (object_ptr_kinds h2))) \<rightarrow>\<^sub>r []"
by(auto intro!: filter_M_empty_I bind_pure_I)
qed
qed
end
locale l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma remove_child_parent_child_rel_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "parent_child_rel h' \<subseteq> parent_child_rel h"
proof (standard, safe)
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(2)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure]
split: if_splits)[1]
using pure_returns_heap_eq by fastforce
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_eq: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq2: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
using select_result_eq by force
then have node_ptr_kinds_eq2: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by auto
then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
using node_ptr_kinds_M_eq by auto
have document_ptr_kinds_eq2: "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
using document_ptr_kinds_M_eq by auto
have children_eq:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow>
h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children =h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
then have children_eq2:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq:
"\<And>document_ptr disconnected_nodes. document_ptr \<noteq> owner_document
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes"
apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers
by (metis (no_types, lifting) Un_iff owner_document select_result_I2)
then have disconnected_nodes_eq2:
"\<And>document_ptr. document_ptr \<noteq> owner_document
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes
h2 children_h] )
by (simp add: set_disconnected_nodes_get_child_nodes)
have "known_ptr ptr"
using assms(3)
using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes
h2]
using set_disconnected_nodes_types_preserved type_wf
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_child_nodes_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r remove1 child children_h"
using assms(2) owner_document h2 disconnected_nodes_h children_h
apply(auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto split: if_splits)[1]
apply(simp)
apply(auto split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E4)
apply(auto)[1]
apply(simp)
using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h'
by blast
fix parent child
assume a1: "(parent, child) \<in> parent_child_rel h'"
then show "(parent, child) \<in> parent_child_rel h"
proof (cases "parent = ptr")
case True
then show ?thesis
using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h'
get_child_nodes_ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1]
by (metis notin_set_remove1)
next
case False
then show ?thesis
using a1
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2)
qed
qed
lemma remove_child_heap_is_wellformed_preserved:
assumes "heap_is_wellformed h"
and "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'"
proof -
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(2)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1]
using pure_returns_heap_eq by fastforce
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_eq: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq2: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
using select_result_eq by force
then have node_ptr_kinds_eq2: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by auto
then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
using node_ptr_kinds_M_eq by auto
have document_ptr_kinds_eq2: "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
using document_ptr_kinds_M_eq by auto
have children_eq:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow>
h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
then have children_eq2:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq: "\<And>document_ptr disconnected_nodes. document_ptr \<noteq> owner_document
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes"
apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers
by (metis (no_types, lifting) Un_iff owner_document select_result_I2)
then have disconnected_nodes_eq2:
"\<And>document_ptr. document_ptr \<noteq> owner_document
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes
h2 children_h] )
by (simp add: set_disconnected_nodes_get_child_nodes)
show "known_ptrs h'"
using object_ptr_kinds_eq3 known_ptrs_preserved \<open>known_ptrs h\<close> by blast
have "known_ptr ptr"
using assms(3)
using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF set_disconnected_nodes_writes h2]
using set_disconnected_nodes_types_preserved type_wf
by(auto simp add: reflp_def transp_def)
then show "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_child_nodes_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r remove1 child children_h"
using assms(2) owner_document h2 disconnected_nodes_h children_h
apply(auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto split: if_splits)[1]
apply(simp)
apply(auto split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E4)
apply(auto)[1]
apply simp
using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h'
by blast
have disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
using owner_document assms(2) h2 disconnected_nodes_h
apply (auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E2)
apply(auto split: if_splits)[1]
apply(simp)
by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits)
then have disconnected_nodes_h': "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h'])
by (simp add: set_child_nodes_get_disconnected_nodes)
moreover have "a_acyclic_heap h"
using assms(1) by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h"
proof (standard, safe)
fix parent child
assume a1: "(parent, child) \<in> parent_child_rel h'"
then show "(parent, child) \<in> parent_child_rel h"
proof (cases "parent = ptr")
case True
then show ?thesis
using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h'
get_child_nodes_ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1]
by (metis imageI notin_set_remove1)
next
case False
then show ?thesis
using a1
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2)
qed
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h"
using assms(1) by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1]
- apply (metis (no_types, lifting) \<open>type_wf h'\<close> assms(2) assms(3) local.get_child_nodes_ok
- local.known_ptrs_known_ptr local.remove_child_children_subset notin_fset object_ptr_kinds_eq3
+ apply (metis (no_types, opaque_lifting) \<open>type_wf h'\<close> assms(2) assms(3) local.get_child_nodes_ok
+ local.known_ptrs_known_ptr local.remove_child_children_subset object_ptr_kinds_eq3
returns_result_select_result subset_code(1) type_wf)
- apply (metis (no_types, lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h
- disconnected_nodes_h' document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap
+ apply (metis (no_types, opaque_lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h
+ disconnected_nodes_h' document_ptr_kinds_eq3 local.remove_child_child_in_heap
node_ptr_kinds_eq3 select_result_I2 set_ConsD subset_code(1))
done
moreover have "a_owner_document_valid h"
using assms(1) by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3
node_ptr_kinds_eq3)[1]
proof -
fix node_ptr
assume 0: "\<forall>node_ptr\<in>fset (node_ptr_kinds h'). (\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
and 1: "node_ptr |\<in>| node_ptr_kinds h'"
and 2: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<longrightarrow>
node_ptr \<notin> set |h' \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
then show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h'
\<and> node_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
proof (cases "node_ptr = child")
case True
show ?thesis
apply(rule exI[where x=owner_document])
using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True
by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I
list.set_intros(1) select_result_I2)
next
case False
then show ?thesis
using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h
disconnected_nodes_h'
apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1]
- by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2))
+ by (metis children_eq2 disconnected_nodes_eq2 in_set_remove1 list.set_intros(2))
qed
qed
moreover
{
have h0: "a_distinct_lists h"
using assms(1) by (simp add: heap_is_wellformed_def)
moreover have ha1: "(\<Union>x\<in>set |h \<turnstile> object_ptr_kinds_M|\<^sub>r. set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r. set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
using \<open>a_distinct_lists h\<close>
unfolding a_distinct_lists_def
by(auto)
have ha2: "ptr |\<in>| object_ptr_kinds h"
using children_h get_child_nodes_ptr_in_heap by blast
have ha3: "child \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r"
using child_in_children_h children_h
by(simp)
have child_not_in: "\<And>document_ptr. document_ptr |\<in>| document_ptr_kinds h
\<Longrightarrow> child \<notin> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using ha1 ha2 ha3
apply(simp)
using IntI by fastforce
moreover have "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
apply(rule select_result_I)
by(auto simp add: object_ptr_kinds_M_defs)
moreover have "distinct |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
apply(rule select_result_I)
by(auto simp add: document_ptr_kinds_M_defs)
ultimately have "a_distinct_lists h'"
proof(simp (no_asm) add: a_distinct_lists_def, safe)
assume 1: "a_distinct_lists h"
and 3: "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
assume 1: "a_distinct_lists h"
and 3: "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
have 4: "distinct (concat ((map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r)))"
using 1 by(auto simp add: a_distinct_lists_def)
show "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
proof(rule distinct_concat_map_I[OF 3[unfolded object_ptr_kinds_eq2], simplified])
fix x
assume 5: "x |\<in>| object_ptr_kinds h'"
then have 6: "distinct |h \<turnstile> get_child_nodes x|\<^sub>r"
using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce
obtain children where children: "h \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children"
and distinct_children: "distinct children"
by (metis "5" "6" type_wf assms(3) get_child_nodes_ok local.known_ptrs_known_ptr
object_ptr_kinds_eq3 select_result_I)
obtain children' where children': "h' \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children'"
using children children_eq children_h' by fastforce
then have "distinct children'"
proof (cases "ptr = x")
case True
then show ?thesis
using children distinct_children children_h children_h'
by (metis children' distinct_remove1 returns_result_eq)
next
case False
then show ?thesis
using children distinct_children children_eq[OF False]
using children' distinct_lists_children h0
using select_result_I2 by fastforce
qed
then show "distinct |h' \<turnstile> get_child_nodes x|\<^sub>r"
using children' by(auto simp add: )
next
fix x y
assume 5: "x |\<in>| object_ptr_kinds h'" and 6: "y |\<in>| object_ptr_kinds h'" and 7: "x \<noteq> y"
obtain children_x where children_x: "h \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children_x"
by (metis "5" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children_y where children_y: "h \<turnstile> get_child_nodes y \<rightarrow>\<^sub>r children_y"
by (metis "6" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children_x' where children_x': "h' \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children_x'"
using children_eq children_h' children_x by fastforce
obtain children_y' where children_y': "h' \<turnstile> get_child_nodes y \<rightarrow>\<^sub>r children_y'"
using children_eq children_h' children_y by fastforce
have "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r))"
using h0 by(auto simp add: a_distinct_lists_def)
then have 8: "set children_x \<inter> set children_y = {}"
using "7" assms(1) children_x children_y local.heap_is_wellformed_one_parent by blast
have "set children_x' \<inter> set children_y' = {}"
proof (cases "ptr = x")
case True
then have "ptr \<noteq> y"
by(simp add: 7)
have "children_x' = remove1 child children_x"
using children_h children_h' children_x children_x' True returns_result_eq by fastforce
moreover have "children_y' = children_y"
using children_y children_y' children_eq[OF \<open>ptr \<noteq> y\<close>] by auto
ultimately show ?thesis
using 8 set_remove1_subset by fastforce
next
case False
then show ?thesis
proof (cases "ptr = y")
case True
have "children_y' = remove1 child children_y"
using children_h children_h' children_y children_y' True returns_result_eq by fastforce
moreover have "children_x' = children_x"
using children_x children_x' children_eq[OF \<open>ptr \<noteq> x\<close>] by auto
ultimately show ?thesis
using 8 set_remove1_subset by fastforce
next
case False
have "children_x' = children_x"
using children_x children_x' children_eq[OF \<open>ptr \<noteq> x\<close>] by auto
moreover have "children_y' = children_y"
using children_y children_y' children_eq[OF \<open>ptr \<noteq> y\<close>] by auto
ultimately show ?thesis
using 8 by simp
qed
qed
then show "set |h' \<turnstile> get_child_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_child_nodes y|\<^sub>r = {}"
using children_x' children_y'
by (metis (no_types, lifting) select_result_I2)
qed
next
assume 2: "distinct |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
then have 4: "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))"
by simp
have 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))))"
using h0
by(simp add: a_distinct_lists_def document_ptr_kinds_eq3)
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))))"
proof(rule distinct_concat_map_I[OF 4[unfolded document_ptr_kinds_eq3]])
fix x
assume 4: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 5: "distinct |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using distinct_lists_disconnected_nodes[OF h0] 4 get_disconnected_nodes_ok
by (simp add: type_wf document_ptr_kinds_eq3 select_result_I)
show "distinct |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "x = owner_document")
case True
have "child \<notin> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using child_not_in document_ptr_kinds_eq2 "4" by fastforce
moreover have "|h' \<turnstile> get_disconnected_nodes x|\<^sub>r = child # |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using disconnected_nodes_h' disconnected_nodes_h unfolding True
by(simp)
ultimately show ?thesis
using 5 unfolding True
by simp
next
case False
show ?thesis
using "5" False disconnected_nodes_eq2 by auto
qed
next
fix x y
assume 4: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and 5: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))" and "x \<noteq> y"
obtain disc_nodes_x where disc_nodes_x: "h \<turnstile> get_disconnected_nodes x \<rightarrow>\<^sub>r disc_nodes_x"
using 4 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of x] document_ptr_kinds_eq2
by auto
obtain disc_nodes_y where disc_nodes_y: "h \<turnstile> get_disconnected_nodes y \<rightarrow>\<^sub>r disc_nodes_y"
using 5 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of y] document_ptr_kinds_eq2
by auto
obtain disc_nodes_x' where disc_nodes_x': "h' \<turnstile> get_disconnected_nodes x \<rightarrow>\<^sub>r disc_nodes_x'"
using 4 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of x] document_ptr_kinds_eq2
by auto
obtain disc_nodes_y' where disc_nodes_y': "h' \<turnstile> get_disconnected_nodes y \<rightarrow>\<^sub>r disc_nodes_y'"
using 5 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of y] document_ptr_kinds_eq2
by auto
have "distinct
(concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) |h \<turnstile> document_ptr_kinds_M|\<^sub>r))"
using h0 by (simp add: a_distinct_lists_def)
then have 6: "set disc_nodes_x \<inter> set disc_nodes_y = {}"
using \<open>x \<noteq> y\<close> assms(1) disc_nodes_x disc_nodes_y local.heap_is_wellformed_one_disc_parent
by blast
have "set disc_nodes_x' \<inter> set disc_nodes_y' = {}"
proof (cases "x = owner_document")
case True
then have "y \<noteq> owner_document"
using \<open>x \<noteq> y\<close> by simp
then have "disc_nodes_y' = disc_nodes_y"
using disconnected_nodes_eq[OF \<open>y \<noteq> owner_document\<close>] disc_nodes_y disc_nodes_y'
by auto
have "disc_nodes_x' = child # disc_nodes_x"
using disconnected_nodes_h' disc_nodes_x disc_nodes_x' True disconnected_nodes_h
returns_result_eq
by fastforce
have "child \<notin> set disc_nodes_y"
using child_not_in disc_nodes_y 5
using document_ptr_kinds_eq2 by fastforce
then show ?thesis
apply(unfold \<open>disc_nodes_x' = child # disc_nodes_x\<close> \<open>disc_nodes_y' = disc_nodes_y\<close>)
using 6 by auto
next
case False
then show ?thesis
proof (cases "y = owner_document")
case True
then have "disc_nodes_x' = disc_nodes_x"
using disconnected_nodes_eq[OF \<open>x \<noteq> owner_document\<close>] disc_nodes_x disc_nodes_x'
by auto
have "disc_nodes_y' = child # disc_nodes_y"
using disconnected_nodes_h' disc_nodes_y disc_nodes_y' True disconnected_nodes_h
returns_result_eq
by fastforce
have "child \<notin> set disc_nodes_x"
using child_not_in disc_nodes_x 4
using document_ptr_kinds_eq2 by fastforce
then show ?thesis
apply(unfold \<open>disc_nodes_y' = child # disc_nodes_y\<close> \<open>disc_nodes_x' = disc_nodes_x\<close>)
using 6 by auto
next
case False
have "disc_nodes_x' = disc_nodes_x"
using disconnected_nodes_eq[OF \<open>x \<noteq> owner_document\<close>] disc_nodes_x disc_nodes_x'
by auto
have "disc_nodes_y' = disc_nodes_y"
using disconnected_nodes_eq[OF \<open>y \<noteq> owner_document\<close>] disc_nodes_y disc_nodes_y'
by auto
then show ?thesis
apply(unfold \<open>disc_nodes_y' = disc_nodes_y\<close> \<open>disc_nodes_x' = disc_nodes_x\<close>)
using 6 by auto
qed
qed
then show "set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using disc_nodes_x' disc_nodes_y' by auto
qed
next
fix x xa xb
assume 1: "xa \<in> fset (object_ptr_kinds h')"
and 2: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 3: "xb \<in> fset (document_ptr_kinds h')"
and 4: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
obtain disc_nodes where disc_nodes: "h \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r disc_nodes"
using 3 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of xb] document_ptr_kinds_eq2 by auto
obtain disc_nodes' where disc_nodes': "h' \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r disc_nodes'"
using 3 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of xb] document_ptr_kinds_eq2 by auto
obtain children where children: "h \<turnstile> get_child_nodes xa \<rightarrow>\<^sub>r children"
- by (metis "1" type_wf assms(3) finite_set_in get_child_nodes_ok is_OK_returns_result_E
+ by (metis "1" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children' where children': "h' \<turnstile> get_child_nodes xa \<rightarrow>\<^sub>r children'"
using children children_eq children_h' by fastforce
have "\<And>x. x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r \<Longrightarrow> x \<in> set |h \<turnstile> get_disconnected_nodes xb|\<^sub>r \<Longrightarrow> False"
using 1 3
apply(fold \<open> object_ptr_kinds h = object_ptr_kinds h'\<close>)
apply(fold \<open> document_ptr_kinds h = document_ptr_kinds h'\<close>)
using children disc_nodes h0 apply(auto simp add: a_distinct_lists_def)[1]
by (metis (no_types, lifting) h0 local.distinct_lists_no_parent select_result_I2)
then have 5: "\<And>x. x \<in> set children \<Longrightarrow> x \<in> set disc_nodes \<Longrightarrow> False"
using children disc_nodes by fastforce
have 6: "|h' \<turnstile> get_child_nodes xa|\<^sub>r = children'"
using children' by simp
have 7: "|h' \<turnstile> get_disconnected_nodes xb|\<^sub>r = disc_nodes'"
using disc_nodes' by simp
have "False"
proof (cases "xa = ptr")
case True
have "distinct children_h"
using children_h distinct_lists_children h0 \<open>known_ptr ptr\<close> by blast
have "|h' \<turnstile> get_child_nodes ptr|\<^sub>r = remove1 child children_h"
using children_h'
by simp
have "children = children_h"
using True children children_h by auto
show ?thesis
using disc_nodes' children' 5 2 4 children_h \<open>distinct children_h\<close> disconnected_nodes_h'
apply(auto simp add: 6 7
\<open>xa = ptr\<close> \<open>|h' \<turnstile> get_child_nodes ptr|\<^sub>r = remove1 child children_h\<close> \<open>children = children_h\<close>)[1]
by (metis (no_types, lifting) disc_nodes disconnected_nodes_eq2 disconnected_nodes_h
select_result_I2 set_ConsD)
next
case False
have "children' = children"
using children' children children_eq[OF False[symmetric]]
by auto
then show ?thesis
proof (cases "xb = owner_document")
case True
then show ?thesis
using disc_nodes disconnected_nodes_h disconnected_nodes_h'
using "2" "4" "5" "6" "7" False \<open>children' = children\<close> assms(1) child_in_children_h
child_parent_dual children children_h disc_nodes' get_child_nodes_ptr_in_heap
list.set_cases list.simps(3) option.simps(1) returns_result_eq set_ConsD
by (metis (no_types, opaque_lifting) assms(3) type_wf)
next
case False
then show ?thesis
using "2" "4" "5" "6" "7" \<open>children' = children\<close> disc_nodes disc_nodes'
disconnected_nodes_eq returns_result_eq
by metis
qed
qed
then show "x \<in> {}"
by simp
qed
}
ultimately show "heap_is_wellformed h'"
using heap_is_wellformed_def by blast
qed
lemma remove_heap_is_wellformed_preserved:
assumes "heap_is_wellformed h"
and "h \<turnstile> remove child \<rightarrow>\<^sub>h h'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'"
using assms
by(auto simp add: remove_def intro: remove_child_heap_is_wellformed_preserved
elim!: bind_returns_heap_E2 split: option.splits)
lemma remove_child_removes_child:
assumes wellformed: "heap_is_wellformed h"
and remove_child: "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'"
and children: "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "child \<notin> set children"
proof -
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr' (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(2)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure]
split: if_splits)[1]
using pure_returns_heap_eq
by fastforce
have "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes remove_child])
unfolding remove_child_locs_def
using set_child_nodes_pointers_preserved set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
moreover have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF remove_child_writes assms(2)]
using set_child_nodes_types_preserved set_disconnected_nodes_types_preserved type_wf
unfolding remove_child_locs_def
apply(auto simp add: reflp_def transp_def)[1]
by blast
ultimately show ?thesis
using remove_child_removes_parent remove_child_heap_is_wellformed_preserved child_parent_dual
by (meson children known_ptrs local.known_ptrs_preserved option.distinct(1) remove_child
returns_result_eq type_wf wellformed)
qed
lemma remove_child_removes_first_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children"
assumes "h \<turnstile> remove_child ptr node_ptr \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
proof -
obtain h2 disc_nodes owner_document where
"h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document" and
"h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (node_ptr # disc_nodes) \<rightarrow>\<^sub>h h2" and
"h2 \<turnstile> set_child_nodes ptr children \<rightarrow>\<^sub>h h'"
using assms(5)
apply(auto simp add: remove_child_def
dest!: bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])[1]
by(auto elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated,OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
have "known_ptr ptr"
by (meson assms(3) assms(4) is_OK_returns_result_I get_child_nodes_ptr_in_heap known_ptrs_known_ptr)
moreover have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 assms(4)])
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
moreover have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h2]
using \<open>type_wf h\<close> set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
ultimately show ?thesis
using set_child_nodes_get_child_nodes\<open>h2 \<turnstile> set_child_nodes ptr children \<rightarrow>\<^sub>h h'\<close>
by fast
qed
lemma remove_removes_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children"
assumes "h \<turnstile> remove node_ptr \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
proof -
have "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some ptr"
using child_parent_dual assms by fastforce
show ?thesis
using assms remove_child_removes_first_child
by(auto simp add: remove_def
dest!: bind_returns_heap_E3[rotated, OF \<open>h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some ptr\<close>, rotated]
bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])
qed
lemma remove_for_all_empty_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
using assms
proof(induct children arbitrary: h h')
case Nil
then show ?case
by simp
next
case (Cons a children)
have "h \<turnstile> get_parent a \<rightarrow>\<^sub>r Some ptr"
using child_parent_dual Cons by fastforce
with Cons show ?case
proof(auto elim!: bind_returns_heap_E)[1]
fix h2
assume 0: "(\<And>h h'. heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r [])"
and 1: "heap_is_wellformed h"
and 2: "type_wf h"
and 3: "known_ptrs h"
and 4: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r a # children"
and 5: "h \<turnstile> get_parent a \<rightarrow>\<^sub>r Some ptr"
and 7: "h \<turnstile> remove a \<rightarrow>\<^sub>h h2"
and 8: "h2 \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h'"
then have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using remove_removes_child by blast
moreover have "heap_is_wellformed h2"
using 7 1 2 3 remove_child_heap_is_wellformed_preserved(3)
by(auto simp add: remove_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
split: option.splits)
moreover have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF remove_writes 7]
using \<open>type_wf h\<close> remove_child_types_preserved
by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
moreover have "object_ptr_kinds h = object_ptr_kinds h2"
using 7
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_writes])
using remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have "known_ptrs h2"
using 3 known_ptrs_preserved by blast
ultimately show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
using 0 8 by fast
qed
qed
end
locale l_remove_child_wf2 = l_type_wf + l_known_ptrs + l_remove_child_defs + l_heap_is_wellformed_defs
+ l_get_child_nodes_defs + l_remove_defs +
assumes remove_child_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes remove_child_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes remove_child_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
assumes remove_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove child \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes remove_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove child \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes remove_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> remove child \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
assumes remove_child_removes_child:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> child \<notin> set children"
assumes remove_child_removes_first_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children
\<Longrightarrow> h \<turnstile> remove_child ptr node_ptr \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes remove_removes_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children
\<Longrightarrow> h \<turnstile> remove node_ptr \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes remove_for_all_empty_children:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs
set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs
heap_is_wellformed parent_child_rel
by unfold_locales
lemma remove_child_wf2_is_l_remove_child_wf2 [instances]:
"l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove"
apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1]
using remove_child_heap_is_wellformed_preserved apply(fast, fast, fast)
using remove_heap_is_wellformed_preserved apply(fast, fast, fast)
using remove_child_removes_child apply fast
using remove_child_removes_first_child apply fast
using remove_removes_child apply fast
using remove_for_all_empty_children apply fast
done
subsection \<open>adopt\_node\<close>
locale l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf +
l_get_owner_document_wf +
l_remove_child_wf2 +
l_heap_is_wellformed
begin
lemma adopt_node_removes_first_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children"
shows "h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
proof -
obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> do { remove_child parent node }
| None \<Rightarrow> do { return ()}) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node # disc_nodes)
} else do { return () }) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: adopt_node_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_parent_pure])
have "h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using h2 remove_child_removes_first_child assms(1) assms(2) assms(3) assms(5)
by (metis list.set_intros(1) local.child_parent_dual option.simps(5) parent_opt returns_result_eq)
then
show ?thesis
using h'
by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes]
split: if_splits)
qed
lemma adopt_node_document_in_heap:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> ok (adopt_node owner_document node)"
shows "owner_document |\<in>| document_ptr_kinds h"
proof -
obtain old_document parent_opt h2 h' where
old_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> do { remove_child parent node } | None \<Rightarrow> do { return ()}) \<rightarrow>\<^sub>h h2"
and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node # disc_nodes)
} else do { return () }) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: adopt_node_def
elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_parent_pure])
show ?thesis
proof (cases "owner_document = old_document")
case True
then show ?thesis
using old_document get_owner_document_owner_document_in_heap assms(1) assms(2) assms(3)
by auto
next
case False
then obtain h3 old_disc_nodes disc_nodes where
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 node old_disc_nodes) \<rightarrow>\<^sub>h h3" and
old_disc_nodes: "h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h': "h3 \<turnstile> set_disconnected_nodes owner_document (node # disc_nodes) \<rightarrow>\<^sub>h h'"
using h'
by(auto elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "owner_document |\<in>| document_ptr_kinds h3"
by (meson is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
moreover have "object_ptr_kinds h = object_ptr_kinds h2"
using h2 apply(simp split: option.splits)
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes])
using remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def)
moreover have "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_disconnected_nodes_writes h3])
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
ultimately show ?thesis
by(auto simp add: document_ptr_kinds_def)
qed
qed
end
locale l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node +
l_get_owner_document_wf +
l_remove_child_wf2 +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma adopt_node_removes_child_step:
assumes wellformed: "heap_is_wellformed h"
and adopt_node: "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h2"
and children: "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<notin> set children"
proof -
obtain old_document parent_opt h' where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h': "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return () ) \<rightarrow>\<^sub>h h'"
using adopt_node get_parent_pure
by(auto simp add: adopt_node_def
elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
split: if_splits)
then have "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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
\<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close> returns_result_eq
by fastforce
next
case (Some option)
then show ?case
using remove_child_removes_child \<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close> known_ptrs type_wf
wellformed
by auto
qed
qed
lemma adopt_node_removes_child:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'"
shows "\<And>ptr' children'.
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children' \<Longrightarrow> node_ptr \<notin> set children'"
using adopt_node_removes_child_step assms by blast
lemma adopt_node_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^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 \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r old_document"
and
parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt"
and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent child | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2"
and
h': "h2 \<turnstile> (if document_ptr \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 child old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes document_ptr;
set_disconnected_nodes document_ptr (child # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^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="\<lambda>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:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have wellformed_h2: "heap_is_wellformed h2"
using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "type_wf h2"
using h2 remove_child_preserves_type_wf known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "known_ptrs h2"
using h2 remove_child_preserves_known_ptrs known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "heap_is_wellformed h' \<and> known_ptrs h' \<and> type_wf h'"
proof(cases "document_ptr = old_document")
case True
then show ?thesis
using h' wellformed_h2 \<open>type_wf h2\<close> \<open>known_ptrs h2\<close> by auto
next
case False
then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where
docs_neq: "document_ptr \<noteq> old_document" and
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 child old_disc_nodes) \<rightarrow>\<^sub>h h3" and
disc_nodes_document_ptr_h3:
"h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \<rightarrow>\<^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="\<lambda>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:
"\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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:
"\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r = |h3 \<turnstile> 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="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h3: "|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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:
"\<And>ptr children. h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r = |h' \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. old_document \<noteq> doc_ptr
\<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. old_document \<noteq> doc_ptr
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
using old_disc_nodes by blast
then have disc_nodes_old_document_h3:
"h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF set_disconnected_nodes_writes h']
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have "known_ptrs h3"
using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3
by blast
then have "known_ptrs h'"
using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2:
"h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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 |\<in>| 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 |\<in>| 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 \<in> set disc_nodes_old_document_h2"
proof (insert parent_opt h2, induct parent_opt)
case None
then have "h = h2"
by(auto)
moreover have "a_owner_document_valid h"
using assms(1) heap_is_wellformed_def by(simp add: heap_is_wellformed_def)
ultimately show ?case
using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)]
in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast
next
case (Some option)
then show ?case
apply(simp split: option.splits)
using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes
known_ptrs
by blast
qed
have "child \<notin> set (remove1 child disc_nodes_old_document_h2)"
using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \<open>distinct disc_nodes_old_document_h2\<close>
by auto
have "child \<notin> set disc_nodes_document_ptr_h3"
proof -
have "a_distinct_lists h2"
using heap_is_wellformed_def wellformed_h2 by blast
then have 0: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r))"
by(simp add: a_distinct_lists_def)
show ?thesis
using distinct_concat_map_E(1)[OF 0] \<open>child \<in> set disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h2 disc_nodes_document_ptr_h2
by (meson \<open>type_wf h2\<close> 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 |\<in>| node_ptr_kinds h"
using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]]
node_ptr_kinds_commutes by blast
have "a_acyclic_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h2"
proof
fix x
assume "x \<in> parent_child_rel h'"
then show "x \<in> parent_child_rel h2"
using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3
mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong
unfolding parent_child_rel_def
by(simp)
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h2\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h3"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1]
apply (simp add: children_eq2_h2 object_ptr_kinds_h2_eq3 subset_code(1))
by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> \<open>type_wf h2\<close>
disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2
document_ptr_kinds_eq3_h2 in_set_remove1 local.get_disconnected_nodes_ok
local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 returns_result_select_result
select_result_I2 wellformed_h2)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1]
apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1))
- by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close>
+ by (metis (no_types, opaque_lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close>
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
+ disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3
local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
select_result_I2 set_ConsD subset_code(1) wellformed_h2)
moreover have "a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2
document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 )
by (smt (verit) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2
disc_nodes_old_document_h2 disc_nodes_old_document_h3
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap
document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1
list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2
node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3
select_result_I2)
have a_distinct_lists_h2: "a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h'"
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2
children_eq2_h2 children_eq2_h3)[1]
proof -
assume 1: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 3: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> 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 \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 4: "distinct |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "old_document \<noteq> x")
case True
then show ?thesis
proof (cases "document_ptr \<noteq> x")
case True
then show ?thesis
using disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>]
disconnected_nodes_eq2_h3[OF \<open>document_ptr \<noteq> x\<close>] 4
by(auto)
next
case False
then show ?thesis
using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4
\<open>child \<notin> set disc_nodes_document_ptr_h3\<close>
by(auto simp add: disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>] )
qed
next
case False
then show ?thesis
by (metis (no_types, opaque_lifting) \<open>distinct disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h3 disconnected_nodes_eq2_h3
distinct_remove1 docs_neq select_result_I2)
qed
next
fix x y
assume a0: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a1: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a2: "x \<noteq> y"
moreover have 5: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
proof(cases "old_document = x")
case True
have "old_document \<noteq> y"
using \<open>x \<noteq> y\<close> \<open>old_document = x\<close> by simp
have "document_ptr \<noteq> x"
using docs_neq \<open>old_document = x\<close> 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] \<open>old_document = x\<close>
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
\<open>document_ptr \<noteq> x\<close> 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 \<open>old_document = x\<close>
docs_neq \<open>old_document \<noteq> y\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr = x\<close>
apply(simp)
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr \<noteq> x\<close>
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 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False
\<open>type_wf h2\<close> 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 \<noteq> y"
using \<open>x \<noteq> y\<close> by auto
have "set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
using \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document \<noteq> y\<close> \<open>document_ptr = x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3
\<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
by(auto)
next
case False
then show ?thesis
proof(cases "document_ptr = y")
case True
have f1: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set disc_nodes_document_ptr_h3 = {}"
using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>document_ptr \<noteq> x\<close> 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 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = {}"
using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>old_document \<noteq> x\<close>
by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2
- document_ptr_kinds_eq3_h3 finite_fset fmember_iff_member_fset set_sorted_list_of_set)
+ document_ptr_kinds_eq3_h3 finite_fset set_sorted_list_of_set)
ultimately show ?thesis
using 5 select_result_I2[OF disc_nodes_document_ptr_h']
select_result_I2[OF disc_nodes_old_document_h2] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr = y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> 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] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close>
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
by (metis \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
empty_iff inf.idem)
qed
qed
qed
qed
qed
next
fix x xa xb
assume 0: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 2: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h'"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h'"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
then show False
using \<open>child \<in> set disc_nodes_old_document_h2\<close> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a3: "h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r remove1 child disc_nodes_old_document_h2"
assume a4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a5: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f6: "old_document |\<in>| document_ptr_kinds h'"
using a1 \<open>xb |\<in>| document_ptr_kinds h'\<close> by blast
have f7: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a2 by simp
have "x \<in> set disc_nodes_old_document_h2"
using f6 a3 a1 by (metis (no_types) \<open>type_wf h'\<close> \<open>x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r\<close>
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' \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using f7 f6 a5 a4 \<open>xa |\<in>| object_ptr_kinds h'\<close>
by fastforce
then show ?thesis
using \<open>x \<in> set disc_nodes_old_document_h2\<close> a1 a4 f7 by blast
next
assume a1: "xb \<noteq> old_document"
assume a2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3"
assume a3: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a4: "xa |\<in>| object_ptr_kinds h'"
assume a5: "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r child # disc_nodes_document_ptr_h3"
assume a6: "old_document |\<in>| document_ptr_kinds h'"
assume a7: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
assume a8: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a10: "\<And>doc_ptr. old_document \<noteq> doc_ptr
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a11: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a12: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f13: "\<And>d. d \<notin> set |h' \<turnstile> document_ptr_kinds_M|\<^sub>r \<or> h2 \<turnstile> ok get_disconnected_nodes d"
using a9 \<open>type_wf h2\<close> get_disconnected_nodes_ok
by simp
then have f14: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a6 a3 by simp
have "x \<notin> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
using a12 a8 a4 \<open>xb |\<in>| document_ptr_kinds h'\<close>
- by (meson UN_I disjoint_iff_not_equal fmember_iff_member_fset)
+ by (meson UN_I disjoint_iff_not_equal)
then have "x = child"
using f13 a11 a10 a7 a5 a2 a1
by (metis (no_types, lifting) select_result_I2 set_ConsD)
then have "child \<notin> set disc_nodes_old_document_h2"
using f14 a12 a8 a6 a4
by (metis \<open>type_wf h'\<close> 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 \<open>child \<in> set disc_nodes_old_document_h2\<close> by fastforce
qed
qed
ultimately show ?thesis
using \<open>type_wf h'\<close> \<open>known_ptrs h'\<close> \<open>a_owner_document_valid h'\<close> heap_is_wellformed_def by blast
qed
then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'"
by auto
qed
lemma adopt_node_node_in_disconnected_nodes:
assumes wellformed: "heap_is_wellformed h"
and adopt_node: "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'"
and "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2"
and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node_ptr # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
assume a2: "\<And>h document_ptr disc_nodes h'. h \<turnstile> set_disconnected_nodes document_ptr disc_nodes \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume "h'a \<turnstile> set_disconnected_nodes owner_document (node_ptr # xb) \<rightarrow>\<^sub>h h'"
then have "node_ptr # xb = disc_nodes"
using a2 a1 by (meson returns_result_eq)
then show ?thesis
by (meson list.set_intros(1))
qed
qed
qed
end
interpretation i_adopt_node_wf?: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs
remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr
type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs
remove heap_is_wellformed parent_child_rel
by(simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs
remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr
type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs
remove heap_is_wellformed parent_child_rel get_root_node get_root_node_locs
by(simp add: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
locale l_adopt_node_wf = l_heap_is_wellformed + l_known_ptrs + l_type_wf + l_adopt_node_defs
+ l_get_child_nodes_defs + l_get_disconnected_nodes_defs +
assumes adopt_node_preserves_wellformedness:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> heap_is_wellformed h'"
assumes adopt_node_removes_child:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h2
\<Longrightarrow> h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> node_ptr \<notin> set children"
assumes adopt_node_node_in_disconnected_nodes:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> node_ptr \<in> set disc_nodes"
assumes adopt_node_removes_first_child: "heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
assumes adopt_node_document_in_heap: "heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> h \<turnstile> ok (adopt_node owner_document node)
\<Longrightarrow> owner_document |\<in>| document_ptr_kinds h"
assumes adopt_node_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> type_wf h'"
assumes adopt_node_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h'"
lemma adopt_node_wf_is_l_adopt_node_wf [instances]:
"l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes
get_disconnected_nodes known_ptrs adopt_node"
using heap_is_wellformed_is_l_heap_is_wellformed known_ptrs_is_l_known_ptrs
apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def)[1]
using adopt_node_preserves_wellformedness apply blast
using adopt_node_removes_child apply blast
using adopt_node_node_in_disconnected_nodes apply blast
using adopt_node_removes_first_child apply blast
using adopt_node_document_in_heap apply blast
using adopt_node_preserves_wellformedness apply blast
using adopt_node_preserves_wellformedness apply blast
done
subsection \<open>insert\_before\<close>
locale l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node_wf +
l_set_disconnected_nodes_get_child_nodes +
l_heap_is_wellformed
begin
lemma insert_before_removes_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "ptr \<noteq> ptr'"
assumes "h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children"
shows "h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
proof -
obtain owner_document h2 h3 disc_nodes reference_child where
"h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
"h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
"h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disc_nodes) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
split: if_splits option.splits)
have "h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using h2 adopt_node_removes_first_child assms(1) assms(2) assms(3) assms(6)
by simp
then have "h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using h3
by(auto simp add: set_disconnected_nodes_get_child_nodes
dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes])
then show ?thesis
using h' assms(4)
apply(auto simp add: a_insert_node_def
elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated])[1]
by(auto simp add: set_child_nodes_get_child_nodes_different_pointers
elim!: reads_writes_separate_forwards[OF get_child_nodes_reads set_child_nodes_writes])
qed
end
locale l_insert_before_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
+ l_insert_before_defs + l_get_child_nodes_defs +
assumes insert_before_removes_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> ptr \<noteq> ptr'
\<Longrightarrow> h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs
get_child_nodes get_child_nodes_locs set_child_nodes
set_child_nodes_locs get_ancestors get_ancestors_locs
adopt_node adopt_node_locs set_disconnected_nodes
set_disconnected_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_owner_document insert_before
insert_before_locs append_child type_wf known_ptr known_ptrs
heap_is_wellformed parent_child_rel
by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma insert_before_wf_is_l_insert_before_wf [instances]:
"l_insert_before_wf heap_is_wellformed type_wf known_ptr known_ptrs insert_before get_child_nodes"
apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1]
using insert_before_removes_child apply fast
done
locale l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_child_nodes_get_disconnected_nodes +
l_remove_child +
l_get_root_node_wf +
l_set_disconnected_nodes_get_disconnected_nodes_wf +
l_set_disconnected_nodes_get_ancestors +
l_get_ancestors_wf +
l_get_owner_document +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document_wf
begin
lemma insert_before_preserves_acyclitity:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'"
shows "acyclic (parent_child_rel h')"
proof -
obtain ancestors reference_child owner_document h2 h3
disconnected_nodes_h2
where
ancestors: "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
reference_child:
"h \<turnstile> (if Some node = child then a_next_sibling node
else return child) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document
\<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document
(remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
have "known_ptr ptr"
by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I assms
l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document)
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF adopt_node_writes h2]
using assms adopt_node_types_preserved
by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF insert_node_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs )
then have object_ptr_kinds_M_eq2_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have "known_ptrs h2"
using assms object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast
have wellformed_h2: "heap_is_wellformed h2"
using adopt_node_preserves_wellformedness[OF assms(1) h2] assms by simp
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h2: "|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto
have "known_ptrs h3"
using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \<open>known_ptrs h2\<close> by blast
have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h3:
"|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h3: "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto
have "known_ptrs h'"
using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \<open>known_ptrs h3\<close> by blast
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. owner_document \<noteq> doc_ptr
\<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. doc_ptr \<noteq> owner_document
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_h3:
"h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r remove1 node disconnected_nodes_h2"
using h3 set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h3:
"\<And>ptr' children. ptr \<noteq> ptr'
\<Longrightarrow> h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr \<noteq> ptr' \<Longrightarrow> |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
obtain children_h3 where children_h3: "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h3"
using h' a_insert_node_def by auto
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r insert_before_list node reference_child children_h3"
using h' \<open>type_wf h3\<close> \<open>known_ptr ptr\<close>
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 |\<in>| object_ptr_kinds h3"
using children_h3 get_child_nodes_ptr_in_heap by blast
have node_in_heap: "node |\<in>| node_ptr_kinds h"
using h2 adopt_node_child_in_heap by fast
have child_not_in_any_children:
"\<And>p children. h2 \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children \<Longrightarrow> node \<notin> set children"
using assms h2 adopt_node_removes_child by auto
have "node \<in> set disconnected_nodes_h2"
using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1)
\<open>type_wf h\<close> \<open>known_ptrs h\<close> by blast
have node_not_in_disconnected_nodes:
"\<And>d. d |\<in>| document_ptr_kinds h3 \<Longrightarrow> node \<notin> set |h3 \<turnstile> get_disconnected_nodes d|\<^sub>r"
proof -
fix d
assume "d |\<in>| document_ptr_kinds h3"
show "node \<notin> set |h3 \<turnstile> 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 \<open>d |\<in>| document_ptr_kinds h3\<close> disconnected_nodes_h3
by fastforce
next
case False
then have
"set |h2 \<turnstile> get_disconnected_nodes d|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes owner_document|\<^sub>r = {}"
using distinct_concat_map_E(1) wellformed_h2
by (metis (no_types, lifting) \<open>d |\<in>| document_ptr_kinds h3\<close> \<open>type_wf h2\<close>
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] \<open>node \<in> set disconnected_nodes_h2\<close>
disconnected_nodes_h2 by fastforce
qed
qed
have "cast node \<noteq> ptr"
using ancestors node_not_in_ancestors get_ancestors_ptr
by fast
obtain ancestors_h2 where ancestors_h2: "h2 \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_ok object_ptr_kinds_M_eq2_h2 \<open>known_ptrs h2\<close> \<open>type_wf h2\<close>
by (metis is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2)
have ancestors_h3: "h3 \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_reads set_disconnected_nodes_writes h3
apply(rule reads_writes_separate_forwards)
using \<open>heap_is_wellformed h2\<close> ancestors_h2
by (auto simp add: set_disconnected_nodes_get_ancestors)
have node_not_in_ancestors_h2: "cast node \<notin> set ancestors_h2"
apply(rule get_ancestors_remains_not_in_ancestors[OF assms(1) wellformed_h2 ancestors ancestors_h2])
using adopt_node_children_subset using h2 \<open>known_ptrs h\<close> \<open> type_wf h\<close> apply(blast)
using node_not_in_ancestors apply(blast)
using object_ptr_kinds_M_eq3_h apply(blast)
using \<open>known_ptrs h\<close> apply(blast)
using \<open>type_wf h\<close> apply(blast)
using \<open>type_wf h2\<close> by blast
have "acyclic (parent_child_rel h2)"
using wellformed_h2 by (simp add: heap_is_wellformed_def acyclic_heap_def)
then have "acyclic (parent_child_rel h3)"
by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
moreover
have "cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h2)\<^sup>*}"
using adopt_node_removes_child
using ancestors node_not_in_ancestors
using \<open>known_ptrs h2\<close> \<open>type_wf h2\<close> ancestors_h2 local.get_ancestors_parent_child_rel
node_not_in_ancestors_h2 wellformed_h2
by blast
then have "cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h3)\<^sup>*}"
by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
moreover have "parent_child_rel h'
= insert (ptr, cast node) ((parent_child_rel h3))"
using children_h3 children_h' ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3
insert_before_list_node_in_set)[1]
apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2)
by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2)
ultimately show "acyclic (parent_child_rel h')"
by (auto simp add: heap_is_wellformed_def)
qed
lemma insert_before_heap_is_wellformed_preserved:
assumes wellformed: "heap_is_wellformed h"
and insert_before: "h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
proof -
obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where
ancestors: "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
reference_child:
"h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
using assms(2)
by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
have "known_ptr ptr"
by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I known_ptrs
l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document)
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF adopt_node_writes h2]
using type_wf adopt_node_types_preserved
by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF insert_node_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs )
then have object_ptr_kinds_M_eq2_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have "known_ptrs h2"
using known_ptrs object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast
have wellformed_h2: "heap_is_wellformed h2"
using adopt_node_preserves_wellformedness[OF wellformed h2] known_ptrs type_wf .
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h2: "|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto
have "known_ptrs h3"
using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \<open>known_ptrs h2\<close> by blast
have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h3:
"|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h3: "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto
show "known_ptrs h'"
using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \<open>known_ptrs h3\<close> by blast
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. owner_document \<noteq> doc_ptr
\<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. doc_ptr \<noteq> owner_document
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_h3:
"h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r remove1 node disconnected_nodes_h2"
using h3 set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h3:
"\<And>ptr' children. ptr \<noteq> ptr'
\<Longrightarrow> h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr \<noteq> ptr' \<Longrightarrow> |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
obtain children_h3 where children_h3: "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h3"
using h' a_insert_node_def by auto
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r insert_before_list node reference_child children_h3"
using h' \<open>type_wf h3\<close> \<open>known_ptr ptr\<close>
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 |\<in>| object_ptr_kinds h3"
using children_h3 get_child_nodes_ptr_in_heap by blast
have node_in_heap: "node |\<in>| node_ptr_kinds h"
using h2 adopt_node_child_in_heap by fast
have child_not_in_any_children:
"\<And>p children. h2 \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children \<Longrightarrow> node \<notin> set children"
using wellformed h2 adopt_node_removes_child \<open>type_wf h\<close> \<open>known_ptrs h\<close> by auto
have "node \<in> set disconnected_nodes_h2"
using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1)
\<open>type_wf h\<close> \<open>known_ptrs h\<close> by blast
have node_not_in_disconnected_nodes:
"\<And>d. d |\<in>| document_ptr_kinds h3 \<Longrightarrow> node \<notin> set |h3 \<turnstile> get_disconnected_nodes d|\<^sub>r"
proof -
fix d
assume "d |\<in>| document_ptr_kinds h3"
show "node \<notin> set |h3 \<turnstile> 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 \<open>d |\<in>| document_ptr_kinds h3\<close> disconnected_nodes_h3
by fastforce
next
case False
then have
"set |h2 \<turnstile> get_disconnected_nodes d|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes owner_document|\<^sub>r = {}"
using distinct_concat_map_E(1) wellformed_h2
by (metis (no_types, lifting) \<open>d |\<in>| document_ptr_kinds h3\<close> \<open>type_wf h2\<close>
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] \<open>node \<in> set disconnected_nodes_h2\<close>
disconnected_nodes_h2 by fastforce
qed
qed
have "cast node \<noteq> ptr"
using ancestors node_not_in_ancestors get_ancestors_ptr
by fast
obtain ancestors_h2 where ancestors_h2: "h2 \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_ok object_ptr_kinds_M_eq2_h2 \<open>known_ptrs h2\<close> \<open>type_wf h2\<close>
by (metis is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2)
have ancestors_h3: "h3 \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_reads set_disconnected_nodes_writes h3
apply(rule reads_writes_separate_forwards)
using \<open>heap_is_wellformed h2\<close> ancestors_h2
by (auto simp add: set_disconnected_nodes_get_ancestors)
have node_not_in_ancestors_h2: "cast node \<notin> set ancestors_h2"
apply(rule get_ancestors_remains_not_in_ancestors[OF assms(1) wellformed_h2 ancestors ancestors_h2])
using adopt_node_children_subset using h2 \<open>known_ptrs h\<close> \<open> type_wf h\<close> apply(blast)
using node_not_in_ancestors apply(blast)
using object_ptr_kinds_M_eq3_h apply(blast)
using \<open>known_ptrs h\<close> apply(blast)
using \<open>type_wf h\<close> apply(blast)
using \<open>type_wf h2\<close> by blast
moreover have "a_acyclic_heap h'"
proof -
have "acyclic (parent_child_rel h2)"
using wellformed_h2 by (simp add: heap_is_wellformed_def acyclic_heap_def)
then have "acyclic (parent_child_rel h3)"
by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
moreover have "cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h2)\<^sup>*}"
using get_ancestors_parent_child_rel node_not_in_ancestors_h2 \<open>known_ptrs h2\<close> \<open>type_wf h2\<close>
using ancestors_h2 wellformed_h2 by blast
then have "cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h3)\<^sup>*}"
by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
moreover have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))"
using children_h3 children_h' ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3
insert_before_list_node_in_set)[1]
apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2)
by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2)
ultimately show ?thesis
by(auto simp add: acyclic_heap_def)
qed
moreover have "a_all_ptrs_in_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
have "a_all_ptrs_in_heap h'"
proof -
have "a_all_ptrs_in_heap h3"
using \<open>a_all_ptrs_in_heap h2\<close>
apply(auto simp add: a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2
children_eq_h2)[1]
using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
using node_ptr_kinds_eq2_h2 apply auto[1]
apply (metis \<open>known_ptrs h2\<close> \<open>type_wf h3\<close> children_eq_h2 local.get_child_nodes_ok
local.heap_is_wellformed_children_in_heap local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h2
returns_result_select_result wellformed_h2)
- by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2
- disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in node_ptr_kinds_commutes
+ by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2
+ disconnected_nodes_h3 document_ptr_kinds_commutes node_ptr_kinds_commutes
object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD)
have "set children_h3 \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using children_h3 \<open>a_all_ptrs_in_heap h3\<close>
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1]
by (metis children_eq_h2 l_heap_is_wellformed.heap_is_wellformed_children_in_heap
local.l_heap_is_wellformed_axioms node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h'
object_ptr_kinds_M_eq3_h2 wellformed_h2)
then have "set (insert_before_list node reference_child children_h3) \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_in_heap
apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1]
- by (metis (no_types, opaque_lifting) contra_subsetD finite_set_in insert_before_list_in_set
+ by (metis (no_types, opaque_lifting) contra_subsetD insert_before_list_in_set
node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h'
object_ptr_kinds_M_eq3_h2)
then show ?thesis
using \<open>a_all_ptrs_in_heap h3\<close>
apply(auto simp add: object_ptr_kinds_M_eq3_h' a_all_ptrs_in_heap_def node_ptr_kinds_def
node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1]
using children_eq_h3 children_h'
- apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD)
+ apply (metis (no_types, lifting) children_eq2_h3 select_result_I2 subsetD)
by (metis (no_types) \<open>type_wf h'\<close> disconnected_nodes_eq2_h3 disconnected_nodes_eq_h3
- finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok
+ is_OK_returns_result_I local.get_disconnected_nodes_ok
local.get_disconnected_nodes_ptr_in_heap returns_result_select_result subsetD)
qed
moreover have "a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h3"
proof(auto simp add: a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2
children_eq2_h2 intro!: distinct_concat_map_I)[1]
fix x
assume 1: "x |\<in>| document_ptr_kinds h3"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
show "distinct |h3 \<turnstile> 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_iff_member_fset set_sorted_list_of_set)
+ by (metis (full_types) distinct_remove1 finite_fset set_sorted_list_of_set)
next
fix x y xa
assume 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and 2: "x |\<in>| document_ptr_kinds h3"
and 3: "y |\<in>| document_ptr_kinds h3"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r"
and 6: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r"
show False
proof (cases "x = owner_document")
case True
then have "y \<noteq> 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 \<open>y \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
next
case False
then show ?thesis
proof (cases "y = owner_document")
case True
then show ?thesis
using distinct_concat_map_E(1)[OF 1]
using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2]
apply(auto simp add: True disconnected_nodes_eq2_h2[OF \<open>x \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
next
case False
then show ?thesis
using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6
using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
- disjoint_iff_not_equal finite_fset fmember_iff_member_fset notin_set_remove1 select_result_I2
- set_sorted_list_of_set
- by (metis (no_types, lifting))
+ disjoint_iff_not_equal notin_set_remove1 select_result_I2
+ by (metis (no_types, opaque_lifting))
qed
qed
next
fix x xa xb
assume 1: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h3 \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 2: "xa |\<in>| object_ptr_kinds h3"
and 3: "x \<in> set |h3 \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h3"
and 5: "x \<in> set |h3 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 4
by (metis \<open>type_wf h2\<close> children_eq2_h2 document_ptr_kinds_commutes known_ptrs
local.get_child_nodes_ok local.get_disconnected_nodes_ok
local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr
object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result
wellformed_h2)
show False
proof (cases "xb = owner_document")
case True
then show ?thesis
using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]]
by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1)
next
case False
show ?thesis
using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto
qed
qed
then have "a_distinct_lists h'"
proof(auto simp add: a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3
disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I)[1]
fix x
assume 1: "distinct (concat (map (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))" and
2: "x |\<in>| object_ptr_kinds h'"
have 3: "\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> distinct |h3 \<turnstile> get_child_nodes p|\<^sub>r"
using 1 by (auto elim: distinct_concat_map_E)
show "distinct |h' \<turnstile> 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 (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "x |\<in>| object_ptr_kinds h'"
and 3: "y |\<in>| object_ptr_kinds h'"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h' \<turnstile> get_child_nodes x|\<^sub>r"
and 6: "xa \<in> set |h' \<turnstile> get_child_nodes y|\<^sub>r"
have 7:"set |h3 \<turnstile> get_child_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 \<noteq> 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 \<open>ptr \<noteq> y\<close>])[1]
by (metis (no_types, opaque_lifting) "3" "7" \<open>type_wf h3\<close> children_eq2_h3 disjoint_iff_not_equal
get_child_nodes_ok insert_before_list_in_set known_ptrs local.known_ptrs_known_ptr
object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h'
object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2)
next
case False
then show ?thesis
proof (cases "ptr = y")
case True
then show ?thesis
using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6
apply(auto simp add: True children_eq2_h3[OF \<open>ptr \<noteq> x\<close>])[1]
by (metis (no_types, opaque_lifting) "2" "4" "7" IntI \<open>known_ptrs h3\<close> \<open>type_wf h'\<close>
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 \<open>ptr \<noteq> x\<close>] children_eq2_h3[OF \<open>ptr \<noteq> y\<close>] 5 6 7 by auto
qed
qed
next
fix x xa xb
assume 1: " (\<Union>x\<in>fset (object_ptr_kinds h'). set |h3 \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r) = {} "
and 2: "xa |\<in>| object_ptr_kinds h'"
and 3: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h'"
and 5: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 3 4 5
proof -
have "\<forall>h d. \<not> type_wf h \<or> d |\<notin>| document_ptr_kinds h \<or> h \<turnstile> ok get_disconnected_nodes d"
using local.get_disconnected_nodes_ok by satx
then have "h' \<turnstile> ok get_disconnected_nodes xb"
using "4" \<open>type_wf h'\<close> by fastforce
then have f1: "h3 \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
by (simp add: disconnected_nodes_eq_h3)
have "xa |\<in>| object_ptr_kinds h3"
using "2" object_ptr_kinds_M_eq3_h' by blast
then show ?thesis
using f1 \<open>local.a_distinct_lists h3\<close> local.distinct_lists_no_parent by fastforce
qed
show False
proof (cases "ptr = xa")
case True
show ?thesis
using 6 node_not_in_disconnected_nodes 3 4 5 select_result_I2[OF children_h']
select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3
by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M
\<open>a_distinct_lists h3\<close> \<open>type_wf h'\<close> disconnected_nodes_eq_h3
distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok
insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result)
next
case False
then show ?thesis
using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce
qed
qed
moreover have "a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_M_eq2_h2
object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3
document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2)[1]
apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified]
object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified]
node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1]
apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1]
by (smt (verit) children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2
- disconnected_nodes_h3 finite_set_in in_set_remove1 insert_before_list_in_set object_ptr_kinds_M_eq3_h'
+ disconnected_nodes_h3 in_set_remove1 insert_before_list_in_set object_ptr_kinds_M_eq3_h'
ptr_in_heap select_result_I2)
ultimately show "heap_is_wellformed h'"
by (simp add: heap_is_wellformed_def)
qed
lemma adopt_node_children_remain_distinct:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'"
shows "\<And>ptr' children'.
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children' \<Longrightarrow> distinct children'"
using assms(1) assms(2) assms(3) assms(4) local.adopt_node_preserves_wellformedness
local.heap_is_wellformed_children_distinct
by blast
lemma insert_node_children_remain_distinct:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> a_insert_node ptr new_child reference_child_opt \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "new_child \<notin> set children"
shows "\<And>children'.
h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children' \<Longrightarrow> distinct children'"
proof -
fix children'
assume a1: "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'"
have "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r (insert_before_list new_child reference_child_opt children)"
using assms(4) assms(5) apply(auto simp add: a_insert_node_def elim!: bind_returns_heap_E)[1]
using returns_result_eq set_child_nodes_get_child_nodes assms(2) assms(3)
by (metis is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_child_nodes_pure
local.known_ptrs_known_ptr pure_returns_heap_eq)
moreover have "a_distinct_lists h"
using assms local.heap_is_wellformed_def by blast
then have "\<And>children. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> distinct children"
using assms local.heap_is_wellformed_children_distinct by blast
ultimately show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children' \<Longrightarrow> distinct children'"
using assms(5) assms(6) insert_before_list_distinct returns_result_eq by fastforce
qed
lemma insert_before_children_remain_distinct:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> insert_before ptr new_child child_opt \<rightarrow>\<^sub>h h'"
shows "\<And>ptr' children'.
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children' \<Longrightarrow> distinct children'"
proof -
obtain reference_child owner_document h2 h3 disconnected_nodes_h2 where
reference_child:
"h \<turnstile> (if Some new_child = child_opt then a_next_sibling new_child else return child_opt) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document new_child \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 new_child disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr new_child reference_child \<rightarrow>\<^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 "\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> distinct children"
using adopt_node_children_remain_distinct
using assms(1) assms(2) assms(3) h2
by blast
moreover have "\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> new_child \<notin> set children"
using adopt_node_removes_child
using assms(1) assms(2) assms(3) h2
by blast
moreover have "\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h3
apply(rule reads_writes_preserved)
by (auto simp add: set_disconnected_nodes_get_child_nodes)
ultimately show "\<And>ptr children. h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> distinct children"
using insert_node_children_remain_distinct
by (meson assms(1) assms(2) assms(3) assms(4) insert_before_heap_is_wellformed_preserved(1)
local.heap_is_wellformed_children_distinct)
qed
lemma insert_before_removes_child:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'"
assumes "ptr \<noteq> ptr'"
shows "\<And>children'. h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children' \<Longrightarrow> node \<notin> set children'"
proof -
fix children'
assume a1: "h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'"
obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where
ancestors: "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
reference_child:
"h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
have "known_ptr ptr"
by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I assms(2)
l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document)
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF adopt_node_writes h2]
using assms(3) adopt_node_types_preserved
by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF insert_node_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs )
then have object_ptr_kinds_M_eq2_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have "known_ptrs h2"
using assms object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast
have wellformed_h2: "heap_is_wellformed h2"
using adopt_node_preserves_wellformedness[OF assms(1) h2] assms by simp
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h2: "|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto
have "known_ptrs h3"
using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \<open>known_ptrs h2\<close> by blast
have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h3:
"|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h3: "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto
have "known_ptrs h'"
using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \<open>known_ptrs h3\<close> by blast
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. owner_document \<noteq> doc_ptr
\<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. doc_ptr \<noteq> owner_document
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_h3:
"h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r remove1 node disconnected_nodes_h2"
using h3 set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h3:
"\<And>ptr' children. ptr \<noteq> ptr'
\<Longrightarrow> h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr \<noteq> ptr' \<Longrightarrow> |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
obtain children_h3 where children_h3: "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h3"
using h' a_insert_node_def by auto
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r insert_before_list node reference_child children_h3"
using h' \<open>type_wf h3\<close> \<open>known_ptr ptr\<close>
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 |\<in>| object_ptr_kinds h3"
using children_h3 get_child_nodes_ptr_in_heap by blast
have node_in_heap: "node |\<in>| node_ptr_kinds h"
using h2 adopt_node_child_in_heap by fast
have child_not_in_any_children:
"\<And>p children. h2 \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children \<Longrightarrow> node \<notin> set children"
using assms(1) assms(2) assms(3) h2 local.adopt_node_removes_child by blast
show "node \<notin> set children'"
using a1 assms(5) child_not_in_any_children children_eq_h2 children_eq_h3 by blast
qed
lemma ensure_pre_insertion_validity_ok:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "\<not>is_character_data_ptr_kind parent"
assumes "cast node \<notin> set |h \<turnstile> get_ancestors parent|\<^sub>r"
assumes "h \<turnstile> get_parent ref \<rightarrow>\<^sub>r Some parent"
assumes "is_document_ptr parent \<Longrightarrow> h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r []"
assumes "is_document_ptr parent \<Longrightarrow> \<not>is_character_data_ptr_kind node"
shows "h \<turnstile> ok (a_ensure_pre_insertion_validity node parent (Some ref))"
proof -
have "h \<turnstile> (if is_character_data_ptr_kind parent
then error HierarchyRequestError else return ()) \<rightarrow>\<^sub>r ()"
using assms
by (simp add: assms(4))
moreover have "h \<turnstile> do {
ancestors \<leftarrow> get_ancestors parent;
(if cast node \<in> set ancestors then error HierarchyRequestError else return ())
} \<rightarrow>\<^sub>r ()"
using assms(6)
apply(auto intro!: bind_pure_returns_result_I)[1]
using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap
by auto
moreover have "h \<turnstile> do {
(case Some ref of
Some child \<Rightarrow> do {
child_parent \<leftarrow> get_parent child;
(if child_parent \<noteq> Some parent then error NotFoundError else return ())}
| None \<Rightarrow> return ())
} \<rightarrow>\<^sub>r ()"
using assms(7)
by(auto split: option.splits)
moreover have "h \<turnstile> do {
children \<leftarrow> get_child_nodes parent;
(if children \<noteq> [] \<and> is_document_ptr parent
then error HierarchyRequestError else return ())
} \<rightarrow>\<^sub>r ()"
by (smt (verit, best) assms(5) assms(7) assms(8) bind_pure_returns_result_I2 calculation(1)
is_OK_returns_result_I local.get_child_nodes_pure local.get_parent_child_dual returns_result_eq)
moreover have "h \<turnstile> do {
(if is_character_data_ptr node \<and> is_document_ptr parent
then error HierarchyRequestError else return ())
} \<rightarrow>\<^sub>r ()"
using assms
using is_character_data_ptr_kind_none by force
ultimately show ?thesis
unfolding a_ensure_pre_insertion_validity_def
apply(intro bind_is_OK_pure_I)
apply auto[1]
apply auto[1]
apply auto[1]
using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap
apply blast
apply auto[1]
apply auto[1]
using assms(6)
apply auto[1]
using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap
apply auto[1]
apply (smt (verit) bind_returns_heap_E is_OK_returns_heap_E local.get_parent_pure pure_def
pure_returns_heap_eq return_returns_heap returns_result_eq)
apply(blast)
using local.get_child_nodes_pure
apply blast
apply (meson assms(7) is_OK_returns_result_I local.get_parent_child_dual)
apply (simp)
apply (smt (verit) assms(5) assms(8) is_OK_returns_result_I returns_result_eq)
by(auto)
qed
end
locale l_insert_before_wf2 = l_type_wf + l_known_ptrs + l_insert_before_defs
+ l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs +
assumes insert_before_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> insert_before ptr child ref \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes insert_before_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> insert_before ptr child ref \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes insert_before_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> insert_before ptr child ref \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs
get_child_nodes get_child_nodes_locs set_child_nodes
set_child_nodes_locs get_ancestors get_ancestors_locs
adopt_node adopt_node_locs set_disconnected_nodes
set_disconnected_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_owner_document insert_before
insert_before_locs append_child type_wf known_ptr known_ptrs
heap_is_wellformed parent_child_rel remove_child
remove_child_locs get_root_node get_root_node_locs
by(simp add: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma insert_before_wf2_is_l_insert_before_wf2 [instances]:
"l_insert_before_wf2 type_wf known_ptr known_ptrs insert_before heap_is_wellformed"
apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1]
using insert_before_heap_is_wellformed_preserved apply(fast, fast, fast)
done
locale l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_child_nodes_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_remove_child_wf2
begin
lemma next_sibling_ok:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "node_ptr |\<in>| node_ptr_kinds h"
shows "h \<turnstile> ok (a_next_sibling node_ptr)"
proof -
have "known_ptr (cast node_ptr)"
using assms(2) assms(4) local.known_ptrs_known_ptr node_ptr_kinds_commutes by blast
then show ?thesis
using assms
apply(auto simp add: a_next_sibling_def intro!: bind_is_OK_pure_I split: option.splits list.splits)[1]
using get_child_nodes_ok local.get_parent_parent_in_heap local.known_ptrs_known_ptr by blast
qed
lemma remove_child_ok:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "h \<turnstile> ok (remove_child ptr child)"
proof -
have "ptr |\<in>| object_ptr_kinds h"
using assms(4) local.get_child_nodes_ptr_in_heap by blast
have "child |\<in>| node_ptr_kinds h"
using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap by blast
have "\<not>is_character_data_ptr ptr"
proof (rule ccontr, simp)
assume "is_character_data_ptr ptr"
then have "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
using \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(simp add: get_child_nodes_def a_get_child_nodes_tups_def)
apply(split invoke_splits)+
by(auto simp add: get_child_nodes\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits)
then
show False
using assms returns_result_eq by fastforce
qed
have "is_character_data_ptr child \<Longrightarrow> \<not>is_document_ptr_kind ptr"
proof (rule ccontr, simp)
assume "is_character_data_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child"
and "is_document_ptr_kind ptr"
then show False
using assms
using \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(simp add: get_child_nodes_def a_get_child_nodes_tups_def)
apply(split invoke_splits)+
apply(auto split: option.splits)[1]
apply (meson invoke_empty is_OK_returns_result_I)
apply (meson invoke_empty is_OK_returns_result_I)
by(auto simp add: get_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits)
qed
obtain owner_document where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document"
by (meson \<open>child |\<in>| node_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.get_owner_document_ok node_ptr_kinds_commutes)
obtain disconnected_nodes_h where
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h"
by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_disconnected_nodes_ok
local.get_owner_document_owner_document_in_heap owner_document)
obtain h2 where
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2"
by (meson assms(1) assms(2) assms(3) is_OK_returns_heap_E
l_set_disconnected_nodes.set_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap
local.l_set_disconnected_nodes_axioms owner_document)
have "known_ptr ptr"
using assms(2) assms(4) local.known_ptrs_known_ptr
using \<open>ptr |\<in>| object_ptr_kinds h\<close> by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h2]
using set_disconnected_nodes_types_preserved assms(3)
by(auto simp add: reflp_def transp_def)
have "object_ptr_kinds h = object_ptr_kinds h2"
using h2
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_disconnected_nodes_writes])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
have "h2 \<turnstile> ok (set_child_nodes ptr (remove1 child children))"
proof (cases "is_element_ptr_kind ptr")
case True
then show ?thesis
using set_child_nodes_element_ok \<open>known_ptr ptr\<close> \<open>object_ptr_kinds h = object_ptr_kinds h2\<close>
\<open>type_wf h2\<close> assms(4)
using \<open>ptr |\<in>| object_ptr_kinds h\<close> by blast
next
case False
then have "is_document_ptr_kind ptr"
using \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> \<open>\<not>is_character_data_ptr ptr\<close>
by(auto simp add:known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
moreover have "is_document_ptr ptr"
using \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False \<open>\<not>is_character_data_ptr ptr\<close>
by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
ultimately show ?thesis
using assms(4)
apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def)[1]
apply(split invoke_splits)+
apply(auto elim!: bind_returns_result_E2 split: option.splits)[1]
apply(auto simp add: get_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits)[1]
using assms(5) apply auto[1]
using \<open>is_document_ptr_kind ptr\<close> \<open>known_ptr ptr\<close> \<open>object_ptr_kinds h = object_ptr_kinds h2\<close>
\<open>ptr |\<in>| object_ptr_kinds h\<close> \<open>type_wf h2\<close> local.set_child_nodes_document1_ok apply blast
using \<open>is_document_ptr_kind ptr\<close> \<open>known_ptr ptr\<close> \<open>object_ptr_kinds h = object_ptr_kinds h2\<close>
\<open>ptr |\<in>| object_ptr_kinds h\<close> \<open>type_wf h2\<close> is_element_ptr_kind_cast local.set_child_nodes_document2_ok
apply blast
using \<open>\<not> is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\<close> apply blast
by (metis False is_element_ptr_implies_kind option.case_eq_if)
qed
then
obtain h' where
h': "h2 \<turnstile> set_child_nodes ptr (remove1 child children) \<rightarrow>\<^sub>h h'"
by auto
show ?thesis
using assms
apply(auto simp add: remove_child_def
simp add: is_OK_returns_heap_I[OF h2] is_OK_returns_heap_I[OF h']
is_OK_returns_result_I[OF assms(4)] is_OK_returns_result_I[OF owner_document]
is_OK_returns_result_I[OF disconnected_nodes_h]
intro!: bind_is_OK_pure_I[OF get_owner_document_pure]
bind_is_OK_pure_I[OF get_child_nodes_pure]
bind_is_OK_pure_I[OF get_disconnected_nodes_pure]
bind_is_OK_I[rotated, OF h2]
dest!: returns_result_eq[OF assms(4)] returns_result_eq[OF owner_document]
returns_result_eq[OF disconnected_nodes_h]
)[1]
using h2 returns_result_select_result by force
qed
lemma adopt_node_ok:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "document_ptr |\<in>| document_ptr_kinds h"
assumes "child |\<in>| node_ptr_kinds h"
shows "h \<turnstile> ok (adopt_node document_ptr child)"
proof -
obtain old_document where
old_document: "h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r old_document"
by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E local.get_owner_document_ok
node_ptr_kinds_commutes)
then have "h \<turnstile> ok (get_owner_document (cast child))"
by auto
obtain parent_opt where
parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt"
by (meson assms(2) assms(3) is_OK_returns_result_I l_get_owner_document.get_owner_document_ptr_in_heap
local.get_parent_ok local.l_get_owner_document_axioms node_ptr_kinds_commutes old_document
returns_result_select_result)
then have "h \<turnstile> ok (get_parent child)"
by auto
have "h \<turnstile> ok (case parent_opt of Some parent \<Rightarrow> remove_child parent child | None \<Rightarrow> return ())"
apply(auto split: option.splits)[1]
using remove_child_ok
by (metis assms(1) assms(2) assms(3) local.get_parent_child_dual parent_opt)
then
obtain h2 where
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent child | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2"
by auto
have "object_ptr_kinds h = object_ptr_kinds h2"
using h2 apply(simp split: option.splits)
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes])
using remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def)
then
have "old_document |\<in>| document_ptr_kinds h2"
using assms(1) assms(2) assms(3) document_ptr_kinds_commutes
local.get_owner_document_owner_document_in_heap old_document
by blast
have wellformed_h2: "heap_is_wellformed h2"
using h2 remove_child_heap_is_wellformed_preserved assms
by(auto split: option.splits)
have "type_wf h2"
using h2 remove_child_preserves_type_wf assms
by(auto split: option.splits)
have "known_ptrs h2"
using h2 remove_child_preserves_known_ptrs assms
by(auto split: option.splits)
have "object_ptr_kinds h = object_ptr_kinds h2"
using h2 apply(simp split: option.splits)
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes])
using remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have "document_ptr_kinds h = document_ptr_kinds h2"
by(auto simp add: document_ptr_kinds_def)
have "h2 \<turnstile> ok (if document_ptr \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 child old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes document_ptr;
set_disconnected_nodes document_ptr (child # disc_nodes)
} else do {
return ()
})"
proof(cases "document_ptr = old_document")
case True
then show ?thesis
by simp
next
case False
then have "h2 \<turnstile> ok (get_disconnected_nodes old_document)"
by (simp add: \<open>old_document |\<in>| document_ptr_kinds h2\<close> \<open>type_wf h2\<close> local.get_disconnected_nodes_ok)
then obtain old_disc_nodes where
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes"
by auto
have "h2 \<turnstile> ok (set_disconnected_nodes old_document (remove1 child old_disc_nodes))"
by (simp add: \<open>old_document |\<in>| document_ptr_kinds h2\<close> \<open>type_wf h2\<close> local.set_disconnected_nodes_ok)
then obtain h3 where
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 child old_disc_nodes) \<rightarrow>\<^sub>h h3"
by auto
have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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:
"\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h3"
using \<open>type_wf h2\<close>
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h3]
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
moreover have "document_ptr |\<in>| document_ptr_kinds h3"
using \<open>document_ptr_kinds h = document_ptr_kinds h2\<close> assms(4) document_ptr_kinds_eq3_h2 by auto
ultimately have "h3 \<turnstile> ok (get_disconnected_nodes document_ptr)"
by (simp add: local.get_disconnected_nodes_ok)
then obtain disc_nodes where
disc_nodes: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
by auto
have "h3 \<turnstile> ok (set_disconnected_nodes document_ptr (child # disc_nodes))"
using \<open>document_ptr |\<in>| document_ptr_kinds h3\<close> \<open>type_wf h3\<close> local.set_disconnected_nodes_ok by auto
then obtain h' where
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (child # disc_nodes) \<rightarrow>\<^sub>h h'"
by auto
then show ?thesis
using False
using \<open>h2 \<turnstile> ok get_disconnected_nodes old_document\<close>
using \<open>h3 \<turnstile> ok get_disconnected_nodes document_ptr\<close>
apply(auto dest!: returns_result_eq[OF old_disc_nodes] returns_result_eq[OF disc_nodes]
intro!: bind_is_OK_I[rotated, OF h3] bind_is_OK_pure_I[OF get_disconnected_nodes_pure] )[1]
using \<open>h2 \<turnstile> ok set_disconnected_nodes old_document (remove1 child old_disc_nodes)\<close> by auto
qed
then obtain h' where
h': "h2 \<turnstile> (if document_ptr \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 child old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes document_ptr;
set_disconnected_nodes document_ptr (child # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^sub>h h'"
by auto
show ?thesis
using \<open>h \<turnstile> ok (get_owner_document (cast child))\<close>
using \<open>h \<turnstile> ok (get_parent child)\<close>
using h2 h'
apply(auto simp add: adopt_node_def
simp add: is_OK_returns_heap_I[OF h2]
intro!: bind_is_OK_pure_I[OF get_owner_document_pure]
bind_is_OK_pure_I[OF get_parent_pure]
bind_is_OK_I[rotated, OF h2]
dest!: returns_result_eq[OF parent_opt] returns_result_eq[OF old_document])[1]
using \<open>h \<turnstile> ok (case parent_opt of None \<Rightarrow> return () | Some parent \<Rightarrow> remove_child parent child)\<close>
by auto
qed
lemma insert_node_ok:
assumes "known_ptr parent" and "type_wf h"
assumes "parent |\<in>| object_ptr_kinds h"
assumes "\<not>is_character_data_ptr_kind parent"
assumes "is_document_ptr parent \<Longrightarrow> h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r []"
assumes "is_document_ptr parent \<Longrightarrow> \<not>is_character_data_ptr_kind node"
assumes "known_ptr (cast node)"
shows "h \<turnstile> ok (a_insert_node parent node ref)"
proof(auto simp add: a_insert_node_def get_child_nodes_ok[OF assms(1) assms(2) assms(3)]
intro!: bind_is_OK_pure_I)
fix children'
assume "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children'"
show "h \<turnstile> ok set_child_nodes parent (insert_before_list node ref children')"
proof (cases "is_element_ptr_kind parent")
case True
then show ?thesis
using set_child_nodes_element_ok
using assms(1) assms(2) assms(3) by blast
next
case False
then have "is_document_ptr_kind parent"
using assms(4) assms(1)
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then have "is_document_ptr parent"
using assms(1)
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children" and "children = []"
using assms(5) by blast
have "insert_before_list node ref children' = [node]"
by (metis \<open>children = []\<close> \<open>h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children'\<close> append.left_neutral
children insert_Nil l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.insert_before_list.elims
l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.insert_before_list.simps(3) neq_Nil_conv returns_result_eq)
moreover have "\<not>is_character_data_ptr_kind node"
using \<open>is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r parent\<close> assms(6) by blast
then have "is_element_ptr_kind node"
by (metis (no_types, lifting) CharacterDataClass.a_known_ptr_def DocumentClass.a_known_ptr_def
ElementClass.a_known_ptr_def NodeClass.a_known_ptr_def assms(7) cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject
document_ptr_no_node_ptr_cast is_character_data_ptr_kind_none is_document_ptr_kind_none
is_element_ptr_implies_kind is_node_ptr_kind_cast local.known_ptr_impl node_ptr_casts_commute3
option.case_eq_if)
ultimately
show ?thesis
using set_child_nodes_document2_ok
by (metis \<open>is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r parent\<close> assms(1) assms(2) assms(3) assms(5)
is_document_ptr_kind_none option.case_eq_if)
qed
qed
lemma insert_before_ok:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "parent |\<in>| object_ptr_kinds h"
assumes "node |\<in>| node_ptr_kinds h"
assumes "\<not>is_character_data_ptr_kind parent"
assumes "cast node \<notin> set |h \<turnstile> get_ancestors parent|\<^sub>r"
assumes "h \<turnstile> get_parent ref \<rightarrow>\<^sub>r Some parent"
assumes "is_document_ptr parent \<Longrightarrow> h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r []"
assumes "is_document_ptr parent \<Longrightarrow> \<not>is_character_data_ptr_kind node"
shows "h \<turnstile> ok (insert_before parent node (Some ref))"
proof -
have "h \<turnstile> ok (a_ensure_pre_insertion_validity node parent (Some ref))"
using assms ensure_pre_insertion_validity_ok by blast
have "h \<turnstile> ok (if Some node = Some ref
then a_next_sibling node
else return (Some ref))" (is "h \<turnstile> ok ?P")
apply(auto split: if_splits)[1]
using assms(1) assms(2) assms(3) assms(5) next_sibling_ok by blast
then obtain reference_child where
reference_child: "h \<turnstile> ?P \<rightarrow>\<^sub>r reference_child"
by auto
obtain owner_document where
owner_document: "h \<turnstile> get_owner_document parent \<rightarrow>\<^sub>r owner_document"
using assms get_owner_document_ok
by (meson returns_result_select_result)
then have "h \<turnstile> ok (get_owner_document parent)"
by auto
have "owner_document |\<in>| document_ptr_kinds h"
using assms(1) assms(2) assms(3) local.get_owner_document_owner_document_in_heap owner_document
by blast
obtain h2 where
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2"
by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_heap_E adopt_node_ok
l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms
local.get_owner_document_owner_document_in_heap owner_document)
then have "h \<turnstile> ok (adopt_node owner_document node)"
by auto
have "object_ptr_kinds h = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF adopt_node_writes h2])
using adopt_node_pointers_preserved
apply blast
by (auto simp add: reflp_def transp_def)
then have "document_ptr_kinds h = document_ptr_kinds h2"
by(auto simp add: document_ptr_kinds_def)
have "heap_is_wellformed h2"
using h2 adopt_node_preserves_wellformedness assms by blast
have "known_ptrs h2"
using h2 adopt_node_preserves_known_ptrs assms by blast
have "type_wf h2"
using h2 adopt_node_preserves_type_wf assms by blast
obtain disconnected_nodes_h2 where
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2"
by (metis \<open>document_ptr_kinds h = document_ptr_kinds h2\<close> \<open>type_wf h2\<close> assms(1) assms(2) assms(3)
is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap
owner_document)
obtain h3 where
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3"
by (metis \<open>document_ptr_kinds h = document_ptr_kinds h2\<close> \<open>owner_document |\<in>| document_ptr_kinds h\<close>
\<open>type_wf h2\<close> document_ptr_kinds_def is_OK_returns_heap_E
l_set_disconnected_nodes.set_disconnected_nodes_ok local.l_set_disconnected_nodes_axioms)
have "type_wf h3"
using \<open>type_wf h2\<close>
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h3]
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_disconnected_nodes_writes h3])
unfolding a_remove_child_locs_def
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
have "parent |\<in>| object_ptr_kinds h3"
using \<open>object_ptr_kinds h = object_ptr_kinds h2\<close> assms(4) object_ptr_kinds_M_eq3_h2 by blast
moreover have "known_ptr parent"
using assms(2) assms(4) local.known_ptrs_known_ptr by blast
moreover have "known_ptr (cast node)"
using assms(2) assms(5) local.known_ptrs_known_ptr node_ptr_kinds_commutes by blast
moreover have "is_document_ptr parent \<Longrightarrow> h3 \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r []"
by (metis assms(8) assms(9) distinct.simps(2) distinct_singleton local.get_parent_child_dual
returns_result_eq)
ultimately obtain h' where
h': "h3 \<turnstile> a_insert_node parent node reference_child \<rightarrow>\<^sub>h h'"
using insert_node_ok \<open>type_wf h3\<close> assms by blast
show ?thesis
using \<open>h \<turnstile> ok (a_ensure_pre_insertion_validity node parent (Some ref))\<close>
using reference_child \<open>h \<turnstile> ok (get_owner_document parent)\<close> \<open>h \<turnstile> ok (adopt_node owner_document node)\<close>
h3 h'
apply(auto simp add: insert_before_def
simp add: is_OK_returns_result_I[OF disconnected_nodes_h2]
simp add: is_OK_returns_heap_I[OF h3] is_OK_returns_heap_I[OF h']
intro!: bind_is_OK_I2
bind_is_OK_pure_I[OF ensure_pre_insertion_validity_pure]
bind_is_OK_pure_I[OF next_sibling_pure]
bind_is_OK_pure_I[OF get_owner_document_pure]
bind_is_OK_pure_I[OF get_disconnected_nodes_pure]
dest!: returns_result_eq[OF owner_document] returns_result_eq[OF disconnected_nodes_h2]
returns_heap_eq[OF h2] returns_heap_eq[OF h3]
dest!: sym[of node ref]
)[1]
using returns_result_eq by fastforce
qed
end
interpretation i_insert_before_wf3?: l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs
get_ancestors get_ancestors_locs adopt_node adopt_node_locs set_disconnected_nodes
set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document
insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed
parent_child_rel remove_child remove_child_locs get_root_node get_root_node_locs remove
by(auto simp add: l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_insert_before_wf +
l_insert_before_wf2 +
l_get_child_nodes
begin
lemma append_child_heap_is_wellformed_preserved:
assumes wellformed: "heap_is_wellformed h"
and append_child: "h \<turnstile> append_child ptr node \<rightarrow>\<^sub>h h'"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
using assms
by(auto simp add: append_child_def intro: insert_before_preserves_type_wf
insert_before_preserves_known_ptrs insert_before_heap_is_wellformed_preserved)
lemma append_child_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
assumes "h \<turnstile> append_child ptr node \<rightarrow>\<^sub>h h'"
assumes "node \<notin> set xs"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ [node]"
proof -
obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where
ancestors: "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node None \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: append_child_def insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
have "\<And>parent. |h \<turnstile> get_parent node|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr"
using assms(1) assms(4) assms(6)
by (metis (no_types, lifting) assms(2) assms(3) h2 is_OK_returns_heap_I is_OK_returns_result_E
local.adopt_node_child_in_heap local.get_parent_child_dual local.get_parent_ok
select_result_I2)
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
using get_child_nodes_reads adopt_node_writes h2 assms(4)
apply(rule reads_writes_separate_forwards)
using \<open>\<And>parent. |h \<turnstile> get_parent node|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>
apply(auto simp add: adopt_node_locs_def remove_child_locs_def)[1]
by (meson local.set_child_nodes_get_child_nodes_different_pointers)
have "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
using get_child_nodes_reads set_disconnected_nodes_writes h3 \<open>h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs\<close>
apply(rule reads_writes_separate_forwards)
by(auto)
have "ptr |\<in>| object_ptr_kinds h"
by (meson ancestors is_OK_returns_result_I local.get_ancestors_ptr_in_heap)
then
have "known_ptr ptr"
using assms(3)
using local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF adopt_node_writes h2]
using adopt_node_types_preserved \<open>type_wf h\<close>
by(auto simp add: adopt_node_locs_def remove_child_locs_def reflp_def transp_def split: if_splits)
then
have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h3]
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs@[node]"
using h'
apply(auto simp add: a_insert_node_def
dest!: bind_returns_heap_E3[rotated, OF \<open>h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs\<close>
get_child_nodes_pure, rotated])[1]
using \<open>type_wf h3\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close>
by metis
qed
lemma append_child_for_all_on_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
assumes "h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
assumes "set nodes \<inter> set xs = {}"
assumes "distinct nodes"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs@nodes"
using assms
apply(induct nodes arbitrary: h xs)
apply(simp)
proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a
assume 0: "(\<And>h xs. heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs \<Longrightarrow> h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'
\<Longrightarrow> set nodes \<inter> set xs = {} \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ nodes)"
and 1: "heap_is_wellformed h"
and 2: "type_wf h"
and 3: "known_ptrs h"
and 4: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
and 5: "h \<turnstile> append_child ptr a \<rightarrow>\<^sub>r ()"
and 6: "h \<turnstile> append_child ptr a \<rightarrow>\<^sub>h h'a"
and 7: "h'a \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
and 8: "a \<notin> set xs"
and 9: "set nodes \<inter> set xs = {}"
and 10: "a \<notin> set nodes"
and 11: "distinct nodes"
then have "h'a \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ [a]"
using append_child_children 6
using "1" "2" "3" "4" "8" by blast
moreover have "heap_is_wellformed h'a" and "type_wf h'a" and "known_ptrs h'a"
using insert_before_heap_is_wellformed_preserved insert_before_preserves_known_ptrs
insert_before_preserves_type_wf 1 2 3 6 append_child_def
by metis+
moreover have "set nodes \<inter> set (xs @ [a]) = {}"
using 9 10
by auto
ultimately show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ a # nodes"
using 0 7
by fastforce
qed
lemma append_child_for_all_on_no_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
assumes "h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
assumes "distinct nodes"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r nodes"
using assms append_child_for_all_on_children
by force
end
locale l_append_child_wf = l_type_wf + l_known_ptrs + l_append_child_defs + l_heap_is_wellformed_defs +
assumes append_child_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes append_child_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes append_child_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent
get_parent_locs remove_child remove_child_locs
get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs
adopt_node adopt_node_locs known_ptr type_wf get_child_nodes
get_child_nodes_locs known_ptrs set_child_nodes
set_child_nodes_locs remove get_ancestors get_ancestors_locs
insert_before insert_before_locs append_child heap_is_wellformed
parent_child_rel
by(auto simp add: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr
known_ptrs append_child heap_is_wellformed"
apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1]
using append_child_heap_is_wellformed_preserved by fast+
subsection \<open>create\_element\<close>
locale l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel +
l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs +
l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs
get_disconnected_nodes get_disconnected_nodes_locs +
l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr +
l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs +
l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr
get_child_nodes get_child_nodes_locs +
l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs
get_child_nodes get_child_nodes_locs +
l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs +
l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs +
l_new_element type_wf +
l_known_ptrs known_ptr known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
begin
lemma create_element_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(2)
by(auto simp add: create_element_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_element_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_element_ptr)"
using \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr
by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []"
using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr]
new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes
by blast
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_child_nodes)
then have children_eq2_h2: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_element_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3: "\<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_element_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close>
using \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def acyclic_heap_def)
also have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting)
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally have "a_acyclic_heap h'"
by (simp add: acyclic_heap_def)
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h2"
apply(auto simp add: a_all_ptrs_in_heap_def)[1]
apply (metis \<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> assms(1)
assms(3) funion_iff local.get_child_nodes_ok local.known_ptrs_known_ptr
local.parent_child_rel_child_in_heap local.parent_child_rel_child_nodes2 node_ptr_kinds_commutes
node_ptr_kinds_eq_h returns_result_select_result)
by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funion_iff
local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h
returns_result_select_result)
then have "a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "a_all_ptrs_in_heap h'"
by (smt (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2
- disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 element_ptr_kinds_commutes finite_set_in
+ disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 element_ptr_kinds_commutes
h' h2 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
new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3
object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subsetD subsetI)
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_element_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
heap_is_wellformed_children_in_heap
by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp
fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result)
then have "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_element_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_element_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_element_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_element_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
by (metis \<open>local.a_distinct_lists h\<close> \<open>type_wf h2\<close> disconnected_nodes_eq_h document_ptr_kinds_eq_h
local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
children_eq2_h2 object_ptr_kinds_eq_h2)
then have "a_distinct_lists h'"
proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3
object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3
intro!: distinct_concat_map_I)[1]
fix x
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> 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) \<open>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 \<notin> set disc_nodes_h3\<close>
\<open>a_distinct_lists h3\<close> \<open>type_wf h'\<close> disc_nodes_h3 distinct.simps(2)
distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq
returns_result_select_result)
next
fix x y xa
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using calculation by(auto dest: distinct_concat_map_E(1))
ultimately show "False"
apply(-)
apply(cases "x = document_ptr")
apply (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>local.a_all_ptrs_in_heap h\<close>
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'
+ disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
by (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>local.a_all_ptrs_in_heap h\<close>
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'
+ disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply -
apply(cases "xb = document_ptr")
apply (metis (no_types, opaque_lifting) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3
\<Longrightarrow> 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 \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
\<open>a_distinct_lists h3\<close> children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h'
select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes)
by (metis "3" "4" "5" "6" \<open>a_distinct_lists h3\<close> \<open>type_wf h3\<close> children_eq2_h3
distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result)
qed
have "a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(auto simp add: a_owner_document_valid_def)[1]
apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1]
apply(auto simp add: object_ptr_kinds_eq_h2)[1]
apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1]
apply(auto simp add: document_ptr_kinds_eq_h2)[1]
apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1]
apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1]
apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric]
disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3)[1]
apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1)
local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
apply(simp add: object_ptr_kinds_eq_h)
by (smt (verit) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h
children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
- disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2)
+ disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h h' list.set_intros(2)
local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_create_element_wf?: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf
get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
set_tag_name set_tag_name_locs
set_disconnected_nodes set_disconnected_nodes_locs create_element
using instances
by(auto simp add: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>create\_character\_data\<close>
locale l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
+ l_new_character_data_get_disconnected_nodes
get_disconnected_nodes get_disconnected_nodes_locs
+ l_set_val_get_disconnected_nodes
type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs
+ l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr
+ l_new_character_data_get_child_nodes
type_wf known_ptr get_child_nodes get_child_nodes_locs
+ l_set_val_get_child_nodes
type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs
+ l_set_disconnected_nodes_get_child_nodes
set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs
+ l_set_disconnected_nodes
type_wf set_disconnected_nodes set_disconnected_nodes_locs
+ l_set_disconnected_nodes_get_disconnected_nodes
type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs
+ l_new_character_data
type_wf
+ l_known_ptrs
known_ptr known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes ::
"(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) character_data_ptr) prog"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
begin
lemma create_character_data_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(2)
by(auto simp add: create_character_data_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then
have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_character_data_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2
get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h:
"\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2
get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_child_nodes)
then have children_eq2_h2:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_character_data_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_val_writes h3]
using set_val_types_preserved
by(auto simp add: reflp_def transp_def)
then show "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_character_data_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close> using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def acyclic_heap_def)
also have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally have "a_acyclic_heap h'"
by (simp add: acyclic_heap_def)
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h2"
apply(auto simp add: a_all_ptrs_in_heap_def)[1]
using node_ptr_kinds_eq_h \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
- apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \<open>parent_child_rel h = parent_child_rel h2\<close>
- children_eq2_h finite_set_in finsert_iff funion_finsert_right local.parent_child_rel_child
+ apply (metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M \<open>parent_child_rel h = parent_child_rel h2\<close>
+ children_eq2_h finsert_iff funion_finsert_right local.parent_child_rel_child
local.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h
select_result_I2 subsetD sup_bot.right_neutral)
by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funionI1
local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap
node_ptr_kinds_eq_h returns_result_select_result)
then have "a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "a_all_ptrs_in_heap h'"
by (smt (verit) character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3
- finite_set_in h' h2 local.a_all_ptrs_in_heap_def
+ h' h2 local.a_all_ptrs_in_heap_def
local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr
new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3
object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_character_data_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
heap_is_wellformed_children_in_heap
by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp
fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result)
then have "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_character_data_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_character_data_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr
returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
by (metis \<open>local.a_distinct_lists h\<close> \<open>type_wf h2\<close> disconnected_nodes_eq_h document_ptr_kinds_eq_h
local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
children_eq2_h2 object_ptr_kinds_eq_h2)[1]
then have "a_distinct_lists h'"
proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3
object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1]
fix x
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> 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) \<open>cast new_character_data_ptr \<notin> set disc_nodes_h3\<close>
\<open>a_distinct_lists h3\<close> \<open>type_wf h'\<close> disc_nodes_h3 distinct.simps(2)
distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq
returns_result_select_result)
next
fix x y xa
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using calculation by(auto dest: distinct_concat_map_E(1))
ultimately show "False"
by (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>local.a_all_ptrs_in_heap h\<close> 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'
+ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply(cases "xb = document_ptr")
apply (metis (no_types, opaque_lifting) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
\<open>a_distinct_lists h3\<close> children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h'
select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes)
by (metis "3" "4" "5" "6" \<open>a_distinct_lists h3\<close> \<open>type_wf h3\<close> children_eq2_h3
distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result)
qed
have "a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(simp add: a_owner_document_valid_def)
apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )
apply(simp add: object_ptr_kinds_eq_h2)
apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )
apply(simp add: document_ptr_kinds_eq_h2)
apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )
apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )
apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1]
apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1)
local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
apply(simp add: object_ptr_kinds_eq_h)
by (smt (verit) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h
- disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2)
+ disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h h' list.set_intros(2)
local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
have "known_ptr (cast new_character_data_ptr)"
using \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
local.create_character_data_known_ptr by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2
by blast
then
have "known_ptrs h3"
using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast
then
show "known_ptrs h'"
using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast
show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_create_character_data_wf?: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes
set_disconnected_nodes_locs create_character_data known_ptrs
using instances
by (auto simp add: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>create\_document\<close>
locale l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
+ l_new_document_get_disconnected_nodes
get_disconnected_nodes get_disconnected_nodes_locs
+ l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
create_document
+ l_new_document_get_child_nodes
type_wf known_ptr get_child_nodes get_child_nodes_locs
+ l_new_document
type_wf
+ l_known_ptrs
known_ptr known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_document :: "((_) heap, exception, (_) document_ptr) prog"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
begin
lemma create_document_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_document \<rightarrow>\<^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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr" and
h': "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
using assms(2)
apply(simp add: create_document_def)
using new_document_ok by blast
have "new_document_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have "new_document_ptr |\<notin>| 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 |\<notin>| object_ptr_kinds h"
by simp
have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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)
+ by (auto simp add: document_ptr_kinds_def)
have children_eq:
"\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h' get_child_nodes_new_document[rotated, OF new_document_ptr h']
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2: "\<And>ptr'. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by (metis(full_types) \<open>\<And>thesis. (\<And>new_document_ptr.
\<lbrakk>h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr; h \<turnstile> new_document \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+
then have disconnected_nodes_eq2_h: "\<And>doc_ptr. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
using h' local.new_document_no_disconnected_nodes new_document_ptr by blast
have "type_wf h'"
using \<open>type_wf h\<close> new_document_types_preserved h' by blast
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def acyclic_heap_def)
also have "parent_child_rel h = parent_child_rel h'"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h'"
by (simp add: object_ptr_kinds_eq)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_document_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2 empty_iff empty_set image_eqI select_result_I2)
qed
finally have "a_acyclic_heap h'"
by (simp add: acyclic_heap_def)
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def)[1]
using ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close> assms(1) children_eq fset_of_list_elem
local.heap_is_wellformed_children_in_heap local.parent_child_rel_child
local.parent_child_rel_parent_in_heap node_ptr_kinds_eq
apply (metis (no_types, lifting) \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
- children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq
+ children_eq2 finsert_iff funion_finsert_right object_ptr_kinds_eq
select_result_I2 subsetD sup_bot.right_neutral)
by (metis (no_types, lifting) \<open>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 |\<notin>| object_ptr_kinds h\<close>
\<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
\<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close> \<open>type_wf h'\<close> assms(1) disconnected_nodes_eq_h
local.get_disconnected_nodes_ok
local.heap_is_wellformed_disc_nodes_in_heap local.parent_child_rel_child
local.parent_child_rel_parent_in_heap
node_ptr_kinds_eq returns_result_select_result select_result_I2)
have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h'"
using \<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close>
\<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: children_eq2[symmetric] a_distinct_lists_def insort_split object_ptr_kinds_eq
document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(auto simp add: dest: distinct_concat_map_E)[1]
apply(auto simp add: dest: distinct_concat_map_E)[1]
using \<open>new_document_ptr |\<notin>| document_ptr_kinds h\<close>
apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1]
using disconnected_nodes_eq_h
apply (metis assms(1) assms(3) disconnected_nodes_eq2_h local.get_disconnected_nodes_ok
local.heap_is_wellformed_disconnected_nodes_distinct
returns_result_select_result)
proof -
fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr"
assume a1: "x \<noteq> y"
assume a2: "x |\<in>| document_ptr_kinds h"
assume a3: "x \<noteq> new_document_ptr"
assume a4: "y |\<in>| document_ptr_kinds h"
assume a5: "y \<noteq> new_document_ptr"
assume a6: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
assume a7: "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
assume a8: "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
have f9: "xa \<in> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using a7 a3 disconnected_nodes_eq2_h by presburger
have f10: "xa \<in> set |h \<turnstile> get_disconnected_nodes y|\<^sub>r"
using a8 a5 disconnected_nodes_eq2_h by presburger
have f11: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))"
using a4 by simp
have "x \<in> 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' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
and 1: "h' \<turnstile> 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) \<rightarrow>\<^sub>r []"
and 2: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
and 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
and 4: "(\<Union>x\<in>fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h). set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 5: "x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
and 7: "xa |\<in>| object_ptr_kinds h"
and 8: "xa \<noteq> 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 |\<in>| document_ptr_kinds h"
and 10: "xb \<noteq> new_document_ptr"
then show "False"
by (metis \<open>local.a_distinct_lists h\<close> assms(3) disconnected_nodes_eq2_h
local.distinct_lists_no_parent local.get_disconnected_nodes_ok
returns_result_select_result)
qed
have "a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def)[1]
by (metis \<open>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 |\<notin>| object_ptr_kinds h\<close>
- children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in
+ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes
funion_iff node_ptr_kinds_eq object_ptr_kinds_eq)
show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_create_document_wf?: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
set_val set_val_locs set_disconnected_nodes
set_disconnected_nodes_locs create_document known_ptrs
using instances
by (auto simp add: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
end
diff --git a/thys/Core_DOM/standard/classes/ElementClass.thy b/thys/Core_DOM/standard/classes/ElementClass.thy
--- a/thys/Core_DOM/standard/classes/ElementClass.thy
+++ b/thys/Core_DOM/standard/classes/ElementClass.thy
@@ -1,328 +1,335 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Element\<close>
text\<open>In this theory, we introduce the types for the Element class.\<close>
theory ElementClass
imports
"NodeClass"
"ShadowRootPointer"
begin
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, define
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
type_synonym attr_key = DOMString
type_synonym attr_value = DOMString
type_synonym attrs = "(attr_key, attr_value) fmap"
type_synonym tag_name = DOMString
record ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr) RElement = RNode +
nothing :: unit
tag_name :: tag_name
child_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
attrs :: attrs
shadow_root_opt :: "'shadow_root_ptr shadow_root_ptr option"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_scheme"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node
= "(('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext
+ 'Node) Node"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object
= "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_ext + 'Node) Object"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object"
type_synonym
('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap
= "('document_ptr document_ptr + 'shadow_root_ptr shadow_root_ptr + 'object_ptr,
'element_ptr element_ptr + 'character_data_ptr character_data_ptr + 'node_ptr, 'Object,
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext +
'Node) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition element_ptr_kinds :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where
"element_ptr_kinds heap =
the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_element_ptr_kind (node_ptr_kinds heap)))"
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) =
{|element_ptr|} |\<union>| element_ptr_kinds h"
- apply(auto simp add: element_ptr_kinds_def)[1]
- by force
+ by (auto simp add: element_ptr_kinds_def)
definition element_ptrs :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where
"element_ptrs heap = ffilter is_element_ptr (element_ptr_kinds heap)"
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Node \<Rightarrow> (_) Element option"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node =
(case RNode.more node of Inl element \<Rightarrow> Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Object \<Rightarrow> (_) Element option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj \<equiv> (case cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj of Some node \<Rightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node | None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
definition cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Element \<Rightarrow> (_) Node"
where
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = RNode.extend (RNode.truncate element) (Inl (RNode.more element))"
adhoc_overloading cast cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e
abbreviation cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) Element \<Rightarrow> (_) Object"
where
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr)"
adhoc_overloading cast cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
consts is_element_kind :: 'a
definition is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Node \<Rightarrow> bool"
where
"is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr \<longleftrightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
adhoc_overloading is_element_kind is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e
lemmas is_element_kind_def = is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
abbreviation is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) Object \<Rightarrow> bool"
where
"is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
adhoc_overloading is_element_kind is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma element_ptr_kinds_commutes [simp]:
"cast element_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> element_ptr |\<in>| element_ptr_kinds h"
- apply(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)[1]
- by (metis (no_types, lifting) element_ptr_casts_commute2 ffmember_filter fimage_eqI
- fset.map_comp is_element_ptr_kind_none node_ptr_casts_commute3
- node_ptr_kinds_commutes node_ptr_kinds_def option.sel option.simps(3))
+proof (rule iffI)
+ show "cast element_ptr |\<in>| node_ptr_kinds h \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h"
+ by (metis (no_types, opaque_lifting) element_ptr_casts_commute2 element_ptr_kinds_def
+ ffmember_filter fimage_eqI is_element_ptr_kind_cast option.sel)
+next
+ show "element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> cast element_ptr |\<in>| node_ptr_kinds h"
+ by (auto simp: node_ptr_kinds_def element_ptr_kinds_def)
+qed
definition get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Element option"
where
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h = Option.bind (get\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) h) cast"
adhoc_overloading get get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_type_wf_def\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (NodeClass.type_wf h \<and> (\<forall>element_ptr \<in> fset (element_ptr_kinds h).
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t: "type_wf h \<Longrightarrow> ElementClass.type_wf h"
sublocale l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t \<subseteq> l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
apply(unfold_locales)
using NodeClass.a_type_wf_def
by (meson ElementClass.a_type_wf_def l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
locale l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas by unfold_locales
lemma get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf:
assumes "type_wf h"
shows "element_ptr |\<in>| element_ptr_kinds h \<longleftrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h \<noteq> None"
- using l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
- apply(simp add: type_wf_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf bind_eq_None_conv element_ptr_kinds_commutes notin_fset
- option.distinct(1))
+proof (rule iffI)
+ show "element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> get element_ptr h \<noteq> None"
+ using ElementClass.a_type_wf_def assms type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast
+next
+ show "get element_ptr h \<noteq> None \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h"
+ by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf assms bind.bind_lzero element_ptr_kinds_commutes
+ get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def local.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e)
+qed
+
end
global_interpretation l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf
by unfold_locales
definition put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) Element \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) (cast element)"
adhoc_overloading put put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h = h'"
shows "element_ptr |\<in>| element_ptr_kinds h'"
using assms
unfolding put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def element_ptr_kinds_def
by (metis element_ptr_kinds_commutes element_ptr_kinds_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap)
lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs:
assumes "put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast element_ptr|}"
using assms
by (simp add: put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs)
lemma cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]:
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e x = cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def)
by (metis (full_types) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = None \<longleftrightarrow> \<not> (\<exists>element. cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node)"
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)[1]
by (metis (full_types) RNode.select_convs(2) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = Some element \<longleftrightarrow> cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node"
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv [simp]: "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element) = Some element"
by simp
lemma get_elment_ptr_simp1 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h) = Some element"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_elment_ptr_simp2 [simp]:
"element_ptr \<noteq> element_ptr'
\<Longrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' element h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
abbreviation "create_element_obj tag_name_arg child_nodes_arg attrs_arg shadow_root_opt_arg
\<equiv> \<lparr> RObject.nothing = (), RNode.nothing = (), RElement.nothing = (),
tag_name = tag_name_arg, Element.child_nodes = child_nodes_arg, attrs = attrs_arg,
shadow_root_opt = shadow_root_opt_arg, \<dots> = None \<rparr>"
definition new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) heap \<Rightarrow> ((_) element_ptr \<times> (_) heap)"
where
"new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_element_ptr = element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref
|`| (element_ptrs h)))))
in
(new_element_ptr, put new_element_ptr (create_element_obj '''' [] fmempty None) h))"
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "new_element_ptr |\<in>| element_ptr_kinds h'"
using assms
unfolding new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
using put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap by blast
lemma new_element_ptr_new:
"element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref |`| element_ptrs h)))) |\<notin>| element_ptrs h"
by (metis Suc_n_not_le_n element_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "new_element_ptr |\<notin>| element_ptr_kinds h"
using assms
unfolding new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis Pair_inject element_ptrs_def ffmember_filter new_element_ptr_new is_element_ptr_ref)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using assms
by (metis Pair_inject new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_element_ptr:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "is_element_ptr new_element_ptr"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
assumes "ptr \<noteq> cast new_element_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
assumes "ptr \<noteq> cast new_element_ptr"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
assumes "ptr \<noteq> new_element_ptr"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
locale l_known_ptr\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_element_ptr ptr)"
lemma known_ptr_not_element_ptr: "\<not>is_element_ptr ptr \<Longrightarrow> a_known_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def by blast
end
diff --git a/thys/Core_SC_DOM/common/classes/CharacterDataClass.thy b/thys/Core_SC_DOM/common/classes/CharacterDataClass.thy
--- a/thys/Core_SC_DOM/common/classes/CharacterDataClass.thy
+++ b/thys/Core_SC_DOM/common/classes/CharacterDataClass.thy
@@ -1,355 +1,359 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>CharacterData\<close>
text\<open>In this theory, we introduce the types for the CharacterData class.\<close>
theory CharacterDataClass
imports
ElementClass
begin
subsubsection\<open>CharacterData\<close>
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, defined
\autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
record RCharacterData = RNode +
nothing :: unit
val :: DOMString
register_default_tvars "'CharacterData RCharacterData_ext"
type_synonym 'CharacterData CharacterData = "'CharacterData option RCharacterData_scheme"
register_default_tvars "'CharacterData CharacterData"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
'Element, 'CharacterData) Node
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
'CharacterData option RCharacterData_ext + 'Node, 'Element) Node"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
'Element, 'CharacterData) Node"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'CharacterData option RCharacterData_ext + 'Node,
'Element) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'CharacterData option RCharacterData_ext + 'Node, 'Element) heap"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition character_data_ptr_kinds :: "(_) heap \<Rightarrow> (_) character_data_ptr fset"
where
"character_data_ptr_kinds heap = the |`| (cast |`| (ffilter is_character_data_ptr_kind
(node_ptr_kinds heap)))"
lemma character_data_ptr_kinds_simp [simp]:
"character_data_ptr_kinds (Heap (fmupd (cast character_data_ptr) character_data (the_heap h)))
= {|character_data_ptr|} |\<union>| character_data_ptr_kinds h"
- apply(auto simp add: character_data_ptr_kinds_def)[1]
- by force
+ by (auto simp add: character_data_ptr_kinds_def)
definition character_data_ptrs :: "(_) heap \<Rightarrow> _ character_data_ptr fset"
where
"character_data_ptrs heap = ffilter is_character_data_ptr (character_data_ptr_kinds heap)"
abbreviation "character_data_ptr_exts heap \<equiv> character_data_ptr_kinds heap - character_data_ptrs heap"
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) Node \<Rightarrow> (_) CharacterData option"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node = (case RNode.more node of
Inr (Inl character_data) \<Rightarrow> Some (RNode.extend (RNode.truncate node) character_data)
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) Object \<Rightarrow> (_) CharacterData option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a obj \<equiv> (case cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj of Some node \<Rightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node
| None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
definition cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) CharacterData \<Rightarrow> (_) Node"
where
"cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = RNode.extend (RNode.truncate character_data)
(Inr (Inl (RNode.more character_data)))"
adhoc_overloading cast cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e
abbreviation cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) CharacterData \<Rightarrow> (_) Object"
where
"cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr)"
adhoc_overloading cast cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
consts is_character_data_kind :: 'a
definition is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Node \<Rightarrow> bool"
where
"is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr \<longleftrightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr \<noteq> None"
adhoc_overloading is_character_data_kind is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e
lemmas is_character_data_kind_def = is_character_data_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
abbreviation is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) Object \<Rightarrow> bool"
where
"is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr \<noteq> None"
adhoc_overloading is_character_data_kind is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma character_data_ptr_kinds_commutes [simp]:
"cast character_data_ptr |\<in>| node_ptr_kinds h
\<longleftrightarrow> character_data_ptr |\<in>| character_data_ptr_kinds h"
- apply(auto simp add: character_data_ptr_kinds_def)[1]
- by (metis character_data_ptr_casts_commute2 comp_eq_dest_lhs ffmember_filter fimage_eqI
- is_character_data_ptr_kind_none
- option.distinct(1) option.sel)
+proof (rule iffI)
+ show "cast character_data_ptr |\<in>| node_ptr_kinds h \<Longrightarrow>
+ character_data_ptr |\<in>| character_data_ptr_kinds h"
+ by (metis (no_types, opaque_lifting) character_data_ptr_casts_commute2
+ character_data_ptr_kinds_def ffmember_filter fimageI is_character_data_ptr_kind\<^sub>_cast option.sel)
+next
+ show "character_data_ptr |\<in>| character_data_ptr_kinds h \<Longrightarrow>
+ cast character_data_ptr |\<in>| node_ptr_kinds h"
+ by (auto simp add: character_data_ptr_kinds_def)
+qed
definition get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) character_data_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) CharacterData option"
where
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h = Option.bind (get\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast character_data_ptr) h) cast"
adhoc_overloading get get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
locale l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (ElementClass.type_wf h
\<and> (\<forall>character_data_ptr \<in> fset (character_data_ptr_kinds h).
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a: "type_wf h \<Longrightarrow> CharacterDataClass.type_wf h"
sublocale l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a \<subseteq> l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
using ElementClass.a_type_wf_def
by (meson CharacterDataClass.a_type_wf_def l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
locale l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
sublocale l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
lemma get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf:
assumes "type_wf h"
shows "character_data_ptr |\<in>| character_data_ptr_kinds h
\<longleftrightarrow> get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h \<noteq> None"
using l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms assms
apply(simp add: type_wf_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
- by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes fmember_iff_member_fset
+ by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes
local.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf option.exhaust option.simps(3))
end
global_interpretation l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf
by unfold_locales
definition put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) character_data_ptr \<Rightarrow> (_) CharacterData \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast character_data_ptr)
(cast character_data)"
adhoc_overloading put put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
lemma put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap:
assumes "put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data h = h'"
shows "character_data_ptr |\<in>| character_data_ptr_kinds h'"
using assms put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap
unfolding put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def character_data_ptr_kinds_def
by (metis character_data_ptr_kinds_commutes character_data_ptr_kinds_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap)
lemma put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_put_ptrs:
assumes "put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast character_data_ptr|}"
using assms
by (simp add: put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs)
lemma cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]: "cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e x = cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def)
by (metis (full_types) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_none [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node = None \<longleftrightarrow> \<not> (\<exists>character_data. cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = node)"
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)[1]
by (metis (full_types) RNode.select_convs(2) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_some [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a node = Some character_data \<longleftrightarrow> cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = node"
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_inv [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a (cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data) = Some character_data"
by simp
lemma cast_element_not_character_data [simp]:
"(cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element \<noteq> cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data)"
"(cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data \<noteq> cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element)"
by(auto simp add: cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RNode.extend_def)
lemma get_CharacterData_simp1 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr character_data h)
= Some character_data"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma get_CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr' \<Longrightarrow> get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr
(put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' character_data h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma get_CharacterData_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma get_CharacterData_simp4 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t character_data_ptr f h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a element_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
abbreviation "create_character_data_obj val_arg
\<equiv> \<lparr> RObject.nothing = (), RNode.nothing = (), RCharacterData.nothing = (), val = val_arg, \<dots> = None \<rparr>"
definition new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a :: "(_) heap \<Rightarrow> ((_) character_data_ptr \<times> (_) heap)"
where
"new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h =
(let new_character_data_ptr = character_data_ptr.Ref (Suc (fMax (character_data_ptr.the_ref
|`| (character_data_ptrs h)))) in
(new_character_data_ptr, put new_character_data_ptr (create_character_data_obj '''') h))"
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "new_character_data_ptr |\<in>| character_data_ptr_kinds h'"
using assms
unfolding new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def
using put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap by blast
lemma new_character_data_ptr_new:
"character_data_ptr.Ref (Suc (fMax (finsert 0 (character_data_ptr.the_ref |`| character_data_ptrs h))))
|\<notin>| character_data_ptrs h"
by (metis Suc_n_not_le_n character_data_ptr.sel(1) fMax_ge fimage_finsert finsertI1
finsertI2 set_finsert)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "new_character_data_ptr |\<notin>| character_data_ptr_kinds h"
using assms
unfolding new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
by (metis Pair_inject character_data_ptrs_def fMax_finsert fempty_iff ffmember_filter
fimage_is_fempty is_character_data_ptr_ref max_0L new_character_data_ptr_new)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_new_ptr:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using assms
by (metis Pair_inject new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_put_ptrs)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_is_character_data_ptr:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "is_character_data_ptr new_character_data_ptr"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
assumes "ptr \<noteq> cast new_character_data_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
assumes "ptr \<noteq> cast new_character_data_ptr"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
assumes "ptr \<noteq> new_character_data_ptr"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
locale l_known_ptr\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_character_data_ptr ptr)"
lemma known_ptr_not_character_data_ptr:
"\<not>is_character_data_ptr ptr \<Longrightarrow> a_known_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def
by blast
end
diff --git a/thys/Core_SC_DOM/common/classes/DocumentClass.thy b/thys/Core_SC_DOM/common/classes/DocumentClass.thy
--- a/thys/Core_SC_DOM/common/classes/DocumentClass.thy
+++ b/thys/Core_SC_DOM/common/classes/DocumentClass.thy
@@ -1,345 +1,348 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Document\<close>
text\<open>In this theory, we introduce the types for the Document class.\<close>
theory DocumentClass
imports
CharacterDataClass
begin
text\<open>The type @{type "doctype"} is a type synonym for @{type "string"}, defined
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
record ('node_ptr, 'element_ptr, 'character_data_ptr) RDocument = RObject +
nothing :: unit
doctype :: doctype
document_element :: "(_) element_ptr option"
disconnected_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document) Document
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option) RDocument_scheme"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'Document) Document"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData, 'Document) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option)
RDocument_ext + 'Object, 'Node, 'Element, 'CharacterData) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'Document) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr,
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option) RDocument_ext + 'Object, 'Node,
'Element, 'CharacterData) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition document_ptr_kinds :: "(_) heap \<Rightarrow> (_) document_ptr fset"
where
"document_ptr_kinds heap = 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 |`|
(ffilter is_document_ptr_kind (object_ptr_kinds heap)))"
definition document_ptrs :: "(_) heap \<Rightarrow> (_) document_ptr fset"
where
"document_ptrs heap = ffilter is_document_ptr (document_ptr_kinds heap)"
definition cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Object \<Rightarrow> (_) Document option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = (case RObject.more obj of
Inr (Inl document) \<Rightarrow> Some (RObject.extend (RObject.truncate obj) document)
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
definition cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:: "(_) Document \<Rightarrow> (_) Object"
where
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document = (RObject.extend (RObject.truncate document)
(Inr (Inl (RObject.more document))))"
adhoc_overloading cast cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
definition is_document_kind :: "(_) Object \<Rightarrow> bool"
where
"is_document_kind ptr \<longleftrightarrow> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
lemma document_ptr_kinds_simp [simp]:
"document_ptr_kinds (Heap (fmupd (cast document_ptr) document (the_heap h)))
= {|document_ptr|} |\<union>| document_ptr_kinds h"
- apply(auto simp add: document_ptr_kinds_def)[1]
- by force
+ by (auto simp add: document_ptr_kinds_def)
lemma document_ptr_kinds_commutes [simp]:
"cast document_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> document_ptr |\<in>| document_ptr_kinds h"
- apply(auto simp add: object_ptr_kinds_def document_ptr_kinds_def)[1]
- by (metis (no_types, lifting) document_ptr_casts_commute2 document_ptr_document_ptr_cast
- ffmember_filter fimage_eqI fset.map_comp option.sel)
+proof (rule iffI)
+ show "cast document_ptr |\<in>| object_ptr_kinds h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h"
+ by (metis (no_types, opaque_lifting) document_ptr_casts_commute2 document_ptr_document_ptr_cast
+ document_ptr_kinds_def ffmember_filter fimage_eqI option.sel)
+next
+ show "document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow> cast document_ptr |\<in>| object_ptr_kinds h"
+ by (auto simp: document_ptr_kinds_def)
+qed
definition get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) document_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Document option"
where
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h = Option.bind (get (cast document_ptr) h) cast"
adhoc_overloading get get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_type_wf_def\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (CharacterDataClass.type_wf h \<and>
(\<forall>document_ptr \<in> fset (document_ptr_kinds h). get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t: "type_wf h \<Longrightarrow> DocumentClass.type_wf h"
sublocale l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t \<subseteq> l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
apply(unfold_locales)
by (metis (full_types) type_wf_defs l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
locale l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
lemma get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf:
assumes "type_wf h"
shows "document_ptr |\<in>| document_ptr_kinds h \<longleftrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None"
using l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- by (metis document_ptr_kinds_commutes fmember_iff_member_fset is_none_bind is_none_simps(1)
+ by (metis document_ptr_kinds_commutes is_none_bind is_none_simps(1)
is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
end
global_interpretation l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
definition put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) document_ptr \<Rightarrow> (_) Document \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document = put (cast document_ptr) (cast document)"
adhoc_overloading put put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
lemma put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h = h'"
shows "document_ptr |\<in>| document_ptr_kinds h'"
using assms
unfolding put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis document_ptr_kinds_commutes put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap)
lemma put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs:
assumes "put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast document_ptr|}"
using assms
by (simp add: put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs)
lemma cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_inject [simp]: "cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x = cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
by (metis (full_types) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = None \<longleftrightarrow> \<not> (\<exists>document. cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document = obj)"
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
split: sum.splits)[1]
by (metis (full_types) RObject.select_convs(2) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = Some document \<longleftrightarrow> cast document = obj"
by(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
split: sum.splits)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document) = Some document"
by simp
lemma cast_document_not_node [simp]:
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document \<noteq> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node"
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node \<noteq> cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document"
by(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
lemma get_document_ptr_simp1 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h) = Some document"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp2 [simp]:
"document_ptr \<noteq> document_ptr'
\<Longrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' document h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp4 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_document_ptr_simp5 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp6 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
abbreviation
create_document_obj :: "char list \<Rightarrow> (_) element_ptr option \<Rightarrow> (_) node_ptr list \<Rightarrow> (_) Document"
where
"create_document_obj doctype_arg document_element_arg disconnected_nodes_arg
\<equiv> \<lparr> RObject.nothing = (), RDocument.nothing = (), doctype = doctype_arg,
document_element = document_element_arg,
disconnected_nodes = disconnected_nodes_arg, \<dots> = None \<rparr>"
definition new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_)heap \<Rightarrow> ((_) document_ptr \<times> (_) heap)"
where
"new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_document_ptr = document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| (document_ptrs h)))))
in
(new_document_ptr, put new_document_ptr (create_document_obj '''' None []) h))"
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "new_document_ptr |\<in>| document_ptr_kinds h'"
using assms
unfolding new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
using put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap by blast
lemma new_document_ptr_new:
"document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| document_ptrs h))))
|\<notin>| document_ptrs h"
by (metis Suc_n_not_le_n document_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "new_document_ptr |\<notin>| document_ptr_kinds h"
using assms
unfolding new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis Pair_inject document_ptrs_def fMax_finsert fempty_iff ffmember_filter
fimage_is_fempty is_document_ptr_ref max_0L new_document_ptr_new)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_document_ptr|}"
using assms
by (metis Pair_inject new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_document_ptr:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "is_document_ptr new_document_ptr"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
assumes "ptr \<noteq> cast new_document_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
apply(simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
assumes "ptr \<noteq> new_document_ptr"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
locale l_known_ptr\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_document_ptr ptr)"
lemma known_ptr_not_document_ptr: "\<not>is_document_ptr ptr \<Longrightarrow> a_known_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs [instances]: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
end
diff --git a/thys/Core_SC_DOM/common/classes/NodeClass.thy b/thys/Core_SC_DOM/common/classes/NodeClass.thy
--- a/thys/Core_SC_DOM/common/classes/NodeClass.thy
+++ b/thys/Core_SC_DOM/common/classes/NodeClass.thy
@@ -1,209 +1,212 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Node\<close>
text\<open>In this theory, we introduce the types for the Node class.\<close>
theory NodeClass
imports
ObjectClass
"../pointers/NodePointer"
begin
subsubsection\<open>Node\<close>
record RNode = RObject
+ nothing :: unit
register_default_tvars "'Node RNode_ext"
type_synonym 'Node Node = "'Node RNode_scheme"
register_default_tvars "'Node Node"
type_synonym ('Object, 'Node) Object = "('Node RNode_ext + 'Object) Object"
register_default_tvars "('Object, 'Node) Object"
type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node) heap
= "('node_ptr node_ptr + 'object_ptr, 'Node RNode_ext + 'Object) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'Object, 'Node) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit) heap"
definition node_ptr_kinds :: "(_) heap \<Rightarrow> (_) node_ptr fset"
where
"node_ptr_kinds heap =
(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 |`| (ffilter is_node_ptr_kind (object_ptr_kinds heap))))"
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (Heap (fmupd (cast node_ptr) node (the_heap h)))
= {|node_ptr|} |\<union>| node_ptr_kinds h"
- apply(auto simp add: node_ptr_kinds_def)[1]
- by force
+ by (auto simp add: node_ptr_kinds_def)
definition cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Object \<Rightarrow> (_) Node option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = (case RObject.more obj of Inl node
\<Rightarrow> Some (RObject.extend (RObject.truncate obj) node) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:: "(_) Node \<Rightarrow> (_) Object"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node = (RObject.extend (RObject.truncate node) (Inl (RObject.more node)))"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
definition is_node_kind :: "(_) Object \<Rightarrow> bool"
where
"is_node_kind ptr \<longleftrightarrow> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr \<noteq> None"
definition get\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) node_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Node option"
where
"get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h = Option.bind (get (cast node_ptr) h) cast"
adhoc_overloading get get\<^sub>N\<^sub>o\<^sub>d\<^sub>e
locale l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (ObjectClass.type_wf h
\<and> (\<forall>node_ptr \<in> fset( node_ptr_kinds h). get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e: "type_wf h \<Longrightarrow> NodeClass.type_wf h"
sublocale l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e \<subseteq> l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
apply(unfold_locales)
using ObjectClass.a_type_wf_def by auto
locale l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
sublocale l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales
lemma get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf:
assumes "type_wf h"
shows "node_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None"
using l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_axioms assms
apply(simp add: type_wf_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
- by (metis bind_eq_None_conv ffmember_filter fimage_eqI fmember_iff_member_fset is_node_ptr_kind_cast
+ by (metis bind_eq_None_conv ffmember_filter fimage_eqI is_node_ptr_kind_cast
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf node_ptr_casts_commute2 node_ptr_kinds_def option.sel option.simps(3))
end
global_interpretation l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf
by unfold_locales
definition put\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) node_ptr \<Rightarrow> (_) Node \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node = put (cast node_ptr) (cast node)"
adhoc_overloading put put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
lemma put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap:
assumes "put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h = h'"
shows "node_ptr |\<in>| node_ptr_kinds h'"
using assms
unfolding put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def node_ptr_kinds_def
by (metis ffmember_filter fimage_eqI is_node_ptr_kind_cast node_ptr_casts_commute2
option.sel put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap)
lemma put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs:
assumes "put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast node_ptr|}"
using assms
by (simp add: put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs)
lemma node_ptr_kinds_commutes [simp]:
"cast node_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> node_ptr |\<in>| node_ptr_kinds h"
- apply(auto simp add: node_ptr_kinds_def split: option.splits)[1]
- by (metis (no_types, lifting) ffmember_filter fimage_eqI fset.map_comp
- is_node_ptr_kind_none node_ptr_casts_commute2
- option.distinct(1) option.sel)
+proof (rule iffI)
+ show "cast node_ptr |\<in>| object_ptr_kinds h \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h"
+ by (metis ffmember_filter fimage.rep_eq image_eqI is_node_ptr_kind_cast node_ptr_casts_commute2
+ node_ptr_kinds_def option.sel)
+next
+ show "node_ptr |\<in>| node_ptr_kinds h \<Longrightarrow> cast node_ptr |\<in>| object_ptr_kinds h"
+ by (auto simp: node_ptr_kinds_def)
+qed
lemma node_empty [simp]:
"\<lparr>RObject.nothing = (), RNode.nothing = (), \<dots> = RNode.more node\<rparr> = node"
by simp
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_inject [simp]: "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x = cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
by (metis (full_types) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = None \<longleftrightarrow> \<not> (\<exists>node. cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node = obj)"
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def split: sum.splits)[1]
by (metis (full_types) RObject.select_convs(2) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_some [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = Some node \<longleftrightarrow> cast node = obj"
by(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def split: sum.splits)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node) = Some node"
by simp
locale l_known_ptr\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = False"
end
global_interpretation l_known_ptr\<^sub>N\<^sub>o\<^sub>d\<^sub>e defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
+
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset
known_ptrs_new_ptr
by blast
lemma get_node_ptr_simp1 [simp]: "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h) = Some node"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_node_ptr_simp2 [simp]:
"node_ptr \<noteq> node_ptr' \<Longrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr' node h) = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
end
diff --git a/thys/Core_SC_DOM/common/classes/ObjectClass.thy b/thys/Core_SC_DOM/common/classes/ObjectClass.thy
--- a/thys/Core_SC_DOM/common/classes/ObjectClass.thy
+++ b/thys/Core_SC_DOM/common/classes/ObjectClass.thy
@@ -1,225 +1,224 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Object\<close>
text\<open>In this theory, we introduce the definition of the class Object. This class is the
common superclass of our class model.\<close>
theory ObjectClass
imports
BaseClass
"../pointers/ObjectPointer"
begin
record RObject =
nothing :: unit
register_default_tvars "'Object RObject_ext"
type_synonym 'Object Object = "'Object RObject_scheme"
register_default_tvars "'Object Object"
datatype ('object_ptr, 'Object) heap = Heap (the_heap: "((_) object_ptr, (_) Object) fmap")
register_default_tvars "('object_ptr, 'Object) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit) heap"
definition object_ptr_kinds :: "(_) heap \<Rightarrow> (_) object_ptr fset"
where
"object_ptr_kinds = fmdom \<circ> the_heap"
lemma object_ptr_kinds_simp [simp]:
"object_ptr_kinds (Heap (fmupd object_ptr object (the_heap h)))
= {|object_ptr|} |\<union>| object_ptr_kinds h"
by(auto simp add: object_ptr_kinds_def)
definition get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Object option"
where
"get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = fmlookup (the_heap h) ptr"
adhoc_overloading get get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
locale l_type_wf_def\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = True"
end
global_interpretation l_type_wf_def\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t: "type_wf h \<Longrightarrow> ObjectClass.type_wf h"
locale l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas = l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
lemma get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf:
assumes "type_wf h"
shows "object_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h \<noteq> None"
using l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms assms
apply(simp add: type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
by (simp add: fmlookup_dom_iff object_ptr_kinds_def)
end
global_interpretation l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas type_wf
by (simp add: l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas.intro l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t.intro)
definition put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) Object \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h = Heap (fmupd ptr obj (the_heap h))"
adhoc_overloading put put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap:
assumes "put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h = h'"
shows "object_ptr |\<in>| object_ptr_kinds h'"
using assms
unfolding put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
by auto
lemma put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs:
assumes "put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|object_ptr|}"
using assms
by (metis comp_apply fmdom_fmupd funion_finsert_right heap.sel object_ptr_kinds_def
sup_bot.right_neutral put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemma object_more_extend_id [simp]: "more (extend x y) = y"
by(simp add: extend_def)
lemma object_empty [simp]: "\<lparr>nothing = (), \<dots> = more x\<rparr> = x"
by simp
locale l_known_ptr\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = False"
lemma known_ptr_not_object_ptr:
"a_known_ptr ptr \<Longrightarrow> \<not>is_object_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool" +
fixes known_ptrs :: "(_) heap \<Rightarrow> bool"
assumes known_ptrs_known_ptr: "known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
assumes known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h = known_ptrs h'"
assumes known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'"
assumes known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
known_ptrs h \<Longrightarrow> known_ptrs h'"
locale l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
"a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
lemma get_object_ptr_simp1 [simp]: "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h) = Some object"
by(simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemma get_object_ptr_simp2 [simp]:
"object_ptr \<noteq> object_ptr'
\<Longrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr' object h) = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h"
by(simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
subsection\<open>Limited Heap Modifications\<close>
definition heap_unchanged_except :: "(_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"heap_unchanged_except S h h' = (\<forall>ptr \<in> (fset (object_ptr_kinds h)
\<union> (fset (object_ptr_kinds h'))) - S. get ptr h = get ptr h')"
definition delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) heap option" where
"delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = (if ptr |\<in>| object_ptr_kinds h then Some (Heap (fmdrop ptr (the_heap h)))
else None)"
lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_removed:
assumes "delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = Some h'"
shows "ptr |\<notin>| object_ptr_kinds h'"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def object_ptr_kinds_def split: if_splits)
lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_ptr_in_heap:
assumes "delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = Some h'"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def object_ptr_kinds_def split: if_splits)
lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok:
assumes "ptr |\<in>| object_ptr_kinds h"
shows "delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h \<noteq> None"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def object_ptr_kinds_def split: if_splits)
subsection \<open>Code Generator Setup\<close>
definition "create_heap xs = Heap (fmap_of_list xs)"
code_datatype ObjectClass.heap.Heap create_heap
lemma object_ptr_kinds_code3 [code]:
"fmlookup (the_heap (create_heap xs)) x = map_of xs x"
by(auto simp add: create_heap_def fmlookup_of_list)
lemma object_ptr_kinds_code4 [code]:
"the_heap (create_heap xs) = fmap_of_list xs"
by(simp add: create_heap_def)
lemma object_ptr_kinds_code5 [code]:
"the_heap (Heap x) = x"
by simp
end
diff --git a/thys/Core_SC_DOM/common/monads/BaseMonad.thy b/thys/Core_SC_DOM/common/monads/BaseMonad.thy
--- a/thys/Core_SC_DOM/common/monads/BaseMonad.thy
+++ b/thys/Core_SC_DOM/common/monads/BaseMonad.thy
@@ -1,384 +1,381 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>The Monad Infrastructure\<close>
text\<open>In this theory, we introduce the basic infrastructure for our monadic class encoding.\<close>
theory BaseMonad
imports
"../classes/BaseClass"
"../preliminaries/Heap_Error_Monad"
begin
subsection\<open>Datatypes\<close>
datatype exception = NotFoundError | HierarchyRequestError | NotSupportedError | SegmentationFault
| AssertException | NonTerminationException | InvokeError | TypeError
-lemma finite_set_in [simp]: "x \<in> fset FS \<longleftrightarrow> x |\<in>| FS"
- by (meson notin_fset)
-
consts put_M :: 'a
consts get_M :: 'a
consts delete_M :: 'a
lemma sorted_list_of_set_eq [dest]:
"sorted_list_of_set (fset x) = sorted_list_of_set (fset y) \<Longrightarrow> x = y"
by (metis finite_fset fset_inject sorted_list_of_set(1))
locale l_ptr_kinds_M =
fixes ptr_kinds :: "'heap \<Rightarrow> 'ptr::linorder fset"
begin
definition a_ptr_kinds_M :: "('heap, exception, 'ptr list) prog"
where
"a_ptr_kinds_M = do {
h \<leftarrow> get_heap;
return (sorted_list_of_set (fset (ptr_kinds h)))
}"
lemma ptr_kinds_M_ok [simp]: "h \<turnstile> ok a_ptr_kinds_M"
by(simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_pure [simp]: "pure a_ptr_kinds_M h"
by (auto simp add: a_ptr_kinds_M_def intro: bind_pure_I)
lemma ptr_kinds_ptr_kinds_M [simp]: "ptr \<in> set |h \<turnstile> a_ptr_kinds_M|\<^sub>r \<longleftrightarrow> ptr |\<in>| ptr_kinds h"
by(simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds [simp]:
"h \<turnstile> a_ptr_kinds_M \<rightarrow>\<^sub>r xa \<longleftrightarrow> xa = sorted_list_of_set (fset (ptr_kinds h))"
by(auto simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds_returns_result [simp]:
"h \<turnstile> a_ptr_kinds_M \<bind> f \<rightarrow>\<^sub>r x \<longleftrightarrow> h \<turnstile> f (sorted_list_of_set (fset (ptr_kinds h))) \<rightarrow>\<^sub>r x"
by(auto simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds_returns_heap [simp]:
"h \<turnstile> a_ptr_kinds_M \<bind> f \<rightarrow>\<^sub>h h' \<longleftrightarrow> h \<turnstile> f (sorted_list_of_set (fset (ptr_kinds h))) \<rightarrow>\<^sub>h h'"
by(auto simp add: a_ptr_kinds_M_def)
end
locale l_get_M =
fixes get :: "'ptr \<Rightarrow> 'heap \<Rightarrow> 'obj option"
fixes type_wf :: "'heap \<Rightarrow> bool"
fixes ptr_kinds :: "'heap \<Rightarrow> 'ptr fset"
assumes "type_wf h \<Longrightarrow> ptr |\<in>| ptr_kinds h \<Longrightarrow> get ptr h \<noteq> None"
assumes "get ptr h \<noteq> None \<Longrightarrow> ptr |\<in>| ptr_kinds h"
begin
definition a_get_M :: "'ptr \<Rightarrow> ('obj \<Rightarrow> 'result) \<Rightarrow> ('heap, exception, 'result) prog"
where
"a_get_M ptr getter = (do {
h \<leftarrow> get_heap;
(case get ptr h of
Some res \<Rightarrow> return (getter res)
| None \<Rightarrow> error SegmentationFault)
})"
lemma get_M_pure [simp]: "pure (a_get_M ptr getter) h"
by(auto simp add: a_get_M_def bind_pure_I split: option.splits)
lemma get_M_ok:
"type_wf h \<Longrightarrow> ptr |\<in>| ptr_kinds h \<Longrightarrow> h \<turnstile> ok (a_get_M ptr getter)"
apply(simp add: a_get_M_def)
by (metis l_get_M_axioms l_get_M_def option.case_eq_if return_ok)
lemma get_M_ptr_in_heap:
"h \<turnstile> ok (a_get_M ptr getter) \<Longrightarrow> ptr |\<in>| ptr_kinds h"
apply(simp add: a_get_M_def)
by (metis error_returns_result is_OK_returns_result_E l_get_M_axioms l_get_M_def option.simps(4))
end
locale l_put_M = l_get_M get for get :: "'ptr \<Rightarrow> 'heap \<Rightarrow> 'obj option" +
fixes put :: "'ptr \<Rightarrow> 'obj \<Rightarrow> 'heap \<Rightarrow> 'heap"
begin
definition a_put_M :: "'ptr \<Rightarrow> (('v \<Rightarrow> 'v) \<Rightarrow> 'obj \<Rightarrow> 'obj) \<Rightarrow> 'v \<Rightarrow> ('heap, exception, unit) prog"
where
"a_put_M ptr setter v = (do {
obj \<leftarrow> a_get_M ptr id;
h \<leftarrow> get_heap;
return_heap (put ptr (setter (\<lambda>_. v) obj) h)
})"
lemma put_M_ok:
"type_wf h \<Longrightarrow> ptr |\<in>| ptr_kinds h \<Longrightarrow> h \<turnstile> ok (a_put_M ptr setter v)"
by(auto simp add: a_put_M_def intro!: bind_is_OK_I2 dest: get_M_ok elim!: bind_is_OK_E)
lemma put_M_ptr_in_heap:
"h \<turnstile> ok (a_put_M ptr setter v) \<Longrightarrow> ptr |\<in>| ptr_kinds h"
by(auto simp add: a_put_M_def intro!: bind_is_OK_I2 elim: get_M_ptr_in_heap
dest: is_OK_returns_result_I elim!: bind_is_OK_E)
end
subsection \<open>Setup for Defining Partial Functions\<close>
lemma execute_admissible:
"ccpo.admissible (fun_lub (flat_lub (Inl (e::'e)))) (fun_ord (flat_ord (Inl e)))
((\<lambda>a. \<forall>(h::'heap) h2 (r::'result). h \<turnstile> a = Inr (r, h2) \<longrightarrow> P h h2 r) \<circ> Prog)"
proof (unfold comp_def, rule ccpo.admissibleI, clarify)
fix A :: "('heap \<Rightarrow> 'e + 'result \<times> 'heap) set"
let ?lub = "Prog (fun_lub (flat_lub (Inl e)) A)"
fix h h2 r
assume 1: "Complete_Partial_Order.chain (fun_ord (flat_ord (Inl e))) A"
and 2: "\<forall>xa\<in>A. \<forall>h h2 r. h \<turnstile> Prog xa = Inr (r, h2) \<longrightarrow> P h h2 r"
and 4: "h \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r, h2)"
have h1:"\<And>a. Complete_Partial_Order.chain (flat_ord (Inl e)) {y. \<exists>f\<in>A. y = f a}"
by (rule chain_fun[OF 1])
show "P h h2 r"
using CollectD Inl_Inr_False prog.sel chain_fun[OF 1] flat_lub_in_chain[OF chain_fun[OF 1]] 2 4
unfolding execute_def fun_lub_def
proof -
assume a1: "the_prog (Prog (\<lambda>x. flat_lub (Inl e) {y. \<exists>f\<in>A. y = f x})) h = Inr (r, h2)"
assume a2: "\<forall>xa\<in>A. \<forall>h h2 r. the_prog (Prog xa) h = Inr (r, h2) \<longrightarrow> P h h2 r"
have "Inr (r, h2) \<in> {s. \<exists>f. f \<in> A \<and> s = f h} \<or> Inr (r, h2) = Inl e"
using a1 by (metis (lifting) \<open>\<And>aa a. flat_lub (Inl e) {y. \<exists>f\<in>A. y = f aa} = a \<Longrightarrow> a = Inl e \<or> a \<in> {y. \<exists>f\<in>A. y = f aa}\<close> prog.sel)
then show ?thesis
using a2 by fastforce
qed
qed
lemma execute_admissible2:
"ccpo.admissible (fun_lub (flat_lub (Inl (e::'e)))) (fun_ord (flat_ord (Inl e)))
((\<lambda>a. \<forall>(h::'heap) h' h2 h2' (r::'result) r'.
h \<turnstile> a = Inr (r, h2) \<longrightarrow> h' \<turnstile> a = Inr (r', h2') \<longrightarrow> P h h' h2 h2' r r') \<circ> Prog)"
proof (unfold comp_def, rule ccpo.admissibleI, clarify)
fix A :: "('heap \<Rightarrow> 'e + 'result \<times> 'heap) set"
let ?lub = "Prog (fun_lub (flat_lub (Inl e)) A)"
fix h h' h2 h2' r r'
assume 1: "Complete_Partial_Order.chain (fun_ord (flat_ord (Inl e))) A"
and 2 [rule_format]: "\<forall>xa\<in>A. \<forall>h h' h2 h2' r r'. h \<turnstile> Prog xa = Inr (r, h2)
\<longrightarrow> h' \<turnstile> Prog xa = Inr (r', h2') \<longrightarrow> P h h' h2 h2' r r'"
and 4: "h \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r, h2)"
and 5: "h' \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r', h2')"
have h1:"\<And>a. Complete_Partial_Order.chain (flat_ord (Inl e)) {y. \<exists>f\<in>A. y = f a}"
by (rule chain_fun[OF 1])
have "h \<turnstile> ?lub \<in> {y. \<exists>f\<in>A. y = f h}"
using flat_lub_in_chain[OF h1] 4
unfolding execute_def fun_lub_def
by (metis (mono_tags, lifting) Collect_cong Inl_Inr_False prog.sel)
moreover have "h' \<turnstile> ?lub \<in> {y. \<exists>f\<in>A. y = f h'}"
using flat_lub_in_chain[OF h1] 5
unfolding execute_def fun_lub_def
by (metis (no_types, lifting) Collect_cong Inl_Inr_False prog.sel)
ultimately obtain f where
"f \<in> A" and
"h \<turnstile> Prog f = Inr (r, h2)" and
"h' \<turnstile> Prog f = Inr (r', h2')"
using 1 4 5
apply(auto simp add: chain_def fun_ord_def flat_ord_def execute_def)[1]
by (metis Inl_Inr_False)
then show "P h h' h2 h2' r r'"
by(fact 2)
qed
definition dom_prog_ord ::
"('heap, exception, 'result) prog \<Rightarrow> ('heap, exception, 'result) prog \<Rightarrow> bool" where
"dom_prog_ord = img_ord (\<lambda>a b. execute b a) (fun_ord (flat_ord (Inl NonTerminationException)))"
definition dom_prog_lub ::
"('heap, exception, 'result) prog set \<Rightarrow> ('heap, exception, 'result) prog" where
"dom_prog_lub = img_lub (\<lambda>a b. execute b a) Prog (fun_lub (flat_lub (Inl NonTerminationException)))"
lemma dom_prog_lub_empty: "dom_prog_lub {} = error NonTerminationException"
by(simp add: dom_prog_lub_def img_lub_def fun_lub_def flat_lub_def error_def)
lemma dom_prog_interpretation: "partial_function_definitions dom_prog_ord dom_prog_lub"
proof -
have "partial_function_definitions (fun_ord (flat_ord (Inl NonTerminationException)))
(fun_lub (flat_lub (Inl NonTerminationException)))"
by (rule partial_function_lift) (rule flat_interpretation)
then show ?thesis
apply (simp add: dom_prog_lub_def dom_prog_ord_def flat_interpretation execute_def)
using partial_function_image prog.expand prog.sel by blast
qed
interpretation dom_prog: partial_function_definitions dom_prog_ord dom_prog_lub
rewrites "dom_prog_lub {} \<equiv> error NonTerminationException"
by (fact dom_prog_interpretation)(simp add: dom_prog_lub_empty)
lemma admissible_dom_prog:
"dom_prog.admissible (\<lambda>f. \<forall>x h h' r. h \<turnstile> f x \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h' \<longrightarrow> P x h h' r)"
proof (rule admissible_fun[OF dom_prog_interpretation])
fix x
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h' r. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h'
\<longrightarrow> P x h h' r)"
unfolding dom_prog_ord_def dom_prog_lub_def
proof (intro admissible_image partial_function_lift flat_interpretation)
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
(fun_ord (flat_ord (Inl NonTerminationException)))
((\<lambda>a. \<forall>h h' r. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> P x h h' r) \<circ> Prog)"
by(auto simp add: execute_admissible returns_result_def returns_heap_def split: sum.splits)
next
show "\<And>x y. (\<lambda>b. b \<turnstile> x) = (\<lambda>b. b \<turnstile> y) \<Longrightarrow> x = y"
by(simp add: execute_def prog.expand)
next
show "\<And>x. (\<lambda>b. b \<turnstile> Prog x) = x"
by(simp add: execute_def)
qed
qed
lemma admissible_dom_prog2:
"dom_prog.admissible (\<lambda>f. \<forall>x h h2 h' h2' r r2. h \<turnstile> f x \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h'
\<longrightarrow> h2 \<turnstile> f x \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> f x \<rightarrow>\<^sub>h h2' \<longrightarrow> P x h h2 h' h2' r r2)"
proof (rule admissible_fun[OF dom_prog_interpretation])
fix x
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r
\<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>h h2' \<longrightarrow> P x h h2 h' h2' r r2)"
unfolding dom_prog_ord_def dom_prog_lub_def
proof (intro admissible_image partial_function_lift flat_interpretation)
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
(fun_ord (flat_ord (Inl NonTerminationException)))
((\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>h h2'
\<longrightarrow> P x h h2 h' h2' r r2) \<circ> Prog)"
by(auto simp add: returns_result_def returns_heap_def intro!: ccpo.admissibleI
dest!: ccpo.admissibleD[OF execute_admissible2[where P="P x"]]
split: sum.splits)
next
show "\<And>x y. (\<lambda>b. b \<turnstile> x) = (\<lambda>b. b \<turnstile> y) \<Longrightarrow> x = y"
by(simp add: execute_def prog.expand)
next
show "\<And>x. (\<lambda>b. b \<turnstile> Prog x) = x"
by(simp add: execute_def)
qed
qed
lemma fixp_induct_dom_prog:
fixes F :: "'c \<Rightarrow> 'c" and
U :: "'c \<Rightarrow> 'b \<Rightarrow> ('heap, exception, 'result) prog" and
C :: "('b \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> 'c" and
P :: "'b \<Rightarrow> 'heap \<Rightarrow> 'heap \<Rightarrow> 'result \<Rightarrow> bool"
assumes mono: "\<And>x. monotone (fun_ord dom_prog_ord) dom_prog_ord (\<lambda>f. U (F (C f)) x)"
assumes eq: "f \<equiv> C (ccpo.fixp (fun_lub dom_prog_lub) (fun_ord dom_prog_ord) (\<lambda>f. U (F (C f))))"
assumes inverse2: "\<And>f. U (C f) = f"
assumes step: "\<And>f x h h' r. (\<And>x h h' r. h \<turnstile> (U f x) \<rightarrow>\<^sub>r r \<Longrightarrow> h \<turnstile> (U f x) \<rightarrow>\<^sub>h h' \<Longrightarrow> P x h h' r)
\<Longrightarrow> h \<turnstile> (U (F f) x) \<rightarrow>\<^sub>r r \<Longrightarrow> h \<turnstile> (U (F f) x) \<rightarrow>\<^sub>h h' \<Longrightarrow> P x h h' r"
assumes defined: "h \<turnstile> (U f x) \<rightarrow>\<^sub>r r" and "h \<turnstile> (U f x) \<rightarrow>\<^sub>h h'"
shows "P x h h' r"
using step defined dom_prog.fixp_induct_uc[of U F C, OF mono eq inverse2 admissible_dom_prog, of P]
by (metis assms(6) error_returns_heap)
declaration \<open>Partial_Function.init "dom_prog" @{term dom_prog.fixp_fun}
@{term dom_prog.mono_body} @{thm dom_prog.fixp_rule_uc} @{thm dom_prog.fixp_induct_uc}
(SOME @{thm fixp_induct_dom_prog})\<close>
abbreviation "mono_dom_prog \<equiv> monotone (fun_ord dom_prog_ord) dom_prog_ord"
lemma dom_prog_ordI:
assumes "\<And>h. h \<turnstile> f \<rightarrow>\<^sub>e NonTerminationException \<or> h \<turnstile> f = h \<turnstile> g"
shows "dom_prog_ord f g"
proof(auto simp add: dom_prog_ord_def img_ord_def fun_ord_def flat_ord_def)[1]
fix x
assume "x \<turnstile> f \<noteq> x \<turnstile> g"
then show "x \<turnstile> f = Inl NonTerminationException"
using assms[where h=x]
by(auto simp add: returns_error_def split: sum.splits)
qed
lemma dom_prog_ordE:
assumes "dom_prog_ord x y"
obtains "h \<turnstile> x \<rightarrow>\<^sub>e NonTerminationException" | " h \<turnstile> x = h \<turnstile> y"
using assms unfolding dom_prog_ord_def img_ord_def fun_ord_def flat_ord_def
using returns_error_def by force
lemma bind_mono [partial_function_mono]:
fixes B :: "('a \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> ('heap, exception, 'result2) prog"
assumes mf: "mono_dom_prog B" and mg: "\<And>y. mono_dom_prog (\<lambda>f. C y f)"
shows "mono_dom_prog (\<lambda>f. B f \<bind> (\<lambda>y. C y f))"
proof (rule monotoneI)
fix f g :: "'a \<Rightarrow> ('heap, exception, 'result) prog"
assume fg: "dom_prog.le_fun f g"
from mf
have 1: "dom_prog_ord (B f) (B g)" by (rule monotoneD) (rule fg)
from mg
have 2: "\<And>y'. dom_prog_ord (C y' f) (C y' g)" by (rule monotoneD) (rule fg)
have "dom_prog_ord (B f \<bind> (\<lambda>y. C y f)) (B g \<bind> (\<lambda>y. C y f))"
(is "dom_prog_ord ?L ?R")
proof (rule dom_prog_ordI)
fix h
from 1 show "h \<turnstile> ?L \<rightarrow>\<^sub>e NonTerminationException \<or> h \<turnstile> ?L = h \<turnstile> ?R"
apply(rule dom_prog_ordE)
apply(auto)[1]
using bind_cong by fastforce
qed
also
have h1: "dom_prog_ord (B g \<bind> (\<lambda>y'. C y' f)) (B g \<bind> (\<lambda>y'. C y' g))"
(is "dom_prog_ord ?L ?R")
proof (rule dom_prog_ordI)
(* { *)
fix h
show "h \<turnstile> ?L \<rightarrow>\<^sub>e NonTerminationException \<or> h \<turnstile> ?L = h \<turnstile> ?R"
proof (cases "h \<turnstile> ok (B g)")
case True
then obtain x h' where x: "h \<turnstile> B g \<rightarrow>\<^sub>r x" and h': "h \<turnstile> B g \<rightarrow>\<^sub>h h'"
by blast
then have "dom_prog_ord (C x f) (C x g)"
using 2 by simp
then show ?thesis
using x h'
apply(auto intro!: bind_returns_error_I3 dest: returns_result_eq dest!: dom_prog_ordE)[1]
apply(auto simp add: execute_bind_simp)[1]
using "2" dom_prog_ordE by metis
next
case False
then obtain e where e: "h \<turnstile> B g \<rightarrow>\<^sub>e e"
by(simp add: is_OK_def returns_error_def split: sum.splits)
have "h \<turnstile> B g \<bind> (\<lambda>y'. C y' f) \<rightarrow>\<^sub>e e"
using e by(auto)
moreover have "h \<turnstile> B g \<bind> (\<lambda>y'. C y' g) \<rightarrow>\<^sub>e e"
using e by auto
ultimately show ?thesis
using bind_returns_error_eq by metis
qed
qed
finally (dom_prog.leq_trans)
show "dom_prog_ord (B f \<bind> (\<lambda>y. C y f)) (B g \<bind> (\<lambda>y'. C y' g))" .
qed
lemma mono_dom_prog1 [partial_function_mono]:
fixes g :: "('a \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> 'b \<Rightarrow> ('heap, exception, 'result) prog"
assumes "\<And>x. (mono_dom_prog (\<lambda>f. g f x))"
shows "mono_dom_prog (\<lambda>f. map_M (g f) xs)"
using assms
apply (induct xs)
by(auto simp add: call_mono dom_prog.const_mono intro!: bind_mono)
lemma mono_dom_prog2 [partial_function_mono]:
fixes g :: "('a \<Rightarrow> ('heap, exception, 'result) prog) \<Rightarrow> 'b \<Rightarrow> ('heap, exception, 'result) prog"
assumes "\<And>x. (mono_dom_prog (\<lambda>f. g f x))"
shows "mono_dom_prog (\<lambda>f. forall_M (g f) xs)"
using assms
apply (induct xs)
by(auto simp add: call_mono dom_prog.const_mono intro!: bind_mono)
lemma sorted_list_set_cong [simp]:
"sorted_list_of_set (fset FS) = sorted_list_of_set (fset FS') \<longleftrightarrow> FS = FS'"
by auto
end
diff --git a/thys/Core_SC_DOM/common/monads/CharacterDataMonad.thy b/thys/Core_SC_DOM/common/monads/CharacterDataMonad.thy
--- a/thys/Core_SC_DOM/common/monads/CharacterDataMonad.thy
+++ b/thys/Core_SC_DOM/common/monads/CharacterDataMonad.thy
@@ -1,534 +1,534 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>CharacterData\<close>
text\<open>In this theory, we introduce the monadic method setup for the CharacterData class.\<close>
theory CharacterDataMonad
imports
ElementMonad
"../classes/CharacterDataClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog"
global_interpretation l_ptr_kinds_M character_data_ptr_kinds
defines character_data_ptr_kinds_M = a_ptr_kinds_M .
lemmas character_data_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma character_data_ptr_kinds_M_eq:
assumes "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> character_data_ptr_kinds_M|\<^sub>r = |h' \<turnstile> character_data_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: character_data_ptr_kinds_M_defs node_ptr_kinds_M_defs
character_data_ptr_kinds_def)
lemma character_data_ptr_kinds_M_reads:
"reads (\<Union>node_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node_ptr RObject.nothing)}) character_data_ptr_kinds_M h h'"
using node_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs character_data_ptr_kinds_M_defs
character_data_ptr_kinds_def preserved_def)
by (smt node_ptr_kinds_small preserved_def unit_all_impI)
global_interpretation l_dummy defines get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = "l_get_M.a_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf character_data_ptr_kinds"
apply(simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf l_get_M_def)
by (metis (no_types, opaque_lifting) NodeMonad.get_M_is_l_get_M bind_eq_Some_conv
character_data_ptr_kinds_commutes get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def l_get_M_def option.distinct(1))
lemmas get_M_defs = get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
locale l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
sublocale l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
interpretation l_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf character_data_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf local.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a)
by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ok = get_M_ok[folded get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def]
end
global_interpretation l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf character_data_ptr_kinds get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
rewrites "a_get_M = get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a" defines put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
locale l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
begin
sublocale l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
interpretation l_put_M type_wf character_data_ptr_kinds get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
apply(unfold_locales)
using get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a local.l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms
apply blast
by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ok = put_M_ok[folded put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def]
end
global_interpretation l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf by unfold_locales
lemma CharacterData_simp1 [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = v) \<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp3 [simp]: "
(\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' getter) h h'"
apply(cases "character_data_ptr = character_data_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp4 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp5 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp6 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply (cases "cast character_data_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv split: option.splits)
lemma CharacterData_simp7 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast character_data_ptr = node_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv split: option.splits)
lemma CharacterData_simp8 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def NodeMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp9 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast character_data_ptr \<noteq> node_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
NodeMonad.get_M_defs preserved_def split: option.splits bind_splits
dest: get_heap_E)
lemma CharacterData_simp10 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: NodeMonad.put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def NodeMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp11 [simp]:
"cast character_data_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp12 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast character_data_ptr \<noteq> object_ptr")
apply(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)[1]
lemma CharacterData_simp13 [simp]:
"cast character_data_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma new_element_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def split: prod.splits option.splits
elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Creating CharacterData\<close>
definition new_character_data :: "(_, (_) character_data_ptr) dom_prog"
where
"new_character_data = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h);
return_heap h';
return new_ptr
}"
lemma new_character_data_ok [simp]:
"h \<turnstile> ok new_character_data"
by(auto simp add: new_character_data_def split: prod.splits)
lemma new_character_data_ptr_in_heap:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "new_character_data_ptr |\<in>| character_data_ptr_kinds h'"
using assms
unfolding new_character_data_def
by(auto simp add: new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap
is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_ptr_not_in_heap:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "new_character_data_ptr |\<notin>| character_data_ptr_kinds h"
using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap
by(auto simp add: new_character_data_def split: prod.splits
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_new_ptr:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_new_ptr
by(auto simp add: new_character_data_def split: prod.splits
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_is_character_data_ptr:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "is_character_data_ptr new_character_data_ptr"
using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_is_character_data_ptr
by(auto simp add: new_character_data_def elim!: bind_returns_result_E split: prod.splits)
lemma new_character_data_child_nodes:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "h' \<turnstile> get_M new_character_data_ptr val \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> cast new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> cast new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_character_data_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_CharacterData_ptr_simp [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast character_data_ptr then cast obj else get character_data_ptr h)"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def split: option.splits Option.bind_splits)
lemma Character_data_ptr_kinds_simp [simp]:
"character_data_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = character_data_ptr_kinds h |\<union>|
(if is_character_data_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: character_data_ptr_kinds_def is_node_ptr_kind_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "ElementClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_character_data_ptr_kind ptr \<Longrightarrow> is_character_data_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: ElementMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)[1]
using assms(2) node_ptr_kinds_commutes by blast
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "ElementClass.type_wf h"
assumes "is_character_data_ptr_kind ptr \<Longrightarrow> is_character_data_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
by (metis (no_types, lifting) ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf assms(2) bind.bind_lunit
- cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def notin_fset option.collapse)
+ cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def option.collapse)
subsection\<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I split: if_splits)[1]
using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast
- using element_ptrs_def apply fastforce
+ using element_ptrs_def apply force
using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_element_ptr_ref)
lemma new_element_is_l_new_element: "l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
+ apply metis
done
lemma new_character_data_type_wf_preserved [simp]:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(simp_all add: type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs is_node_kind_def)
by (meson new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap)
locale l_new_character_data = l_type_wf +
assumes new_character_data_types_preserved: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_character_data_is_l_new_character_data: "l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_val_type_wf_preserved [simp]:
"h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
CharacterDataClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e CharacterDataClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs
split: option.splits)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs
ObjectClass.a_type_wf_def
split: option.splits)[1]
- apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
- apply (metis finite_set_in)
+ apply (metis (no_types, lifting) bind_eq_Some_conv get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
+ apply metis
done
lemma character_data_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "character_data_ptr_kinds h = character_data_ptr_kinds h'"
by(simp add: character_data_ptr_kinds_def node_ptr_kinds_def preserved_def
object_ptr_kinds_preserved_small[OF assms])
lemma character_data_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "character_data_ptr_kinds h = character_data_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def character_data_ptr_kinds_def)
by (metis assms node_ptr_kinds_preserved)
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr
RCharacterData.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3)]
allI[OF assms(4), of id, simplified] character_data_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs character_data_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply(force)
by force
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr
RCharacterData.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_def ElementMonad.type_wf_drop
l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a.a_type_wf_def)[1]
using type_wf_drop
by (metis (no_types, lifting) ElementClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
- character_data_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
+ character_data_ptr_kinds_commutes fmlookup_drop get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def node_ptr_kinds_commutes object_ptr_kinds_code5)
end
diff --git a/thys/Core_SC_DOM/common/monads/DocumentMonad.thy b/thys/Core_SC_DOM/common/monads/DocumentMonad.thy
--- a/thys/Core_SC_DOM/common/monads/DocumentMonad.thy
+++ b/thys/Core_SC_DOM/common/monads/DocumentMonad.thy
@@ -1,614 +1,606 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Document\<close>
text\<open>In this theory, we introduce the monadic method setup for the Document class.\<close>
theory DocumentMonad
imports
CharacterDataMonad
"../classes/DocumentClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'result) dom_prog"
global_interpretation l_ptr_kinds_M document_ptr_kinds defines document_ptr_kinds_M = a_ptr_kinds_M .
lemmas document_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma document_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: document_ptr_kinds_M_defs object_ptr_kinds_M_defs document_ptr_kinds_def)
lemma document_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) document_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def object_ptr_kinds_M_defs document_ptr_kinds_M_defs
document_ptr_kinds_def preserved_def cong del: image_cong_simp)
apply (metis (mono_tags, opaque_lifting) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def)
done
global_interpretation l_dummy defines get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = "l_get_M.a_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf document_ptr_kinds"
apply(simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf l_get_M_def)
by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs bind_eq_None_conv
document_ptr_kinds_commutes get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.simps(3))
lemmas get_M_defs = get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf document_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson DocumentMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = get_M_ok[folded get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf document_ptr_kinds get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
rewrites "a_get_M = get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t" defines put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_put_M type_wf document_ptr_kinds get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
apply (simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson DocumentMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = put_M_ok[folded put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
lemma document_put_get [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Mdocument_preserved1 [simp]:
"document_ptr \<noteq> document_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma document_put_get_preserved [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' getter) h h'"
apply(cases "document_ptr = document_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved2 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved3 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved4 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast document_ptr \<noteq> object_ptr")[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Mdocument_preserved5 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved7 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved8 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved10 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast document_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection \<open>Creating Documents\<close>
definition new_document :: "(_, (_) document_ptr) dom_prog"
where
"new_document = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h);
return_heap h';
return new_ptr
}"
lemma new_document_ok [simp]:
"h \<turnstile> ok new_document"
by(auto simp add: new_document_def split: prod.splits)
lemma new_document_ptr_in_heap:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "new_document_ptr |\<in>| document_ptr_kinds h'"
using assms
unfolding new_document_def
by(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_ptr_not_in_heap:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "new_document_ptr |\<notin>| document_ptr_kinds h"
using assms new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap
by(auto simp add: new_document_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_new_ptr:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_document_ptr|}"
using assms new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr
by(auto simp add: new_document_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_is_document_ptr:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "is_document_ptr new_document_ptr"
using assms new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_document_ptr
by(auto simp add: new_document_def elim!: bind_returns_result_E split: prod.splits)
lemma new_document_doctype:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr doctype \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_document_element:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr document_element \<rightarrow>\<^sub>r None"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_disconnected_nodes:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr disconnected_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> ptr \<noteq> cast new_document_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_document_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new_document_def CharacterDataMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> ptr \<noteq> new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection \<open>Modified Heaps\<close>
lemma get_document_ptr_simp [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast document_ptr then cast obj else get document_ptr h)"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: option.splits Option.bind_splits)
lemma document_ptr_kidns_simp [simp]:
"document_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= document_ptr_kinds h |\<union>| (if is_document_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: document_ptr_kinds_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "CharacterDataClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_document_ptr_kind ptr \<Longrightarrow> is_document_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs is_document_kind_def split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "CharacterDataClass.type_wf h"
assumes "is_document_ptr_kind ptr \<Longrightarrow> is_document_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1]
- by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
- is_document_kind_def notin_fset option.exhaust_sel)
+ by (metis (no_types, opaque_lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
+ is_document_kind_def option.exhaust_sel)
subsection \<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def element_ptrs_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
- apply fastforce
+ apply force
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_element_ptr_ref)
lemma new_element_is_l_new_element [instances]:
"l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
- bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
- option.distinct(1) option.simps(3))
- by (metis fmember_iff_member_fset)
+ apply (metis bind.bind_lzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.collapse option.simps(3))
+ by metis
lemma new_character_data_type_wf_preserved [simp]:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2 bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
by (meson new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap)
lemma new_character_data_is_l_new_character_data [instances]:
"l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_val_type_wf_preserved [simp]:
"h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t is_node_kind_def
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- apply (metis bind.bind_lzero finite_set_in get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def option.distinct(1) option.exhaust_sel)
- by (metis finite_set_in)
+ apply (metis bind.bind_lzero get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def option.distinct(1) option.exhaust_sel)
+ by metis
lemma new_document_type_wf_preserved [simp]: "h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(auto simp add: type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_document_kind_def
split: option.splits)[1]
- using document_ptrs_def apply fastforce
+ using document_ptrs_def apply force
apply (simp add: is_document_kind_def)
apply (metis Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_document_ptr_ref)
done
locale l_new_document = l_type_wf +
assumes new_document_types_preserved: "h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_document_is_l_new_document [instances]: "l_new_document type_wf"
using l_new_document.intro new_document_type_wf_preserved
by blast
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doctype_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr doctype_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: get_M_defs)[1]
- by (metis (mono_tags) error_returns_result finite_set_in option.exhaust_sel option.simps(4))
+ by (metis (mono_tags) error_returns_result option.exhaust_sel option.simps(4))
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr document_element_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t is_node_ptr_kind_none
cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none is_document_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: get_M_defs is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs
split: option.splits)[1]
- by (metis finite_set_in)
+ by metis
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr disconnected_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a
DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none
cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none is_document_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def get_M_defs type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
- by (metis finite_set_in)
+ by metis
lemma document_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "document_ptr_kinds h = document_ptr_kinds h'"
by(simp add: document_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms])
lemma document_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "document_ptr_kinds h = document_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def document_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved
(get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
shows "DocumentClass.type_wf h = DocumentClass.type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3) assms(4)]
allI[OF assms(5), of id, simplified] document_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs )[1]
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply force
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
by force
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved
(get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
shows "DocumentClass.type_wf h = DocumentClass.type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> DocumentClass.type_wf h = DocumentClass.type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs)[1]
using type_wf_drop
apply blast
- by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf CharacterDataMonad.type_wf_drop
- document_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel)
+ by (metis (no_types, opaque_lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf CharacterDataMonad.type_wf_drop
+ document_ptr_kinds_commutes fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel)
end
diff --git a/thys/Core_SC_DOM/common/monads/ElementMonad.thy b/thys/Core_SC_DOM/common/monads/ElementMonad.thy
--- a/thys/Core_SC_DOM/common/monads/ElementMonad.thy
+++ b/thys/Core_SC_DOM/common/monads/ElementMonad.thy
@@ -1,445 +1,445 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Element\<close>
text\<open>In this theory, we introduce the monadic method setup for the Element class.\<close>
theory ElementMonad
imports
NodeMonad
"ElementClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element,'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element,'result) dom_prog"
global_interpretation l_ptr_kinds_M element_ptr_kinds defines element_ptr_kinds_M = a_ptr_kinds_M .
lemmas element_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma element_ptr_kinds_M_eq:
assumes "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> element_ptr_kinds_M|\<^sub>r = |h' \<turnstile> element_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: element_ptr_kinds_M_defs node_ptr_kinds_M_defs element_ptr_kinds_def)
lemma element_ptr_kinds_M_reads:
"reads (\<Union>element_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t element_ptr RObject.nothing)}) element_ptr_kinds_M h h'"
apply (simp add: reads_def node_ptr_kinds_M_defs element_ptr_kinds_M_defs element_ptr_kinds_def
node_ptr_kinds_M_reads preserved_def cong del: image_cong_simp)
apply (metis (mono_tags, opaque_lifting) node_ptr_kinds_small old.unit.exhaust preserved_def)
done
global_interpretation l_dummy defines get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = "l_get_M.a_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf element_ptr_kinds"
apply(simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf l_get_M_def)
by (metis (no_types, lifting) ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs
bind_eq_Some_conv bind_eq_Some_conv element_ptr_kinds_commutes get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def node_ptr_kinds_commutes option.simps(3))
lemmas get_M_defs = get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas by unfold_locales
interpretation l_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf element_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson ElementMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = get_M_ok[folded get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
lemmas get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap = get_M_ptr_in_heap[folded get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf element_ptr_kinds get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
rewrites "a_get_M = get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t"
defines put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas by unfold_locales
interpretation l_put_M type_wf element_ptr_kinds get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
apply (simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf local.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
by (meson ElementMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok = put_M_ok[folded put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
lemma element_put_get [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Element_preserved1 [simp]:
"element_ptr \<noteq> element_ptr' \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma element_put_get_preserved [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = getter x) \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' getter) h h'"
apply(cases "element_ptr = element_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved3 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast element_ptr = object_ptr")
by (auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma get_M_Element_preserved4 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast element_ptr = node_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma get_M_Element_preserved5 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast element_ptr \<noteq> node_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Element_preserved7 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: NodeMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved8 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast element_ptr \<noteq> object_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Element_preserved10 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
subsection\<open>Creating Elements\<close>
definition new_element :: "(_, (_) element_ptr) dom_prog"
where
"new_element = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h);
return_heap h';
return new_ptr
}"
lemma new_element_ok [simp]:
"h \<turnstile> ok new_element"
by(auto simp add: new_element_def split: prod.splits)
lemma new_element_ptr_in_heap:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "new_element_ptr |\<in>| element_ptr_kinds h'"
using assms
unfolding new_element_def
by(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_ptr_not_in_heap:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "new_element_ptr |\<notin>| element_ptr_kinds h"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
bind_returns_heap_E)
lemma new_element_new_ptr:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
bind_returns_heap_E)
lemma new_element_is_element_ptr:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "is_element_ptr new_element_ptr"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_element_ptr
by(auto simp add: new_element_def elim!: bind_returns_result_E split: prod.splits)
lemma new_element_child_nodes:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr child_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_tag_name:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr tag_name \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_attrs:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr attrs \<rightarrow>\<^sub>r fmempty"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_shadow_root_opt:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr shadow_root_opt \<rightarrow>\<^sub>r None"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_element_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_Element_ptr_simp [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast element_ptr then cast obj else get element_ptr h)"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: option.splits Option.bind_splits)
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= element_ptr_kinds h |\<union>| (if is_element_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: element_ptr_kinds_def is_node_ptr_kind_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "NodeClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_element_ptr_kind ptr \<Longrightarrow> is_element_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: NodeMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)[1]
using assms(2) node_ptr_kinds_commutes by blast
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "NodeClass.type_wf h"
assumes "is_element_ptr_kind ptr \<Longrightarrow> is_element_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
by (metis (no_types, lifting) NodeClass.l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_axioms assms(2) bind.bind_lunit
- cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
+ cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf option.collapse)
subsection\<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]: "h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
new_element_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits if_splits elim!: bind_returns_heap_E)[1]
- apply (metis element_ptr_kinds_commutes element_ptrs_def fempty_iff ffmember_filter finite_set_in
+ apply (metis element_ptr_kinds_commutes element_ptrs_def fempty_iff ffmember_filter
is_element_ptr_ref)
- apply (metis element_ptrs_def fempty_iff ffmember_filter finite_set_in is_element_ptr_ref)
- apply (metis (no_types, lifting) Suc_n_not_le_n element_ptr.sel(1) element_ptr_kinds_commutes
- element_ptrs_def fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref notin_fset)
+ apply (metis element_ptrs_def fempty_iff ffmember_filter is_element_ptr_ref)
+ apply (metis (no_types, opaque_lifting) Suc_n_not_le_n element_ptr.sel(1) element_ptr_kinds_commutes
+ element_ptrs_def fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref)
apply (metis (no_types, lifting) Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def
- fMax_ge ffmember_filter fimage_eqI finite_set_in is_element_ptr_ref)
+ fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref)
done
locale l_new_element = l_type_wf +
assumes new_element_types_preserved: "h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_element_is_l_new_element: "l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs Let_def
put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
- apply (metis finite_set_in option.inject)
- apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
+ apply (metis option.inject)
+ apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv option.sel)
done
lemma put_M_pointers_preserved:
assumes "h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
using assms
apply(auto simp add: put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
elim!: bind_returns_heap_E2 dest!: get_heap_E)[1]
by (meson get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap is_OK_returns_result_I)
lemma element_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "element_ptr_kinds h = element_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def element_ptr_kinds_def)
by (metis assms node_ptr_kinds_preserved)
lemma element_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "element_ptr_kinds h = element_ptr_kinds h'"
by(simp add: element_ptr_kinds_def node_ptr_kinds_def preserved_def
object_ptr_kinds_preserved_small[OF assms])
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2)] allI[OF assms(3), of id, simplified]
apply(auto simp add: type_wf_defs )[1]
apply(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
split: option.splits,force)[1]
by(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
split: option.splits,force)
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
node_ptr_kinds_def object_ptr_kinds_def is_node_ptr_kind_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)[1]
- apply (metis (no_types, lifting) element_ptr_kinds_commutes finite_set_in fmdom_notD fmdom_notI
+ apply (metis (no_types, lifting) element_ptr_kinds_commutes fmdom_notD fmdom_notI
fmlookup_drop heap.sel node_ptr_kinds_commutes o_apply object_ptr_kinds_def)
by (metis element_ptr_kinds_commutes fmdom_notI fmdrop_lookup heap.sel node_ptr_kinds_commutes
o_apply object_ptr_kinds_def)
end
diff --git a/thys/Core_SC_DOM/common/monads/NodeMonad.thy b/thys/Core_SC_DOM/common/monads/NodeMonad.thy
--- a/thys/Core_SC_DOM/common/monads/NodeMonad.thy
+++ b/thys/Core_SC_DOM/common/monads/NodeMonad.thy
@@ -1,218 +1,218 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Node\<close>
text\<open>In this theory, we introduce the monadic method setup for the Node class.\<close>
theory NodeMonad
imports
ObjectMonad
"../classes/NodeClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog"
global_interpretation l_ptr_kinds_M node_ptr_kinds defines node_ptr_kinds_M = a_ptr_kinds_M .
lemmas node_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma node_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: node_ptr_kinds_M_defs object_ptr_kinds_M_defs node_ptr_kinds_def)
global_interpretation l_dummy defines get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = "l_get_M.a_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf node_ptr_kinds"
apply(simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf l_get_M_def)
by (metis ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind_eq_None_conv get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
node_ptr_kinds_commutes option.simps(3))
lemmas get_M_defs = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e
locale l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
sublocale l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales
interpretation l_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf node_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf local.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e)
by (meson NodeMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ok = get_M_ok[folded get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def]
end
global_interpretation l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales
lemma node_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) node_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_def
object_ptr_kinds_M_reads preserved_def)
by (smt object_ptr_kinds_preserved_small preserved_def unit_all_impI)
global_interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
rewrites "a_get_M = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e"
defines put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e
locale l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
sublocale l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales
interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
apply(unfold_locales)
apply (simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf local.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e)
by (meson NodeMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ok = put_M_ok[folded put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def]
end
global_interpretation l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales
lemma get_M_Object_preserved1 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x)) \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast node_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv
split: option.splits)
lemma get_M_Object_preserved2 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Object_preserved3 [simp]:
"h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast node_ptr \<noteq> object_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Object_preserved4 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_node_ptr_simp [simp]:
"get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = cast node_ptr then cast obj else get node_ptr h)"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= node_ptr_kinds h |\<union>| (if is_node_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: node_ptr_kinds_def)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "ObjectClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_node_ptr_kind ptr \<Longrightarrow> is_node_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
apply(auto simp add: type_wf_defs split: option.splits)[1]
using cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none is_node_kind_def apply blast
using cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none is_node_kind_def apply blast
done
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: ObjectMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "ObjectClass.type_wf h"
assumes "is_node_ptr_kind ptr \<Longrightarrow> is_node_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
- by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit finite_set_in get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def is_node_kind_def option.exhaust_sel)
+ by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def is_node_kind_def option.exhaust_sel)
subsection\<open>Preserving Types\<close>
lemma node_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "node_ptr_kinds h = node_ptr_kinds h'"
by(simp add: node_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms])
lemma node_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "node_ptr_kinds h = node_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def node_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved allI[OF assms(2), of id, simplified]
apply(auto simp add: type_wf_defs)[1]
apply(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
- apply (metis notin_fset option.simps(3))
+ apply (metis option.simps(3))
by(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits, force)[1]
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
end
diff --git a/thys/Core_SC_DOM/common/monads/ObjectMonad.thy b/thys/Core_SC_DOM/common/monads/ObjectMonad.thy
--- a/thys/Core_SC_DOM/common/monads/ObjectMonad.thy
+++ b/thys/Core_SC_DOM/common/monads/ObjectMonad.thy
@@ -1,258 +1,258 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Object\<close>
text\<open>In this theory, we introduce the monadic method setup for the Object class.\<close>
theory ObjectMonad
imports
BaseMonad
"../classes/ObjectClass"
begin
type_synonym ('object_ptr, 'Object, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'Object, 'result) dom_prog"
global_interpretation l_ptr_kinds_M object_ptr_kinds defines object_ptr_kinds_M = a_ptr_kinds_M .
lemmas object_ptr_kinds_M_defs = a_ptr_kinds_M_def
global_interpretation l_dummy defines get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = "l_get_M.a_get_M get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t type_wf object_ptr_kinds"
by (simp add: a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf l_get_M_def)
lemmas get_M_defs = get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
locale l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas = l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
interpretation l_get_M get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t type_wf object_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf local.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
by (simp add: a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
lemmas get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok = get_M_ok[folded get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
lemmas get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap = get_M_ptr_in_heap[folded get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas type_wf
by (simp add: l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_def l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms)
lemma object_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) object_ptr_kinds_M h h'"
apply(auto simp add: object_ptr_kinds_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf type_wf_defs reads_def
preserved_def get_M_defs
split: option.splits)[1]
using a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf by blast+
global_interpretation l_put_M type_wf object_ptr_kinds get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
rewrites "a_get_M = get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"
defines put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
locale l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas = l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
interpretation l_put_M type_wf object_ptr_kinds get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
apply(unfold_locales)
using get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t local.l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms apply blast
by (simp add: a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
lemmas put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok = put_M_ok[folded put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
lemmas put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap = put_M_ptr_in_heap[folded put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas type_wf
by (simp add: l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_def l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms)
definition check_in_heap :: "(_) object_ptr \<Rightarrow> (_, unit) dom_prog"
where
"check_in_heap ptr = do {
h \<leftarrow> get_heap;
(if ptr |\<in>| object_ptr_kinds h then
return ()
else
error SegmentationFault
)}"
lemma check_in_heap_ptr_in_heap: "ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> h \<turnstile> ok (check_in_heap ptr)"
by(auto simp add: check_in_heap_def)
lemma check_in_heap_pure [simp]: "pure (check_in_heap ptr) h"
by(auto simp add: check_in_heap_def intro!: bind_pure_I)
lemma check_in_heap_is_OK [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (check_in_heap ptr \<bind> f) = h \<turnstile> ok (f ())"
by(simp add: check_in_heap_def)
lemma check_in_heap_returns_result [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> (check_in_heap ptr \<bind> f) \<rightarrow>\<^sub>r x = h \<turnstile> f () \<rightarrow>\<^sub>r x"
by(simp add: check_in_heap_def)
lemma check_in_heap_returns_heap [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> (check_in_heap ptr \<bind> f) \<rightarrow>\<^sub>h h' = h \<turnstile> f () \<rightarrow>\<^sub>h h'"
by(simp add: check_in_heap_def)
lemma check_in_heap_reads:
"reads {preserved (get_M object_ptr nothing)} (check_in_heap object_ptr) h h'"
apply(simp add: check_in_heap_def reads_def preserved_def)
by (metis a_type_wf_def get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap is_OK_returns_result_E
is_OK_returns_result_I unit_all_impI)
subsection\<open>Invoke\<close>
fun invoke_rec :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list \<Rightarrow> (_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog"
where
"invoke_rec ((P, f)#xs) ptr args = (if P ptr then f ptr args else invoke_rec xs ptr args)"
| "invoke_rec [] ptr args = error InvokeError"
definition invoke :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list
\<Rightarrow> (_) object_ptr \<Rightarrow> 'args \<Rightarrow> (_, 'result) dom_prog"
where
"invoke xs ptr args = do { check_in_heap ptr; invoke_rec xs ptr args}"
lemma invoke_split: "P (invoke ((Pred, f) # xs) ptr args) =
((\<not>(Pred ptr) \<longrightarrow> P (invoke xs ptr args))
\<and> (Pred ptr \<longrightarrow> P (do {check_in_heap ptr; f ptr args})))"
by(simp add: invoke_def)
lemma invoke_split_asm: "P (invoke ((Pred, f) # xs) ptr args) =
(\<not>((\<not>(Pred ptr) \<and> (\<not> P (invoke xs ptr args)))
\<or> (Pred ptr \<and> (\<not> P (do {check_in_heap ptr; f ptr args})))))"
by(simp add: invoke_def)
lemmas invoke_splits = invoke_split invoke_split_asm
lemma invoke_ptr_in_heap: "h \<turnstile> ok (invoke xs ptr args) \<Longrightarrow> ptr |\<in>| object_ptr_kinds h"
by (metis bind_is_OK_E check_in_heap_ptr_in_heap invoke_def is_OK_returns_heap_I)
lemma invoke_pure [simp]: "pure (invoke [] ptr args) h"
by(auto simp add: invoke_def intro!: bind_pure_I)
lemma invoke_is_OK [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> ok (invoke ((Pred, f) # xs) ptr args) = h \<turnstile> ok (f ptr args)"
by(simp add: invoke_def)
lemma invoke_returns_result [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> (invoke ((Pred, f) # xs) ptr args) \<rightarrow>\<^sub>r x = h \<turnstile> f ptr args \<rightarrow>\<^sub>r x"
by(simp add: invoke_def)
lemma invoke_returns_heap [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> (invoke ((Pred, f) # xs) ptr args) \<rightarrow>\<^sub>h h' = h \<turnstile> f ptr args \<rightarrow>\<^sub>h h'"
by(simp add: invoke_def)
lemma invoke_not [simp]: "\<not>Pred ptr \<Longrightarrow> invoke ((Pred, f) # xs) ptr args = invoke xs ptr args"
by(auto simp add: invoke_def)
lemma invoke_empty [simp]: "\<not>h \<turnstile> ok (invoke [] ptr args)"
by(auto simp add: invoke_def check_in_heap_def)
lemma invoke_empty_reads [simp]: "\<forall>P \<in> S. reflp P \<and> transp P \<Longrightarrow> reads S (invoke [] ptr args) h h'"
apply(simp add: invoke_def reads_def preserved_def)
by (meson bind_returns_result_E error_returns_result)
subsection\<open>Modified Heaps\<close>
lemma get_object_ptr_simp [simp]:
"get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = object_ptr then Some obj else get object_ptr h)"
by(auto simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: option.splits Option.bind_splits)
lemma object_ptr_kinds_simp [simp]: "object_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = object_ptr_kinds h |\<union>| {|ptr|}"
by(auto simp add: object_ptr_kinds_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs split: option.splits if_splits)
subsection\<open>Preserving Types\<close>
lemma type_wf_preserved: "type_wf h = type_wf h'"
by(auto simp add: type_wf_defs)
lemma object_ptr_kinds_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
using assms
apply(auto simp add: object_ptr_kinds_def preserved_def get_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: option.splits)[1]
- apply (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember_iff_member_fset
+ apply (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq
old.unit.exhaust option.case_eq_if return_returns_result)
- by (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember_iff_member_fset
+ by (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq
old.unit.exhaust option.case_eq_if return_returns_result)
lemma object_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w object_ptr. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
proof -
{
fix object_ptr w
have "preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
apply(rule writes_small_big[OF assms])
by auto
}
then show ?thesis
using object_ptr_kinds_preserved_small by blast
qed
lemma reads_writes_preserved2:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' x. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
shows "preserved (get_M ptr getter) h h'"
apply(clarsimp simp add: preserved_def)
using reads_singleton assms(1) assms(2)
apply(rule reads_writes_preserved)
using assms(3)
by(auto simp add: preserved_def)
end
diff --git a/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy b/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy
--- a/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy
+++ b/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy
@@ -1,6965 +1,6961 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Wellformedness\<close>
text\<open>In this theory, we discuss the wellformedness of the DOM. First, we define
wellformedness and, second, we show for all functions for querying and modifying the
DOM to what extend they preserve wellformendess.\<close>
theory Core_DOM_Heap_WF
imports
"Core_DOM_Functions"
begin
locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs =
l_get_child_nodes_defs get_child_nodes get_child_nodes_locs +
l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs
for get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_owner_document_valid :: "(_) heap \<Rightarrow> bool"
where
"a_owner_document_valid h \<longleftrightarrow> (\<forall>node_ptr \<in> fset (node_ptr_kinds h).
((\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h
\<and> node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
\<or> (\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h
\<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)))"
lemma a_owner_document_valid_code [code]: "a_owner_document_valid h \<longleftrightarrow> node_ptr_kinds h |\<subseteq>|
fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)) @
map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))
"
apply(auto simp add: a_owner_document_valid_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_owner_document_valid_def)[1]
proof -
fix x
assume 1: " \<forall>node_ptr\<in>fset (node_ptr_kinds h).
(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
assume 2: "x |\<in>| node_ptr_kinds h"
assume 3: "x |\<notin>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
have "\<not>(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
using 1 2 3
- by (smt UN_I fset_of_list_elem image_eqI notin_fset set_concat set_map sorted_list_of_fset_simps(1))
+ by (smt UN_I fset_of_list_elem image_eqI set_concat set_map sorted_list_of_fset_simps(1))
then
have "(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
using 1 2
by auto
then obtain parent_ptr where parent_ptr: "parent_ptr |\<in>| object_ptr_kinds h \<and>
x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
by auto
moreover have "parent_ptr \<in> set (sorted_list_of_fset (object_ptr_kinds h))"
using parent_ptr by auto
moreover have "|h \<turnstile> get_child_nodes parent_ptr|\<^sub>r \<in> set (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))"
using calculation(2) by auto
ultimately
show "x |\<in>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h))))"
using fset_of_list_elem by fastforce
next
fix node_ptr
assume 1: "node_ptr_kinds h |\<subseteq>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))) |\<union>|
fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
assume 2: "node_ptr |\<in>| node_ptr_kinds h"
assume 3: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<longrightarrow> node_ptr \<notin> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
have "node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))) \<or>
node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
using 1 2
by (meson fin_mono fset_of_list_elem funion_iff)
then
show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using 3
by auto
qed
definition a_parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
where
"a_parent_child_rel h = {(parent, child). parent |\<in>| object_ptr_kinds h
\<and> child \<in> cast ` set |h \<turnstile> get_child_nodes parent|\<^sub>r}"
lemma a_parent_child_rel_code [code]: "a_parent_child_rel h = set (concat (map
(\<lambda>parent. map
(\<lambda>child. (parent, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child))
|h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))
)"
by(auto simp add: a_parent_child_rel_def)
definition a_acyclic_heap :: "(_) heap \<Rightarrow> bool"
where
"a_acyclic_heap h = acyclic (a_parent_child_rel h)"
definition a_all_ptrs_in_heap :: "(_) heap \<Rightarrow> bool"
where
"a_all_ptrs_in_heap h \<longleftrightarrow>
(\<forall>ptr \<in> fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h)) \<and>
(\<forall>document_ptr \<in> fset (document_ptr_kinds h).
set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h))"
definition a_distinct_lists :: "(_) heap \<Rightarrow> bool"
where
"a_distinct_lists h = distinct (concat (
(map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r)
@ (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) |h \<turnstile> document_ptr_kinds_M|\<^sub>r)
))"
definition a_heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
where
"a_heap_is_wellformed h \<longleftrightarrow>
a_acyclic_heap h \<and> a_all_ptrs_in_heap h \<and> a_distinct_lists h \<and> a_owner_document_valid h"
end
locale l_heap_is_wellformed_defs =
fixes heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
fixes parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
global_interpretation l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
defines heap_is_wellformed = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_heap_is_wellformed get_child_nodes
get_disconnected_nodes"
and parent_child_rel = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_parent_child_rel get_child_nodes"
and acyclic_heap = a_acyclic_heap
and all_ptrs_in_heap = a_all_ptrs_in_heap
and distinct_lists = a_distinct_lists
and owner_document_valid = a_owner_document_valid
.
locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs
+ l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs
+ l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel
+ l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set" +
assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed"
assumes parent_child_rel_impl: "parent_child_rel = a_parent_child_rel"
begin
lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def]
lemmas parent_child_rel_def = parent_child_rel_impl[unfolded a_parent_child_rel_def]
lemmas acyclic_heap_def = a_acyclic_heap_def[folded parent_child_rel_impl]
lemma parent_child_rel_node_ptr:
"(parent, child) \<in> parent_child_rel h \<Longrightarrow> is_node_ptr_kind child"
by(auto simp add: parent_child_rel_def)
lemma parent_child_rel_child_nodes:
assumes "known_ptr parent"
and "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
and "child \<in> set children"
shows "(parent, cast child) \<in> parent_child_rel h"
using assms
apply(auto simp add: parent_child_rel_def is_OK_returns_result_I )[1]
using get_child_nodes_ptr_in_heap by blast
lemma parent_child_rel_child_nodes2:
assumes "known_ptr parent"
and "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
and "child \<in> set children"
and "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = child_obj"
shows "(parent, child_obj) \<in> parent_child_rel h"
using assms parent_child_rel_child_nodes by blast
lemma parent_child_rel_finite: "finite (parent_child_rel h)"
proof -
have "parent_child_rel h = (\<Union>ptr \<in> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r.
(\<Union>child \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r. {(ptr, cast child)}))"
by(auto simp add: parent_child_rel_def)
moreover have "finite (\<Union>ptr \<in> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r.
(\<Union>child \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r. {(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)}))"
by simp
ultimately show ?thesis
by simp
qed
lemma distinct_lists_no_parent:
assumes "a_distinct_lists h"
assumes "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "node_ptr \<in> set disc_nodes"
shows "\<not>(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h
\<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
using assms
apply(auto simp add: a_distinct_lists_def)[1]
proof -
fix parent_ptr :: "(_) object_ptr"
assume a1: "parent_ptr |\<in>| object_ptr_kinds h"
assume a2: "(\<Union>x\<in>fset (object_ptr_kinds h).
set |h \<turnstile> get_child_nodes x|\<^sub>r) \<inter> (\<Union>x\<in>fset (document_ptr_kinds h).
set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
assume a3: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume a4: "node_ptr \<in> set disc_nodes"
assume a5: "node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
have f6: "parent_ptr \<in> fset (object_ptr_kinds h)"
using a1 by auto
have f7: "document_ptr \<in> fset (document_ptr_kinds h)"
- using a3 by (meson fmember_iff_member_fset get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I)
+ using a3 by (meson get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I)
have "|h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes"
using a3 by simp
then show False
using f7 f6 a5 a4 a2 by blast
qed
lemma distinct_lists_disconnected_nodes:
assumes "a_distinct_lists h"
and "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
shows "distinct disc_nodes"
proof -
have h1: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
|h \<turnstile> document_ptr_kinds_M|\<^sub>r))"
using assms(1)
by(simp add: a_distinct_lists_def)
then show ?thesis
using concat_map_all_distinct[OF h1] assms(2) is_OK_returns_result_I get_disconnected_nodes_ok
by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M
l_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap
l_get_disconnected_nodes_axioms select_result_I2)
qed
lemma distinct_lists_children:
assumes "a_distinct_lists h"
and "known_ptr ptr"
and "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
shows "distinct children"
proof (cases "children = []", simp)
assume "children \<noteq> []"
have h1: "distinct (concat ((map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r)))"
using assms(1)
by(simp add: a_distinct_lists_def)
show ?thesis
using concat_map_all_distinct[OF h1] assms(2) assms(3)
by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M get_child_nodes_ptr_in_heap
is_OK_returns_result_I select_result_I2)
qed
lemma heap_is_wellformed_children_in_heap:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "child |\<in>| node_ptr_kinds h"
using assms
apply(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def)[1]
- by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I
+ by (metis (no_types, opaque_lifting) is_OK_returns_result_I
local.get_child_nodes_ptr_in_heap select_result_I2 subsetD)
lemma heap_is_wellformed_one_parent:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'"
assumes "set children \<inter> set children' \<noteq> {}"
shows "ptr = ptr'"
using assms
proof (auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1]
fix x :: "(_) node_ptr"
assume a1: "ptr \<noteq> ptr'"
assume a2: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assume a3: "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'"
assume a4: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
have f5: "|h \<turnstile> get_child_nodes ptr|\<^sub>r = children"
using a2 by simp
have "|h \<turnstile> get_child_nodes ptr'|\<^sub>r = children'"
using a3 by (meson select_result_I2)
then have "ptr \<notin> set (sorted_list_of_set (fset (object_ptr_kinds h)))
\<or> ptr' \<notin> set (sorted_list_of_set (fset (object_ptr_kinds h)))
\<or> set children \<inter> set children' = {}"
using f5 a4 a1 by (meson distinct_concat_map_E(1))
then show False
- using a3 a2 by (metis (no_types) assms(4) finite_fset fmember_iff_member_fset is_OK_returns_result_I
+ using a3 a2 by (metis (no_types) assms(4) finite_fset is_OK_returns_result_I
local.get_child_nodes_ptr_in_heap set_sorted_list_of_set)
qed
lemma parent_child_rel_child:
"h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children \<longleftrightarrow> (ptr, cast child) \<in> parent_child_rel h"
by (simp add: is_OK_returns_result_I get_child_nodes_ptr_in_heap parent_child_rel_def)
lemma parent_child_rel_acyclic: "heap_is_wellformed h \<Longrightarrow> acyclic (parent_child_rel h)"
by (simp add: acyclic_heap_def local.heap_is_wellformed_def)
lemma heap_is_wellformed_disconnected_nodes_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow> distinct disc_nodes"
using distinct_lists_disconnected_nodes local.heap_is_wellformed_def by blast
lemma parent_child_rel_parent_in_heap:
"(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> parent |\<in>| object_ptr_kinds h"
using local.parent_child_rel_def by blast
lemma parent_child_rel_child_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptr parent
\<Longrightarrow> (parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> child_ptr |\<in>| object_ptr_kinds h"
apply(auto simp add: heap_is_wellformed_def parent_child_rel_def a_all_ptrs_in_heap_def)[1]
using get_child_nodes_ok
- by (meson finite_set_in subsetD)
+ by (meson subsetD)
lemma heap_is_wellformed_disc_nodes_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> node \<in> set disc_nodes \<Longrightarrow> node |\<in>| node_ptr_kinds h"
- by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I local.a_all_ptrs_in_heap_def
+ by (metis (no_types, opaque_lifting) is_OK_returns_result_I local.a_all_ptrs_in_heap_def
local.get_disconnected_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2 subsetD)
lemma heap_is_wellformed_one_disc_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes'
\<Longrightarrow> set disc_nodes \<inter> set disc_nodes' \<noteq> {} \<Longrightarrow> document_ptr = document_ptr'"
using DocumentMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append distinct_concat_map_E(1)
is_OK_returns_result_I local.a_distinct_lists_def local.get_disconnected_nodes_ptr_in_heap
local.heap_is_wellformed_def select_result_I2
proof -
assume a1: "heap_is_wellformed h"
assume a2: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume a3: "h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes'"
assume a4: "set disc_nodes \<inter> set disc_nodes' \<noteq> {}"
have f5: "|h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes"
using a2 by (meson select_result_I2)
have f6: "|h \<turnstile> get_disconnected_nodes document_ptr'|\<^sub>r = disc_nodes'"
using a3 by (meson select_result_I2)
have "\<And>nss nssa. \<not> distinct (concat (nss @ nssa)) \<or> distinct (concat nssa::(_) node_ptr list)"
by (metis (no_types) concat_append distinct_append)
then have "distinct (concat (map (\<lambda>d. |h \<turnstile> get_disconnected_nodes d|\<^sub>r) |h \<turnstile> document_ptr_kinds_M|\<^sub>r))"
using a1 local.a_distinct_lists_def local.heap_is_wellformed_def by blast
then show ?thesis
using f6 f5 a4 a3 a2 by (meson DocumentMonad.ptr_kinds_ptr_kinds_M distinct_concat_map_E(1)
is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
qed
lemma heap_is_wellformed_children_disc_nodes_different:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> set children \<inter> set disc_nodes = {}"
by (metis (no_types, opaque_lifting) disjoint_iff_not_equal distinct_lists_no_parent
is_OK_returns_result_I local.get_child_nodes_ptr_in_heap
local.heap_is_wellformed_def select_result_I2)
lemma heap_is_wellformed_children_disc_nodes:
"heap_is_wellformed h \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h
\<Longrightarrow> \<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)
\<Longrightarrow> (\<exists>document_ptr \<in> fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
- apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)[1]
- by (meson fmember_iff_member_fset)
+ by (auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)
+
lemma heap_is_wellformed_children_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> distinct children"
by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append
distinct_concat_map_E(2) is_OK_returns_result_I local.a_distinct_lists_def
local.get_child_nodes_ptr_in_heap local.heap_is_wellformed_def
select_result_I2)
end
locale l_heap_is_wellformed = l_type_wf + l_known_ptr + l_heap_is_wellformed_defs
+ l_get_child_nodes_defs + l_get_disconnected_nodes_defs +
assumes heap_is_wellformed_children_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children
\<Longrightarrow> child |\<in>| node_ptr_kinds h"
assumes heap_is_wellformed_disc_nodes_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> node \<in> set disc_nodes \<Longrightarrow> node |\<in>| node_ptr_kinds h"
assumes heap_is_wellformed_one_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'
\<Longrightarrow> set children \<inter> set children' \<noteq> {} \<Longrightarrow> ptr = ptr'"
assumes heap_is_wellformed_one_disc_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes'
\<Longrightarrow> set disc_nodes \<inter> set disc_nodes' \<noteq> {} \<Longrightarrow> document_ptr = document_ptr'"
assumes heap_is_wellformed_children_disc_nodes_different:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> set children \<inter> set disc_nodes = {}"
assumes heap_is_wellformed_disconnected_nodes_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> distinct disc_nodes"
assumes heap_is_wellformed_children_distinct:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> distinct children"
assumes heap_is_wellformed_children_disc_nodes:
"heap_is_wellformed h \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h
\<Longrightarrow> \<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)
\<Longrightarrow> (\<exists>document_ptr \<in> fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
assumes parent_child_rel_child:
"h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<longleftrightarrow> (ptr, cast child) \<in> parent_child_rel h"
assumes parent_child_rel_finite:
"heap_is_wellformed h \<Longrightarrow> finite (parent_child_rel h)"
assumes parent_child_rel_acyclic:
"heap_is_wellformed h \<Longrightarrow> acyclic (parent_child_rel h)"
assumes parent_child_rel_node_ptr:
"(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> is_node_ptr_kind child_ptr"
assumes parent_child_rel_parent_in_heap:
"(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> parent |\<in>| object_ptr_kinds h"
assumes parent_child_rel_child_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptr parent
\<Longrightarrow> (parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> child_ptr |\<in>| object_ptr_kinds h"
interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel
apply(unfold_locales)
by(auto simp add: heap_is_wellformed_def parent_child_rel_def)
declare l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]:
"l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes
get_disconnected_nodes"
apply(auto simp add: l_heap_is_wellformed_def)[1]
using heap_is_wellformed_children_in_heap
apply blast
using heap_is_wellformed_disc_nodes_in_heap
apply blast
using heap_is_wellformed_one_parent
apply blast
using heap_is_wellformed_one_disc_parent
apply blast
using heap_is_wellformed_children_disc_nodes_different
apply blast
using heap_is_wellformed_disconnected_nodes_distinct
apply blast
using heap_is_wellformed_children_distinct
apply blast
using heap_is_wellformed_children_disc_nodes
apply blast
using parent_child_rel_child
apply (blast)
using parent_child_rel_child
apply(blast)
using parent_child_rel_finite
apply blast
using parent_child_rel_acyclic
apply blast
using parent_child_rel_node_ptr
apply blast
using parent_child_rel_parent_in_heap
apply blast
using parent_child_rel_child_in_heap
apply blast
done
subsection \<open>get\_parent\<close>
locale l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
+ l_heap_is_wellformed
type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma child_parent_dual:
assumes heap_is_wellformed: "heap_is_wellformed h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
assumes "known_ptrs h"
assumes type_wf: "type_wf h"
shows "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr"
proof -
obtain ptrs where ptrs: "h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have h1: "ptr \<in> set ptrs"
using get_child_nodes_ok assms(2) is_OK_returns_result_I
by (metis (no_types, opaque_lifting) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>\<And>thesis. (\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
get_child_nodes_ptr_in_heap returns_result_eq select_result_I2)
let ?P = "(\<lambda>ptr. get_child_nodes ptr \<bind> (\<lambda>children. return (child \<in> set children)))"
let ?filter = "filter_M ?P ptrs"
have "h \<turnstile> ok ?filter"
using ptrs type_wf
using get_child_nodes_ok
apply(auto intro!: filter_M_is_OK_I bind_is_OK_pure_I get_child_nodes_ok simp add: bind_pure_I)[1]
using assms(4) local.known_ptrs_known_ptr by blast
then obtain parent_ptrs where parent_ptrs: "h \<turnstile> ?filter \<rightarrow>\<^sub>r parent_ptrs"
by auto
have h5: "\<exists>!x. x \<in> set ptrs \<and> h \<turnstile> Heap_Error_Monad.bind (get_child_nodes x)
(\<lambda>children. return (child \<in> set children)) \<rightarrow>\<^sub>r True"
apply(auto intro!: bind_pure_returns_result_I)[1]
using heap_is_wellformed_one_parent
proof -
have "h \<turnstile> (return (child \<in> set children)::((_) heap, exception, bool) prog) \<rightarrow>\<^sub>r True"
by (simp add: assms(3))
then show
"\<exists>z. z \<in> set ptrs \<and> h \<turnstile> Heap_Error_Monad.bind (get_child_nodes z)
(\<lambda>ns. return (child \<in> set ns)) \<rightarrow>\<^sub>r True"
by (metis (no_types) assms(2) bind_pure_returns_result_I2 h1 is_OK_returns_result_I
local.get_child_nodes_pure select_result_I2)
next
fix x y
assume 0: "x \<in> set ptrs"
and 1: "h \<turnstile> Heap_Error_Monad.bind (get_child_nodes x)
(\<lambda>children. return (child \<in> set children)) \<rightarrow>\<^sub>r True"
and 2: "y \<in> set ptrs"
and 3: "h \<turnstile> Heap_Error_Monad.bind (get_child_nodes y)
(\<lambda>children. return (child \<in> set children)) \<rightarrow>\<^sub>r True"
and 4: "(\<And>h ptr children ptr' children'. heap_is_wellformed h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children'
\<Longrightarrow> set children \<inter> set children' \<noteq> {} \<Longrightarrow> ptr = ptr')"
then show "x = y"
by (metis (no_types, lifting) bind_returns_result_E disjoint_iff_not_equal heap_is_wellformed
return_returns_result)
qed
have "child |\<in>| node_ptr_kinds h"
using heap_is_wellformed_children_in_heap heap_is_wellformed assms(2) assms(3)
by fast
moreover have "parent_ptrs = [ptr]"
apply(rule filter_M_ex1[OF parent_ptrs h1 h5])
using ptrs assms(2) assms(3)
by(auto simp add: object_ptr_kinds_M_defs bind_pure_I intro!: bind_pure_returns_result_I)
ultimately show ?thesis
using ptrs parent_ptrs
by(auto simp add: bind_pure_I get_parent_def
elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I) (*slow, ca 1min *)
qed
lemma parent_child_rel_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent"
shows "(parent, cast child_node) \<in> parent_child_rel h"
using assms parent_child_rel_child get_parent_child_dual by auto
lemma heap_wellformed_induct [consumes 1, case_names step]:
assumes "heap_is_wellformed h"
and step: "\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<Longrightarrow> P (cast child)) \<Longrightarrow> P parent"
shows "P ptr"
proof -
fix ptr
have "wf ((parent_child_rel h)\<inverse>)"
by (simp add: assms(1) finite_acyclic_wf_converse parent_child_rel_acyclic parent_child_rel_finite)
then show "?thesis"
proof (induct rule: wf_induct_rule)
case (less parent)
then show ?case
using assms parent_child_rel_child
by (meson converse_iff)
qed
qed
lemma heap_wellformed_induct2 [consumes 3, case_names not_in_heap empty_children step]:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
and not_in_heap: "\<And>parent. parent |\<notin>| object_ptr_kinds h \<Longrightarrow> P parent"
and empty_children: "\<And>parent. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r [] \<Longrightarrow> P parent"
and step: "\<And>parent children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<Longrightarrow> P (cast child) \<Longrightarrow> P parent"
shows "P ptr"
proof(insert assms(1), induct rule: heap_wellformed_induct)
case (step parent)
then show ?case
proof(cases "parent |\<in>| object_ptr_kinds h")
case True
then obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
using get_child_nodes_ok assms(2) assms(3)
by (meson is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?thesis
proof (cases "children = []")
case True
then show ?thesis
using children empty_children
by simp
next
case False
then show ?thesis
using assms(6) children last_in_set step.hyps by blast
qed
next
case False
then show ?thesis
by (simp add: not_in_heap)
qed
qed
lemma heap_wellformed_induct_rev [consumes 1, case_names step]:
assumes "heap_is_wellformed h"
and step: "\<And>child. (\<And>parent child_node. cast child_node = child
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent) \<Longrightarrow> P child"
shows "P ptr"
proof -
fix ptr
have "wf ((parent_child_rel h))"
by (simp add: assms(1) local.parent_child_rel_acyclic local.parent_child_rel_finite
wf_iff_acyclic_if_finite)
then show "?thesis"
proof (induct rule: wf_induct_rule)
case (less child)
show ?case
using assms get_parent_child_dual
by (metis less.hyps parent_child_rel_parent)
qed
qed
end
interpretation i_get_parent_wf?: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed
parent_child_rel get_disconnected_nodes
using instances
by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
locale l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs
+ l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma preserves_wellformedness_writes_needed:
assumes heap_is_wellformed: "heap_is_wellformed h"
and "h \<turnstile> f \<rightarrow>\<^sub>h h'"
and "writes SW f h h'"
and preserved_get_child_nodes:
"\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. \<forall>r \<in> get_child_nodes_locs object_ptr. r h h'"
and preserved_get_disconnected_nodes:
"\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>document_ptr. \<forall>r \<in> get_disconnected_nodes_locs document_ptr. r h h'"
and preserved_object_pointers:
"\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "heap_is_wellformed h'"
proof -
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
using assms(2) assms(3) object_ptr_kinds_preserved preserved_object_pointers by blast
then have object_ptr_kinds_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq2: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
using select_result_eq by force
then have node_ptr_kinds_eq2: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by auto
then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
by auto
have document_ptr_kinds_eq2: "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
by auto
have children_eq:
"\<And>ptr children. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads assms(3) assms(2)])
using preserved_get_child_nodes by fast
then have children_eq2: "\<And>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r = |h' \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq:
"\<And>document_ptr disconnected_nodes.
h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes"
apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads assms(3) assms(2)])
using preserved_get_disconnected_nodes by fast
then have disconnected_nodes_eq2:
"\<And>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r
= |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using select_result_eq by force
have get_parent_eq: "\<And>ptr parent. h \<turnstile> get_parent ptr \<rightarrow>\<^sub>r parent = h' \<turnstile> get_parent ptr \<rightarrow>\<^sub>r parent"
apply(rule reads_writes_preserved[OF get_parent_reads assms(3) assms(2)])
using preserved_get_child_nodes preserved_object_pointers unfolding get_parent_locs_def by fast
have "a_acyclic_heap h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h"
proof
fix x
assume "x \<in> parent_child_rel h'"
then show "x \<in> parent_child_rel h"
by(simp add: parent_child_rel_def children_eq2 object_ptr_kinds_eq3)
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'"
by (simp add: children_eq2 disconnected_nodes_eq2 document_ptr_kinds_eq3
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_all_ptrs_in_heap_def node_ptr_kinds_eq3 object_ptr_kinds_eq3)
moreover have h0: "a_distinct_lists h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
have h1: "map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))
= map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))"
by (simp add: children_eq2 object_ptr_kinds_eq3)
have h2: "map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))
= map (\<lambda>document_ptr. |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))"
using disconnected_nodes_eq document_ptr_kinds_eq2 select_result_eq by force
have "a_distinct_lists h'"
using h0
by(simp add: a_distinct_lists_def h1 h2)
moreover have "a_owner_document_valid h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
by(auto simp add: a_owner_document_valid_def children_eq2 disconnected_nodes_eq2
object_ptr_kinds_eq3 node_ptr_kinds_eq3 document_ptr_kinds_eq3)
ultimately show ?thesis
by (simp add: heap_is_wellformed_def)
qed
end
interpretation i_get_parent_wf2?: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs known_ptrs get_parent get_parent_locs
heap_is_wellformed parent_child_rel get_disconnected_nodes
get_disconnected_nodes_locs
using l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms
by (simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
locale l_get_parent_wf = l_type_wf + l_known_ptrs + l_heap_is_wellformed_defs
+ l_get_child_nodes_defs + l_get_parent_defs +
assumes child_parent_dual:
"heap_is_wellformed h
\<Longrightarrow> type_wf h
\<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children
\<Longrightarrow> h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr"
assumes heap_wellformed_induct [consumes 1, case_names step]:
"heap_is_wellformed h
\<Longrightarrow> (\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children
\<Longrightarrow> child \<in> set children \<Longrightarrow> P (cast child)) \<Longrightarrow> P parent)
\<Longrightarrow> P ptr"
assumes heap_wellformed_induct_rev [consumes 1, case_names step]:
"heap_is_wellformed h
\<Longrightarrow> (\<And>child. (\<And>parent child_node. cast child_node = child
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent) \<Longrightarrow> P child)
\<Longrightarrow> P ptr"
assumes parent_child_rel_parent: "heap_is_wellformed h
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent
\<Longrightarrow> (parent, cast child_node) \<in> parent_child_rel h"
lemma get_parent_wf_is_l_get_parent_wf [instances]:
"l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel
get_child_nodes get_parent"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def)[1]
using child_parent_dual heap_wellformed_induct heap_wellformed_induct_rev parent_child_rel_parent
by metis+
subsection \<open>get\_disconnected\_nodes\<close>
subsection \<open>set\_disconnected\_nodes\<close>
subsubsection \<open>get\_disconnected\_nodes\<close>
locale l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_set_disconnected_nodes_get_disconnected_nodes
type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs
+ l_heap_is_wellformed
type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma remove_from_disconnected_nodes_removes:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "h \<turnstile> set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \<rightarrow>\<^sub>h h'"
assumes "h' \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes'"
shows "node_ptr \<notin> set disc_nodes'"
using assms
by (metis distinct_remove1_removeAll heap_is_wellformed_disconnected_nodes_distinct
set_disconnected_nodes_get_disconnected_nodes member_remove remove_code(1)
returns_result_eq)
end
locale l_set_disconnected_nodes_get_disconnected_nodes_wf = l_heap_is_wellformed
+ l_set_disconnected_nodes_get_disconnected_nodes +
assumes remove_from_disconnected_nodes_removes:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> h \<turnstile> set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes'
\<Longrightarrow> node_ptr \<notin> set disc_nodes'"
interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M?:
l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed
parent_child_rel get_child_nodes
using instances
by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma set_disconnected_nodes_get_disconnected_nodes_wf_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]:
"l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed parent_child_rel
get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs"
apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def
l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1]
using remove_from_disconnected_nodes_removes apply fast
done
subsection \<open>get\_root\_node\<close>
locale l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed
type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
+ l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
+ l_get_parent_wf
type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes
get_child_nodes_locs get_parent get_parent_locs
+ l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_ancestors get_ancestors_locs get_root_node get_root_node_locs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_ancestors_reads:
assumes "heap_is_wellformed h"
shows "reads get_ancestors_locs (get_ancestors node_ptr) h h'"
proof (insert assms(1), induct rule: heap_wellformed_induct_rev)
case (step child)
then show ?case
using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def]
apply(simp (no_asm) add: get_ancestors_def)
by(auto simp add: get_ancestors_locs_def reads_subset[OF return_reads] get_parent_reads_pointers
intro!: reads_bind_pure reads_subset[OF check_in_heap_reads]
reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads]
split: option.splits)
qed
lemma get_ancestors_ok:
assumes "heap_is_wellformed h"
and "ptr |\<in>| object_ptr_kinds h"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "h \<turnstile> ok (get_ancestors ptr)"
proof (insert assms(1) assms(2), induct rule: heap_wellformed_induct_rev)
case (step child)
then show ?case
using assms(3) assms(4)
apply(simp (no_asm) add: get_ancestors_def)
apply(simp add: assms(1) get_parent_parent_in_heap)
by(auto intro!: bind_is_OK_pure_I bind_pure_I get_parent_ok split: option.splits)
qed
lemma get_root_node_ptr_in_heap:
assumes "h \<turnstile> ok (get_root_node ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
unfolding get_root_node_def
using get_ancestors_ptr_in_heap
by auto
lemma get_root_node_ok:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
and "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_root_node ptr)"
unfolding get_root_node_def
using assms get_ancestors_ok
by auto
lemma get_ancestors_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
shows "h \<turnstile> get_ancestors (cast child) \<rightarrow>\<^sub>r (cast child) # parent # ancestors
\<longleftrightarrow> h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r parent # ancestors"
proof
assume a1: "h \<turnstile> get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors"
then have "h \<turnstile> Heap_Error_Monad.bind (check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child))
(\<lambda>_. Heap_Error_Monad.bind (get_parent child)
(\<lambda>x. Heap_Error_Monad.bind (case x of None \<Rightarrow> return [] | Some x \<Rightarrow> get_ancestors x)
(\<lambda>ancestors. return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # ancestors))))
\<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors"
by(simp add: get_ancestors_def)
then show "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r parent # ancestors"
using assms(2) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1]
using returns_result_eq by fastforce
next
assume "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r parent # ancestors"
then show "h \<turnstile> get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors"
using assms(2)
apply(simp (no_asm) add: get_ancestors_def)
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
by (metis (full_types) assms(2) check_in_heap_ptr_in_heap is_OK_returns_result_I
local.get_parent_ptr_in_heap node_ptr_kinds_commutes old.unit.exhaust
select_result_I)
qed
lemma get_ancestors_never_empty:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors"
shows "ancestors \<noteq> []"
proof(insert assms(2), induct arbitrary: ancestors rule: heap_wellformed_induct_rev[OF assms(1)])
case (1 child)
then show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then show ?case
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
next
case (Some child_node)
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
with Some show ?case
proof(induct parent_opt)
case None
then show ?case
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
next
case (Some option)
then show ?case
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
qed
qed
qed
lemma get_ancestors_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and "ancestor \<in> set ancestors"
and "h \<turnstile> get_ancestors ancestor \<rightarrow>\<^sub>r ancestor_ancestors"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "set ancestor_ancestors \<subseteq> set ancestors"
proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors
rule: heap_wellformed_induct_rev)
case (step child)
have "child |\<in>| object_ptr_kinds h"
using get_ancestors_ptr_in_heap step(2) by auto
(* then have "h \<turnstile> check_in_heap child \<rightarrow>\<^sub>r ()"
using returns_result_select_result by force *)
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then have "ancestors = [child]"
using step(2) step(3)
by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2)
show ?case
using step(2) step(3)
apply(auto simp add: \<open>ancestors = [child]\<close>)[1]
using assms(4) returns_result_eq by fastforce
next
case (Some child_node)
note s1 = Some
obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using \<open>child |\<in>| object_ptr_kinds h\<close> assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs]
by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes)
then show ?case
proof (induct parent_opt)
case None
then have "ancestors = [child]"
using step(2) step(3) s1
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
show ?case
using step(2) step(3)
apply(auto simp add: \<open>ancestors = [child]\<close>)[1]
using assms(4) returns_result_eq by fastforce
next
case (Some parent)
have "h \<turnstile> Heap_Error_Monad.bind (check_in_heap child)
(\<lambda>_. Heap_Error_Monad.bind
(case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \<Rightarrow> return []
| Some node_ptr \<Rightarrow> Heap_Error_Monad.bind (get_parent node_ptr)
(\<lambda>parent_ptr_opt. case parent_ptr_opt of None \<Rightarrow> return []
| Some x \<Rightarrow> get_ancestors x))
(\<lambda>ancestors. return (child # ancestors)))
\<rightarrow>\<^sub>r ancestors"
using step(2)
by(simp add: get_ancestors_def)
moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
using calculation
by(auto elim!: bind_returns_result_E2 split: option.splits)
ultimately have "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors"
using s1 Some
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
show ?case
using step(1)[OF s1[symmetric, simplified] Some \<open>h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors\<close>]
step(3)
apply(auto simp add: tl_ancestors)[1]
by (metis assms(4) insert_iff list.simps(15) local.step(2) returns_result_eq tl_ancestors)
qed
qed
qed
lemma get_ancestors_also_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast child \<in> set ancestors"
and "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "parent \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors: "h \<turnstile> get_ancestors (cast child) \<rightarrow>\<^sub>r child_ancestors"
by (meson assms(1) assms(4) get_ancestors_ok is_OK_returns_result_I known_ptrs
local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result
type_wf)
then have "parent \<in> set child_ancestors"
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)]
get_ancestors_ptr)
then show ?thesis
using assms child_ancestors get_ancestors_subset by blast
qed
lemma get_ancestors_obtains_children:
assumes "heap_is_wellformed h"
and "ancestor \<noteq> ptr"
and "ancestor \<in> set ancestors"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
obtains children ancestor_child where "h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children"
and "ancestor_child \<in> set children" and "cast ancestor_child \<in> set ancestors"
proof -
assume 0: "(\<And>children ancestor_child.
h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<Longrightarrow>
ancestor_child \<in> set children \<Longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set ancestors
\<Longrightarrow> thesis)"
have "\<exists>child. h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ancestor \<and> cast child \<in> set ancestors"
proof (insert assms(1) assms(2) assms(3) assms(4), induct ptr arbitrary: ancestors
rule: heap_wellformed_induct_rev)
case (step child)
have "child |\<in>| object_ptr_kinds h"
using get_ancestors_ptr_in_heap step(4) by auto
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then have "ancestors = [child]"
using step(3) step(4)
by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2)
show ?case
using step(2) step(3) step(4)
by(auto simp add: \<open>ancestors = [child]\<close>)
next
case (Some child_node)
note s1 = Some
obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using \<open>child |\<in>| object_ptr_kinds h\<close> assms(1) Some[symmetric]
using get_parent_ok known_ptrs type_wf
by (metis (no_types, lifting) is_OK_returns_result_E node_ptr_casts_commute
node_ptr_kinds_commutes)
then show ?case
proof (induct parent_opt)
case None
then have "ancestors = [child]"
using step(2) step(3) step(4) s1
apply(simp add: get_ancestors_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
show ?case
using step(2) step(3) step(4)
by(auto simp add: \<open>ancestors = [child]\<close>)
next
case (Some parent)
have "h \<turnstile> Heap_Error_Monad.bind (check_in_heap child)
(\<lambda>_. Heap_Error_Monad.bind
(case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \<Rightarrow> return []
| Some node_ptr \<Rightarrow> Heap_Error_Monad.bind (get_parent node_ptr)
(\<lambda>parent_ptr_opt. case parent_ptr_opt of None \<Rightarrow> return []
| Some x \<Rightarrow> get_ancestors x))
(\<lambda>ancestors. return (child # ancestors)))
\<rightarrow>\<^sub>r ancestors"
using step(4)
by(simp add: get_ancestors_def)
moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
using calculation
by(auto elim!: bind_returns_result_E2 split: option.splits)
ultimately have "h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors"
using s1 Some
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
(* have "ancestor \<noteq> parent" *)
have "ancestor \<in> set tl_ancestors"
using tl_ancestors step(2) step(3) by auto
show ?case
proof (cases "ancestor \<noteq> parent")
case True
show ?thesis
using step(1)[OF s1[symmetric, simplified] Some True
\<open>ancestor \<in> set tl_ancestors\<close> \<open>h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r tl_ancestors\<close>]
using tl_ancestors by auto
next
case False
have "child \<in> set ancestors"
using step(4) get_ancestors_ptr by simp
then show ?thesis
using Some False s1[symmetric] by(auto)
qed
qed
qed
qed
then obtain child where child: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ancestor"
and in_ancestors: "cast child \<in> set ancestors"
by auto
then obtain children where
children: "h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children" and
child_in_children: "child \<in> set children"
using get_parent_child_dual by blast
show thesis
using 0[OF children child_in_children] child assms(3) in_ancestors by blast
qed
lemma get_ancestors_parent_child_rel:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "(ptr, child) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> ptr \<in> set ancestors"
proof (safe)
assume 3: "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
show "ptr \<in> set ancestors"
proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by (metis (no_types, lifting) assms(2) bind_returns_result_E get_ancestors_def
in_set_member member_rec(1) return_returns_result)
next
case False
obtain ptr_child where
ptr_child: "(ptr, ptr_child) \<in> (parent_child_rel h) \<and> (ptr_child, child) \<in> (parent_child_rel h)\<^sup>*"
using converse_rtranclE[OF 1(2)] \<open>ptr \<noteq> child\<close>
by metis
then obtain ptr_child_node
where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node"
using ptr_child node_ptr_casts_commute3 parent_child_rel_node_ptr
by (metis )
then obtain children where
children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children"
proof -
assume a1: "\<And>children. \<lbrakk>h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children; ptr_child_node \<in> set children\<rbrakk>
\<Longrightarrow> thesis"
have "ptr |\<in>| object_ptr_kinds h"
using local.parent_child_rel_parent_in_heap ptr_child by blast
moreover have "ptr_child_node \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r"
by (metis calculation known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr
local.parent_child_rel_child ptr_child ptr_child_ptr_child_node
returns_result_select_result type_wf)
ultimately show ?thesis
using a1 get_child_nodes_ok type_wf known_ptrs
by (meson local.known_ptrs_known_ptr returns_result_select_result)
qed
moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \<in> (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 \<in> set ancestors"
using 1 by auto
moreover have "h \<turnstile> get_parent ptr_child_node \<rightarrow>\<^sub>r Some ptr"
using assms(1) children ptr_child_node child_parent_dual
using known_ptrs type_wf by blast
ultimately show ?thesis
using get_ancestors_also_parent assms type_wf by blast
qed
qed
next
assume 3: "ptr \<in> set ancestors"
show "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by simp
next
case False
then obtain children ptr_child_node where
children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children" and
ptr_child_node_in_ancestors: "cast ptr_child_node \<in> set ancestors"
using 1(2) assms(2) get_ancestors_obtains_children assms(1)
using known_ptrs type_wf by blast
then have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \<in> (parent_child_rel h)\<^sup>*"
using 1(1) by blast
moreover have "(ptr, cast ptr_child_node) \<in> parent_child_rel h"
using children ptr_child_node assms(1) parent_child_rel_child_nodes2
using child_parent_dual known_ptrs parent_child_rel_parent type_wf
by blast
ultimately show ?thesis
by auto
qed
qed
qed
lemma get_root_node_parent_child_rel:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r root"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "(root, child) \<in> (parent_child_rel h)\<^sup>*"
using assms get_ancestors_parent_child_rel
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
using get_ancestors_never_empty last_in_set by blast
lemma get_ancestors_eq:
assumes "heap_is_wellformed h"
and "heap_is_wellformed h'"
and "\<And>object_ptr w. object_ptr \<noteq> ptr \<Longrightarrow> w \<in> get_child_nodes_locs object_ptr \<Longrightarrow> w h h'"
and pointers_preserved: "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
and known_ptrs: "known_ptrs h"
and known_ptrs': "known_ptrs h'"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and type_wf: "type_wf h"
and type_wf': "type_wf h'"
shows "h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
proof -
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
using pointers_preserved object_ptr_kinds_preserved_small by blast
then have object_ptr_kinds_M_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
have "h' \<turnstile> ok (get_ancestors ptr)"
using get_ancestors_ok get_ancestors_ptr_in_heap object_ptr_kinds_eq3 assms(1) known_ptrs
known_ptrs' assms(2) assms(7) type_wf'
by blast
then obtain ancestors' where ancestors': "h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors'"
by auto
obtain root where root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
proof -
assume 0: "(\<And>root. h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow> thesis)"
show thesis
apply(rule 0)
using assms(7)
by(auto simp add: get_root_node_def elim!: bind_returns_result_E2 split: option.splits)
qed
have children_eq:
"\<And>p children. p \<noteq> ptr \<Longrightarrow> h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
using get_child_nodes_reads assms(3)
apply(simp add: reads_def reflp_def transp_def preserved_def)
by blast
have "acyclic (parent_child_rel h)"
using assms(1) local.parent_child_rel_acyclic by auto
have "acyclic (parent_child_rel h')"
using assms(2) local.parent_child_rel_acyclic by blast
have 2: "\<And>c parent_opt. cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \<in> set ancestors \<inter> set ancestors'
\<Longrightarrow> h \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt = h' \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt"
proof -
fix c parent_opt
assume 1: " cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \<in> set ancestors \<inter> set ancestors'"
obtain ptrs where ptrs: "h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by simp
let ?P = "(\<lambda>ptr. Heap_Error_Monad.bind (get_child_nodes ptr) (\<lambda>children. return (c \<in> set children)))"
have children_eq_True: "\<And>p. p \<in> set ptrs \<Longrightarrow> h \<turnstile> ?P p \<rightarrow>\<^sub>r True \<longleftrightarrow> h' \<turnstile> ?P p \<rightarrow>\<^sub>r True"
proof -
fix p
assume "p \<in> set ptrs"
then show "h \<turnstile> ?P p \<rightarrow>\<^sub>r True \<longleftrightarrow> h' \<turnstile> ?P p \<rightarrow>\<^sub>r True"
proof (cases "p = ptr")
case True
have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h)\<^sup>*"
using get_ancestors_parent_child_rel 1 assms by blast
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h)"
proof (cases "cast c = ptr")
case True
then show ?thesis
using \<open>acyclic (parent_child_rel h)\<close> by(auto simp add: acyclic_def)
next
case False
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h)\<^sup>*"
using \<open>acyclic (parent_child_rel h)\<close> False rtrancl_eq_or_trancl rtrancl_trancl_trancl
\<open>(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h)\<^sup>*\<close>
by (metis acyclic_def)
then show ?thesis
using r_into_rtrancl by auto
qed
obtain children where children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using type_wf
by (metis \<open>h' \<turnstile> ok get_ancestors ptr\<close> assms(1) get_ancestors_ptr_in_heap get_child_nodes_ok
heap_is_wellformed_def is_OK_returns_result_E known_ptrs local.known_ptrs_known_ptr
object_ptr_kinds_eq3)
then have "c \<notin> set children"
using \<open>(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h)\<close> assms(1)
using parent_child_rel_child_nodes2
using child_parent_dual known_ptrs parent_child_rel_parent
type_wf by blast
with children have "h \<turnstile> ?P p \<rightarrow>\<^sub>r False"
by(auto simp add: True)
moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h')\<^sup>*"
using get_ancestors_parent_child_rel assms(2) ancestors' 1 known_ptrs' type_wf
type_wf' by blast
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')"
proof (cases "cast c = ptr")
case True
then show ?thesis
using \<open>acyclic (parent_child_rel h')\<close> by(auto simp add: acyclic_def)
next
case False
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')\<^sup>*"
using \<open>acyclic (parent_child_rel h')\<close> False rtrancl_eq_or_trancl rtrancl_trancl_trancl
\<open>(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \<in> (parent_child_rel h')\<^sup>*\<close>
by (metis acyclic_def)
then show ?thesis
using r_into_rtrancl by auto
qed
then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')"
using r_into_rtrancl by auto
obtain children' where children': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'"
using type_wf type_wf'
by (meson \<open>h' \<turnstile> ok (get_ancestors ptr)\<close> assms(2) get_ancestors_ptr_in_heap
get_child_nodes_ok is_OK_returns_result_E known_ptrs'
local.known_ptrs_known_ptr)
then have "c \<notin> set children'"
using \<open>(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<notin> (parent_child_rel h')\<close> assms(2) type_wf type_wf'
using parent_child_rel_child_nodes2 child_parent_dual known_ptrs' parent_child_rel_parent
by auto
with children' have "h' \<turnstile> ?P p \<rightarrow>\<^sub>r False"
by(auto simp add: True)
ultimately show ?thesis
by (metis returns_result_eq)
next
case False
then show ?thesis
using children_eq ptrs
by (metis (no_types, lifting) bind_pure_returns_result_I bind_returns_result_E
get_child_nodes_pure return_returns_result)
qed
qed
have "\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> ok (get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children))) = h' \<turnstile> ok ( get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)))"
using assms(1) assms(2) object_ptr_kinds_eq ptrs type_wf type_wf'
by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M bind_is_OK_pure_I
get_child_nodes_ok get_child_nodes_pure known_ptrs'
local.known_ptrs_known_ptr return_ok select_result_I2)
have children_eq_False:
"\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False = h' \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
proof
fix pa
assume "pa \<in> set ptrs"
and "h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
have "h \<turnstile> ok (get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))
\<Longrightarrow> h' \<turnstile> ok ( get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))"
using \<open>pa \<in> set ptrs\<close> \<open>\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> ok (get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children))) = h' \<turnstile> ok ( get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)))\<close>
by auto
moreover have "h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False
\<Longrightarrow> h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
by (metis (mono_tags, lifting) \<open>\<And>pa. pa \<in> set ptrs
\<Longrightarrow> h \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True = h' \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True\<close> \<open>pa \<in> set ptrs\<close>
calculation is_OK_returns_result_I returns_result_eq returns_result_select_result)
ultimately show "h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
using \<open>h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False\<close>
by auto
next
fix pa
assume "pa \<in> set ptrs"
and "h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
have "h' \<turnstile> ok (get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))
\<Longrightarrow> h \<turnstile> ok ( get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)))"
using \<open>pa \<in> set ptrs\<close> \<open>\<And>pa. pa \<in> set ptrs
\<Longrightarrow> h \<turnstile> ok (get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children))) = h' \<turnstile> ok ( get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)))\<close>
by auto
moreover have "h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False
\<Longrightarrow> h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
by (metis (mono_tags, lifting)
\<open>\<And>pa. pa \<in> set ptrs \<Longrightarrow> h \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True = h' \<turnstile> get_child_nodes pa
\<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r True\<close> \<open>pa \<in> set ptrs\<close>
calculation is_OK_returns_result_I returns_result_eq returns_result_select_result)
ultimately show "h \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False"
using \<open>h' \<turnstile> get_child_nodes pa \<bind> (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r False\<close> by blast
qed
have filter_eq: "\<And>xs. h \<turnstile> filter_M ?P ptrs \<rightarrow>\<^sub>r xs = h' \<turnstile> filter_M ?P ptrs \<rightarrow>\<^sub>r xs"
proof (rule filter_M_eq)
show
"\<And>xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children))) h"
by(auto intro!: bind_pure_I)
next
show
"\<And>xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children))) h'"
by(auto intro!: bind_pure_I)
next
fix xs b x
assume 0: "x \<in> set ptrs"
then show "h \<turnstile> Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r b
= h' \<turnstile> Heap_Error_Monad.bind (get_child_nodes x) (\<lambda>children. return (c \<in> set children)) \<rightarrow>\<^sub>r b"
apply(induct b)
using children_eq_True apply blast
using children_eq_False apply blast
done
qed
show "h \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt = h' \<turnstile> get_parent c \<rightarrow>\<^sub>r parent_opt"
apply(simp add: get_parent_def)
apply(rule bind_cong_2)
apply(simp)
apply(simp)
apply(simp add: check_in_heap_def node_ptr_kinds_def object_ptr_kinds_eq3)
apply(rule bind_cong_2)
apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
apply(rule bind_cong_2)
apply(auto intro!: filter_M_pure_I bind_pure_I)[1]
apply(auto intro!: filter_M_pure_I bind_pure_I)[1]
apply(auto simp add: filter_eq (* dest!: returns_result_eq[OF ptrs] *))[1]
using filter_eq ptrs apply auto[1]
using filter_eq ptrs by auto
qed
have "ancestors = ancestors'"
proof(insert assms(1) assms(7) ancestors' 2, induct ptr arbitrary: ancestors ancestors'
rule: heap_wellformed_induct_rev)
case (step child)
show ?case
using step(2) step(3) step(4)
apply(simp add: get_ancestors_def)
apply(auto intro!: elim!: bind_returns_result_E2 split: option.splits)[1]
using returns_result_eq apply fastforce
apply (meson option.simps(3) returns_result_eq)
by (metis IntD1 IntD2 option.inject returns_result_eq step.hyps)
qed
then show ?thesis
using assms(5) ancestors'
by simp
qed
lemma get_ancestors_remains_not_in_ancestors:
assumes "heap_is_wellformed h"
and "heap_is_wellformed h'"
and "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
and "h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors'"
and "\<And>p children children'. h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children
\<Longrightarrow> h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children' \<Longrightarrow> set children' \<subseteq> set children"
and "node \<notin> 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 \<notin> set ancestors'"
proof -
have object_ptr_kinds_M_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
using object_ptr_kinds_eq3
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
show ?thesis
proof (insert assms(1) assms(3) assms(4) assms(6), induct ptr arbitrary: ancestors ancestors'
rule: heap_wellformed_induct_rev)
case (step child)
have 1: "\<And>p parent. h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent \<Longrightarrow> h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
proof -
fix p parent
assume "h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
then obtain children' where
children': "h' \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children'" and
p_in_children': "p \<in> set children'"
using get_parent_child_dual by blast
obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^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 \<open>h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent\<close> get_parent_parent_in_heap is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
have "p \<in> set children"
using assms(5) children children' p_in_children'
by blast
then show "h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
using child_parent_dual assms(1) children known_ptrs type_wf by blast
qed
have "node \<noteq> child"
using assms(1) get_ancestors_parent_child_rel step.prems(1) step.prems(3) known_ptrs
using type_wf type_wf'
by blast
then show ?case
using step(2) step(3)
apply(simp add: get_ancestors_def)
using step(4)
apply(auto elim!: bind_returns_result_E2 split: option.splits)[1]
using 1
apply (meson option.distinct(1) returns_result_eq)
by (metis "1" option.inject returns_result_eq step.hyps)
qed
qed
lemma get_ancestors_ptrs_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
shows "ptr' |\<in>| 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 \<turnstile> get_ancestors x \<rightarrow>\<^sub>r a # ancestors"
by(auto simp add: get_ancestors_def[of a] elim!: bind_returns_result_E2 split: option.splits)
then have "x = a"
by(auto simp add: get_ancestors_def[of x] elim!: bind_returns_result_E2 split: option.splits)
then show ?case
using Cons.hyps Cons.prems(2) get_ancestors_ptr_in_heap x
by (metis assms(1) assms(2) assms(3) get_ancestors_obtains_children get_child_nodes_ptr_in_heap
is_OK_returns_result_I)
qed
lemma get_ancestors_prefix:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
assumes "h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'"
shows "\<exists>pre. ancestors = pre @ ancestors'"
proof (insert assms(1) assms(5) assms(6), induct ptr' arbitrary: ancestors'
rule: heap_wellformed_induct)
case (step parent)
then show ?case
proof (cases "parent \<noteq> ptr" )
case True
then obtain children ancestor_child where "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
and "ancestor_child \<in> set children" and "cast ancestor_child \<in> set ancestors"
using assms(1) assms(2) assms(3) assms(4) get_ancestors_obtains_children step.prems(1) by blast
then have "h \<turnstile> get_parent ancestor_child \<rightarrow>\<^sub>r Some parent"
using assms(1) assms(2) assms(3) child_parent_dual by blast
then have "h \<turnstile> get_ancestors (cast ancestor_child) \<rightarrow>\<^sub>r cast ancestor_child # ancestors'"
apply(simp add: get_ancestors_def)
using \<open>h \<turnstile> get_ancestors parent \<rightarrow>\<^sub>r ancestors'\<close> get_parent_ptr_in_heap
by(auto simp add: check_in_heap_def is_OK_returns_result_I intro!: bind_pure_returns_result_I)
then show ?thesis
using step(1) \<open>h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children\<close> \<open>ancestor_child \<in> set children\<close>
\<open>cast ancestor_child \<in> set ancestors\<close> \<open>h \<turnstile> get_ancestors (cast ancestor_child) \<rightarrow>\<^sub>r cast ancestor_child # ancestors'\<close>
by fastforce
next
case False
then show ?thesis
by (metis append_Nil assms(4) returns_result_eq step.prems(2))
qed
qed
lemma get_ancestors_same_root_node:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
assumes "ptr'' \<in> set ancestors"
shows "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr \<longleftrightarrow> h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
proof -
have "ptr' |\<in>| object_ptr_kinds h"
by (metis assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_obtains_children
get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I)
then obtain ancestors' where ancestors': "h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'"
by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E)
then have "\<exists>pre. ancestors = pre @ ancestors'"
using get_ancestors_prefix assms by blast
moreover have "ptr'' |\<in>| object_ptr_kinds h"
by (metis assms(1) assms(2) assms(3) assms(4) assms(6) get_ancestors_obtains_children
get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I)
then obtain ancestors'' where ancestors'': "h \<turnstile> get_ancestors ptr'' \<rightarrow>\<^sub>r ancestors''"
by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E)
then have "\<exists>pre. ancestors = pre @ ancestors''"
using get_ancestors_prefix assms by blast
ultimately show ?thesis
using ancestors' ancestors''
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I)[1]
apply (metis (no_types, lifting) assms(1) get_ancestors_never_empty last_appendR
returns_result_eq)
by (metis assms(1) get_ancestors_never_empty last_appendR returns_result_eq)
qed
lemma get_root_node_parent_same:
assumes "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr"
shows "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root \<longleftrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
proof
assume 1: " h \<turnstile> get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r root"
show "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
using 1[unfolded get_root_node_def] assms
apply(simp add: get_ancestors_def)
apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I split: option.splits)[1]
using returns_result_eq apply fastforce
using get_ancestors_ptr by fastforce
next
assume 1: " h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
show "h \<turnstile> get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r root"
apply(simp add: get_root_node_def)
using assms 1
apply(simp add: get_ancestors_def)
apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I split: option.splits)[1]
apply (simp add: check_in_heap_def is_OK_returns_result_I)
using get_ancestors_ptr get_parent_ptr_in_heap
apply (simp add: is_OK_returns_result_I)
by (meson list.distinct(1) list.set_cases local.get_ancestors_ptr)
qed
lemma get_root_node_same_no_parent:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r cast child"
shows "h \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev)
case (step c)
then show ?case
proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c")
case None
then have "c = cast child"
using step(2)
by(auto simp add: get_root_node_def get_ancestors_def[of c] elim!: bind_returns_result_E2)
then show ?thesis
using None by auto
next
case (Some child_node)
note s = this
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
by (metis (no_types, lifting) assms(2) assms(3) get_root_node_ptr_in_heap
is_OK_returns_result_I local.get_parent_ok node_ptr_casts_commute
node_ptr_kinds_commutes returns_result_select_result step.prems)
then show ?thesis
proof(induct parent_opt)
case None
then show ?case
using Some get_root_node_no_parent returns_result_eq step.prems by fastforce
next
case (Some parent)
then show ?case
using step s
apply(auto simp add: get_root_node_def get_ancestors_def[of c]
elim!: bind_returns_result_E2 split: option.splits list.splits)[1]
using get_root_node_parent_same step.hyps step.prems by auto
qed
qed
qed
lemma get_root_node_not_node_same:
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "\<not>is_node_ptr_kind ptr"
shows "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r ptr"
using assms
apply(simp add: get_root_node_def get_ancestors_def)
by(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I split: option.splits)
lemma get_root_node_root_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
shows "root |\<in>| object_ptr_kinds h"
using assms
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
by (simp add: get_ancestors_never_empty get_ancestors_ptrs_in_heap)
lemma get_root_node_same_no_parent_parent_child_rel:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r ptr'"
shows "\<not>(\<exists>p. (p, ptr') \<in> (parent_child_rel h))"
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) get_root_node_same_no_parent
l_heap_is_wellformed.parent_child_rel_child local.child_parent_dual local.get_child_nodes_ok
local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms local.parent_child_rel_node_ptr
local.parent_child_rel_parent_in_heap node_ptr_casts_commute3 option.simps(3) returns_result_eq
returns_result_select_result)
end
locale l_get_ancestors_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_ancestors_defs
+ l_get_child_nodes_defs + l_get_parent_defs +
assumes get_ancestors_never_empty:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors \<Longrightarrow> ancestors \<noteq> []"
assumes get_ancestors_ok:
"heap_is_wellformed h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> h \<turnstile> ok (get_ancestors ptr)"
assumes get_ancestors_reads:
"heap_is_wellformed h \<Longrightarrow> reads get_ancestors_locs (get_ancestors node_ptr) h h'"
assumes get_ancestors_ptrs_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors \<Longrightarrow> ptr' \<in> set ancestors
\<Longrightarrow> ptr' |\<in>| object_ptr_kinds h"
assumes get_ancestors_remains_not_in_ancestors:
"heap_is_wellformed h \<Longrightarrow> heap_is_wellformed h' \<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors
\<Longrightarrow> h' \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors'
\<Longrightarrow> (\<And>p children children'. h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children
\<Longrightarrow> h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children'
\<Longrightarrow> set children' \<subseteq> set children)
\<Longrightarrow> node \<notin> set ancestors
\<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> type_wf h' \<Longrightarrow> node \<notin> set ancestors'"
assumes get_ancestors_also_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_ancestors some_ptr \<rightarrow>\<^sub>r ancestors
\<Longrightarrow> cast child_node \<in> set ancestors
\<Longrightarrow> h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> type_wf h
\<Longrightarrow> known_ptrs h \<Longrightarrow> parent \<in> set ancestors"
assumes get_ancestors_obtains_children:
"heap_is_wellformed h \<Longrightarrow> ancestor \<noteq> ptr \<Longrightarrow> ancestor \<in> set ancestors
\<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> (\<And>children ancestor_child . h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children
\<Longrightarrow> ancestor_child \<in> set children
\<Longrightarrow> cast ancestor_child \<in> set ancestors
\<Longrightarrow> thesis)
\<Longrightarrow> thesis"
assumes get_ancestors_parent_child_rel:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_ancestors child \<rightarrow>\<^sub>r ancestors \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> (ptr, child) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> ptr \<in> set ancestors"
locale l_get_root_node_wf = l_heap_is_wellformed_defs + l_get_root_node_defs + l_type_wf
+ l_known_ptrs + l_get_ancestors_defs + l_get_parent_defs +
assumes get_root_node_ok:
"heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h
\<Longrightarrow> h \<turnstile> ok (get_root_node ptr)"
assumes get_root_node_ptr_in_heap:
"h \<turnstile> ok (get_root_node ptr) \<Longrightarrow> ptr |\<in>| object_ptr_kinds h"
assumes get_root_node_root_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow> root |\<in>| object_ptr_kinds h"
assumes get_ancestors_same_root_node:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors \<Longrightarrow> ptr' \<in> set ancestors
\<Longrightarrow> ptr'' \<in> set ancestors
\<Longrightarrow> h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr \<longleftrightarrow> h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
assumes get_root_node_same_no_parent:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r cast child \<Longrightarrow> h \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
assumes get_root_node_parent_same:
"h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some ptr
\<Longrightarrow> h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root \<longleftrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
interpretation i_get_root_node_wf?:
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs
using instances
by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]:
"l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors
get_ancestors_locs get_child_nodes get_parent"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def)[1]
using get_ancestors_never_empty apply blast
using get_ancestors_ok apply blast
using get_ancestors_reads apply blast
using get_ancestors_ptrs_in_heap apply blast
using get_ancestors_remains_not_in_ancestors apply blast
using get_ancestors_also_parent apply blast
using get_ancestors_obtains_children apply blast
using get_ancestors_parent_child_rel apply blast
using get_ancestors_parent_child_rel apply blast
done
lemma get_root_node_wf_is_l_get_root_node_wf [instances]:
"l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs
get_ancestors get_parent"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1]
using get_root_node_ok apply blast
using get_root_node_ptr_in_heap apply blast
using get_root_node_root_in_heap apply blast
using get_ancestors_same_root_node apply(blast, blast)
using get_root_node_same_no_parent apply blast
using get_root_node_parent_same apply (blast, blast)
done
subsection \<open>to\_tree\_order\<close>
locale l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent +
l_get_parent_wf +
l_heap_is_wellformed
(* l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *)
begin
lemma to_tree_order_ptr_in_heap:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> ok (to_tree_order ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct)
case (step parent)
then show ?case
apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_is_OK_E3)[1]
using get_child_nodes_ptr_in_heap by blast
qed
lemma to_tree_order_either_ptr_or_in_children:
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
and "node \<in> set nodes"
and "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and "node \<noteq> ptr"
obtains child child_to where "child \<in> set children"
and "h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r child_to" and "node \<in> set child_to"
proof -
obtain treeorders where treeorders: "h \<turnstile> map_M to_tree_order (map cast children) \<rightarrow>\<^sub>r treeorders"
using assms
apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1]
using pure_returns_heap_eq returns_result_eq by fastforce
then have "node \<in> set (concat treeorders)"
using assms[simplified to_tree_order_def]
by(auto elim!: bind_returns_result_E4 dest: pure_returns_heap_eq)
then obtain treeorder where "treeorder \<in> set treeorders"
and node_in_treeorder: "node \<in> set treeorder"
by auto
then obtain child where "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r treeorder"
and "child \<in> set children"
using assms[simplified to_tree_order_def] treeorders
by(auto elim!: map_M_pure_E2)
then show ?thesis
using node_in_treeorder returns_result_eq that by auto
qed
lemma to_tree_order_ptrs_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "ptr' \<in> set to"
shows "ptr' |\<in>| object_ptr_kinds h"
proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wellformed_induct)
case (step parent)
have "parent |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) step.prems(1) to_tree_order_ptr_in_heap by blast
then obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then have "to = [parent]"
using step(2) children
apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_returns_result_E2)[1]
by (metis list.distinct(1) list.map_disc_iff list.set_cases map_M_pure_E2 returns_result_eq)
then show ?thesis
using \<open>parent |\<in>| object_ptr_kinds h\<close> step.prems(2) by auto
next
case False
note f = this
then show ?thesis
using children step to_tree_order_either_ptr_or_in_children
proof (cases "ptr' = parent")
case True
then show ?thesis
using \<open>parent |\<in>| object_ptr_kinds h\<close> by blast
next
case False
then show ?thesis
using children step.hyps to_tree_order_either_ptr_or_in_children
by (metis step.prems(1) step.prems(2))
qed
qed
qed
lemma to_tree_order_ok:
assumes wellformed: "heap_is_wellformed h"
and "ptr |\<in>| object_ptr_kinds h"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "h \<turnstile> ok (to_tree_order ptr)"
proof(insert assms(1) assms(2), induct rule: heap_wellformed_induct)
case (step parent)
then show ?case
using assms(3) type_wf
apply(simp add: to_tree_order_def)
apply(auto simp add: heap_is_wellformed_def intro!: map_M_ok_I bind_is_OK_pure_I map_M_pure_I)[1]
using get_child_nodes_ok known_ptrs_known_ptr apply blast
by (simp add: local.heap_is_wellformed_children_in_heap local.to_tree_order_def wellformed)
qed
lemma to_tree_order_child_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
and "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and "node \<in> set children"
and "h \<turnstile> to_tree_order (cast node) \<rightarrow>\<^sub>r nodes'"
shows "set nodes' \<subseteq> set nodes"
proof
fix x
assume a1: "x \<in> set nodes'"
moreover obtain treeorders
where treeorders: "h \<turnstile> map_M to_tree_order (map cast children) \<rightarrow>\<^sub>r treeorders"
using assms(2) assms(3)
apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1]
using pure_returns_heap_eq returns_result_eq by fastforce
then have "nodes' \<in> set treeorders"
using assms(4) assms(5)
by(auto elim!: map_M_pure_E dest: returns_result_eq)
moreover have "set (concat treeorders) \<subseteq> set nodes"
using treeorders assms(2) assms(3)
by(auto simp add: to_tree_order_def elim!: bind_returns_result_E4 dest: pure_returns_heap_eq)
ultimately show "x \<in> set nodes"
by auto
qed
lemma to_tree_order_ptr_in_result:
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
shows "ptr \<in> set nodes"
using assms
apply(simp add: to_tree_order_def)
by(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I bind_pure_I)
lemma to_tree_order_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
and "node \<in> set nodes"
and "h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "set nodes' \<subseteq> set nodes"
proof -
have "\<forall>nodes. h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes \<longrightarrow> (\<forall>node. node \<in> set nodes
\<longrightarrow> (\<forall>nodes'. h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes' \<longrightarrow> set nodes' \<subseteq> set nodes))"
proof(insert assms(1), induct ptr rule: heap_wellformed_induct)
case (step parent)
then show ?case
proof safe
fix nodes node nodes' x
assume 1: "(\<And>children child.
h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> \<forall>nodes. h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes
\<longrightarrow> (\<forall>node. node \<in> set nodes \<longrightarrow> (\<forall>nodes'. h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'
\<longrightarrow> set nodes' \<subseteq> set nodes)))"
and 2: "h \<turnstile> to_tree_order parent \<rightarrow>\<^sub>r nodes"
and 3: "node \<in> set nodes"
and "h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'"
and "x \<in> set nodes'"
have h1: "(\<And>children child nodes node nodes'.
h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes
\<longrightarrow> (node \<in> set nodes \<longrightarrow> (h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes' \<longrightarrow> set nodes' \<subseteq> set nodes)))"
using 1
by blast
obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
using 2
by(auto simp add: to_tree_order_def elim!: bind_returns_result_E)
then have "set nodes' \<subseteq> set nodes"
proof (cases "children = []")
case True
then show ?thesis
by (metis "2" "3" \<open>h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'\<close> children empty_iff list.set(1)
subsetI to_tree_order_either_ptr_or_in_children)
next
case False
then show ?thesis
proof (cases "node = parent")
case True
then show ?thesis
using "2" \<open>h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'\<close> returns_result_eq by fastforce
next
case False
then obtain child nodes_of_child where
"child \<in> set children" and
"h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes_of_child" and
"node \<in> set nodes_of_child"
using 2[simplified to_tree_order_def] 3
to_tree_order_either_ptr_or_in_children[where node=node and ptr=parent] children
apply(auto elim!: bind_returns_result_E2 intro: map_M_pure_I)[1]
using is_OK_returns_result_E 2 a_all_ptrs_in_heap_def assms(1) heap_is_wellformed_def
using "3" by blast
then have "set nodes' \<subseteq> set nodes_of_child"
using h1
using \<open>h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes'\<close> children by blast
moreover have "set nodes_of_child \<subseteq> set nodes"
using "2" \<open>child \<in> set children\<close> \<open>h \<turnstile> to_tree_order (cast child) \<rightarrow>\<^sub>r nodes_of_child\<close>
assms children to_tree_order_child_subset by auto
ultimately show ?thesis
by blast
qed
qed
then show "x \<in> set nodes"
using \<open>x \<in> set nodes'\<close> by blast
qed
qed
then show ?thesis
using assms by blast
qed
lemma to_tree_order_parent:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
assumes "parent \<in> set nodes"
shows "cast child \<in> set nodes"
proof -
obtain nodes' where nodes': "h \<turnstile> to_tree_order parent \<rightarrow>\<^sub>r nodes'"
using assms to_tree_order_ok get_parent_parent_in_heap
by (meson get_parent_parent_in_heap is_OK_returns_result_E)
then have "set nodes' \<subseteq> set nodes"
using to_tree_order_subset assms
by blast
moreover obtain children where
children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children" and
child: "child \<in> set children"
using assms get_parent_child_dual by blast
then obtain child_to where child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r child_to"
by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E is_OK_returns_result_I
get_parent_ptr_in_heap node_ptr_kinds_commutes to_tree_order_ok)
then have "cast child \<in> set child_to"
apply(simp add: to_tree_order_def)
by(auto elim!: bind_returns_result_E2 map_M_pure_E
dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I)
have "cast child \<in> set nodes'"
using nodes' child
apply(simp add: to_tree_order_def)
apply(auto elim!: bind_returns_result_E2 map_M_pure_E
dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I)[1]
using child_to \<open>cast child \<in> set child_to\<close> returns_result_eq by fastforce
ultimately show ?thesis
by auto
qed
lemma to_tree_order_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
assumes "cast child \<noteq> ptr"
assumes "child \<in> set children"
assumes "cast child \<in> set nodes"
shows "parent \<in> set nodes"
proof(insert assms(1) assms(4) assms(6) assms(8), induct ptr arbitrary: nodes
rule: heap_wellformed_induct)
case (step p)
have "p |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> to_tree_order p \<rightarrow>\<^sub>r nodes\<close> to_tree_order_ptr_in_heap
using assms(1) assms(2) assms(3) by blast
then obtain children where children: "h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then show ?thesis
using step(2) step(3) step(4) children
by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])
next
case False
then obtain c child_to where
child: "c \<in> set children" and
child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<rightarrow>\<^sub>r child_to" and
"cast child \<in> set child_to"
using step(2) children
apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
by (metis (full_types) assms(1) assms(2) assms(3) get_parent_ptr_in_heap
is_OK_returns_result_I l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_kinds_commutes
returns_result_select_result step.prems(1) step.prems(2) step.prems(3)
to_tree_order_either_ptr_or_in_children to_tree_order_ok)
then have "set child_to \<subseteq> set nodes"
using assms(1) child children step.prems(1) to_tree_order_child_subset by auto
show ?thesis
proof (cases "c = child")
case True
then have "parent = p"
using step(3) children child assms(5) assms(7)
by (meson assms(1) assms(2) assms(3) child_parent_dual option.inject returns_result_eq)
then show ?thesis
using step.prems(1) to_tree_order_ptr_in_result by blast
next
case False
then show ?thesis
using step(1)[OF children child child_to] step(3) step(4)
using \<open>set child_to \<subseteq> set nodes\<close>
using \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \<in> set child_to\<close> by auto
qed
qed
qed
lemma to_tree_order_node_ptrs:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "ptr' \<noteq> ptr"
assumes "ptr' \<in> set nodes"
shows "is_node_ptr_kind ptr'"
proof(insert assms(1) assms(4) assms(5) assms(6), induct ptr arbitrary: nodes
rule: heap_wellformed_induct)
case (step p)
have "p |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> to_tree_order p \<rightarrow>\<^sub>r nodes\<close> to_tree_order_ptr_in_heap
using assms(1) assms(2) assms(3) by blast
then obtain children where children: "h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then show ?thesis
using step(2) step(3) step(4) children
by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
next
case False
then obtain c child_to where
child: "c \<in> set children" and
child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<rightarrow>\<^sub>r child_to" and
"ptr' \<in> set child_to"
using step(2) children
apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children by blast
then have "set child_to \<subseteq> set nodes"
using assms(1) child children step.prems(1) to_tree_order_child_subset by auto
show ?thesis
proof (cases "cast c = ptr")
case True
then show ?thesis
using step \<open>ptr' \<in> set child_to\<close> assms(5) child child_to children by blast
next
case False
then show ?thesis
using \<open>ptr' \<in> set child_to\<close> child child_to children is_node_ptr_kind_cast step.hyps by blast
qed
qed
qed
lemma to_tree_order_child2:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "cast child \<noteq> ptr"
assumes "cast child \<in> set nodes"
obtains parent where "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent" and "parent \<in> set nodes"
proof -
assume 1: "(\<And>parent. h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent \<Longrightarrow> parent \<in> set nodes \<Longrightarrow> thesis)"
show thesis
proof(insert assms(1) assms(4) assms(5) assms(6) 1, induct ptr arbitrary: nodes
rule: heap_wellformed_induct)
case (step p)
have "p |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> to_tree_order p \<rightarrow>\<^sub>r nodes\<close> to_tree_order_ptr_in_heap
using assms(1) assms(2) assms(3) by blast
then obtain children where children: "h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children"
by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
then show ?case
proof (cases "children = []")
case True
then show ?thesis
using step(2) step(3) step(4) children
by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])
next
case False
then obtain c child_to where
child: "c \<in> set children" and
child_to: "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \<rightarrow>\<^sub>r child_to" and
"cast child \<in> set child_to"
using step(2) children
apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children
by blast
then have "set child_to \<subseteq> set nodes"
using assms(1) child children step.prems(1) to_tree_order_child_subset by auto
have "cast child |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) assms(4) assms(6) to_tree_order_ptrs_in_heap by blast
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt"
by (meson assms(2) assms(3) is_OK_returns_result_E get_parent_ok node_ptr_kinds_commutes)
then show ?thesis
proof (induct parent_opt)
case None
then show ?case
by (metis \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \<in> set child_to\<close> assms(1) assms(2) assms(3)
cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject child child_parent_dual child_to children
option.distinct(1) returns_result_eq step.hyps)
next
case (Some option)
then show ?case
by (meson assms(1) assms(2) assms(3) get_parent_child_dual step.prems(1) step.prems(2)
step.prems(3) step.prems(4) to_tree_order_child)
qed
qed
qed
qed
lemma to_tree_order_parent_child_rel:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
shows "(ptr, child) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> child \<in> set to"
proof
assume 3: "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
show "child \<in> set to"
proof (insert 3, induct child rule: heap_wellformed_induct_rev[OF assms(1)])
case (1 child)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
using assms(4)
apply(simp add: to_tree_order_def)
by(auto simp add: map_M_pure_I elim!: bind_returns_result_E2)
next
case False
obtain child_parent where
"(ptr, child_parent) \<in> (parent_child_rel h)\<^sup>*" and
"(child_parent, child) \<in> (parent_child_rel h)"
using \<open>ptr \<noteq> child\<close>
by (metis "1.prems" rtranclE)
obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child"
using \<open>(child_parent, child) \<in> parent_child_rel h\<close> node_ptr_casts_commute3
parent_child_rel_node_ptr
by blast
then have "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some child_parent"
using \<open>(child_parent, child) \<in> (parent_child_rel h)\<close>
by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E l_get_parent_wf.child_parent_dual
l_heap_is_wellformed.parent_child_rel_child local.get_child_nodes_ok
local.known_ptrs_known_ptr local.l_get_parent_wf_axioms
local.l_heap_is_wellformed_axioms local.parent_child_rel_parent_in_heap)
then show ?thesis
using 1(1) child_node \<open>(ptr, child_parent) \<in> (parent_child_rel h)\<^sup>*\<close>
using assms(1) assms(2) assms(3) assms(4) to_tree_order_parent by blast
qed
qed
next
assume "child \<in> set to"
then show "(ptr, child) \<in> (parent_child_rel h)\<^sup>*"
proof (induct child rule: heap_wellformed_induct_rev[OF assms(1)])
case (1 child)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by simp
next
case False
then have "\<exists>parent. (parent, child) \<in> (parent_child_rel h)"
using 1(2) assms(4) to_tree_order_child2[OF assms(1) assms(2) assms(3) assms(4)]
to_tree_order_node_ptrs
by (metis assms(1) assms(2) assms(3) node_ptr_casts_commute3 parent_child_rel_parent)
then obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child"
using node_ptr_casts_commute3 parent_child_rel_node_ptr by blast
then obtain child_parent where child_parent: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some child_parent"
using \<open>\<exists>parent. (parent, child) \<in> (parent_child_rel h)\<close>
by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) to_tree_order_child2)
then have "(child_parent, child) \<in> (parent_child_rel h)"
using assms(1) child_node parent_child_rel_parent by blast
moreover have "child_parent \<in> set to"
by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) child_node child_parent
get_parent_child_dual to_tree_order_child)
then have "(ptr, child_parent) \<in> (parent_child_rel h)\<^sup>*"
using 1 child_node child_parent by blast
ultimately show ?thesis
by auto
qed
qed
qed
end
interpretation i_to_tree_order_wf?: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs to_tree_order known_ptrs get_parent
get_parent_locs heap_is_wellformed parent_child_rel
get_disconnected_nodes get_disconnected_nodes_locs
using instances
apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
done
declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_to_tree_order_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
+ l_to_tree_order_defs
+ l_get_parent_defs + l_get_child_nodes_defs +
assumes to_tree_order_ok:
"heap_is_wellformed h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> h \<turnstile> ok (to_tree_order ptr)"
assumes to_tree_order_ptrs_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to
\<Longrightarrow> ptr' \<in> set to \<Longrightarrow> ptr' |\<in>| object_ptr_kinds h"
assumes to_tree_order_parent_child_rel:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to
\<Longrightarrow> (ptr, child_ptr) \<in> (parent_child_rel h)\<^sup>* \<longleftrightarrow> child_ptr \<in> set to"
assumes to_tree_order_child2:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> cast child \<noteq> ptr \<Longrightarrow> cast child \<in> set nodes
\<Longrightarrow> (\<And>parent. h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent
\<Longrightarrow> parent \<in> set nodes \<Longrightarrow> thesis)
\<Longrightarrow> thesis"
assumes to_tree_order_node_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> ptr' \<noteq> ptr \<Longrightarrow> ptr' \<in> set nodes \<Longrightarrow> is_node_ptr_kind ptr'"
assumes to_tree_order_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow> cast child \<noteq> ptr
\<Longrightarrow> child \<in> set children \<Longrightarrow> cast child \<in> set nodes
\<Longrightarrow> parent \<in> set nodes"
assumes to_tree_order_ptr_in_result:
"h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes \<Longrightarrow> ptr \<in> set nodes"
assumes to_tree_order_parent:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes
\<Longrightarrow> h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent \<Longrightarrow> parent \<in> set nodes
\<Longrightarrow> cast child \<in> set nodes"
assumes to_tree_order_subset:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes \<Longrightarrow> node \<in> set nodes
\<Longrightarrow> h \<turnstile> to_tree_order node \<rightarrow>\<^sub>r nodes' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> set nodes' \<subseteq> set nodes"
lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]:
"l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
to_tree_order get_parent get_child_nodes"
using instances
apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def)[1]
using to_tree_order_ok
apply blast
using to_tree_order_ptrs_in_heap
apply blast
using to_tree_order_parent_child_rel
apply(blast, blast)
using to_tree_order_child2
apply blast
using to_tree_order_node_ptrs
apply blast
using to_tree_order_child
apply blast
using to_tree_order_ptr_in_result
apply blast
using to_tree_order_parent
apply blast
using to_tree_order_subset
apply blast
done
subsubsection \<open>get\_root\_node\<close>
locale l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
+ l_to_tree_order_wf
begin
lemma to_tree_order_get_root_node:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "ptr' \<in> set to"
assumes "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
assumes "ptr'' \<in> set to"
shows "h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
proof -
obtain ancestors' where ancestors': "h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'"
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_ok is_OK_returns_result_E
to_tree_order_ptrs_in_heap )
moreover have "ptr \<in> set ancestors'"
using \<open>h \<turnstile> get_ancestors ptr' \<rightarrow>\<^sub>r ancestors'\<close>
using assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_parent_child_rel
to_tree_order_parent_child_rel by blast
ultimately have "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
using \<open>h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr\<close>
using assms(1) assms(2) assms(3) get_ancestors_ptr get_ancestors_same_root_node by blast
obtain ancestors'' where ancestors'': "h \<turnstile> get_ancestors ptr'' \<rightarrow>\<^sub>r ancestors''"
by (meson assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_ok is_OK_returns_result_E
to_tree_order_ptrs_in_heap)
moreover have "ptr \<in> set ancestors''"
using \<open>h \<turnstile> get_ancestors ptr'' \<rightarrow>\<^sub>r ancestors''\<close>
using assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_parent_child_rel
to_tree_order_parent_child_rel by blast
ultimately show ?thesis
using \<open>h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr\<close> assms(1) assms(2) assms(3) get_ancestors_ptr
get_ancestors_same_root_node by blast
qed
lemma to_tree_order_same_root:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
assumes "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r to"
assumes "ptr' \<in> set to"
shows "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
proof (insert assms(1)(* assms(4) assms(5) *) assms(6), induct ptr' rule: heap_wellformed_induct_rev)
case (step child)
then show ?case
proof (cases "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child")
case True
then have "child = root_ptr"
using assms(1) assms(2) assms(3) assms(5) step.prems
by (metis (no_types, lifting) get_root_node_same_no_parent node_ptr_casts_commute3
option.simps(3) returns_result_eq to_tree_order_child2 to_tree_order_node_ptrs)
then show ?thesis
using True by blast
next
case False
then obtain child_node parent where "cast child_node = child"
and "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent"
by (metis assms(1) assms(2) assms(3) assms(4) assms(5) local.get_root_node_no_parent
local.get_root_node_not_node_same local.get_root_node_same_no_parent
local.to_tree_order_child2 local.to_tree_order_ptrs_in_heap node_ptr_casts_commute3
step.prems)
then show ?thesis
proof (cases "child = root_ptr")
case True
then have "h \<turnstile> get_root_node root_ptr \<rightarrow>\<^sub>r root_ptr"
using assms(4)
using \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\<close> assms(1) assms(2) assms(3)
get_root_node_no_parent get_root_node_same_no_parent
by blast
then show ?thesis
using step assms(4)
using True by blast
next
case False
then have "parent \<in> set to"
using assms(5) step(2) to_tree_order_child \<open>h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent\<close>
\<open>cast child_node = child\<close>
by (metis False assms(1) assms(2) assms(3) get_parent_child_dual)
then show ?thesis
using \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\<close> \<open>h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent\<close>
get_root_node_parent_same
using step.hyps by blast
qed
qed
qed
end
interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors
get_ancestors_locs get_root_node get_root_node_locs to_tree_order
using instances
by(simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
locale l_to_tree_order_wf_get_root_node_wf = l_type_wf + l_known_ptrs + l_to_tree_order_defs
+ l_get_root_node_defs + l_heap_is_wellformed_defs +
assumes to_tree_order_get_root_node:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to
\<Longrightarrow> ptr' \<in> set to \<Longrightarrow> h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr
\<Longrightarrow> ptr'' \<in> set to \<Longrightarrow> h \<turnstile> get_root_node ptr'' \<rightarrow>\<^sub>r root_ptr"
assumes to_tree_order_same_root:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr
\<Longrightarrow> h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r to \<Longrightarrow> ptr' \<in> set to
\<Longrightarrow> h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]:
"l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order
get_root_node heap_is_wellformed"
using instances
apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def
l_to_tree_order_wf_get_root_node_wf_axioms_def)[1]
using to_tree_order_get_root_node apply blast
using to_tree_order_same_root apply blast
done
subsection \<open>get\_owner\_document\<close>
locale l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_known_ptrs
+ l_heap_is_wellformed
+ l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
+ l_get_ancestors
+ l_get_ancestors_wf
+ l_get_parent
+ l_get_parent_wf
+ l_get_root_node_wf
+ l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_owner_document_disconnected_nodes:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "node_ptr \<in> set disc_nodes"
assumes known_ptrs: "known_ptrs h"
assumes type_wf: "type_wf h"
shows "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r document_ptr"
proof -
have 2: "node_ptr |\<in>| node_ptr_kinds h"
using assms heap_is_wellformed_disc_nodes_in_heap
by blast
have 3: "document_ptr |\<in>| document_ptr_kinds h"
using assms(2) get_disconnected_nodes_ptr_in_heap by blast
have 0:
"\<exists>!document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r. node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
by (metis (no_types, lifting) "3" DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(2) assms(3)
disjoint_iff_not_equal l_heap_is_wellformed.heap_is_wellformed_one_disc_parent
local.get_disconnected_nodes_ok local.l_heap_is_wellformed_axioms
returns_result_select_result select_result_I2 type_wf)
have "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
using heap_is_wellformed_children_disc_nodes_different child_parent_dual assms
using "2" disjoint_iff_not_equal local.get_parent_child_dual local.get_parent_ok
returns_result_select_result split_option_ex
by (metis (no_types, lifting))
then have 4: "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
using 2 get_root_node_no_parent
by blast
obtain document_ptrs where document_ptrs: "h \<turnstile> document_ptr_kinds_M \<rightarrow>\<^sub>r document_ptrs"
by simp
then
have "h \<turnstile> ok (filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs)"
using assms(1) get_disconnected_nodes_ok type_wf unfolding heap_is_wellformed_def
by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I)
then obtain candidates where
candidates: "h \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^sub>r candidates"
by auto
have eq: "\<And>document_ptr. document_ptr |\<in>| document_ptr_kinds h
\<Longrightarrow> node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r \<longleftrightarrow> |h \<turnstile> do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}|\<^sub>r"
apply(auto dest!: get_disconnected_nodes_ok[OF type_wf]
intro!: select_result_I[where P=id, simplified] elim!: bind_returns_result_E2)[1]
apply(drule select_result_E[where P=id, simplified])
by(auto elim!: bind_returns_result_E2)
have filter: "filter (\<lambda>document_ptr. |h \<turnstile> do {
disconnected_nodes \<leftarrow> 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 \<in> cast ` set disconnected_nodes)
}|\<^sub>r) document_ptrs = [document_ptr]"
apply(rule filter_ex1)
using 0 document_ptrs apply(simp)[1]
using eq
using local.get_disconnected_nodes_ok apply auto[1]
using assms(2) assms(3)
apply(auto intro!: intro!: select_result_I[where P=id, simplified]
elim!: bind_returns_result_E2)[1]
using returns_result_eq apply fastforce
using document_ptrs 3 apply(simp)
using document_ptrs
by simp
have "h \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^sub>r [document_ptr]"
apply(rule filter_M_filter2)
using get_disconnected_nodes_ok document_ptrs 3 assms(1) type_wf filter
unfolding heap_is_wellformed_def
by(auto intro: bind_pure_I bind_is_OK_I2)
with 4 document_ptrs have "h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r document_ptr"
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I
split: option.splits)[1]
moreover have "known_ptr (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)"
using "4" assms(1) known_ptrs type_wf known_ptrs_known_ptr "2" node_ptr_kinds_commutes by blast
ultimately show ?thesis
using 2
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
by(auto split: option.splits intro!: bind_pure_returns_result_I)
qed
lemma in_disconnected_nodes_no_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
and "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document"
and "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
have 2: "cast node_ptr |\<in>| object_ptr_kinds h"
using assms(3) get_owner_document_ptr_in_heap by fast
then have 3: "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
using assms(2) local.get_root_node_no_parent by blast
have "\<not>(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
apply(auto)[1]
using assms(2) child_parent_dual[OF assms(1)] type_wf
assms(1) assms(5) get_child_nodes_ok known_ptrs_known_ptr option.simps(3)
returns_result_eq returns_result_select_result
by (metis (no_types, opaque_lifting))
moreover have "node_ptr |\<in>| node_ptr_kinds h"
using assms(2) get_parent_ptr_in_heap by blast
ultimately
have 0: "\<exists>document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r. node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
- by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) finite_set_in heap_is_wellformed_children_disc_nodes)
+ by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) heap_is_wellformed_children_disc_nodes)
then obtain document_ptr where
document_ptr: "document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r" and
node_ptr_in_disc_nodes: "node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
by auto
then show ?thesis
using get_owner_document_disconnected_nodes known_ptrs type_wf assms
using DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) assms(4) get_disconnected_nodes_ok
returns_result_select_result select_result_I2
by (metis (no_types, opaque_lifting) )
qed
lemma get_owner_document_owner_document_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
shows "owner_document |\<in>| document_ptr_kinds h"
using assms
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_split_asm)+
proof -
assume "h \<turnstile> invoke [] ptr () \<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| document_ptr_kinds h"
by (meson invoke_empty is_OK_returns_result_I)
next
assume "h \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ())
\<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| document_ptr_kinds h"
by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits)
next
assume 0: "heap_is_wellformed h"
and 1: "type_wf h"
and 2: "known_ptrs h"
and 3: "\<not> 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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^sub>r owner_document"
then obtain root where
root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
split: option.splits)
then show ?thesis
proof (cases "is_document_ptr root")
case True
then show ?thesis
using 4 5 root
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply(drule(1) returns_result_eq) apply(auto)[1]
using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
next
case False
have "known_ptr root"
using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast
have "root |\<in>| object_ptr_kinds h"
using root
using "0" "1" "2" local.get_root_node_root_in_heap
by blast
then have "is_node_ptr_kind root"
using False \<open>known_ptr root\<close>
apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs)
using is_node_ptr_kind_none by force
then
have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
- by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close> local.child_parent_dual
+ by (metis (no_types, opaque_lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close> local.child_parent_dual
local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes
local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes
- notin_fset option.distinct(1) returns_result_eq returns_result_select_result root)
+ option.distinct(1) returns_result_eq returns_result_select_result root)
then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
then
obtain candidates where
candidates: "h \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (root \<in> 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)))
\<rightarrow>\<^sub>r candidates"
by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset
- is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset
+ is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure
return_ok return_pure sorted_list_of_set(1))
then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
apply (simp add: \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>)
using "1" \<open>root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
\<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
local.get_disconnected_nodes_ok by auto
then have "candidates \<noteq> []"
by auto
then have "owner_document \<in> set candidates"
using 5 root 4
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply (metis candidates list.set_sel(1) returns_result_eq)
by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq)
then show ?thesis
using candidates
by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I
local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure)
qed
next
assume 0: "heap_is_wellformed h"
and 1: "type_wf h"
and 2: "known_ptrs h"
and 3: "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
and 4: "h \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr) (\<lambda>_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^sub>r owner_document"
then obtain root where
root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits)
then show ?thesis
proof (cases "is_document_ptr root")
case True
then show ?thesis
using 3 4 root
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply(drule(1) returns_result_eq) apply(auto)[1]
using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
next
case False
have "known_ptr root"
using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast
have "root |\<in>| object_ptr_kinds h"
using root
using "0" "1" "2" local.get_root_node_root_in_heap
by blast
then have "is_node_ptr_kind root"
using False \<open>known_ptr root\<close>
apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs)
using is_node_ptr_kind_none by force
then
have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
- by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close>
+ by (metis (no_types, opaque_lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close>
local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent
local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3
- node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq
+ node_ptr_inclusion node_ptr_kinds_commutes option.distinct(1) returns_result_eq
returns_result_select_result root)
then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
then
obtain candidates where
candidates: "h \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (root \<in> 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)))
\<rightarrow>\<^sub>r candidates"
by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset
- is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset
+ is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure
return_ok return_pure sorted_list_of_set(1))
then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
apply (simp add: \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>)
using "1" \<open>root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
\<open>some_owner_document |\<in>| document_ptr_kinds h\<close> local.get_disconnected_nodes_ok
by auto
then have "candidates \<noteq> []"
by auto
then have "owner_document \<in> set candidates"
using 4 root 3
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply (metis candidates list.set_sel(1) returns_result_eq)
by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq)
then show ?thesis
using candidates
by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I
local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure)
qed
qed
lemma get_owner_document_ok:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_owner_document ptr)"
proof -
have "known_ptr ptr"
using assms(2) assms(4) local.known_ptrs_known_ptr
by blast
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(auto simp add: known_ptr_impl)[1]
using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr
known_ptr_not_element_ptr
apply blast
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply (metis (no_types, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes
is_document_ptr_kind_none option.case_eq_if)
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply (metis (no_types, lifting) assms(1) assms(2) assms(3) is_node_ptr_kind_none
local.get_root_node_ok node_ptr_casts_commute3 option.case_eq_if)
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
filter_M_is_OK_I)[1]
using assms(3) local.get_disconnected_nodes_ok
apply blast
apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)
using assms(4)
apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I)[1]
apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
filter_M_is_OK_I)[1]
apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)[1]
apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
filter_M_is_OK_I)[1]
using assms(3) local.get_disconnected_nodes_ok by blast
qed
lemma get_owner_document_child_same:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document"
proof -
have "ptr |\<in>| object_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_child_nodes_ptr_in_heap)
then have "known_ptr ptr"
using assms(2) local.known_ptrs_known_ptr by blast
have "cast child |\<in>| object_ptr_kinds h"
using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes
by blast
then
have "known_ptr (cast child)"
using assms(2) local.known_ptrs_known_ptr by blast
obtain root where root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.get_root_node_ok)
then have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root"
using assms(1) assms(2) assms(3) assms(4) assms(5) local.child_parent_dual
local.get_root_node_parent_same
by blast
have "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr ptr")
case True
then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = ptr"
using case_optionE document_ptr_casts_commute by blast
then have "root = cast document_ptr"
using root
by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
split: option.splits)
then have "h \<turnstile> a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using document_ptr
\<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> document_ptr]
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated,
OF \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> document_ptr], rotated]
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1]
using \<open>ptr |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes by blast
then show ?thesis
using \<open>known_ptr ptr\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>ptr |\<in>| object_ptr_kinds h\<close> True
by(auto simp add: document_ptr[symmetric] intro!: bind_pure_returns_result_I
split: option.splits)
next
case False
then obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = ptr"
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then have "h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using root \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>
unfolding a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure)
then show ?thesis
using \<open>known_ptr ptr\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
apply (meson invoke_empty is_OK_returns_result_I)
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
by(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
qed
then show ?thesis
using \<open>known_ptr (cast child)\<close>
apply(auto simp add: get_owner_document_def[of "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child"]
a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
by (smt \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\<in>| object_ptr_kinds h\<close> cast_document_ptr_not_node_ptr(1)
comp_apply invoke_empty invoke_not invoke_returns_result is_OK_returns_result_I
node_ptr_casts_commute2 option.sel)
qed
end
locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
+ l_get_disconnected_nodes_defs + l_get_owner_document_defs
+ l_get_parent_defs + l_get_child_nodes_defs +
assumes get_owner_document_disconnected_nodes:
"heap_is_wellformed h \<Longrightarrow>
known_ptrs h \<Longrightarrow>
type_wf h \<Longrightarrow>
h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
node_ptr \<in> set disc_nodes \<Longrightarrow>
h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r document_ptr"
assumes in_disconnected_nodes_no_parent:
"heap_is_wellformed h \<Longrightarrow>
h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None\<Longrightarrow>
h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document \<Longrightarrow>
h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
known_ptrs h \<Longrightarrow>
type_wf h\<Longrightarrow>
node_ptr \<in> set disc_nodes"
assumes get_owner_document_owner_document_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow>
h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<Longrightarrow>
owner_document |\<in>| document_ptr_kinds h"
assumes get_owner_document_ok:
"heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h
\<Longrightarrow> h \<turnstile> ok (get_owner_document ptr)"
assumes get_owner_document_child_same:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document"
interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr known_ptrs type_wf heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors
get_ancestors_locs get_root_node get_root_node_locs get_owner_document
by(auto simp add: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]:
"l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes
get_owner_document get_parent get_child_nodes"
using known_ptrs_is_l_known_ptrs
apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def)[1]
using get_owner_document_disconnected_nodes apply fast
using in_disconnected_nodes_no_parent apply fast
using get_owner_document_owner_document_in_heap apply fast
using get_owner_document_ok apply fast
using get_owner_document_child_same apply (fast, fast)
done
subsubsection \<open>get\_root\_node\<close>
locale l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node_wf +
l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document_wf
begin
lemma get_root_node_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
assumes "is_document_ptr_kind root"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r the (cast root)"
proof -
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
by (meson is_OK_returns_result_I local.get_root_node_ptr_in_heap)
then have "known_ptr ptr"
using assms(3) local.known_ptrs_known_ptr by blast
{
assume "is_document_ptr_kind ptr"
then have "ptr = root"
using assms(4)
by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
split: option.splits)
then have ?thesis
using \<open>is_document_ptr_kind ptr\<close> \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I
split: option.splits)
}
moreover
{
assume "is_node_ptr_kind ptr"
then have ?thesis
using \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
apply(auto split: option.splits)[1]
using \<open>h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root\<close> assms(5)
by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def
intro!: bind_pure_returns_result_I split: option.splits)[2]
}
ultimately
show ?thesis
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
qed
lemma get_root_node_same_owner_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
proof -
have "ptr |\<in>| object_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_root_node_ptr_in_heap)
have "root |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) assms(4) local.get_root_node_root_in_heap by blast
have "known_ptr ptr"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(3) local.known_ptrs_known_ptr by blast
have "known_ptr root"
using \<open>root |\<in>| object_ptr_kinds h\<close> assms(3) local.known_ptrs_known_ptr by blast
show ?thesis
proof (cases "is_document_ptr_kind ptr")
case True
then
have "ptr = root"
using assms(4)
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
by (metis document_ptr_casts_commute3 last_ConsL local.get_ancestors_not_node node_ptr_no_document_ptr_cast)
then show ?thesis
by auto
next
case False
then have "is_node_ptr_kind ptr"
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
by (metis node_ptr_casts_commute3)
show ?thesis
proof
assume "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
using node_ptr
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
by(auto elim!: bind_returns_result_E2 split: option.splits)
show "h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr_kind root")
case True
have "is_document_ptr root"
using True \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
have "root = cast owner_document"
using True
by (smt \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3) assms(4)
document_ptr_casts_commute3 get_root_node_document returns_result_eq)
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\<close> apply blast
using \<open>root |\<in>| object_ptr_kinds h\<close>
by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none)
next
case False
then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> assms(4)
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
apply (metis assms(1) assms(2) assms(3) local.get_root_node_no_parent
local.get_root_node_same_no_parent node_ptr returns_result_eq)
using \<open>is_node_ptr_kind root\<close> node_ptr returns_result_eq by fastforce
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_node_ptr_kind root\<close> \<open>known_ptr root\<close>
apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)[1]
apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)[1]
using \<open>root |\<in>| object_ptr_kinds h\<close>
by(auto simp add: root_node_ptr)
qed
next
assume "h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
show "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr_kind root")
case True
have "root = cast owner_document"
using \<open>h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
apply(auto simp add: True a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
split: if_splits)[1]
apply (metis True cast_document_ptr_not_node_ptr(2) is_document_ptr_kind_obtains
is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
by (metis True cast_document_ptr_not_node_ptr(1) document_ptr_casts_commute3
is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) get_root_node_document
by fastforce
next
case False
then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
using assms(1) assms(2) assms(3) assms(4) local.get_root_node_no_parent
local.get_root_node_same_no_parent node_ptr returns_result_eq root_node_ptr
by fastforce+
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using node_ptr \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
intro!: bind_pure_returns_result_I split: option.splits)
qed
qed
qed
qed
end
interpretation get_owner_document_wf_get_root_node_wf?: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_ancestors get_ancestors_locs get_root_node get_root_node_locs heap_is_wellformed
parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs get_owner_document
by(auto simp add: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_get_owner_document_wf_get_root_node_wf = l_heap_is_wellformed_defs + l_type_wf +
l_known_ptrs + l_get_root_node_defs + l_get_owner_document_defs +
assumes get_root_node_document:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow>
is_document_ptr_kind root \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r the (cast root)"
assumes get_root_node_same_owner_document:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow>
h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]:
"l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs
get_root_node get_owner_document"
apply(auto simp add: l_get_owner_document_wf_get_root_node_wf_def
l_get_owner_document_wf_get_root_node_wf_axioms_def instances)[1]
using get_root_node_document apply blast
using get_root_node_same_owner_document apply (blast, blast)
done
subsection \<open>Preserving heap-wellformedness\<close>
subsection \<open>set\_attribute\<close>
locale l_set_attribute_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_attribute\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_attribute_get_disconnected_nodes +
l_set_attribute_get_child_nodes
begin
lemma set_attribute_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> set_attribute element_ptr k v \<rightarrow>\<^sub>h h'"
shows "heap_is_wellformed h'"
thm preserves_wellformedness_writes_needed
apply(rule preserves_wellformedness_writes_needed[OF assms set_attribute_writes])
using set_attribute_get_child_nodes
apply(fast)
using set_attribute_get_disconnected_nodes apply(fast)
by(auto simp add: all_args_def set_attribute_locs_def)
end
subsection \<open>remove\_child\<close>
locale l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed +
l_set_disconnected_nodes_get_child_nodes
begin
lemma remove_child_removes_parent:
assumes wellformed: "heap_is_wellformed h"
and remove_child: "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h2"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "h2 \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
proof -
obtain children where children: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using remove_child remove_child_def by auto
then have "child \<in> set children"
using remove_child remove_child_def
by(auto elim!: bind_returns_heap_E dest: returns_result_eq split: if_splits)
then have h1: "\<And>other_ptr other_children. other_ptr \<noteq> ptr
\<Longrightarrow> h \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children \<Longrightarrow> child \<notin> set other_children"
using assms(1) known_ptrs type_wf child_parent_dual
by (meson child_parent_dual children option.inject returns_result_eq)
have known_ptr: "known_ptr ptr"
using known_ptrs
by (meson is_OK_returns_heap_I l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms
remove_child remove_child_ptr_in_heap)
obtain owner_document disc_nodes h' where
owner_document: "h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document" and
disc_nodes: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h': "h \<turnstile> set_disconnected_nodes owner_document (child # disc_nodes) \<rightarrow>\<^sub>h h'" and
h2: "h' \<turnstile> set_child_nodes ptr (remove1 child children) \<rightarrow>\<^sub>h h2"
using assms children unfolding remove_child_def
apply(auto split: if_splits elim!: bind_returns_heap_E)[1]
by (metis (full_types) get_child_nodes_pure get_disconnected_nodes_pure
get_owner_document_pure pure_returns_heap_eq returns_result_eq)
have "object_ptr_kinds h = object_ptr_kinds h2"
using remove_child_writes remove_child unfolding remove_child_locs_def
apply(rule writes_small_big)
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by(auto simp add: reflp_def transp_def)
then have "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
unfolding object_ptr_kinds_M_defs by simp
have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF set_disconnected_nodes_writes h']
using set_disconnected_nodes_types_preserved type_wf
by(auto simp add: reflp_def transp_def)
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF remove_child_writes remove_child] unfolding remove_child_locs_def
using set_disconnected_nodes_types_preserved set_child_nodes_types_preserved type_wf
apply(auto simp add: reflp_def transp_def)[1]
by blast
then obtain children' where children': "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'"
using h2 set_child_nodes_get_child_nodes known_ptr
by (metis \<open>object_ptr_kinds h = object_ptr_kinds h2\<close> children get_child_nodes_ok
get_child_nodes_ptr_in_heap is_OK_returns_result_E is_OK_returns_result_I)
have "child \<notin> set children'"
by (metis (mono_tags, lifting) \<open>type_wf h'\<close> children children' distinct_remove1_removeAll h2
known_ptr local.heap_is_wellformed_children_distinct
local.set_child_nodes_get_child_nodes member_remove remove_code(1) select_result_I2
wellformed)
moreover have "\<And>other_ptr other_children. other_ptr \<noteq> ptr
\<Longrightarrow> h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children \<Longrightarrow> child \<notin> set other_children"
proof -
fix other_ptr other_children
assume a1: "other_ptr \<noteq> ptr" and a3: "h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
have "h \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
using get_child_nodes_reads set_disconnected_nodes_writes h' a3
apply(rule reads_writes_separate_backwards)
using set_disconnected_nodes_get_child_nodes by fast
show "child \<notin> set other_children"
using \<open>h \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children\<close> a1 h1 by blast
qed
then have "\<And>other_ptr other_children. other_ptr \<noteq> ptr
\<Longrightarrow> h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children \<Longrightarrow> child \<notin> set other_children"
proof -
fix other_ptr other_children
assume a1: "other_ptr \<noteq> ptr" and a3: "h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
have "h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children"
using get_child_nodes_reads set_child_nodes_writes h2 a3
apply(rule reads_writes_separate_backwards)
using set_disconnected_nodes_get_child_nodes a1 set_child_nodes_get_child_nodes_different_pointers
by metis
then show "child \<notin> set other_children"
using \<open>\<And>other_ptr other_children. \<lbrakk>other_ptr \<noteq> ptr; h' \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children\<rbrakk>
\<Longrightarrow> child \<notin> set other_children\<close> a1 by blast
qed
ultimately have ha: "\<And>other_ptr other_children. h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children
\<Longrightarrow> child \<notin> set other_children"
by (metis (full_types) children' returns_result_eq)
moreover obtain ptrs where ptrs: "h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by (simp add: object_ptr_kinds_M_defs)
moreover have "\<And>ptr. ptr \<in> set ptrs \<Longrightarrow> h2 \<turnstile> ok (get_child_nodes ptr)"
using \<open>type_wf h2\<close> ptrs get_child_nodes_ok known_ptr
using \<open>object_ptr_kinds h = object_ptr_kinds h2\<close> known_ptrs local.known_ptrs_known_ptr by auto
ultimately show "h2 \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I)[1]
proof -
have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\<in>| object_ptr_kinds h"
using get_owner_document_ptr_in_heap owner_document by blast
then show "h2 \<turnstile> check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r ()"
by (simp add: \<open>object_ptr_kinds h = object_ptr_kinds h2\<close> check_in_heap_def)
next
show "(\<And>other_ptr other_children. h2 \<turnstile> get_child_nodes other_ptr \<rightarrow>\<^sub>r other_children
\<Longrightarrow> child \<notin> set other_children) \<Longrightarrow>
ptrs = sorted_list_of_set (fset (object_ptr_kinds h2)) \<Longrightarrow>
(\<And>ptr. ptr |\<in>| object_ptr_kinds h2 \<Longrightarrow> h2 \<turnstile> ok get_child_nodes ptr) \<Longrightarrow>
h2 \<turnstile> filter_M (\<lambda>ptr. Heap_Error_Monad.bind (get_child_nodes ptr)
(\<lambda>children. return (child \<in> set children))) (sorted_list_of_set (fset (object_ptr_kinds h2))) \<rightarrow>\<^sub>r []"
by(auto intro!: filter_M_empty_I bind_pure_I)
qed
qed
end
locale l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma remove_child_parent_child_rel_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "parent_child_rel h' \<subseteq> parent_child_rel h"
proof (standard, safe)
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(2)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure]
split: if_splits)[1]
using pure_returns_heap_eq by fastforce
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_eq: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq2: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
using select_result_eq by force
then have node_ptr_kinds_eq2: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by auto
then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
using node_ptr_kinds_M_eq by auto
have document_ptr_kinds_eq2: "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
using document_ptr_kinds_M_eq by auto
have children_eq:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
then have children_eq2:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq:
"\<And>document_ptr disconnected_nodes. document_ptr \<noteq> owner_document
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes"
apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers
by (metis (no_types, lifting) Un_iff owner_document select_result_I2)
then have disconnected_nodes_eq2:
"\<And>document_ptr. document_ptr \<noteq> owner_document
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] )
by (simp add: set_disconnected_nodes_get_child_nodes)
have "known_ptr ptr"
using assms(3)
using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h2]
using set_disconnected_nodes_types_preserved type_wf
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_child_nodes_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r remove1 child children_h"
using assms(2) owner_document h2 disconnected_nodes_h children_h
apply(auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto split: if_splits)[1]
apply(simp)
apply(auto split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E4)
apply(auto)[1]
apply(simp)
using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h'
by blast
fix parent child
assume a1: "(parent, child) \<in> parent_child_rel h'"
then show "(parent, child) \<in> parent_child_rel h"
proof (cases "parent = ptr")
case True
then show ?thesis
using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h'
get_child_nodes_ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1]
by (metis notin_set_remove1)
next
case False
then show ?thesis
using a1
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2)
qed
qed
lemma remove_child_heap_is_wellformed_preserved:
assumes "heap_is_wellformed h"
and "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'"
proof -
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(2)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1]
using pure_returns_heap_eq by fastforce
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_eq: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq2: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
using select_result_eq by force
then have node_ptr_kinds_eq2: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by auto
then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
using node_ptr_kinds_M_eq by auto
have document_ptr_kinds_eq2: "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
using document_ptr_kinds_M_eq by auto
have children_eq:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children =
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
then have children_eq2:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq: "\<And>document_ptr disconnected_nodes. document_ptr \<noteq> owner_document
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes"
apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def
using set_child_nodes_get_disconnected_nodes
set_disconnected_nodes_get_disconnected_nodes_different_pointers
by (metis (no_types, lifting) Un_iff owner_document select_result_I2)
then have disconnected_nodes_eq2:
"\<And>document_ptr. document_ptr \<noteq> owner_document
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads
set_disconnected_nodes_writes h2 children_h] )
by (simp add: set_disconnected_nodes_get_child_nodes)
show "known_ptrs h'"
using object_ptr_kinds_eq3 known_ptrs_preserved \<open>known_ptrs h\<close> by blast
have "known_ptr ptr"
using assms(3)
using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF set_disconnected_nodes_writes h2]
using set_disconnected_nodes_types_preserved type_wf
by(auto simp add: reflp_def transp_def)
then show "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_child_nodes_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r remove1 child children_h"
using assms(2) owner_document h2 disconnected_nodes_h children_h
apply(auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto split: if_splits)[1]
apply(simp)
apply(auto split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E4)
apply(auto)[1]
apply simp
using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h'
by blast
have disconnected_nodes_h2:
"h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
using owner_document assms(2) h2 disconnected_nodes_h
apply (auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E2)
apply(auto split: if_splits)[1]
apply(simp)
by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits)
then have disconnected_nodes_h':
"h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h'])
by (simp add: set_child_nodes_get_disconnected_nodes)
moreover have "a_acyclic_heap h"
using assms(1) by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h"
proof (standard, safe)
fix parent child
assume a1: "(parent, child) \<in> parent_child_rel h'"
then show "(parent, child) \<in> parent_child_rel h"
proof (cases "parent = ptr")
case True
then show ?thesis
using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h'
get_child_nodes_ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1]
by (metis imageI notin_set_remove1)
next
case False
then show ?thesis
using a1
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2)
qed
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h"
using assms(1) by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1]
- apply (metis (no_types, lifting) \<open>type_wf h'\<close> assms(2) assms(3) local.get_child_nodes_ok
- local.known_ptrs_known_ptr local.remove_child_children_subset notin_fset object_ptr_kinds_eq3
+ apply (metis (no_types, opaque_lifting) \<open>type_wf h'\<close> assms(2) assms(3) local.get_child_nodes_ok
+ local.known_ptrs_known_ptr local.remove_child_children_subset object_ptr_kinds_eq3
returns_result_select_result subset_code(1) type_wf)
- apply (metis (no_types, lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h
- disconnected_nodes_h' document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap
+ apply (metis (no_types, opaque_lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h
+ disconnected_nodes_h' document_ptr_kinds_eq3 local.remove_child_child_in_heap
node_ptr_kinds_eq3 select_result_I2 set_ConsD subset_code(1))
done
moreover have "a_owner_document_valid h"
using assms(1) by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3
node_ptr_kinds_eq3)[1]
proof -
fix node_ptr
assume 0: "\<forall>node_ptr\<in>fset (node_ptr_kinds h'). (\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
and 1: "node_ptr |\<in>| node_ptr_kinds h'"
and 2: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<longrightarrow>
node_ptr \<notin> set |h' \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
then show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h'
\<and> node_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
proof (cases "node_ptr = child")
case True
show ?thesis
apply(rule exI[where x=owner_document])
using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True
by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I
list.set_intros(1) select_result_I2)
next
case False
then show ?thesis
using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h
disconnected_nodes_h'
apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1]
- by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2))
+ by (metis children_eq2 disconnected_nodes_eq2 in_set_remove1 list.set_intros(2))
qed
qed
moreover
{
have h0: "a_distinct_lists h"
using assms(1) by (simp add: heap_is_wellformed_def)
moreover have ha1: "(\<Union>x\<in>set |h \<turnstile> object_ptr_kinds_M|\<^sub>r. set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r. set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
using \<open>a_distinct_lists h\<close>
unfolding a_distinct_lists_def
by(auto)
have ha2: "ptr |\<in>| object_ptr_kinds h"
using children_h get_child_nodes_ptr_in_heap by blast
have ha3: "child \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r"
using child_in_children_h children_h
by(simp)
have child_not_in: "\<And>document_ptr. document_ptr |\<in>| document_ptr_kinds h
\<Longrightarrow> child \<notin> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using ha1 ha2 ha3
apply(simp)
using IntI by fastforce
moreover have "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
apply(rule select_result_I)
by(auto simp add: object_ptr_kinds_M_defs)
moreover have "distinct |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
apply(rule select_result_I)
by(auto simp add: document_ptr_kinds_M_defs)
ultimately have "a_distinct_lists h'"
proof(simp (no_asm) add: a_distinct_lists_def, safe)
assume 1: "a_distinct_lists h"
and 3: "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
assume 1: "a_distinct_lists h"
and 3: "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
have 4: "distinct (concat ((map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r)))"
using 1 by(auto simp add: a_distinct_lists_def)
show "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
proof(rule distinct_concat_map_I[OF 3[unfolded object_ptr_kinds_eq2], simplified])
fix x
assume 5: "x |\<in>| object_ptr_kinds h'"
then have 6: "distinct |h \<turnstile> get_child_nodes x|\<^sub>r"
using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce
obtain children where children: "h \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children"
and distinct_children: "distinct children"
by (metis "5" "6" type_wf assms(3) get_child_nodes_ok local.known_ptrs_known_ptr
object_ptr_kinds_eq3 select_result_I)
obtain children' where children': "h' \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children'"
using children children_eq children_h' by fastforce
then have "distinct children'"
proof (cases "ptr = x")
case True
then show ?thesis
using children distinct_children children_h children_h'
by (metis children' distinct_remove1 returns_result_eq)
next
case False
then show ?thesis
using children distinct_children children_eq[OF False]
using children' distinct_lists_children h0
using select_result_I2 by fastforce
qed
then show "distinct |h' \<turnstile> get_child_nodes x|\<^sub>r"
using children' by(auto simp add: )
next
fix x y
assume 5: "x |\<in>| object_ptr_kinds h'" and 6: "y |\<in>| object_ptr_kinds h'" and 7: "x \<noteq> y"
obtain children_x where children_x: "h \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children_x"
by (metis "5" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children_y where children_y: "h \<turnstile> get_child_nodes y \<rightarrow>\<^sub>r children_y"
by (metis "6" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children_x' where children_x': "h' \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children_x'"
using children_eq children_h' children_x by fastforce
obtain children_y' where children_y': "h' \<turnstile> get_child_nodes y \<rightarrow>\<^sub>r children_y'"
using children_eq children_h' children_y by fastforce
have "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r))"
using h0 by(auto simp add: a_distinct_lists_def)
then have 8: "set children_x \<inter> set children_y = {}"
using "7" assms(1) children_x children_y local.heap_is_wellformed_one_parent by blast
have "set children_x' \<inter> set children_y' = {}"
proof (cases "ptr = x")
case True
then have "ptr \<noteq> y"
by(simp add: 7)
have "children_x' = remove1 child children_x"
using children_h children_h' children_x children_x' True returns_result_eq by fastforce
moreover have "children_y' = children_y"
using children_y children_y' children_eq[OF \<open>ptr \<noteq> y\<close>] by auto
ultimately show ?thesis
using 8 set_remove1_subset by fastforce
next
case False
then show ?thesis
proof (cases "ptr = y")
case True
have "children_y' = remove1 child children_y"
using children_h children_h' children_y children_y' True returns_result_eq by fastforce
moreover have "children_x' = children_x"
using children_x children_x' children_eq[OF \<open>ptr \<noteq> x\<close>] by auto
ultimately show ?thesis
using 8 set_remove1_subset by fastforce
next
case False
have "children_x' = children_x"
using children_x children_x' children_eq[OF \<open>ptr \<noteq> x\<close>] by auto
moreover have "children_y' = children_y"
using children_y children_y' children_eq[OF \<open>ptr \<noteq> y\<close>] by auto
ultimately show ?thesis
using 8 by simp
qed
qed
then show "set |h' \<turnstile> get_child_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_child_nodes y|\<^sub>r = {}"
using children_x' children_y'
by (metis (no_types, lifting) select_result_I2)
qed
next
assume 2: "distinct |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
then have 4: "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))"
by simp
have 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))))"
using h0
by(simp add: a_distinct_lists_def document_ptr_kinds_eq3)
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))))"
proof(rule distinct_concat_map_I[OF 4[unfolded document_ptr_kinds_eq3]])
fix x
assume 4: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 5: "distinct |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using distinct_lists_disconnected_nodes[OF h0] 4 get_disconnected_nodes_ok
by (simp add: type_wf document_ptr_kinds_eq3 select_result_I)
show "distinct |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "x = owner_document")
case True
have "child \<notin> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using child_not_in document_ptr_kinds_eq2 "4" by fastforce
moreover have "|h' \<turnstile> get_disconnected_nodes x|\<^sub>r = child # |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using disconnected_nodes_h' disconnected_nodes_h unfolding True
by(simp)
ultimately show ?thesis
using 5 unfolding True
by simp
next
case False
show ?thesis
using "5" False disconnected_nodes_eq2 by auto
qed
next
fix x y
assume 4: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and 5: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))" and "x \<noteq> y"
obtain disc_nodes_x where disc_nodes_x: "h \<turnstile> get_disconnected_nodes x \<rightarrow>\<^sub>r disc_nodes_x"
using 4 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of x] document_ptr_kinds_eq2
by auto
obtain disc_nodes_y where disc_nodes_y: "h \<turnstile> get_disconnected_nodes y \<rightarrow>\<^sub>r disc_nodes_y"
using 5 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of y] document_ptr_kinds_eq2
by auto
obtain disc_nodes_x' where disc_nodes_x': "h' \<turnstile> get_disconnected_nodes x \<rightarrow>\<^sub>r disc_nodes_x'"
using 4 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of x] document_ptr_kinds_eq2
by auto
obtain disc_nodes_y' where disc_nodes_y': "h' \<turnstile> get_disconnected_nodes y \<rightarrow>\<^sub>r disc_nodes_y'"
using 5 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of y] document_ptr_kinds_eq2
by auto
have "distinct
(concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) |h \<turnstile> document_ptr_kinds_M|\<^sub>r))"
using h0 by (simp add: a_distinct_lists_def)
then have 6: "set disc_nodes_x \<inter> set disc_nodes_y = {}"
using \<open>x \<noteq> y\<close> assms(1) disc_nodes_x disc_nodes_y local.heap_is_wellformed_one_disc_parent
by blast
have "set disc_nodes_x' \<inter> set disc_nodes_y' = {}"
proof (cases "x = owner_document")
case True
then have "y \<noteq> owner_document"
using \<open>x \<noteq> y\<close> by simp
then have "disc_nodes_y' = disc_nodes_y"
using disconnected_nodes_eq[OF \<open>y \<noteq> owner_document\<close>] disc_nodes_y disc_nodes_y'
by auto
have "disc_nodes_x' = child # disc_nodes_x"
using disconnected_nodes_h' disc_nodes_x disc_nodes_x' True disconnected_nodes_h returns_result_eq
by fastforce
have "child \<notin> set disc_nodes_y"
using child_not_in disc_nodes_y 5
using document_ptr_kinds_eq2 by fastforce
then show ?thesis
apply(unfold \<open>disc_nodes_x' = child # disc_nodes_x\<close> \<open>disc_nodes_y' = disc_nodes_y\<close>)
using 6 by auto
next
case False
then show ?thesis
proof (cases "y = owner_document")
case True
then have "disc_nodes_x' = disc_nodes_x"
using disconnected_nodes_eq[OF \<open>x \<noteq> owner_document\<close>] disc_nodes_x disc_nodes_x' by auto
have "disc_nodes_y' = child # disc_nodes_y"
using disconnected_nodes_h' disc_nodes_y disc_nodes_y' True disconnected_nodes_h returns_result_eq
by fastforce
have "child \<notin> set disc_nodes_x"
using child_not_in disc_nodes_x 4
using document_ptr_kinds_eq2 by fastforce
then show ?thesis
apply(unfold \<open>disc_nodes_y' = child # disc_nodes_y\<close> \<open>disc_nodes_x' = disc_nodes_x\<close>)
using 6 by auto
next
case False
have "disc_nodes_x' = disc_nodes_x"
using disconnected_nodes_eq[OF \<open>x \<noteq> owner_document\<close>] disc_nodes_x disc_nodes_x' by auto
have "disc_nodes_y' = disc_nodes_y"
using disconnected_nodes_eq[OF \<open>y \<noteq> owner_document\<close>] disc_nodes_y disc_nodes_y' by auto
then show ?thesis
apply(unfold \<open>disc_nodes_y' = disc_nodes_y\<close> \<open>disc_nodes_x' = disc_nodes_x\<close>)
using 6 by auto
qed
qed
then show "set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using disc_nodes_x' disc_nodes_y' by auto
qed
next
fix x xa xb
assume 1: "xa \<in> fset (object_ptr_kinds h')"
and 2: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 3: "xb \<in> fset (document_ptr_kinds h')"
and 4: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
obtain disc_nodes where disc_nodes: "h \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r disc_nodes"
using 3 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of xb] document_ptr_kinds_eq2 by auto
obtain disc_nodes' where disc_nodes': "h' \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r disc_nodes'"
using 3 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of xb] document_ptr_kinds_eq2 by auto
obtain children where children: "h \<turnstile> get_child_nodes xa \<rightarrow>\<^sub>r children"
- by (metis "1" type_wf assms(3) finite_set_in get_child_nodes_ok is_OK_returns_result_E
+ by (metis "1" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children' where children': "h' \<turnstile> get_child_nodes xa \<rightarrow>\<^sub>r children'"
using children children_eq children_h' by fastforce
have "\<And>x. x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r \<Longrightarrow> x \<in> set |h \<turnstile> get_disconnected_nodes xb|\<^sub>r \<Longrightarrow> False"
using 1 3
apply(fold \<open> object_ptr_kinds h = object_ptr_kinds h'\<close>)
apply(fold \<open> document_ptr_kinds h = document_ptr_kinds h'\<close>)
using children disc_nodes h0 apply(auto simp add: a_distinct_lists_def)[1]
by (metis (no_types, lifting) h0 local.distinct_lists_no_parent select_result_I2)
then have 5: "\<And>x. x \<in> set children \<Longrightarrow> x \<in> set disc_nodes \<Longrightarrow> False"
using children disc_nodes by fastforce
have 6: "|h' \<turnstile> get_child_nodes xa|\<^sub>r = children'"
using children' by simp
have 7: "|h' \<turnstile> get_disconnected_nodes xb|\<^sub>r = disc_nodes'"
using disc_nodes' by simp
have "False"
proof (cases "xa = ptr")
case True
have "distinct children_h"
using children_h distinct_lists_children h0 \<open>known_ptr ptr\<close> by blast
have "|h' \<turnstile> get_child_nodes ptr|\<^sub>r = remove1 child children_h"
using children_h'
by simp
have "children = children_h"
using True children children_h by auto
show ?thesis
using disc_nodes' children' 5 2 4 children_h \<open>distinct children_h\<close> disconnected_nodes_h'
apply(auto simp add: 6 7
\<open>xa = ptr\<close> \<open>|h' \<turnstile> get_child_nodes ptr|\<^sub>r = remove1 child children_h\<close> \<open>children = children_h\<close>)[1]
by (metis (no_types, lifting) disc_nodes disconnected_nodes_eq2 disconnected_nodes_h
select_result_I2 set_ConsD)
next
case False
have "children' = children"
using children' children children_eq[OF False[symmetric]]
by auto
then show ?thesis
proof (cases "xb = owner_document")
case True
then show ?thesis
using disc_nodes disconnected_nodes_h disconnected_nodes_h'
using "2" "4" "5" "6" "7" False \<open>children' = children\<close> assms(1) child_in_children_h
child_parent_dual children children_h disc_nodes' get_child_nodes_ptr_in_heap
list.set_cases list.simps(3) option.simps(1) returns_result_eq set_ConsD
by (metis (no_types, opaque_lifting) assms(3) type_wf)
next
case False
then show ?thesis
using "2" "4" "5" "6" "7" \<open>children' = children\<close> disc_nodes disc_nodes'
disconnected_nodes_eq returns_result_eq
by metis
qed
qed
then show "x \<in> {}"
by simp
qed
}
ultimately show "heap_is_wellformed h'"
using heap_is_wellformed_def by blast
qed
lemma remove_heap_is_wellformed_preserved:
assumes "heap_is_wellformed h"
and "h \<turnstile> remove child \<rightarrow>\<^sub>h h'"
and "known_ptrs h"
and type_wf: "type_wf h"
shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'"
using assms
by(auto simp add: remove_def intro: remove_child_heap_is_wellformed_preserved
elim!: bind_returns_heap_E2 split: option.splits)
lemma remove_child_removes_child:
assumes wellformed: "heap_is_wellformed h"
and remove_child: "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'"
and children: "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "child \<notin> set children"
proof -
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr' (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(2)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure]
split: if_splits)[1]
using pure_returns_heap_eq
by fastforce
have "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes remove_child])
unfolding remove_child_locs_def
using set_child_nodes_pointers_preserved set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
moreover have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF remove_child_writes assms(2)]
using set_child_nodes_types_preserved set_disconnected_nodes_types_preserved type_wf
unfolding remove_child_locs_def
apply(auto simp add: reflp_def transp_def)[1]
by blast
ultimately show ?thesis
using remove_child_removes_parent remove_child_heap_is_wellformed_preserved child_parent_dual
by (meson children known_ptrs local.known_ptrs_preserved option.distinct(1) remove_child
returns_result_eq type_wf wellformed)
qed
lemma remove_child_removes_first_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children"
assumes "h \<turnstile> remove_child ptr node_ptr \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
proof -
obtain h2 disc_nodes owner_document where
"h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document" and
"h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (node_ptr # disc_nodes) \<rightarrow>\<^sub>h h2" and
"h2 \<turnstile> set_child_nodes ptr children \<rightarrow>\<^sub>h h'"
using assms(5)
apply(auto simp add: remove_child_def
dest!: bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])[1]
by(auto elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated,OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
have "known_ptr ptr"
by (meson assms(3) assms(4) is_OK_returns_result_I get_child_nodes_ptr_in_heap known_ptrs_known_ptr)
moreover have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 assms(4)])
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
moreover have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h2]
using \<open>type_wf h\<close> set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
ultimately show ?thesis
using set_child_nodes_get_child_nodes\<open>h2 \<turnstile> set_child_nodes ptr children \<rightarrow>\<^sub>h h'\<close>
by fast
qed
lemma remove_removes_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children"
assumes "h \<turnstile> remove node_ptr \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
proof -
have "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some ptr"
using child_parent_dual assms by fastforce
show ?thesis
using assms remove_child_removes_first_child
by(auto simp add: remove_def
dest!: bind_returns_heap_E3[rotated, OF \<open>h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some ptr\<close>, rotated]
bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])
qed
lemma remove_for_all_empty_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
using assms
proof(induct children arbitrary: h h')
case Nil
then show ?case
by simp
next
case (Cons a children)
have "h \<turnstile> get_parent a \<rightarrow>\<^sub>r Some ptr"
using child_parent_dual Cons by fastforce
with Cons show ?case
proof(auto elim!: bind_returns_heap_E)[1]
fix h2
assume 0: "(\<And>h h'. heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r [])"
and 1: "heap_is_wellformed h"
and 2: "type_wf h"
and 3: "known_ptrs h"
and 4: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r a # children"
and 5: "h \<turnstile> get_parent a \<rightarrow>\<^sub>r Some ptr"
and 7: "h \<turnstile> remove a \<rightarrow>\<^sub>h h2"
and 8: "h2 \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h'"
then have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using remove_removes_child by blast
moreover have "heap_is_wellformed h2"
using 7 1 2 3 remove_child_heap_is_wellformed_preserved(3)
by(auto simp add: remove_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
split: option.splits)
moreover have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF remove_writes 7]
using \<open>type_wf h\<close> remove_child_types_preserved
by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
moreover have "object_ptr_kinds h = object_ptr_kinds h2"
using 7
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_writes])
using remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have "known_ptrs h2"
using 3 known_ptrs_preserved by blast
ultimately show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
using 0 8 by fast
qed
qed
end
locale l_remove_child_wf2 = l_type_wf + l_known_ptrs + l_remove_child_defs + l_heap_is_wellformed_defs
+ l_get_child_nodes_defs + l_remove_defs +
assumes remove_child_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes remove_child_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes remove_child_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
assumes remove_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove child \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes remove_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> remove child \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes remove_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> remove child \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
assumes remove_child_removes_child:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> child \<notin> set children"
assumes remove_child_removes_first_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children
\<Longrightarrow> h \<turnstile> remove_child ptr node_ptr \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes remove_removes_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children
\<Longrightarrow> h \<turnstile> remove node_ptr \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes remove_for_all_empty_children:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs
set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs
heap_is_wellformed parent_child_rel
by unfold_locales
declare l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma remove_child_wf2_is_l_remove_child_wf2 [instances]:
"l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove"
apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1]
using remove_child_heap_is_wellformed_preserved apply(fast, fast, fast)
using remove_heap_is_wellformed_preserved apply(fast, fast, fast)
using remove_child_removes_child apply fast
using remove_child_removes_first_child apply fast
using remove_removes_child apply fast
using remove_for_all_empty_children apply fast
done
subsection \<open>adopt\_node\<close>
locale l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf +
l_get_owner_document_wf +
l_remove_child_wf2 +
l_heap_is_wellformed
begin
lemma adopt_node_removes_first_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children"
shows "h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
proof -
obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> do { remove_child parent node }
| None \<Rightarrow> do { return ()}) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node # disc_nodes)
} else do { return () }) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: adopt_node_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_parent_pure])
have "h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using h2 remove_child_removes_first_child assms(1) assms(2) assms(3) assms(5)
by (metis list.set_intros(1) local.child_parent_dual option.simps(5) parent_opt returns_result_eq)
then
show ?thesis
using h'
by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes]
split: if_splits)
qed
lemma adopt_node_document_in_heap:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> ok (adopt_node owner_document node)"
shows "owner_document |\<in>| document_ptr_kinds h"
proof -
obtain old_document parent_opt h2 h' where
old_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> do { remove_child parent node } | None \<Rightarrow> do { return ()}) \<rightarrow>\<^sub>h h2"
and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node # disc_nodes)
} else do { return () }) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: adopt_node_def
elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_parent_pure])
show ?thesis
proof (cases "owner_document = old_document")
case True
then show ?thesis
using old_document get_owner_document_owner_document_in_heap assms(1) assms(2) assms(3)
by auto
next
case False
then obtain h3 old_disc_nodes disc_nodes where
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 node old_disc_nodes) \<rightarrow>\<^sub>h h3" and
old_disc_nodes: "h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h': "h3 \<turnstile> set_disconnected_nodes owner_document (node # disc_nodes) \<rightarrow>\<^sub>h h'"
using h'
by(auto elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "owner_document |\<in>| document_ptr_kinds h3"
by (meson is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
moreover have "object_ptr_kinds h = object_ptr_kinds h2"
using h2 apply(simp split: option.splits)
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes])
using remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def)
moreover have "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_disconnected_nodes_writes h3])
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
ultimately show ?thesis
by(auto simp add: document_ptr_kinds_def)
qed
qed
end
locale l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node +
l_get_owner_document_wf +
l_remove_child_wf2 +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma adopt_node_removes_child_step:
assumes wellformed: "heap_is_wellformed h"
and adopt_node: "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h2"
and children: "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<notin> set children"
proof -
obtain old_document parent_opt h' where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h': "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return () ) \<rightarrow>\<^sub>h h'"
using adopt_node get_parent_pure
by(auto simp add: adopt_node_def
elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
split: if_splits)
then have "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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
\<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close> returns_result_eq
by fastforce
next
case (Some option)
then show ?case
using remove_child_removes_child \<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close> known_ptrs type_wf
wellformed
by auto
qed
qed
lemma adopt_node_removes_child:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'"
shows "\<And>ptr' children'.
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children' \<Longrightarrow> node_ptr \<notin> set children'"
using adopt_node_removes_child_step assms by blast
lemma adopt_node_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^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 \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r old_document"
and
parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt"
and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent child | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2"
and
h': "h2 \<turnstile> (if document_ptr \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 child old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes document_ptr;
set_disconnected_nodes document_ptr (child # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^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="\<lambda>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:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have wellformed_h2: "heap_is_wellformed h2"
using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "type_wf h2"
using h2 remove_child_preserves_type_wf known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "known_ptrs h2"
using h2 remove_child_preserves_known_ptrs known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "heap_is_wellformed h' \<and> known_ptrs h' \<and> type_wf h'"
proof(cases "document_ptr = old_document")
case True
then show ?thesis
using h' wellformed_h2 \<open>type_wf h2\<close> \<open>known_ptrs h2\<close> by auto
next
case False
then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where
docs_neq: "document_ptr \<noteq> old_document" and
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 child old_disc_nodes) \<rightarrow>\<^sub>h h3" and
disc_nodes_document_ptr_h3:
"h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \<rightarrow>\<^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="\<lambda>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:
"\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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:
"\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r = |h3 \<turnstile> 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="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h3: "|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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:
"\<And>ptr children. h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r = |h' \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. old_document \<noteq> doc_ptr
\<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. old_document \<noteq> doc_ptr
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
using old_disc_nodes by blast
then have disc_nodes_old_document_h3:
"h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h']
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have "known_ptrs h3"
using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 by blast
then have "known_ptrs h'"
using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2:
"h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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 |\<in>| 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 |\<in>| 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 \<in> set disc_nodes_old_document_h2"
proof (insert parent_opt h2, induct parent_opt)
case None
then have "h = h2"
by(auto)
moreover have "a_owner_document_valid h"
using assms(1) heap_is_wellformed_def by(simp add: heap_is_wellformed_def)
ultimately show ?case
using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)]
in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast
next
case (Some option)
then show ?case
apply(simp split: option.splits)
using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs
by blast
qed
have "child \<notin> set (remove1 child disc_nodes_old_document_h2)"
using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \<open>distinct disc_nodes_old_document_h2\<close>
by auto
have "child \<notin> set disc_nodes_document_ptr_h3"
proof -
have "a_distinct_lists h2"
using heap_is_wellformed_def wellformed_h2 by blast
then have 0: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r))"
by(simp add: a_distinct_lists_def)
show ?thesis
using distinct_concat_map_E(1)[OF 0] \<open>child \<in> set disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h2 disc_nodes_document_ptr_h2
by (meson \<open>type_wf h2\<close> 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 |\<in>| node_ptr_kinds h"
using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]]
node_ptr_kinds_commutes by blast
have "a_acyclic_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h2"
proof
fix x
assume "x \<in> parent_child_rel h'"
then show "x \<in> parent_child_rel h2"
using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3
mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong
unfolding parent_child_rel_def
by(simp)
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h2\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h3"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1]
apply (simp add: children_eq2_h2 object_ptr_kinds_h2_eq3 subset_code(1))
by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> \<open>type_wf h2\<close>
disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2
document_ptr_kinds_eq3_h2 in_set_remove1 local.get_disconnected_nodes_ok
local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 returns_result_select_result
select_result_I2 wellformed_h2)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1]
apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1))
- by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> disc_nodes_document_ptr_h'
+ by (metis (no_types, opaque_lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> 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
+ local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
select_result_I2 set_ConsD subset_code(1) wellformed_h2)
moreover have "a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2
document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 )
by (smt disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2
disc_nodes_old_document_h2 disc_nodes_old_document_h3
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap
document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1
list.set_intros(1) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2
set_subset_Cons subset_code(1))
have a_distinct_lists_h2: "a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h'"
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2
children_eq2_h2 children_eq2_h3)[1]
proof -
assume 1: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 3: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> 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 \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 4: "distinct |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "old_document \<noteq> x")
case True
then show ?thesis
proof (cases "document_ptr \<noteq> x")
case True
then show ?thesis
using disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>]
disconnected_nodes_eq2_h3[OF \<open>document_ptr \<noteq> x\<close>] 4
by(auto)
next
case False
then show ?thesis
using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4
\<open>child \<notin> set disc_nodes_document_ptr_h3\<close>
by(auto simp add: disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>] )
qed
next
case False
then show ?thesis
by (metis (no_types, opaque_lifting) \<open>distinct disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h3 disconnected_nodes_eq2_h3
distinct_remove1 docs_neq select_result_I2)
qed
next
fix x y
assume a0: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a1: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a2: "x \<noteq> y"
moreover have 5: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
proof(cases "old_document = x")
case True
have "old_document \<noteq> y"
using \<open>x \<noteq> y\<close> \<open>old_document = x\<close> by simp
have "document_ptr \<noteq> x"
using docs_neq \<open>old_document = x\<close> 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] \<open>old_document = x\<close>
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
\<open>document_ptr \<noteq> x\<close> 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 \<open>old_document = x\<close>
docs_neq \<open>old_document \<noteq> y\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr = x\<close>
apply(simp)
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr \<noteq> x\<close>
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 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False
\<open>type_wf h2\<close> 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 \<noteq> y"
using \<open>x \<noteq> y\<close> by auto
have "set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
using \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document \<noteq> y\<close> \<open>document_ptr = x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3
\<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
by(auto)
next
case False
then show ?thesis
proof(cases "document_ptr = y")
case True
have f1: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set disc_nodes_document_ptr_h3 = {}"
using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>document_ptr \<noteq> x\<close> 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 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = {}"
using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>old_document \<noteq> x\<close>
by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2
- document_ptr_kinds_eq3_h3 finite_fset fmember_iff_member_fset set_sorted_list_of_set)
+ document_ptr_kinds_eq3_h3 finite_fset set_sorted_list_of_set)
ultimately show ?thesis
using 5 select_result_I2[OF disc_nodes_document_ptr_h']
select_result_I2[OF disc_nodes_old_document_h2] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr = y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> 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] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close>
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
by (metis \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
empty_iff inf.idem)
qed
qed
qed
qed
qed
next
fix x xa xb
assume 0: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 2: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h'"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h'"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
then show False
using \<open>child \<in> set disc_nodes_old_document_h2\<close> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a3: "h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r remove1 child disc_nodes_old_document_h2"
assume a4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a5: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f6: "old_document |\<in>| document_ptr_kinds h'"
using a1 \<open>xb |\<in>| document_ptr_kinds h'\<close> by blast
have f7: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a2 by simp
have "x \<in> set disc_nodes_old_document_h2"
using f6 a3 a1 by (metis (no_types) \<open>type_wf h'\<close> \<open>x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r\<close>
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' \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using f7 f6 a5 a4 \<open>xa |\<in>| object_ptr_kinds h'\<close>
by fastforce
then show ?thesis
using \<open>x \<in> set disc_nodes_old_document_h2\<close> a1 a4 f7 by blast
next
assume a1: "xb \<noteq> old_document"
assume a2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3"
assume a3: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a4: "xa |\<in>| object_ptr_kinds h'"
assume a5: "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r child # disc_nodes_document_ptr_h3"
assume a6: "old_document |\<in>| document_ptr_kinds h'"
assume a7: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
assume a8: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a10: "\<And>doc_ptr. old_document \<noteq> doc_ptr
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a11: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a12: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f13: "\<And>d. d \<notin> set |h' \<turnstile> document_ptr_kinds_M|\<^sub>r \<or> h2 \<turnstile> ok get_disconnected_nodes d"
using a9 \<open>type_wf h2\<close> get_disconnected_nodes_ok
by simp
then have f14: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a6 a3 by simp
have "x \<notin> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
using a12 a8 a4 \<open>xb |\<in>| document_ptr_kinds h'\<close>
- by (meson UN_I disjoint_iff_not_equal fmember_iff_member_fset)
+ by (meson UN_I disjoint_iff_not_equal)
then have "x = child"
using f13 a11 a10 a7 a5 a2 a1
by (metis (no_types, lifting) select_result_I2 set_ConsD)
then have "child \<notin> set disc_nodes_old_document_h2"
using f14 a12 a8 a6 a4
by (metis \<open>type_wf h'\<close> 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 \<open>child \<in> set disc_nodes_old_document_h2\<close> by fastforce
qed
qed
ultimately show ?thesis
using \<open>type_wf h'\<close> \<open>known_ptrs h'\<close> \<open>a_owner_document_valid h'\<close> heap_is_wellformed_def by blast
qed
then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'"
by auto
qed
lemma adopt_node_node_in_disconnected_nodes:
assumes wellformed: "heap_is_wellformed h"
and adopt_node: "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'"
and "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2"
and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node_ptr # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
assume a2: "\<And>h document_ptr disc_nodes h'. h \<turnstile> set_disconnected_nodes document_ptr disc_nodes \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume "h'a \<turnstile> set_disconnected_nodes owner_document (node_ptr # xb) \<rightarrow>\<^sub>h h'"
then have "node_ptr # xb = disc_nodes"
using a2 a1 by (meson returns_result_eq)
then show ?thesis
by (meson list.set_intros(1))
qed
qed
qed
end
interpretation i_adopt_node_wf?: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs
remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr
type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs
remove heap_is_wellformed parent_child_rel
by(simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs
remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr
type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs
remove heap_is_wellformed parent_child_rel get_root_node get_root_node_locs
by(simp add: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
locale l_adopt_node_wf = l_heap_is_wellformed + l_known_ptrs + l_type_wf + l_adopt_node_defs
+ l_get_child_nodes_defs + l_get_disconnected_nodes_defs +
assumes adopt_node_preserves_wellformedness:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h' \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> heap_is_wellformed h'"
assumes adopt_node_removes_child:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h2
\<Longrightarrow> h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> known_ptrs h
\<Longrightarrow> type_wf h \<Longrightarrow> node_ptr \<notin> set children"
assumes adopt_node_node_in_disconnected_nodes:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> node_ptr \<in> set disc_nodes"
assumes adopt_node_removes_first_child: "heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
assumes adopt_node_document_in_heap: "heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> h \<turnstile> ok (adopt_node owner_document node)
\<Longrightarrow> owner_document |\<in>| document_ptr_kinds h"
lemma adopt_node_wf_is_l_adopt_node_wf [instances]:
"l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes
get_disconnected_nodes known_ptrs adopt_node"
using heap_is_wellformed_is_l_heap_is_wellformed known_ptrs_is_l_known_ptrs
apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def)[1]
using adopt_node_preserves_wellformedness apply blast
using adopt_node_removes_child apply blast
using adopt_node_node_in_disconnected_nodes apply blast
using adopt_node_removes_first_child apply blast
using adopt_node_document_in_heap apply blast
done
subsection \<open>insert\_before\<close>
locale l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node_wf +
l_set_disconnected_nodes_get_child_nodes +
l_heap_is_wellformed
begin
lemma insert_before_removes_child:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "ptr \<noteq> ptr'"
assumes "h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children"
shows "h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
proof -
obtain owner_document h2 h3 disc_nodes reference_child where
"h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
"h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
"h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disc_nodes) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
split: if_splits option.splits)
have "h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using h2 adopt_node_removes_first_child assms(1) assms(2) assms(3) assms(6)
by simp
then have "h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using h3
by(auto simp add: set_disconnected_nodes_get_child_nodes
dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes])
then show ?thesis
using h' assms(4)
apply(auto simp add: a_insert_node_def
elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated])[1]
by(auto simp add: set_child_nodes_get_child_nodes_different_pointers
elim!: reads_writes_separate_forwards[OF get_child_nodes_reads set_child_nodes_writes])
qed
end
locale l_insert_before_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
+ l_insert_before_defs + l_get_child_nodes_defs +
assumes insert_before_removes_child:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> ptr \<noteq> ptr'
\<Longrightarrow> h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs
get_child_nodes get_child_nodes_locs set_child_nodes
set_child_nodes_locs get_ancestors get_ancestors_locs
adopt_node adopt_node_locs set_disconnected_nodes
set_disconnected_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_owner_document insert_before
insert_before_locs append_child type_wf known_ptr known_ptrs
heap_is_wellformed parent_child_rel
by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma insert_before_wf_is_l_insert_before_wf [instances]:
"l_insert_before_wf heap_is_wellformed type_wf known_ptr known_ptrs insert_before get_child_nodes"
apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1]
using insert_before_removes_child apply fast
done
locale l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_child_nodes_get_disconnected_nodes +
l_remove_child +
l_get_root_node_wf +
l_set_disconnected_nodes_get_disconnected_nodes_wf +
l_set_disconnected_nodes_get_ancestors +
l_get_ancestors_wf +
l_get_owner_document +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma insert_before_heap_is_wellformed_preserved:
assumes wellformed: "heap_is_wellformed h"
and insert_before: "h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
proof -
obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where
ancestors: "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
reference_child:
"h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
using assms(2)
by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
have "known_ptr ptr"
by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I known_ptrs
l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document)
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF adopt_node_writes h2]
using type_wf adopt_node_types_preserved
by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF insert_node_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs )
then have object_ptr_kinds_M_eq2_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have "known_ptrs h2"
using known_ptrs object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast
have wellformed_h2: "heap_is_wellformed h2"
using adopt_node_preserves_wellformedness[OF wellformed h2] known_ptrs type_wf .
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h2: "|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto
have "known_ptrs h3"
using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \<open>known_ptrs h2\<close> by blast
have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h3:
"|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h3: "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto
show "known_ptrs h'"
using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \<open>known_ptrs h3\<close> by blast
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. owner_document \<noteq> doc_ptr
\<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. doc_ptr \<noteq> owner_document
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_h3:
"h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r remove1 node disconnected_nodes_h2"
using h3 set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h3:
"\<And>ptr' children. ptr \<noteq> ptr'
\<Longrightarrow> h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr \<noteq> ptr' \<Longrightarrow> |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
obtain children_h3 where children_h3: "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h3"
using h' a_insert_node_def by auto
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r insert_before_list node reference_child children_h3"
using h' \<open>type_wf h3\<close> \<open>known_ptr ptr\<close>
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 |\<in>| object_ptr_kinds h3"
using children_h3 get_child_nodes_ptr_in_heap by blast
have node_in_heap: "node |\<in>| node_ptr_kinds h"
using h2 adopt_node_child_in_heap by fast
have child_not_in_any_children:
"\<And>p children. h2 \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children \<Longrightarrow> node \<notin> set children"
using wellformed h2 adopt_node_removes_child \<open>type_wf h\<close> \<open>known_ptrs h\<close> by auto
have "node \<in> set disconnected_nodes_h2"
using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1)
\<open>type_wf h\<close> \<open>known_ptrs h\<close> by blast
have node_not_in_disconnected_nodes:
"\<And>d. d |\<in>| document_ptr_kinds h3 \<Longrightarrow> node \<notin> set |h3 \<turnstile> get_disconnected_nodes d|\<^sub>r"
proof -
fix d
assume "d |\<in>| document_ptr_kinds h3"
show "node \<notin> set |h3 \<turnstile> 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 \<open>d |\<in>| document_ptr_kinds h3\<close> disconnected_nodes_h3
by fastforce
next
case False
then have
"set |h2 \<turnstile> get_disconnected_nodes d|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes owner_document|\<^sub>r = {}"
using distinct_concat_map_E(1) wellformed_h2
by (metis (no_types, lifting) \<open>d |\<in>| document_ptr_kinds h3\<close> \<open>type_wf h2\<close>
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] \<open>node \<in> set disconnected_nodes_h2\<close>
disconnected_nodes_h2 by fastforce
qed
qed
have "cast node \<noteq> ptr"
using ancestors node_not_in_ancestors get_ancestors_ptr
by fast
obtain ancestors_h2 where ancestors_h2: "h2 \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_ok object_ptr_kinds_M_eq2_h2 \<open>known_ptrs h2\<close> \<open>type_wf h2\<close>
by (metis is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2)
have ancestors_h3: "h3 \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_reads set_disconnected_nodes_writes h3
apply(rule reads_writes_separate_forwards)
using \<open>heap_is_wellformed h2\<close> ancestors_h2
by (auto simp add: set_disconnected_nodes_get_ancestors)
have node_not_in_ancestors_h2: "cast node \<notin> set ancestors_h2"
apply(rule get_ancestors_remains_not_in_ancestors[OF assms(1) wellformed_h2 ancestors ancestors_h2])
using adopt_node_children_subset using h2 \<open>known_ptrs h\<close> \<open> type_wf h\<close> apply(blast)
using node_not_in_ancestors apply(blast)
using object_ptr_kinds_M_eq3_h apply(blast)
using \<open>known_ptrs h\<close> apply(blast)
using \<open>type_wf h\<close> apply(blast)
using \<open>type_wf h2\<close> by blast
moreover have "a_acyclic_heap h'"
proof -
have "acyclic (parent_child_rel h2)"
using wellformed_h2 by (simp add: heap_is_wellformed_def acyclic_heap_def)
then have "acyclic (parent_child_rel h3)"
by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
moreover have "cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h2)\<^sup>*}"
using get_ancestors_parent_child_rel node_not_in_ancestors_h2 \<open>known_ptrs h2\<close> \<open>type_wf h2\<close>
using ancestors_h2 wellformed_h2 by blast
then have "cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h3)\<^sup>*}"
by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
moreover have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))"
using children_h3 children_h' ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3
insert_before_list_node_in_set)[1]
apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2)
by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2)
ultimately show ?thesis
by(auto simp add: acyclic_heap_def)
qed
moreover have "a_all_ptrs_in_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
have "a_all_ptrs_in_heap h'"
proof -
have "a_all_ptrs_in_heap h3"
using \<open>a_all_ptrs_in_heap h2\<close>
apply(auto simp add: a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2
children_eq_h2)[1]
using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
using node_ptr_kinds_eq2_h2 apply auto[1]
apply (metis \<open>known_ptrs h2\<close> \<open>type_wf h3\<close> children_eq_h2 local.get_child_nodes_ok
local.heap_is_wellformed_children_in_heap local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h2
returns_result_select_result wellformed_h2)
- by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2
- disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in node_ptr_kinds_commutes
+ by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2
+ disconnected_nodes_h3 document_ptr_kinds_commutes node_ptr_kinds_commutes
object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD)
have "set children_h3 \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using children_h3 \<open>a_all_ptrs_in_heap h3\<close>
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1]
by (metis children_eq_h2 l_heap_is_wellformed.heap_is_wellformed_children_in_heap
local.l_heap_is_wellformed_axioms node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h'
object_ptr_kinds_M_eq3_h2 wellformed_h2)
then have "set (insert_before_list node reference_child children_h3) \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_in_heap
apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1]
- by (metis (no_types, opaque_lifting) contra_subsetD finite_set_in insert_before_list_in_set
+ by (metis (no_types, opaque_lifting) contra_subsetD insert_before_list_in_set
node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h'
object_ptr_kinds_M_eq3_h2)
then show ?thesis
using \<open>a_all_ptrs_in_heap h3\<close>
apply(auto simp add: object_ptr_kinds_M_eq3_h' a_all_ptrs_in_heap_def node_ptr_kinds_def
node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1]
using children_eq_h3 children_h'
- apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD)
+ apply (metis (no_types, lifting) children_eq2_h3 select_result_I2 subsetD)
by (metis (no_types) \<open>type_wf h'\<close> disconnected_nodes_eq2_h3 disconnected_nodes_eq_h3
- finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok
+ is_OK_returns_result_I local.get_disconnected_nodes_ok
local.get_disconnected_nodes_ptr_in_heap returns_result_select_result subsetD)
qed
moreover have "a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h3"
proof(auto simp add: a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2
children_eq2_h2 intro!: distinct_concat_map_I)[1]
fix x
assume 1: "x |\<in>| document_ptr_kinds h3"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
show "distinct |h3 \<turnstile> 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_iff_member_fset set_sorted_list_of_set)
+ by (metis (full_types) distinct_remove1 finite_fset set_sorted_list_of_set)
next
fix x y xa
assume 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and 2: "x |\<in>| document_ptr_kinds h3"
and 3: "y |\<in>| document_ptr_kinds h3"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r"
and 6: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r"
show False
proof (cases "x = owner_document")
case True
then have "y \<noteq> 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 \<open>y \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
next
case False
then show ?thesis
proof (cases "y = owner_document")
case True
then show ?thesis
using distinct_concat_map_E(1)[OF 1]
using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2]
apply(auto simp add: True disconnected_nodes_eq2_h2[OF \<open>x \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
next
case False
then show ?thesis
using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6
using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
- disjoint_iff_not_equal finite_fset fmember_iff_member_fset notin_set_remove1 select_result_I2
- set_sorted_list_of_set
- by (metis (no_types, lifting))
+ disjoint_iff_not_equal notin_set_remove1 select_result_I2
+ by (metis (no_types, opaque_lifting))
qed
qed
next
fix x xa xb
assume 1: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h3 \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 2: "xa |\<in>| object_ptr_kinds h3"
and 3: "x \<in> set |h3 \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h3"
and 5: "x \<in> set |h3 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 4
by (metis \<open>type_wf h2\<close> children_eq2_h2 document_ptr_kinds_commutes known_ptrs
local.get_child_nodes_ok local.get_disconnected_nodes_ok
local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr
object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result
wellformed_h2)
show False
proof (cases "xb = owner_document")
case True
then show ?thesis
using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]]
by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1)
next
case False
show ?thesis
using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto
qed
qed
then have "a_distinct_lists h'"
proof(auto simp add: a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3
disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I)[1]
fix x
assume 1: "distinct (concat (map (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))" and
2: "x |\<in>| object_ptr_kinds h'"
have 3: "\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> distinct |h3 \<turnstile> get_child_nodes p|\<^sub>r"
using 1 by (auto elim: distinct_concat_map_E)
show "distinct |h' \<turnstile> 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 (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "x |\<in>| object_ptr_kinds h'"
and 3: "y |\<in>| object_ptr_kinds h'"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h' \<turnstile> get_child_nodes x|\<^sub>r"
and 6: "xa \<in> set |h' \<turnstile> get_child_nodes y|\<^sub>r"
have 7:"set |h3 \<turnstile> get_child_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 \<noteq> 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 \<open>ptr \<noteq> y\<close>])[1]
by (metis (no_types, opaque_lifting) "3" "7" \<open>type_wf h3\<close> children_eq2_h3 disjoint_iff_not_equal
get_child_nodes_ok insert_before_list_in_set known_ptrs local.known_ptrs_known_ptr
object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h'
object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2)
next
case False
then show ?thesis
proof (cases "ptr = y")
case True
then show ?thesis
using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6
apply(auto simp add: True children_eq2_h3[OF \<open>ptr \<noteq> x\<close>])[1]
by (metis (no_types, opaque_lifting) "2" "4" "7" IntI \<open>known_ptrs h3\<close> \<open>type_wf h'\<close>
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 \<open>ptr \<noteq> x\<close>] children_eq2_h3[OF \<open>ptr \<noteq> y\<close>] 5 6 7 by auto
qed
qed
next
fix x xa xb
assume 1: " (\<Union>x\<in>fset (object_ptr_kinds h'). set |h3 \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r) = {} "
and 2: "xa |\<in>| object_ptr_kinds h'"
and 3: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h'"
and 5: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 3 4 5
proof -
have "\<forall>h d. \<not> type_wf h \<or> d |\<notin>| document_ptr_kinds h \<or> h \<turnstile> ok get_disconnected_nodes d"
using local.get_disconnected_nodes_ok by satx
then have "h' \<turnstile> ok get_disconnected_nodes xb"
using "4" \<open>type_wf h'\<close> by fastforce
then have f1: "h3 \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
by (simp add: disconnected_nodes_eq_h3)
have "xa |\<in>| object_ptr_kinds h3"
using "2" object_ptr_kinds_M_eq3_h' by blast
then show ?thesis
using f1 \<open>local.a_distinct_lists h3\<close> local.distinct_lists_no_parent by fastforce
qed
show False
proof (cases "ptr = xa")
case True
show ?thesis
using 6 node_not_in_disconnected_nodes 3 4 5 select_result_I2[OF children_h']
select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3
by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M
\<open>a_distinct_lists h3\<close> \<open>type_wf h'\<close> disconnected_nodes_eq_h3
distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok
insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result)
next
case False
then show ?thesis
using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce
qed
qed
moreover have "a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_M_eq2_h2
object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3
document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2)[1]
apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified]
object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified]
node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1]
apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1]
- by (smt children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2
- disconnected_nodes_h3 finite_set_in in_set_remove1 insert_before_list_in_set
+ by (smt (verit) children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2
+ disconnected_nodes_h3 in_set_remove1 insert_before_list_in_set
object_ptr_kinds_M_eq3_h' ptr_in_heap select_result_I2)
ultimately show "heap_is_wellformed h'"
by (simp add: heap_is_wellformed_def)
qed
end
locale l_insert_before_wf2 = l_type_wf + l_known_ptrs + l_insert_before_defs
+ l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs +
assumes insert_before_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> insert_before ptr child ref \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes insert_before_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> insert_before ptr child ref \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes insert_before_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> insert_before ptr child ref \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs
get_child_nodes get_child_nodes_locs set_child_nodes
set_child_nodes_locs get_ancestors get_ancestors_locs
adopt_node adopt_node_locs set_disconnected_nodes
set_disconnected_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_owner_document insert_before
insert_before_locs append_child type_wf known_ptr known_ptrs
heap_is_wellformed parent_child_rel remove_child
remove_child_locs get_root_node get_root_node_locs
by(simp add: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma insert_before_wf2_is_l_insert_before_wf2 [instances]:
"l_insert_before_wf2 type_wf known_ptr known_ptrs insert_before heap_is_wellformed"
apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1]
using insert_before_heap_is_wellformed_preserved apply(fast, fast, fast)
done
locale l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_insert_before_wf +
l_insert_before_wf2 +
l_get_child_nodes
begin
lemma append_child_heap_is_wellformed_preserved:
assumes wellformed: "heap_is_wellformed h"
and append_child: "h \<turnstile> append_child ptr node \<rightarrow>\<^sub>h h'"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
using assms
by(auto simp add: append_child_def intro: insert_before_preserves_type_wf
insert_before_preserves_known_ptrs insert_before_heap_is_wellformed_preserved)
lemma append_child_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
assumes "h \<turnstile> append_child ptr node \<rightarrow>\<^sub>h h'"
assumes "node \<notin> set xs"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ [node]"
proof -
obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where
ancestors: "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node None \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: append_child_def insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
have "\<And>parent. |h \<turnstile> get_parent node|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr"
using assms(1) assms(4) assms(6)
by (metis (no_types, lifting) assms(2) assms(3) h2 is_OK_returns_heap_I is_OK_returns_result_E
local.adopt_node_child_in_heap local.get_parent_child_dual local.get_parent_ok
select_result_I2)
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
using get_child_nodes_reads adopt_node_writes h2 assms(4)
apply(rule reads_writes_separate_forwards)
using \<open>\<And>parent. |h \<turnstile> get_parent node|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>
apply(auto simp add: adopt_node_locs_def remove_child_locs_def)[1]
by (meson local.set_child_nodes_get_child_nodes_different_pointers)
have "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
using get_child_nodes_reads set_disconnected_nodes_writes h3 \<open>h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs\<close>
apply(rule reads_writes_separate_forwards)
by(auto)
have "ptr |\<in>| object_ptr_kinds h"
by (meson ancestors is_OK_returns_result_I local.get_ancestors_ptr_in_heap)
then
have "known_ptr ptr"
using assms(3)
using local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF adopt_node_writes h2]
using adopt_node_types_preserved \<open>type_wf h\<close>
by(auto simp add: adopt_node_locs_def remove_child_locs_def reflp_def transp_def split: if_splits)
then
have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h3]
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs@[node]"
using h'
apply(auto simp add: a_insert_node_def
dest!: bind_returns_heap_E3[rotated, OF \<open>h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs\<close>
get_child_nodes_pure, rotated])[1]
using \<open>type_wf h3\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close>
by metis
qed
lemma append_child_for_all_on_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
assumes "h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
assumes "set nodes \<inter> set xs = {}"
assumes "distinct nodes"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs@nodes"
using assms
apply(induct nodes arbitrary: h xs)
apply(simp)
proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a
assume 0: "(\<And>h xs. heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs \<Longrightarrow> h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'
\<Longrightarrow> set nodes \<inter> set xs = {} \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ nodes)"
and 1: "heap_is_wellformed h"
and 2: "type_wf h"
and 3: "known_ptrs h"
and 4: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
and 5: "h \<turnstile> append_child ptr a \<rightarrow>\<^sub>r ()"
and 6: "h \<turnstile> append_child ptr a \<rightarrow>\<^sub>h h'a"
and 7: "h'a \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
and 8: "a \<notin> set xs"
and 9: "set nodes \<inter> set xs = {}"
and 10: "a \<notin> set nodes"
and 11: "distinct nodes"
then have "h'a \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ [a]"
using append_child_children 6
using "1" "2" "3" "4" "8" by blast
moreover have "heap_is_wellformed h'a" and "type_wf h'a" and "known_ptrs h'a"
using insert_before_heap_is_wellformed_preserved insert_before_preserves_known_ptrs
insert_before_preserves_type_wf 1 2 3 6 append_child_def
by metis+
moreover have "set nodes \<inter> set (xs @ [a]) = {}"
using 9 10
by auto
ultimately show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ a # nodes"
using 0 7
by fastforce
qed
lemma append_child_for_all_on_no_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
assumes "h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
assumes "distinct nodes"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r nodes"
using assms append_child_for_all_on_children
by force
end
locale l_append_child_wf = l_type_wf + l_known_ptrs + l_append_child_defs + l_heap_is_wellformed_defs +
assumes append_child_preserves_type_wf:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> type_wf h'"
assumes append_child_preserves_known_ptrs:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> known_ptrs h'"
assumes append_child_heap_is_wellformed_preserved:
"type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> heap_is_wellformed h \<Longrightarrow> h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'
\<Longrightarrow> heap_is_wellformed h'"
interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent
get_parent_locs remove_child remove_child_locs
get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs
adopt_node adopt_node_locs known_ptr type_wf get_child_nodes
get_child_nodes_locs known_ptrs set_child_nodes
set_child_nodes_locs remove get_ancestors get_ancestors_locs
insert_before insert_before_locs append_child heap_is_wellformed
parent_child_rel
by(auto simp add: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr
known_ptrs append_child heap_is_wellformed"
apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1]
using append_child_heap_is_wellformed_preserved by fast+
subsection \<open>create\_element\<close>
locale l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel +
l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs +
l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs
get_disconnected_nodes get_disconnected_nodes_locs +
l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr +
l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs +
l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr
get_child_nodes get_child_nodes_locs +
l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs
get_child_nodes get_child_nodes_locs +
l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs +
l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs +
l_new_element type_wf +
l_known_ptrs known_ptr known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
begin
lemma create_element_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(2)
by(auto simp add: create_element_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_element_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_element_ptr)"
using \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr
by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []"
using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr]
new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes
by blast
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_child_nodes)
then have children_eq2_h2: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_element_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3: "\<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_element_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close>
using \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def acyclic_heap_def)
also have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting)
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally have "a_acyclic_heap h'"
by (simp add: acyclic_heap_def)
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h2"
apply(auto simp add: a_all_ptrs_in_heap_def)[1]
apply (metis \<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> assms(1)
assms(3) funion_iff local.get_child_nodes_ok local.known_ptrs_known_ptr
local.parent_child_rel_child_in_heap local.parent_child_rel_child_nodes2 node_ptr_kinds_commutes
node_ptr_kinds_eq_h returns_result_select_result)
by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funion_iff
local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h
returns_result_select_result)
then have "a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "a_all_ptrs_in_heap h'"
- by (smt \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> children_eq2_h3
+ by (smt (verit) \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> children_eq2_h3
disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
- document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I
+ document_ptr_kinds_eq_h3 h' is_OK_returns_result_I
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap
local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_element_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
heap_is_wellformed_children_in_heap
by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp
fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result)
then have "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_element_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_element_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_element_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_element_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
by (metis \<open>local.a_distinct_lists h\<close> \<open>type_wf h2\<close> disconnected_nodes_eq_h document_ptr_kinds_eq_h
local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
children_eq2_h2 object_ptr_kinds_eq_h2)
then have "a_distinct_lists h'"
proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3
object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3
intro!: distinct_concat_map_I)[1]
fix x
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> 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) \<open>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 \<notin> set disc_nodes_h3\<close>
\<open>a_distinct_lists h3\<close> \<open>type_wf h'\<close> disc_nodes_h3 distinct.simps(2)
distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq
returns_result_select_result)
next
fix x y xa
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
+ apply (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>local.a_all_ptrs_in_heap h\<close>
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'
+ disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
- by (smt NodeMonad.ptr_kinds_ptr_kinds_M
+ by (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>local.a_all_ptrs_in_heap h\<close>
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'
+ disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply -
apply(cases "xb = document_ptr")
apply (metis (no_types, opaque_lifting) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3
\<Longrightarrow> 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 \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
\<open>a_distinct_lists h3\<close> children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h'
select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes)
by (metis "3" "4" "5" "6" \<open>a_distinct_lists h3\<close> \<open>type_wf h3\<close> children_eq2_h3
distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result)
qed
have "a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(auto simp add: a_owner_document_valid_def)[1]
apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1]
apply(auto simp add: object_ptr_kinds_eq_h2)[1]
apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1]
apply(auto simp add: document_ptr_kinds_eq_h2)[1]
apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1]
apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1]
apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric]
disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3)[1]
apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1)
local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
apply(simp add: object_ptr_kinds_eq_h)
- by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M
+ by(metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> 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'
+ document_ptr_kinds_eq_h h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
node_ptr_kinds_commutes select_result_I2)
show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_create_element_wf?: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf
get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
set_tag_name set_tag_name_locs
set_disconnected_nodes set_disconnected_nodes_locs create_element
using instances
by(auto simp add: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>create\_character\_data\<close>
locale l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
+ l_new_character_data_get_disconnected_nodes
get_disconnected_nodes get_disconnected_nodes_locs
+ l_set_val_get_disconnected_nodes
type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs
+ l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr
+ l_new_character_data_get_child_nodes
type_wf known_ptr get_child_nodes get_child_nodes_locs
+ l_set_val_get_child_nodes
type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs
+ l_set_disconnected_nodes_get_child_nodes
set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs
+ l_set_disconnected_nodes
type_wf set_disconnected_nodes set_disconnected_nodes_locs
+ l_set_disconnected_nodes_get_disconnected_nodes
type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs
+ l_new_character_data
type_wf
+ l_known_ptrs
known_ptr known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes ::
"(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) character_data_ptr) prog"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
begin
lemma create_character_data_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(2)
by(auto simp add: create_character_data_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_character_data_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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 \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
local.create_character_data_known_ptr by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h:
"\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2
get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_child_nodes)
then have children_eq2_h2:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_character_data_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_val_writes h3]
using set_val_types_preserved
by(auto simp add: reflp_def transp_def)
then show "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_character_data_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close> using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def acyclic_heap_def)
also have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally have "a_acyclic_heap h'"
by (simp add: acyclic_heap_def)
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h2"
apply(auto simp add: a_all_ptrs_in_heap_def)[1]
using node_ptr_kinds_eq_h \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \<open>parent_child_rel h = parent_child_rel h2\<close>
- children_eq2_h finite_set_in finsert_iff funion_finsert_right local.parent_child_rel_child
+ children_eq2_h finsert_iff funion_finsert_right local.parent_child_rel_child
local.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h
select_result_I2 subsetD sup_bot.right_neutral)
by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funionI1
local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap
node_ptr_kinds_eq_h returns_result_select_result)
then have "a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "a_all_ptrs_in_heap h'"
- by (smt character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2
+ by (smt (verit) character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3
- finite_set_in h' h2 local.a_all_ptrs_in_heap_def
+ h' h2 local.a_all_ptrs_in_heap_def
local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr
new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3
object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_character_data_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
heap_is_wellformed_children_in_heap
by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp
fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result)
then have "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_character_data_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_character_data_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr
returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
by (metis \<open>local.a_distinct_lists h\<close> \<open>type_wf h2\<close> disconnected_nodes_eq_h document_ptr_kinds_eq_h
local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
children_eq2_h2 object_ptr_kinds_eq_h2)[1]
then have "a_distinct_lists h'"
proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3
object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1]
fix x
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> 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) \<open>cast new_character_data_ptr \<notin> set disc_nodes_h3\<close>
\<open>a_distinct_lists h3\<close> \<open>type_wf h'\<close> disc_nodes_h3 distinct.simps(2)
distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq
returns_result_select_result)
next
fix x y xa
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using calculation by(auto dest: distinct_concat_map_E(1))
ultimately show "False"
- by (smt NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
+ by (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>local.a_all_ptrs_in_heap h\<close> 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'
+ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply(cases "xb = document_ptr")
apply (metis (no_types, opaque_lifting) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
\<open>a_distinct_lists h3\<close> children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h'
select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes)
by (metis "3" "4" "5" "6" \<open>a_distinct_lists h3\<close> \<open>type_wf h3\<close> children_eq2_h3
distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result)
qed
have "a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(simp add: a_owner_document_valid_def)
apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )
apply(simp add: object_ptr_kinds_eq_h2)
apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )
apply(simp add: document_ptr_kinds_eq_h2)
apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )
apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )
apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1]
apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1)
local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
apply(simp add: object_ptr_kinds_eq_h)
- by (metis (mono_tags, lifting) \<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
- children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h'
+ by (metis (mono_tags, opaque_lifting) \<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
+ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h h'
l_ptr_kinds_M.ptr_kinds_ptr_kinds_M
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms object_ptr_kinds_M_def
select_result_I2)
show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_create_character_data_wf?: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes
set_disconnected_nodes_locs create_character_data known_ptrs
using instances
by (auto simp add: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>create\_document\<close>
locale l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
+ l_new_document_get_disconnected_nodes
get_disconnected_nodes get_disconnected_nodes_locs
+ l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
create_document
+ l_new_document_get_child_nodes
type_wf known_ptr get_child_nodes get_child_nodes_locs
+ l_new_document
type_wf
+ l_known_ptrs
known_ptr known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_document :: "((_) heap, exception, (_) document_ptr) prog"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
begin
lemma create_document_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_document \<rightarrow>\<^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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr" and
h': "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
using assms(2)
apply(simp add: create_document_def)
using new_document_ok by blast
have "new_document_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have "new_document_ptr |\<notin>| 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 |\<notin>| object_ptr_kinds h"
by simp
have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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)
+ by (auto simp add: document_ptr_kinds_def)
have children_eq:
"\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h' get_child_nodes_new_document[rotated, OF new_document_ptr h']
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2: "\<And>ptr'. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
- by (metis(full_types) \<open>\<And>thesis. (\<And>new_document_ptr.
- \<lbrakk>h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr; h \<turnstile> new_document \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
- local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+
+ by (smt (verit))+
then have disconnected_nodes_eq2_h: "\<And>doc_ptr. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
using h' local.new_document_no_disconnected_nodes new_document_ptr by blast
have "type_wf h'"
using \<open>type_wf h\<close> new_document_types_preserved h' by blast
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def acyclic_heap_def)
also have "parent_child_rel h = parent_child_rel h'"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h'"
by (simp add: object_ptr_kinds_eq)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_document_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2 empty_iff empty_set image_eqI select_result_I2)
qed
finally have "a_acyclic_heap h'"
by (simp add: acyclic_heap_def)
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def)[1]
using ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close> assms(1) children_eq fset_of_list_elem
local.heap_is_wellformed_children_in_heap local.parent_child_rel_child
local.parent_child_rel_parent_in_heap node_ptr_kinds_eq
- apply (metis (no_types, lifting) \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
- children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq
+ apply (metis (no_types, opaque_lifting) \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
+ children_eq2 finsert_iff funion_finsert_right object_ptr_kinds_eq
select_result_I2 subsetD sup_bot.right_neutral)
by (metis (no_types, lifting) \<open>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 |\<notin>| object_ptr_kinds h\<close>
\<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
\<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close> \<open>type_wf h'\<close> assms(1) disconnected_nodes_eq_h
local.get_disconnected_nodes_ok
local.heap_is_wellformed_disc_nodes_in_heap local.parent_child_rel_child
local.parent_child_rel_parent_in_heap
node_ptr_kinds_eq returns_result_select_result select_result_I2)
have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h'"
using \<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close>
\<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: children_eq2[symmetric] a_distinct_lists_def insort_split object_ptr_kinds_eq
document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(auto simp add: dest: distinct_concat_map_E)[1]
apply(auto simp add: dest: distinct_concat_map_E)[1]
using \<open>new_document_ptr |\<notin>| document_ptr_kinds h\<close>
apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1]
using disconnected_nodes_eq_h
apply (metis assms(1) assms(3) disconnected_nodes_eq2_h local.get_disconnected_nodes_ok
local.heap_is_wellformed_disconnected_nodes_distinct
returns_result_select_result)
proof -
fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr"
assume a1: "x \<noteq> y"
assume a2: "x |\<in>| document_ptr_kinds h"
assume a3: "x \<noteq> new_document_ptr"
assume a4: "y |\<in>| document_ptr_kinds h"
assume a5: "y \<noteq> new_document_ptr"
assume a6: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
assume a7: "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
assume a8: "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
have f9: "xa \<in> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using a7 a3 disconnected_nodes_eq2_h by presburger
have f10: "xa \<in> set |h \<turnstile> get_disconnected_nodes y|\<^sub>r"
using a8 a5 disconnected_nodes_eq2_h by presburger
have f11: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))"
using a4 by simp
have "x \<in> 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' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
and 1: "h' \<turnstile> 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) \<rightarrow>\<^sub>r []"
and 2: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
and 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
and 4: "(\<Union>x\<in>fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h). set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 5: "x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
and 7: "xa |\<in>| object_ptr_kinds h"
and 8: "xa \<noteq> 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 |\<in>| document_ptr_kinds h"
and 10: "xb \<noteq> new_document_ptr"
then show "False"
by (metis \<open>local.a_distinct_lists h\<close> assms(3) disconnected_nodes_eq2_h
local.distinct_lists_no_parent local.get_disconnected_nodes_ok
returns_result_select_result)
qed
have "a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def)[1]
by (metis \<open>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 |\<notin>| object_ptr_kinds h\<close>
- children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in
+ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes
funion_iff node_ptr_kinds_eq object_ptr_kinds_eq)
show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_create_document_wf?: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed parent_child_rel
set_val set_val_locs set_disconnected_nodes
set_disconnected_nodes_locs create_document known_ptrs
using instances
by (auto simp add: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def)
declare l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
end
diff --git a/thys/Core_SC_DOM/safely_composable/classes/ElementClass.thy b/thys/Core_SC_DOM/safely_composable/classes/ElementClass.thy
--- a/thys/Core_SC_DOM/safely_composable/classes/ElementClass.thy
+++ b/thys/Core_SC_DOM/safely_composable/classes/ElementClass.thy
@@ -1,327 +1,329 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Element\<close>
text\<open>In this theory, we introduce the types for the Element class.\<close>
theory ElementClass
imports
"NodeClass"
"ShadowRootPointer"
begin
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, define
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
type_synonym attr_key = DOMString
type_synonym attr_value = DOMString
type_synonym attrs = "(attr_key, attr_value) fmap"
type_synonym tag_name = DOMString
record ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr) RElement = RNode +
nothing :: unit
tag_name :: tag_name
child_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
attrs :: attrs
shadow_root_opt :: "'shadow_root_ptr shadow_root_ptr option"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_scheme"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node
= "(('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) Node"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object
= "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_ext + 'Node) Object"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object"
type_synonym
('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap
= "(('document_ptr, 'shadow_root_ptr) document_ptr + 'object_ptr, 'element_ptr element_ptr +
'character_data_ptr character_data_ptr + 'node_ptr, 'Object,
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_ext + 'Node) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition element_ptr_kinds :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where
"element_ptr_kinds heap = the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`|
(ffilter is_element_ptr_kind (node_ptr_kinds heap)))"
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) =
{|element_ptr|} |\<union>| element_ptr_kinds h"
- apply(auto simp add: element_ptr_kinds_def)[1]
- by force
+ by (auto simp add: element_ptr_kinds_def)
definition element_ptrs :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where
"element_ptrs heap = ffilter is_element_ptr (element_ptr_kinds heap)"
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Node \<Rightarrow> (_) Element option"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = (case RNode.more node of Inl element \<Rightarrow>
Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Object \<Rightarrow> (_) Element option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj \<equiv> (case cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj of Some node \<Rightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node | None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
definition cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Element \<Rightarrow> (_) Node"
where
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = RNode.extend (RNode.truncate element) (Inl (RNode.more element))"
adhoc_overloading cast cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e
abbreviation cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) Element \<Rightarrow> (_) Object"
where
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr)"
adhoc_overloading cast cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
consts is_element_kind :: 'a
definition is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Node \<Rightarrow> bool"
where
"is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr \<longleftrightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
adhoc_overloading is_element_kind is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e
lemmas is_element_kind_def = is_element_kind\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
abbreviation is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) Object \<Rightarrow> bool"
where
"is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
adhoc_overloading is_element_kind is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma element_ptr_kinds_commutes [simp]:
"cast element_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> element_ptr |\<in>| element_ptr_kinds h"
- apply(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)[1]
- by (metis (no_types, lifting) element_ptr_casts_commute2 ffmember_filter fimage_eqI
- fset.map_comp is_element_ptr_kind_none node_ptr_casts_commute3
- node_ptr_kinds_commutes node_ptr_kinds_def option.sel option.simps(3))
+proof (rule iffI)
+ show "cast element_ptr |\<in>| node_ptr_kinds h \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h"
+ by (metis (no_types, opaque_lifting) element_ptr_casts_commute2 element_ptr_kinds_def
+ ffmember_filter fimage_eqI is_element_ptr_kind_cast option.sel)
+next
+ show "element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> cast element_ptr |\<in>| node_ptr_kinds h"
+ by (auto simp: element_ptr_kinds_def)
+qed
definition get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Element option"
where
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h = Option.bind (get\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) h) cast"
adhoc_overloading get get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
locale l_type_wf_def\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (NodeClass.type_wf h \<and> (\<forall>element_ptr \<in> fset (element_ptr_kinds h).
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t: "type_wf h \<Longrightarrow> ElementClass.type_wf h"
sublocale l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t \<subseteq> l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
apply(unfold_locales)
using NodeClass.a_type_wf_def
by (meson ElementClass.a_type_wf_def l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
locale l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas = l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
sublocale l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas by unfold_locales
lemma get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf:
assumes "type_wf h"
shows "element_ptr |\<in>| element_ptr_kinds h \<longleftrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h \<noteq> None"
using l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
- by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf bind_eq_None_conv element_ptr_kinds_commutes notin_fset
+ by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf bind_eq_None_conv element_ptr_kinds_commutes
option.distinct(1))
end
global_interpretation l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf
by unfold_locales
definition put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) Element \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) (cast element)"
adhoc_overloading put put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h = h'"
shows "element_ptr |\<in>| element_ptr_kinds h'"
using assms
unfolding put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def element_ptr_kinds_def
by (metis element_ptr_kinds_commutes element_ptr_kinds_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap)
lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs:
assumes "put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast element_ptr|}"
using assms
by (simp add: put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs)
lemma cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]:
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e x = cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def)
by (metis (full_types) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = None \<longleftrightarrow> \<not> (\<exists>element. cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node)"
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)[1]
by (metis (full_types) RNode.select_convs(2) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = Some element \<longleftrightarrow> cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node"
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv [simp]: "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element) = Some element"
by simp
lemma get_elment_ptr_simp1 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h) = Some element"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_elment_ptr_simp2 [simp]:
"element_ptr \<noteq> element_ptr'
\<Longrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' element h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
abbreviation "create_element_obj tag_name_arg child_nodes_arg attrs_arg shadow_root_opt_arg
\<equiv> \<lparr> RObject.nothing = (), RNode.nothing = (), RElement.nothing = (),
tag_name = tag_name_arg, Element.child_nodes = child_nodes_arg, attrs = attrs_arg,
shadow_root_opt = shadow_root_opt_arg, \<dots> = None \<rparr>"
definition new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) heap \<Rightarrow> ((_) element_ptr \<times> (_) heap)"
where
"new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_element_ptr = element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref
|`| (element_ptrs h)))))
in
(new_element_ptr, put new_element_ptr (create_element_obj '''' [] fmempty None) h))"
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "new_element_ptr |\<in>| element_ptr_kinds h'"
using assms
unfolding new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
using put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap by blast
lemma new_element_ptr_new:
"element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref |`| element_ptrs h)))) |\<notin>| element_ptrs h"
by (metis Suc_n_not_le_n element_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "new_element_ptr |\<notin>| element_ptr_kinds h"
using assms
unfolding new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis Pair_inject element_ptrs_def ffmember_filter new_element_ptr_new is_element_ptr_ref)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using assms
by (metis Pair_inject new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_is_element_ptr:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "is_element_ptr new_element_ptr"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
assumes "ptr \<noteq> cast new_element_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
assumes "ptr \<noteq> cast new_element_ptr"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
assumes "ptr \<noteq> new_element_ptr"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
locale l_known_ptr\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_element_ptr ptr)"
lemma known_ptr_not_element_ptr: "\<not>is_element_ptr ptr \<Longrightarrow> a_known_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
global_interpretation l_known_ptr\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def
by blast
end
diff --git a/thys/DOM_Components/Core_DOM_Components.thy b/thys/DOM_Components/Core_DOM_Components.thy
--- a/thys/DOM_Components/Core_DOM_Components.thy
+++ b/thys/DOM_Components/Core_DOM_Components.thy
@@ -1,2154 +1,2154 @@
(***********************************************************************************
* 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 \<open>DOM Components\<close>
theory Core_DOM_Components
imports Core_DOM.Core_DOM
begin
locale l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs =
l_get_root_node_defs get_root_node get_root_node_locs +
l_to_tree_order_defs to_tree_order
for get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
begin
definition a_get_component :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_get_component ptr = do {
root \<leftarrow> get_root_node ptr;
to_tree_order root
}"
definition a_is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"a_is_strongly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = (
let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
let arg_components =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_component ptr|\<^sub>r) \<inter> fset (object_ptr_kinds h).
set |h \<turnstile> a_get_component ptr|\<^sub>r) in
let arg_components' =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_component ptr|\<^sub>r) \<inter> fset (object_ptr_kinds h').
set |h' \<turnstile> a_get_component ptr|\<^sub>r) in
removed_pointers \<subseteq> arg_components \<and>
added_pointers \<subseteq> arg_components' \<and>
S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \<subseteq> arg_components' \<and>
(\<forall>outside_ptr \<in> fset (object_ptr_kinds h) \<inter> fset (object_ptr_kinds h') -
(\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))"
definition a_is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"a_is_weakly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = (
let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
let arg_components =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_component ptr|\<^sub>r) \<inter> fset (object_ptr_kinds h).
set |h \<turnstile> a_get_component ptr|\<^sub>r) in
let arg_components' =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_component ptr|\<^sub>r) \<inter> fset (object_ptr_kinds h').
set |h' \<turnstile> a_get_component ptr|\<^sub>r) in
removed_pointers \<subseteq> arg_components \<and>
S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \<subseteq> arg_components' \<union> added_pointers \<and>
(\<forall>outside_ptr \<in> fset (object_ptr_kinds h) \<inter> fset (object_ptr_kinds h') -
(\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))"
lemma "a_is_strongly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' \<Longrightarrow> a_is_weakly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h'"
by(auto simp add: a_is_strongly_dom_component_safe_def a_is_weakly_dom_component_safe_def Let_def)
definition is_document_component :: "(_) object_ptr list \<Rightarrow> bool"
where
"is_document_component c = is_document_ptr_kind (hd c)"
definition is_disconnected_component :: "(_) object_ptr list \<Rightarrow> bool"
where
"is_disconnected_component c = is_node_ptr_kind (hd c)"
end
global_interpretation l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs to_tree_order
defines get_component = a_get_component
and is_strongly_dom_component_safe = a_is_strongly_dom_component_safe
and is_weakly_dom_component_safe = a_is_weakly_dom_component_safe
.
locale l_get_component_defs =
fixes get_component :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
fixes is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
fixes is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
locale l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_to_tree_order_wf +
l_get_component_defs +
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +
l_get_ancestors +
l_get_ancestors_wf +
l_get_root_node +
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent +
l_get_parent_wf +
l_get_element_by +
l_to_tree_order_wf_get_root_node_wf +
(* l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ get_child_nodes +
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ get_child_nodes+
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ "l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.a_to_tree_order get_child_nodes"
for get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog" *)
assumes get_component_impl: "get_component = a_get_component"
assumes is_strongly_dom_component_safe_impl:
"is_strongly_dom_component_safe = a_is_strongly_dom_component_safe"
assumes is_weakly_dom_component_safe_impl:
"is_weakly_dom_component_safe = a_is_weakly_dom_component_safe"
begin
lemmas get_component_def = a_get_component_def[folded get_component_impl]
lemmas is_strongly_dom_component_safe_def =
a_is_strongly_dom_component_safe_def[folded is_strongly_dom_component_safe_impl]
lemmas is_weakly_dom_component_safe_def =
a_is_weakly_dom_component_safe_def[folded is_weakly_dom_component_safe_impl]
lemma get_dom_component_ptr_in_heap:
assumes "h \<turnstile> ok (get_component ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms get_root_node_ptr_in_heap
by(auto simp add: get_component_def)
lemma get_component_ok:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_component ptr)"
using assms
apply(auto simp add: get_component_def a_get_root_node_def intro!: bind_is_OK_pure_I)[1]
using get_root_node_ok to_tree_order_ok get_root_node_ptr_in_heap
apply blast
by (simp add: local.get_root_node_root_in_heap local.to_tree_order_ok)
lemma get_dom_component_ptr:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
shows "ptr \<in> set c"
proof(insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev )
case (step child)
then show ?case
proof (cases "is_node_ptr_kind child")
case True
obtain node_ptr where
node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child"
using \<open>is_node_ptr_kind child\<close> node_ptr_casts_commute3 by blast
have "child |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> get_component child \<rightarrow>\<^sub>r c\<close> get_dom_component_ptr_in_heap by fast
with node_ptr have "node_ptr |\<in>| node_ptr_kinds h"
by auto
then obtain parent_opt where
parent: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt"
using get_parent_ok \<open>type_wf h\<close> \<open>known_ptrs h\<close>
by fast
then show ?thesis
proof (induct parent_opt)
case None
then have "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
by (simp add: local.get_root_node_no_parent)
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> node_ptr step(2)
apply(auto simp add: get_component_def a_get_root_node_def elim!: bind_returns_result_E2)[1]
using to_tree_order_ptr_in_result returns_result_eq by fastforce
next
case (Some parent_ptr)
then have "h \<turnstile> get_component parent_ptr \<rightarrow>\<^sub>r c"
using step(2) node_ptr \<open>type_wf h\<close> \<open>known_ptrs h\<close> \<open>heap_is_wellformed h\<close>
get_root_node_parent_same
by(auto simp add: get_component_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)
then have "parent_ptr \<in> set c"
using step node_ptr Some by blast
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> \<open>heap_is_wellformed h\<close> step(2) node_ptr Some
apply(auto simp add: get_component_def elim!: bind_returns_result_E2)[1]
using to_tree_order_parent by blast
qed
next
case False
then show ?thesis
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> step(2)
apply(auto simp add: get_component_def elim!: bind_returns_result_E2)[1]
by (metis is_OK_returns_result_I local.get_root_node_not_node_same
local.get_root_node_ptr_in_heap local.to_tree_order_ptr_in_result returns_result_eq)
qed
qed
lemma get_component_parent_inside:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "cast node_ptr \<in> set c"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some parent"
shows "parent \<in> set c"
proof -
have "parent |\<in>| object_ptr_kinds h"
using assms(6) get_parent_parent_in_heap by blast
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr" and
c: "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using assms(4)
by (metis (no_types, opaque_lifting) bind_returns_result_E2 get_component_def get_root_node_pure)
then have "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r root_ptr"
using assms(1) assms(2) assms(3) assms(5) to_tree_order_same_root by blast
then have "h \<turnstile> get_root_node parent \<rightarrow>\<^sub>r root_ptr"
using assms(6) get_root_node_parent_same by blast
then have "h \<turnstile> get_component parent \<rightarrow>\<^sub>r c"
using c get_component_def by auto
then show ?thesis
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
qed
lemma get_component_subset:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "ptr' \<in> set c"
shows "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
proof(insert assms(1) assms(5), induct ptr' rule: heap_wellformed_induct_rev )
case (step child)
then show ?case
proof (cases "is_node_ptr_kind child")
case True
obtain node_ptr where
node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child"
using \<open>is_node_ptr_kind child\<close> node_ptr_casts_commute3 by blast
have "child |\<in>| object_ptr_kinds h"
using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2)
unfolding get_component_def
by (meson bind_returns_result_E2 get_root_node_pure)
with node_ptr have "node_ptr |\<in>| node_ptr_kinds h"
by auto
then obtain parent_opt where
parent: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt"
using get_parent_ok \<open>type_wf h\<close> \<open>known_ptrs h\<close>
by fast
then show ?thesis
proof (induct parent_opt)
case None
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) get_root_node_no_parent node_ptr by blast
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> node_ptr step(2) assms(4) assms(1)
by (metis (no_types) bind_pure_returns_result_I2 bind_returns_result_E2 get_component_def
get_root_node_pure is_OK_returns_result_I returns_result_eq to_tree_order_same_root)
next
case (Some parent_ptr)
then have "h \<turnstile> get_component parent_ptr \<rightarrow>\<^sub>r c"
using step get_component_parent_inside assms node_ptr by blast
then show ?case
using Some node_ptr
apply(auto simp add: get_component_def elim!: bind_returns_result_E2
del: bind_pure_returns_result_I intro!: bind_pure_returns_result_I)[1]
using get_root_node_parent_same by blast
qed
next
case False
then have "child |\<in>| object_ptr_kinds h"
using assms(1) assms(4) step(2)
by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure
get_component_def to_tree_order_ptrs_in_heap)
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) False get_root_node_not_node_same by blast
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) step.prems
by (metis (no_types) False \<open>child |\<in>| object_ptr_kinds h\<close> bind_pure_returns_result_I2
bind_returns_result_E2 get_component_def get_root_node_ok get_root_node_pure returns_result_eq
to_tree_order_node_ptrs)
qed
qed
lemma get_component_to_tree_order_subset:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
shows "set nodes \<subseteq> set c"
using assms
apply(auto simp add: get_component_def elim!: bind_returns_result_E2)[1]
by (meson to_tree_order_subset assms(5) contra_subsetD get_dom_component_ptr)
lemma get_component_to_tree_order:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to"
assumes "ptr \<in> set to"
shows "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
get_component_ok get_component_subset get_component_to_tree_order_subset is_OK_returns_result_E
local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap select_result_I2 subsetCE)
lemma get_component_root_node_same:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
assumes "x \<in> set c"
shows "h \<turnstile> get_root_node x \<rightarrow>\<^sub>r root_ptr"
proof(insert assms(1) assms(6), induct x rule: heap_wellformed_induct_rev )
case (step child)
then show ?case
proof (cases "is_node_ptr_kind child")
case True
obtain node_ptr where
node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child"
using \<open>is_node_ptr_kind child\<close> node_ptr_casts_commute3 by blast
have "child |\<in>| object_ptr_kinds h"
using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2)
unfolding get_component_def
by (meson bind_returns_result_E2 get_root_node_pure)
with node_ptr have "node_ptr |\<in>| node_ptr_kinds h"
by auto
then obtain parent_opt where
parent: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt"
using get_parent_ok \<open>type_wf h\<close> \<open>known_ptrs h\<close>
by fast
then show ?thesis
proof (induct parent_opt)
case None
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) get_root_node_no_parent node_ptr by blast
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> node_ptr step(2) assms(4) assms(1) assms(5)
by (metis (no_types) \<open>child |\<in>| object_ptr_kinds h\<close> bind_pure_returns_result_I
get_component_def get_dom_component_ptr get_component_subset get_root_node_pure is_OK_returns_result_E
returns_result_eq to_tree_order_ok to_tree_order_same_root)
next
case (Some parent_ptr)
then have "h \<turnstile> get_component parent_ptr \<rightarrow>\<^sub>r c"
using step get_component_parent_inside assms node_ptr
by (meson get_component_subset)
then show ?case
using Some node_ptr
apply(auto simp add: get_component_def elim!: bind_returns_result_E2)[1]
using get_root_node_parent_same
using \<open>h \<turnstile> get_component parent_ptr \<rightarrow>\<^sub>r c\<close> assms(1) assms(2) assms(3) get_dom_component_ptr
step.hyps
by blast
qed
next
case False
then have "child |\<in>| object_ptr_kinds h"
using assms(1) assms(4) step(2)
by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure
get_component_def to_tree_order_ptrs_in_heap)
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) False get_root_node_not_node_same by auto
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) step.prems assms(5)
by (metis (no_types, opaque_lifting) bind_returns_result_E2 get_component_def get_root_node_pure
returns_result_eq to_tree_order_same_root)
qed
qed
lemma get_dom_component_no_overlap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c'"
shows "set c \<inter> set c' = {} \<or> c = c'"
proof (rule ccontr, auto)
fix x
assume 1: "c \<noteq> c'" and 2: "x \<in> set c" and 3: "x \<in> set c'"
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
using assms(4) unfolding get_component_def
by (meson bind_is_OK_E is_OK_returns_result_I)
moreover obtain root_ptr' where root_ptr': "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr'"
using assms(5) unfolding get_component_def
by (meson bind_is_OK_E is_OK_returns_result_I)
ultimately have "root_ptr \<noteq> root_ptr'"
using 1 assms
unfolding get_component_def
by (meson bind_returns_result_E3 get_root_node_pure returns_result_eq)
moreover have "h \<turnstile> get_root_node x \<rightarrow>\<^sub>r root_ptr"
using 2 root_ptr get_component_root_node_same assms by blast
moreover have "h \<turnstile> get_root_node x \<rightarrow>\<^sub>r root_ptr'"
using 3 root_ptr' get_component_root_node_same assms by blast
ultimately show False
using select_result_I2 by force
qed
lemma get_component_separates_tree_order:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c'"
assumes "ptr' \<notin> set c"
shows "set to \<inter> set c' = {}"
proof -
have "c \<noteq> c'"
using assms get_dom_component_ptr by blast
then have "set c \<inter> set c' = {}"
using assms get_dom_component_no_overlap by blast
moreover have "set to \<subseteq> set c"
using assms get_component_to_tree_order_subset by blast
ultimately show ?thesis
by blast
qed
lemma get_component_separates_tree_order_general:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> to_tree_order ptr'' \<rightarrow>\<^sub>r to''"
assumes "ptr'' \<in> set c"
assumes "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c'"
assumes "ptr' \<notin> set c"
shows "set to'' \<inter> set c' = {}"
proof -
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
using assms(4)
by (metis bind_is_OK_E get_component_def is_OK_returns_result_I)
then have c: "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using assms(4) unfolding get_component_def
by (simp add: bind_returns_result_E3)
with root_ptr show ?thesis
using assms get_component_separates_tree_order get_component_subset
by meson
qed
end
interpretation i_get_component?: l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name
by(auto simp add: l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_component_def
is_strongly_dom_component_safe_def is_weakly_dom_component_safe_def instances)
declare l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_child\_nodes\<close>
locale l_get_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_dom_component_get_child_nodes:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "cast child \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast child \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual
local.get_root_node_parent_same
by blast
then have "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual
local.get_component_parent_inside local.get_component_subset
by blast
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr
by blast
next
assume 1: "ptr' \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root_ptr"
using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual local.get_root_node_parent_same
by blast
then have "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual
local.get_component_parent_inside local.get_component_subset
by blast
then show "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \<in> set c"
by (smt \<open>h \<turnstile> get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r root_ptr\<close> assms(1) assms(2) assms(3)
assms(5) assms(6) disjoint_iff_not_equal is_OK_returns_result_E is_OK_returns_result_I
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_no_overlap local.child_parent_dual local.get_component_ok
local.get_component_parent_inside local.get_dom_component_ptr local.get_root_node_ptr_in_heap
local.l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms)
qed
lemma get_child_nodes_get_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
shows "cast ` set children \<subseteq> set c"
using assms get_dom_component_get_child_nodes
using local.get_dom_component_ptr by auto
lemma get_child_nodes_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} (cast ` set children) h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_child_nodes_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
- by (smt IntI finite_set_in get_dom_component_get_child_nodes is_OK_returns_result_E
+ by (smt IntI get_dom_component_get_child_nodes is_OK_returns_result_E
is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_component_impl
local.get_component_ok local.get_dom_component_ptr select_result_I2)
qed
end
interpretation i_get_component_get_child_nodes?: l_get_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_parent\<close>
locale l_get_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_parent_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_parent ptr' \<rightarrow>\<^sub>r Some parent"
shows "parent \<in> set c \<longleftrightarrow> cast ptr' \<in> set c"
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) is_OK_returns_result_E
l_to_tree_order_wf.to_tree_order_parent local.get_component_parent_inside local.get_component_subset
local.get_component_to_tree_order_subset local.get_parent_parent_in_heap local.l_to_tree_order_wf_axioms
local.to_tree_order_ok local.to_tree_order_ptr_in_result subsetCE)
lemma get_parent_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some parent"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {cast node_ptr} {parent} h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_parent_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
using get_dom_component_get_child_nodes local.get_component_impl local.get_dom_component_ptr
- by (metis (no_types, opaque_lifting) Int_iff finite_set_in get_parent_is_strongly_dom_component_safe_step
+ by (metis (no_types, opaque_lifting) Int_iff get_parent_is_strongly_dom_component_safe_step
local.get_component_ok local.get_parent_parent_in_heap local.to_tree_order_ok local.to_tree_order_parent
local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap returns_result_select_result)
qed
end
interpretation i_get_component_get_parent?: l_get_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_root\_node\<close>
locale l_get_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_root_node_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root"
shows "root \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "root \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node root \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_component_root_node_same root_ptr
by blast
moreover have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
by (metis (no_types, lifting) calculation assms(1) assms(2) assms(3) assms(5)
is_OK_returns_result_E local.get_root_node_root_in_heap local.to_tree_order_ok
local.to_tree_order_ptr_in_result local.to_tree_order_same_root select_result_I2)
ultimately have "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
apply(auto simp add: get_component_def)[1]
by (smt "1" assms(1) assms(2) assms(3) assms(4) bind_pure_returns_result_I bind_returns_result_E3
local.get_component_def local.get_component_subset local.get_root_node_pure)
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume 1: "ptr' \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node root \<rightarrow>\<^sub>r root_ptr"
by (metis assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E local.get_root_node_root_in_heap
local.to_tree_order_ok local.to_tree_order_ptr_in_result local.to_tree_order_same_root returns_result_eq)
then have "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
using "1" assms(1) assms(2) assms(3) assms(4) local.get_component_subset by blast
then show "root \<in> set c"
using assms(5) bind_returns_result_E3 local.get_component_def local.to_tree_order_ptr_in_result
by fastforce
qed
lemma get_root_node_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} {root} h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_root_node_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
- by (metis (no_types, lifting) IntI get_root_node_is_strongly_dom_component_safe_step is_OK_returns_result_I
+ by (metis (no_types, opaque_lifting) IntI get_root_node_is_strongly_dom_component_safe_step is_OK_returns_result_I
local.get_component_impl local.get_component_ok local.get_dom_component_ptr
- local.get_root_node_ptr_in_heap notin_fset returns_result_select_result)
+ local.get_root_node_ptr_in_heap returns_result_select_result)
qed
end
interpretation i_get_component_get_root_node?: l_get_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_element\_by\_id\<close>
locale l_get_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_element_by_id_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_element_by_id ptr' idd \<rightarrow>\<^sub>r Some result"
shows "cast result \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast result \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_component_def is_OK_returns_result_I)
then have "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c\<close>
by (simp add: get_component_def bind_returns_result_E3)
obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
using \<open>h \<turnstile> get_element_by_id ptr' idd \<rightarrow>\<^sub>r Some result\<close>
apply(simp add: get_element_by_id_def first_in_tree_order_def)
by (meson bind_is_OK_E is_OK_returns_result_I)
then have "cast result \<in> set to'"
using \<open>h \<turnstile> get_element_by_id ptr' idd \<rightarrow>\<^sub>r Some result\<close> get_element_by_id_result_in_tree_order
by blast
have "h \<turnstile> get_root_node (cast result) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using \<open>cast result \<in> set to'\<close> \<open>h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'\<close>
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_component_root_node_same
get_component_subset get_component_to_tree_order
by blast
then have "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c\<close>
using get_component_def by auto
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume "ptr' \<in> set c"
moreover obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
get_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
ultimately have "set to' \<subseteq> set c"
using assms(1) assms(2) assms(3) assms(4) get_component_subset get_component_to_tree_order_subset
by blast
moreover have "cast result \<in> set to'"
using assms(5) get_element_by_id_result_in_tree_order to' by blast
ultimately show "cast result \<in> set c"
by blast
qed
lemma get_element_by_id_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_element_by_id ptr idd \<rightarrow>\<^sub>r Some result"
assumes "h \<turnstile> get_element_by_id ptr idd \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} {cast result} h h'"
proof -
have "h = h'"
using assms(5)
by(auto simp add: preserved_def get_element_by_id_def first_in_tree_order_def
elim!: bind_returns_heap_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_element_by_id_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast result \<in> set to"
using assms(4) local.get_element_by_id_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_component_impl
local.get_component_ok
by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def get_element_by_id_def
first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I
split: option.splits list.splits)[1]
- by (metis (no_types, lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4) finite_set_in
+ by (metis (no_types, opaque_lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4)
get_element_by_id_is_strongly_dom_component_safe_step local.get_component_impl local.get_dom_component_ptr
select_result_I2)
qed
end
interpretation i_get_component_get_element_by_id?: l_get_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_elements\_by\_class\_name\<close>
locale l_get_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_elements_by_class_name_result_in_tree_order:
assumes "h \<turnstile> get_elements_by_class_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "result \<in> set results"
shows "cast result \<in> set to"
using assms
by(auto simp add: get_elements_by_class_name_def first_in_tree_order_def
elim!: map_filter_M_pure_E[where y=result] bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF assms(2), rotated]
intro!: map_filter_M_pure map_M_pure_I bind_pure_I
split: option.splits list.splits if_splits)
lemma get_elements_by_class_name_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_elements_by_class_name ptr' name \<rightarrow>\<^sub>r results"
assumes "result \<in> set results"
shows "cast result \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast result \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_component_def is_OK_returns_result_I)
then have "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c\<close>
by (simp add: get_component_def bind_returns_result_E3)
obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
using \<open>h \<turnstile> get_elements_by_class_name ptr' name \<rightarrow>\<^sub>r results\<close>
apply(simp add: get_elements_by_class_name_def first_in_tree_order_def)
by (meson bind_is_OK_E is_OK_returns_result_I)
then have "cast result \<in> set to'"
using get_elements_by_class_name_result_in_tree_order assms by blast
have "h \<turnstile> get_root_node (cast result) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using \<open>cast result \<in> set to'\<close> \<open>h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'\<close>
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_component_root_node_same
get_component_subset get_component_to_tree_order
by blast
then have "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c\<close>
using get_component_def by auto
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume "ptr' \<in> set c"
moreover obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
get_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
ultimately have "set to' \<subseteq> set c"
using assms(1) assms(2) assms(3) assms(4) get_component_subset get_component_to_tree_order_subset
by blast
moreover have "cast result \<in> set to'"
using assms get_elements_by_class_name_result_in_tree_order to' by blast
ultimately show "cast result \<in> set c"
by blast
qed
lemma get_elements_by_class_name_pure [simp]:
"pure (get_elements_by_class_name ptr name) h"
by(auto simp add: get_elements_by_class_name_def intro!: bind_pure_I map_filter_M_pure
split: option.splits)
lemma get_elements_by_class_name_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_elements_by_class_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> get_elements_by_class_name ptr name \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'"
proof -
have "h = h'"
using assms(5)
by (meson get_elements_by_class_name_pure pure_returns_heap_eq)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_elements_by_class_name_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast ` set results \<subseteq> set to"
using assms(4) local.get_elements_by_class_name_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_component_impl local.get_component_ok
by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def
get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2
intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
- by (metis (no_types, lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4) finite_set_in
+ by (metis (no_types, opaque_lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4)
get_elements_by_class_name_is_strongly_dom_component_safe_step local.get_component_impl
local.get_dom_component_ptr select_result_I2)
qed
end
interpretation i_get_component_get_elements_by_class_name?: l_get_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_elements\_by\_tag\_name\<close>
locale l_get_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_elements_by_tag_name_result_in_tree_order:
assumes "h \<turnstile> get_elements_by_tag_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "result \<in> set results"
shows "cast result \<in> set to"
using assms
by(auto simp add: get_elements_by_tag_name_def first_in_tree_order_def
elim!: map_filter_M_pure_E[where y=result] bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF assms(2), rotated]
intro!: map_filter_M_pure map_M_pure_I bind_pure_I
split: option.splits list.splits if_splits)
lemma get_elements_by_tag_name_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_elements_by_tag_name ptr' name \<rightarrow>\<^sub>r results"
assumes "result \<in> set results"
shows "cast result \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast result \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_component_def is_OK_returns_result_I)
then have "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> get_component ptr \<rightarrow>\<^sub>r c\<close>
by (simp add: get_component_def bind_returns_result_E3)
obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
using \<open>h \<turnstile> get_elements_by_tag_name ptr' name \<rightarrow>\<^sub>r results\<close>
apply(simp add: get_elements_by_tag_name_def first_in_tree_order_def)
by (meson bind_is_OK_E is_OK_returns_result_I)
then have "cast result \<in> set to'"
using get_elements_by_tag_name_result_in_tree_order assms by blast
have "h \<turnstile> get_root_node (cast result) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using \<open>cast result \<in> set to'\<close> \<open>h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'\<close>
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_component_root_node_same
get_component_subset get_component_to_tree_order
by blast
then have "h \<turnstile> get_component ptr' \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c\<close>
using get_component_def by auto
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume "ptr' \<in> set c"
moreover obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
get_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
ultimately have "set to' \<subseteq> set c"
using assms(1) assms(2) assms(3) assms(4) get_component_subset get_component_to_tree_order_subset
by blast
moreover have "cast result \<in> set to'"
using assms get_elements_by_tag_name_result_in_tree_order to' by blast
ultimately show "cast result \<in> set c"
by blast
qed
lemma get_elements_by_tag_name_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_elements_by_tag_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> get_elements_by_tag_name ptr name \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'"
proof -
have "h = h'"
using assms(5)
by (meson get_elements_by_tag_name_pure pure_returns_heap_eq)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_elements_by_tag_name_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast ` set results \<subseteq> set to"
using assms(4) local.get_elements_by_tag_name_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_component_impl local.get_component_ok
by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def
get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2
intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
- by (metis (no_types, lifting) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close> finite_set_in
+ by (metis (no_types, opaque_lifting) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close>
get_elements_by_tag_name_is_strongly_dom_component_safe_step local.get_component_impl
local.get_dom_component_ptr select_result_I2)
qed
end
interpretation i_get_component_get_elements_by_tag_name?: l_get_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>remove\_child\<close>
lemma remove_child_not_strongly_dom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and ptr and child where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe { ptr, cast child} {} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
e2 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
return (e1, e2)
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e1 = "fst |?h0 \<turnstile> ?P|\<^sub>r"
let ?e2 = "snd |?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1" and child="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e2"])
by code_simp+
qed
subsection \<open>adopt\_node\<close>
lemma adopt_node_not_strongly_dom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and document_ptr and child where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {cast document_ptr, cast child} {} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
document_ptr2 \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
adopt_node document_ptr2 (cast e1);
return (document_ptr, e1)
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "fst |?h0 \<turnstile> ?P|\<^sub>r"
let ?e1 = "snd |?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and document_ptr="?document_ptr" and child="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1"])
by code_simp+
qed
subsection \<open>create\_element\<close>
lemma create_element_not_strongly_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and document_ptr and new_element_ptr and tag where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {cast document_ptr} {cast new_element_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "create_document"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and document_ptr="?document_ptr"])
by code_simp+
qed
locale l_get_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order
get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name +
l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr +
l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs +
l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_tag_name set_tag_name_locs +
l_new_element_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs +
l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_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_disconnected_nodes type_wf set_tag_name set_tag_name_locs
get_disconnected_nodes get_disconnected_nodes_locs +
l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_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_get_disconnected_nodes type_wf get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs +
l_new_element type_wf +
l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name
set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs
create_element
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_component :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_element_by_id :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_elements_by_class_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_elements_by_tag_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
begin
lemma create_element_is_weakly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<noteq> cast |h \<turnstile> create_element document_ptr tag|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain new_element_ptr h2 h3 disc_nodes where
new_element_ptr: "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile>set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: create_element_def elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
using new_element_ptr h2 h3 disc_nodes h'
apply(auto simp add: create_element_def intro!: bind_returns_result_I
bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "preserved (get_M ptr getter) h h2"
using h2 new_element_ptr
apply(rule new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
using new_element_ptr assms(6) \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close>
by simp
have "preserved (get_M ptr getter) h2 h3"
using set_tag_name_writes h3
apply(rule reads_writes_preserved2)
apply(auto simp add: set_tag_name_locs_impl a_set_tag_name_locs_def all_args_def)[1]
by (metis (no_types, lifting) \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> assms(6)
get_M_Element_preserved8 select_result_I2)
have "document_ptr |\<in>| document_ptr_kinds h"
using create_element_document_in_heap
using assms(4)
by blast
then
have "ptr \<noteq> (cast document_ptr)"
using assms(5) assms(1) assms(2) assms(3) local.get_component_ok local.get_dom_component_ptr
by auto
have "preserved (get_M ptr getter) h3 h'"
using set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved2)
apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1]
by (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
show ?thesis
using \<open>preserved (get_M ptr getter) h h2\<close> \<open>preserved (get_M ptr getter) h2 h3\<close>
\<open>preserved (get_M ptr getter) h3 h'\<close>
by(auto simp add: preserved_def)
qed
lemma create_element_is_weakly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r result \<rightarrow>\<^sub>h h'"
shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'"
proof -
obtain new_element_ptr h2 h3 disc_nodes_h3 where
new_element_ptr: "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: create_element_def returns_result_heap_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'"
by (meson assms(4) returns_result_heap_def)
have "new_element_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_element_ptr)"
using \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr
by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2
by blast
then
have "known_ptrs h3"
using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast
then
have "known_ptrs h'"
using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []"
using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr]
new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes
by blast
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_child_nodes)
then have children_eq2_h2: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_element_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_tag_name_writes h3]
using set_tag_name_types_preserved
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3: "\<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_element_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close>
using \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
a_all_ptrs_in_heap_def heap_is_wellformed_def
using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast
have "parent_child_rel h = parent_child_rel h'"
proof -
have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting)
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally show ?thesis
by simp
qed
have root: "h \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> local.get_root_node_not_node_same)
then
have root': "h' \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h
local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3)
have "heap_is_wellformed h'"
using create_element_preserves_wellformedness assms returns_result_heap_def
by metis
have "cast result |\<notin>| object_ptr_kinds h"
using \<open>cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> returns_result_heap_def
by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> assms(4) returns_result_eq)
obtain to where to: "h \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to"
by (meson \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(1) assms(2) assms(3)
document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok)
then
have "h \<turnstile> a_get_component (cast document_ptr) \<rightarrow>\<^sub>r to"
using root
by(auto simp add: a_get_component_def)
moreover
obtain to' where to': "h' \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to'"
by (meson \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> is_OK_returns_result_E
local.get_root_node_root_in_heap local.to_tree_order_ok root')
then
have "h' \<turnstile> a_get_component (cast document_ptr) \<rightarrow>\<^sub>r to'"
using root'
by(auto simp add: a_get_component_def)
moreover
have "\<And>child. child \<in> set to \<longleftrightarrow> child \<in> set to'"
by (metis \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>parent_child_rel h = parent_child_rel h'\<close>
\<open>type_wf h'\<close> assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to')
ultimately
have "set |h \<turnstile> local.a_get_component (cast document_ptr)|\<^sub>r = set |h' \<turnstile> local.a_get_component (cast document_ptr)|\<^sub>r"
by(auto simp add: a_get_component_def)
show ?thesis
apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1]
apply (metis \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> assms(2) assms(3)
children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result)
apply (metis \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> assms(4)
element_ptr_kinds_commutes h2 new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2
node_ptr_kinds_eq_h3 returns_result_eq returns_result_heap_def)
using \<open>cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r result |\<notin>| object_ptr_kinds h\<close> element_ptr_kinds_commutes
node_ptr_kinds_commutes apply blast
using assms(1) assms(2) assms(3) \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'\<close>
apply(rule create_element_is_weakly_dom_component_safe_step)
apply (simp add: local.get_component_impl)
using \<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> by auto
qed
end
interpretation i_get_component_create_element?: l_get_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name
set_tag_name_locs create_element
by(auto simp add: l_get_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>create\_character\_data\<close>
lemma create_character_data_not_strongly_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and document_ptr and new_character_data_ptr and val where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> create_character_data document_ptr val \<rightarrow>\<^sub>r new_character_data_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {cast document_ptr} {cast new_character_data_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "create_document"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and document_ptr="?document_ptr"])
by code_simp+
qed
locale l_get_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order
get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name +
l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr +
l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs +
l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_val set_val_locs +
l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val
set_val_locs set_disconnected_nodes set_disconnected_nodes_locs
create_character_data known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_component :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_element_by_id ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_elements_by_class_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_elements_by_tag_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) character_data_ptr) prog"
begin
lemma create_character_data_is_weakly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<noteq> cast |h \<turnstile> create_character_data document_ptr text|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain new_character_data_ptr h2 h3 disc_nodes where
new_character_data_ptr: "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: create_character_data_def
elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
using new_character_data_ptr h2 h3 disc_nodes h'
apply(auto simp add: create_character_data_def
intro!: bind_returns_result_I bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "preserved (get_M ptr getter) h h2"
using h2 new_character_data_ptr
apply(rule new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
using new_character_data_ptr assms(6)
\<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
by simp
have "preserved (get_M ptr getter) h2 h3"
using set_val_writes h3
apply(rule reads_writes_preserved2)
apply(auto simp add: set_val_locs_impl a_set_val_locs_def all_args_def)[1]
by (metis (mono_tags) CharacterData_simp11
\<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> assms(4) assms(6)
is_OK_returns_heap_I is_OK_returns_result_E returns_result_eq select_result_I2)
have "document_ptr |\<in>| document_ptr_kinds h"
using create_character_data_document_in_heap
using assms(4)
by blast
then
have "ptr \<noteq> (cast document_ptr)"
using assms(5) assms(1) assms(2) assms(3) local.get_component_ok local.get_dom_component_ptr
by auto
have "preserved (get_M ptr getter) h3 h'"
using set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved2)
apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1]
by (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
show ?thesis
using \<open>preserved (get_M ptr getter) h h2\<close> \<open>preserved (get_M ptr getter) h2 h3\<close>
\<open>preserved (get_M ptr getter) h3 h'\<close>
by(auto simp add: preserved_def)
qed
lemma create_character_data_is_weakly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>h h'"
shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'"
proof -
obtain new_character_data_ptr h2 h3 disc_nodes_h3 where
new_character_data_ptr: "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: create_character_data_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then
have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_character_data_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h:
"\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2
get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2
get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_child_nodes)
then have children_eq2_h2:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_character_data_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_val_writes h3]
using set_val_types_preserved
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_character_data_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close> using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
a_all_ptrs_in_heap_def heap_is_wellformed_def
using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast
have "parent_child_rel h = parent_child_rel h'"
proof -
have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally show ?thesis
by simp
qed
have root: "h \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> local.get_root_node_not_node_same)
then
have root': "h' \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h
local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3)
have "heap_is_wellformed h'" and "known_ptrs h'"
using create_character_data_preserves_wellformedness assms
by blast+
have "cast result |\<notin>| object_ptr_kinds h"
using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> assms(4) returns_result_eq)
obtain to where to: "h \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to"
by (meson \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(1) assms(2) assms(3)
document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok)
then
have "h \<turnstile> a_get_component (cast document_ptr) \<rightarrow>\<^sub>r to"
using root
by(auto simp add: a_get_component_def)
moreover
obtain to' where to': "h' \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to'"
by (meson \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> is_OK_returns_result_E
local.get_root_node_root_in_heap local.to_tree_order_ok root')
then
have "h' \<turnstile> a_get_component (cast document_ptr) \<rightarrow>\<^sub>r to'"
using root'
by(auto simp add: a_get_component_def)
moreover
have "\<And>child. child \<in> set to \<longleftrightarrow> child \<in> set to'"
by (metis \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>parent_child_rel h = parent_child_rel h'\<close>
\<open>type_wf h'\<close> assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to')
ultimately
have "set |h \<turnstile> local.a_get_component (cast document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_component (cast document_ptr)|\<^sub>r"
by(auto simp add: a_get_component_def)
show ?thesis
apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1]
using assms(2) assms(3) children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap
local.known_ptrs_known_ptr object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result
apply (metis \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
is_OK_returns_result_I)
apply (metis \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> assms(4)
character_data_ptr_kinds_commutes h2 new_character_data_ptr new_character_data_ptr_in_heap
node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 returns_result_eq)
using \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
\<open>new_character_data_ptr \<notin> set |h \<turnstile> character_data_ptr_kinds_M|\<^sub>r\<close> assms(4) returns_result_eq
apply fastforce
using assms(2) assms(3) children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap
local.known_ptrs_known_ptr object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result
apply (smt ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> assms(1) assms(5)
create_character_data_is_weakly_dom_component_safe_step local.get_component_impl select_result_I2)
done
qed
end
interpretation i_get_component_create_character_data?: l_get_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_val set_val_locs
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs
create_character_data
by(auto simp add: l_get_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>create\_document\<close>
lemma create_document_not_strongly_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and new_document_ptr where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> create_document \<rightarrow>\<^sub>r new_document_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {} {cast new_document_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "create_document"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1"])
by code_simp+
qed
locale l_get_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order
get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name +
l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_component :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_element_by_id ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_elements_by_class_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_elements_by_tag_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and create_document :: "((_) heap, exception, (_) document_ptr) prog"
begin
lemma create_document_is_weakly_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>h h'"
assumes "ptr \<noteq> cast |h \<turnstile> create_document|\<^sub>r"
shows "preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
using assms
apply(auto simp add: create_document_def)[1]
by (metis assms(4) assms(5) is_OK_returns_heap_I local.create_document_def new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
select_result_I)
lemma create_document_is_weakly_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>h h'"
shows "is_weakly_dom_component_safe {} {cast result} h h'"
proof -
have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast result|}"
using assms(4) assms(5) local.create_document_def new_document_new_ptr by auto
moreover have "result |\<notin>| document_ptr_kinds h"
using assms(4) assms(5) local.create_document_def new_document_ptr_not_in_heap by auto
ultimately show ?thesis
using assms
apply(auto simp add: is_weakly_dom_component_safe_def Let_def local.create_document_def
new_document_ptr_not_in_heap)[1]
by (metis \<open>result |\<notin>| document_ptr_kinds h\<close> document_ptr_kinds_commutes new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
qed
end
interpretation i_get_component_create_document?: l_get_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name create_document
by(auto simp add: l_get_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>insert\_before\<close>
lemma insert_before_not_strongly_dom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and ptr and child and ref where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> insert_before ptr child ref \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe ({ptr, cast child} \<union> (cast ` set_option ref)) {} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
e2 \<leftarrow> create_element document_ptr ''div'';
return (e1, e2)
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e1 = "fst |?h0 \<turnstile> ?P|\<^sub>r"
let ?e2 = "snd |?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1" and
child="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e2" and ref=None])
by code_simp+
qed
lemma append_child_not_strongly_dom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and ptr and child where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {ptr, cast child} {} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
e2 \<leftarrow> create_element document_ptr ''div'';
return (e1, e2)
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e1 = "fst |?h0 \<turnstile> ?P|\<^sub>r"
let ?e2 = "snd |?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1" and child="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e2"])
by code_simp+
qed
subsection \<open>get\_owner\_document\<close>
lemma get_owner_document_not_strongly_dom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and ptr and owner_document where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {ptr} {cast owner_document} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "do {
doc \<leftarrow> create_document;
create_element doc ''div''
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e1 = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1"])
by code_simp+
qed
end
diff --git a/thys/Decl_Sem_Fun_PL/DenotCompleteFSet.thy b/thys/Decl_Sem_Fun_PL/DenotCompleteFSet.thy
--- a/thys/Decl_Sem_Fun_PL/DenotCompleteFSet.thy
+++ b/thys/Decl_Sem_Fun_PL/DenotCompleteFSet.thy
@@ -1,323 +1,323 @@
section "Completeness of the declarative semantics wrt. operational"
theory DenotCompleteFSet
imports ChangeEnv SmallStepLam DenotSoundFSet
begin
subsection "Reverse substitution preserves denotation"
fun join :: "val \<Rightarrow> val \<Rightarrow> val option" (infix "\<squnion>" 60) where
"(VNat n) \<squnion> (VNat n') = (if n = n' then Some (VNat n) else None)" |
"(VFun f) \<squnion> (VFun f') = Some (VFun (f |\<union>| f'))" |
"v \<squnion> v' = None"
lemma combine_values:
assumes vv: "isval v" and v1v: "v1 \<in> E v \<rho>" and v2v: "v2 \<in> E v \<rho>"
shows " \<exists> v3. v3 \<in> E v \<rho> \<and> (v1 \<squnion> v2 = Some v3)"
using vv v1v v2v by (induction v arbitrary: v1 v2 \<rho>) auto
lemma le_union1: fixes v1::val assumes v12: "v1 \<squnion> v2 = Some v12" shows "v1 \<sqsubseteq> v12"
proof (cases v1)
case (VNat n1) hence v1: "v1=VNat n1" by simp
show ?thesis
proof (cases v2)
case (VNat n2) with v1 v12 show ?thesis by (cases "n1=n2") auto
next
case (VFun x2) with v1 v12 show ?thesis by auto
qed
next
case (VFun t2) from VFun have v1: "v1=VFun t2" by simp
show ?thesis
proof (cases v2)
case (VNat n1) with v1 v12 show ?thesis by auto
next
case (VFun n2) with v1 v12 show ?thesis by auto
qed
qed
lemma le_union2: "v1 \<squnion> v2 = Some v12 \<Longrightarrow> v2 \<sqsubseteq> v12"
apply (cases v1)
apply (cases v2)
apply auto
apply (rename_tac x1 x1')
apply (case_tac "x1 = x1'")
apply auto
apply (cases v2)
apply auto
done
lemma le_union_left: "\<lbrakk> v1 \<squnion> v2 = Some v12; v1 \<sqsubseteq> v3; v2 \<sqsubseteq> v3 \<rbrakk> \<Longrightarrow> v12 \<sqsubseteq> v3"
apply (cases v1) apply (cases v2) apply force+ done
lemma e_val: "isval v \<Longrightarrow> \<exists> v'. v' \<in> E v \<rho>"
- apply (case_tac v) apply auto apply (rule_tac x="{||}" in exI) apply force done
+ by (cases v) auto
lemma reverse_subst_lam:
assumes fl: "VFun f \<in> E (ELam x e) \<rho>"
and vv: "is_val v" and ls: "ELam x e = ELam x (subst y v e')" and xy: "x \<noteq> y"
and IH: "\<forall> v1 v2. v2 \<in> E (subst y v e') ((x,v1)#\<rho>)
\<longrightarrow> (\<exists> \<rho>' v'. v' \<in> E v [] \<and> v2 \<in> E e' \<rho>' \<and> \<rho>' \<approx> (y,v')#(x,v1)#\<rho>)"
shows "\<exists> \<rho>' v''. v'' \<in> E v [] \<and> VFun f \<in> E (ELam x e') \<rho>' \<and> \<rho>' \<approx> ((y,v'')#\<rho>)"
using fl vv ls IH xy
proof (induction f arbitrary: x e e' \<rho> v y)
case empty
from empty(2) is_val_def obtain v' where vp_v: "v' \<in> E v []" using e_val[of v "[]"] by blast
let ?R = "(y,v')#\<rho>"
have 1: "VFun {||} \<in> E (ELam x e') ?R" by simp
have 2: "?R \<approx> (y, v') # \<rho>" by auto
from vp_v 1 2 show ?case by blast
next
case (insert a f x e e' \<rho> v y)
from insert(3) have 1: "VFun f \<in> E (ELam x e) \<rho>" by auto
obtain v1 v2 where a: "a = (v1,v2)" by (cases a) simp
from insert 1 have "\<exists> \<rho>' v''. v'' \<in> E v [] \<and> VFun f \<in> E (ELam x e') \<rho>' \<and> \<rho>' \<approx> ((y,v'')#\<rho>)"
- by blast
+ by metis
from this obtain \<rho>'' v'' where vpp_v: "v'' \<in> E v []" and f_l: "VFun f \<in> E (ELam x e') \<rho>''"
and rpp_r: "\<rho>'' \<approx> ((y,v'')#\<rho>)" by blast
from insert(3) a have v2_e: "v2 \<in> E e ((x,v1)#\<rho>)" using e_lam_elim2 by blast
from insert v2_e have "\<exists>\<rho>'' v'. v' \<in> E v [] \<and> v2 \<in> E e' \<rho>'' \<and> \<rho>'' \<approx> (y, v')#(x, v1)#\<rho>" by auto
from this obtain \<rho>3 v' where vp_v: "v' \<in> E v []" and v2_ep: "v2 \<in> E e' \<rho>3"
and r3: "\<rho>3 \<approx> (y,v') # (x,v1) # \<rho>" by blast
from insert(4) have "isval v" by auto
from this vp_v vpp_v obtain v3 where v3_v: "v3 \<in> E v []" and vp_vpp: "v' \<squnion> v'' = Some v3"
using combine_values by blast
have 4: "VFun (finsert a f) \<in> E (ELam x e') ((y, v3) # \<rho>)"
proof -
from vp_vpp have v3_vpp: "v'' \<sqsubseteq> v3" using le_union2 by simp
from rpp_r v3_vpp have "\<rho>'' \<sqsubseteq> (y,v3)#\<rho>" by (simp add: env_eq_def env_le_def)
from f_l this have 2: "VFun f \<in> E (ELam x e') ((y, v3) # \<rho>)" by (rule raise_env)
from vp_vpp have vp_v3: "v' \<sqsubseteq> v3" using le_union1 by simp
from vp_v3 r3 insert have "\<rho>3 \<sqsubseteq> (x,v1)#(y,v3)#\<rho>" by (simp add: env_eq_def env_le_def)
from v2_ep this have 3: "v2 \<in> E e' ((x,v1)#(y,v3)#\<rho>)" by (rule raise_env)
from 2 3 a show "?thesis" using e_lam_intro2 by blast
qed
have 5: "(y, v3) # \<rho> \<approx> (y, v3) # \<rho>" by auto
from v3_v 4 5 show ?case by blast
qed
lemma lookup_ext_none: "\<lbrakk> lookup \<rho> y = None; x \<noteq> y \<rbrakk> \<Longrightarrow> lookup ((x,v)#\<rho>) y = None"
by auto
\<comment> \<open>For reverse subst lemma, the variable case shows up over and over, so we prove it as a lemma\<close>
lemma rev_subst_var:
assumes ev: "e = EVar y \<and> v = e'" and vv: "is_val v" and vp_E: "v' \<in> E e' \<rho>"
shows "\<exists> \<rho>' v''. v'' \<in> E v [] \<and> v' \<in> E e \<rho>' \<and> \<rho>' \<approx> ((y,v'')#\<rho>)"
proof -
from vv have lx: "\<forall> x. x \<in> FV v \<longrightarrow> lookup [] x = lookup \<rho> x" by auto
from ev vp_E lx env_strengthen[of v' v \<rho> "[]"] have n_Ev: "v' \<in> E v []" by blast
have ly: "lookup ((y,v')#\<rho>) y = Some v'" by simp
from env_eq_def have rr: "((y,v')#\<rho>) \<approx> ((y,v')#\<rho>)" by simp
from ev ly have n_Ee: "v' \<in> E e ((y,v')#\<rho>)" by simp
from n_Ev rr n_Ee show ?thesis by blast
qed
lemma reverse_subst_pres_denot:
assumes vep: "v' \<in> E e' \<rho>" and vv: "is_val v" and ep: "e' = subst y v e"
shows "\<exists> \<rho>' v''. v'' \<in> E v [] \<and> v' \<in> E e \<rho>' \<and> \<rho>' \<approx> ((y,v'')#\<rho>)"
using vep vv ep
proof (induction arbitrary: v' y v e rule: E.induct)
case (1 n \<rho>) \<comment> \<open>e' = ENat n\<close>
from 1(1) have vp: "v' = VNat n" by auto
from 1(3) have "e = ENat n \<or> (e = EVar y \<and> v = ENat n)" by (cases e, auto)
then show ?case
proof
assume e: "e = ENat n"
from 1(2) e_val is_val_def obtain v'' where vpp_E: "v'' \<in> E v []" by force
from env_eq_def have rr: "((y,v'')#\<rho>) \<approx> ((y,v'')#\<rho>)" by simp
from vp e have vp_E: "v' \<in> E e ((y,v'')#\<rho>)" by simp
from vpp_E vp_E rr show ?thesis by blast
next
assume ev: "e = EVar y \<and> v = ENat n"
from ev 1(2) 1(1) rev_subst_var show ?thesis by blast
qed
next
case (2 x \<rho>) \<comment> \<open>e' = EVar x\<close>
from 2 have e: "e = EVar x" by (cases e, auto)
from 2 e have xy: "x \<noteq> y" by force
from 2(2) e_val is_val_def obtain v'' where vpp_E: "v'' \<in> E v []" by force
from env_eq_def have rr: "((y,v'')#\<rho>) \<approx> ((y,v'')#\<rho>)" by simp
from 2(1) obtain vx where lx: "lookup \<rho> x = Some vx" and vp_vx: "v' \<sqsubseteq> vx" by auto
from e lx vp_vx xy have vp_E: "v' \<in> E e ((y,v'')#\<rho>)" by simp
from vpp_E rr vp_E show ?case by blast
next
case (3 x eb \<rho>)
{ assume ev: "e = EVar y \<and> v = ELam x eb"
- from ev 3(3) 3(2) rev_subst_var have ?case by blast
+ from ev 3(3) 3(2) rev_subst_var have ?case by metis
} also { assume ex: "e = ELam x eb \<and> x = y"
from 3(3) e_val is_val_def obtain v'' where vpp_E: "v'' \<in> E v []" by force
from env_eq_def have rr: "((y,v'')#\<rho>) \<approx> ((y,v'')#\<rho>)" by simp
from ex have lz: "\<forall> z. z \<in> FV (ELam x eb) \<longrightarrow> lookup ((y,v'')#\<rho>) z = lookup \<rho> z" by auto
from ex 3(2) lz env_strengthen[of v' "ELam x eb" \<rho> "(y,v'')#\<rho>"]
have vp_E: "v' \<in> E e ((y,v'')#\<rho>)" by blast
from vpp_E vp_E rr have ?case by blast
} moreover { assume exb: "\<exists> e'. e = ELam x e' \<and> x \<noteq> y \<and> eb = subst y v e'"
from exb obtain e'' where e: "e = ELam x e''" and xy: "x \<noteq> y"
and eb: "eb = subst y v e''" by blast
from 3(2) obtain f where vp: "v' = VFun f" by auto
from 3(2) vp have f_E: "VFun f \<in> E (ELam x eb) \<rho>" by simp
from 3(4) e xy have ls: "ELam x eb = ELam x (subst y v e'')" by simp
from 3(3) eb have IH: "\<forall> v1 v2. v2 \<in> E (subst y v e'') ((x,v1)#\<rho>)
\<longrightarrow> (\<exists> \<rho>' v'. v' \<in> E v [] \<and> v2 \<in> E e'' \<rho>' \<and> \<rho>' \<approx> (y,v')#(x,v1)#\<rho>)"
apply clarify apply (subgoal_tac "(v1,v2) \<in> fset {|(v1,v2)|}") prefer 2 apply simp
apply (rule 3(1)) apply assumption apply simp+ done
from f_E 3(3) ls xy IH e vp have ?case apply clarify apply (rule reverse_subst_lam)
apply blast+ done
} moreover from 3(4) have "(e = EVar y \<and> v = ELam x eb)
\<or> (e = ELam x eb \<and> x = y)
\<or> (\<exists> e'. e = ELam x e' \<and> x \<noteq> y \<and> eb = subst y v e')" by (cases e) auto
ultimately show ?case by blast
next
case (4 e1 e2 \<rho>) \<comment> \<open>e' = EApp e1 e2\<close>
from 4(4) 4(5) obtain e1' e2' where
e: "e = EApp e1' e2'" and e1:"e1 = subst y v e1'" and e2: "e2 = subst y v e2'"
apply (cases e) apply (rename_tac x) apply auto apply (case_tac "y = x") apply auto
apply (rename_tac x1 x2) apply (case_tac "y = x1") apply auto done
from 4(3) obtain f v2 and v2'::val and v3' where
f_E: "VFun f \<in> E e1 \<rho>" and v2_E: "v2 \<in> E e2 \<rho>" and v23: "(v2',v3') \<in> fset f"
and v2p_v2: "v2' \<sqsubseteq> v2" and vp_v3: "v' \<sqsubseteq> v3'" by blast
from 4(1) f_E 4(4) e1 obtain \<rho>1 w1 where v1_Ev: "w1 \<in> E v []" and f_E1: "VFun f \<in> E e1' \<rho>1"
and r1: "\<rho>1 \<approx> (y,w1)#\<rho>" by blast
from 4(2) v2_E 4(4) e2 obtain \<rho>2 w2 where v2_Ev: "w2 \<in> E v []" and v2_E2: "v2 \<in> E e2' \<rho>2"
and r2: "\<rho>2 \<approx> (y,w2)#\<rho>" by blast
from 4(4) v1_Ev v2_Ev combine_values obtain w3 where
w3_Ev: "w3 \<in> E v []" and w123: "w1 \<squnion> w2 = Some w3" by (simp only: is_val_def) blast
from w123 le_union1 have w13: "w1 \<sqsubseteq> w3" by blast
from w123 le_union2 have w23: "w2 \<sqsubseteq> w3" by blast
from w13 have r13: "((y,w1)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from w23 have r23: "((y,w2)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from r1 f_E1 have f_E1b: "VFun f \<in> E e1' ((y,w1)#\<rho>)" by (rule env_swap)
from f_E1b r13 have f_E1c: "VFun f \<in> E e1' ((y,w3)#\<rho>)" by (rule raise_env)
from r2 v2_E2 have v2_E2b: "v2 \<in> E e2' ((y,w2)#\<rho>)" by (rule env_swap)
from v2_E2b r23 have v2_E2c: "v2 \<in> E e2' ((y,w3)#\<rho>)" by (rule raise_env)
from f_E1c v2_E2c v23 v2p_v2 vp_v3 have vp_E2: "v' \<in> E (EApp e1' e2') ((y,w3)#\<rho>)" by blast
have rr3: "((y,w3)#\<rho>) \<approx> ((y,w3)#\<rho>)" by (simp add: env_eq_def)
from w3_Ev vp_E2 rr3 e show ?case by blast
next
case (5 f e1 e2 \<rho>) \<comment> \<open>e' = EPrim f e1 e2, very similar to case for EApp\<close>
from 5(4) 5(5) obtain e1' e2' where
e: "e = EPrim f e1' e2'" and e1:"e1 = subst y v e1'" and e2: "e2 = subst y v e2'"
apply (cases e) apply auto apply (rename_tac x) apply (case_tac "y = x") apply auto
apply (rename_tac x1 x2) apply (case_tac "y = x1") apply auto done
from 5(3) obtain n1 n2 where
n1_E: "VNat n1 \<in> E e1 \<rho>" and n2_E: "VNat n2 \<in> E e2 \<rho>" and vp: "v' = VNat (f n1 n2)" by blast
from 5(1) n1_E 5(4) e1 obtain \<rho>1 w1 where v1_Ev: "w1 \<in> E v []" and n1_E1: "VNat n1 \<in> E e1' \<rho>1"
and r1: "\<rho>1 \<approx> (y,w1)#\<rho>" by blast
from 5(2) n2_E 5(4) e2 obtain \<rho>2 w2 where v2_Ev: "w2 \<in> E v []" and n2_E2: "VNat n2 \<in> E e2' \<rho>2"
and r2: "\<rho>2 \<approx> (y,w2)#\<rho>" by blast
from 5(4) v1_Ev v2_Ev combine_values obtain w3 where
w3_Ev: "w3 \<in> E v []" and w123: "w1 \<squnion> w2 = Some w3" by (simp only: is_val_def) blast
from w123 le_union1 have w13: "w1 \<sqsubseteq> w3" by blast
from w123 le_union2 have w23: "w2 \<sqsubseteq> w3" by blast
from w13 have r13: "((y,w1)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from w23 have r23: "((y,w2)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from r1 n1_E1 have n1_E1b: "VNat n1 \<in> E e1' ((y,w1)#\<rho>)" by (rule env_swap)
from n1_E1b r13 have n1_E1c: "VNat n1 \<in> E e1' ((y,w3)#\<rho>)" by (rule raise_env)
from r2 n2_E2 have n2_E2b: "VNat n2 \<in> E e2' ((y,w2)#\<rho>)" by (rule env_swap)
from n2_E2b r23 have v2_E2c: "VNat n2 \<in> E e2' ((y,w3)#\<rho>)" by (rule raise_env)
from n1_E1c v2_E2c vp have vp_E2: "v' \<in> E (EPrim f e1' e2') ((y,w3)#\<rho>)" by blast
have rr3: "((y,w3)#\<rho>) \<approx> ((y,w3)#\<rho>)" by (simp add: env_eq_def)
from w3_Ev vp_E2 rr3 e show ?case by blast
next
case (6 e1 e2 e3 \<rho>) \<comment> \<open>e' = EIf e1 e2 e3\<close>
from 6(5) 6(6) obtain e1' e2' e3' where
e: "e = EIf e1' e2' e3'" and e1:"e1 = subst y v e1'" and e2: "e2 = subst y v e2'"
and e3: "e3 = subst y v e3'"
apply (cases e) apply auto apply (case_tac "y=x1") apply auto apply (case_tac "y=x31") by auto
from 6(4) e_if_elim obtain n where n_E: "VNat n \<in> E e1 \<rho>" and
els: "n = 0 \<longrightarrow> v' \<in> E e3 \<rho>" and thn: "n \<noteq> 0 \<longrightarrow> v' \<in> E e2 \<rho>" by blast
from 6 n_E e1 obtain \<rho>1 w1 where w1_Ev: "w1 \<in> E v []" and n_E2: "VNat n \<in> E e1' \<rho>1"
and r1: "\<rho>1 \<approx> (y,w1)#\<rho>" by blast
show ?case
proof (cases "n = 0")
case True with els have vp_E2: "v' \<in> E e3 \<rho>" by simp
from 6 vp_E2 e3 obtain \<rho>2 w2 where w2_Ev: "w2 \<in> E v []" and vp_E2: "v' \<in> E e3' \<rho>2"
and r2: "\<rho>2 \<approx> (y,w2)#\<rho>" by blast
from 6(5) w1_Ev w2_Ev combine_values obtain w3 where
w3_Ev: "w3 \<in> E v []" and w123: "w1 \<squnion> w2 = Some w3" by (simp only: is_val_def) blast
from w123 le_union1 have w13: "w1 \<sqsubseteq> w3" by blast
from w123 le_union2 have w23: "w2 \<sqsubseteq> w3" by blast
from w13 have r13: "((y,w1)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from w23 have r23: "((y,w2)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from r1 n_E2 have n_E1b: "VNat n \<in> E e1' ((y,w1)#\<rho>)" by (rule env_swap)
from n_E1b r13 have n_E1c: "VNat n \<in> E e1' ((y,w3)#\<rho>)" by (rule raise_env)
from r2 vp_E2 have vp_E2b: "v' \<in> E e3' ((y,w2)#\<rho>)" by (rule env_swap)
from vp_E2b r23 have vp_E2c: "v' \<in> E e3' ((y,w3)#\<rho>)" by (rule raise_env)
have rr3: "((y,w3)#\<rho>) \<approx> ((y,w3)#\<rho>)" by (simp add: env_eq_def)
from True n_E1c vp_E2c e have vp_E3: "v' \<in> E e ((y,w3)#\<rho>)" by auto
from w3_Ev rr3 vp_E3 show ?thesis by blast
next
case False with thn have vp_E2: "v' \<in> E e2 \<rho>" by simp
from 6 vp_E2 e2 obtain \<rho>2 w2 where w2_Ev: "w2 \<in> E v []" and vp_E2: "v' \<in> E e2' \<rho>2"
and r2: "\<rho>2 \<approx> (y,w2)#\<rho>" by blast
from 6(5) w1_Ev w2_Ev combine_values obtain w3 where
w3_Ev: "w3 \<in> E v []" and w123: "w1 \<squnion> w2 = Some w3" by (simp only: is_val_def) blast
from w123 le_union1 have w13: "w1 \<sqsubseteq> w3" by blast
from w123 le_union2 have w23: "w2 \<sqsubseteq> w3" by blast
from w13 have r13: "((y,w1)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from w23 have r23: "((y,w2)#\<rho>) \<sqsubseteq> ((y,w3)#\<rho>)" by (simp add: env_le_def)
from r1 n_E2 have n_E1b: "VNat n \<in> E e1' ((y,w1)#\<rho>)" by (rule env_swap)
from n_E1b r13 have n_E1c: "VNat n \<in> E e1' ((y,w3)#\<rho>)" by (rule raise_env)
from r2 vp_E2 have vp_E2b: "v' \<in> E e2' ((y,w2)#\<rho>)" by (rule env_swap)
from vp_E2b r23 have vp_E2c: "v' \<in> E e2' ((y,w3)#\<rho>)" by (rule raise_env)
have rr3: "((y,w3)#\<rho>) \<approx> ((y,w3)#\<rho>)" by (simp add: env_eq_def)
from False n_E1c vp_E2c e have vp_E3: "v' \<in> E e ((y,w3)#\<rho>)" by auto
from w3_Ev rr3 vp_E3 show ?thesis by blast
qed
qed
subsection "Reverse reduction preserves denotation"
lemma reverse_step_pres_denot:
fixes e::exp assumes e_ep: "e \<longrightarrow> e'" and v_ep: "v \<in> E e' \<rho>"
shows "v \<in> E e \<rho>"
using e_ep v_ep
proof (induction arbitrary: v \<rho> rule: reduce.induct)
case (beta v x e v' \<rho>)
from beta obtain \<rho>' v'' where 1: "v'' \<in> E v []" and 2: "v' \<in> E e \<rho>'" and 3: "\<rho>' \<approx> (x, v'') # \<rho>"
using reverse_subst_pres_denot[of v' "subst x v e" \<rho> v x e] by blast
from beta 1 2 3 show ?case
apply simp apply (rule_tac x="{|(v'',v')|}" in exI) apply (rule conjI)
- apply clarify apply simp apply clarify apply (rule env_swap) apply blast apply blast
+ apply clarify apply simp apply (rule env_swap) apply blast apply blast
apply (rule_tac x=v'' in exI) apply (rule conjI) apply (rule env_strengthen)
apply assumption apply force apply force done
qed auto
lemma reverse_multi_step_pres_denot:
fixes e::exp assumes e_ep: "e \<longrightarrow>* e'" and v_ep: "v \<in> E e' \<rho>" shows "v \<in> E e \<rho>"
using e_ep v_ep reverse_step_pres_denot
by (induction arbitrary: v \<rho> rule: multi_step.induct) auto
subsection "Completeness"
theorem completeness:
assumes ev: "e \<longrightarrow>* v"and vv: "is_val v"
shows "\<exists> v'. v' \<in> E e \<rho> \<and> v' \<in> E v []"
proof -
from vv have "\<exists> v'. v' \<in> E v []" using e_val by auto
from this obtain v' where vp_v: "v' \<in> E v []" by blast
from vp_v vv have vp_v2: "v' \<in> E v \<rho>" using env_strengthen by force
from ev vp_v2 reverse_multi_step_pres_denot[of e v v' \<rho>]
have "v' \<in> E e \<rho>" by simp
from this vp_v show ?thesis by blast
qed
theorem reduce_pres_denot: fixes e::exp assumes r: "e \<longrightarrow> e'" shows "E e = E e'"
apply (rule ext) apply (rule equalityI) apply (rule subsetI)
apply (rule subject_reduction) apply assumption using r apply assumption
apply (rule subsetI)
using r apply (rule reverse_step_pres_denot) apply assumption
done
theorem multi_reduce_pres_denot: fixes e::exp assumes r: "e \<longrightarrow>* e'" shows "E e = E e'"
using r reduce_pres_denot by induction auto
theorem complete_wrt_op_sem:
assumes e_n: "e \<Down> ONat n" shows "E e [] = E (ENat n) []"
proof -
from e_n have 1: "e \<longrightarrow>* ENat n"
unfolding run_def apply simp apply (erule exE)
apply (rename_tac v) apply (case_tac v) apply auto done
from 1 show ?thesis using multi_reduce_pres_denot by simp
qed
end
diff --git a/thys/Decl_Sem_Fun_PL/DenotSoundFSet.thy b/thys/Decl_Sem_Fun_PL/DenotSoundFSet.thy
--- a/thys/Decl_Sem_Fun_PL/DenotSoundFSet.thy
+++ b/thys/Decl_Sem_Fun_PL/DenotSoundFSet.thy
@@ -1,410 +1,411 @@
section "Soundness of the declarative semantics wrt. operational"
theory DenotSoundFSet
imports SmallStepLam BigStepLam ChangeEnv
begin
subsection "Substitution preserves denotation"
lemma subst_app: "subst x v (EApp e1 e2) = EApp (subst x v e1) (subst x v e2)"
by auto
lemma subst_prim: "subst x v (EPrim f e1 e2) = EPrim f (subst x v e1) (subst x v e2)"
by auto
lemma subst_lam_eq: "subst x v (ELam x e) = ELam x e" by auto
lemma subst_lam_neq: "y \<noteq> x \<Longrightarrow> subst x v (ELam y e) = ELam y (subst x v e)" by simp
lemma subst_if: "subst x v (EIf e1 e2 e3) = EIf (subst x v e1) (subst x v e2) (subst x v e3)"
by auto
lemma substitution:
fixes \<Gamma>::env and A::val
assumes wte: "B \<in> E e \<Gamma>'" and wtv: "A \<in> E v []"
and gp: "\<Gamma>' \<approx> (x,A)#\<Gamma>" and v: "is_val v"
shows "B \<in> E (subst x v e) \<Gamma>"
using wte wtv gp v
proof (induction arbitrary: v A B \<Gamma> x rule: E.induct)
case (1 n \<rho>)
then show ?case by auto
next
case (2 x \<rho> v A B \<Gamma> x')
then show ?case
apply (simp only: env_eq_def)
apply (cases "x = x'")
apply simp apply clarify
apply (rule env_strengthen)
apply (rule e_sub)
apply auto
done
next
case (3 x e \<rho> v A B \<Gamma> x')
then show ?case
apply (case_tac "x' = x") apply (simp only: subst_lam_eq)
apply (rule env_strengthen) apply assumption apply (simp add: env_eq_def)
apply (simp only: subst_lam_neq) apply (erule e_lam_elim)
apply (rule e_lam_intro)
apply assumption apply clarify apply (erule_tac x=v1 in allE) apply (erule_tac x=v2 in allE)
apply clarify
apply (subgoal_tac "(x,v1)#\<rho> \<approx> (x',A)#(x,v1)#\<Gamma>")
prefer 2 apply (simp add: env_eq_def)
apply blast
done
next
case (4 e1 e2 \<rho>)
then show ?case
apply (simp only: subst_app)
apply (erule e_app_elim)
apply (rule e_app_intro)
apply auto
done
next
case (5 f e1 e2 \<rho>)
then show ?case
apply (simp only: subst_prim) apply (erule e_prim_elim) apply simp
apply (rule_tac x=n1 in exI) apply (rule conjI)
apply force
apply (rule_tac x=n2 in exI)
apply auto
done
next
case (6 e1 e2 e3 \<rho>)
then show ?case
apply (simp only: subst_if) apply (erule e_if_elim) apply (rename_tac n)
apply simp
apply (case_tac "n = 0") apply (rule_tac x=0 in exI)
apply force
apply (rule_tac x=n in exI) apply simp done
qed
subsection "Reduction preserves denotation"
lemma subject_reduction: fixes e::exp assumes v: "v \<in> E e \<rho>" and r: "e \<longrightarrow> e'" shows "v \<in> E e' \<rho>"
using r v
proof (induction arbitrary: v \<rho> rule: reduce.induct)
case (beta v x e v' \<rho>)
then show ?case apply (simp only: is_val_def)
apply (erule e_app_elim) apply (erule e_lam_elim) apply clarify
apply (rename_tac f v2 v2' v3' f')
apply (erule_tac x=v2' in allE) apply (erule_tac x=v3' in allE) apply clarify
apply (subgoal_tac "v3' \<in> E (subst x v e) \<rho>") prefer 2 apply (rule substitution)
apply (subgoal_tac "v3' \<in> E e ((x,v2)#\<rho>)") prefer 2 apply (rule raise_env)
apply assumption apply (simp add: env_le_def) prefer 2 apply (rule env_strengthen)
apply assumption apply force prefer 2 apply (subgoal_tac "(x,v2)#\<rho> \<approx> (x,v2)#\<rho>") prefer 2
apply (simp add: env_eq_def) apply assumption apply assumption
apply simp
apply simp
apply (rule e_sub)
apply assumption
apply (rule val_le_trans)
apply blast
apply force
done
qed force+
theorem preservation: assumes v: "v \<in> E e \<rho>" and rr: "e \<longrightarrow>* e'" shows "v \<in> E e' \<rho>"
using rr v subject_reduction by (induction arbitrary: \<rho> v) auto
lemma canonical_nat: assumes v: "VNat n \<in> E v \<rho>" and vv: "isval v" shows "v = ENat n"
using v vv by (cases v) auto
lemma canonical_fun: assumes v: "VFun f \<in> E v \<rho>" and vv: "isval v" shows "\<exists> x e. v = ELam x e"
using v vv by (cases v) auto
subsection "Progress"
theorem progress: assumes v: "v \<in> E e \<rho>" and r: "\<rho> = []" and fve: "FV e = {}"
shows "is_val e \<or> (\<exists> e'. e \<longrightarrow> e')"
using v r fve
proof (induction arbitrary: v rule: E.induct)
case (4 e1 e2 \<rho>)
show ?case
apply (rule e_app_elim) using 4(3) apply assumption
apply (cases "is_val e1")
apply (cases "is_val e2")
apply (frule canonical_fun) apply force apply (erule exE)+ apply simp apply (rule disjI2)
apply (rename_tac x e)
apply (rule_tac x="subst x e2 e" in exI)
apply (rule beta) apply simp
using 4 apply simp
apply blast
using 4 apply simp
apply blast done
next
case (5 f e1 e2 \<rho>)
show ?case
apply (rule e_prim_elim) using 5(3) apply assumption
using 5 apply (case_tac "isval e1")
apply (case_tac "isval e2")
apply (subgoal_tac "e1 = ENat n1") prefer 2 using canonical_nat apply blast
apply (subgoal_tac "e2 = ENat n2") prefer 2 using canonical_nat apply blast
apply force
apply force
apply force done
next
case (6 e1 e2 e3 \<rho>)
show ?case
apply (rule e_if_elim)
using 6(4) apply assumption
apply (cases "isval e1")
apply (rename_tac n)
apply (subgoal_tac "e1 = ENat n") prefer 2 apply (rule canonical_nat) apply blast apply blast
apply (rule disjI2) apply (case_tac "n = 0") apply force apply force
apply (rule disjI2)
using 6 apply (subgoal_tac "\<exists> e1'. e1 \<longrightarrow> e1'") prefer 2 apply force
apply clarify apply (rename_tac e1')
apply (rule_tac x="EIf e1' e2 e3" in exI)
apply (rule if_cond) apply assumption
done
qed auto
subsection "Logical relation between values and big-step values"
fun good_entry :: "name \<Rightarrow> exp \<Rightarrow> benv \<Rightarrow> (val \<times> bval set) \<times> (val \<times> bval set) \<Rightarrow> bool \<Rightarrow> bool" where
"good_entry x e \<rho> ((v1,g1),(v2,g2)) r = ((\<forall> v \<in> g1. \<exists> v'. (x,v)#\<rho> \<turnstile> e \<Down> v' \<and> v' \<in> g2) \<and> r)"
primrec good :: "val \<Rightarrow> bval set" where
Gnat: "good (VNat n) = { BNat n }" |
Gfun: "good (VFun f) = { vc. \<exists> x e \<rho>. vc = BClos x e \<rho>
\<and> (ffold (good_entry x e \<rho>) True (fimage (map_prod (\<lambda>v. (v,good v)) (\<lambda>v. (v,good v))) f)) }"
inductive good_env :: "benv \<Rightarrow> env \<Rightarrow> bool" where
genv_nil[intro!]: "good_env [] []" |
genv_cons[intro!]: "\<lbrakk> v \<in> good v'; good_env \<rho> \<rho>' \<rbrakk> \<Longrightarrow> good_env ((x,v)#\<rho>) ((x,v')#\<rho>')"
inductive_cases
genv_any_nil_inv: "good_env \<rho> []" and
genv_any_cons_inv: "good_env \<rho> (b#\<rho>')"
lemma lookup_good:
assumes l: "lookup \<rho>' x = Some A" and EE: "good_env \<rho> \<rho>'"
shows "\<exists> v. lookup \<rho> x = Some v \<and> v \<in> good A"
using l EE
proof (induction \<rho>' arbitrary: x A \<rho>)
case Nil
show ?case apply (rule genv_any_nil_inv) using Nil by auto
next
case (Cons a \<rho>')
show ?case
apply (rule genv_any_cons_inv)
using Cons apply force
apply (rename_tac x') apply clarify
using Cons apply (case_tac "x = x'")
apply force
apply force
done
qed
abbreviation good_prod :: "val \<times> val \<Rightarrow> (val \<times> bval set) \<times> (val \<times> bval set)" where
"good_prod \<equiv> map_prod (\<lambda>v. (v,good v)) (\<lambda>v. (v,good v))"
lemma good_prod_inj: "inj_on good_prod (fset A)"
unfolding inj_on_def apply auto done
definition good_fun :: "func \<Rightarrow> name \<Rightarrow> exp \<Rightarrow> benv \<Rightarrow> bool" where
"good_fun f x e \<rho> \<equiv> (ffold (good_entry x e \<rho>) True (fimage good_prod f))"
lemma good_fun_def2:
"good_fun f x e \<rho> = ffold (good_entry x e \<rho> \<circ> good_prod) True f"
proof -
interpret ge: comp_fun_commute "(good_entry x e \<rho>) \<circ> good_prod"
unfolding comp_fun_commute_def by auto
show "good_fun f x e \<rho>
= ffold ((good_entry x e \<rho>) \<circ> good_prod) True f"
using good_prod_inj[of "f"] good_fun_def
ffold_fimage[of good_prod "f" "good_entry x e \<rho>" True] by auto
qed
lemma gfun_elim: "w \<in> good (VFun f) \<Longrightarrow> \<exists> x e \<rho>. w = BClos x e \<rho> \<and> good_fun f x e \<rho>"
using good_fun_def by auto
lemma gfun_mem_iff: "good_fun f x e \<rho> = (\<forall> v1 v2. (v1,v2) \<in> fset f \<longrightarrow>
(\<forall> v \<in> good v1. \<exists> v'. (x,v)#\<rho> \<turnstile> e \<Down> v' \<and> v' \<in> good v2))"
proof (induction f arbitrary: x e \<rho>)
case empty
interpret ge: comp_fun_commute "(good_entry x e \<rho>)"
unfolding comp_fun_commute_def by auto
- from empty show ?case using good_fun_def2 by simp
+ from empty show ?case using good_fun_def2
+ by (simp add: comp_fun_commute.ffold_empty ge.comp_comp_fun_commute)
next
case (insert p f)
interpret ge: comp_fun_commute "(good_entry x e \<rho>) \<circ> good_prod"
unfolding comp_fun_commute_def by auto
have "good_fun (finsert p f) x e \<rho>
= ffold ((good_entry x e \<rho>) \<circ> good_prod) True (finsert p f)" by (simp add: good_fun_def2)
also from insert(1) have "... = ((good_entry x e \<rho>) \<circ> good_prod) p
(ffold ((good_entry x e \<rho>) \<circ> good_prod) True f)" by simp
finally have 1: "good_fun (finsert p f) x e \<rho>
= ((good_entry x e \<rho>) \<circ> good_prod) p (ffold ((good_entry x e \<rho>) \<circ> good_prod) True f)" .
show ?case
proof
assume 2: "good_fun (finsert p f) x e \<rho>"
show "\<forall>v1 v2. (v1, v2) \<in> fset (finsert p f) \<longrightarrow>
(\<forall>v\<in>good v1. \<exists>v'. (x, v) # \<rho> \<turnstile> e \<Down> v' \<and> v' \<in> good v2)"
proof clarify
fix v1 v2 v assume 3: "(v1, v2) \<in> fset (finsert p f)" and 4: "v \<in> good v1"
from 3 have "(v1,v2) = p \<or> (v1,v2) \<in> fset f" by auto
from this show "\<exists>v'. (x, v) # \<rho> \<turnstile> e \<Down> v' \<and> v' \<in> good v2"
proof
assume v12_p: "(v1,v2) = p"
from 1 v12_p[THEN sym] 2 4 show ?thesis by simp
next
assume v12_f: "(v1,v2) \<in> fset f"
from 1 2 have 5: "good_fun f x e \<rho>" apply simp
apply (cases "(good_prod p)") by (auto simp: good_fun_def2)
from v12_f 5 4 insert(2)[of x e \<rho>] show ?thesis by auto
qed
qed
next
assume 2: "\<forall>v1 v2. (v1, v2) \<in> fset (finsert p f) \<longrightarrow>
(\<forall>v\<in>good v1. \<exists>v'. (x, v) # \<rho> \<turnstile> e \<Down> v' \<and> v' \<in> good v2)"
have 3: "good_entry x e \<rho> (good_prod p) True"
apply (cases p) apply simp apply clarify
proof -
fix v1 v2 v
assume p: "p = (v1,v2)" and v_v1: "v \<in> good v1"
from p have "(v1,v2) \<in> fset (finsert p f)" by simp
from this 2 v_v1 show "\<exists>v'. (x, v) # \<rho> \<turnstile> e \<Down> v' \<and> v' \<in> good v2" by blast
qed
from insert(2) 2 have 4: "good_fun f x e \<rho>" by auto
have "(good_entry x e \<rho> \<circ> good_prod) p
(ffold (good_entry x e \<rho> \<circ> good_prod) True f)"
apply simp apply (cases "good_prod p")
apply (rename_tac a b c)
apply (case_tac a) apply simp
apply (rule conjI) prefer 2 using 4 good_fun_def2 apply force
using 3 apply force done
from this 1 show "good_fun (finsert p f) x e \<rho>"
unfolding good_fun_def by simp
qed
qed
lemma gfun_mem: "\<lbrakk> (v1,v2) \<in> fset f; good_fun f x e \<rho> \<rbrakk>
\<Longrightarrow> \<forall> v \<in> good v1. \<exists> v'. (x,v)#\<rho> \<turnstile> e \<Down> v' \<and> v' \<in> good v2"
using gfun_mem_iff by blast
lemma gfun_intro: "(\<forall> v1 v2.(v1,v2)\<in>fset f\<longrightarrow>(\<forall>v\<in>good v1.\<exists> v'.(x,v)#\<rho> \<turnstile> e \<Down> v'\<and>v'\<in>good v2))
\<Longrightarrow> good_fun f x e \<rho>" using gfun_mem_iff[of f x e \<rho>] by simp
lemma sub_good: fixes v::val assumes wv: "w \<in> good v" and vp_v: "v' \<sqsubseteq> v" shows "w \<in> good v'"
proof (cases v)
case (VNat n)
from this wv vp_v show ?thesis by auto
next
case (VFun t1)
from vp_v VFun obtain t2 where b: "v' = VFun t2" and t2_t1: "fset t2 \<subseteq> fset t1" by auto
from wv VFun obtain x e \<rho> where w: "w = BClos x e \<rho>" by auto
from w wv VFun have gt1: "good_fun t1 x e \<rho>" by (simp add: good_fun_def)
have gt2: "good_fun t2 x e \<rho>" apply (rule gfun_intro) apply clarify
proof -
fix v1 v2 w1
assume v12: "(v1,v2) \<in> fset t2" and w1_v1: "w1 \<in> good v1"
from v12 t2_t1 have v12_t1: "(v1,v2) \<in> fset t1" by blast
from gt1 v12_t1 w1_v1 show "\<exists>v'. (x, w1) # \<rho> \<turnstile> e \<Down> v' \<and> v' \<in> good v2"
by (simp add: gfun_mem)
qed
from gt2 b w show ?thesis by (simp add: good_fun_def)
qed
subsection "Denotational semantics sound wrt. big-step"
lemma denot_terminates: assumes vp_e: "v' \<in> E e \<rho>'" and ge: "good_env \<rho> \<rho>'"
shows "\<exists> v. \<rho> \<turnstile> e \<Down> v \<and> v \<in> good v'"
using vp_e ge
proof (induction arbitrary: v' \<rho> rule: E.induct)
case (1 n \<rho>) \<comment> \<open>ENat\<close>
then show ?case by auto
next \<comment> \<open>EVar\<close>
case (2 x \<rho> v' \<rho>')
from 2 obtain v1 where lx_vpp: "lookup \<rho> x = Some v1" and vp_v1: "v' \<sqsubseteq> v1" by auto
from lx_vpp 2(2) obtain v2 where lx: "lookup \<rho>' x = Some v2" and v2_v1: "v2 \<in> good v1"
using lookup_good[of \<rho> x v1 \<rho>'] by blast
from lx have x_v2: "\<rho>' \<turnstile> EVar x \<Down> v2" by auto
from v2_v1 vp_v1 have v2_vp: "v2 \<in> good v'" using sub_good by blast
from x_v2 v2_vp show ?case by blast
next \<comment> \<open>ELam\<close>
case (3 x e \<rho> v' \<rho>')
have 1: "\<rho>' \<turnstile> ELam x e \<Down> BClos x e \<rho>'" by auto
have 2: "BClos x e \<rho>' \<in> good v'"
proof -
from 3(2) obtain t where vp: "v' = VFun t" and
body: "\<forall>v1 v2. (v1, v2) \<in> fset t \<longrightarrow> v2 \<in> E e ((x, v1) # \<rho>)" by blast
have gt: "good_fun t x e \<rho>'" apply (rule gfun_intro) apply clarify
proof -
fix v1 v2 w1 assume v12_t: "(v1,v2) \<in> fset t" and w1_v1: "w1 \<in> good v1"
from v12_t body have v2_Ee: "v2 \<in> E e ((x, v1) # \<rho>)" by blast
from 3(3) w1_v1 have ge: "good_env ((x,w1)#\<rho>') ((x,v1)#\<rho>)" by auto
from v12_t v2_Ee ge 3(1)[of v1 v2 t v2]
show "\<exists>v'. (x, w1) # \<rho>' \<turnstile> e \<Down> v' \<and> v' \<in> good v2" by blast
qed
from vp gt show ?thesis unfolding good_fun_def by simp
qed
from 1 2 show ?case by blast
next \<comment> \<open>EApp\<close>
case (4 e1 e2 \<rho> v' \<rho>')
from 4(3) show ?case
proof
fix t v2 and v2'::val and v3' assume t_Ee1: "VFun t \<in> E e1 \<rho>" and v2_Ee2: "v2 \<in> E e2 \<rho>" and
v23_t: "(v2',v3') \<in> fset t" and v2p_v2: "v2' \<sqsubseteq> v2" and vp_v3p: "v' \<sqsubseteq> v3'"
from 4(1) t_Ee1 4(4) obtain w1 where e1_w1: "\<rho>' \<turnstile> e1 \<Down> w1" and
w1_t: "w1 \<in> good (VFun t)" by blast
from 4(2) v2_Ee2 4(4) obtain w2 where e2_w2: "\<rho>' \<turnstile> e2 \<Down> w2" and w2_v2: "w2 \<in> good v2" by blast
from w1_t obtain x e \<rho>1 where w1: "w1 = BClos x e \<rho>1" and gt: "good_fun t x e \<rho>1"
by (auto simp: good_fun_def)
from w2_v2 v2p_v2 have w2_v2p: "w2 \<in> good v2'" by (rule sub_good)
from v23_t gt w2_v2p obtain w3 where e_w3: "(x,w2)#\<rho>1 \<turnstile> e \<Down> w3" and w3_v3p: "w3 \<in> good v3'"
using gfun_mem[of v2' v3' t x e \<rho>1] by blast
from w3_v3p vp_v3p have w3_vp: "w3 \<in> good v'" by (rule sub_good)
from e1_w1 e2_w2 w1 e_w3 w3_vp show "\<exists>v. \<rho>' \<turnstile> EApp e1 e2 \<Down> v \<and> v \<in> good v'" by blast
qed
next \<comment> \<open>EPrim\<close>
case (5 f e1 e2 \<rho> v' \<rho>')
from 5(3) show ?case
proof
fix n1 n2 assume n1_e1: "VNat n1 \<in> E e1 \<rho>" and n2_e2: "VNat n2 \<in> E e2 \<rho>" and
vp: "v' = VNat (f n1 n2)"
from 5(1)[of "VNat n1" \<rho>'] n1_e1 5(4) have e1_w1: "\<rho>' \<turnstile> e1 \<Down> BNat n1" by auto
from 5(2)[of "VNat n2" \<rho>'] n2_e2 5(4) have e2_w2: "\<rho>' \<turnstile> e2 \<Down> BNat n2" by auto
from e1_w1 e2_w2 have 1: "\<rho>' \<turnstile> EPrim f e1 e2 \<Down> BNat (f n1 n2)" by blast
from vp have 2: "BNat (f n1 n2) \<in> good v'" by auto
from 1 2 show "\<exists>v. \<rho>' \<turnstile> EPrim f e1 e2 \<Down> v \<and> v \<in> good v'" by auto
qed
next \<comment> \<open>EIf\<close>
case (6 e1 e2 e3 \<rho> v' \<rho>')
from 6(4) show ?case
proof
fix n assume n_e1: "VNat n \<in> E e1 \<rho>" and els: "n = 0 \<longrightarrow> v' \<in> E e3 \<rho>" and
thn: "n \<noteq> 0 \<longrightarrow> v' \<in> E e2 \<rho>"
from 6(1)[of "VNat n" \<rho>'] n_e1 6(5) have e1_w1: "\<rho>' \<turnstile> e1 \<Down> BNat n" by auto
show "\<exists>v. \<rho>' \<turnstile> EIf e1 e2 e3 \<Down> v \<and> v \<in> good v'"
proof (cases "n = 0")
case True
from 6(2)[of n v' \<rho>'] True els 6(5) obtain w3 where
e3_w3: "\<rho>' \<turnstile> e3 \<Down> w3" and w3_vp: "w3 \<in> good v'" by blast
from e1_w1 True e3_w3 w3_vp show ?thesis by blast
next
case False
from 6(3)[of n v' \<rho>'] False thn 6(5) obtain w2 where
e2_w2: "\<rho>' \<turnstile> e2 \<Down> w2" and w2_vp: "w2 \<in> good v'" by blast
from e1_w1 False e2_w2 have "\<rho>' \<turnstile> EIf e1 e2 e3 \<Down> w2"
using eval_if1[of \<rho>' e1 n e2 w2 e3] by simp
from this w2_vp show ?thesis by (rule_tac x=w2 in exI) simp
qed
qed
qed
theorem sound_wrt_op_sem:
assumes E_e_n: "E e [] = E (ENat n) []" and fv_e: "FV e = {}" shows "e \<Down> ONat n"
proof -
have "VNat n \<in> E (ENat n) []" by simp
with E_e_n have 1: "VNat n \<in> E e []" by simp
have 2: "good_env [] []" by auto
from 1 2 obtain v where e_v: "[] \<turnstile> e \<Down> v" and v_n: "v \<in> good (VNat n)" using denot_terminates by blast
from v_n have v: "v = BNat n" by auto
from e_v fv_e obtain v' ob where e_vp: "e \<longrightarrow>* v'" and
vp_ob: "observe v' ob" and v_ob: "bs_observe v ob" using sound_wrt_small_step by blast
from e_vp vp_ob v_ob v show ?thesis unfolding run_def by (case_tac ob) auto
qed
end
diff --git a/thys/Decl_Sem_Fun_PL/ValuesFSetProps.thy b/thys/Decl_Sem_Fun_PL/ValuesFSetProps.thy
--- a/thys/Decl_Sem_Fun_PL/ValuesFSetProps.thy
+++ b/thys/Decl_Sem_Fun_PL/ValuesFSetProps.thy
@@ -1,39 +1,39 @@
theory ValuesFSetProps
imports ValuesFSet
begin
inductive_cases
vfun_le_inv[elim!]: "VFun t1 \<sqsubseteq> VFun t2" and
le_fun_nat_inv[elim!]: "VFun t2 \<sqsubseteq> VNat x1" and
le_any_nat_inv[elim!]: "v \<sqsubseteq> VNat n" and
le_nat_any_inv[elim!]: "VNat n \<sqsubseteq> v" and
le_fun_any_inv[elim!]: "VFun t \<sqsubseteq> v" and
le_any_fun_inv[elim!]: "v \<sqsubseteq> VFun t"
proposition val_le_refl[simp]: fixes v::val shows "v \<sqsubseteq> v" by (induction v) auto
proposition val_le_trans[trans]: fixes v2::val shows "\<lbrakk> v1 \<sqsubseteq> v2; v2 \<sqsubseteq> v3 \<rbrakk> \<Longrightarrow> v1 \<sqsubseteq> v3"
by (induction v2 arbitrary: v1 v3) blast+
lemma fsubset[intro!]: "fset A \<subseteq> fset B \<Longrightarrow> A |\<subseteq>| B"
proof (rule fsubsetI)
fix x assume ab: "fset A \<subseteq> fset B" and xa: "x |\<in>| A"
- from xa have "x \<in> fset A" using fmember_iff_member_fset[of x A] by simp
+ from xa have "x \<in> fset A" by simp
from this ab have "x \<in> fset B" by blast
- from this show "x |\<in>| B" using fmember_iff_member_fset[of x B] by simp
+ from this show "x |\<in>| B" by simp
qed
proposition val_le_antisymm: fixes v1::val shows "\<lbrakk> v1 \<sqsubseteq> v2; v2 \<sqsubseteq> v1 \<rbrakk> \<Longrightarrow> v1 = v2"
by (induction v1 arbitrary: v2) auto
lemma le_nat_any[simp]: "VNat n \<sqsubseteq> v \<Longrightarrow> v = VNat n"
by (cases v) auto
lemma le_any_nat[simp]: "v \<sqsubseteq> VNat n \<Longrightarrow> v = VNat n"
by (cases v) auto
lemma le_nat_nat[simp]: "VNat n \<sqsubseteq> VNat n' \<Longrightarrow> n = n'"
by auto
end
diff --git a/thys/Extended_Finite_State_Machine_Inference/Inference.thy b/thys/Extended_Finite_State_Machine_Inference/Inference.thy
--- a/thys/Extended_Finite_State_Machine_Inference/Inference.thy
+++ b/thys/Extended_Finite_State_Machine_Inference/Inference.thy
@@ -1,588 +1,586 @@
chapter\<open>EFSM Inference\<close>
text\<open>This chapter presents the definitions necessary for EFSM inference by state-merging.\<close>
section\<open>Inference by State-Merging\<close>
text\<open>This theory sets out the key definitions for the inference of EFSMs from system traces.\<close>
theory Inference
imports
Subsumption
"Extended_Finite_State_Machines.Transition_Lexorder"
"HOL-Library.Product_Lexorder"
begin
declare One_nat_def [simp del]
subsection\<open>Transition Identifiers\<close>
text\<open>We first need to define the \texttt{iEFSM} data type which assigns each transition a unique identity.
This is necessary because transitions may not occur uniquely in an EFSM. Assigning transitions a unique
identifier enables us to look up the origin and destination states of transitions without having to
pass them around in the inference functions.\<close>
type_synonym tid = nat
type_synonym tids = "tid list"
type_synonym iEFSM = "(tids \<times> (cfstate \<times> cfstate) \<times> transition) fset"
definition origin :: "tids \<Rightarrow> iEFSM \<Rightarrow> nat" where
"origin uid t = fst (fst (snd (fthe_elem (ffilter (\<lambda>x. set uid \<subseteq> set (fst x)) t))))"
definition dest :: "tids \<Rightarrow> iEFSM \<Rightarrow> nat" where
"dest uid t = snd (fst (snd (fthe_elem (ffilter (\<lambda>x. set uid \<subseteq> set (fst x)) t))))"
definition get_by_id :: "iEFSM \<Rightarrow> tid \<Rightarrow> transition" where
"get_by_id e uid = (snd \<circ> snd) (fthe_elem (ffilter (\<lambda>(tids, _). uid \<in> set tids) e))"
definition get_by_ids :: "iEFSM \<Rightarrow> tids \<Rightarrow> transition" where
"get_by_ids e uid = (snd \<circ> snd) (fthe_elem (ffilter (\<lambda>(tids, _). set uid \<subseteq> set tids) e))"
definition uids :: "iEFSM \<Rightarrow> nat fset" where
"uids e = ffUnion (fimage (fset_of_list \<circ> fst) e)"
definition max_uid :: "iEFSM \<Rightarrow> nat option" where
"max_uid e = (let uids = uids e in if uids = {||} then None else Some (fMax uids))"
definition tm :: "iEFSM \<Rightarrow> transition_matrix" where
"tm e = fimage snd e"
definition all_regs :: "iEFSM \<Rightarrow> nat set" where
"all_regs e = EFSM.all_regs (tm e)"
definition max_reg :: "iEFSM \<Rightarrow> nat option" where
"max_reg e = EFSM.max_reg (tm e)"
definition "max_reg_total e = (case max_reg e of None \<Rightarrow> 0 | Some r \<Rightarrow> r)"
definition max_output :: "iEFSM \<Rightarrow> nat" where
"max_output e = EFSM.max_output (tm e)"
definition max_int :: "iEFSM \<Rightarrow> int" where
"max_int e = EFSM.max_int (tm e)"
definition S :: "iEFSM \<Rightarrow> nat fset" where
"S m = (fimage (\<lambda>(uid, (s, s'), t). s) m) |\<union>| fimage (\<lambda>(uid, (s, s'), t). s') m"
lemma S_alt: "S t = EFSM.S (tm t)"
apply (simp add: S_def EFSM.S_def tm_def)
by force
lemma to_in_S:
"(\<exists>to from uid. (uid, (from, to), t) |\<in>| xb \<longrightarrow> to |\<in>| S xb)"
apply (simp add: S_def)
by blast
lemma from_in_S:
"(\<exists>to from uid. (uid, (from, to), t) |\<in>| xb \<longrightarrow> from |\<in>| S xb)"
apply (simp add: S_def)
by blast
subsection\<open>Building the PTA\<close>
text\<open>The first step in EFSM inference is to construct a PTA from the observed traces in the same way
as for classical FSM inference. Beginning with the empty EFSM, we iteratively attempt to walk each
observed trace in the model. When we reach a point where there is no available transition, one is
added. For classical FSMs, this is simply an atomic label. EFSMs deal with data, so we need to add
guards which test for the observed input values and outputs which produce the observed values.\<close>
primrec make_guard :: "value list \<Rightarrow> nat \<Rightarrow> vname gexp list" where
"make_guard [] _ = []" |
"make_guard (h#t) n = (gexp.Eq (V (vname.I n)) (L h))#(make_guard t (n+1))"
primrec make_outputs :: "value list \<Rightarrow> output_function list" where
"make_outputs [] = []" |
"make_outputs (h#t) = (L h)#(make_outputs t)"
definition max_uid_total :: "iEFSM \<Rightarrow> nat" where
"max_uid_total e = (case max_uid e of None \<Rightarrow> 0 | Some u \<Rightarrow> u)"
definition add_transition :: "iEFSM \<Rightarrow> cfstate \<Rightarrow> label \<Rightarrow> value list \<Rightarrow> value list \<Rightarrow> iEFSM" where
"add_transition e s label inputs outputs = finsert ([max_uid_total e + 1], (s, (maxS (tm e))+1), \<lparr>Label=label, Arity=length inputs, Guards=(make_guard inputs 0), Outputs=(make_outputs outputs), Updates=[]\<rparr>) e"
fun make_branch :: "iEFSM \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> trace \<Rightarrow> iEFSM" where
"make_branch e _ _ [] = e" |
"make_branch e s r ((label, inputs, outputs)#t) =
(case (step (tm e) s r label inputs) of
Some (transition, s', outputs', updated) \<Rightarrow>
if outputs' = (map Some outputs) then
make_branch e s' updated t
else
make_branch (add_transition e s label inputs outputs) ((maxS (tm e))+1) r t |
None \<Rightarrow>
make_branch (add_transition e s label inputs outputs) ((maxS (tm e))+1) r t
)"
primrec make_pta_aux :: "log \<Rightarrow> iEFSM \<Rightarrow> iEFSM" where
"make_pta_aux [] e = e" |
"make_pta_aux (h#t) e = make_pta_aux t (make_branch e 0 <> h)"
definition "make_pta log = make_pta_aux log {||}"
lemma make_pta_aux_fold [code]:
"make_pta_aux l e = fold (\<lambda>h e. make_branch e 0 <> h) l e"
by(induct l arbitrary: e, auto)
subsection\<open>Integrating Heuristics\<close>
text\<open>A key contribution of the inference technique presented in \<^cite>\<open>"foster2019"\<close> is the ability to
introduce \emph{internal variables} to the model to generalise behaviours and allow transitions to
be merged. This is done by providing the inference technique with a set of \emph{heuristics}. The
aim here is not to create a ``one size fits all'' magic oracle, rather to recognise particular
\emph{data usage patterns} which can be abstracted.\<close>
type_synonym update_modifier = "tids \<Rightarrow> tids \<Rightarrow> cfstate \<Rightarrow> iEFSM \<Rightarrow> iEFSM \<Rightarrow> iEFSM \<Rightarrow> (transition_matrix \<Rightarrow> bool) \<Rightarrow> iEFSM option"
definition null_modifier :: update_modifier where
"null_modifier f _ _ _ _ _ _ = None"
definition replace_transition :: "iEFSM \<Rightarrow> tids \<Rightarrow> transition \<Rightarrow> iEFSM" where
"replace_transition e uid new = (fimage (\<lambda>(uids, (from, to), t). if set uid \<subseteq> set uids then (uids, (from, to), new) else (uids, (from, to), t)) e)"
definition replace_all :: "iEFSM \<Rightarrow> tids list \<Rightarrow> transition \<Rightarrow> iEFSM" where
"replace_all e ids new = fold (\<lambda>id acc. replace_transition acc id new) ids e"
definition replace_transitions :: "iEFSM \<Rightarrow> (tids \<times> transition) list \<Rightarrow> iEFSM" where
"replace_transitions e ts = fold (\<lambda>(uid, new) acc. replace_transition acc uid new) ts e"
primrec try_heuristics_check :: "(transition_matrix \<Rightarrow> bool) \<Rightarrow> update_modifier list \<Rightarrow> update_modifier" where
"try_heuristics_check _ [] = null_modifier" |
"try_heuristics_check check (h#t) = (\<lambda>a b c d e f ch.
case h a b c d e f ch of
Some e' \<Rightarrow> Some e' |
None \<Rightarrow> (try_heuristics_check check t) a b c d e f ch
)"
subsection\<open>Scoring State Merges\<close>
text\<open>To tackle the state merging challenge, we need some means of determining which states are
compatible for merging. Because states are merged pairwise, we additionally require a way of
ordering the state merges. The potential merges are then sorted highest to lowest according to this
score such that we can merge states in order of their merge score.
We want to sort first by score (highest to lowest) and then by state pairs (lowest to highest) so we
endup merging the states with the highest scores first and then break ties by those state pairs
which are closest to the origin.\<close>
record score =
Score :: nat
S1 :: cfstate
S2 :: cfstate
instantiation score_ext :: (linorder) linorder begin
definition less_score_ext :: "'a::linorder score_ext \<Rightarrow> 'a score_ext \<Rightarrow> bool" where
"less_score_ext t1 t2 = ((Score t2, S1 t1, S2 t1, more t1) < (Score t1, S1 t2, S2 t2, more t2) )"
definition less_eq_score_ext :: "'a::linorder score_ext \<Rightarrow> 'a::linorder score_ext \<Rightarrow> bool" where
"less_eq_score_ext s1 s2 = (s1 < s2 \<or> s1 = s2)"
instance
apply standard prefer 5
unfolding less_score_ext_def less_eq_score_ext_def
using score.equality apply fastforce
by auto
end
type_synonym scoreboard = "score fset"
type_synonym strategy = "tids \<Rightarrow> tids \<Rightarrow> iEFSM \<Rightarrow> nat"
definition outgoing_transitions :: "cfstate \<Rightarrow> iEFSM \<Rightarrow> (cfstate \<times> transition \<times> tids) fset" where
"outgoing_transitions s e = fimage (\<lambda>(uid, (from, to), t'). (to, t', uid)) ((ffilter (\<lambda>(uid, (origin, dest), t). origin = s)) e)"
primrec paths_of_length :: "nat \<Rightarrow> iEFSM \<Rightarrow> cfstate \<Rightarrow> tids list fset" where
"paths_of_length 0 _ _ = {|[]|}" |
"paths_of_length (Suc m) e s = (
let
outgoing = outgoing_transitions s e;
paths = ffUnion (fimage (\<lambda>(d, t, id). fimage (\<lambda>p. id#p) (paths_of_length m e d)) outgoing)
in
ffilter (\<lambda>l. length l = Suc m) paths
)"
lemma paths_of_length_1: "paths_of_length 1 e s = fimage (\<lambda>(d, t, id). [id]) (outgoing_transitions s e)"
apply (simp add: One_nat_def)
apply (simp add: outgoing_transitions_def comp_def One_nat_def[symmetric])
apply (rule fBall_ffilter2)
defer
apply (simp add: ffilter_def ffUnion_def fBall_def Abs_fset_inverse)
apply auto[1]
apply (simp add: ffilter_def ffUnion_def fBall_def Abs_fset_inverse fset_both_sides)
by force
fun step_score :: "(tids \<times> tids) list \<Rightarrow> iEFSM \<Rightarrow> strategy \<Rightarrow> nat" where
"step_score [] _ _ = 0" |
"step_score ((id1, id2)#t) e s = (
let score = s id1 id2 e in
if score = 0 then
0
else
score + (step_score t e s)
)"
lemma step_score_foldr [code]:
"step_score xs e s = foldr (\<lambda>(id1, id2) acc. let score = s id1 id2 e in
if score = 0 then
0
else
score + acc) xs 0"
proof(induct xs)
case Nil
then show ?case
by simp
next
case (Cons a xs)
then show ?case
apply (cases a, clarify)
by (simp add: Let_def)
qed
definition score_from_list :: "tids list fset \<Rightarrow> tids list fset \<Rightarrow> iEFSM \<Rightarrow> strategy \<Rightarrow> nat" where
"score_from_list P1 P2 e s = (
let
pairs = fimage (\<lambda>(l1, l2). zip l1 l2) (P1 |\<times>| P2);
scored_pairs = fimage (\<lambda>l. step_score l e s) pairs
in
fSum scored_pairs
)"
definition k_score :: "nat \<Rightarrow> iEFSM \<Rightarrow> strategy \<Rightarrow> scoreboard" where
"k_score k e strat = (
let
states = S e;
pairs_to_score = (ffilter (\<lambda>(x, y). x < y) (states |\<times>| states));
paths = fimage (\<lambda>(s1, s2). (s1, s2, paths_of_length k e s1, paths_of_length k e s2)) pairs_to_score;
scores = fimage (\<lambda>(s1, s2, p1, p2). \<lparr>Score = score_from_list p1 p2 e strat, S1 = s1, S2 = s2\<rparr>) paths
in
ffilter (\<lambda>x. Score x > 0) scores
)"
definition score_state_pair :: "strategy \<Rightarrow> iEFSM \<Rightarrow> cfstate \<Rightarrow> cfstate \<Rightarrow> nat" where
"score_state_pair strat e s1 s2 = (
let
T1 = outgoing_transitions s1 e;
T2 = outgoing_transitions s2 e
in
fSum (fimage (\<lambda>((_, _, t1), (_, _, t2)). strat t1 t2 e) (T1 |\<times>| T2))
)"
definition score_1 :: "iEFSM \<Rightarrow> strategy \<Rightarrow> scoreboard" where
"score_1 e strat = (
let
states = S e;
pairs_to_score = (ffilter (\<lambda>(x, y). x < y) (states |\<times>| states));
scores = fimage (\<lambda>(s1, s2). \<lparr>Score = score_state_pair strat e s1 s2, S1 = s1, S2 = s2\<rparr>) pairs_to_score
in
ffilter (\<lambda>x. Score x > 0) scores
)"
lemma score_1: "score_1 e s = k_score 1 e s"
proof-
have fprod_fimage:
"\<And>a b. ((\<lambda>(_, _, id). [id]) |`| a |\<times>| (\<lambda>(_, _, id). [id]) |`| b) =
fimage (\<lambda>((_, _, id1), (_, _, id2)). ([id1], [id2])) (a |\<times>| b)"
apply (simp add: fimage_def fprod_def Abs_fset_inverse fset_both_sides)
by force
show ?thesis
apply (simp add: score_1_def k_score_def Let_def comp_def)
apply (rule arg_cong[of _ _ "ffilter (\<lambda>x. 0 < Score x)"])
apply (rule fun_cong[of _ _ "(Inference.S e |\<times>| Inference.S e)"])
apply (rule ext)
subgoal for x
apply (rule fun_cong[of _ _ "ffilter (\<lambda>a. case a of (a, b) \<Rightarrow> a < b) x"])
apply (rule arg_cong[of _ _ fimage])
apply (rule ext)
subgoal for x
apply (case_tac x)
apply simp
apply (simp add: paths_of_length_1)
apply (simp add: score_state_pair_def Let_def score_from_list_def comp_def)
subgoal for a b
apply (rule arg_cong[of _ _ fSum])
apply (simp add: fprod_fimage)
apply (rule fun_cong[of _ _ "(outgoing_transitions a e |\<times>| outgoing_transitions b e)"])
apply (rule arg_cong[of _ _ fimage])
apply (rule ext)
apply clarify
by (simp add: Let_def)
done
done
done
qed
fun bool2nat :: "bool \<Rightarrow> nat" where
"bool2nat True = 1" |
"bool2nat False = 0"
definition score_transitions :: "transition \<Rightarrow> transition \<Rightarrow> nat" where
"score_transitions t1 t2 = (
if Label t1 = Label t2 \<and> Arity t1 = Arity t2 \<and> length (Outputs t1) = length (Outputs t2) then
1 + bool2nat (t1 = t2) + card ((set (Guards t2)) \<inter> (set (Guards t2))) + card ((set (Updates t2)) \<inter> (set (Updates t2))) + card ((set (Outputs t2)) \<inter> (set (Outputs t2)))
else
0
)"
subsection\<open>Merging States\<close>
definition merge_states_aux :: "nat \<Rightarrow> nat \<Rightarrow> iEFSM \<Rightarrow> iEFSM" where
"merge_states_aux s1 s2 e = fimage (\<lambda>(uid, (origin, dest), t). (uid, (if origin = s1 then s2 else origin , if dest = s1 then s2 else dest), t)) e"
definition merge_states :: "nat \<Rightarrow> nat \<Rightarrow> iEFSM \<Rightarrow> iEFSM" where
"merge_states x y t = (if x > y then merge_states_aux x y t else merge_states_aux y x t)"
lemma merge_states_symmetry: "merge_states x y t = merge_states y x t"
by (simp add: merge_states_def)
lemma merge_state_self: "merge_states s s t = t"
apply (simp add: merge_states_def merge_states_aux_def)
by force
lemma merge_states_self_simp [code]:
"merge_states x y t = (if x = y then t else if x > y then merge_states_aux x y t else merge_states_aux y x t)"
apply (simp add: merge_states_def merge_states_aux_def)
by force
subsection\<open>Resolving Nondeterminism\<close>
text\<open>Because EFSM transitions are not simply atomic actions, duplicated behaviours cannot be
resolved into a single transition by simply merging destination states, as it can in classical FSM
inference. It is now possible for attempts to resolve the nondeterminism introduced by merging
states to fail, meaning that two states which initially seemed compatible cannot actually be merged.
This is not the case in classical FSM inference.\<close>
type_synonym nondeterministic_pair = "(cfstate \<times> (cfstate \<times> cfstate) \<times> ((transition \<times> tids) \<times> (transition \<times> tids)))"
definition state_nondeterminism :: "nat \<Rightarrow> (cfstate \<times> transition \<times> tids) fset \<Rightarrow> nondeterministic_pair fset" where
"state_nondeterminism og nt = (if size nt < 2 then {||} else ffUnion (fimage (\<lambda>x. let (dest, t) = x in fimage (\<lambda>y. let (dest', t') = y in (og, (dest, dest'), (t, t'))) (nt - {|x|})) nt))"
lemma state_nondeterminism_empty [simp]: "state_nondeterminism a {||} = {||}"
by (simp add: state_nondeterminism_def ffilter_def Set.filter_def)
lemma state_nondeterminism_singledestn [simp]: "state_nondeterminism a {|x|} = {||}"
by (simp add: state_nondeterminism_def ffilter_def Set.filter_def)
(* For each state, get its outgoing transitions and see if there's any nondeterminism there *)
definition nondeterministic_pairs :: "iEFSM \<Rightarrow> nondeterministic_pair fset" where
"nondeterministic_pairs t = ffilter (\<lambda>(_, _, (t, _), (t', _)). Label t = Label t' \<and> Arity t = Arity t' \<and> choice t t') (ffUnion (fimage (\<lambda>s. state_nondeterminism s (outgoing_transitions s t)) (S t)))"
definition nondeterministic_pairs_labar_dest :: "iEFSM \<Rightarrow> nondeterministic_pair fset" where
"nondeterministic_pairs_labar_dest t = ffilter
(\<lambda>(_, (d, d'), (t, _), (t', _)).
Label t = Label t' \<and> Arity t = Arity t' \<and> (choice t t' \<or> (Outputs t = Outputs t' \<and> d = d')))
(ffUnion (fimage (\<lambda>s. state_nondeterminism s (outgoing_transitions s t)) (S t)))"
definition nondeterministic_pairs_labar :: "iEFSM \<Rightarrow> nondeterministic_pair fset" where
"nondeterministic_pairs_labar t = ffilter
(\<lambda>(_, (d, d'), (t, _), (t', _)).
Label t = Label t' \<and> Arity t = Arity t' \<and> (choice t t' \<or> Outputs t = Outputs t'))
(ffUnion (fimage (\<lambda>s. state_nondeterminism s (outgoing_transitions s t)) (S t)))"
definition deterministic :: "iEFSM \<Rightarrow> (iEFSM \<Rightarrow> nondeterministic_pair fset) \<Rightarrow> bool" where
"deterministic t np = (np t = {||})"
definition nondeterministic :: "iEFSM \<Rightarrow> (iEFSM \<Rightarrow> nondeterministic_pair fset) \<Rightarrow> bool" where
"nondeterministic t np = (\<not> deterministic t np)"
definition insert_transition :: "tids \<Rightarrow> cfstate \<Rightarrow> cfstate \<Rightarrow> transition \<Rightarrow> iEFSM \<Rightarrow> iEFSM" where
"insert_transition uid from to t e = (
if \<nexists>(uid, (from', to'), t') |\<in>| e. from = from' \<and> to = to' \<and> t = t' then
finsert (uid, (from, to), t) e
else
fimage (\<lambda>(uid', (from', to'), t').
if from = from' \<and> to = to' \<and> t = t' then
(List.union uid' uid, (from', to'), t')
else
(uid', (from', to'), t')
) e
)"
definition make_distinct :: "iEFSM \<Rightarrow> iEFSM" where
"make_distinct e = ffold_ord (\<lambda>(uid, (from, to), t) acc. insert_transition uid from to t acc) e {||}"
\<comment> \<open>When we replace one transition with another, we need to merge their uids to keep track of which\<close>
\<comment> \<open>transition accounts for which action in the original traces \<close>
definition merge_transitions_aux :: "iEFSM \<Rightarrow> tids \<Rightarrow> tids \<Rightarrow> iEFSM" where
"merge_transitions_aux e oldID newID = (let
(uids1, (origin, dest), old) = fthe_elem (ffilter (\<lambda>(uids, _). oldID = uids) e);
(uids2, (origin, dest), new) = fthe_elem (ffilter (\<lambda>(uids, _). newID = uids) e) in
make_distinct (finsert (List.union uids1 uids2, (origin, dest), new) (e - {|(uids1, (origin, dest), old), (uids2, (origin, dest), new)|}))
)"
(* merge_transitions - Try dest merge transitions t1 and t2 dest help resolve nondeterminism in
newEFSM. If either subsumes the other directly then the subsumed transition
can simply be replaced with the subsuming one, else we try dest apply the
modifier function dest resolve nondeterminism that way. *)
(* @param oldEFSM - the EFSM before merging the states which caused the nondeterminism *)
(* @param preDestMerge - the EFSM after merging the states which caused the nondeterminism *)
(* @param newEFSM - the current EFSM with nondeterminism *)
(* @param t1 - a transition dest be merged with t2 *)
(* @param u1 - the unique identifier of t1 *)
(* @param t2 - a transition dest be merged with t1 *)
(* @param u2 - the unique identifier of t2 *)
(* @param modifier - an update modifier function which tries dest generalise transitions *)
definition merge_transitions :: "(cfstate \<times> cfstate) set \<Rightarrow> iEFSM \<Rightarrow> iEFSM \<Rightarrow> iEFSM \<Rightarrow> transition \<Rightarrow> tids \<Rightarrow> transition \<Rightarrow> tids \<Rightarrow> update_modifier \<Rightarrow> (transition_matrix \<Rightarrow> bool) \<Rightarrow> iEFSM option" where
"merge_transitions failedMerges oldEFSM preDestMerge destMerge t1 u1 t2 u2 modifier check = (
if \<forall>id \<in> set u1. directly_subsumes (tm oldEFSM) (tm destMerge) (origin [id] oldEFSM) (origin u1 destMerge) t2 t1 then
\<comment> \<open>Replace t1 with t2\<close>
Some (merge_transitions_aux destMerge u1 u2)
else if \<forall>id \<in> set u2. directly_subsumes (tm oldEFSM) (tm destMerge) (origin [id] oldEFSM) (origin u2 destMerge) t1 t2 then
\<comment> \<open>Replace t2 with t1\<close>
Some (merge_transitions_aux destMerge u2 u1)
else
case modifier u1 u2 (origin u1 destMerge) destMerge preDestMerge oldEFSM check of
None \<Rightarrow> None |
Some e \<Rightarrow> Some (make_distinct e)
)"
definition outgoing_transitions_from :: "iEFSM \<Rightarrow> cfstate \<Rightarrow> transition fset" where
"outgoing_transitions_from e s = fimage (\<lambda>(_, _, t). t) (ffilter (\<lambda>(_, (orig, _), _). orig = s) e)"
definition order_nondeterministic_pairs :: "nondeterministic_pair fset \<Rightarrow> nondeterministic_pair list" where
"order_nondeterministic_pairs s = map snd (sorted_list_of_fset (fimage (\<lambda>s. let (_, _, (t1, _), (t2, _)) = s in (score_transitions t1 t2, s)) s))"
(* resolve_nondeterminism - tries dest resolve nondeterminism in a given iEFSM *)
(* @param ((from, (dest1, dest2), ((t1, u1), (t2, u2)))#ss) - a list of nondeterministic pairs where
from - nat - the state from which t1 and t2 eminate
dest1 - nat - the destination state of t1
dest2 - nat - the destination state of t2
t1 - transition - a transition dest be merged with t2
t2 - transition - a transition dest be merged with t1
u1 - nat - the unique identifier of t1
u2 - nat - the unique identifier of t2
ss - list - the rest of the list *)
(* @param oldEFSM - the EFSM before merging the states which caused the nondeterminism *)
(* @param newEFSM - the current EFSM with nondeterminism *)
(* @param m - an update modifier function which tries dest generalise transitions *)
(* @param check - a function which takes an EFSM and returns a bool dest ensure that certain
properties hold in the new iEFSM *)
function resolve_nondeterminism :: "(cfstate \<times> cfstate) set \<Rightarrow> nondeterministic_pair list \<Rightarrow> iEFSM \<Rightarrow> iEFSM \<Rightarrow> update_modifier \<Rightarrow> (transition_matrix \<Rightarrow> bool) \<Rightarrow> (iEFSM \<Rightarrow> nondeterministic_pair fset) \<Rightarrow> (iEFSM option \<times> (cfstate \<times> cfstate) set)" where
"resolve_nondeterminism failedMerges [] _ newEFSM _ check np = (
if deterministic newEFSM np \<and> check (tm newEFSM) then Some newEFSM else None, failedMerges
)" |
"resolve_nondeterminism failedMerges ((from, (dest1, dest2), ((t1, u1), (t2, u2)))#ss) oldEFSM newEFSM m check np = (
if (dest1, dest2) \<in> failedMerges \<or> (dest2, dest1) \<in> failedMerges then
(None, failedMerges)
else
let destMerge = merge_states dest1 dest2 newEFSM in
case merge_transitions failedMerges oldEFSM newEFSM destMerge t1 u1 t2 u2 m check of
None \<Rightarrow> resolve_nondeterminism (insert (dest1, dest2) failedMerges) ss oldEFSM newEFSM m check np |
Some new \<Rightarrow> (
let newScores = order_nondeterministic_pairs (np new) in
if (size new, size (S new), size (newScores)) < (size newEFSM, size (S newEFSM), size ss) then
case resolve_nondeterminism failedMerges newScores oldEFSM new m check np of
(Some new', failedMerges) \<Rightarrow> (Some new', failedMerges) |
(None, failedMerges) \<Rightarrow> resolve_nondeterminism (insert (dest1, dest2) failedMerges) ss oldEFSM newEFSM m check np
else
(None, failedMerges)
)
)"
apply (clarify, metis neq_Nil_conv prod_cases3 surj_pair)
by auto
termination
by (relation "measures [\<lambda>(_, _, _, newEFSM, _). size newEFSM,
\<lambda>(_, _, _, newEFSM, _). size (S newEFSM),
\<lambda>(_, ss, _, _, _). size ss]", auto)
subsection\<open>EFSM Inference\<close>
(* Merge - tries dest merge two states in a given iEFSM and resolve the resulting nondeterminism *)
(* @param e - an iEFSM *)
(* @param s1 - a state dest be merged with s2 *)
(* @param s2 - a state dest be merged with s1 *)
(* @param m - an update modifier function which tries dest generalise transitions *)
(* @param check - a function which takes an EFSM and returns a bool dest ensure that certain
properties hold in the new iEFSM *)
definition merge :: "(cfstate \<times> cfstate) set \<Rightarrow> iEFSM \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> update_modifier \<Rightarrow> (transition_matrix \<Rightarrow> bool) \<Rightarrow> (iEFSM \<Rightarrow> nondeterministic_pair fset) \<Rightarrow> (iEFSM option \<times> (cfstate \<times> cfstate) set)" where
"merge failedMerges e s1 s2 m check np = (
if s1 = s2 \<or> (s1, s2) \<in> failedMerges \<or> (s2, s1) \<in> failedMerges then
(None, failedMerges)
else
let e' = make_distinct (merge_states s1 s2 e) in
resolve_nondeterminism failedMerges (order_nondeterministic_pairs (np e')) e e' m check np
)"
(* inference_step - attempt dest carry out a single step of the inference process by merging the *)
(* @param e - an iEFSM dest be generalised *)
(* @param ((s, s1, s2)#t) - a list of triples of the form (score, state, state) dest be merged *)
(* @param m - an update modifier function which tries dest generalise transitions *)
(* @param check - a function which takes an EFSM and returns a bool dest ensure that certain
properties hold in the new iEFSM *)
function inference_step :: "(cfstate \<times> cfstate) set \<Rightarrow> iEFSM \<Rightarrow> score fset \<Rightarrow> update_modifier \<Rightarrow> (transition_matrix \<Rightarrow> bool) \<Rightarrow> (iEFSM \<Rightarrow> nondeterministic_pair fset) \<Rightarrow> (iEFSM option \<times> (cfstate \<times> cfstate) set)" where
"inference_step failedMerges e s m check np = (
if s = {||} then (None, failedMerges) else
let
h = fMin s;
t = s - {|h|}
in
case merge failedMerges e (S1 h) (S2 h) m check np of
(Some new, failedMerges) \<Rightarrow> (Some new, failedMerges) |
(None, failedMerges) \<Rightarrow> inference_step (insert ((S1 h), (S2 h)) failedMerges) e t m check np
)"
by auto
termination
- apply (relation "measures [\<lambda>(_, _, s, _, _, _). size s]")
- apply simp
- by (simp add: card_minus_fMin)
+ by (relation "measures [\<lambda>(_, _, s, _, _, _). size s]") (auto dest!: card_minus_fMin)
(* Takes an iEFSM and iterates inference_step until no further states can be successfully merged *)
(* @param e - an iEFSM dest be generalised *)
(* @param r - a strategy dest identify and prioritise pairs of states dest merge *)
(* @param m - an update modifier function which tries dest generalise transitions *)
(* @param check - a function which takes an EFSM and returns a bool dest ensure that certain
properties hold in the new iEFSM *)
function infer :: "(cfstate \<times> cfstate) set \<Rightarrow> nat \<Rightarrow> iEFSM \<Rightarrow> strategy \<Rightarrow> update_modifier \<Rightarrow> (transition_matrix \<Rightarrow> bool) \<Rightarrow> (iEFSM \<Rightarrow> nondeterministic_pair fset) \<Rightarrow> iEFSM" where
"infer failedMerges k e r m check np = (
let scores = if k = 1 then score_1 e r else (k_score k e r) in
case inference_step failedMerges e (ffilter (\<lambda>s. (S1 s, S2 s) \<notin> failedMerges \<and> (S2 s, S1 s) \<notin> failedMerges) scores) m check np of
(None, _) \<Rightarrow> e |
(Some new, failedMerges) \<Rightarrow> if (S new) |\<subset>| (S e) then infer failedMerges k new r m check np else e
)"
by auto
termination
apply (relation "measures [\<lambda>(_, _, e, _). size (S e)]")
apply simp
by (metis (no_types, lifting) case_prod_conv measures_less size_fsubset)
fun get_ints :: "trace \<Rightarrow> int list" where
"get_ints [] = []" |
"get_ints ((_, inputs, outputs)#t) = (map (\<lambda>x. case x of Num n \<Rightarrow> n) (filter is_Num (inputs@outputs)))"
definition learn :: "nat \<Rightarrow> iEFSM \<Rightarrow> log \<Rightarrow> strategy \<Rightarrow> update_modifier \<Rightarrow> (iEFSM \<Rightarrow> nondeterministic_pair fset) \<Rightarrow> iEFSM" where
"learn n pta l r m np = (
let check = accepts_log (set l) in
(infer {} n pta r m check np)
)"
subsection\<open>Evaluating Inferred Models\<close>
text\<open>We need a function to test the EFSMs we infer. The \texttt{test\_trace} function executes a
trace in the model and outputs a more comprehensive trace such that the expected outputs and actual
outputs can be compared. If a point is reached where the model does not recognise an action, the
remainder of the trace forms the second element of the output pair such that we know the exact point
at which the model stopped processing.\<close>
definition i_possible_steps :: "iEFSM \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> label \<Rightarrow> inputs \<Rightarrow> (tids \<times> cfstate \<times> transition) fset" where
"i_possible_steps e s r l i = fimage (\<lambda>(uid, (origin, dest), t). (uid, dest, t))
(ffilter (\<lambda>(uid, (origin, dest::nat), t::transition).
origin = s
\<and> (Label t) = l
\<and> (length i) = (Arity t)
\<and> apply_guards (Guards t) (join_ir i r)
)
e)"
fun test_trace :: "trace \<Rightarrow> iEFSM \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> ((label \<times> inputs \<times> cfstate \<times> cfstate \<times> registers \<times> tids \<times> value list \<times> outputs) list \<times> trace)" where
"test_trace [] _ _ _ = ([], [])" |
"test_trace ((l, i, expected)#es) e s r = (
let
ps = i_possible_steps e s r l i
in
if fis_singleton ps then
let
(id, s', t) = fthe_elem ps;
r' = evaluate_updates t i r;
actual = evaluate_outputs t i r;
(est, fail) = (test_trace es e s' r')
in
((l, i, s, s', r, id, expected, actual)#est, fail)
else
([], (l, i, expected)#es)
)"
text\<open>The \texttt{test\_log} function executes the \texttt{test\_trace} function on a collection of
traces known as the \emph{test set.}\<close>
definition test_log :: "log \<Rightarrow> iEFSM \<Rightarrow> ((label \<times> inputs \<times> cfstate \<times> cfstate \<times> registers \<times> tids \<times> value list \<times> outputs) list \<times> trace) list" where
"test_log l e = map (\<lambda>t. test_trace t e 0 <>) l"
end
diff --git a/thys/Extended_Finite_State_Machines/EFSM.thy b/thys/Extended_Finite_State_Machines/EFSM.thy
--- a/thys/Extended_Finite_State_Machines/EFSM.thy
+++ b/thys/Extended_Finite_State_Machines/EFSM.thy
@@ -1,1496 +1,1492 @@
section \<open>Extended Finite State Machines\<close>
text\<open>This theory defines extended finite state machines as presented in \<^cite>\<open>"foster2018"\<close>. States
are indexed by natural numbers, however, since transition matrices are implemented by finite sets,
the number of reachable states in $S$ is necessarily finite. For ease of implementation, we
implicitly make the initial state zero for all EFSMs. This allows EFSMs to be represented purely by
their transition matrix which, in this implementation, is a finite set of tuples of the form
$((s_1, s_2), t)$ in which $s_1$ is the origin state, $s_2$ is the destination state, and $t$ is a
transition.\<close>
theory EFSM
imports "HOL-Library.FSet" Transition FSet_Utils
begin
declare One_nat_def [simp del]
type_synonym cfstate = nat
type_synonym inputs = "value list"
type_synonym outputs = "value option list"
type_synonym action = "(label \<times> inputs)"
type_synonym execution = "action list"
type_synonym observation = "outputs list"
type_synonym transition_matrix = "((cfstate \<times> cfstate) \<times> transition) fset"
no_notation relcomp (infixr "O" 75) and comp (infixl "o" 55)
type_synonym event = "(label \<times> inputs \<times> value list)"
type_synonym trace = "event list"
type_synonym log = "trace list"
definition Str :: "string \<Rightarrow> value" where
"Str s \<equiv> value.Str (String.implode s)"
lemma str_not_num: "Str s \<noteq> Num x1"
by (simp add: Str_def)
definition S :: "transition_matrix \<Rightarrow> nat fset" where
"S m = (fimage (\<lambda>((s, s'), t). s) m) |\<union>| fimage (\<lambda>((s, s'), t). s') m"
lemma S_ffUnion: "S e = ffUnion (fimage (\<lambda>((s, s'), _). {|s, s'|}) e)"
unfolding S_def
by(induct e, auto)
subsection\<open>Possible Steps\<close>
text\<open>From a given state, the possible steps for a given action are those transitions with labels
which correspond to the action label, arities which correspond to the number of inputs, and guards
which are satisfied by those inputs.\<close>
definition possible_steps :: "transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> label \<Rightarrow> inputs \<Rightarrow> (cfstate \<times> transition) fset" where
"possible_steps e s r l i = fimage (\<lambda>((origin, dest), t). (dest, t)) (ffilter (\<lambda>((origin, dest), t). origin = s \<and> (Label t) = l \<and> (length i) = (Arity t) \<and> apply_guards (Guards t) (join_ir i r)) e)"
lemma possible_steps_finsert:
"possible_steps (finsert ((s, s'), t) e) ss r l i = (
if s = ss \<and> (Label t) = l \<and> (length i) = (Arity t) \<and> apply_guards (Guards t) (join_ir i r) then
finsert (s', t) (possible_steps e s r l i)
else
possible_steps e ss r l i
)"
by (simp add: possible_steps_def ffilter_finsert)
lemma split_origin:
"ffilter (\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> can_take_transition t i r) e =
ffilter (\<lambda>((origin, dest), t). Label t = l \<and> can_take_transition t i r) (ffilter (\<lambda>((origin, dest), t). origin = s) e)"
by auto
lemma split_label:
"ffilter (\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> can_take_transition t i r) e =
ffilter (\<lambda>((origin, dest), t). origin = s \<and> can_take_transition t i r) (ffilter (\<lambda>((origin, dest), t). Label t = l) e)"
by auto
lemma possible_steps_empty_guards_false:
"\<forall>((s1, s2), t) |\<in>| ffilter (\<lambda>((origin, dest), t). Label t = l) e. \<not>can_take_transition t i r \<Longrightarrow>
possible_steps e s r l i = {||}"
apply (simp add: possible_steps_def can_take[symmetric] split_label)
by (simp add: Abs_ffilter fBall_def Ball_def)
lemma fmember_possible_steps: "(s', t) |\<in>| possible_steps e s r l i = (((s, s'), t) \<in> {((origin, dest), t) \<in> fset e. origin = s \<and> Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r)})"
- apply (simp add: possible_steps_def ffilter_def fimage_def fmember_def Abs_fset_inverse)
+ apply (simp add: possible_steps_def ffilter_def fimage_def Abs_fset_inverse)
by force
lemma possible_steps_alt_aux:
"possible_steps e s r l i = {|(d, t)|} \<Longrightarrow>
ffilter (\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r)) e = {|((s, d), t)|}"
proof(induct e)
case empty
then show ?case
by (simp add: fempty_not_finsert possible_steps_def)
next
case (insert x e)
then show ?case
apply (case_tac x)
subgoal for a b
apply (case_tac a)
subgoal for aa _
apply (simp add: possible_steps_def)
apply (simp add: ffilter_finsert)
apply (case_tac "aa = s \<and> Label b = l \<and> length i = Arity b \<and> apply_guards (Guards b) (join_ir i r)")
by auto
done
done
qed
lemma possible_steps_alt: "(possible_steps e s r l i = {|(d, t)|}) = (ffilter
(\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r))
e = {|((s, d), t)|})"
apply standard
apply (simp add: possible_steps_alt_aux)
by (simp add: possible_steps_def)
lemma possible_steps_alt3: "(possible_steps e s r l i = {|(d, t)|}) = (ffilter
(\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> can_take_transition t i r)
e = {|((s, d), t)|})"
apply standard
apply (simp add: possible_steps_alt_aux can_take)
by (simp add: possible_steps_def can_take)
lemma possible_steps_alt_atom: "(possible_steps e s r l i = {|dt|}) = (ffilter
(\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> can_take_transition t i r)
e = {|((s, fst dt), snd dt)|})"
apply (cases dt)
by (simp add: possible_steps_alt can_take_transition_def can_take_def)
lemma possible_steps_alt2: "(possible_steps e s r l i = {|(d, t)|}) = (
(ffilter (\<lambda>((origin, dest), t). Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r)) (ffilter (\<lambda>((origin, dest), t). origin = s) e) = {|((s, d), t)|}))"
apply (simp add: possible_steps_alt)
apply (simp only: filter_filter)
apply (rule arg_cong [of "(\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r))"])
by (rule ext, auto)
lemma possible_steps_single_out:
"ffilter (\<lambda>((origin, dest), t). origin = s) e = {|((s, d), t)|} \<Longrightarrow>
Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r) \<Longrightarrow>
possible_steps e s r l i = {|(d, t)|}"
apply (simp add: possible_steps_alt2 Abs_ffilter)
by blast
lemma possible_steps_singleton: "(possible_steps e s r l i = {|(d, t)|}) =
({((origin, dest), t) \<in> fset e. origin = s \<and> Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r)} = {((s, d), t)})"
apply (simp add: possible_steps_alt Abs_ffilter Set.filter_def)
by fast
lemma possible_steps_apply_guards:
"possible_steps e s r l i = {|(s', t)|} \<Longrightarrow>
apply_guards (Guards t) (join_ir i r)"
apply (simp add: possible_steps_singleton)
by auto
lemma possible_steps_empty:
"(possible_steps e s r l i = {||}) = (\<forall>((origin, dest), t) \<in> fset e. origin \<noteq> s \<or> Label t \<noteq> l \<or> \<not> can_take_transition t i r)"
apply (simp add: can_take_transition_def can_take_def)
apply (simp add: possible_steps_def Abs_ffilter Set.filter_def)
by auto
lemma singleton_dest:
assumes "fis_singleton (possible_steps e s r aa b)"
and "fthe_elem (possible_steps e s r aa b) = (baa, aba)"
shows "((s, baa), aba) |\<in>| e"
using assms
apply (simp add: fis_singleton_fthe_elem)
using possible_steps_alt_aux by force
lemma no_outgoing_transitions:
"ffilter (\<lambda>((s', _), _). s = s') e = {||} \<Longrightarrow>
possible_steps e s r l i = {||}"
apply (simp add: possible_steps_def)
- by auto
+ by (smt (verit, best) case_prod_beta eq_ffilter ffilter_empty ffmember_filter)
lemma ffilter_split: "ffilter (\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r)) e =
ffilter (\<lambda>((origin, dest), t). Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r)) (ffilter (\<lambda>((origin, dest), t). origin = s) e)"
by auto
lemma one_outgoing_transition:
defines "outgoing s \<equiv> (\<lambda>((origin, dest), t). origin = s)"
assumes prem: "size (ffilter (outgoing s) e) = 1"
shows "size (possible_steps e s r l i) \<le> 1"
proof-
have less_eq_1: "\<And>x::nat. (x \<le> 1) = (x = 1 \<or> x = 0)"
by auto
have size_empty: "\<And>f. (size f = 0) = (f = {||})"
subgoal for f
by (induct f, auto)
done
show ?thesis
using prem
apply (simp only: possible_steps_def)
apply (rule fimage_size_le)
apply (simp only: ffilter_split outgoing_def[symmetric])
by (metis (no_types, lifting) size_ffilter)
qed
subsection\<open>Choice\<close>
text\<open>Here we define the \texttt{choice} operator which determines whether or not two transitions are
nondeterministic.\<close>
definition choice :: "transition \<Rightarrow> transition \<Rightarrow> bool" where
"choice t t' = (\<exists> i r. apply_guards (Guards t) (join_ir i r) \<and> apply_guards (Guards t') (join_ir i r))"
definition choice_alt :: "transition \<Rightarrow> transition \<Rightarrow> bool" where
"choice_alt t t' = (\<exists> i r. apply_guards (Guards t@Guards t') (join_ir i r))"
lemma choice_alt: "choice t t' = choice_alt t t'"
by (simp add: choice_def choice_alt_def apply_guards_append)
lemma choice_symmetry: "choice x y = choice y x"
using choice_def by auto
definition deterministic :: "transition_matrix \<Rightarrow> bool" where
"deterministic e = (\<forall>s r l i. size (possible_steps e s r l i) \<le> 1)"
lemma deterministic_alt_aux: "size (possible_steps e s r l i) \<le> 1 =(
possible_steps e s r l i = {||} \<or>
(\<exists>s' t.
ffilter
(\<lambda>((origin, dest), t). origin = s \<and> Label t = l \<and> length i = Arity t \<and> apply_guards (Guards t) (join_ir i r)) e =
{|((s, s'), t)|}))"
apply (case_tac "size (possible_steps e s r l i) = 0")
apply (simp add: fset_equiv)
apply (case_tac "possible_steps e s r l i = {||}")
apply simp
apply (simp only: possible_steps_alt[symmetric])
by (metis le_neq_implies_less le_numeral_extra(4) less_one prod.collapse size_fsingleton)
lemma deterministic_alt: "deterministic e = (
\<forall>s r l i.
possible_steps e s r l i = {||} \<or>
(\<exists>s' t. ffilter (\<lambda>((origin, dest), t). origin = s \<and> (Label t) = l \<and> (length i) = (Arity t) \<and> apply_guards (Guards t) (join_ir i r)) e = {|((s, s'), t)|})
)"
using deterministic_alt_aux
by (simp add: deterministic_def)
lemma size_le_1: "size f \<le> 1 = (f = {||} \<or> (\<exists>e. f = {|e|}))"
apply standard
apply (metis bot.not_eq_extremum gr_implies_not0 le_neq_implies_less less_one size_fsingleton size_fsubset)
by auto
lemma ffilter_empty_if: "\<forall>x |\<in>| xs. \<not> P x \<Longrightarrow> ffilter P xs = {||}"
by auto
lemma empty_ffilter: "ffilter P xs = {||} = (\<forall>x |\<in>| xs. \<not> P x)"
by auto
lemma all_states_deterministic:
"(\<forall>s l i r.
ffilter (\<lambda>((origin, dest), t). origin = s \<and> (Label t) = l \<and> can_take_transition t i r) e = {||} \<or>
(\<exists>x. ffilter (\<lambda>((origin, dest), t). origin = s \<and> (Label t) = l \<and> can_take_transition t i r) e = {|x|})
) \<Longrightarrow> deterministic e"
unfolding deterministic_def
apply clarify
subgoal for s r l i
apply (erule_tac x=s in allE)
apply (erule_tac x=l in allE)
apply (erule_tac x=i in allE)
apply (erule_tac x=r in allE)
apply (simp only: size_le_1)
apply (erule disjE)
apply (rule_tac disjI1)
apply (simp add: possible_steps_def can_take_transition_def can_take_def)
apply (erule exE)
subgoal for x
apply (case_tac x)
subgoal for a b
apply (case_tac a)
apply simp
apply (induct e)
apply auto[1]
subgoal for _ _ _ ba
apply (rule disjI2)
apply (rule_tac x=ba in exI)
apply (rule_tac x=b in exI)
by (simp add: possible_steps_def can_take_transition_def[symmetric] can_take_def[symmetric])
done
done
done
done
lemma deterministic_finsert:
"\<forall>i r l.
\<forall>((a, b), t) |\<in>| ffilter (\<lambda>((origin, dest), t). origin = s) (finsert ((s, s'), t') e).
Label t = l \<and> can_take_transition t i r \<longrightarrow> \<not> can_take_transition t' i r \<Longrightarrow>
deterministic e \<Longrightarrow>
deterministic (finsert ((s, s'), t') e)"
apply (simp add: deterministic_def possible_steps_finsert can_take del: size_fset_overloaded_simps)
apply clarify
subgoal for r i
apply (erule_tac x=s in allE)
apply (erule_tac x=r in allE)
apply (erule_tac x="Label t'" in allE)
apply (erule_tac x=i in allE)
apply (erule_tac x=r in allE)
apply (erule_tac x=i in allE)
apply (erule_tac x="Label t'" in allE)
by auto
done
lemma ffilter_fBall: "(\<forall>x |\<in>| xs. P x) = (ffilter P xs = xs)"
by auto
lemma fsubset_if: "\<forall>x. x |\<in>| f1 \<longrightarrow> x |\<in>| f2 \<Longrightarrow> f1 |\<subseteq>| f2"
by auto
lemma in_possible_steps: "(((s, s'), t)|\<in>|e \<and> Label t = l \<and> can_take_transition t i r) = ((s', t) |\<in>| possible_steps e s r l i)"
apply (simp add: fmember_possible_steps)
- by (simp add: can_take_def can_take_transition_def fmember_iff_member_fset)
+ by (simp add: can_take_def can_take_transition_def)
lemma possible_steps_can_take_transition:
"(s2, t1) |\<in>| possible_steps e1 s1 r l i \<Longrightarrow> can_take_transition t1 i r"
using in_possible_steps by blast
lemma not_deterministic:
"\<exists>s l i r.
\<exists>d1 d2 t1 t2.
d1 \<noteq> d2 \<and> t1 \<noteq> t2 \<and>
((s, d1), t1) |\<in>| e \<and>
((s, d2), t2) |\<in>| e \<and>
Label t1 = Label t2 \<and>
can_take_transition t1 i r \<and>
can_take_transition t2 i r \<Longrightarrow>
\<not>deterministic e"
apply (simp add: deterministic_def not_le del: size_fset_overloaded_simps)
apply clarify
subgoal for s i r d1 d2 t1 t2
apply (rule_tac x=s in exI)
apply (rule_tac x=r in exI)
apply (rule_tac x="Label t1" in exI)
apply (rule_tac x=i in exI)
apply (case_tac "(d1, t1) |\<in>| possible_steps e s r (Label t1) i")
defer using in_possible_steps apply blast
apply (case_tac "(d2, t2) |\<in>| possible_steps e s r (Label t1) i")
apply (metis fempty_iff fsingleton_iff not_le_imp_less prod.inject size_le_1)
using in_possible_steps by force
done
lemma not_deterministic_conv:
"\<not>deterministic e \<Longrightarrow>
\<exists>s l i r.
\<exists>d1 d2 t1 t2.
(d1 \<noteq> d2 \<or> t1 \<noteq> t2) \<and>
((s, d1), t1) |\<in>| e \<and>
((s, d2), t2) |\<in>| e \<and>
Label t1 = Label t2 \<and>
can_take_transition t1 i r \<and>
can_take_transition t2 i r"
apply (simp add: deterministic_def not_le del: size_fset_overloaded_simps)
apply clarify
subgoal for s r l i
apply (case_tac "\<exists>e1 e2 f'. e1 \<noteq> e2 \<and> possible_steps e s r l i = finsert e1 (finsert e2 f')")
defer using size_gt_1 apply blast
apply (erule exE)+
subgoal for e1 e2 f'
apply (case_tac e1, case_tac e2)
subgoal for a b aa ba
apply (simp del: size_fset_overloaded_simps)
apply (rule_tac x=s in exI)
apply (rule_tac x=i in exI)
apply (rule_tac x=r in exI)
apply (rule_tac x=a in exI)
apply (rule_tac x=aa in exI)
apply (rule_tac x=b in exI)
apply (rule_tac x=ba in exI)
- by (metis finsertI1 finsert_commute in_possible_steps)
+ by (metis finsertCI in_possible_steps)
done
done
done
lemma deterministic_if:
"\<nexists>s l i r.
\<exists>d1 d2 t1 t2.
(d1 \<noteq> d2 \<or> t1 \<noteq> t2) \<and>
((s, d1), t1) |\<in>| e \<and>
((s, d2), t2) |\<in>| e \<and>
Label t1 = Label t2 \<and>
can_take_transition t1 i r \<and>
can_take_transition t2 i r \<Longrightarrow>
deterministic e"
using not_deterministic_conv by blast
lemma "\<forall>l i r.
(\<forall>((s, s'), t) |\<in>| e. Label t = l \<and> can_take_transition t i r \<and>
(\<nexists>t' s''. ((s, s''), t') |\<in>| e \<and> (s' \<noteq> s'' \<or> t' \<noteq> t) \<and> Label t' = l \<and> can_take_transition t' i r))
\<Longrightarrow> deterministic e"
apply (simp add: deterministic_def del: size_fset_overloaded_simps)
apply (rule allI)+
apply (simp only: size_le_1 possible_steps_empty)
apply (case_tac "\<exists>t s'. ((s, s'), t)|\<in>|e \<and> Label t = l \<and> can_take_transition t i r")
- defer using notin_fset apply fastforce
+ defer apply fastforce
apply (rule disjI2)
apply clarify
apply (rule_tac x="(s', t)" in exI)
apply standard
defer apply (meson fempty_fsubsetI finsert_fsubset in_possible_steps)
apply standard
apply (case_tac x)
apply (simp add: in_possible_steps[symmetric])
apply (erule_tac x="Label t" in allE)
apply (erule_tac x=i in allE)
apply (erule_tac x=r in allE)
apply (erule_tac x="((s, s'), t)" in fBallE)
defer apply simp
apply simp
apply (erule_tac x=b in allE)
apply simp
apply (erule_tac x=a in allE)
by simp
definition "outgoing_transitions e s = ffilter (\<lambda>((o, _), _). o = s) e"
lemma in_outgoing: "((s1, s2), t) |\<in>| outgoing_transitions e s = (((s1, s2), t) |\<in>| e \<and> s1 = s)"
- by (simp add: outgoing_transitions_def)
+ by (auto simp add: outgoing_transitions_def)
lemma outgoing_transitions_deterministic:
"\<forall>s.
\<forall>((s1, s2), t) |\<in>| outgoing_transitions e s.
\<forall>((s1', s2'), t') |\<in>| outgoing_transitions e s.
s2 \<noteq> s2' \<or> t \<noteq> t' \<longrightarrow> Label t = Label t' \<longrightarrow> \<not> choice t t' \<Longrightarrow> deterministic e"
apply (rule deterministic_if)
apply simp
apply (rule allI)
subgoal for s
apply (erule_tac x=s in allE)
apply (simp add: fBall_def Ball_def)
apply (rule allI)+
subgoal for i r d1 d2 t1
apply (erule_tac x=s in allE)
apply (erule_tac x=d1 in allE)
apply (erule_tac x=t1 in allE)
apply (rule impI, rule allI)
subgoal for t2
apply (case_tac "((s, d1), t1) \<in> fset (outgoing_transitions e s)")
apply simp
apply (erule_tac x=s in allE)
apply (erule_tac x=d2 in allE)
apply (erule_tac x=t2 in allE)
apply (simp add: outgoing_transitions_def choice_def can_take)
- apply (meson fmember_implies_member)
- apply (simp add: outgoing_transitions_def)
- by (meson fmember_implies_member)
+ apply meson
+ by (simp add: outgoing_transitions_def)
done
done
done
lemma outgoing_transitions_deterministic2: "(\<And>s a b ba aa bb bc.
((a, b), ba) |\<in>| outgoing_transitions e s \<Longrightarrow>
((aa, bb), bc) |\<in>| (outgoing_transitions e s) - {|((a, b), ba)|} \<Longrightarrow> b \<noteq> bb \<or> ba \<noteq> bc \<Longrightarrow> \<not>choice ba bc)
\<Longrightarrow> deterministic e"
apply (rule outgoing_transitions_deterministic)
by blast
lemma outgoing_transitions_fprod_deterministic:
"(\<And>s b ba bb bc.
(((s, b), ba), ((s, bb), bc)) \<in> fset (outgoing_transitions e s) \<times> fset (outgoing_transitions e s)
\<Longrightarrow> b \<noteq> bb \<or> ba \<noteq> bc \<Longrightarrow> Label ba = Label bc \<Longrightarrow> \<not>choice ba bc)
\<Longrightarrow> deterministic e"
apply (rule outgoing_transitions_deterministic)
apply clarify
- by (metis SigmaI fmember_implies_member in_outgoing)
+ by (metis SigmaI in_outgoing)
text\<open>The \texttt{random\_member} function returns a random member from a finite set, or
\texttt{None}, if the set is empty.\<close>
definition random_member :: "'a fset \<Rightarrow> 'a option" where
"random_member f = (if f = {||} then None else Some (Eps (\<lambda>x. x |\<in>| f)))"
lemma random_member_nonempty: "s \<noteq> {||} = (random_member s \<noteq> None)"
by (simp add: random_member_def)
lemma random_member_singleton [simp]: "random_member {|a|} = Some a"
by (simp add: random_member_def)
lemma random_member_is_member:
"random_member ss = Some s \<Longrightarrow> s |\<in>| ss"
apply (simp add: random_member_def)
by (metis equalsffemptyI option.distinct(1) option.inject verit_sko_ex_indirect)
lemma random_member_None[simp]: "random_member ss = None = (ss = {||})"
by (simp add: random_member_def)
lemma random_member_empty[simp]: "random_member {||} = None"
by simp
definition step :: "transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> label \<Rightarrow> inputs \<Rightarrow> (transition \<times> cfstate \<times> outputs \<times> registers) option" where
"step e s r l i = (case random_member (possible_steps e s r l i) of
None \<Rightarrow> None |
Some (s', t) \<Rightarrow> Some (t, s', evaluate_outputs t i r, evaluate_updates t i r)
)"
lemma possible_steps_not_empty_iff:
"step e s r a b \<noteq> None \<Longrightarrow>
\<exists>aa ba. (aa, ba) |\<in>| possible_steps e s r a b"
apply (simp add: step_def)
apply (case_tac "possible_steps e s r a b")
apply (simp add: random_member_def)
by auto
lemma step_member: "step e s r l i = Some (t, s', p, r') \<Longrightarrow> (s', t) |\<in>| possible_steps e s r l i"
apply (simp add: step_def)
apply (case_tac "random_member (possible_steps e s r l i)")
apply simp
subgoal for a by (case_tac a, simp add: random_member_is_member)
done
lemma step_outputs: "step e s r l i = Some (t, s', p, r') \<Longrightarrow> evaluate_outputs t i r = p"
apply (simp add: step_def)
apply (case_tac "random_member (possible_steps e s r l i)")
by auto
lemma step:
"possibilities = (possible_steps e s r l i) \<Longrightarrow>
random_member possibilities = Some (s', t) \<Longrightarrow>
evaluate_outputs t i r = p \<Longrightarrow>
evaluate_updates t i r = r' \<Longrightarrow>
step e s r l i = Some (t, s', p, r')"
by (simp add: step_def)
lemma step_None: "step e s r l i = None = (possible_steps e s r l i = {||})"
by (simp add: step_def prod.case_eq_if random_member_def)
lemma step_Some: "step e s r l i = Some (t, s', p, r') =
(
random_member (possible_steps e s r l i) = Some (s', t) \<and>
evaluate_outputs t i r = p \<and>
evaluate_updates t i r = r'
)"
apply (simp add: step_def)
apply (case_tac "random_member (possible_steps e s r l i)")
apply simp
subgoal for a by (case_tac a, auto)
done
lemma no_possible_steps_1:
"possible_steps e s r l i = {||} \<Longrightarrow> step e s r l i = None"
by (simp add: step_def random_member_def)
subsection\<open>Execution Observation\<close>
text\<open>One of the key features of this formalisation of EFSMs is their ability to produce
\emph{outputs}, which represent function return values. When action sequences are executed in an
EFSM, they produce a corresponding \emph{observation}.\<close>
text_raw\<open>\snip{observe}{1}{2}{%\<close>
fun observe_execution :: "transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> execution \<Rightarrow> outputs list" where
"observe_execution _ _ _ [] = []" |
"observe_execution e s r ((l, i)#as) = (
let viable = possible_steps e s r l i in
if viable = {||} then
[]
else
let (s', t) = Eps (\<lambda>x. x |\<in>| viable) in
(evaluate_outputs t i r)#(observe_execution e s' (evaluate_updates t i r) as)
)"
text_raw\<open>}%endsnip\<close>
lemma observe_execution_step_def: "observe_execution e s r ((l, i)#as) = (
case step e s r l i of
None \<Rightarrow> []|
Some (t, s', p, r') \<Rightarrow> p#(observe_execution e s' r' as)
)"
apply (simp add: step_def)
apply (case_tac "possible_steps e s r l i")
apply simp
subgoal for x S'
apply (simp add: random_member_def)
apply (case_tac "SOME xa. xa = x \<or> xa |\<in>| S'")
by simp
done
lemma observe_execution_first_outputs_equiv:
"observe_execution e1 s1 r1 ((l, i) # ts) = observe_execution e2 s2 r2 ((l, i) # ts) \<Longrightarrow>
step e1 s1 r1 l i = Some (t, s', p, r') \<Longrightarrow>
\<exists>(s2', t2)|\<in>|possible_steps e2 s2 r2 l i. evaluate_outputs t2 i r2 = p"
apply (simp only: observe_execution_step_def)
apply (case_tac "step e2 s2 r2 l i")
apply simp
subgoal for a
apply simp
apply (case_tac a)
apply clarsimp
by (meson step_member case_prodI rev_fBexI step_outputs)
done
lemma observe_execution_step:
"step e s r (fst h) (snd h) = Some (t, s', p, r') \<Longrightarrow>
observe_execution e s' r' es = obs \<Longrightarrow>
observe_execution e s r (h#es) = p#obs"
apply (cases h, simp add: step_def)
apply (case_tac "possible_steps e s r a b = {||}")
apply simp
subgoal for a b
apply (case_tac "SOME x. x |\<in>| possible_steps e s r a b")
by (simp add: random_member_def)
done
lemma observe_execution_possible_step:
"possible_steps e s r (fst h) (snd h) = {|(s', t)|} \<Longrightarrow>
apply_outputs (Outputs t) (join_ir (snd h) r) = p \<Longrightarrow>
apply_updates (Updates t) (join_ir (snd h) r) r = r' \<Longrightarrow>
observe_execution e s' r' es = obs \<Longrightarrow>
observe_execution e s r (h#es) = p#obs"
by (simp add: observe_execution_step step)
lemma observe_execution_no_possible_step:
"possible_steps e s r (fst h) (snd h) = {||} \<Longrightarrow>
observe_execution e s r (h#es) = []"
by (cases h, simp)
lemma observe_execution_no_possible_steps:
"possible_steps e1 s1 r1 (fst h) (snd h) = {||} \<Longrightarrow>
possible_steps e2 s2 r2 (fst h) (snd h) = {||} \<Longrightarrow>
(observe_execution e1 s1 r1 (h#t)) = (observe_execution e2 s2 r2 (h#t))"
by (simp add: observe_execution_no_possible_step)
lemma observe_execution_one_possible_step:
"possible_steps e1 s1 r (fst h) (snd h) = {|(s1', t1)|} \<Longrightarrow>
possible_steps e2 s2 r (fst h) (snd h) = {|(s2', t2)|} \<Longrightarrow>
apply_outputs (Outputs t1) (join_ir (snd h) r) = apply_outputs (Outputs t2) (join_ir (snd h) r) \<Longrightarrow>
apply_updates (Updates t1) (join_ir (snd h) r) r = r' \<Longrightarrow>
apply_updates (Updates t2) (join_ir (snd h) r) r = r' \<Longrightarrow>
(observe_execution e1 s1' r' t) = (observe_execution e2 s2' r' t) \<Longrightarrow>
(observe_execution e1 s1 r (h#t)) = (observe_execution e2 s2 r (h#t))"
by (simp add: observe_execution_possible_step)
subsubsection\<open>Utilities\<close>
text\<open>Here we define some utility functions to access the various key properties of a given EFSM.\<close>
definition max_reg :: "transition_matrix \<Rightarrow> nat option" where
"max_reg e = (let maxes = (fimage (\<lambda>(_, t). Transition.max_reg t) e) in if maxes = {||} then None else fMax maxes)"
definition enumerate_ints :: "transition_matrix \<Rightarrow> int set" where
"enumerate_ints e = \<Union> (image (\<lambda>(_, t). Transition.enumerate_ints t) (fset e))"
definition max_int :: "transition_matrix \<Rightarrow> int" where
"max_int e = Max (insert 0 (enumerate_ints e))"
definition max_output :: "transition_matrix \<Rightarrow> nat" where
"max_output e = fMax (fimage (\<lambda>(_, t). length (Outputs t)) e)"
definition all_regs :: "transition_matrix \<Rightarrow> nat set" where
"all_regs e = \<Union> (image (\<lambda>(_, t). enumerate_regs t) (fset e))"
text_raw\<open>\snip{finiteRegs}{1}{2}{%\<close>
lemma finite_all_regs: "finite (all_regs e)"
text_raw\<open>}%endsnip\<close>
apply (simp add: all_regs_def enumerate_regs_def)
apply clarify
apply standard
apply (rule finite_UnI)+
using GExp.finite_enumerate_regs apply blast
using AExp.finite_enumerate_regs apply blast
apply (simp add: AExp.finite_enumerate_regs prod.case_eq_if)
by auto
definition max_input :: "transition_matrix \<Rightarrow> nat option" where
"max_input e = fMax (fimage (\<lambda>(_, t). Transition.max_input t) e)"
fun maxS :: "transition_matrix \<Rightarrow> nat" where
"maxS t = (if t = {||} then 0 else fMax ((fimage (\<lambda>((origin, dest), t). origin) t) |\<union>| (fimage (\<lambda>((origin, dest), t). dest) t)))"
subsection\<open>Execution Recognition\<close>
text\<open>The \texttt{recognises} function returns true if the given EFSM recognises a given execution.
That is, the EFSM is able to respond to each event in sequence. There is no restriction on the
outputs produced. When a recognised execution is observed, it produces an accepted trace of the
EFSM.\<close>
text_raw\<open>\snip{recognises}{1}{2}{%\<close>
inductive recognises_execution :: "transition_matrix \<Rightarrow> nat \<Rightarrow> registers \<Rightarrow> execution \<Rightarrow> bool" where
base [simp]: "recognises_execution e s r []" |
step: "\<exists>(s', T) |\<in>| possible_steps e s r l i.
recognises_execution e s' (evaluate_updates T i r) t \<Longrightarrow>
recognises_execution e s r ((l, i)#t)"
text_raw\<open>}%endsnip\<close>
abbreviation "recognises e t \<equiv> recognises_execution e 0 <> t"
definition "E e = {x. recognises e x}"
lemma no_possible_steps_rejects:
"possible_steps e s r l i = {||} \<Longrightarrow> \<not> recognises_execution e s r ((l, i)#t)"
apply clarify
by (rule recognises_execution.cases, auto)
lemma recognises_step_equiv: "recognises_execution e s r ((l, i)#t) =
(\<exists>(s', T) |\<in>| possible_steps e s r l i. recognises_execution e s' (evaluate_updates T i r) t)"
apply standard
apply (rule recognises_execution.cases)
by (auto simp: recognises_execution.step)
fun recognises_prim :: "transition_matrix \<Rightarrow> nat \<Rightarrow> registers \<Rightarrow> execution \<Rightarrow> bool" where
"recognises_prim e s r [] = True" |
"recognises_prim e s r ((l, i)#t) = (
let poss_steps = possible_steps e s r l i in
(\<exists>(s', T) |\<in>| poss_steps. recognises_prim e s' (evaluate_updates T i r) t)
)"
lemma recognises_prim [code]: "recognises_execution e s r t = recognises_prim e s r t"
proof(induct t arbitrary: r s)
case Nil
then show ?case
by simp
next
case (Cons h t)
then show ?case
apply (cases h)
apply simp
apply standard
apply (rule recognises_execution.cases, simp)
apply simp
apply auto[1]
using recognises_execution.step by blast
qed
lemma recognises_single_possible_step:
assumes "possible_steps e s r l i = {|(s', t)|}"
and "recognises_execution e s' (evaluate_updates t i r) trace"
shows "recognises_execution e s r ((l, i)#trace)"
apply (rule recognises_execution.step)
using assms by auto
lemma recognises_single_possible_step_atomic:
assumes "possible_steps e s r (fst h) (snd h) = {|(s', t)|}"
and "recognises_execution e s' (apply_updates (Updates t) (join_ir (snd h) r) r) trace"
shows "recognises_execution e s r (h#trace)"
by (metis assms prod.collapse recognises_single_possible_step)
lemma recognises_must_be_possible_step:
"recognises_execution e s r (h # t) \<Longrightarrow>
\<exists>aa ba. (aa, ba) |\<in>| possible_steps e s r (fst h) (snd h)"
using recognises_step_equiv by fastforce
lemma recognises_possible_steps_not_empty:
"recognises_execution e s r (h#t) \<Longrightarrow> possible_steps e s r (fst h) (snd h) \<noteq> {||}"
apply (rule recognises_execution.cases)
by auto
lemma recognises_must_be_step:
"recognises_execution e s r (h#ts) \<Longrightarrow>
\<exists>t s' p d'. step e s r (fst h) (snd h) = Some (t, s', p, d')"
apply (cases h)
subgoal for a b
apply (simp add: recognises_step_equiv step_def)
apply clarify
apply (case_tac "(possible_steps e s r a b)")
apply (simp add: random_member_def)
apply (simp add: random_member_def)
subgoal for _ _ x S' apply (case_tac "SOME xa. xa = x \<or> xa |\<in>| S'")
by simp
done
done
lemma recognises_cons_step:
"recognises_execution e s r (h # t) \<Longrightarrow> step e s r (fst h) (snd h) \<noteq> None"
by (simp add: recognises_must_be_step)
lemma no_step_none:
"step e s r aa ba = None \<Longrightarrow> \<not> recognises_execution e s r ((aa, ba) # p)"
using recognises_cons_step by fastforce
lemma step_none_rejects:
"step e s r (fst h) (snd h) = None \<Longrightarrow> \<not> recognises_execution e s r (h#t)"
using no_step_none surjective_pairing by fastforce
lemma trace_reject:
"(\<not> recognises_execution e s r ((l, i)#t)) = (possible_steps e s r l i = {||} \<or> (\<forall>(s', T) |\<in>| possible_steps e s r l i. \<not> recognises_execution e s' (evaluate_updates T i r) t))"
using recognises_prim by fastforce
lemma trace_reject_no_possible_steps_atomic:
"possible_steps e s r (fst a) (snd a) = {||} \<Longrightarrow> \<not> recognises_execution e s r (a#t)"
using recognises_possible_steps_not_empty by auto
lemma trace_reject_later:
"\<forall>(s', T) |\<in>| possible_steps e s r l i. \<not> recognises_execution e s' (evaluate_updates T i r) t \<Longrightarrow>
\<not> recognises_execution e s r ((l, i)#t)"
using trace_reject by auto
lemma recognition_prefix_closure: "recognises_execution e s r (t@t') \<Longrightarrow> recognises_execution e s r t"
proof(induct t arbitrary: s r)
case (Cons a t)
then show ?case
apply (cases a, clarsimp)
apply (rule recognises_execution.cases)
apply simp
apply simp
by (rule recognises_execution.step, auto)
qed auto
lemma rejects_prefix: "\<not> recognises_execution e s r t \<Longrightarrow> \<not> recognises_execution e s r (t @ t')"
using recognition_prefix_closure by blast
lemma recognises_head: "recognises_execution e s r (h#t) \<Longrightarrow> recognises_execution e s r [h]"
by (simp add: recognition_prefix_closure)
subsubsection\<open>Trace Acceptance\<close>
text\<open>The \texttt{accepts} function returns true if the given EFSM accepts a given trace. That is,
the EFSM is able to respond to each event in sequence \emph{and} is able to produce the expected
output. Accepted traces represent valid runs of an EFSM.\<close>
text_raw\<open>\snip{accepts}{1}{2}{%\<close>
inductive accepts_trace :: "transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> trace \<Rightarrow> bool" where
base [simp]: "accepts_trace e s r []" |
step: "\<exists>(s', T) |\<in>| possible_steps e s r l i.
evaluate_outputs T i r = map Some p \<and> accepts_trace e s' (evaluate_updates T i r) t \<Longrightarrow>
accepts_trace e s r ((l, i, p)#t)"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{T}{1}{2}{%\<close>
definition T :: "transition_matrix \<Rightarrow> trace set" where
"T e = {t. accepts_trace e 0 <> t}"
text_raw\<open>}%endsnip\<close>
abbreviation "rejects_trace e s r t \<equiv> \<not> accepts_trace e s r t"
lemma accepts_trace_step:
"accepts_trace e s r ((l, i, p)#t) = (\<exists>(s', T) |\<in>| possible_steps e s r l i.
evaluate_outputs T i r = map Some p \<and>
accepts_trace e s' (evaluate_updates T i r) t)"
apply standard
by (rule accepts_trace.cases, auto simp: accepts_trace.step)
lemma accepts_trace_exists_possible_step:
"accepts_trace e1 s1 r1 ((aa, b, c) # t) \<Longrightarrow>
\<exists>(s1', t1)|\<in>|possible_steps e1 s1 r1 aa b.
evaluate_outputs t1 b r1 = map Some c"
using accepts_trace_step by auto
lemma rejects_trace_step:
"rejects_trace e s r ((l, i, p)#t) = (
(\<forall>(s', T) |\<in>| possible_steps e s r l i. evaluate_outputs T i r \<noteq> map Some p \<or> rejects_trace e s' (evaluate_updates T i r) t)
)"
apply (simp add: accepts_trace_step)
by auto
definition accepts_log :: "trace set \<Rightarrow> transition_matrix \<Rightarrow> bool" where
"accepts_log l e = (\<forall>t \<in> l. accepts_trace e 0 <> t)"
text_raw\<open>\snip{prefixClosure}{1}{2}{%\<close>
lemma prefix_closure: "accepts_trace e s r (t@t') \<Longrightarrow> accepts_trace e s r t"
text_raw\<open>}%endsnip\<close>
proof(induct t arbitrary: s r)
next
case (Cons a t)
then show ?case
apply (cases a, clarsimp)
apply (simp add: accepts_trace_step)
by auto
qed auto
text\<open>For code generation, it is much more efficient to re-implement the \texttt{accepts\_trace}
function primitively than to use the code generator's default setup for inductive definitions.\<close>
fun accepts_trace_prim :: "transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> trace \<Rightarrow> bool" where
"accepts_trace_prim _ _ _ [] = True" |
"accepts_trace_prim e s r ((l, i, p)#t) = (
let poss_steps = possible_steps e s r l i in
if fis_singleton poss_steps then
let (s', T) = fthe_elem poss_steps in
if evaluate_outputs T i r = map Some p then
accepts_trace_prim e s' (evaluate_updates T i r) t
else False
else
(\<exists>(s', T) |\<in>| poss_steps.
evaluate_outputs T i r = (map Some p) \<and>
accepts_trace_prim e s' (evaluate_updates T i r) t))"
lemma accepts_trace_prim [code]: "accepts_trace e s r l = accepts_trace_prim e s r l"
proof(induct l arbitrary: s r)
case (Cons a l)
then show ?case
apply (cases a)
apply (simp add: accepts_trace_step Let_def fis_singleton_alt)
by auto
qed auto
subsection\<open>EFSM Comparison\<close>
text\<open>Here, we define some different metrics of EFSM equality.\<close>
subsubsection\<open>State Isomporphism\<close>
text\<open>Two EFSMs are isomorphic with respect to states if there exists a bijective function between
the state names of the two EFSMs, i.e. the only difference between the two models is the way the
states are indexed.\<close>
definition isomorphic :: "transition_matrix \<Rightarrow> transition_matrix \<Rightarrow> bool" where
"isomorphic e1 e2 = (\<exists>f. bij f \<and> (\<forall>((s1, s2), t) |\<in>| e1. ((f s1, f s2), t) |\<in>| e2))"
subsubsection\<open>Register Isomporphism\<close>
text\<open>Two EFSMs are isomorphic with respect to registers if there exists a bijective function between
the indices of the registers in the two EFSMs, i.e. the only difference between the two models is
the way the registers are indexed.\<close>
definition rename_regs :: "(nat \<Rightarrow> nat) \<Rightarrow> transition_matrix \<Rightarrow> transition_matrix" where
"rename_regs f e = fimage (\<lambda>(tf, t). (tf, Transition.rename_regs f t)) e"
definition eq_upto_rename_strong :: "transition_matrix \<Rightarrow> transition_matrix \<Rightarrow> bool" where
"eq_upto_rename_strong e1 e2 = (\<exists>f. bij f \<and> rename_regs f e1 = e2)"
subsubsection\<open>Trace Simulation\<close>
text\<open>An EFSM, $e_1$ simulates another EFSM $e_2$ if there is a function between the states of the
states of $e_1$ and $e_1$ such that in each state, if $e_1$ can respond to the event and produce
the correct output, so can $e_2$.\<close>
text_raw\<open>\snip{traceSim}{1}{2}{%\<close>
inductive trace_simulation :: "(cfstate \<Rightarrow> cfstate) \<Rightarrow> transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow>
transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> trace \<Rightarrow> bool" where
base: "s2 = f s1 \<Longrightarrow> trace_simulation f e1 s1 r1 e2 s2 r2 []" |
step: "s2 = f s1 \<Longrightarrow>
\<forall>(s1', t1) |\<in>| ffilter (\<lambda>(s1', t1). evaluate_outputs t1 i r1 = map Some o) (possible_steps e1 s1 r1 l i).
\<exists>(s2', t2) |\<in>| possible_steps e2 s2 r2 l i. evaluate_outputs t2 i r2 = map Some o \<and>
trace_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \<Longrightarrow>
trace_simulation f e1 s1 r1 e2 s2 r2 ((l, i, o)#es)"
text_raw\<open>}%endsnip\<close>
lemma trace_simulation_step:
"trace_simulation f e1 s1 r1 e2 s2 r2 ((l, i, o)#es) = (
(s2 = f s1) \<and> (\<forall>(s1', t1) |\<in>| ffilter (\<lambda>(s1', t1). evaluate_outputs t1 i r1 = map Some o) (possible_steps e1 s1 r1 l i).
(\<exists>(s2', t2) |\<in>| possible_steps e2 s2 r2 l i. evaluate_outputs t2 i r2 = map Some o \<and>
trace_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es))
)"
apply standard
apply (rule trace_simulation.cases, simp+)
apply (rule trace_simulation.step)
apply simp
by blast
lemma trace_simulation_step_none:
"s2 = f s1 \<Longrightarrow>
\<nexists>(s1', t1) |\<in>| possible_steps e1 s1 r1 l i. evaluate_outputs t1 i r1 = map Some o \<Longrightarrow>
trace_simulation f e1 s1 r1 e2 s2 r2 ((l, i, o)#es)"
apply (rule trace_simulation.step)
apply simp
apply (case_tac "ffilter (\<lambda>(s1', t1). evaluate_outputs t1 i r1 = map Some o) (possible_steps e1 s1 r1 l i)")
apply simp
by fastforce
definition "trace_simulates e1 e2 = (\<exists>f. \<forall>t. trace_simulation f e1 0 <> e2 0 <> t)"
lemma rejects_trace_simulation:
"rejects_trace e2 s2 r2 t \<Longrightarrow>
accepts_trace e1 s1 r1 t \<Longrightarrow>
\<not>trace_simulation f e1 s1 r1 e2 s2 r2 t"
proof(induct t arbitrary: s1 r1 s2 r2)
case Nil
then show ?case
using accepts_trace.base by blast
next
case (Cons a t)
then show ?case
apply (cases a)
apply (simp add: rejects_trace_step)
apply (simp add: accepts_trace_step)
apply clarify
apply (rule trace_simulation.cases)
apply simp
apply simp
apply clarsimp
subgoal for l o _ _ i
- apply (case_tac "ffilter (\<lambda>(s1', t1). evaluate_outputs t1 i r1 = map Some o) (possible_steps e1 s1 r1 l i) = {||}")
- apply auto[1]
by fastforce
done
qed
lemma accepts_trace_simulation:
"accepts_trace e1 s1 r1 t \<Longrightarrow>
trace_simulation f e1 s1 r1 e2 s2 r2 t \<Longrightarrow>
accepts_trace e2 s2 r2 t"
using rejects_trace_simulation by blast
lemma simulates_trace_subset: "trace_simulates e1 e2 \<Longrightarrow> T e1 \<subseteq> T e2"
using T_def accepts_trace_simulation trace_simulates_def by fastforce
subsubsection\<open>Trace Equivalence\<close>
text\<open>Two EFSMs are trace equivalent if they accept the same traces. This is the intuitive definition
of ``observable equivalence'' between the behaviours of the two models. If two EFSMs are trace
equivalent, there is no trace which can distinguish the two.\<close>
text_raw\<open>\snip{traceEquiv}{1}{2}{%\<close>
definition "trace_equivalent e1 e2 = (T e1 = T e2)"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{simEquiv}{1}{2}{%\<close>
lemma simulation_implies_trace_equivalent:
"trace_simulates e1 e2 \<Longrightarrow> trace_simulates e2 e1 \<Longrightarrow> trace_equivalent e1 e2"
text_raw\<open>}%endsnip\<close>
using simulates_trace_subset trace_equivalent_def by auto
lemma trace_equivalent_reflexive: "trace_equivalent e1 e1"
by (simp add: trace_equivalent_def)
lemma trace_equivalent_symmetric:
"trace_equivalent e1 e2 = trace_equivalent e2 e1"
using trace_equivalent_def by auto
lemma trace_equivalent_transitive:
"trace_equivalent e1 e2 \<Longrightarrow>
trace_equivalent e2 e3 \<Longrightarrow>
trace_equivalent e1 e3"
by (simp add: trace_equivalent_def)
text\<open>Two EFSMs are trace equivalent if they accept the same traces.\<close>
lemma trace_equivalent:
"\<forall>t. accepts_trace e1 0 <> t = accepts_trace e2 0 <> t \<Longrightarrow> trace_equivalent e1 e2"
by (simp add: T_def trace_equivalent_def)
lemma accepts_trace_step_2: "(s2', t2) |\<in>| possible_steps e2 s2 r2 l i \<Longrightarrow>
accepts_trace e2 s2' (evaluate_updates t2 i r2) t \<Longrightarrow>
evaluate_outputs t2 i r2 = map Some p \<Longrightarrow>
accepts_trace e2 s2 r2 ((l, i, p)#t)"
by (rule accepts_trace.step, auto)
subsubsection\<open>Execution Simulation\<close>
text\<open>Execution simulation is similar to trace simulation but for executions rather than traces.
Execution simulation has no notion of ``expected'' output. It simply requires that the simulating
EFSM must be able to produce equivalent output for each action.\<close>
text_raw\<open>\snip{execSim}{1}{2}{%\<close>
inductive execution_simulation :: "(cfstate \<Rightarrow> cfstate) \<Rightarrow> transition_matrix \<Rightarrow> cfstate \<Rightarrow>
registers \<Rightarrow> transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> execution \<Rightarrow> bool" where
base: "s2 = f s1 \<Longrightarrow> execution_simulation f e1 s1 r1 e2 s2 r2 []" |
step: "s2 = f s1 \<Longrightarrow>
\<forall>(s1', t1) |\<in>| (possible_steps e1 s1 r1 l i).
\<exists>(s2', t2) |\<in>| possible_steps e2 s2 r2 l i.
evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \<and>
execution_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \<Longrightarrow>
execution_simulation f e1 s1 r1 e2 s2 r2 ((l, i)#es)"
text_raw\<open>}%endsnip\<close>
definition "execution_simulates e1 e2 = (\<exists>f. \<forall>t. execution_simulation f e1 0 <> e2 0 <> t)"
lemma execution_simulation_step:
"execution_simulation f e1 s1 r1 e2 s2 r2 ((l, i)#es) =
(s2 = f s1 \<and>
(\<forall>(s1', t1) |\<in>| (possible_steps e1 s1 r1 l i).
(\<exists>(s2', t2) |\<in>| possible_steps e2 s2 r2 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \<and>
execution_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es))
)"
apply standard
apply (rule execution_simulation.cases)
apply simp
apply simp
apply simp
apply (rule execution_simulation.step)
apply simp
by blast
text_raw\<open>\snip{execTraceSim}{1}{2}{%\<close>
lemma execution_simulation_trace_simulation:
"execution_simulation f e1 s1 r1 e2 s2 r2 (map (\<lambda>(l, i, o). (l, i)) t) \<Longrightarrow>
trace_simulation f e1 s1 r1 e2 s2 r2 t"
text_raw\<open>}%endsnip\<close>
proof(induct t arbitrary: s1 s2 r1 r2)
case Nil
then show ?case
apply (rule execution_simulation.cases)
apply (simp add: trace_simulation.base)
by simp
next
case (Cons a t)
then show ?case
apply (cases a, clarsimp)
apply (rule execution_simulation.cases)
apply simp
apply simp
apply (rule trace_simulation.step)
apply simp
apply clarsimp
subgoal for _ _ _ aa ba
apply (erule_tac x="(aa, ba)" in fBallE)
apply clarsimp
apply blast
by simp
done
qed
lemma execution_simulates_trace_simulates:
"execution_simulates e1 e2 \<Longrightarrow> trace_simulates e1 e2"
apply (simp add: execution_simulates_def trace_simulates_def)
using execution_simulation_trace_simulation by blast
subsubsection\<open>Executional Equivalence\<close>
text\<open>Two EFSMs are executionally equivalent if there is no execution which can distinguish between
the two. That is, for every execution, they must produce equivalent outputs.\<close>
text_raw\<open>\snip{execEquiv}{1}{2}{%\<close>
inductive executionally_equivalent :: "transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow>
transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> execution \<Rightarrow> bool" where
base [simp]: "executionally_equivalent e1 s1 r1 e2 s2 r2 []" |
step: "\<forall>(s1', t1) |\<in>| possible_steps e1 s1 r1 l i.
\<exists>(s2', t2) |\<in>| possible_steps e2 s2 r2 l i.
evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \<and>
executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \<Longrightarrow>
\<forall>(s2', t2) |\<in>| possible_steps e2 s2 r2 l i.
\<exists>(s1', t1) |\<in>| possible_steps e1 s1 r1 l i.
evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \<and>
executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \<Longrightarrow>
executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es)"
text_raw\<open>}%endsnip\<close>
lemma executionally_equivalent_step:
"executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es) = (
(\<forall>(s1', t1) |\<in>| (possible_steps e1 s1 r1 l i). (\<exists>(s2', t2) |\<in>| possible_steps e2 s2 r2 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \<and>
executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es)) \<and>
(\<forall>(s2', t2) |\<in>| (possible_steps e2 s2 r2 l i). (\<exists>(s1', t1) |\<in>| possible_steps e1 s1 r1 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \<and>
executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es)))"
apply standard
apply (rule executionally_equivalent.cases)
apply simp
apply simp
apply simp
by (rule executionally_equivalent.step, auto)
lemma execution_end:
"possible_steps e1 s1 r1 l i = {||} \<Longrightarrow>
possible_steps e2 s2 r2 l i = {||} \<Longrightarrow>
executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es)"
by (simp add: executionally_equivalent_step)
lemma possible_steps_disparity:
"possible_steps e1 s1 r1 l i \<noteq> {||} \<Longrightarrow>
possible_steps e2 s2 r2 l i = {||} \<Longrightarrow>
\<not>executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es)"
by (simp add: executionally_equivalent_step, auto)
lemma executionally_equivalent_acceptance_map:
"executionally_equivalent e1 s1 r1 e2 s2 r2 (map (\<lambda>(l, i, o). (l, i)) t) \<Longrightarrow>
accepts_trace e2 s2 r2 t = accepts_trace e1 s1 r1 t"
proof(induct t arbitrary: s1 s2 r1 r2)
case (Cons a t)
then show ?case
apply (cases a, simp)
apply (rule executionally_equivalent.cases)
apply simp
apply simp
apply clarsimp
apply standard
subgoal for i p l
apply (rule accepts_trace.cases)
apply simp
apply simp
apply clarsimp
subgoal for aa b
apply (rule accepts_trace.step)
apply (erule_tac x="(aa, b)" in fBallE[of "possible_steps e2 s2 r2 l i"])
defer apply simp
apply simp
by blast
done
apply (rule accepts_trace.cases)
apply simp
apply simp
apply clarsimp
subgoal for _ _ _ aa b
apply (rule accepts_trace.step)
apply (erule_tac x="(aa, b)" in fBallE)
defer apply simp
apply simp
by fastforce
done
qed auto
lemma executionally_equivalent_acceptance:
"\<forall>x. executionally_equivalent e1 s1 r1 e2 s2 r2 x \<Longrightarrow> accepts_trace e1 s1 r1 t \<Longrightarrow> accepts_trace e2 s2 r2 t"
using executionally_equivalent_acceptance_map by blast
lemma executionally_equivalent_trace_equivalent:
"\<forall>x. executionally_equivalent e1 0 <> e2 0 <> x \<Longrightarrow> trace_equivalent e1 e2"
apply (rule trace_equivalent)
apply clarify
subgoal for t apply (erule_tac x="map (\<lambda>(l, i, o). (l, i)) t" in allE)
by (simp add: executionally_equivalent_acceptance_map)
done
lemma executionally_equivalent_symmetry:
"executionally_equivalent e1 s1 r1 e2 s2 r2 x \<Longrightarrow>
executionally_equivalent e2 s2 r2 e1 s1 r1 x"
proof(induct x arbitrary: s1 s2 r1 r2)
case (Cons a x)
then show ?case
apply (cases a, clarsimp)
apply (simp add: executionally_equivalent_step)
apply standard
apply (rule fBallI)
apply clarsimp
subgoal for aa b aaa ba
apply (erule_tac x="(aaa, ba)" in fBallE[of "possible_steps e2 s2 r2 aa b"])
by (force, simp)
apply (rule fBallI)
apply clarsimp
subgoal for aa b aaa ba
apply (erule_tac x="(aaa, ba)" in fBallE)
by (force, simp)
done
qed auto
lemma executionally_equivalent_transitivity:
"executionally_equivalent e1 s1 r1 e2 s2 r2 x \<Longrightarrow>
executionally_equivalent e2 s2 r2 e3 s3 r3 x \<Longrightarrow>
executionally_equivalent e1 s1 r1 e3 s3 r3 x"
proof(induct x arbitrary: s1 s2 s3 r1 r2 r3)
case (Cons a x)
then show ?case
apply (cases a, clarsimp)
apply (simp add: executionally_equivalent_step)
apply clarsimp
apply standard
apply (rule fBallI)
apply clarsimp
subgoal for aa b ab ba
apply (erule_tac x="(ab, ba)" in fBallE[of "possible_steps e1 s1 r1 aa b"])
prefer 2 apply simp
apply simp
apply (erule fBexE)
subgoal for x apply (case_tac x)
apply simp
by blast
done
apply (rule fBallI)
apply clarsimp
subgoal for aa b ab ba
apply (erule_tac x="(ab, ba)" in fBallE[of "possible_steps e3 s3 r3 aa b"])
prefer 2 apply simp
apply simp
apply (erule fBexE)
subgoal for x apply (case_tac x)
apply clarsimp
subgoal for aaa baa
apply (erule_tac x="(aaa, baa)" in fBallE[of "possible_steps e2 s2 r2 aa b"])
prefer 2 apply simp
apply simp
by blast
done
done
done
qed auto
subsection\<open>Reachability\<close>
text\<open>Here, we define the function \texttt{visits} which returns true if the given execution
leaves the given EFSM in the given state.\<close>
text_raw\<open>\snip{reachable}{1}{2}{%\<close>
inductive visits :: "cfstate \<Rightarrow> transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> execution \<Rightarrow> bool" where
base [simp]: "visits s e s r []" |
step: "\<exists>(s', T) |\<in>| possible_steps e s r l i. visits target e s' (evaluate_updates T i r) t \<Longrightarrow>
visits target e s r ((l, i)#t)"
definition "reachable s e = (\<exists>t. visits s e 0 <> t)"
text_raw\<open>}%endsnip\<close>
lemma no_further_steps:
"s \<noteq> s' \<Longrightarrow> \<not> visits s e s' r []"
apply safe
apply (rule visits.cases)
by auto
lemma visits_base: "visits target e s r [] = (s = target)"
by (metis visits.base no_further_steps)
lemma visits_step:
"visits target e s r (h#t) = (\<exists>(s', T) |\<in>| possible_steps e s r (fst h) (snd h). visits target e s' (evaluate_updates T (snd h) r) t)"
apply standard
apply (rule visits.cases)
apply simp+
apply (cases h)
using visits.step by auto
lemma reachable_initial: "reachable 0 e"
apply (simp add: reachable_def)
apply (rule_tac x="[]" in exI)
by simp
lemma visits_finsert:
"visits s e s' r t \<Longrightarrow> visits s (finsert ((aa, ba), b) e) s' r t"
proof(induct t arbitrary: s' r)
case Nil
then show ?case
by (simp add: visits_base)
next
case (Cons a t)
then show ?case
apply (simp add: visits_step)
apply (erule fBexE)
apply (rule_tac x=x in fBexI)
apply auto[1]
by (simp add: possible_steps_finsert)
qed
lemma reachable_finsert:
"reachable s e \<Longrightarrow> reachable s (finsert ((aa, ba), b) e)"
apply (simp add: reachable_def)
by (meson visits_finsert)
lemma reachable_finsert_contra:
"\<not> reachable s (finsert ((aa, ba), b) e) \<Longrightarrow> \<not>reachable s e"
using reachable_finsert by blast
lemma visits_empty: "visits s e s' r [] = (s = s')"
apply standard
by (rule visits.cases, auto)
definition "remove_state s e = ffilter (\<lambda>((from, to), t). from \<noteq> s \<and> to \<noteq> s) e"
text_raw\<open>\snip{obtainable}{1}{2}{%\<close>
inductive "obtains" :: "cfstate \<Rightarrow> registers \<Rightarrow> transition_matrix \<Rightarrow> cfstate \<Rightarrow> registers \<Rightarrow> execution \<Rightarrow> bool" where
base [simp]: "obtains s r e s r []" |
step: "\<exists>(s'', T) |\<in>| possible_steps e s' r' l i. obtains s r e s'' (evaluate_updates T i r') t \<Longrightarrow>
obtains s r e s' r' ((l, i)#t)"
definition "obtainable s r e = (\<exists>t. obtains s r e 0 <> t)"
text_raw\<open>}%endsnip\<close>
lemma obtains_obtainable:
"obtains s r e 0 <> t \<Longrightarrow> obtainable s r e"
apply (simp add: obtainable_def)
by auto
lemma obtains_base: "obtains s r e s' r' [] = (s = s' \<and> r = r')"
apply standard
by (rule obtains.cases, auto)
lemma obtains_step: "obtains s r e s' r' ((l, i)#t) = (\<exists>(s'', T) |\<in>| possible_steps e s' r' l i. obtains s r e s'' (evaluate_updates T i r') t)"
apply standard
by (rule obtains.cases, auto simp add: obtains.step)
lemma obtains_recognises:
"obtains s c e s' r t \<Longrightarrow> recognises_execution e s' r t"
proof(induct t arbitrary: s' r)
case Nil
then show ?case
by (simp add: obtains_base)
next
case (Cons a t)
then show ?case
apply (cases a)
apply simp
apply (rule obtains.cases)
apply simp
apply simp
apply clarsimp
using recognises_execution.step by fastforce
qed
lemma ex_comm4:
"(\<exists>c1 s a b. (a, b) \<in> fset (possible_steps e s' r l i) \<and> obtains s c1 e a (evaluate_updates b i r) t) =
(\<exists>a b s c1. (a, b) \<in> fset (possible_steps e s' r l i) \<and> obtains s c1 e a (evaluate_updates b i r) t)"
by auto
lemma recognises_execution_obtains:
"recognises_execution e s' r t \<Longrightarrow> \<exists>c1 s. obtains s c1 e s' r t"
proof(induct t arbitrary: s' r)
case Nil
then show ?case
by (simp add: obtains_base)
next
case (Cons a t)
then show ?case
apply (cases a)
apply (simp add: obtains_step)
apply (rule recognises_execution.cases)
apply simp
apply simp
apply clarsimp
apply (simp add: fBex_def Bex_def ex_comm4)
subgoal for _ _ aa ba
apply (rule_tac x=aa in exI)
apply (rule_tac x=ba in exI)
- apply (simp add: fmember_implies_member)
+ apply simp
by blast
done
qed
lemma obtainable_empty_efsm:
"obtainable s c {||} = (s=0 \<and> c = <>)"
apply (simp add: obtainable_def)
apply standard
apply (metis ffilter_empty no_outgoing_transitions no_step_none obtains.cases obtains_recognises step_None)
using obtains_base by blast
lemma obtains_visits: "obtains s r e s' r' t \<Longrightarrow> visits s e s' r' t"
proof(induct t arbitrary: s' r')
case Nil
then show ?case
by (simp add: obtains_base)
next
case (Cons a t)
then show ?case
apply (cases a)
apply (rule obtains.cases)
apply simp
apply simp
apply clarsimp
apply (rule visits.step)
by auto
qed
lemma unobtainable_if: "\<not> visits s e s' r' t \<Longrightarrow> \<not> obtains s r e s' r' t"
using obtains_visits by blast
lemma obtainable_if_unreachable: "\<not>reachable s e \<Longrightarrow> \<not>obtainable s r e"
by (simp add: reachable_def obtainable_def unobtainable_if)
lemma obtains_step_append:
"obtains s r e s' r' t \<Longrightarrow>
(s'', ta) |\<in>| possible_steps e s r l i \<Longrightarrow>
obtains s'' (evaluate_updates ta i r) e s' r' (t @ [(l, i)])"
proof(induct t arbitrary: s' r')
case Nil
then show ?case
apply (simp add: obtains_base)
apply (rule obtains.step)
apply (rule_tac x="(s'', ta)" in fBexI)
by auto
next
case (Cons a t)
then show ?case
apply simp
apply (rule obtains.cases)
apply simp
apply simp
apply clarsimp
apply (rule obtains.step)
by auto
qed
lemma reachable_if_obtainable_step:
"obtainable s r e \<Longrightarrow> \<exists>l i t. (s', t) |\<in>| possible_steps e s r l i \<Longrightarrow> reachable s' e"
apply (simp add: reachable_def obtainable_def)
apply clarify
subgoal for t l i
apply (rule_tac x="t@[(l, i)]" in exI)
using obtains_step_append unobtainable_if by blast
done
lemma possible_steps_remove_unreachable:
"obtainable s r e \<Longrightarrow>
\<not> reachable s' e \<Longrightarrow>
possible_steps (remove_state s' e) s r l i = possible_steps e s r l i"
apply standard
apply (simp add: fsubset_eq)
apply (rule fBallI)
apply clarsimp
apply (metis ffmember_filter in_possible_steps remove_state_def)
apply (simp add: fsubset_eq)
apply (rule fBallI)
apply clarsimp
subgoal for a b
apply (case_tac "a = s'")
using reachable_if_obtainable_step apply blast
apply (simp add: remove_state_def)
by (metis (mono_tags, lifting) ffmember_filter in_possible_steps obtainable_if_unreachable old.prod.case)
done
text_raw\<open>\snip{removeUnreachableArb}{1}{2}{%\<close>
lemma executionally_equivalent_remove_unreachable_state_arbitrary:
"obtainable s r e \<Longrightarrow> \<not> reachable s' e \<Longrightarrow> executionally_equivalent e s r (remove_state s' e) s r x"
text_raw\<open>}%endsnip\<close>
proof(induct x arbitrary: s r)
case (Cons a x)
then show ?case
apply (cases a, simp)
apply (rule executionally_equivalent.step)
apply (simp add: possible_steps_remove_unreachable)
apply standard
apply clarsimp
subgoal for aa b ab ba
apply (rule_tac x="(ab, ba)" in fBexI)
apply (metis (mono_tags, lifting) obtainable_def obtains_step_append case_prodI)
apply simp
done
apply (rule fBallI)
apply clarsimp
apply (rule_tac x="(ab, ba)" in fBexI)
apply simp
apply (metis obtainable_def obtains_step_append possible_steps_remove_unreachable)
by (simp add: possible_steps_remove_unreachable)
qed auto
text_raw\<open>\snip{removeUnreachable}{1}{2}{%\<close>
lemma executionally_equivalent_remove_unreachable_state:
"\<not> reachable s' e \<Longrightarrow> executionally_equivalent e 0 <> (remove_state s' e) 0 <> x"
text_raw\<open>}%endsnip\<close>
by (meson executionally_equivalent_remove_unreachable_state_arbitrary
obtains.simps obtains_obtainable)
subsection\<open>Transition Replacement\<close>
text\<open>Here, we define the function \texttt{replace} to replace one transition with another, and prove
some of its properties.\<close>
definition "replace e1 old new = fimage (\<lambda>x. if x = old then new else x) e1"
lemma replace_finsert:
"replace (finsert ((aaa, baa), b) e1) old new = (if ((aaa, baa), b) = old then (finsert new (replace e1 old new)) else (finsert ((aaa, baa), b) (replace e1 old new)))"
by (simp add: replace_def)
lemma possible_steps_replace_unchanged:
"((s, aa), ba) \<noteq> ((s1, s2), t1) \<Longrightarrow>
(aa, ba) |\<in>| possible_steps e1 s r l i \<Longrightarrow>
(aa, ba) |\<in>| possible_steps (replace e1 ((s1, s2), t1) ((s1, s2), t2)) s r l i"
- apply (simp add: in_possible_steps[symmetric] replace_def)
- by fastforce
+ by (simp add: in_possible_steps[symmetric] replace_def)
end
diff --git a/thys/Extended_Finite_State_Machines/EFSM_LTL.thy b/thys/Extended_Finite_State_Machines/EFSM_LTL.thy
--- a/thys/Extended_Finite_State_Machines/EFSM_LTL.thy
+++ b/thys/Extended_Finite_State_Machines/EFSM_LTL.thy
@@ -1,281 +1,281 @@
section\<open>LTL for EFSMs\<close>
text\<open>This theory builds off the \texttt{Linear\_Temporal\_Logic\_on\_Streams} theory from the HOL
library and defines functions to ease the expression of LTL properties over EFSMs. Since the LTL
operators effectively act over traces of models we must find a way to express models as streams.\<close>
theory EFSM_LTL
imports "Extended_Finite_State_Machines.EFSM" "HOL-Library.Linear_Temporal_Logic_on_Streams"
begin
text_raw\<open>\snip{statedef}{1}{2}{%\<close>
record state =
statename :: "nat option"
datastate :: registers
action :: action
"output" :: outputs
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{whitebox}{1}{2}{%\<close>
type_synonym whitebox_trace = "state stream"
text_raw\<open>}%endsnip\<close>
type_synonym property = "whitebox_trace \<Rightarrow> bool"
abbreviation label :: "state \<Rightarrow> String.literal" where
"label s \<equiv> fst (action s)"
abbreviation inputs :: "state \<Rightarrow> value list" where
"inputs s \<equiv> snd (action s)"
text_raw\<open>\snip{ltlStep}{1}{2}{%\<close>
fun ltl_step :: "transition_matrix \<Rightarrow> cfstate option \<Rightarrow> registers \<Rightarrow> action \<Rightarrow> (nat option \<times> outputs \<times> registers)" where
"ltl_step _ None r _ = (None, [], r)" |
"ltl_step e (Some s) r (l, i) = (let possibilities = possible_steps e s r l i in
if possibilities = {||} then (None, [], r)
else
let (s', t) = Eps (\<lambda>x. x |\<in>| possibilities) in
(Some s', (evaluate_outputs t i r), (evaluate_updates t i r))
)"
text_raw\<open>}%endsnip\<close>
lemma ltl_step_singleton:
"\<exists>t. possible_steps e n r (fst v) (snd v) = {|(aa, t)|} \<and> evaluate_outputs t (snd v) r = b \<and> evaluate_updates t (snd v) r = c\<Longrightarrow>
ltl_step e (Some n) r v = (Some aa, b, c)"
apply (cases v)
by auto
lemma ltl_step_none: "possible_steps e s r a b = {||} \<Longrightarrow> ltl_step e (Some s) r (a, b) = (None, [], r)"
by simp
lemma ltl_step_none_2: "possible_steps e s r (fst ie) (snd ie) = {||} \<Longrightarrow> ltl_step e (Some s) r ie = (None, [], r)"
by (metis ltl_step_none prod.exhaust_sel)
lemma ltl_step_alt: "ltl_step e (Some s) r t = (
let possibilities = possible_steps e s r (fst t) (snd t) in
if possibilities = {||} then
(None, [], r)
else
let (s', t') = Eps (\<lambda>x. x |\<in>| possibilities) in
(Some s', (apply_outputs (Outputs t') (join_ir (snd t) r)), (apply_updates (Updates t') (join_ir (snd t) r) r))
)"
by (case_tac t, simp add: Let_def)
lemma ltl_step_some:
assumes "possible_steps e s r l i = {|(s', t)|}"
and "evaluate_outputs t i r = p"
and "evaluate_updates t i r = r'"
shows "ltl_step e (Some s) r (l, i) = (Some s', p, r')"
by (simp add: assms)
lemma ltl_step_cases:
assumes invalid: "P (None, [], r)"
and valid: "\<forall>(s', t) |\<in>| (possible_steps e s r l i). P (Some s', (evaluate_outputs t i r), (evaluate_updates t i r))"
shows "P (ltl_step e (Some s) r (l, i))"
apply simp
apply (case_tac "possible_steps e s r l i")
apply (simp add: invalid)
apply simp
subgoal for x S'
apply (case_tac "SOME xa. xa = x \<or> xa |\<in>| S'")
apply simp
apply (insert assms(2))
- apply (simp add: fBall_def Ball_def fmember_def)
+ apply (simp add: fBall_def Ball_def)
by (metis (mono_tags, lifting) fst_conv prod.case_eq_if snd_conv someI_ex)
done
text\<open>The \texttt{make\_full\_observation} function behaves similarly to \texttt{observe\_execution}
from the \texttt{EFSM} theory. The main difference in behaviour is what is recorded. While the
observe execution function simply observes an execution of the EFSM to produce the corresponding
output for each action, the intention here is to record every detail of execution, including the
values of internal variables.
Thinking of each action as a step forward in time, there are five components which characterise
a given point in the execution of an EFSM. At each point, the model has a current control state and
data state. Each action has a label and some input parameters, and its execution may produce
some observableoutput. It is therefore sufficient to provide a stream of 5-tuples containing the
current control state, data state, the label and inputs of the action, and computed output. The
make full observation function can then be defined as in Figure 9.1, with an additional
function watch defined on top of this which starts the make full observation off in the
initial control state with the empty data state.
Careful inspection of the definition reveals another way that \texttt{make\_full\_observation}
differs from \texttt{observe\_execution}. Rather than taking a cfstate, it takes a cfstate option.
The reason for this is that we need to make our EFSM models complete. That is, we need them to be
able to respond to every action from every state like a DFA. If a model does not recognise a given
action in a given state, we cannot simply stop processing because we are working with necessarily
infinite traces. Since these traces are generated by observing action sequences, the make full
observation function must keep processing whether there is a viable transition or not.
To support this, the make full observation adds an implicit ``sink state'' to every EFSM it
processes by lifting control flow state indices from \texttt{nat} to \texttt{nat option} such that
state $n$ is seen as state \texttt{Some} $n$. The control flow state \texttt{None} represents a sink
state. If a model is unable to recognise a particular action from its current state, it moves into
the \texttt{None} state. From here, the behaviour is constant for the rest of the time --- the
control flow state remains None; the data state does not change, and no output is produced.\<close>
text_raw\<open>\snip{makeFullObservation}{1}{2}{%\<close>
primcorec make_full_observation :: "transition_matrix \<Rightarrow> cfstate option \<Rightarrow> registers \<Rightarrow> outputs \<Rightarrow> action stream \<Rightarrow> whitebox_trace" where
"make_full_observation e s d p i = (
let (s', o', d') = ltl_step e s d (shd i) in
\<lparr>statename = s, datastate = d, action=(shd i), output = p\<rparr>##(make_full_observation e s' d' o' (stl i))
)"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{watch}{1}{2}{%\<close>
abbreviation watch :: "transition_matrix \<Rightarrow> action stream \<Rightarrow> whitebox_trace" where
"watch e i \<equiv> (make_full_observation e (Some 0) <> [] i)"
text_raw\<open>}%endsnip\<close>
subsection\<open>Expressing Properties\<close>
text\<open>In order to simplify the expression and understanding of properties, this theory defines a
number of named functions which can be used to express certain properties of EFSMs.\<close>
subsubsection\<open>State Equality\<close>
text\<open>The \textsc{state\_eq} takes a cfstate option representing a control flow state index and
returns true if this is the control flow state at the head of the full observation.\<close>
abbreviation state_eq :: "cfstate option \<Rightarrow> whitebox_trace \<Rightarrow> bool" where
"state_eq v s \<equiv> statename (shd s) = v"
lemma state_eq_holds: "state_eq s = holds (\<lambda>x. statename x = s)"
apply (rule ext)
by (simp add: holds_def)
lemma state_eq_None_not_Some: "state_eq None s \<Longrightarrow> \<not> state_eq (Some n) s"
by simp
subsubsection\<open>Label Equality\<close>
text\<open>The \textsc{label\_eq} function takes a string and returns true if this is equal to the label
at the head of the full observation.\<close>
abbreviation "label_eq v s \<equiv> fst (action (shd s)) = (String.implode v)"
lemma watch_label: "label_eq l (watch e t) = (fst (shd t) = String.implode l)"
by simp
subsubsection\<open>Input Equality\<close>
text\<open>The \textsc{input\_eq} function takes a value list and returns true if this is equal to the
input at the head of the full observation.\<close>
abbreviation "input_eq v s \<equiv> inputs (shd s) = v"
subsubsection\<open>Action Equality\<close>
text\<open>The \textsc{action\_eq} function takes a (label, value list) pair and returns true if this is
equal to the action at the head of the full observation. This effectively combines
\texttt{label\_eq} and \texttt{input\_eq} into one function.\<close>
abbreviation "action_eq e \<equiv> label_eq (fst e) aand input_eq (snd e)"
subsubsection\<open>Output Equality\<close>
text\<open>The \textsc{output\_eq} function takes a takes a value option list and returns true if this is
equal to the output at the head of the full observation.\<close>
abbreviation "output_eq v s \<equiv> output (shd s) = v"
text_raw\<open>\snip{ltlVName}{1}{2}{%\<close>
datatype ltl_vname = Ip nat | Op nat | Rg nat
text_raw\<open>}%endsnip\<close>
subsubsection\<open>Checking Arbitrary Expressions\<close>
text\<open>The \textsc{check\_exp} function takes a guard expression and returns true if the guard
expression evaluates to true in the given state.\<close>
type_synonym ltl_gexp = "ltl_vname gexp"
definition join_iro :: "value list \<Rightarrow> registers \<Rightarrow> outputs \<Rightarrow> ltl_vname datastate" where
"join_iro i r p = (\<lambda>x. case x of
Rg n \<Rightarrow> r $ n |
Ip n \<Rightarrow> Some (i ! n) |
Op n \<Rightarrow> p ! n
)"
lemma join_iro_R [simp]: "join_iro i r p (Rg n) = r $ n"
by (simp add: join_iro_def)
abbreviation "check_exp g s \<equiv> (gval g (join_iro (snd (action (shd s))) (datastate (shd s)) (output (shd s))) = trilean.true)"
lemma alw_ev: "alw f = not (ev (\<lambda>s. \<not>f s))"
by simp
lemma alw_state_eq_smap:
"alw (state_eq s) ss = alw (\<lambda>ss. shd ss = s) (smap statename ss)"
apply standard
apply (simp add: alw_iff_sdrop )
by (simp add: alw_mono alw_smap )
subsection\<open>Sink State\<close>
text\<open>Once the sink state is entered, it cannot be left and there are no outputs or updates
henceforth.\<close>
lemma shd_state_is_none: "(state_eq None) (make_full_observation e None r p t)"
by simp
lemma unfold_observe_none: "make_full_observation e None d p t = (\<lparr>statename = None, datastate = d, action=(shd t), output = p\<rparr>##(make_full_observation e None d [] (stl t)))"
by (simp add: stream.expand)
lemma once_none_always_none_aux:
assumes "\<exists> p r i. j = (make_full_observation e None r p) i"
shows "alw (state_eq None) j"
using assms apply coinduct
apply simp
by fastforce
lemma once_none_always_none: "alw (state_eq None) (make_full_observation e None r p t)"
using once_none_always_none_aux by blast
lemma once_none_nxt_always_none: "alw (nxt (state_eq None)) (make_full_observation e None r p t)"
using once_none_always_none
by (simp add: alw_iff_sdrop del: sdrop.simps)
lemma snth_sconst: "(\<forall>i. s !! i = h) = (s = sconst h)"
by (auto simp add: sconst_alt sset_range)
lemma alw_sconst: "(alw (\<lambda>xs. shd xs = h) t) = (t = sconst h)"
by (simp add: snth_sconst[symmetric] alw_iff_sdrop)
lemma smap_statename_None: "smap statename (make_full_observation e None r p i) = sconst None"
by (meson EFSM_LTL.alw_sconst alw_state_eq_smap once_none_always_none)
lemma alw_not_some: "alw (\<lambda>xs. statename (shd xs) \<noteq> Some s) (make_full_observation e None r p t)"
by (metis (mono_tags, lifting) alw_mono once_none_always_none option.distinct(1) )
lemma state_none: "((state_eq None) impl nxt (state_eq None)) (make_full_observation e s r p t)"
by simp
lemma state_none_2:
"(state_eq None) (make_full_observation e s r p t) \<Longrightarrow>
(state_eq None) (make_full_observation e s r p (stl t))"
by simp
lemma no_output_none_aux:
assumes "\<exists> p r i. j = (make_full_observation e None r []) i"
shows "alw (output_eq []) j"
using assms apply coinduct
apply simp
by fastforce
lemma no_output_none: "nxt (alw (output_eq [])) (make_full_observation e None r p t)"
using no_output_none_aux by auto
lemma nxt_alw: "nxt (alw P) s \<Longrightarrow> alw (nxt P) s"
by (simp add: alw_iff_sdrop)
lemma no_output_none_nxt: "alw (nxt (output_eq [])) (make_full_observation e None r p t)"
using nxt_alw no_output_none by blast
lemma no_output_none_if_empty: "alw (output_eq []) (make_full_observation e None r [] t)"
by (metis (mono_tags, lifting) alw_nxt make_full_observation.simps(1) no_output_none state.select_convs(4))
lemma no_updates_none_aux:
assumes "\<exists> p i. j = (make_full_observation e None r p) i"
shows "alw (\<lambda>x. datastate (shd x) = r) j"
using assms apply coinduct
by fastforce
lemma no_updates_none: "alw (\<lambda>x. datastate (shd x) = r) (make_full_observation e None r p t)"
using no_updates_none_aux by blast
lemma action_components: "(label_eq l aand input_eq i) s = (action (shd s) = (String.implode l, i))"
by (metis fst_conv prod.collapse snd_conv)
end
diff --git a/thys/Extended_Finite_State_Machines/FSet_Utils.thy b/thys/Extended_Finite_State_Machines/FSet_Utils.thy
--- a/thys/Extended_Finite_State_Machines/FSet_Utils.thy
+++ b/thys/Extended_Finite_State_Machines/FSet_Utils.thy
@@ -1,341 +1,338 @@
section\<open>FSet Utilities\<close>
text\<open>This theory provides various additional lemmas, definitions, and syntax over the fset data type.\<close>
theory FSet_Utils
imports "HOL-Library.FSet"
begin
notation (latex output)
"FSet.fempty" ("\<emptyset>") and
"FSet.fmember" ("\<in>")
syntax (ASCII)
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3ALL (_/:_)./ _)" [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3EX (_/:_)./ _)" [0, 0, 10] 10)
"_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3EX! (_/:_)./ _)" [0, 0, 10] 10)
syntax (input)
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3! (_/:_)./ _)" [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3? (_/:_)./ _)" [0, 0, 10] 10)
"_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3?! (_/:_)./ _)" [0, 0, 10] 10)
syntax
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<forall>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<exists>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
"_fBnex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<nexists>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
"_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<exists>!(_/|\<in>|_)./ _)" [0, 0, 10] 10)
translations
"\<forall>x|\<in>|A. P" \<rightleftharpoons> "CONST fBall A (\<lambda>x. P)"
"\<exists>x|\<in>|A. P" \<rightleftharpoons> "CONST fBex A (\<lambda>x. P)"
"\<nexists>x|\<in>|A. P" \<rightleftharpoons> "CONST fBall A (\<lambda>x. \<not>P)"
"\<exists>!x|\<in>|A. P" \<rightharpoonup> "\<exists>!x. x |\<in>| A \<and> P"
lemma fset_of_list_remdups [simp]: "fset_of_list (remdups l) = fset_of_list l"
apply (induct l)
apply simp
by (simp add: finsert_absorb fset_of_list_elem)
definition "fSum \<equiv> fsum (\<lambda>x. x)"
lemma fset_both_sides: "(Abs_fset s = f) = (fset (Abs_fset s) = fset f)"
by (simp add: fset_inject)
lemma Abs_ffilter: "(ffilter f s = s') = ({e \<in> (fset s). f e} = (fset s'))"
by (simp add: ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def)
lemma size_ffilter_card: "size (ffilter f s) = card ({e \<in> (fset s). f e})"
by (simp add: ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def)
lemma ffilter_empty [simp]: "ffilter f {||} = {||}"
by auto
lemma ffilter_finsert:
"ffilter f (finsert a s) = (if f a then finsert a (ffilter f s) else (ffilter f s))"
apply simp
apply standard
apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse)
apply auto[1]
apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse)
by auto
lemma fset_equiv: "(f1 = f2) = (fset f1 = fset f2)"
by (simp add: fset_inject)
lemma finsert_equiv: "(finsert e f = f') = (insert e (fset f) = (fset f'))"
by (simp add: finsert_def fset_both_sides Abs_fset_inverse)
lemma filter_elements:
"x |\<in>| Abs_fset (Set.filter f (fset s)) = (x \<in> (Set.filter f (fset s)))"
- by (metis ffilter.rep_eq fset_inverse notin_fset)
+ by (metis ffilter.rep_eq fset_inverse)
lemma sorted_list_of_fempty [simp]: "sorted_list_of_fset {||} = []"
by (simp add: sorted_list_of_fset_def)
-lemma fmember_implies_member: "e |\<in>| f \<Longrightarrow> e \<in> fset f"
- by (simp add: fmember_def)
-
lemma fold_union_ffUnion: "fold (|\<union>|) l {||} = ffUnion (fset_of_list l)"
by(induct l rule: rev_induct, auto)
lemma filter_filter:
"ffilter P (ffilter Q xs) = ffilter (\<lambda>x. Q x \<and> P x) xs"
by auto
lemma fsubset_strict:
"x2 |\<subset>| x1 \<Longrightarrow> \<exists>e. e |\<in>| x1 \<and> e |\<notin>| x2"
by auto
lemma fsubset:
"x2 |\<subset>| x1 \<Longrightarrow> \<nexists>e. e |\<in>| x2 \<and> e |\<notin>| x1"
by auto
lemma size_fsubset_elem:
assumes "\<exists>e. e |\<in>| x1 \<and> e |\<notin>| x2"
and "\<nexists>e. e |\<in>| x2 \<and> e |\<notin>| x1"
shows "size x2 < size x1"
using assms
- apply (simp add: fmember_def)
+ apply simp
by (metis card_seteq finite_fset linorder_not_le subsetI)
lemma size_fsubset: "x2 |\<subset>| x1 \<Longrightarrow> size x2 < size x1"
by (metis fsubset fsubset_strict size_fsubset_elem)
definition fremove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
where [code_abbrev]: "fremove x A = A - {|x|}"
lemma arg_cong_ffilter:
"\<forall>e |\<in>| f. p e = p' e \<Longrightarrow> ffilter p f = ffilter p' f"
by auto
lemma ffilter_singleton: "f e \<Longrightarrow> ffilter f {|e|} = {|e|}"
apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse)
by auto
lemma fset_eq_alt: "(x = y) = (x |\<subseteq>| y \<and> size x = size y)"
by (metis exists_least_iff le_less size_fsubset)
lemma ffold_empty [simp]: "ffold f b {||} = b"
by (simp add: ffold_def)
lemma sorted_list_of_fset_sort:
"sorted_list_of_fset (fset_of_list l) = sort (remdups l)"
by (simp add: fset_of_list.rep_eq sorted_list_of_fset.rep_eq sorted_list_of_set_sort_remdups)
lemma fMin_Min: "fMin (fset_of_list l) = Min (set l)"
by (simp add: fMin.F.rep_eq fset_of_list.rep_eq)
lemma sorted_hd_Min:
"sorted l \<Longrightarrow>
l \<noteq> [] \<Longrightarrow>
hd l = Min (set l)"
by (metis List.finite_set Min_eqI eq_iff hd_Cons_tl insertE list.set_sel(1) list.simps(15) sorted_simps(2))
lemma hd_sort_Min: "l \<noteq> [] \<Longrightarrow> hd (sort l) = Min (set l)"
by (metis sorted_hd_Min set_empty set_sort sorted_sort)
lemma hd_sort_remdups: "hd (sort (remdups l)) = hd (sort l)"
by (metis hd_sort_Min remdups_eq_nil_iff set_remdups)
lemma exists_fset_of_list: "\<exists>l. f = fset_of_list l"
using exists_fset_of_list by fastforce
lemma hd_sorted_list_of_fset:
"s \<noteq> {||} \<Longrightarrow> hd (sorted_list_of_fset s) = (fMin s)"
apply (insert exists_fset_of_list[of s])
apply (erule exE)
apply simp
apply (simp add: sorted_list_of_fset_sort fMin_Min hd_sort_remdups)
by (metis fset_of_list_simps(1) hd_sort_Min)
lemma fminus_filter_singleton:
"fset_of_list l |-| {|x|} = fset_of_list (filter (\<lambda>e. e \<noteq> x) l)"
by auto
lemma card_minus_fMin:
"s \<noteq> {||} \<Longrightarrow> card (fset s - {fMin s}) < card (fset s)"
by (metis Min_in bot_fset.rep_eq card_Diff1_less fMin.F.rep_eq finite_fset fset_equiv)
(* Provides a deterministic way to fold fsets similar to List.fold that works with the code generator *)
function ffold_ord :: "(('a::linorder) \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b \<Rightarrow> 'b" where
"ffold_ord f s b = (
if s = {||} then
b
else
let
h = fMin s;
t = s - {|h|}
in
ffold_ord f t (f h b)
)"
by auto
termination
apply (relation "measures [\<lambda>(a, s, ab). size s]")
apply simp
- by (simp add: card_minus_fMin)
+ by (simp add: card_gt_0_iff fset_equiv)
lemma sorted_list_of_fset_Cons:
"\<exists>h t. (sorted_list_of_fset (finsert s ss)) = h#t"
apply (simp add: sorted_list_of_fset_def)
by (cases "insort s (sorted_list_of_set (fset ss - {s}))", auto)
lemma list_eq_hd_tl:
"l \<noteq> [] \<Longrightarrow>
hd l = h \<Longrightarrow>
tl l = t \<Longrightarrow>
l = (h#t)"
by auto
lemma fset_of_list_sort: "fset_of_list l = fset_of_list (sort l)"
by (simp add: fset_of_list.abs_eq)
lemma exists_sorted_distinct_fset_of_list:
"\<exists>l. sorted l \<and> distinct l \<and> f = fset_of_list l"
by (metis distinct_sorted_list_of_set sorted_list_of_fset.rep_eq sorted_list_of_fset_simps(2) sorted_sorted_list_of_set)
lemma fset_of_list_empty [simp]: "(fset_of_list l = {||}) = (l = [])"
by (metis fset_of_list.rep_eq fset_of_list_simps(1) set_empty)
lemma ffold_ord_cons: assumes sorted: "sorted (h#t)"
and distinct: "distinct (h#t)"
shows "ffold_ord f (fset_of_list (h#t)) b = ffold_ord f (fset_of_list t) (f h b)"
proof-
have h_is_min: "h = fMin (fset_of_list (h#t))"
by (metis sorted fMin_Min list.sel(1) list.simps(3) sorted_hd_Min)
have remove_min: "fset_of_list t = (fset_of_list (h#t)) - {|h|}"
using distinct fset_of_list_elem by force
show ?thesis
apply (simp only: ffold_ord.simps[of f "fset_of_list (h#t)"])
by (metis h_is_min remove_min fset_of_list_empty list.distinct(1))
qed
lemma sorted_distinct_ffold_ord: assumes "sorted l"
and "distinct l"
shows "ffold_ord f (fset_of_list l) b = fold f l b"
using assms
apply (induct l arbitrary: b)
apply simp
by (metis distinct.simps(2) ffold_ord_cons fold_simps(2) sorted_simps(2))
lemma ffold_ord_fold_sorted: "ffold_ord f s b = fold f (sorted_list_of_fset s) b"
by (metis exists_sorted_distinct_fset_of_list sorted_distinct_ffold_ord distinct_remdups_id sorted_list_of_fset_sort sorted_sort_id)
context includes fset.lifting begin
lift_definition fprod :: "'a fset \<Rightarrow> 'b fset \<Rightarrow> ('a \<times> 'b) fset " (infixr "|\<times>|" 80) is "\<lambda>a b. fset a \<times> fset b"
by simp
lift_definition fis_singleton :: "'a fset \<Rightarrow> bool" is "\<lambda>A. is_singleton (fset A)".
end
lemma fprod_empty_l: "{||} |\<times>| a = {||}"
using bot_fset_def fprod.abs_eq by force
lemma fprod_empty_r: "a |\<times>| {||} = {||}"
by (simp add: fprod_def bot_fset_def Abs_fset_inverse)
lemmas fprod_empty = fprod_empty_l fprod_empty_r
lemma fprod_finsert: "(finsert a as) |\<times>| (finsert b bs) =
finsert (a, b) (fimage (\<lambda>b. (a, b)) bs |\<union>| fimage (\<lambda>a. (a, b)) as |\<union>| (as |\<times>| bs))"
apply (simp add: fprod_def fset_both_sides Abs_fset_inverse)
by auto
lemma fprod_member:
"x |\<in>| xs \<Longrightarrow>
y |\<in>| ys \<Longrightarrow>
(x, y) |\<in>| xs |\<times>| ys"
- by (simp add: fmember_def fprod_def Abs_fset_inverse)
+ by (simp add: fprod_def Abs_fset_inverse)
lemma fprod_subseteq:
"x |\<subseteq>| x' \<and> y |\<subseteq>| y' \<Longrightarrow> x |\<times>| y |\<subseteq>| x' |\<times>| y'"
apply (simp add: fprod_def less_eq_fset_def Abs_fset_inverse)
by auto
lemma fimage_fprod:
"(a, b) |\<in>| A |\<times>| B \<Longrightarrow> f a b |\<in>| (\<lambda>(x, y). f x y) |`| (A |\<times>| B)"
by force
lemma fprod_singletons: "{|a|} |\<times>| {|b|} = {|(a, b)|}"
apply (simp add: fprod_def)
by (metis fset_inverse fset_simps(1) fset_simps(2))
lemma fprod_equiv:
"(fset (f |\<times>| f') = s) = (((fset f) \<times> (fset f')) = s)"
by (simp add: fprod_def Abs_fset_inverse)
lemma fis_singleton_alt: "fis_singleton f = (\<exists>e. f = {|e|})"
by (metis fis_singleton.rep_eq fset_inverse fset_simps(1) fset_simps(2) is_singleton_def)
lemma singleton_singleton [simp]: "fis_singleton {|a|}"
by (simp add: fis_singleton_def)
lemma not_singleton_empty [simp]: "\<not> fis_singleton {||}"
apply (simp add: fis_singleton_def)
by (simp add: is_singleton_altdef)
lemma fis_singleton_fthe_elem:
"fis_singleton A \<longleftrightarrow> A = {|fthe_elem A|}"
by (metis fis_singleton_alt fthe_felem_eq)
lemma fBall_ffilter:
"\<forall>x |\<in>| X. f x \<Longrightarrow> ffilter f X = X"
by auto
lemma fBall_ffilter2:
"X = Y \<Longrightarrow>
\<forall>x |\<in>| X. f x \<Longrightarrow>
ffilter f X = Y"
by auto
lemma size_fset_of_list: "size (fset_of_list l) = length (remdups l)"
apply (induct l)
apply simp
by (simp add: fset_of_list.rep_eq insert_absorb)
lemma size_fsingleton: "(size f = 1) = (\<exists>e. f = {|e|})"
apply (insert exists_fset_of_list[of f])
apply clarify
apply (simp only: size_fset_of_list)
apply (simp add: fset_of_list_def fset_both_sides Abs_fset_inverse)
by (metis List.card_set One_nat_def card.insert card_1_singletonE card.empty empty_iff finite.intros(1))
lemma ffilter_mono: "(ffilter X xs = f) \<Longrightarrow> \<forall>x |\<in>| xs. X x = Y x \<Longrightarrow> (ffilter Y xs = f)"
by auto
lemma size_fimage: "size (fimage f s) \<le> size s"
apply (induct s)
apply simp
by (simp add: card_insert_if)
lemma size_ffilter: "size (ffilter P f) \<le> size f"
apply (induct f)
apply simp
apply (simp only: ffilter_finsert)
apply (case_tac "P x")
- apply (simp add: fmember_iff_member_fset)
+ apply simp
by (simp add: card_insert_if)
lemma fimage_size_le: "\<And>f s. size s \<le> n \<Longrightarrow> size (fimage f s) \<le> n"
using le_trans size_fimage by blast
lemma ffilter_size_le: "\<And>f s. size s \<le> n \<Longrightarrow> size (ffilter f s) \<le> n"
using dual_order.trans size_ffilter by blast
lemma set_membership_eq: "A = B \<longleftrightarrow> (\<lambda>x. Set.member x A) = (\<lambda>x. Set.member x B)"
apply standard
apply simp
by (meson equalityI subsetI)
lemmas ffilter_eq_iff = Abs_ffilter set_membership_eq fun_eq_iff
lemma size_le_1: "size f \<le> 1 = (f = {||} \<or> (\<exists>e. f = {|e|}))"
apply standard
apply (metis bot.not_eq_extremum gr_implies_not0 le_neq_implies_less less_one size_fsingleton size_fsubset)
by auto
lemma size_gt_1: "1 < size f \<Longrightarrow> \<exists>e1 e2 f'. e1 \<noteq> e2 \<and> f = finsert e1 (finsert e2 f')"
apply (induct f)
apply simp
apply (rule_tac x=x in exI)
by (metis finsertCI leD not_le_imp_less size_le_1)
end
diff --git a/thys/Extended_Finite_State_Machines/examples/Drinks_Machine.thy b/thys/Extended_Finite_State_Machines/examples/Drinks_Machine.thy
--- a/thys/Extended_Finite_State_Machines/examples/Drinks_Machine.thy
+++ b/thys/Extended_Finite_State_Machines/examples/Drinks_Machine.thy
@@ -1,282 +1,282 @@
chapter\<open>Examples\<close>
text\<open>In this chapter, we provide some examples of EFSMs and proofs over them. We first present a
formalisation of a simple drinks machine. Next, we prove observational equivalence of an alternative
model. Finally, we prove some temporal properties of the first example.\<close>
section\<open>Drinks Machine\<close>
text\<open>This theory formalises a simple drinks machine. The \emph{select} operation takes one
argument - the desired beverage. The \emph{coin} operation also takes one parameter representing
the value of the coin. The \emph{vend} operation has two flavours - one which dispenses the drink if
the customer has inserted enough money, and one which dispenses nothing if the user has not inserted
sufficient funds.
We first define a datatype \emph{statemane} which corresponds to $S$ in the formal definition.
Note that, while statename has four elements, the drinks machine presented here only requires three
states. The fourth element is included here so that the \emph{statename} datatype may be used in
the next example.\<close>
theory Drinks_Machine
imports "Extended_Finite_State_Machines.EFSM"
begin
text_raw\<open>\snip{selectdef}{1}{2}{%\<close>
definition select :: "transition" where
"select \<equiv> \<lparr>
Label = STR ''select'',
Arity = 1,
Guards = [],
Outputs = [],
Updates = [
(1, V (I 0)),
(2, L (Num 0))
]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{coindef}{1}{2}{%\<close>
definition coin :: "transition" where
"coin \<equiv> \<lparr>
Label = STR ''coin'',
Arity = 1,
Guards = [],
Outputs = [Plus (V (R 2)) (V (I 0))],
Updates = [
(1, V (R 1)),
(2, Plus (V (R 2)) (V (I 0)))
]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{venddef}{1}{2}{%\<close>
definition vend:: "transition" where
"vend\<equiv> \<lparr>
Label = STR ''vend'',
Arity = 0,
Guards = [(Ge (V (R 2)) (L (Num 100)))],
Outputs = [(V (R 1))],
Updates = [(1, V (R 1)), (2, V (R 2))]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{vendfaildef}{1}{2}{%\<close>
definition vend_fail :: "transition" where
"vend_fail \<equiv> \<lparr>
Label = STR ''vend'',
Arity = 0,
Guards = [(Lt (V (R 2)) (L (Num 100)))],
Outputs = [],
Updates = [(1, V (R 1)), (2, V (R 2))]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{drinksdef}{1}{2}{%\<close>
definition drinks :: "transition_matrix" where
"drinks \<equiv> {|
((0,1), select),
((1,1), coin),
((1,1), vend_fail),
((1,2), vend)
|}"
text_raw\<open>}%endsnip\<close>
lemmas transitions = select_def coin_def vend_def vend_fail_def
lemma apply_updates_vend: "apply_updates (Updates vend) (join_ir [] r) r = r"
by (simp add: vend_def apply_updates_def)
lemma drinks_states: "S drinks = {|0, 1, 2|}"
apply (simp add: S_def drinks_def)
by auto
lemma possible_steps_0:
"length i = 1 \<Longrightarrow>
possible_steps drinks 0 r (STR ''select'') i = {|(1, select)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions apply_guards_def)
lemma first_step_select:
"(s', t) |\<in>| possible_steps drinks 0 r aa b \<Longrightarrow> s' = 1 \<and> t = select"
- apply (simp add: possible_steps_def fimage_def ffilter_def fmember_def Abs_fset_inverse Set.filter_def drinks_def)
+ apply (simp add: possible_steps_def fimage_def ffilter_def Abs_fset_inverse Set.filter_def drinks_def)
apply safe
by (simp_all add: transitions)
lemma drinks_vend_insufficient:
"r $ 2 = Some (Num x1) \<Longrightarrow>
x1 < 100 \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {|(1, vend_fail)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions apply_guards_def value_gt_def join_ir_def connectives)
lemma drinks_vend_invalid:
"\<nexists>n. r $ 2 = Some (Num n) \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {||}"
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (simp add: MaybeBoolInt_not_num_1 value_gt_def)
lemma possible_steps_1_coin:
"length i = 1 \<Longrightarrow> possible_steps drinks 1 r (STR ''coin'') i = {|(1, coin)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions)
lemma possible_steps_2_vend:
"\<exists>n. r $ 2 = Some (Num n) \<and> n \<ge> 100 \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {|(2, vend)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions apply_guards_def value_gt_def join_ir_def connectives)
lemma recognises_from_2:
"recognises_execution drinks 1 <1 $:= d, 2 $:= Some (Num 100)> [(STR ''vend'', [])]"
apply (rule recognises_execution.step)
apply (rule_tac x="(2, vend)" in fBexI)
apply simp
by (simp add: possible_steps_2_vend)
lemma recognises_from_1a:
"recognises_execution drinks 1 <1 $:= d, 2 $:= Some (Num 50)> [(STR ''coin'', [Num 50]), (STR ''vend'', [])]"
apply (rule recognises_execution.step)
apply (rule_tac x="(1, coin)" in fBexI)
apply (simp add: apply_updates_def coin_def finfun_update_twist value_plus_def recognises_from_2)
by (simp add: possible_steps_1_coin)
lemma recognises_from_1: "recognises_execution drinks 1 <2 $:= Some (Num 0), 1 $:= Some d>
[(STR ''coin'', [Num 50]), (STR ''coin'', [Num 50]), (STR ''vend'', [])]"
apply (rule recognises_execution.step)
apply (rule_tac x="(1, coin)" in fBexI)
apply (simp add: apply_updates_def coin_def value_plus_def finfun_update_twist recognises_from_1a)
by (simp add: possible_steps_1_coin)
lemma purchase_coke:
"observe_execution drinks 0 <> [(STR ''select'', [Str ''coke'']), (STR ''coin'', [Num 50]), (STR ''coin'', [Num 50]), (STR ''vend'', [])] =
[[], [Some (Num 50)], [Some (Num 100)], [Some (Str ''coke'')]]"
by (simp add: possible_steps_0 possible_steps_1_coin possible_steps_2_vend transitions
apply_outputs_def apply_updates_def value_plus_def)
lemma rejects_input:
"l \<noteq> STR ''coin'' \<Longrightarrow>
l \<noteq> STR ''vend'' \<Longrightarrow>
\<not> recognises_execution drinks 1 d' [(l, i)]"
apply (rule no_possible_steps_rejects)
by (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
lemma rejects_recognises_prefix: "l \<noteq> STR ''coin'' \<Longrightarrow>
l \<noteq> STR ''vend'' \<Longrightarrow>
\<not> (recognises drinks [(STR ''select'', [Str ''coke'']), (l, i)])"
apply (rule trace_reject_later)
apply (simp add: possible_steps_0 select_def join_ir_def input2state_def)
using rejects_input by blast
lemma rejects_termination:
"observe_execution drinks 0 <> [(STR ''select'', [Str ''coke'']), (STR ''rejects'', [Num 50]), (STR ''coin'', [Num 50])] = [[]]"
apply (rule observe_execution_step)
apply (simp add: step_def possible_steps_0 select_def)
apply (rule observe_execution_no_possible_step)
by (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
(* Part of Example 2 in Foster et. al. *)
lemma r2_0_vend:
"can_take_transition vend i r \<Longrightarrow>
\<exists>n. r $ 2 = Some (Num n) \<and> n \<ge> 100" (* You can't take vendimmediately after taking select *)
apply (simp add: can_take_transition_def can_take_def vend_def apply_guards_def maybe_negate_true maybe_or_false connectives value_gt_def)
using MaybeBoolInt.elims by force
lemma drinks_vend_sufficient: "r $ 2 = Some (Num x1) \<Longrightarrow>
x1 \<ge> 100 \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {|(2, vend)|}"
using possible_steps_2_vend by blast
lemma drinks_end: "possible_steps drinks 2 r a b = {||}"
apply (simp add: possible_steps_def drinks_def transitions)
by force
lemma drinks_vend_r2_String:
"r $ 2 = Some (value.Str x2) \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {||}"
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (simp add: value_gt_def)
lemma drinks_vend_r2_rejects:
"\<nexists>n. r $ 2 = Some (Num n) \<Longrightarrow> step drinks 1 r (STR ''vend'') [] = None"
apply (rule no_possible_steps_1)
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (simp add: MaybeBoolInt_not_num_1 value_gt_def)
lemma drinks_0_rejects:
"\<not> (fst a = STR ''select'' \<and> length (snd a) = 1) \<Longrightarrow>
(possible_steps drinks 0 r (fst a) (snd a)) = {||}"
apply (simp add: drinks_def possible_steps_def transitions)
by force
lemma drinks_vend_empty: "(possible_steps drinks 0 <> (STR ''vend'') []) = {||}"
using drinks_0_rejects
by auto
lemma drinks_1_rejects:
"fst a = STR ''coin'' \<longrightarrow> length (snd a) \<noteq> 1 \<Longrightarrow>
a \<noteq> (STR ''vend'', []) \<Longrightarrow>
possible_steps drinks 1 r (fst a) (snd a) = {||}"
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (metis prod.collapse)
lemma drinks_rejects_future: "\<not> recognises_execution drinks 2 d ((l, i)#t)"
apply (rule no_possible_steps_rejects)
by (simp add: possible_steps_empty drinks_def)
lemma drinks_1_rejects_trace:
assumes not_vend: "e \<noteq> (STR ''vend'', [])"
and not_coin: "\<nexists>i. e = (STR ''coin'', [i])"
shows "\<not> recognises_execution drinks 1 r (e # es)"
proof-
show ?thesis
apply (cases e, simp)
subgoal for a b
apply (rule no_possible_steps_rejects)
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
apply (case_tac b)
using not_vend apply simp
using not_coin by auto
done
qed
lemma rejects_state_step: "s > 1 \<Longrightarrow> step drinks s r l i = None"
apply (rule no_possible_steps_1)
by (simp add: possible_steps_empty drinks_def)
lemma invalid_other_states:
"s > 1 \<Longrightarrow> \<not> recognises_execution drinks s r ((aa, b) # t)"
apply (rule no_possible_steps_rejects)
by (simp add: possible_steps_empty drinks_def)
lemma vend_ge_100:
"possible_steps drinks 1 r l i = {|(2, vend)|} \<Longrightarrow>
\<not>? value_gt (Some (Num 100)) (r $ 2) = trilean.true"
apply (insert possible_steps_apply_guards[of drinks 1 r l i 2 vend])
by (simp add: possible_steps_def apply_guards_def vend_def)
lemma drinks_no_possible_steps_1:
assumes not_coin: "\<not> (a = STR ''coin'' \<and> length b = 1)"
and not_vend: "\<not> (a = STR ''vend'' \<and> b = [])"
shows "possible_steps drinks 1 r a b = {||}"
using drinks_1_rejects not_coin not_vend by auto
lemma possible_steps_0_not_select: "a \<noteq> STR ''select'' \<Longrightarrow>
possible_steps drinks 0 <> a b = {||}"
apply (simp add: possible_steps_def ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def drinks_def)
apply safe
by (simp_all add: select_def)
lemma possible_steps_select_wrong_arity: "a = STR ''select'' \<Longrightarrow>
length b \<noteq> 1 \<Longrightarrow>
possible_steps drinks 0 <> a b = {||}"
apply (simp add: possible_steps_def ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def drinks_def)
apply safe
by (simp_all add: select_def)
lemma possible_steps_0_invalid:
"\<not> (l = STR ''select'' \<and> length i = 1) \<Longrightarrow>
possible_steps drinks 0 <> l i = {||}"
using possible_steps_0_not_select possible_steps_select_wrong_arity by fastforce
end
diff --git a/thys/FO_Theory_Rewriting/Closure/GTT_RRn.thy b/thys/FO_Theory_Rewriting/Closure/GTT_RRn.thy
--- a/thys/FO_Theory_Rewriting/Closure/GTT_RRn.thy
+++ b/thys/FO_Theory_Rewriting/Closure/GTT_RRn.thy
@@ -1,157 +1,157 @@
theory GTT_RRn
imports Regular_Tree_Relations.GTT
TA_Clousure_Const
Context_RR2
Lift_Root_Step
begin
section \<open>Connecting regular tree languages to set/relation specifications\<close>
abbreviation ggtt_lang where
"ggtt_lang F G \<equiv> map_both gterm_of_term ` (Restr (gtt_lang_terms G) {t. funas_term t \<subseteq> fset F})"
lemma ground_mctxt_map_vars_mctxt [simp]:
"ground_mctxt (map_vars_mctxt f C) = ground_mctxt C"
by (induct C) auto
lemma root_single_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec \<A> (lift_root_step \<F> PRoot ESingle R)"
using assms unfolding RR2_spec_def
by (auto simp: lift_root_step.simps)
lemma root_strictparallel_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec \<A> (lift_root_step \<F> PRoot EStrictParallel R)"
using assms unfolding RR2_spec_def
by (auto simp: lift_root_step.simps)
lemma reflcl_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec (reflcl_reg (lift_sig_RR2 |`| \<F>) \<A>) (lift_root_step (fset \<F>) PRoot EParallel R)"
unfolding RR2_spec_def \<L>_reflcl_reg
unfolding lift_root_step.simps \<T>\<^sub>G_equivalent_def assms[unfolded RR2_spec_def]
by (auto simp flip: \<T>\<^sub>G_equivalent_def)
lemma parallel_closure_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec (parallel_closure_reg (lift_sig_RR2 |`| \<F>) \<A>) (lift_root_step (fset \<F>) PAny EParallel R)"
unfolding RR2_spec_def parallelcl_gmctxt_lang lift_root_step.simps
unfolding gmctxtex_onp_gpair_set_conv assms[unfolded RR2_spec_def]
by (intro RR2_gmctxt_cl_to_gmctxt) auto
lemma ctxt_closure_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec (ctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<A>) (lift_root_step (fset \<F>) PAny ESingle R)"
unfolding RR2_spec_def gctxt_closure_lang lift_root_step.simps
unfolding gctxtex_onp_gpair_set_conv assms[unfolded RR2_spec_def]
by (intro RR2_gctxt_cl_to_gctxt) auto
lemma mctxt_closure_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec (mctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<A>) (lift_root_step (fset \<F>) PAny EStrictParallel R)"
unfolding RR2_spec_def gmctxt_closure_lang lift_root_step.simps
unfolding gmctxtex_onp_gpair_set_conv assms[unfolded RR2_spec_def] conj_assoc
by (intro RR2_gmctxt_cl_to_gmctxt[where ?P = "\<lambda> C. 0 < num_gholes C \<and> funas_gmctxt C \<subseteq> fset (lift_sig_RR2 |`| \<F>)" and
?R = "\<lambda> C. 0 < num_gholes C \<and> funas_gmctxt C \<subseteq> fset \<F>", unfolded conj_assoc]) auto
lemma nhole_ctxt_closure_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec (nhole_ctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<A>) (lift_root_step (fset \<F>) PNonRoot ESingle R)"
unfolding RR2_spec_def nhole_ctxtcl_lang lift_root_step.simps
unfolding gctxtex_onp_gpair_set_conv assms[unfolded RR2_spec_def]
by (intro RR2_gctxt_cl_to_gctxt[where
?P = "\<lambda> C. C \<noteq> \<box>\<^sub>G \<and> funas_gctxt C \<subseteq> fset (lift_sig_RR2 |`| \<F>)", unfolded conj_assoc]) auto
lemma nhole_mctxt_closure_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec (nhole_mctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<A>) (lift_root_step (fset \<F>) PNonRoot EStrictParallel R)"
unfolding RR2_spec_def nhole_gmctxt_closure_lang lift_root_step.simps
unfolding gmctxtex_onp_gpair_set_conv assms[unfolded RR2_spec_def]
by (intro RR2_gmctxt_cl_to_gmctxt[where
?P = "\<lambda> C. 0 < num_gholes C \<and> C \<noteq> GMHole \<and> funas_gmctxt C \<subseteq> fset (lift_sig_RR2 |`| \<F>)", unfolded conj_assoc])
auto
lemma nhole_mctxt_reflcl_automaton:
assumes "RR2_spec \<A> R"
shows "RR2_spec (nhole_mctxt_reflcl_reg (lift_sig_RR2 |`| \<F>) \<A>) (lift_root_step (fset \<F>) PNonRoot EParallel R)"
using nhole_mctxt_closure_automaton[OF assms, of \<F>]
unfolding RR2_spec_def lift_root_step_Parallel_conv nhole_mctxt_reflcl_lang
by (auto simp flip: \<T>\<^sub>G_equivalent_def)
definition GTT_to_RR2_root :: "('q, 'f) gtt \<Rightarrow> (_, 'f option \<times> 'f option) ta" where
"GTT_to_RR2_root \<G> = pair_automaton (fst \<G>) (snd \<G>)"
definition GTT_to_RR2_root_reg where
"GTT_to_RR2_root_reg \<G> = Reg (map_both Some |`| fId_on (gtt_states \<G>)) (GTT_to_RR2_root \<G>)"
lemma GTT_to_RR2_root:
"RR2_spec (GTT_to_RR2_root_reg \<G>) (agtt_lang \<G>)"
proof -
{ fix s assume "s \<in> \<L> (GTT_to_RR2_root_reg \<G>)"
then obtain q where q: "q |\<in>| fin (GTT_to_RR2_root_reg \<G>)" "q |\<in>| ta_der (GTT_to_RR2_root \<G>) (term_of_gterm s)"
by (auto simp: \<L>_def gta_lang_def GTT_to_RR2_root_reg_def gta_der_def)
then obtain q' where [simp]: "q = (Some q', Some q')" using q(1) by (auto simp: GTT_to_RR2_root_reg_def)
have "\<exists>t u q. s = gpair t u \<and> q |\<in>| ta_der (fst \<G>) (term_of_gterm t) \<and> q |\<in>| ta_der (snd \<G>) (term_of_gterm u)"
using fsubsetD[OF ta_der_mono' q(2), of "pair_automaton (fst \<G>) (snd \<G>)"]
by (auto simp: GTT_to_RR2_root_def dest!: from_ta_der_pair_automaton(4))
} moreover
{ fix t u q assume q: "q |\<in>| ta_der (fst \<G>) (term_of_gterm t)" "q |\<in>| ta_der (snd \<G>) (term_of_gterm u)"
have "lift_fun q |\<in>| map_both Some |`| fId_on (\<Q> (fst \<G>) |\<union>| \<Q> (snd \<G>))"
using q[THEN fsubsetD[OF ground_ta_der_states[OF ground_term_of_gterm]]]
by (auto simp: fimage_iff fBex_def)
then have "gpair t u \<in> \<L> (GTT_to_RR2_root_reg \<G>)" using q
using fsubsetD[OF ta_der_mono to_ta_der_pair_automaton(3)[OF q], of "GTT_to_RR2_root \<G>"]
by (auto simp: \<L>_def GTT_to_RR2_root_def gta_lang_def image_def gtt_states_def
gta_der_def GTT_to_RR2_root_reg_def)
} ultimately show ?thesis by (auto simp: RR2_spec_def agtt_lang_def \<L>_def gta_der_def)
qed
lemma swap_GTT_to_RR2_root:
"gpair s t \<in> \<L> (GTT_to_RR2_root_reg (prod.swap \<G>)) \<longleftrightarrow>
gpair t s \<in> \<L> (GTT_to_RR2_root_reg \<G>)"
by (auto simp: GTT_to_RR2_root[unfolded RR2_spec_def] agtt_lang_def)
lemma funas_mctxt_map_vars_mctxt [simp]:
"funas_mctxt (map_vars_mctxt f C) = funas_mctxt C"
by (induct C) auto
definition GTT_to_RR2_reg :: "('f \<times> nat) fset \<Rightarrow> ('q, 'f) gtt \<Rightarrow> (_, 'f option \<times> 'f option) reg" where
"GTT_to_RR2_reg F G = parallel_closure_reg (lift_sig_RR2 |`| F) (GTT_to_RR2_root_reg G)"
lemma agtt_lang_syms:
"gtt_syms \<G> |\<subseteq>| \<F> \<Longrightarrow> agtt_lang \<G> \<subseteq> {t. funas_gterm t \<subseteq> fset \<F>} \<times> {t. funas_gterm t \<subseteq> fset \<F>}"
by (auto simp: agtt_lang_def gta_der_def funas_term_of_gterm_conv)
- (metis ffunas_gterm.rep_eq fin_mono notin_fset ta_der_gterm_sig)+
+ (metis ffunas_gterm.rep_eq fin_mono ta_der_gterm_sig)+
lemma gtt_lang_from_agtt_lang:
"gtt_lang \<G> = lift_root_step UNIV PAny EParallel (agtt_lang \<G>)"
unfolding lift_root_step.simps agtt_lang_def
by (auto simp: lift_root_step.simps agtt_lang_def gmctxt_cl_gmctxtex_onp_conv)
lemma GTT_to_RR2:
assumes "gtt_syms \<G> |\<subseteq>| \<F>"
shows "RR2_spec (GTT_to_RR2_reg \<F> \<G>) (ggtt_lang \<F> \<G>)"
proof -
have *: "snd ` (X \<times> X) = X" for X by auto
show ?thesis unfolding gtt_lang_from_agtt_lang GTT_to_RR2_reg_def RR2_spec_def
parallel_closure_automaton[OF GTT_to_RR2_root, of \<F> \<G>, unfolded RR2_spec_def]
proof (intro arg_cong[where f = "\<lambda>X. {gpair t u |t u. (t,u) \<in> X}"] equalityI subrelI, goal_cases)
case (1 s t) then show ?case
using subsetD[OF equalityD2[OF gtt_lang_from_agtt_lang], of "(s, t)" \<G>]
by (intro rev_image_eqI[of "(term_of_gterm s, term_of_gterm t)"])
(auto simp: funas_term_of_gterm_conv subsetD[OF lift_root_step_mono]
dest: subsetD[OF lift_root_step_sig[unfolded \<T>\<^sub>G_equivalent_def, OF agtt_lang_syms[OF assms]]])
next
case (2 s t)
from image_mono[OF agtt_lang_syms[OF assms], of snd, unfolded *]
have *: "snd ` agtt_lang \<G> \<subseteq> gterms UNIV" by auto
show ?case using 2
by (auto intro!: lift_root_step_sig_transfer[unfolded \<T>\<^sub>G_equivalent_def, OF _ *, of _ _ _ "fset \<F>"]
simp: funas_gterm_gterm_of_term funas_term_of_gterm_conv)
qed
qed
end
diff --git a/thys/FO_Theory_Rewriting/Closure/TA_Clousure_Const.thy b/thys/FO_Theory_Rewriting/Closure/TA_Clousure_Const.thy
--- a/thys/FO_Theory_Rewriting/Closure/TA_Clousure_Const.thy
+++ b/thys/FO_Theory_Rewriting/Closure/TA_Clousure_Const.thy
@@ -1,1019 +1,1038 @@
section \<open>(Multihole)Context closure of recognized tree languages\<close>
theory TA_Clousure_Const
imports Tree_Automata_Derivation_Split
begin
subsection \<open>Tree Automata closure constructions\<close>
declare ta_union_def [simp]
subsubsection \<open>Reflexive closure over a given signature\<close>
definition "reflcl_rules \<F> q \<equiv> (\<lambda> (f, n). TA_rule f (replicate n q) q) |`| \<F>"
definition "refl_ta \<F> q = TA (reflcl_rules \<F> q) {||}"
definition gen_reflcl_automaton :: "('f \<times> nat) fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> 'q \<Rightarrow> ('q, 'f) ta" where
"gen_reflcl_automaton \<F> \<A> q = ta_union \<A> (refl_ta \<F> q)"
definition "reflcl_automaton \<F> \<A> = (let \<B> = fmap_states_ta Some \<A> in
gen_reflcl_automaton \<F> \<B> None)"
definition "reflcl_reg \<F> \<A> = Reg (finsert None (Some |`| fin \<A>)) (reflcl_automaton \<F> (ta \<A>))"
subsubsection \<open>Multihole context closure over a given signature\<close>
definition "refl_over_states_ta Q \<F> \<A> q = TA (reflcl_rules \<F> q) ((\<lambda> p. (p, q)) |`| (Q |\<inter>| \<Q> \<A>))"
definition gen_parallel_closure_automaton :: "'q fset \<Rightarrow> ('f \<times> nat) fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> 'q \<Rightarrow> ('q, 'f) ta" where
"gen_parallel_closure_automaton Q \<F> \<A> q = ta_union \<A> (refl_over_states_ta Q \<F> \<A> q)"
definition parallel_closure_reg where
"parallel_closure_reg \<F> \<A> = (let \<B> = fmap_states_reg Some \<A> in
Reg {|None|} (gen_parallel_closure_automaton (fin \<B>) \<F> (ta \<B>) None))"
subsubsection \<open>Context closure of regular tree language\<close>
definition "semantic_path_rules \<F> q\<^sub>c q\<^sub>i q\<^sub>f \<equiv>
|\<Union>| ((\<lambda> (f, n). fset_of_list (map (\<lambda> i. TA_rule f ((replicate n q\<^sub>c)[i := q\<^sub>i]) q\<^sub>f) [0..< n])) |`| \<F>)"
definition "reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f \<equiv>
TA (reflcl_rules \<F> q\<^sub>c |\<union>| semantic_path_rules \<F> q\<^sub>c q\<^sub>f q\<^sub>f) ((\<lambda> p. (p, q\<^sub>f)) |`| Q)"
definition "gen_ctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>f = ta_union \<A> (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f)"
definition "gen_ctxt_closure_reg \<F> \<A> q\<^sub>c q\<^sub>f =
Reg {|q\<^sub>f|} (gen_ctxt_closure_automaton (fin \<A>) \<F> (ta \<A>) q\<^sub>c q\<^sub>f)"
definition "ctxt_closure_reg \<F> \<A> =
(let \<B> = fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>) in
gen_ctxt_closure_reg \<F> \<B> (Inr False) (Inr True))"
subsubsection \<open>Not empty context closure of regular tree language\<close>
datatype cl_states = cl_state | tr_state | fin_state | fin_clstate
definition "reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f \<equiv>
TA (reflcl_rules \<F> q\<^sub>c |\<union>| semantic_path_rules \<F> q\<^sub>c q\<^sub>i q\<^sub>f |\<union>| semantic_path_rules \<F> q\<^sub>c q\<^sub>f q\<^sub>f) ((\<lambda> p. (p, q\<^sub>i)) |`| Q)"
definition "gen_nhole_ctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f =
ta_union \<A> (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f)"
definition "gen_nhole_ctxt_closure_reg \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f =
Reg {|q\<^sub>f|} (gen_nhole_ctxt_closure_automaton (fin \<A>) \<F> (ta \<A>) q\<^sub>c q\<^sub>i q\<^sub>f)"
definition "nhole_ctxt_closure_reg \<F> \<A> =
(let \<B> = fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>) in
(gen_nhole_ctxt_closure_reg \<F> \<B> (Inr cl_state) (Inr tr_state) (Inr fin_state)))"
subsubsection \<open>Non empty multihole context closure of regular tree language\<close>
abbreviation "add_eps \<A> e \<equiv> TA (rules \<A>) (eps \<A> |\<union>| e)"
definition "reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f \<equiv>
add_eps (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) {|(q\<^sub>i, q\<^sub>c)|}"
definition "gen_nhole_mctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f =
ta_union \<A> (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f)"
definition "gen_nhole_mctxt_closure_reg \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f =
Reg {|q\<^sub>f|} (gen_nhole_mctxt_closure_automaton (fin \<A>) \<F> (ta \<A>) q\<^sub>c q\<^sub>i q\<^sub>f)"
definition "nhole_mctxt_closure_reg \<F> \<A> =
(let \<B> = fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>) in
(gen_nhole_mctxt_closure_reg \<F> \<B> (Inr cl_state) (Inr tr_state) (Inr fin_state)))"
subsubsection \<open>Not empty multihole context closure of regular tree language\<close>
definition "gen_mctxt_closure_reg \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f =
Reg {|q\<^sub>f, q\<^sub>i|} (gen_nhole_mctxt_closure_automaton (fin \<A>) \<F> (ta \<A>) q\<^sub>c q\<^sub>i q\<^sub>f)"
definition "mctxt_closure_reg \<F> \<A> =
(let \<B> = fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>) in
(gen_mctxt_closure_reg \<F> \<B> (Inr cl_state) (Inr tr_state) (Inr fin_state)))"
subsubsection \<open>Multihole context closure of regular tree language\<close>
definition "nhole_mctxt_reflcl_reg \<F> \<A> =
reg_union (nhole_mctxt_closure_reg \<F> \<A>) (Reg {|fin_clstate|} (refl_ta \<F> (fin_clstate)))"
subsubsection \<open>Lemmas about @{const ta_der'}\<close>
lemma ta_det'_ground_id:
"t |\<in>| ta_der' \<A> s \<Longrightarrow> ground t \<Longrightarrow> t = s"
by (induct s arbitrary: t) (auto simp add: ta_der'.simps nth_equalityI)
lemma ta_det'_vars_term_id:
"t |\<in>| ta_der' \<A> s \<Longrightarrow> vars_term t \<inter> fset (\<Q> \<A>) = {} \<Longrightarrow> t = s"
proof (induct s arbitrary: t)
case (Fun f ss)
from Fun(2-) obtain ts where [simp]: "t = Fun f ts" and len: "length ts = length ss"
- by (cases t) (auto simp flip: fmember_iff_member_fset dest: rule_statesD eps_dest_all)
+ by (cases t) (auto dest: rule_statesD eps_dest_all)
from Fun(1)[OF nth_mem, of i "ts ! i" for i] show ?case using Fun(2-) len
- by (auto simp add: ta_der'.simps Union_disjoint simp flip: fmember_iff_member_fset
+ by (auto simp add: ta_der'.simps Union_disjoint
dest: rule_statesD eps_dest_all intro!: nth_equalityI)
-qed (auto simp add: ta_der'.simps simp flip: fmember_iff_member_fset dest: rule_statesD eps_dest_all)
+qed (auto simp add: ta_der'.simps dest: rule_statesD eps_dest_all)
lemma fresh_states_ta_der'_pres:
assumes st: "q \<in> vars_term s" "q |\<notin>| \<Q> \<A>"
and reach: "t |\<in>| ta_der' \<A> s"
shows "q \<in> vars_term t" using reach st(1)
proof (induct s arbitrary: t)
case (Var x)
then show ?case using assms(2)
by (cases t) (auto simp: ta_der'.simps dest: eps_trancl_statesD)
next
case (Fun f ss)
from Fun(3) obtain i where w: "i < length ss" "q \<in> vars_term (ss ! i)" by (auto simp: in_set_conv_nth)
have "i < length (args t) \<and> q \<in> vars_term (args t ! i)" using Fun(2) w assms(2) Fun(1)[OF nth_mem[OF w(1)] _ w(2)]
using rule_statesD(3) ta_der_to_ta_der'
by (auto simp: ta_der'.simps dest: rule_statesD(3)) fastforce+
then show ?case by (cases t) auto
qed
lemma ta_der'_states:
"t |\<in>| ta_der' \<A> s \<Longrightarrow> vars_term t \<subseteq> vars_term s \<union> fset (\<Q> \<A>)"
proof (induct s arbitrary: t)
case (Var x) then show ?case
- by (auto simp: ta_der'.simps simp flip: fmember_iff_member_fset dest: eps_dest_all)
+ by (auto simp: ta_der'.simps dest: eps_dest_all)
next
case (Fun f ts) then show ?case
- by (auto simp: ta_der'.simps rule_statesD simp flip: fmember_iff_member_fset dest: eps_dest_all)
- (metis (no_types, opaque_lifting) Un_iff in_set_conv_nth notin_fset subsetD)
+ by (auto simp: ta_der'.simps rule_statesD dest: eps_dest_all)
+ (metis (no_types, opaque_lifting) Un_iff in_set_conv_nth subsetD)
qed
lemma ta_der'_gterm_states:
"t |\<in>| ta_der' \<A> (term_of_gterm s) \<Longrightarrow> vars_term t \<subseteq> fset (\<Q> \<A>)"
using ta_der'_states[of t \<A> "term_of_gterm s"]
by auto
lemma ta_der'_Var_funas:
"Var q |\<in>| ta_der' \<A> s \<Longrightarrow> funas_term s \<subseteq> fset (ta_sig \<A>)"
by (auto simp: less_eq_fset.rep_eq ffunas_term.rep_eq dest!: ta_der_term_sig ta_der'_to_ta_der)
lemma ta_sig_fsubsetI:
assumes "\<And> r. r |\<in>| rules \<A> \<Longrightarrow> (r_root r, length (r_lhs_states r)) |\<in>| \<F>"
shows "ta_sig \<A> |\<subseteq>| \<F>" using assms
by (auto simp: ta_sig_def)
subsubsection \<open>Signature induced by @{const refl_ta} and @{const refl_over_states_ta}\<close>
lemma refl_ta_sig [simp]:
"ta_sig (refl_ta \<F> q) = \<F>"
"ta_sig (refl_over_states_ta Q \<F> \<A> q ) = \<F>"
- by (auto simp: ta_sig_def refl_ta_def reflcl_rules_def refl_over_states_ta_def fimage_iff fBex_def)
+ by (auto simp: ta_sig_def refl_ta_def reflcl_rules_def refl_over_states_ta_def image_iff Bex_def)
subsubsection \<open>Correctness of @{const refl_ta}, @{const gen_reflcl_automaton}, and @{const reflcl_automaton}\<close>
lemma refl_ta_eps [simp]: "eps (refl_ta \<F> q) = {||}"
by (auto simp: refl_ta_def)
lemma refl_ta_sound:
"s \<in> \<T>\<^sub>G (fset \<F>) \<Longrightarrow> q |\<in>| ta_der (refl_ta \<F> q) (term_of_gterm s)"
by (induct rule: \<T>\<^sub>G.induct) (auto simp: refl_ta_def reflcl_rules_def
- fimage_iff fBex_def simp flip: fmember_iff_member_fset)
+ image_iff Bex_def)
lemma reflcl_rules_args:
"length ps = n \<Longrightarrow> f ps \<rightarrow> p |\<in>| reflcl_rules \<F> q \<Longrightarrow> ps = replicate n q"
by (auto simp: reflcl_rules_def)
lemma \<Q>_refl_ta:
"\<Q> (refl_ta \<F> q) |\<subseteq>| {|q|}"
by (auto simp: \<Q>_def refl_ta_def rule_states_def reflcl_rules_def fset_of_list_elem)
lemma refl_ta_complete1:
"Var p |\<in>| ta_der' (refl_ta \<F> q) s \<Longrightarrow> p \<noteq> q \<Longrightarrow> s = Var p"
by (cases s) (auto simp: ta_der'.simps refl_ta_def reflcl_rules_def)
lemma refl_ta_complete2:
"Var q |\<in>| ta_der' (refl_ta \<F> q) s \<Longrightarrow> funas_term s \<subseteq> fset \<F> \<and> vars_term s \<subseteq> {q}"
unfolding ta_der_to_ta_der'[symmetric]
using ta_der_term_sig[of q "refl_ta \<F> q" s] ta_der_states'[of q "refl_ta \<F> q" s]
using fsubsetD[OF \<Q>_refl_ta[of \<F> q]]
- by (auto simp: fmember_iff_member_fset ffunas_term.rep_eq)
- (metis Term.term.simps(17) fresh_states_ta_der'_pres notin_fset singletonD ta_der_to_ta_der')
+ by (auto simp: ffunas_term.rep_eq)
+ (metis Term.term.simps(17) fresh_states_ta_der'_pres singletonD ta_der_to_ta_der')
lemma gen_reflcl_lang:
assumes "q |\<notin>| \<Q> \<A>"
shows "gta_lang (finsert q Q) (gen_reflcl_automaton \<F> \<A> q) = gta_lang Q \<A> \<union> \<T>\<^sub>G (fset \<F>)"
(is "?Ls = ?Rs")
proof -
let ?A = "gen_reflcl_automaton \<F> \<A> q"
interpret sq: derivation_split ?A "\<A>" "refl_ta \<F> q"
using assms unfolding derivation_split_def
by (auto simp: gen_reflcl_automaton_def refl_ta_def reflcl_rules_def \<Q>_def)
show ?thesis
proof
{fix s assume "s \<in> ?Ls" then obtain p u where
seq: "u |\<in>| ta_der' \<A> (term_of_gterm s)" "Var p |\<in>| ta_der' (refl_ta \<F> q) u" and
fin: "p |\<in>| finsert q Q"
by (auto simp: ta_der_to_ta_der' elim!: gta_langE dest!: sq.ta_der'_split)
have "vars_term u \<subseteq> {q} \<Longrightarrow> u = term_of_gterm s" using assms
- by (intro ta_det'_vars_term_id[OF seq(1)]) (auto simp flip: fmember_iff_member_fset)
+ by (intro ta_det'_vars_term_id[OF seq(1)]) auto
then have "s \<in> ?Rs" using assms fin seq funas_term_of_gterm_conv
using refl_ta_complete1[OF seq(2)]
by (cases "p = q") (auto simp: ta_der_to_ta_der' \<T>\<^sub>G_funas_gterm_conv dest!: refl_ta_complete2)}
then show "?Ls \<subseteq> gta_lang Q \<A> \<union> \<T>\<^sub>G (fset \<F>)" by blast
next
show "gta_lang Q \<A> \<union> \<T>\<^sub>G (fset \<F>) \<subseteq> ?Ls"
using sq.ta_der_monos unfolding gta_lang_def gta_der_def
by (auto dest: refl_ta_sound)
qed
qed
lemma reflcl_lang:
"gta_lang (finsert None (Some |`| Q)) (reflcl_automaton \<F> \<A>) = gta_lang Q \<A> \<union> \<T>\<^sub>G (fset \<F>)"
proof -
have st: "None |\<notin>| \<Q> (fmap_states_ta Some \<A>)" by auto
have "gta_lang Q \<A> = gta_lang (Some |`| Q) (fmap_states_ta Some \<A>)"
by (simp add: finj_Some fmap_states_ta_lang2)
then show ?thesis
unfolding reflcl_automaton_def Let_def gen_reflcl_lang[OF st, of "Some |`| Q" \<F>]
by simp
qed
lemma \<L>_reflcl_reg:
"\<L> (reflcl_reg \<F> \<A>) = \<L> \<A> \<union> \<T>\<^sub>G (fset \<F>)"
by (simp add: \<L>_def reflcl_lang reflcl_reg_def )
subsubsection \<open>Correctness of @{const gen_parallel_closure_automaton} and @{const parallel_closure_reg}\<close>
lemma set_list_subset_nth_conv:
"set xs \<subseteq> A \<Longrightarrow> i < length xs \<Longrightarrow> xs ! i \<in> A"
by (metis in_set_conv_nth subset_code(1))
lemma ground_gmctxt_of_mctxt_fill_holes':
"num_holes C = length ss \<Longrightarrow> ground_mctxt C \<Longrightarrow> \<forall>s\<in>set ss. ground s \<Longrightarrow>
fill_gholes (gmctxt_of_mctxt C) (map gterm_of_term ss) = gterm_of_term (fill_holes C ss)"
using ground_gmctxt_of_mctxt_fill_holes
by (metis term_of_gterm_inv)
lemma refl_over_states_ta_eps_trancl [simp]:
"(eps (refl_over_states_ta Q \<F> \<A> q))|\<^sup>+| = eps (refl_over_states_ta Q \<F> \<A> q)"
- using ftranclD ftranclE unfolding refl_over_states_ta_def
- by fastforce
+proof (intro fequalityI fsubsetI)
+ fix x assume "x |\<in>| (eps (refl_over_states_ta Q \<F> \<A> q))|\<^sup>+|"
+ hence "(fst x, snd x) |\<in>| (eps (refl_over_states_ta Q \<F> \<A> q))|\<^sup>+|"
+ by (metis prod.exhaust_sel)
+ thus "x |\<in>| eps (refl_over_states_ta Q \<F> \<A> q)"
+ by (rule ftranclE) (auto simp add: refl_over_states_ta_def image_iff Bex_def dest: ftranclD)
+next
+ fix x assume "x |\<in>| eps (refl_over_states_ta Q \<F> \<A> q)"
+ thus "x |\<in>| (eps (refl_over_states_ta Q \<F> \<A> q))|\<^sup>+|"
+ by (metis fr_into_trancl prod.exhaust_sel)
+qed
lemma refl_over_states_ta_epsD:
"(p, q) |\<in>| (eps (refl_over_states_ta Q \<F> \<A> q)) \<Longrightarrow> p |\<in>| Q"
by (auto simp: refl_over_states_ta_def)
lemma refl_over_states_ta_vars_term:
"q |\<in>| ta_der (refl_over_states_ta Q \<F> \<A> q) u \<Longrightarrow> vars_term u \<subseteq> insert q (fset Q)"
proof (induct u)
case (Fun f ts)
from Fun(2) reflcl_rules_args[of _ "length ts" f _ \<F> q]
have "i < length ts \<Longrightarrow> q |\<in>| ta_der (refl_over_states_ta Q \<F> \<A> q) (ts ! i)" for i
by (fastforce simp: refl_over_states_ta_def)
then have "i < length ts \<Longrightarrow> x \<in> vars_term (ts ! i) \<Longrightarrow> x = q \<or> x |\<in>| Q" for i x
using Fun(1)[OF nth_mem, of i]
- by (meson insert_iff notin_fset subsetD)
- then show ?case by (fastforce simp: in_set_conv_nth fmember_iff_member_fset)
-qed (auto simp flip: fmember_iff_member_fset dest: refl_over_states_ta_epsD)
+ by (meson insert_iff subsetD)
+ then show ?case by (fastforce simp: in_set_conv_nth)
+qed (auto dest: refl_over_states_ta_epsD)
lemmas refl_over_states_ta_vars_term' =
refl_over_states_ta_vars_term[unfolded ta_der_to_ta_der' ta_der'_target_args_vars_term_conv,
- THEN set_list_subset_nth_conv, unfolded fmember_iff_member_fset[symmetric] finsert.rep_eq[symmetric]]
+ THEN set_list_subset_nth_conv, unfolded finsert.rep_eq[symmetric]]
lemma refl_over_states_ta_sound:
"funas_term u \<subseteq> fset \<F> \<Longrightarrow> vars_term u \<subseteq> insert q (fset (Q |\<inter>| \<Q> \<A>)) \<Longrightarrow> q |\<in>| ta_der (refl_over_states_ta Q \<F> \<A> q) u"
proof (induct u)
case (Fun f ts)
have reach: "i < length ts \<Longrightarrow> q |\<in>| ta_der (refl_over_states_ta Q \<F> \<A> q) (ts ! i)" for i
using Fun(2-) by (intro Fun(1)[OF nth_mem]) (auto simp: SUP_le_iff)
from Fun(2) have "TA_rule f (replicate (length ts) q) q |\<in>| rules (refl_over_states_ta Q \<F> \<A> q)"
- by (auto simp: refl_over_states_ta_def reflcl_rules_def fimage_iff fBex_def simp flip: fmember_iff_member_fset)
+ by (auto simp: refl_over_states_ta_def reflcl_rules_def fimage_iff fBex_def)
then show ?case using reach
by force
-qed (auto simp: refl_over_states_ta_def simp flip: fmember_iff_member_fset)
+qed (auto simp: refl_over_states_ta_def)
lemma gen_parallelcl_lang:
fixes \<A> :: "('q, 'f) ta"
assumes "q |\<notin>| \<Q> \<A>"
shows "gta_lang {|q|} (gen_parallel_closure_automaton Q \<F> \<A> q) =
{fill_gholes C ss | C ss. num_gholes C = length ss \<and> funas_gmctxt C \<subseteq> (fset \<F>) \<and> (\<forall> i < length ss. ss ! i \<in> gta_lang Q \<A>)}"
(is "?Ls = ?Rs")
proof -
let ?A = "gen_parallel_closure_automaton Q \<F> \<A> q" let ?B = "refl_over_states_ta Q \<F> \<A> q"
interpret sq: derivation_split "?A" "\<A>" "?B"
using assms unfolding derivation_split_def
by (auto simp: gen_parallel_closure_automaton_def refl_over_states_ta_def \<Q>_def reflcl_rules_def)
{fix s assume "s \<in> ?Ls" then obtain u where
seq: "u |\<in>| ta_der' \<A> (term_of_gterm s)" "Var q |\<in>| ta_der'?B u" and
fin: "q |\<in>| finsert q Q"
by (auto simp: ta_der_to_ta_der' elim!: gta_langE dest!: sq.ta_der'_split)
let ?w = "\<lambda> i. ta_der'_source_args u (term_of_gterm s) ! i"
have "s \<in> ?Rs" using seq(1) ta_der'_Var_funas[OF seq(2)] fin
using ground_ta_der_statesD[of "?w i" "ta_der'_target_args u ! i" \<A> for i] assms
using refl_over_states_ta_vars_term'[OF seq(2)]
using ta_der'_ground_mctxt_structure[OF seq(1)]
by (force simp: ground_gmctxt_of_mctxt_fill_holes' ta_der'_source_args_term_of_gterm
intro!: exI[of _ "gmctxt_of_mctxt (ta_der'_target_mctxt u)"]
exI[of _ "map gterm_of_term (ta_der'_source_args u (term_of_gterm s))"]
gta_langI[of "ta_der'_target_args u ! i" Q \<A>
"gterm_of_term (?w i)" for i])}
then have ls: "?Ls \<subseteq> ?Rs" by blast
{fix t assume "t \<in> ?Rs"
then obtain C ss where len: "num_gholes C = length ss" and
gr_fun: "funas_gmctxt C \<subseteq> fset \<F>" and
reachA: "\<forall> i < length ss. ss ! i \<in> gta_lang Q \<A>" and
const: "t = fill_gholes C ss" by auto
from reachA obtain qs where "length ss = length qs" "\<forall> i < length qs. qs ! i |\<in>| Q |\<inter>| \<Q> \<A>"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> ((map term_of_gterm ss) ! i)"
using Ex_list_of_length_P[of "length ss" "\<lambda> q i. q |\<in>| ta_der \<A> (term_of_gterm (ss ! i)) \<and> q |\<in>| Q"]
by (metis (full_types) finterI gta_langE gterm_ta_der_states length_map map_nth_eq_conv)
then have "q |\<in>| ta_der ?A (fill_holes (mctxt_of_gmctxt C) (map term_of_gterm ss))"
using reachA len gr_fun
by (intro sq.mctxt_const_to_ta_der[of "mctxt_of_gmctxt C" "map term_of_gterm ss" qs q])
- (auto simp: funas_mctxt_of_gmctxt_conv simp flip: fmember_iff_member_fset
+ (auto simp: funas_mctxt_of_gmctxt_conv
dest!: in_set_idx intro!: refl_over_states_ta_sound)
then have "t \<in> ?Ls" unfolding const
by (simp add: fill_holes_mctxt_of_gmctxt_to_fill_gholes gta_langI len)}
then show ?thesis using ls by blast
qed
lemma parallelcl_gmctxt_lang:
fixes \<A> :: "('q, 'f) reg"
shows "\<L> (parallel_closure_reg \<F> \<A>) =
{fill_gholes C ss |
C ss. num_gholes C = length ss \<and> funas_gmctxt C \<subseteq> fset \<F> \<and> (\<forall> i < length ss. ss ! i \<in> \<L> \<A>)}"
proof -
have *: "gta_lang (fin (fmap_states_reg Some \<A>)) (fmap_states_ta Some (ta \<A>)) = gta_lang (fin \<A>) (ta \<A>)"
by (simp add: finj_Some fmap_states_reg_def fmap_states_ta_lang2)
have " None |\<notin>| \<Q> (fmap_states_ta Some (ta \<A>))" by auto
from gen_parallelcl_lang[OF this, of "fin (fmap_states_reg Some \<A>)" \<F>] show ?thesis
unfolding \<L>_def parallel_closure_reg_def Let_def * fmap_states_reg_def
by (simp add: finj_Some fmap_states_ta_lang2)
qed
lemma parallelcl_mctxt_lang:
shows "\<L> (parallel_closure_reg \<F> \<A>) =
{(gterm_of_term :: ('f, 'q option) term \<Rightarrow> 'f gterm) (fill_holes C (map term_of_gterm ss)) |
C ss. num_holes C = length ss \<and> ground_mctxt C \<and> funas_mctxt C \<subseteq> fset \<F> \<and> (\<forall> i < length ss. ss ! i \<in> \<L> \<A>)}"
by (auto simp: parallelcl_gmctxt_lang) (metis funas_gmctxt_of_mctxt num_gholes_gmctxt_of_mctxt
ground_gmctxt_of_gterm_of_term funas_mctxt_of_gmctxt_conv
ground_mctxt_of_gmctxt mctxt_of_gmctxt_fill_holes num_holes_mctxt_of_gmctxt)+
subsubsection \<open>Correctness of @{const gen_ctxt_closure_reg} and @{const ctxt_closure_reg}\<close>
lemma semantic_path_rules_rhs:
"r |\<in>| semantic_path_rules Q q\<^sub>c q\<^sub>i q\<^sub>f \<Longrightarrow> r_rhs r = q\<^sub>f"
by (auto simp: semantic_path_rules_def)
lemma reflcl_over_single_ta_transl [simp]:
"(eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f))|\<^sup>+| = eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f)"
- using ftranclD ftranclE unfolding reflcl_over_single_ta_def
- by fastforce
+proof (intro fequalityI fsubsetI)
+ fix x assume "x |\<in>| (eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f))|\<^sup>+|"
+ hence "(fst x, snd x) |\<in>| (eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f))|\<^sup>+|"
+ by simp
+ thus "x |\<in>| eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f)"
+ by (smt (verit, ccfv_threshold) fimageE ftranclD ftranclE prod.collapse
+ reflcl_over_single_ta_def snd_conv ta.sel(2))
+next
+ show "\<And>x. x |\<in>| eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) \<Longrightarrow>
+ x |\<in>| (eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f))|\<^sup>+|"
+ by auto
+qed
lemma reflcl_over_single_ta_epsD:
"(p, q\<^sub>f) |\<in>| eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) \<Longrightarrow> p |\<in>| Q"
"(p, q) |\<in>| eps (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) \<Longrightarrow> q = q\<^sub>f"
by (auto simp: reflcl_over_single_ta_def)
lemma reflcl_over_single_ta_rules_split:
"r |\<in>| rules (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) \<Longrightarrow>
r |\<in>| reflcl_rules \<F> q\<^sub>c \<or> r |\<in>| semantic_path_rules \<F> q\<^sub>c q\<^sub>f q\<^sub>f"
by (auto simp: reflcl_over_single_ta_def)
lemma reflcl_over_single_ta_rules_semantic_path_rulesI:
"r |\<in>| semantic_path_rules \<F> q\<^sub>c q\<^sub>f q\<^sub>f \<Longrightarrow> r |\<in>| rules (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f)"
by (auto simp: reflcl_over_single_ta_def)
lemma semantic_path_rules_fmember [intro]:
"TA_rule f qs q |\<in>| semantic_path_rules \<F> q\<^sub>c q\<^sub>i q\<^sub>f \<longleftrightarrow> (\<exists> n i. (f, n) |\<in>| \<F> \<and> i < n \<and> q = q\<^sub>f \<and>
(qs = (replicate n q\<^sub>c)[i := q\<^sub>i]))" (is "?Ls \<longleftrightarrow> ?Rs")
by (force simp: semantic_path_rules_def fBex_def fimage_iff fset_of_list_elem)
lemma semantic_path_rules_fmemberD:
"r |\<in>| semantic_path_rules \<F> q\<^sub>c q\<^sub>i q\<^sub>f \<Longrightarrow> (\<exists> n i. (r_root r, n) |\<in>| \<F> \<and> i < n \<and> r_rhs r = q\<^sub>f \<and>
(r_lhs_states r = (replicate n q\<^sub>c)[i := q\<^sub>i]))"
by (cases r) (simp add: semantic_path_rules_fmember)
lemma reflcl_over_single_ta_vars_term_q\<^sub>c:
"q\<^sub>c \<noteq> q\<^sub>f \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) u \<Longrightarrow>
vars_term_list u = replicate (length (vars_term_list u)) q\<^sub>c"
proof (induct u)
case (Fun f ts)
have "i < length ts \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) (ts ! i)" for i using Fun(2, 3)
by (auto dest!: reflcl_over_single_ta_rules_split reflcl_over_single_ta_epsD
reflcl_rules_args semantic_path_rules_rhs)
then have "i < length (concat (map vars_term_list ts)) \<Longrightarrow> concat (map vars_term_list ts) ! i = q\<^sub>c" for i
using Fun(1)[OF nth_mem Fun(2)]
by (metis (no_types, lifting) length_map nth_concat_split nth_map nth_replicate)
then show ?case using Fun(1)[OF nth_mem Fun(2)]
by (auto intro: nth_equalityI)
-qed (auto simp flip: fmember_iff_member_fset dest: reflcl_over_single_ta_epsD)
+qed (auto dest: reflcl_over_single_ta_epsD)
lemma reflcl_over_single_ta_vars_term:
"q\<^sub>c |\<notin>| Q \<Longrightarrow> q\<^sub>c \<noteq> q\<^sub>f \<Longrightarrow> q\<^sub>f |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) u \<Longrightarrow>
length (vars_term_list u) = n \<Longrightarrow> (\<exists> i q. i < n \<and> q |\<in>| finsert q\<^sub>f Q \<and> vars_term_list u = (replicate n q\<^sub>c)[i := q])"
proof (induct u arbitrary: n)
case (Var x) then show ?case
by (intro exI[of _ 0] exI[of _ x]) (auto dest: reflcl_over_single_ta_epsD(1))
next
case (Fun f ts)
from Fun(2, 3, 4) obtain qs where rule: "TA_rule f qs q\<^sub>f |\<in>| semantic_path_rules \<F> q\<^sub>c q\<^sub>f q\<^sub>f"
"length qs = length ts" "\<forall> i < length ts. qs ! i |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) (ts ! i)"
using semantic_path_rules_rhs reflcl_over_single_ta_epsD
by (fastforce simp: reflcl_rules_def dest!: reflcl_over_single_ta_rules_split)
from rule(1, 2) obtain i where states: "i < length ts" "qs = (replicate (length ts) q\<^sub>c)[i := q\<^sub>f]"
by (auto simp: semantic_path_rules_fmember)
then have qc: "j < length ts \<Longrightarrow> j \<noteq> i \<Longrightarrow> vars_term_list (ts ! j) = replicate (length (vars_term_list (ts ! j))) q\<^sub>c" for j
using reflcl_over_single_ta_vars_term_q\<^sub>c[OF Fun(3)] rule
by force
from Fun(1)[OF nth_mem, of i] Fun(2, 3) rule states obtain k q where
qf: "k < length (vars_term_list (ts ! i))" "q |\<in>| finsert q\<^sub>f Q"
"vars_term_list (ts ! i) = (replicate (length (vars_term_list (ts ! i))) q\<^sub>c)[k := q]"
by (auto simp: nth_list_update split: if_splits)
let ?l = "sum_list (map length (take i (map vars_term_list ts))) + k"
show ?case using qc qf rule(2) Fun(5) states(1)
apply (intro exI[of _ ?l] exI[of _ q])
apply (auto simp: concat_nth_length nth_list_update elim!: nth_concat_split' intro!: nth_equalityI)
apply (metis length_replicate nth_list_update_eq nth_list_update_neq nth_replicate)+
done
qed
lemma refl_ta_reflcl_over_single_ta_mono:
"q |\<in>| ta_der (refl_ta \<F> q) t \<Longrightarrow> q |\<in>| ta_der (reflcl_over_single_ta Q \<F> q q\<^sub>f) t"
by (intro ta_der_el_mono[where ?\<B> = "reflcl_over_single_ta Q \<F> q q\<^sub>f"])
(auto simp: refl_ta_def reflcl_over_single_ta_def)
lemma reflcl_over_single_ta_sound:
assumes "funas_gctxt C \<subseteq> fset \<F>" "q |\<in>| Q"
shows "q\<^sub>f |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) (ctxt_of_gctxt C)\<langle>Var q\<rangle>" using assms(1)
proof (induct C)
case GHole then show ?case using assms(2)
by (auto simp add: reflcl_over_single_ta_def)
next
case (GMore f ss C ts)
let ?i = "length ss" let ?n = "Suc (length ss + length ts)"
- from GMore have "(f, ?n) |\<in>| \<F>" by (auto simp flip: fmember_iff_member_fset)
+ from GMore have "(f, ?n) |\<in>| \<F>" by auto
then have "f ((replicate ?n q\<^sub>c)[?i := q\<^sub>f]) \<rightarrow> q\<^sub>f |\<in>| rules (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f)"
using semantic_path_rules_fmember[of f "(replicate ?n q\<^sub>c)[?i := q\<^sub>f]" q\<^sub>f \<F> q\<^sub>c q\<^sub>f q\<^sub>f]
using less_add_Suc1
by (intro reflcl_over_single_ta_rules_semantic_path_rulesI) blast
moreover from GMore(2) have "i < length ss \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) (term_of_gterm (ss ! i))" for i
by (intro refl_ta_reflcl_over_single_ta_mono refl_ta_sound) (auto simp: SUP_le_iff \<T>\<^sub>G_funas_gterm_conv)
moreover from GMore(2) have "i < length ts \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) (term_of_gterm (ts ! i))" for i
by (intro refl_ta_reflcl_over_single_ta_mono refl_ta_sound) (auto simp: SUP_le_iff \<T>\<^sub>G_funas_gterm_conv)
moreover from GMore have "q\<^sub>f |\<in>| ta_der (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) (ctxt_of_gctxt C)\<langle>Var q\<rangle>" by auto
ultimately show ?case
by (auto simp: nth_append_Cons simp del: replicate.simps intro!: exI[of _ "(replicate ?n q\<^sub>c)[?i := q\<^sub>f]"] exI[of _ q\<^sub>f])
qed
lemma reflcl_over_single_ta_sig: "ta_sig (reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f) |\<subseteq>| \<F>"
by (intro ta_sig_fsubsetI)
(auto simp: reflcl_rules_def dest!: semantic_path_rules_fmemberD reflcl_over_single_ta_rules_split)
lemma gen_gctxtcl_lang:
assumes "q\<^sub>c |\<notin>| \<Q> \<A>" and "q\<^sub>f |\<notin>| \<Q> \<A>" and "q\<^sub>c |\<notin>| Q" and "q\<^sub>c \<noteq> q\<^sub>f"
shows "gta_lang {|q\<^sub>f|} (gen_ctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>f) =
{C\<langle>s\<rangle>\<^sub>G | C s. funas_gctxt C \<subseteq> fset \<F> \<and> s \<in> gta_lang Q \<A>}"
(is "?Ls = ?Rs")
proof -
let ?A = "gen_ctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>f" let ?B = "reflcl_over_single_ta Q \<F> q\<^sub>c q\<^sub>f"
interpret sq: derivation_split ?A \<A> ?B
using assms unfolding derivation_split_def
by (auto simp: gen_ctxt_closure_automaton_def reflcl_over_single_ta_def \<Q>_def reflcl_rules_def
dest!: semantic_path_rules_rhs)
{fix s assume "s \<in> ?Ls" then obtain u where
seq: "u |\<in>| ta_der' \<A> (term_of_gterm s)" "Var q\<^sub>f |\<in>| ta_der'?B u" using sq.ta_der'_split
by (force simp: ta_der_to_ta_der' elim!: gta_langE)
have "q\<^sub>c \<notin> vars_term u" "q\<^sub>f \<notin> vars_term u"
using subsetD[OF ta_der'_gterm_states[OF seq(1)]] assms(1, 2)
- by (auto simp flip: set_vars_term_list fmember_iff_member_fset)
+ by (auto simp flip: set_vars_term_list)
then obtain q where vars: "vars_term_list u = [q]" and fin: "q |\<in>| Q" unfolding set_vars_term_list[symmetric]
using reflcl_over_single_ta_vars_term[unfolded ta_der_to_ta_der', OF assms(3, 4) seq(2), of "length (vars_term_list u)"]
by (metis (no_types, lifting) finsertE in_set_conv_nth length_0_conv length_Suc_conv
length_replicate lessI less_Suc_eq_0_disj nth_Cons_0 nth_list_update nth_replicate zero_less_Suc)
have "s \<in> ?Rs" using fin ta_der'_ground_ctxt_structure[OF seq(1) vars]
using ta_der'_Var_funas[OF seq(2), THEN subset_trans, OF reflcl_over_single_ta_sig[unfolded less_eq_fset.rep_eq]]
by (auto intro!: exI[of _ "ta_der'_gctxt u"] exI[of _ "ta_der'_source_gctxt_arg u s"])
(metis Un_iff funas_ctxt_apply funas_ctxt_of_gctxt_conv subset_eq)
}
then have ls: "?Ls \<subseteq> ?Rs" by blast
{fix t assume "t \<in> ?Rs"
then obtain C s where gr_fun: "funas_gctxt C \<subseteq> fset \<F>" and reachA: "s \<in> gta_lang Q \<A>" and
const: "t = C\<langle>s\<rangle>\<^sub>G" by auto
from reachA obtain q where der_A: "q |\<in>| Q |\<inter>| \<Q> \<A>" "q |\<in>| ta_der \<A> (term_of_gterm s)"
by auto
have "q\<^sub>f |\<in>| ta_der ?B (ctxt_of_gctxt C)\<langle>Var q\<rangle>" using gr_fun der_A(1)
using reflcl_over_single_ta_sound[OF gr_fun]
by force
then have "t \<in> ?Ls" unfolding const
by (meson der_A(2) finsertI1 gta_langI sq.gctxt_const_to_ta_der)}
then show ?thesis using ls by blast
qed
lemma gen_gctxt_closure_sound:
fixes \<A> :: "('q, 'f) reg"
assumes "q\<^sub>c |\<notin>| \<Q>\<^sub>r \<A>" and "q\<^sub>f |\<notin>| \<Q>\<^sub>r \<A>" and "q\<^sub>c |\<notin>| fin \<A>" and "q\<^sub>c \<noteq> q\<^sub>f"
shows "\<L> (gen_ctxt_closure_reg \<F> \<A> q\<^sub>c q\<^sub>f) = {C\<langle>s\<rangle>\<^sub>G | C s. funas_gctxt C \<subseteq> fset \<F> \<and> s \<in> \<L> \<A>}"
using gen_gctxtcl_lang[OF assms] unfolding \<L>_def
by (simp add: gen_ctxt_closure_reg_def)
lemma gen_ctxt_closure_sound:
fixes \<A> :: "('q, 'f) reg"
assumes "q\<^sub>c |\<notin>| \<Q>\<^sub>r \<A>" and "q\<^sub>f |\<notin>| \<Q>\<^sub>r \<A>" and "q\<^sub>c |\<notin>| fin \<A>" and "q\<^sub>c \<noteq> q\<^sub>f"
shows "\<L> (gen_ctxt_closure_reg \<F> \<A> q\<^sub>c q\<^sub>f) =
{(gterm_of_term :: ('f, 'q) term \<Rightarrow> 'f gterm) C\<langle>term_of_gterm s\<rangle> | C s. ground_ctxt C \<and> funas_ctxt C \<subseteq> fset \<F> \<and> s \<in> \<L> \<A>}"
unfolding gen_gctxt_closure_sound[OF assms]
by (metis (no_types, opaque_lifting) ctxt_of_gctxt_apply funas_ctxt_of_gctxt_conv gctxt_of_ctxt_inv ground_ctxt_of_gctxt)
lemma gctxt_closure_lang:
shows "\<L> (ctxt_closure_reg \<F> \<A>) =
{ C\<langle>s\<rangle>\<^sub>G | C s. funas_gctxt C \<subseteq> fset \<F> \<and> s \<in> \<L> \<A>}"
proof -
let ?B = "fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>)"
have ts: "Inr False |\<notin>| \<Q>\<^sub>r ?B" "Inr True |\<notin>| \<Q>\<^sub>r ?B" "Inr False |\<notin>| fin (fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>))"
by (auto simp: fmap_states_reg_def fmap_states_ta_def' \<Q>_def rule_states_def)
from gen_gctxt_closure_sound[OF ts] show ?thesis
by (simp add: ctxt_closure_reg_def)
qed
lemma ctxt_closure_lang:
shows "\<L> (ctxt_closure_reg \<F> \<A>) =
{(gterm_of_term :: ('f, 'q + bool) term \<Rightarrow> 'f gterm) C\<langle>term_of_gterm s\<rangle> |
C s. ground_ctxt C \<and> funas_ctxt C \<subseteq> fset \<F> \<and> s \<in> \<L> \<A>}"
unfolding gctxt_closure_lang
by (metis (mono_tags, opaque_lifting) ctxt_of_gctxt_inv funas_gctxt_of_ctxt
ground_ctxt_of_gctxt ground_gctxt_of_ctxt_apply_gterm term_of_gterm_inv)
subsubsection \<open>Correctness of @{const gen_nhole_ctxt_closure_automaton} and @{const nhole_ctxt_closure_reg}\<close>
lemma reflcl_over_nhole_ctxt_ta_vars_term_q\<^sub>c:
"q\<^sub>c \<noteq> q\<^sub>f \<Longrightarrow> q\<^sub>c \<noteq> q\<^sub>i \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) u \<Longrightarrow>
vars_term_list u = replicate (length (vars_term_list u)) q\<^sub>c"
proof (induct u)
case (Fun f ts)
have "i < length ts \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (ts ! i)" for i using Fun(2, 3, 4)
by (auto simp: reflcl_over_nhole_ctxt_ta_def dest!: ftranclD2 reflcl_rules_args semantic_path_rules_rhs)
then have "i < length (concat (map vars_term_list ts)) \<Longrightarrow> concat (map vars_term_list ts) ! i = q\<^sub>c" for i
using Fun(1)[OF nth_mem Fun(2, 3)]
by (metis (no_types, lifting) length_map map_nth_eq_conv nth_concat_split' nth_replicate)
then show ?case using Fun(1)[OF nth_mem Fun(2)]
by (auto intro: nth_equalityI)
-qed (auto simp flip: fmember_iff_member_fset simp: reflcl_over_nhole_ctxt_ta_def dest: ftranclD2)
+qed (auto simp: reflcl_over_nhole_ctxt_ta_def dest: ftranclD2)
lemma reflcl_over_nhole_ctxt_ta_vars_term_Var:
assumes disj: "q\<^sub>c |\<notin>| Q" "q\<^sub>f |\<notin>| Q" "q\<^sub>c \<noteq> q\<^sub>f" "q\<^sub>i \<noteq> q\<^sub>f" "q\<^sub>c \<noteq> q\<^sub>i"
and reach: "q\<^sub>i |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) u"
shows "(\<exists> q. q |\<in>| finsert q\<^sub>i Q \<and> u = Var q)" using assms
by (cases u) (fastforce simp: reflcl_over_nhole_ctxt_ta_def reflcl_rules_def dest: ftranclD semantic_path_rules_rhs)+
lemma reflcl_over_nhole_ctxt_ta_vars_term:
assumes disj: "q\<^sub>c |\<notin>| Q" "q\<^sub>f |\<notin>| Q" "q\<^sub>c \<noteq> q\<^sub>f" "q\<^sub>i \<noteq> q\<^sub>f" "q\<^sub>c \<noteq> q\<^sub>i"
and reach: "q\<^sub>f |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) u"
shows "(\<exists> i q. i < length (vars_term_list u) \<and> q |\<in>| {|q\<^sub>i, q\<^sub>f|} |\<union>| Q \<and> vars_term_list u = (replicate (length (vars_term_list u)) q\<^sub>c)[i := q])"
using assms
proof (induct u)
case (Var q) then show ?case
by (intro exI[of _ 0] exI[of _ q]) (auto simp: reflcl_over_nhole_ctxt_ta_def dest: ftranclD2)
next
case (Fun f ts)
from Fun(2 - 7) obtain q qs where rule: "TA_rule f qs q\<^sub>f |\<in>| semantic_path_rules \<F> q\<^sub>c q q\<^sub>f" "q = q\<^sub>i \<or> q = q\<^sub>f"
"length qs = length ts" "\<forall> i < length ts. qs ! i |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (ts ! i)"
by (auto simp: reflcl_over_nhole_ctxt_ta_def reflcl_rules_def dest: ftranclD2)
from rule(1- 3) obtain i where states: "i < length ts" "qs = (replicate (length ts) q\<^sub>c)[i := q]"
by (auto simp: semantic_path_rules_fmember)
then have qc: "j < length ts \<Longrightarrow> j \<noteq> i \<Longrightarrow> vars_term_list (ts ! j) = replicate (length (vars_term_list (ts ! j))) q\<^sub>c" for j
using reflcl_over_nhole_ctxt_ta_vars_term_q\<^sub>c[OF Fun(4, 6)] rule
by force
from rule states have "q |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (ts ! i)"
by auto
from this Fun(1)[OF nth_mem, of i, OF _ Fun(2 - 6)] rule(2) states(1) obtain k q' where
qf: "k < length (vars_term_list (ts ! i))" "q' |\<in>| {|q\<^sub>i, q\<^sub>f|} |\<union>| Q "
"vars_term_list (ts ! i) = (replicate (length (vars_term_list (ts ! i))) q\<^sub>c)[k := q']"
using reflcl_over_nhole_ctxt_ta_vars_term_Var[OF Fun(2 - 6), of \<F> "ts ! i"]
by (auto simp: nth_list_update split: if_splits) blast
let ?l = "sum_list (map length (take i (map vars_term_list ts))) + k"
show ?case using qc qf rule(3) states(1)
apply (intro exI[of _ ?l] exI[of _ q'])
apply (auto 0 0 simp: concat_nth_length nth_list_update elim!: nth_concat_split' intro!: nth_equalityI)
apply (metis length_replicate nth_list_update_eq nth_list_update_neq nth_replicate)+
done
qed
lemma reflcl_over_nhole_ctxt_ta_mono:
"q |\<in>| ta_der (refl_ta \<F> q) t \<Longrightarrow> q |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q q\<^sub>i q\<^sub>f) t"
by (intro ta_der_el_mono[where ?\<B> = "reflcl_over_nhole_ctxt_ta Q \<F> q q\<^sub>i q\<^sub>f"])
(auto simp: refl_ta_def reflcl_over_nhole_ctxt_ta_def)
lemma reflcl_over_nhole_ctxt_ta_sound:
assumes "funas_gctxt C \<subseteq> fset \<F>" "C \<noteq> GHole" "q |\<in>| Q"
shows "q\<^sub>f |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (ctxt_of_gctxt C)\<langle>Var q\<rangle>" using assms(1, 2)
proof (induct C)
case GHole then show ?case using assms(2)
by (auto simp add: reflcl_over_single_ta_def)
next
case (GMore f ss C ts) note IH = this
let ?i = "length ss" let ?n = "Suc (length ss + length ts)"
- from GMore have funas: "(f, ?n) |\<in>| \<F>" by (auto simp flip: fmember_iff_member_fset)
+ from GMore have funas: "(f, ?n) |\<in>| \<F>" by auto
from GMore(2) have args_ss: "i < length ss \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (term_of_gterm (ss ! i))" for i
by (intro reflcl_over_nhole_ctxt_ta_mono refl_ta_sound) (auto simp: SUP_le_iff \<T>\<^sub>G_funas_gterm_conv)
from GMore(2) have args_ts: "i < length ts \<Longrightarrow> q\<^sub>c |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (term_of_gterm (ts ! i))" for i
by (intro reflcl_over_nhole_ctxt_ta_mono refl_ta_sound) (auto simp: SUP_le_iff \<T>\<^sub>G_funas_gterm_conv)
note args = this
show ?case
proof (cases C)
case [simp]: GHole
from funas have "f ((replicate ?n q\<^sub>c)[?i := q\<^sub>i]) \<rightarrow> q\<^sub>f |\<in>| rules (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f)"
using semantic_path_rules_fmember[of f "(replicate ?n q\<^sub>c)[?i := q\<^sub>i]" q\<^sub>f \<F> q\<^sub>c q\<^sub>i q\<^sub>f]
unfolding reflcl_over_nhole_ctxt_ta_def
by (metis funionCI less_add_Suc1 ta.sel(1))
moreover have "q\<^sub>i |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (ctxt_of_gctxt C)\<langle>Var q\<rangle>"
using assms(3) by (auto simp: reflcl_over_nhole_ctxt_ta_def)
ultimately show ?thesis using args_ss args_ts
by (auto simp: nth_append_Cons simp del: replicate.simps intro!: exI[of _ "(replicate ?n q\<^sub>c)[?i := q\<^sub>i]"] exI[of _ q\<^sub>f])
next
case (GMore x21 x22 x23 x24)
then have "q\<^sub>f |\<in>| ta_der (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) (ctxt_of_gctxt C)\<langle>Var q\<rangle>"
using IH(1, 2) by auto
moreover from funas have "f ((replicate ?n q\<^sub>c)[?i := q\<^sub>f]) \<rightarrow> q\<^sub>f |\<in>| rules (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f)"
using semantic_path_rules_fmember[of f "(replicate ?n q\<^sub>c)[?i := q\<^sub>f]" q\<^sub>f \<F> q\<^sub>c q\<^sub>f q\<^sub>f]
unfolding reflcl_over_nhole_ctxt_ta_def
by (metis funionI2 less_add_Suc1 ta.sel(1))
ultimately show ?thesis using args_ss args_ts
by (auto simp: nth_append_Cons simp del: replicate.simps intro!: exI[of _ "(replicate ?n q\<^sub>c)[?i := q\<^sub>f]"] exI[of _ q\<^sub>f])
qed
qed
lemma reflcl_over_nhole_ctxt_ta_sig: "ta_sig (reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) |\<subseteq>| \<F>"
by (intro ta_sig_fsubsetI)
(auto simp: reflcl_over_nhole_ctxt_ta_def reflcl_rules_def dest!: semantic_path_rules_fmemberD)
lemma gen_nhole_gctxt_closure_lang:
assumes "q\<^sub>c |\<notin>| \<Q> \<A>" "q\<^sub>i |\<notin>| \<Q> \<A>" "q\<^sub>f |\<notin>| \<Q> \<A>"
and "q\<^sub>c |\<notin>| Q" "q\<^sub>f |\<notin>| Q"
and "q\<^sub>c \<noteq> q\<^sub>i" "q\<^sub>c \<noteq> q\<^sub>f" "q\<^sub>i \<noteq> q\<^sub>f"
shows "gta_lang {|q\<^sub>f|} (gen_nhole_ctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f) =
{C\<langle>s\<rangle>\<^sub>G | C s. C \<noteq> GHole \<and> funas_gctxt C \<subseteq> fset \<F> \<and> s \<in> gta_lang Q \<A>}"
(is "?Ls = ?Rs")
proof -
let ?A = "gen_nhole_ctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f" let ?B = "reflcl_over_nhole_ctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f"
interpret sq: derivation_split ?A \<A> ?B
using assms unfolding derivation_split_def
by (auto simp: gen_nhole_ctxt_closure_automaton_def reflcl_over_nhole_ctxt_ta_def \<Q>_def reflcl_rules_def
dest!: semantic_path_rules_rhs)
{fix s assume "s \<in> ?Ls" then obtain u where
seq: "u |\<in>| ta_der' \<A> (term_of_gterm s)" "Var q\<^sub>f |\<in>| ta_der'?B u" using sq.ta_der'_split
by (force simp: ta_der_to_ta_der' elim!: gta_langE)
have "q\<^sub>c \<notin> vars_term u" "q\<^sub>i \<notin> vars_term u" "q\<^sub>f \<notin> vars_term u"
using subsetD[OF ta_der'_gterm_states[OF seq(1)]] assms(1 - 3)
- by (auto simp flip: set_vars_term_list fmember_iff_member_fset)
+ by (auto simp flip: set_vars_term_list)
then obtain q where vars: "vars_term_list u = [q]" and fin: "q |\<in>| Q"
unfolding set_vars_term_list[symmetric]
using reflcl_over_nhole_ctxt_ta_vars_term[unfolded ta_der_to_ta_der', OF assms(4, 5, 7 - 8, 6) seq(2)]
by (metis (no_types, opaque_lifting) finsert_iff funion_commute funion_finsert_right
length_greater_0_conv lessI list.size(3) list_update_code(2) not0_implies_Suc
nth_list_update_neq nth_mem nth_replicate replicate_Suc replicate_empty sup_bot.right_neutral)
from seq(2) have "ta_der'_gctxt u \<noteq> GHole" using ta_der'_ground_ctxt_structure(1)[OF seq(1) vars]
using fin assms(4, 5, 8) by (auto simp: reflcl_over_nhole_ctxt_ta_def dest!: ftranclD2)
then have "s \<in> ?Rs" using fin ta_der'_ground_ctxt_structure[OF seq(1) vars] seq(2)
using ta_der'_Var_funas[OF seq(2), THEN subset_trans, OF reflcl_over_nhole_ctxt_ta_sig[unfolded less_eq_fset.rep_eq]]
by (auto intro!: exI[of _ "ta_der'_gctxt u"] exI[of _ "ta_der'_source_gctxt_arg u s"])
(metis Un_iff funas_ctxt_apply funas_ctxt_of_gctxt_conv in_mono)}
then have ls: "?Ls \<subseteq> ?Rs" by blast
{fix t assume "t \<in> ?Rs"
then obtain C s where gr_fun: "funas_gctxt C \<subseteq> fset \<F>" "C \<noteq> GHole" and reachA: "s \<in> gta_lang Q \<A>" and
const: "t = C\<langle>s\<rangle>\<^sub>G" by auto
from reachA obtain q where der_A: "q |\<in>| Q |\<inter>| \<Q> \<A>" "q |\<in>| ta_der \<A> (term_of_gterm s)"
by auto
have "q\<^sub>f |\<in>| ta_der ?B (ctxt_of_gctxt C)\<langle>Var q\<rangle>" using gr_fun der_A(1)
using reflcl_over_nhole_ctxt_ta_sound[OF gr_fun]
by force
then have "t \<in> ?Ls" unfolding const
by (meson der_A(2) finsertI1 gta_langI sq.gctxt_const_to_ta_der)}
then show ?thesis using ls by blast
qed
lemma gen_nhole_gctxt_closure_sound:
assumes "q\<^sub>c |\<notin>| \<Q>\<^sub>r \<A>" "q\<^sub>i |\<notin>| \<Q>\<^sub>r \<A>" "q\<^sub>f |\<notin>| \<Q>\<^sub>r \<A>"
and "q\<^sub>c |\<notin>| (fin \<A>)" "q\<^sub>f |\<notin>| (fin \<A>)"
and "q\<^sub>c \<noteq> q\<^sub>i" "q\<^sub>c \<noteq> q\<^sub>f" "q\<^sub>i \<noteq> q\<^sub>f"
shows "\<L> (gen_nhole_ctxt_closure_reg \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f) =
{ C\<langle>s\<rangle>\<^sub>G | C s. C \<noteq> GHole \<and> funas_gctxt C \<subseteq> fset \<F> \<and> s \<in> \<L> \<A>}"
using gen_nhole_gctxt_closure_lang[OF assms] unfolding \<L>_def
by (auto simp: gen_nhole_ctxt_closure_reg_def)
lemma nhole_ctxtcl_lang:
"\<L> (nhole_ctxt_closure_reg \<F> \<A>) =
{ C\<langle>s\<rangle>\<^sub>G | C s. C \<noteq> GHole \<and> funas_gctxt C \<subseteq> fset \<F> \<and> s \<in> \<L> \<A>}"
proof -
let ?B = "fmap_states_reg (Inl :: 'b \<Rightarrow> 'b + cl_states) (reg_Restr_Q\<^sub>f \<A>)"
have ts: "Inr cl_state |\<notin>| \<Q>\<^sub>r ?B" "Inr tr_state |\<notin>| \<Q>\<^sub>r ?B" "Inr fin_state |\<notin>| \<Q>\<^sub>r ?B"
by (auto simp: fmap_states_reg_def)
then have "Inr cl_state |\<notin>| fin (fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>))"
"Inr fin_state |\<notin>| fin (fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>))"
using finj_Inl_Inr(1) fmap_states_reg_Restr_Q\<^sub>f_fin by blast+
from gen_nhole_gctxt_closure_sound[OF ts this] show ?thesis
by (simp add: nhole_ctxt_closure_reg_def Let_def)
qed
subsubsection \<open>Correctness of @{const gen_nhole_mctxt_closure_automaton}\<close>
lemmas reflcl_over_nhole_mctxt_ta_simp = reflcl_over_nhole_mctxt_ta_def reflcl_over_nhole_ctxt_ta_def
lemma reflcl_rules_rhsD:
"f ps \<rightarrow> q |\<in>| reflcl_rules \<F> q\<^sub>c \<Longrightarrow> q = q\<^sub>c"
by (auto simp: reflcl_rules_def)
lemma reflcl_over_nhole_mctxt_ta_vars_term:
assumes "q |\<in>| ta_der (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) t"
and "q\<^sub>c |\<notin>| Q" "q \<noteq> q\<^sub>c" "q\<^sub>f \<noteq> q\<^sub>c" "q\<^sub>i \<noteq> q\<^sub>c"
shows "vars_term t \<noteq> {}" using assms
proof (induction t arbitrary: q)
case (Fun f ts)
let ?A = "reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f"
from Fun(2) obtain p ps where rule: "TA_rule f ps p |\<in>| rules ?A"
"length ps = length ts" "\<forall> i < length ts. ps ! i |\<in>| ta_der ?A (ts ! i)"
"p = q \<or> (p, q) |\<in>| (eps ?A)|\<^sup>+|"
by force
from rule(1, 4) Fun(3-) have "p \<noteq> q\<^sub>c"
by (auto simp: reflcl_over_nhole_mctxt_ta_simp dest: ftranclD)
then have "\<exists> i < length ts. ps ! i \<noteq> q\<^sub>c" using rule(1, 2) Fun(4-)
using semantic_path_rules_fmemberD
by (force simp: reflcl_over_nhole_mctxt_ta_simp dest: reflcl_rules_rhsD)
then show ?case using Fun(1)[OF nth_mem _ Fun(3) _ Fun(5, 6)] rule(2, 3)
by fastforce
qed auto
lemma reflcl_over_nhole_mctxt_ta_Fun:
assumes "q\<^sub>f |\<in>| ta_der (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) t" "t \<noteq> Var q\<^sub>f"
and "q\<^sub>f \<noteq> q\<^sub>c" "q\<^sub>f \<noteq> q\<^sub>i"
shows "is_Fun t" using assms
by (cases t) (auto simp: reflcl_over_nhole_mctxt_ta_simp dest: ftranclD2)
lemma rule_states_reflcl_rulesD:
"p |\<in>| rule_states (reflcl_rules \<F> q) \<Longrightarrow> p = q"
by (auto simp: reflcl_rules_def rule_states_def fset_of_list_elem)
lemma rule_states_semantic_path_rulesD:
"p |\<in>| rule_states (semantic_path_rules \<F> q\<^sub>c q\<^sub>i q\<^sub>f) \<Longrightarrow> p = q\<^sub>c \<or> p = q\<^sub>i \<or> p = q\<^sub>f"
by (auto simp: rule_states_def dest!: semantic_path_rules_fmemberD)
(metis in_fset_conv_nth length_list_update length_replicate nth_list_update nth_replicate)
lemma \<Q>_reflcl_over_nhole_mctxt_ta:
"\<Q> (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) |\<subseteq>| Q |\<union>| {|q\<^sub>c, q\<^sub>i, q\<^sub>f|}"
by (auto 0 0 simp: eps_states_def reflcl_over_nhole_mctxt_ta_simp \<Q>_def
dest!: rule_states_reflcl_rulesD rule_states_semantic_path_rulesD)
lemma reflcl_over_nhole_mctxt_ta_vars_term_subset_eq:
assumes "q |\<in>| ta_der (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) t" "q = q\<^sub>f \<or> q = q\<^sub>i"
shows "vars_term t \<subseteq> {q\<^sub>c, q\<^sub>i, q\<^sub>f} \<union> fset Q"
using fresh_states_ta_der'_pres[OF _ _ assms(1)[unfolded ta_der_to_ta_der']] assms(2)
using fsubsetD[OF \<Q>_reflcl_over_nhole_mctxt_ta[of Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f]]
- by auto (meson notin_fset)+
+ by auto
lemma sig_reflcl_over_nhole_mctxt_ta [simp]:
"ta_sig (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) = \<F>"
by (force simp: reflcl_over_nhole_mctxt_ta_simp reflcl_rules_def
dest!: semantic_path_rules_fmemberD intro!: ta_sig_fsubsetI)
lemma reflcl_over_nhole_mctxt_ta_aux_sound:
assumes "funas_term t \<subseteq> fset \<F>" "vars_term t \<subseteq> fset Q"
shows "q\<^sub>c |\<in>| ta_der (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) t" using assms
proof (induct t)
case (Var x)
then show ?case
- by (auto simp: reflcl_over_nhole_mctxt_ta_simp fimage_iff simp flip: fmember_iff_member_fset)
+ by (auto simp: reflcl_over_nhole_mctxt_ta_simp fimage_iff)
(meson finsertI1 finsertI2 fr_into_trancl ftrancl_into_trancl rev_fimage_eqI)
next
case (Fun f ts)
from Fun(2) have "TA_rule f (replicate (length ts) q\<^sub>c) q\<^sub>c |\<in>| rules (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f)"
by (auto simp: reflcl_over_nhole_mctxt_ta_simp reflcl_rules_def fimage_iff fBall_def
- simp flip: fmember_iff_member_fset split: prod.splits)
+ split: prod.splits)
then show ?case using Fun(1)[OF nth_mem] Fun(2-)
by (auto simp: SUP_le_iff) (metis length_replicate nth_replicate)
qed
lemma reflcl_over_nhole_mctxt_ta_sound:
assumes "funas_term t \<subseteq> fset \<F>" "vars_term t \<subseteq> fset Q" "vars_term t \<noteq> {}"
shows "(is_Var t \<longrightarrow> q\<^sub>i |\<in>| ta_der (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) t) \<and>
(is_Fun t \<longrightarrow> q\<^sub>f |\<in>| ta_der (reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f) t)" using assms
proof (induct t)
case (Fun f ts)
let ?A = "reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f"
from Fun(4) obtain i where vars: "i < length ts" "vars_term (ts ! i) \<noteq> {}"
by (metis SUP_le_iff in_set_conv_nth subset_empty term.set(4))
consider (v) "is_Var (ts ! i)" | (f) "is_Fun (ts ! i)" by blast
then show ?case
proof cases
case v
from v Fun(1)[OF nth_mem[OF vars(1)]] have "q\<^sub>i |\<in>| ta_der ?A (ts ! i)"
using vars Fun(2-) by (auto simp: SUP_le_iff)
moreover have "f (replicate (length ts) q\<^sub>c)[i := q\<^sub>i] \<rightarrow> q\<^sub>f |\<in>| rules ?A"
using Fun(2) vars(1)
- by (auto simp: reflcl_over_nhole_mctxt_ta_simp semantic_path_rules_fmember simp flip: fmember_iff_member_fset)
+ by (auto simp: reflcl_over_nhole_mctxt_ta_simp semantic_path_rules_fmember)
moreover have "j < length ts \<Longrightarrow> q\<^sub>c |\<in>| ta_der ?A (ts ! j)" for j using Fun(2-)
by (intro reflcl_over_nhole_mctxt_ta_aux_sound) (auto simp: SUP_le_iff)
ultimately show ?thesis using vars
by auto (metis length_list_update length_replicate nth_list_update nth_replicate)
next
case f
from f Fun(1)[OF nth_mem[OF vars(1)]] have "q\<^sub>f |\<in>| ta_der ?A (ts ! i)"
using vars Fun(2-) by (auto simp: SUP_le_iff)
moreover have "f (replicate (length ts) q\<^sub>c)[i := q\<^sub>f] \<rightarrow> q\<^sub>f |\<in>| rules ?A"
using Fun(2) vars(1)
- by (auto simp: reflcl_over_nhole_mctxt_ta_simp semantic_path_rules_fmember simp flip: fmember_iff_member_fset)
+ by (auto simp: reflcl_over_nhole_mctxt_ta_simp semantic_path_rules_fmember)
moreover have "j < length ts \<Longrightarrow> q\<^sub>c |\<in>| ta_der ?A (ts ! j)" for j using Fun(2-)
by (intro reflcl_over_nhole_mctxt_ta_aux_sound) (auto simp: SUP_le_iff)
ultimately show ?thesis using vars
by auto (metis length_list_update length_replicate nth_list_update nth_replicate)
qed
-qed (auto simp: reflcl_over_nhole_mctxt_ta_simp simp flip: fmember_iff_member_fset dest!: ftranclD2)
+qed (auto simp: reflcl_over_nhole_mctxt_ta_simp dest!: ftranclD2)
lemma gen_nhole_gmctxt_closure_lang:
assumes "q\<^sub>c |\<notin>| \<Q> \<A>" and "q\<^sub>i |\<notin>| \<Q> \<A>" "q\<^sub>f |\<notin>| \<Q> \<A>"
and "q\<^sub>c |\<notin>| Q" "q\<^sub>f \<noteq> q\<^sub>c" "q\<^sub>f \<noteq> q\<^sub>i" "q\<^sub>i \<noteq> q\<^sub>c"
shows "gta_lang {|q\<^sub>f|} (gen_nhole_mctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f) =
{ fill_gholes C ss |
C ss. 0 < num_gholes C \<and> num_gholes C = length ss \<and> C \<noteq> GMHole \<and>
funas_gmctxt C \<subseteq> fset \<F> \<and> (\<forall> i < length ss. ss ! i \<in> gta_lang Q \<A>)}"
(is "?Ls = ?Rs")
proof -
let ?A = "gen_nhole_mctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f" let ?B = "reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f"
interpret sq: derivation_split "?A" "\<A>" "?B"
using assms unfolding derivation_split_def
by (auto simp: gen_nhole_mctxt_closure_automaton_def reflcl_over_nhole_mctxt_ta_def
reflcl_over_nhole_ctxt_ta_def \<Q>_def reflcl_rules_def dest!: semantic_path_rules_rhs)
{fix s assume "s \<in> ?Ls" then obtain u where
seq: "u |\<in>| ta_der' \<A> (term_of_gterm s)" "Var q\<^sub>f |\<in>| ta_der'?B u"
by (auto simp: ta_der_to_ta_der' elim!: gta_langE dest!: sq.ta_der'_split)
note der = seq(2)[unfolded ta_der_to_ta_der'[symmetric]]
have "vars_term u \<subseteq> fset Q" "vars_term u \<noteq> {}"
using ta_der'_gterm_states[OF seq(1)] assms(1 - 3)
using reflcl_over_nhole_mctxt_ta_vars_term[OF der assms(4) assms(5) assms(5) assms(7)]
using reflcl_over_nhole_mctxt_ta_vars_term_subset_eq[OF der]
- by (metis Un_insert_left insert_is_Un notin_fset subset_iff subset_insert)+
+ by (metis Un_insert_left insert_is_Un subset_iff subset_insert)+
then have vars: "\<not> ground u" "i < length (ta_der'_target_args u) \<Longrightarrow> ta_der'_target_args u ! i |\<in>| Q" for i
by (auto simp: ta_der'_target_args_def split_vars_vars_term_list
- fmember_iff_member_fset set_list_subset_nth_conv simp flip: set_vars_term_list)
+ set_list_subset_nth_conv simp flip: set_vars_term_list)
have hole: "ta_der'_target_mctxt u \<noteq> MHole" using vars assms(3-)
using reflcl_over_nhole_mctxt_ta_Fun[OF der]
using ta_der'_mctxt_structure(1, 3)[OF seq(1)]
by auto (metis fill_holes_MHole gterm_ta_der_states length_map lessI map_nth_eq_conv seq(1) ta_der_to_ta_der' term.disc(1))
let ?w = "\<lambda> i. ta_der'_source_args u (term_of_gterm s) ! i"
have "s \<in> ?Rs" using seq(1) ta_der'_Var_funas[OF seq(2)]
using ground_ta_der_statesD[of "?w i" "ta_der'_target_args u ! i" \<A> for i] assms vars
using ta_der'_ground_mctxt_structure[OF seq(1)] hole
by (force simp: ground_gmctxt_of_mctxt_fill_holes' ta_der'_source_args_term_of_gterm
intro!: exI[of _ "gmctxt_of_mctxt (ta_der'_target_mctxt u)"]
exI[of _ "map gterm_of_term (ta_der'_source_args u (term_of_gterm s))"]
gta_langI[of "ta_der'_target_args u ! i" Q \<A>
"gterm_of_term (?w i)" for i])}
then have ls: "?Ls \<subseteq> ?Rs" by blast
{fix t assume "t \<in> ?Rs"
then obtain C ss where len: "0 < num_gholes C" "num_gholes C = length ss" "C \<noteq> GMHole" and
gr_fun: "funas_gmctxt C \<subseteq> fset \<F>" and
reachA: "\<forall> i < length ss. ss ! i \<in> gta_lang Q \<A>" and
const: "t = fill_gholes C ss" by auto
from reachA obtain qs where states: "length ss = length qs" "\<forall> i < length qs. qs ! i |\<in>| Q |\<inter>| \<Q> \<A>"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> ((map term_of_gterm ss) ! i)"
using Ex_list_of_length_P[of "length ss" "\<lambda> q i. q |\<in>| ta_der \<A> (term_of_gterm (ss ! i)) \<and> q |\<in>| Q"]
by (metis (full_types) finterI gta_langE gterm_ta_der_states length_map map_nth_eq_conv)
have [simp]: "is_Fun (fill_holes (mctxt_of_gmctxt C) (map Var qs)) \<longleftrightarrow> True"
"is_Var (fill_holes (mctxt_of_gmctxt C) (map Var qs)) \<longleftrightarrow> False"
using len(3) by (cases C, auto)+
have "q\<^sub>f |\<in>| ta_der ?A (fill_holes (mctxt_of_gmctxt C) (map term_of_gterm ss))"
using reachA len gr_fun states
using reflcl_over_nhole_mctxt_ta_sound[of "fill_holes (mctxt_of_gmctxt C) (map Var qs)"]
by (intro sq.mctxt_const_to_ta_der[of "mctxt_of_gmctxt C" "map term_of_gterm ss" qs q\<^sub>f])
- (auto simp: funas_mctxt_of_gmctxt_conv fmember_iff_member_fset set_list_subset_eq_nth_conv
- simp flip: fmember_iff_member_fset dest!: in_set_idx)
+ (auto simp: funas_mctxt_of_gmctxt_conv set_list_subset_eq_nth_conv
+ dest!: in_set_idx)
then have "t \<in> ?Ls" unfolding const
by (simp add: fill_holes_mctxt_of_gmctxt_to_fill_gholes gta_langI len)}
then show ?thesis using ls by blast
qed
lemma nhole_gmctxt_closure_lang:
"\<L> (nhole_mctxt_closure_reg \<F> \<A>) =
{ fill_gholes C ss | C ss. num_gholes C = length ss \<and> 0 < num_gholes C \<and> C \<noteq> GMHole \<and>
funas_gmctxt C \<subseteq> fset \<F> \<and> (\<forall> i < length ss. ss ! i \<in> \<L> \<A>)}"
(is "?Ls = ?Rs")
proof -
let ?B = "fmap_states_reg (Inl :: 'b \<Rightarrow> 'b + cl_states) (reg_Restr_Q\<^sub>f \<A>)"
have ts: "Inr cl_state |\<notin>| \<Q>\<^sub>r ?B" "Inr tr_state |\<notin>| \<Q>\<^sub>r ?B" "Inr fin_state |\<notin>| \<Q>\<^sub>r ?B"
"Inr cl_state |\<notin>| fin ?B"
by (auto simp: fmap_states_reg_def)
have [simp]: "gta_lang (fin (fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>))) (ta (fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>)))
= gta_lang (fin \<A>) (ta \<A>)"
by (metis \<L>_def \<L>_fmap_states_reg_Inl_Inr(1) reg_Rest_fin_states)
from gen_nhole_gmctxt_closure_lang[OF ts] show ?thesis
by (auto simp add: nhole_mctxt_closure_reg_def gen_nhole_mctxt_closure_reg_def Let_def \<L>_def)
qed
subsubsection \<open>Correctness of @{const gen_mctxt_closure_reg} and @{const mctxt_closure_reg}\<close>
lemma gen_gmctxt_closure_lang:
assumes "q\<^sub>c |\<notin>| \<Q> \<A>" and "q\<^sub>i |\<notin>| \<Q> \<A>" "q\<^sub>f |\<notin>| \<Q> \<A>"
and disj: "q\<^sub>c |\<notin>| Q" "q\<^sub>f \<noteq> q\<^sub>c" "q\<^sub>f \<noteq> q\<^sub>i" "q\<^sub>i \<noteq> q\<^sub>c"
shows "gta_lang {|q\<^sub>f, q\<^sub>i|} (gen_nhole_mctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f) =
{ fill_gholes C ss |
C ss. 0 < num_gholes C \<and> num_gholes C = length ss \<and>
funas_gmctxt C \<subseteq> fset \<F> \<and> (\<forall> i < length ss. ss ! i \<in> gta_lang Q \<A>)}"
(is "?Ls = ?Rs")
proof -
let ?A = "gen_nhole_mctxt_closure_automaton Q \<F> \<A> q\<^sub>c q\<^sub>i q\<^sub>f" let ?B = "reflcl_over_nhole_mctxt_ta Q \<F> q\<^sub>c q\<^sub>i q\<^sub>f"
interpret sq: derivation_split "?A" "\<A>" "?B"
using assms unfolding derivation_split_def
by (auto simp: gen_nhole_mctxt_closure_automaton_def reflcl_over_nhole_mctxt_ta_def
reflcl_over_nhole_ctxt_ta_def \<Q>_def reflcl_rules_def dest!: semantic_path_rules_rhs)
{fix s assume "s \<in> ?Ls" then obtain u q where
seq: "u |\<in>| ta_der' \<A> (term_of_gterm s)" "Var q |\<in>| ta_der'?B u" "q = q\<^sub>f \<or> q = q\<^sub>i"
by (auto simp: ta_der_to_ta_der' elim!: gta_langE dest!: sq.ta_der'_split)
have "vars_term u \<subseteq> fset Q" "vars_term u \<noteq> {}"
using ta_der'_gterm_states[OF seq(1)] assms seq(3)
using reflcl_over_nhole_mctxt_ta_vars_term[OF seq(2)[unfolded ta_der_to_ta_der'[symmetric]] disj(1) _ disj(2, 4)]
using reflcl_over_nhole_mctxt_ta_vars_term_subset_eq[OF seq(2)[unfolded ta_der_to_ta_der'[symmetric]] seq(3)]
- by (metis Un_insert_left notin_fset subsetD subset_insert sup_bot_left)+
+ by (metis Un_insert_left subsetD subset_insert sup_bot_left)+
then have vars: "\<not> ground u" "i < length (ta_der'_target_args u) \<Longrightarrow> ta_der'_target_args u ! i |\<in>| Q" for i
by (auto simp: ta_der'_target_args_def split_vars_vars_term_list
- fmember_iff_member_fset set_list_subset_nth_conv simp flip: set_vars_term_list)
+ set_list_subset_nth_conv simp flip: set_vars_term_list)
let ?w = "\<lambda> i. ta_der'_source_args u (term_of_gterm s) ! i"
have "s \<in> ?Rs" using seq(1) ta_der'_Var_funas[OF seq(2)]
using ground_ta_der_statesD[of "?w i" "ta_der'_target_args u ! i" \<A> for i] assms vars
using ta_der'_ground_mctxt_structure[OF seq(1)]
by (force simp: ground_gmctxt_of_mctxt_fill_holes' ta_der'_source_args_term_of_gterm
intro!: exI[of _ "gmctxt_of_mctxt (ta_der'_target_mctxt u)"]
exI[of _ "map gterm_of_term (ta_der'_source_args u (term_of_gterm s))"]
gta_langI[of "ta_der'_target_args u ! i" Q \<A>
"gterm_of_term (?w i)" for i])}
then have "?Ls \<subseteq> ?Rs" by blast
moreover
{fix t assume "t \<in> ?Rs"
then obtain C ss where len: "0 < num_gholes C" "num_gholes C = length ss" and
gr_fun: "funas_gmctxt C \<subseteq> fset \<F>" and
reachA: "\<forall> i < length ss. ss ! i \<in> gta_lang Q \<A>" and
const: "t = fill_gholes C ss" by auto
from const have const': "term_of_gterm t = fill_holes (mctxt_of_gmctxt C) (map term_of_gterm ss)"
by (simp add: fill_holes_mctxt_of_gmctxt_to_fill_gholes len(2))
from reachA obtain qs where states: "length ss = length qs" "\<forall> i < length qs. qs ! i |\<in>| Q |\<inter>| \<Q> \<A>"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> ((map term_of_gterm ss) ! i)"
using Ex_list_of_length_P[of "length ss" "\<lambda> q i. q |\<in>| ta_der \<A> (term_of_gterm (ss ! i)) \<and> q |\<in>| Q"]
by (metis (full_types) finterI gta_langE gterm_ta_der_states length_map map_nth_eq_conv)
have "C = GMHole \<Longrightarrow> is_Var (fill_holes (mctxt_of_gmctxt C) (map Var qs)) = True" using len states(1)
by (metis fill_holes_MHole length_map mctxt_of_gmctxt.simps(1) nth_map num_gholes.simps(1) term.disc(1))
then have hole: "C = GMHole \<Longrightarrow> q\<^sub>i |\<in>| ta_der ?A (fill_holes (mctxt_of_gmctxt C) (map term_of_gterm ss))"
using reachA len gr_fun states len
using reflcl_over_nhole_mctxt_ta_sound[of "fill_holes (mctxt_of_gmctxt C) (map Var qs)"]
by (intro sq.mctxt_const_to_ta_der[of "mctxt_of_gmctxt C" "map term_of_gterm ss" qs q\<^sub>i])
- (auto simp: funas_mctxt_of_gmctxt_conv fmember_iff_member_fset set_list_subset_eq_nth_conv
- simp flip: fmember_iff_member_fset dest!: in_set_idx)
+ (auto simp: funas_mctxt_of_gmctxt_conv set_list_subset_eq_nth_conv
+ dest!: in_set_idx)
have "C \<noteq> GMHole \<Longrightarrow> is_Fun (fill_holes (mctxt_of_gmctxt C) (map Var qs)) = True"
by (cases C) auto
then have "C \<noteq> GMHole \<Longrightarrow> q\<^sub>f |\<in>| ta_der ?A (fill_holes (mctxt_of_gmctxt C) (map term_of_gterm ss))"
using reachA len gr_fun states
using reflcl_over_nhole_mctxt_ta_sound[of "fill_holes (mctxt_of_gmctxt C) (map Var qs)"]
by (intro sq.mctxt_const_to_ta_der[of "mctxt_of_gmctxt C" "map term_of_gterm ss" qs q\<^sub>f])
- (auto simp: funas_mctxt_of_gmctxt_conv fmember_iff_member_fset set_list_subset_eq_nth_conv
- simp flip: fmember_iff_member_fset dest!: in_set_idx)
+ (auto simp: funas_mctxt_of_gmctxt_conv set_list_subset_eq_nth_conv
+ dest!: in_set_idx)
then have "t \<in> ?Ls" using hole const' unfolding gta_lang_def gta_der_def
by (metis (mono_tags, lifting) fempty_iff finsert_iff finterI mem_Collect_eq)}
ultimately show ?thesis
by (meson subsetI subset_antisym)
qed
lemma gmctxt_closure_lang:
"\<L> (mctxt_closure_reg \<F> \<A>) =
{ fill_gholes C ss | C ss. num_gholes C = length ss \<and> 0 < num_gholes C \<and>
funas_gmctxt C \<subseteq> fset \<F> \<and> (\<forall> i < length ss. ss ! i \<in> \<L> \<A>)}"
(is "?Ls = ?Rs")
proof -
let ?B = "fmap_states_reg (Inl :: 'b \<Rightarrow> 'b + cl_states) (reg_Restr_Q\<^sub>f \<A>)"
have ts: "Inr cl_state |\<notin>| \<Q>\<^sub>r ?B" "Inr tr_state |\<notin>| \<Q>\<^sub>r ?B" "Inr fin_state |\<notin>| \<Q>\<^sub>r ?B"
"Inr cl_state |\<notin>| fin ?B"
by (auto simp: fmap_states_reg_def)
have [simp]: "gta_lang (fin (fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>))) (ta (fmap_states_reg Inl (reg_Restr_Q\<^sub>f \<A>)))
= gta_lang (fin \<A>) (ta \<A>)"
by (metis \<L>_def \<L>_fmap_states_reg_Inl_Inr(1) reg_Rest_fin_states)
from gen_gmctxt_closure_lang[OF ts] show ?thesis
by (auto simp add: mctxt_closure_reg_def gen_mctxt_closure_reg_def Let_def \<L>_def)
qed
subsubsection \<open>Correctness of @{const nhole_mctxt_reflcl_reg}\<close>
lemma nhole_mctxt_reflcl_lang:
"\<L> (nhole_mctxt_reflcl_reg \<F> \<A>) = \<L> (nhole_mctxt_closure_reg \<F> \<A>) \<union> \<T>\<^sub>G (fset \<F>)"
proof -
let ?refl = "Reg {|fin_clstate|} (refl_ta \<F> (fin_clstate))"
{fix t assume "t \<in> \<L> ?refl" then have "t \<in> \<T>\<^sub>G (fset \<F>)"
using reg_funas by fastforce}
moreover
{fix t assume "t \<in> \<T>\<^sub>G (fset \<F>)" then have "t \<in> \<L> ?refl"
by (simp add: \<L>_def gta_langI refl_ta_sound)}
ultimately have *: "\<L> ?refl = \<T>\<^sub>G (fset \<F>)"
by blast
show ?thesis unfolding nhole_mctxt_reflcl_reg_def \<L>_union * by simp
qed
declare ta_union_def [simp del]
end
\ No newline at end of file
diff --git a/thys/FO_Theory_Rewriting/FOR_Check.thy b/thys/FO_Theory_Rewriting/FOR_Check.thy
--- a/thys/FO_Theory_Rewriting/FOR_Check.thy
+++ b/thys/FO_Theory_Rewriting/FOR_Check.thy
@@ -1,1193 +1,1193 @@
theory FOR_Check
imports
FOR_Semantics
FOL_Extra
GTT_RRn
First_Order_Terms.Option_Monad
LV_to_GTT
NF
Regular_Tree_Relations.GTT_Transitive_Closure
Regular_Tree_Relations.AGTT
Regular_Tree_Relations.RR2_Infinite_Q_infinity
Regular_Tree_Relations.RRn_Automata
begin
section \<open>Check inference steps\<close>
type_synonym ('f, 'v) fin_trs = "('f, 'v) rule fset"
lemma tl_drop_conv:
"tl xs = drop 1 xs"
by (induct xs) auto
definition rrn_drop_fst where
"rrn_drop_fst \<A> = relabel_reg (trim_reg (collapse_automaton_reg (fmap_funs_reg (drop_none_rule 1) (trim_reg \<A>))))"
lemma rrn_drop_fst_lang:
assumes "RRn_spec n A T" "1 < n"
shows "RRn_spec (n - 1) (rrn_drop_fst A) (drop 1 ` T)"
using drop_automaton_reg[OF _ assms(2), of "trim_reg A" T] assms(1)
unfolding rrn_drop_fst_def
by (auto simp: trim_ta_reach)
definition liftO1 :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a option \<Rightarrow> 'b option" where
"liftO1 = map_option"
definition liftO2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option" where
"liftO2 f a b = case_option None (\<lambda>a'. liftO1 (f a') b) a"
lemma liftO1_Some [simp]:
"liftO1 f x = Some y \<longleftrightarrow> (\<exists>x'. x = Some x') \<and> y = f (the x)"
by (cases x) (auto simp: liftO1_def)
lemma liftO2_Some [simp]:
"liftO2 f x y = Some z \<longleftrightarrow> (\<exists>x' y'. x = Some x' \<and> y = Some y') \<and> z = f (the x) (the y)"
by (cases x; cases y) (auto simp: liftO2_def)
subsection \<open>Computing TRSs\<close>
lemma is_to_trs_props:
assumes "\<forall> R \<in> set Rs. finite R \<and> lv_trs R \<and> funas_trs R \<subseteq> \<F>" "\<forall>i \<in> set is. case_ftrs id id i < length Rs"
shows "funas_trs (is_to_trs Rs is) \<subseteq> \<F>" "lv_trs (is_to_trs Rs is)" "finite (is_to_trs Rs is)"
proof (goal_cases \<F> lv fin)
case \<F> show ?case using assms nth_mem
apply (auto simp: is_to_trs_def funas_trs_def case_prod_beta split: ftrs.splits)
apply fastforce
apply (metis (no_types, lifting) assms(1) in_mono rhs_wf)
apply (metis (no_types, lifting) assms(1) in_mono rhs_wf)
by (smt (z3) UN_subset_iff fst_conv in_mono le_sup_iff)
qed (insert assms, (fastforce simp: is_to_trs_def funas_trs_def lv_trs_def split: ftrs.splits)+)
definition is_to_fin_trs :: "('f, 'v) fin_trs list \<Rightarrow> ftrs list \<Rightarrow> ('f, 'v) fin_trs" where
"is_to_fin_trs Rs is = |\<Union>| (fset_of_list (map (case_ftrs ((!) Rs) ((|`|) prod.swap \<circ> (!) Rs)) is))"
lemma is_to_fin_trs_conv:
assumes "\<forall>i \<in> set is. case_ftrs id id i < length Rs"
shows "is_to_trs (map fset Rs) is = fset (is_to_fin_trs Rs is)"
using assms unfolding is_to_trs_def is_to_fin_trs_def
by (auto simp: ffUnion.rep_eq fset_of_list.rep_eq split: ftrs.splits)
definition is_to_trs' :: "('f, 'v) fin_trs list \<Rightarrow> ftrs list \<Rightarrow> ('f, 'v) fin_trs option" where
"is_to_trs' Rs is = do {
guard (\<forall>i \<in> set is. case_ftrs id id i < length Rs);
Some (is_to_fin_trs Rs is)
}"
lemma is_to_trs_conv:
"is_to_trs' Rs is = Some S \<Longrightarrow> is_to_trs (map fset Rs) is = fset S"
using is_to_fin_trs_conv unfolding is_to_trs'_def
by (auto simp add: guard_simps split: bind_splits)
lemma is_to_trs'_props:
assumes "\<forall> R \<in> set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>" and "is_to_trs' Rs is = Some S"
shows "ffunas_trs S |\<subseteq>| \<F>" "lv_trs (fset S)"
proof -
from assms(2) have well: "\<forall>i \<in> set is. case_ftrs id id i < length Rs" "is_to_fin_trs Rs is = S"
unfolding is_to_trs'_def
by (auto simp add: guard_simps split: bind_splits)
have "\<forall> R \<in> set Rs. finite (fset R) \<and> lv_trs (fset R) \<and> funas_trs (fset R) \<subseteq> (fset \<F>)"
using assms(1) by (auto simp: ffunas_trs.rep_eq less_eq_fset.rep_eq)
from is_to_trs_props[of "map fset Rs" "fset \<F>" "is"] this well(1)
have "lv_trs (is_to_trs (map fset Rs) is)" "funas_trs (is_to_trs (map fset Rs) is) \<subseteq> fset \<F>"
by auto
then show "lv_trs (fset S)" "ffunas_trs S |\<subseteq>| \<F>"
using is_to_fin_trs_conv[OF well(1)] unfolding well(2)
by (auto simp: ffunas_trs.rep_eq less_eq_fset.rep_eq)
qed
subsection \<open>Computing GTTs\<close>
fun gtt_of_gtt_rel :: "('f \<times> nat) fset \<Rightarrow> ('f :: linorder, 'v) fin_trs list \<Rightarrow> ftrs gtt_rel \<Rightarrow> (nat, 'f) gtt option" where
"gtt_of_gtt_rel \<F> Rs (ARoot is) = liftO1 (\<lambda>R. relabel_gtt (agtt_grrstep R \<F>)) (is_to_trs' Rs is)"
| "gtt_of_gtt_rel \<F> Rs (GInv g) = liftO1 prod.swap (gtt_of_gtt_rel \<F> Rs g)"
| "gtt_of_gtt_rel \<F> Rs (AUnion g1 g2) = liftO2 (\<lambda>g1 g2. relabel_gtt (AGTT_union' g1 g2)) (gtt_of_gtt_rel \<F> Rs g1) (gtt_of_gtt_rel \<F> Rs g2)"
| "gtt_of_gtt_rel \<F> Rs (ATrancl g) = liftO1 (relabel_gtt \<circ> AGTT_trancl) (gtt_of_gtt_rel \<F> Rs g)"
| "gtt_of_gtt_rel \<F> Rs (GTrancl g) = liftO1 GTT_trancl (gtt_of_gtt_rel \<F> Rs g)"
| "gtt_of_gtt_rel \<F> Rs (AComp g1 g2) = liftO2 (\<lambda>g1 g2. relabel_gtt (AGTT_comp' g1 g2)) (gtt_of_gtt_rel \<F> Rs g1) (gtt_of_gtt_rel \<F> Rs g2)"
| "gtt_of_gtt_rel \<F> Rs (GComp g1 g2) = liftO2 (\<lambda>g1 g2. relabel_gtt (GTT_comp' g1 g2)) (gtt_of_gtt_rel \<F> Rs g1) (gtt_of_gtt_rel \<F> Rs g2)"
lemma gtt_of_gtt_rel_correct:
assumes "\<forall>R \<in> set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>"
shows "gtt_of_gtt_rel \<F> Rs g = Some g' \<Longrightarrow> agtt_lang g' = eval_gtt_rel (fset \<F>) (map fset Rs) g"
proof (induct g arbitrary: g')
note [simp] = bind_eq_Some_conv guard_simps
have proj_sq: "fst ` (X \<times> X) = X" "snd ` (X \<times> X) = X" for X by auto
{
case (ARoot "is")
then obtain w where w:"is_to_trs' Rs is = Some w" by auto
then show ?case using ARoot is_to_trs'_props[OF assms w] is_to_trs_conv[OF w]
using agtt_grrstep
by auto
next
case (GInv g) then show ?case by (simp add: agtt_lang_swap gtt_states_def)
next
case (AUnion g1 g2)
from AUnion(3)[simplified, THEN conjunct1] AUnion(3)[simplified, THEN conjunct2, THEN conjunct1]
obtain w1 w2 where
[simp]: "gtt_of_gtt_rel \<F> Rs g1 = Some w1" "gtt_of_gtt_rel \<F> Rs g2 = Some w2"
by blast
then show ?case using AUnion(3)
by (simp add: AGTT_union'_sound AUnion)
next
case (ATrancl g)
from ATrancl[simplified] obtain w1 where
[simp]: "gtt_of_gtt_rel \<F> Rs g = Some w1" "g' = relabel_gtt (AGTT_trancl w1)" by auto
then have fin_lang: "eval_gtt_rel (fset \<F>) (map fset Rs) g = agtt_lang w1"
using ATrancl by auto
from fin_lang show ?case using AGTT_trancl_sound[of w1]
by auto
next
case (GTrancl g) note * = GTrancl(2)[simplified, THEN conjunct2]
show ?case unfolding gtt_of_gtt_rel.simps GTT_trancl_alang * gtrancl_rel_def eval_gtt_rel.simps gmctxt_cl_gmctxtex_onp_conv
proof ((intro conjI equalityI subrelI; (elim relcompE)?), goal_cases LR RL)
case (LR _ _ s _ z s' t' t)
show ?case using lift_root_steps_sig_transfer'[OF LR(2)[folded lift_root_step.simps], of "fset \<F>"]
lift_root_steps_sig_transfer[OF LR(5)[folded lift_root_step.simps], of "fset \<F>"]
image_mono[OF eval_gtt_rel_sig[of "fset \<F>" "map fset Rs" g], of fst, unfolded proj_sq]
image_mono[OF eval_gtt_rel_sig[of "fset \<F>" "map fset Rs" g], of snd, unfolded proj_sq]
subsetD[OF eval_gtt_rel_sig[of "fset \<F>" "map fset Rs" g]] LR(1, 3, 4) GTrancl
by (intro relcompI[OF _ relcompI, of _ s' _ t' _])
(auto simp: \<T>\<^sub>G_funas_gterm_conv lift_root_step.simps)
next
case (RL _ _ s _ z s' t' t)
then show ?case using GTrancl
lift_root_step_mono[of "fset \<F>" UNIV PAny ESingle "eval_gtt_rel (fset \<F>) (map fset Rs) g", THEN rtrancl_mono]
unfolding lift_root_step.simps[symmetric]
by (intro relcompI[OF _ relcompI, of _ s' _ t' _])
(auto simp: \<T>\<^sub>G_funas_gterm_conv lift_root_step_mono trancl_mono)
qed
next
case (AComp g1 g2)
from AComp[simplified] obtain w1 w2 where
[simp]: "gtt_of_gtt_rel \<F> Rs g1 = Some w1" "gtt_of_gtt_rel \<F> Rs g2 = Some w2"
"g' = relabel_gtt (AGTT_comp' w1 w2)" by auto
then have fin_lang: "eval_gtt_rel (fset \<F>) (map fset Rs) g1 = agtt_lang w1"
"eval_gtt_rel (fset \<F>) (map fset Rs) g2 = agtt_lang w2"
using AComp by auto
from fin_lang AGTT_comp'_sound[of w1 w2]
show ?case by simp
next
case (GComp g1 g2)
let ?r = "\<lambda> g. eval_gtt_rel (fset \<F>) (map fset Rs) g"
have *: "gmctxtex_onp (\<lambda>C. True) (?r g1) = lift_root_step UNIV PAny EParallel (?r g1)"
"gmctxtex_onp (\<lambda>C. True) (?r g2) = lift_root_step UNIV PAny EParallel (?r g2)"
by (auto simp: lift_root_step.simps)
show ?case using GComp(3)
apply (intro conjI equalityI subrelI; simp add: gmctxt_cl_gmctxtex_onp_conv GComp(1,2) gtt_comp'_alang gcomp_rel_def * flip: lift_root_step.simps; elim conjE disjE exE relcompE)
subgoal for s t _ _ _ _ _ u
using image_mono[OF eval_gtt_rel_sig, of snd "fset \<F>" "map fset Rs", unfolded proj_sq]
apply (subst relcompI[of _ u "eval_gtt_rel _ _ g1", OF _ lift_root_step_sig_transfer[of _ UNIV PAny EParallel "_ g2" "fset \<F>"]])
apply (force simp add: subsetI \<T>\<^sub>G_equivalent_def)+
done
subgoal for s t _ _ _ _ _ u
using image_mono[OF eval_gtt_rel_sig, of fst "fset \<F>" "map fset Rs", unfolded proj_sq]
apply (subst relcompI[of _ u _ _ "eval_gtt_rel _ _ g2", OF lift_root_step_sig_transfer'[of _ UNIV PAny EParallel "_ g1" "fset \<F>"]])
apply (force simp add: subsetI \<T>\<^sub>G_equivalent_def)+
done
by (auto intro: subsetD[OF lift_root_step_mono[of "fset \<F>" UNIV]])
}
qed
subsection \<open>Computing RR1 and RR2 relations\<close>
definition "simplify_reg \<A> = (relabel_reg (trim_reg \<A>))"
lemma \<L>_simplify_reg [simp]: "\<L> (simplify_reg \<A>) = \<L> \<A>"
by (simp add: simplify_reg_def \<L>_trim)
lemma RR1_spec_simplify_reg[simp]:
"RR1_spec (simplify_reg \<A>) R = RR1_spec \<A> R"
by (auto simp: RR1_spec_def)
lemma RR2_spec_simplify_reg[simp]:
"RR2_spec (simplify_reg \<A>) R = RR2_spec \<A> R"
by (auto simp: RR2_spec_def)
lemma RRn_spec_simplify_reg[simp]:
"RRn_spec n (simplify_reg \<A>) R = RRn_spec n \<A> R"
by (auto simp: RRn_spec_def)
lemma RR1_spec_eps_free_reg[simp]:
"RR1_spec (eps_free_reg \<A>) R = RR1_spec \<A> R"
by (auto simp: RR1_spec_def \<L>_eps_free)
lemma RR2_spec_eps_free_reg[simp]:
"RR2_spec (eps_free_reg \<A>) R = RR2_spec \<A> R"
by (auto simp: RR2_spec_def \<L>_eps_free)
lemma RRn_spec_eps_free_reg[simp]:
"RRn_spec n (eps_free_reg \<A>) R = RRn_spec n \<A> R"
by (auto simp: RRn_spec_def \<L>_eps_free)
fun rr1_of_rr1_rel :: "('f \<times> nat) fset \<Rightarrow> ('f :: linorder, 'v) fin_trs list \<Rightarrow> ftrs rr1_rel \<Rightarrow> (nat, 'f) reg option"
and rr2_of_rr2_rel :: "('f \<times> nat) fset \<Rightarrow> ('f, 'v) fin_trs list \<Rightarrow> ftrs rr2_rel \<Rightarrow> (nat, 'f option \<times> 'f option) reg option" where
"rr1_of_rr1_rel \<F> Rs R1Terms = Some (relabel_reg (term_reg \<F>))"
| "rr1_of_rr1_rel \<F> Rs (R1NF is) = liftO1 (\<lambda>R. (simplify_reg (nf_reg (fst |`| R) \<F>))) (is_to_trs' Rs is)"
| "rr1_of_rr1_rel \<F> Rs (R1Inf r) = liftO1 (\<lambda>R.
let \<A> = trim_reg R in
simplify_reg (proj_1_reg (Inf_reg_impl \<A>))
) (rr2_of_rr2_rel \<F> Rs r)"
| "rr1_of_rr1_rel \<F> Rs (R1Proj i r) = (case i of 0 \<Rightarrow>
liftO1 (trim_reg \<circ> proj_1_reg) (rr2_of_rr2_rel \<F> Rs r)
| _ \<Rightarrow> liftO1 (trim_reg \<circ> proj_2_reg) (rr2_of_rr2_rel \<F> Rs r))"
| "rr1_of_rr1_rel \<F> Rs (R1Union s1 s2) =
liftO2 (\<lambda> x y. relabel_reg (reg_union x y)) (rr1_of_rr1_rel \<F> Rs s1) (rr1_of_rr1_rel \<F> Rs s2)"
| "rr1_of_rr1_rel \<F> Rs (R1Inter s1 s2) =
liftO2 (\<lambda> x y. simplify_reg (reg_intersect x y)) (rr1_of_rr1_rel \<F> Rs s1) (rr1_of_rr1_rel \<F> Rs s2)"
| "rr1_of_rr1_rel \<F> Rs (R1Diff s1 s2) = liftO2 (\<lambda> x y. relabel_reg (trim_reg (difference_reg x y))) (rr1_of_rr1_rel \<F> Rs s1) (rr1_of_rr1_rel \<F> Rs s2)"
| "rr2_of_rr2_rel \<F> Rs (R2GTT_Rel g w x) =
(case w of PRoot \<Rightarrow>
(case x of ESingle \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g)
| EParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> reflcl_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g)
| EStrictParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g))
| PNonRoot \<Rightarrow>
(case x of ESingle \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> nhole_ctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g)
| EParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> nhole_mctxt_reflcl_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g)
| EStrictParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> nhole_mctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g))
| PAny \<Rightarrow>
(case x of ESingle \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> ctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g)
| EParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> parallel_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g)
| EStrictParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_reg \<circ> mctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel \<F> Rs g)))"
| "rr2_of_rr2_rel \<F> Rs (R2Diag s) =
liftO1 (\<lambda> x. fmap_funs_reg (\<lambda>f. (Some f, Some f)) x) (rr1_of_rr1_rel \<F> Rs s)"
| "rr2_of_rr2_rel \<F> Rs (R2Prod s1 s2) =
liftO2 (\<lambda> x y. simplify_reg (pair_automaton_reg x y)) (rr1_of_rr1_rel \<F> Rs s1) (rr1_of_rr1_rel \<F> Rs s2)"
| "rr2_of_rr2_rel \<F> Rs (R2Inv r) = liftO1 (fmap_funs_reg prod.swap) (rr2_of_rr2_rel \<F> Rs r)"
| "rr2_of_rr2_rel \<F> Rs (R2Union r1 r2) =
liftO2 (\<lambda> x y. relabel_reg (reg_union x y)) (rr2_of_rr2_rel \<F> Rs r1) (rr2_of_rr2_rel \<F> Rs r2)"
| "rr2_of_rr2_rel \<F> Rs (R2Inter r1 r2) =
liftO2 (\<lambda> x y. simplify_reg (reg_intersect x y)) (rr2_of_rr2_rel \<F> Rs r1) (rr2_of_rr2_rel \<F> Rs r2)"
| "rr2_of_rr2_rel \<F> Rs (R2Diff r1 r2) = liftO2 (\<lambda> x y. simplify_reg (difference_reg x y)) (rr2_of_rr2_rel \<F> Rs r1) (rr2_of_rr2_rel \<F> Rs r2)"
| "rr2_of_rr2_rel \<F> Rs (R2Comp r1 r2) = liftO2 (\<lambda> x y. simplify_reg (rr2_compositon \<F> x y))
(rr2_of_rr2_rel \<F> Rs r1) (rr2_of_rr2_rel \<F> Rs r2)"
abbreviation lhss where
"lhss R \<equiv> fst |`| R"
lemma rr12_of_rr12_rel_correct:
fixes Rs :: "(('f :: linorder, 'v) Term.term \<times> ('f, 'v) Term.term) fset list"
assumes "\<forall>R \<in> set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>"
shows "\<forall>ta1. rr1_of_rr1_rel \<F> Rs r1 = Some ta1 \<longrightarrow> RR1_spec ta1 (eval_rr1_rel (fset \<F>) (map fset Rs) r1)"
"\<forall>ta2. rr2_of_rr2_rel \<F> Rs r2 = Some ta2 \<longrightarrow> RR2_spec ta2 (eval_rr2_rel (fset \<F>) (map fset Rs) r2)"
proof (induct r1 and r2)
note [simp] = bind_eq_Some_conv guard_simps
let ?F = "fset \<F>" let ?Rs = "map fset Rs"
{
case R1Terms
then show ?case using term_automaton[of \<F>]
by (simp add: \<T>\<^sub>G_equivalent_def)
next
case (R1NF r)
consider (a) "\<exists> R. is_to_trs' Rs r = Some R" | (b) "is_to_trs' Rs r = None" by auto
then show ?case
proof (cases)
case a
from a obtain R where [simp]: "is_to_trs' Rs r = Some R" "is_to_fin_trs Rs r = R"
by (auto simp: is_to_trs'_def)
from is_to_trs'_props[OF assms this(1)] have inv: "ffunas_trs R |\<subseteq>| \<F>" "lv_trs (fset R)" .
from inv have fl: "\<forall> l |\<in>| lhss R. linear_term l"
- by (auto simp: lv_trs_def fmember_iff_member_fset split!: prod.splits)
+ by (auto simp: lv_trs_def split!: prod.splits)
{fix s t assume ass: "(s, t) \<in> grstep (fset R)"
then obtain C l r \<sigma> where step: "(l, r) |\<in>| R" "term_of_gterm s = (C :: ('f, 'v) ctxt) \<langle>l \<cdot> \<sigma>\<rangle>" "term_of_gterm t = C\<langle>r \<cdot> \<sigma>\<rangle>"
- unfolding grstep_def by (auto simp: fmember_iff_member_fset dest!: rstep_imp_C_s_r)
+ unfolding grstep_def by (auto simp: dest!: rstep_imp_C_s_r)
from step ta_nf_lang_sound[of l "lhss R" C \<sigma> \<F>]
have "s \<notin> \<L> (nf_reg (lhss R) \<F>)" unfolding \<L>_def
by (metis fimage_eqI fst_conv nf_reg_def reg.sel(1, 2) term_of_gterm_in_ta_lang_conv)}
note mem = this
have funas: "funas_trs (fset R) \<subseteq> ?F" using inv(1)
by (simp add: ffunas_trs.rep_eq less_eq_fset.rep_eq subsetD)
{fix s assume "s \<in> \<L> (nf_reg (lhss R) \<F>)"
then have "s \<in> NF (Restr (grstep (fset R)) (\<T>\<^sub>G (fset \<F>))) \<inter> \<T>\<^sub>G (fset \<F>)"
by (meson IntI NF_I \<T>\<^sub>G_funas_gterm_conv gta_lang_nf_ta_funas inf.cobounded1 mem subset_iff)}
moreover
{fix s assume ass: "s \<in> NF (Restr (grstep (fset R)) (\<T>\<^sub>G (fset \<F>))) \<inter> \<T>\<^sub>G (fset \<F>)"
then have *: "(term_of_gterm s, term_of_gterm t) \<notin> rstep (fset R)" for t using funas
- by (auto simp: funas_trs_def grstep_def NF_iff_no_step \<T>\<^sub>G_funas_gterm_conv fmember_iff_member_fset)
+ by (auto simp: funas_trs_def grstep_def NF_iff_no_step \<T>\<^sub>G_funas_gterm_conv)
(meson R1NF_reps funas rstep.cases)
then have "s \<in> \<L> (nf_reg (lhss R) \<F>)" using fl ass
using ta_nf_\<L>_complete[OF fl, of _ \<F>] gta_lang_nf_ta_funas[of _ "lhss R" \<F>]
- by (smt (verit, ccfv_SIG) IntE R1NF_reps \<T>\<^sub>G_sound fimageE funas notin_fset surjective_pairing)}
+ by (smt (verit, ccfv_SIG) IntE R1NF_reps \<T>\<^sub>G_sound fimageE funas surjective_pairing)}
ultimately have "\<L> (nf_reg (lhss R) \<F>) = NF (Restr (grstep (fset R)) (\<T>\<^sub>G (fset \<F>))) \<inter> \<T>\<^sub>G (fset \<F>)"
by blast
then show ?thesis using fl(1)
by (simp add: RR1_spec_def is_to_trs_conv)
qed auto
next
case (R1Inf r)
consider (a) "\<exists> A. rr2_of_rr2_rel \<F> Rs r = Some A" | (b) " rr2_of_rr2_rel \<F> Rs r = None" by auto
then show ?case
proof cases
case a
have [simp]: "{u. (t, u) \<in> eval_rr2_rel ?F ?Rs r \<and> funas_gterm u \<subseteq> ?F} =
{u. (t, u) \<in> eval_rr2_rel ?F ?Rs r}" for t
using eval_rr12_rel_sig(2)[of ?F ?Rs r] by (auto simp: \<T>\<^sub>G_equivalent_def)
have [simp]: "infinite {u. (t, u) \<in> eval_rr2_rel ?F ?Rs r} \<Longrightarrow> funas_gterm t \<subseteq> ?F" for t
using eval_rr12_rel_sig(2)[of ?F ?Rs r] not_finite_existsD by (fastforce simp: \<T>\<^sub>G_equivalent_def)
from a obtain A where [simp]: "rr2_of_rr2_rel \<F> Rs r = Some A" by blast
from R1Inf this have spec: "RR2_spec A (eval_rr2_rel ?F ?Rs r)" by auto
then have spec_trim: "RR2_spec (trim_reg A) (eval_rr2_rel ?F ?Rs r)" by auto
let ?B = "(Inf_reg (trim_reg A) (Q_infty (ta (trim_reg A)) \<F>))"
have B: "RR2_spec ?B {(s, t) | s t. gpair s t \<in> \<L> ?B}"
using subset_trans[OF Inf_automata_subseteq[of "trim_reg A" \<F>], of "\<L> A"] spec
by (auto simp: RR2_spec_def \<L>_trim)
have *: "\<L> (Inf_reg_impl (trim_reg A)) = \<L> ?B" using spec
using eval_rr12_rel_sig(2)[of ?F ?Rs r]
by (intro Inf_reg_impl_sound) (auto simp: \<L>_trim RR2_spec_def \<T>\<^sub>G_equivalent_def)
then have **: "RR2_spec (Inf_reg_impl (trim_reg A)) {(s, t) | s t. gpair s t \<in> \<L> ?B}" using B
by (auto simp: RR2_spec_def)
show ?thesis
using spec eval_rr12_rel_sig(2)[of ?F ?Rs r]
using \<L>_Inf_reg[OF spec_trim, of \<F>]
by (auto simp: \<T>\<^sub>G_equivalent_def * RR1_spec_def \<L>_trim \<L>_proj(1)[OF **]
Inf_branching_terms_def fImage_singleton)
(metis (no_types, lifting) SigmaD1 in_mono mem_Collect_eq not_finite_existsD)
qed auto
next
case (R1Proj i r)
then show ?case
proof (cases i)
case [simp]:0 show ?thesis using R1Proj
using proj_automaton_gta_lang(1)[of "the (rr2_of_rr2_rel \<F> Rs r)" "eval_rr2_rel ?F ?Rs r"]
by simp
next
case (Suc nat) then show ?thesis using R1Proj
using proj_automaton_gta_lang(2)[of "the (rr2_of_rr2_rel \<F> Rs r)" "eval_rr2_rel ?F ?Rs r"]
by simp
qed
next
case (R1Union s1 s2)
then show ?case
by (auto simp: RR1_spec_def \<L>_union)
next
case (R1Inter s1 s2)
from R1Inter show ?case
by (auto simp: \<L>_intersect RR1_spec_def)
next
case (R1Diff s1 s2)
then show ?case
by (auto intro: RR1_difference)
next
case (R2GTT_Rel g w x)
note ass = R2GTT_Rel
consider (a) "\<exists> A. gtt_of_gtt_rel \<F> Rs g = Some A" | (b) "gtt_of_gtt_rel \<F> Rs g = None" by blast
then show ?case
proof cases
case a then obtain A where [simp]: "gtt_of_gtt_rel \<F> Rs g = Some A" by blast
from gtt_of_gtt_rel_correct[OF assms this]
have spec [simp]: "eval_gtt_rel ?F ?Rs g = agtt_lang A" by auto
let ?B = "GTT_to_RR2_root_reg A" note [simp] = GTT_to_RR2_root[of A]
show ?thesis
proof (cases w)
case [simp]: PRoot show ?thesis
proof (cases x)
case EParallel
then show ?thesis using reflcl_automaton[of ?B "agtt_lang A" \<F>]
by auto
qed (auto simp: GTT_to_RR2_root)
next
case PNonRoot
then show ?thesis
using nhole_ctxt_closure_automaton[of ?B "agtt_lang A" \<F>]
using nhole_mctxt_reflcl_automaton[of ?B "agtt_lang A" \<F>]
using nhole_mctxt_closure_automaton[of ?B "agtt_lang A" \<F>]
by (cases x) auto
next
case PAny
then show ?thesis
using ctxt_closure_automaton[of ?B "agtt_lang A" \<F>]
using parallel_closure_automaton[of ?B "agtt_lang A" \<F>]
using mctxt_closure_automaton[of ?B "agtt_lang A" \<F>]
by (cases x) auto
qed
qed (cases w; cases x, auto)
next
case (R2Diag s)
then show ?case
by (auto simp: RR2_spec_def RR1_spec_def fmap_funs_\<L> Id_on_iff
fmap_funs_gta_lang map_funs_term_some_gpair)
next
case (R2Prod s1 s2)
then show ?case using pair_automaton[of "the (rr1_of_rr1_rel \<F> Rs s1)" _ "the (rr1_of_rr1_rel \<F> Rs s2)"]
by auto
next
case (R2Inv r)
show ?case using R2Inv by (auto simp: swap_RR2_spec)
next
case (R2Union r1 r2)
then show ?case using union_automaton
by (auto simp: RR2_spec_def \<L>_union)
next
case (R2Inter r1 r2)
then show ?case
by (auto simp: \<L>_intersect RR2_spec_def)
next
case (R2Diff r1 r2)
then show ?case by (auto intro: RR2_difference)
next
case (R2Comp r1 r2)
then show ?case using eval_rr12_rel_sig
by (auto intro!: rr2_compositon) blast+
}
qed
subsection \<open>Misc\<close>
lemma eval_formula_arity_cong:
assumes "\<And>i. i < formula_arity f \<Longrightarrow> \<alpha>' i = \<alpha> i"
shows "eval_formula \<F> Rs \<alpha>' f = eval_formula \<F> Rs \<alpha> f"
proof -
have [simp]: "j < length fs \<Longrightarrow> i < formula_arity (fs ! j) \<Longrightarrow> i < max_list (map formula_arity fs)" for i j fs
by (simp add: less_le_trans max_list)
show ?thesis using assms
proof (induct f arbitrary: \<alpha> \<alpha>')
case (FAnd fs)
show ?case using FAnd(1)[OF nth_mem, of _ \<alpha>' \<alpha>] FAnd(2) by (auto simp: all_set_conv_all_nth)
next
case (FOr fs)
show ?case using FOr(1)[OF nth_mem, of _ \<alpha>' \<alpha>] FOr(2) by (auto simp: ex_set_conv_ex_nth)
next
case (FNot f)
show ?case using FNot(1)[of \<alpha>' \<alpha>] FNot(2) by simp
next
case (FExists f)
show ?case using FExists(1)[of "\<alpha>'\<langle>0 : z\<rangle>" "\<alpha>\<langle>0 : z\<rangle>" for z] FExists(2) by (auto simp: shift_def)
next
case (FForall f)
show ?case using FForall(1)[of "\<alpha>'\<langle>0 : z\<rangle>" "\<alpha>\<langle>0 : z\<rangle>" for z] FForall(2) by (auto simp: shift_def)
qed simp_all
qed
subsection \<open>Connect semantics to FOL-Fitting\<close>
primrec form_of_formula :: "'trs formula \<Rightarrow> (unit, 'trs rr1_rel + 'trs rr2_rel) form" where
"form_of_formula (FRR1 r1 x) = Pred (Inl r1) [Var x]"
| "form_of_formula (FRR2 r2 x y) = Pred (Inr r2) [Var x, Var y]"
| "form_of_formula (FAnd fs) = foldr And (map form_of_formula fs) TT"
| "form_of_formula (FOr fs) = foldr Or (map form_of_formula fs) FF"
| "form_of_formula (FNot f) = Neg (form_of_formula f)"
| "form_of_formula (FExists f) = Exists (And (Pred (Inl R1Terms) [Var 0]) (form_of_formula f))"
| "form_of_formula (FForall f) = Forall (Impl (Pred (Inl R1Terms) [Var 0]) (form_of_formula f))"
fun for_eval_rel :: "('f \<times> nat) set \<Rightarrow> ('f, 'v) trs list \<Rightarrow> ftrs rr1_rel + ftrs rr2_rel \<Rightarrow> 'f gterm list \<Rightarrow> bool" where
"for_eval_rel \<F> Rs (Inl r1) [t] \<longleftrightarrow> t \<in> eval_rr1_rel \<F> Rs r1"
| "for_eval_rel \<F> Rs (Inr r2) [t, u] \<longleftrightarrow> (t, u) \<in> eval_rr2_rel \<F> Rs r2"
lemma eval_formula_conv:
"eval_formula \<F> Rs \<alpha> f = eval \<alpha> undefined (for_eval_rel \<F> Rs) (form_of_formula f)"
proof (induct f arbitrary: \<alpha>)
case (FAnd fs) then show ?case
unfolding eval_formula.simps by (induct fs) auto
next
case (FOr fs) then show ?case
unfolding eval_formula.simps by (induct fs) auto
qed auto
subsection \<open>RRn relations and formulas\<close>
lemma shift_rangeI [intro!]:
"range \<alpha> \<subseteq> T \<Longrightarrow> x \<in> T \<Longrightarrow> range (shift \<alpha> i x) \<subseteq> T"
by (auto simp: shift_def)
definition formula_relevant where
"formula_relevant \<F> Rs vs fm \<longleftrightarrow>
(\<forall>\<alpha> \<alpha>'. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<longrightarrow> range \<alpha>' \<subseteq> \<T>\<^sub>G \<F> \<longrightarrow> map \<alpha> vs = map \<alpha>' vs \<longrightarrow> eval_formula \<F> Rs \<alpha> fm \<longrightarrow> eval_formula \<F> Rs \<alpha>' fm)"
lemma formula_relevant_mono:
"set vs \<subseteq> set ws \<Longrightarrow> formula_relevant \<F> Rs vs fm \<Longrightarrow> formula_relevant \<F> Rs ws fm"
unfolding formula_relevant_def
by (meson map_eq_conv subset_code(1))
lemma formula_relevantD:
"formula_relevant \<F> Rs vs fm \<Longrightarrow>
range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<Longrightarrow> range \<alpha>' \<subseteq> \<T>\<^sub>G \<F> \<Longrightarrow> map \<alpha> vs = map \<alpha>' vs \<Longrightarrow>
eval_formula \<F> Rs \<alpha> fm \<Longrightarrow> eval_formula \<F> Rs \<alpha>' fm"
unfolding formula_relevant_def
by blast
lemma trivial_formula_relevant:
assumes "\<And>\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<Longrightarrow> \<not> eval_formula \<F> Rs \<alpha> fm"
shows "formula_relevant \<F> Rs vs fm"
using assms unfolding formula_relevant_def
by auto
lemma formula_relevant_0_FExists:
assumes "formula_relevant \<F> Rs [0] fm"
shows "formula_relevant \<F> Rs [] (FExists fm)"
unfolding formula_relevant_def
proof (intro allI, intro impI)
fix \<alpha> \<alpha>' assume ass: "range \<alpha> \<subseteq> \<T>\<^sub>G \<F>" "range (\<alpha>' :: fvar \<Rightarrow> 'a gterm) \<subseteq> \<T>\<^sub>G \<F>"
"eval_formula \<F> Rs \<alpha> (FExists fm)"
from ass(3) obtain z where "z \<in> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs (\<alpha>\<langle>0 : z\<rangle>) fm"
by auto
then show "eval_formula \<F> Rs \<alpha>' (FExists fm)"
using ass(1, 2) formula_relevantD[OF assms, of "\<alpha>\<langle>0:z\<rangle>" "\<alpha>'\<langle>0:z\<rangle>"]
by (auto simp: shift_rangeI intro!: exI[of _ z])
qed
definition formula_spec where
"formula_spec \<F> Rs vs A fm \<longleftrightarrow> sorted vs \<and> distinct vs \<and>
formula_relevant \<F> Rs vs fm \<and>
RRn_spec (length vs) A {map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> eval_formula \<F> Rs \<alpha> fm}"
lemma formula_spec_RRn_spec:
"formula_spec \<F> Rs vs A fm \<Longrightarrow> RRn_spec (length vs) A {map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> eval_formula \<F> Rs \<alpha> fm}"
by (simp add: formula_spec_def)
lemma formula_spec_nt_empty_form_sat:
"\<not> reg_empty A \<Longrightarrow> formula_spec \<F> Rs vs A fm \<Longrightarrow> \<exists> \<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> eval_formula \<F> Rs \<alpha> fm"
unfolding formula_spec_def
by (auto simp: RRn_spec_def \<L>_def)
lemma formula_spec_empty:
"reg_empty A \<Longrightarrow> formula_spec \<F> Rs vs A fm \<Longrightarrow> range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<Longrightarrow> eval_formula \<F> Rs \<alpha> fm \<longleftrightarrow> False"
unfolding formula_spec_def
by (auto simp: RRn_spec_def \<L>_def)
text \<open>In each inference step, we obtain a triple consisting of a formula @{term "fm"}, a list of
relevant variables @{term "vs"} (typically a sublist of @{term "[0..<formula_arity fm]"}), and
an RRn automaton @{term "A"}, such that the property @{term "formula_spec \<F> Rs vs A fm"} holds.\<close>
lemma false_formula_spec:
"sorted vs \<Longrightarrow> distinct vs \<Longrightarrow> formula_spec \<F> Rs vs empty_reg FFalse"
by (auto simp: formula_spec_def false_RRn_spec FFalse_def formula_relevant_def)
lemma true_formula_spec:
assumes "vs \<noteq> [] \<or> \<T>\<^sub>G (fset \<F>) \<noteq> {}" "sorted vs" "distinct vs"
shows "formula_spec (fset \<F>) Rs vs (true_RRn \<F> (length vs)) FTrue"
proof -
have "{ts. length ts = length vs \<and> set ts \<subseteq> \<T>\<^sub>G (fset \<F>)} = {map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G (fset \<F>)}"
proof (intro equalityI subsetI CollectI, goal_cases LR RL)
case (LR ts)
moreover obtain t0 where "funas_gterm t0 \<subseteq> fset \<F>" using LR assms(1) unfolding \<T>\<^sub>G_equivalent_def
by (cases vs) fastforce+
ultimately show ?case using `distinct vs`
apply (intro exI[of _ "\<lambda>t. if t \<in> set vs then ts ! inv_into {0..<length vs} ((!) vs) t else t0"])
apply (auto intro!: nth_equalityI dest!: inj_on_nth[of vs "{0..<length vs}"] simp: in_set_conv_nth \<T>\<^sub>G_equivalent_def)
by (metis inv_to_set mem_Collect_eq subsetD)
qed fastforce
then show ?thesis using assms true_RRn_spec[of "length vs" \<F>]
by (auto simp: formula_spec_def FTrue_def formula_relevant_def \<T>\<^sub>G_equivalent_def)
qed
lemma relabel_formula_spec:
"formula_spec \<F> Rs vs A fm \<Longrightarrow> formula_spec \<F> Rs vs (relabel_reg A) fm"
by (simp add: formula_spec_def)
lemma trim_formula_spec:
"formula_spec \<F> Rs vs A fm \<Longrightarrow> formula_spec \<F> Rs vs (trim_reg A) fm"
by (simp add: formula_spec_def)
definition fit_permute :: "nat list \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> nat list" where
"fit_permute vs vs' vs'' = map (\<lambda>v. if v \<in> set vs then the (mem_idx v vs) else length vs + the (mem_idx v vs'')) vs'"
definition fit_rrn :: "('f \<times> nat) fset \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> (nat, 'f option list) reg \<Rightarrow> (_, 'f option list) reg" where
"fit_rrn \<F> vs vs' A = (let vs'' = subtract_list_sorted vs' vs in
fmap_funs_reg (\<lambda>fs. map ((!) fs) (fit_permute vs vs' vs''))
(fmap_funs_reg (pad_with_Nones (length vs) (length vs'')) (pair_automaton_reg A (true_RRn \<F> (length vs'')))))"
lemma the_mem_idx_simp [simp]:
"distinct xs \<Longrightarrow> i < length xs \<Longrightarrow> the (mem_idx (xs ! i) xs) = i"
using mem_idx_sound[THEN iffD1, OF nth_mem, of i xs] mem_idx_sound_output[of "xs ! i" xs] distinct_conv_nth
by fastforce
lemma fit_rrn:
assumes spec: "formula_spec (fset \<F>) Rs vs A fm" and vs: "sorted vs'" "distinct vs'" "set vs \<subseteq> set vs'"
shows "formula_spec (fset \<F>) Rs vs' (fit_rrn \<F> vs vs' A) fm"
using spec unfolding formula_spec_def formula_relevant_def
apply (elim conjE)
proof (intro conjI vs(1,2) allI, goal_cases rel spec)
case (rel \<alpha> \<alpha>') show ?case using vs(3)
by (fastforce intro!: rel(3)[rule_format, of \<alpha> \<alpha>'])
next
case spec
define vs'' where "vs'' = subtract_list_sorted vs' vs"
have evalI: "range \<alpha> \<subseteq> \<T>\<^sub>G (fset \<F>) \<Longrightarrow> range \<alpha>' \<subseteq> \<T>\<^sub>G (fset \<F>) \<Longrightarrow> map \<alpha> vs = map \<alpha>' vs
\<Longrightarrow> eval_formula (fset \<F>) Rs \<alpha> fm \<Longrightarrow> eval_formula (fset \<F>) Rs \<alpha>' fm" for \<alpha> \<alpha>'
using spec(3) by blast
have [simp]: "set vs' = set vs \<union> set vs''" "set vs'' \<inter> set vs = {}" "set vs \<inter> set vs'' = {}" and d: "distinct vs''"
using vs spec(1,2) by (auto simp: vs''_def)
then have [dest]: "v \<in> set vs'' \<Longrightarrow> v \<in> set vs \<Longrightarrow> False" for v by blast
note * = permute_automaton[OF append_automaton[OF spec(4) true_RRn_spec, of "length vs''"]]
have [simp]: "distinct vs \<Longrightarrow> i \<in> set vs \<Longrightarrow> vs ! the (mem_idx i vs) = (i :: nat)" for vs i
by (simp add: mem_idx_sound mem_idx_sound_output)
have [dest]: "distinct vs \<Longrightarrow> i \<in> set vs \<Longrightarrow> \<not> the (mem_idx i vs) < length vs \<Longrightarrow> False" for i
by (meson mem_idx_sound2 mem_idx_sound_output option.exhaust_sel)
show ?case unfolding fit_rrn_def Let_def vs''_def[symmetric] \<T>\<^sub>G_equivalent_def
apply (rule subst[where P = "\<lambda>l. RRn_spec l _ _", OF _ subst[where P = "\<lambda>ta. RRn_spec _ _ ta", OF _ *]])
subgoal by (simp add: fit_permute_def)
subgoal
apply (intro equalityI subsetI CollectI imageI; elim imageE CollectE exE conjE; unfold \<T>\<^sub>G_equivalent_def)
subgoal for x fs ts us \<alpha>
using spec(1, 2) d
apply (intro exI[of _ "\<lambda>v. if v \<in> set vs'' then us ! the (mem_idx v vs'') else \<alpha> v"])
apply (auto simp: fit_permute_def nth_append \<T>\<^sub>G_equivalent_def
intro!: nth_equalityI evalI[of \<alpha> "\<lambda>v. if v \<in> set vs'' then us ! the (mem_idx v vs'') else \<alpha> v"])
apply (metis distinct_Ex1 in_mono mem_Collect_eq nth_mem the_mem_idx_simp)
apply (metis distinct_Ex1 in_mono mem_Collect_eq nth_mem the_mem_idx_simp)
apply blast
by (meson \<open>\<And>va. \<lbrakk>va \<in> set vs''; va \<in> set vs\<rbrakk> \<Longrightarrow> False\<close> nth_mem)
subgoal premises p for xs \<alpha>
apply (intro rev_image_eqI[of "map \<alpha> (vs @ vs'')"])
subgoal using p by (force intro!: exI[of _ "map \<alpha> vs", OF exI[of _ "map \<alpha> vs''"]])
subgoal using p(1)
by (force intro!: nth_equalityI simp: fit_permute_def comp_def nth_append dest: iffD1[OF mem_idx_sound] mem_idx_sound_output)
done
done
subgoal using vs spec(1,2) unfolding fit_permute_def
apply (intro equalityI subsetI)
subgoal by (auto 0 3 dest: iffD1[OF mem_idx_sound] mem_idx_sound_output)
subgoal for x
apply (simp add: Compl_eq[symmetric] Diff_eq[symmetric] Un_Diff Diff_triv Int_absorb1)
apply (simp add: nth_image[symmetric, of "length xs" xs for xs, simplified] image_iff comp_def)
using image_cong[OF refl arg_cong[OF the_mem_idx_simp]] \<open>distinct vs''\<close>
by (smt (z3) add_diff_inverse_nat add_less_cancel_left atLeast0LessThan lessThan_iff the_mem_idx_simp)
done
done
qed
definition fit_rrns :: "('f \<times> nat) fset \<Rightarrow> (ftrs formula \<times> nat list \<times> (nat, 'f option list) reg) list \<Rightarrow>
nat list \<times> ((nat, 'f option list) reg) list" where
"fit_rrns \<F> rrns = (let vs' = fold union_list_sorted (map (fst \<circ> snd) rrns) [] in
(vs', map (\<lambda>(fm, vs, ta). relabel_reg (trim_reg (fit_rrn \<F> vs vs' ta))) rrns))"
lemma sorted_union_list_sortedI [simp]:
"sorted xs \<Longrightarrow> sorted ys \<Longrightarrow> sorted (union_list_sorted xs ys)"
by (induct xs ys rule: union_list_sorted.induct) auto
lemma distinct_union_list_sortedI [simp]:
"sorted xs \<Longrightarrow> sorted ys \<Longrightarrow> distinct xs \<Longrightarrow> distinct ys \<Longrightarrow> distinct (union_list_sorted xs ys)"
by (induct xs ys rule: union_list_sorted.induct) auto
lemma fit_rrns:
assumes infs: "\<And>fvA. fvA \<in> set rrns \<Longrightarrow> formula_spec (fset \<F>) Rs (fst (snd fvA)) (snd (snd fvA)) (fst fvA)"
assumes "(vs', tas') = fit_rrns \<F> rrns"
shows "length tas' = length rrns" "\<And>i. i < length rrns \<Longrightarrow> formula_spec (fset \<F>) Rs vs' (tas' ! i) (fst (rrns ! i))"
"distinct vs'" "sorted vs'"
proof (goal_cases)
have vs': "vs' = fold union_list_sorted (map (fst \<circ> snd) rrns) []" using assms(2) by (simp add: fit_rrns_def Let_def)
have *: "sorted vs'" "distinct vs'" "\<And>fvA. fvA \<in> set rrns \<Longrightarrow> set (fst (snd fvA)) \<subseteq> set vs'"
using infs[unfolded formula_spec_def, THEN conjunct2, THEN conjunct1]
infs[unfolded formula_spec_def, THEN conjunct1]
unfolding vs' by (induct rrns rule: rev_induct) auto
{
case 1 then show ?case using assms(2) by (simp add: fit_rrns_def Let_def)
next
case (2 i)
have tas': "tas' ! i = relabel_reg (trim_reg (fit_rrn \<F> (fst (snd (rrns ! i))) vs' (snd (snd (rrns ! i)))))"
using 2 assms(2) by (simp add: fit_rrns_def Let_def split: prod.splits)
from *(1,2) *(3)[OF nth_mem] show ?case using 2 unfolding tas'
by (auto intro!: relabel_formula_spec trim_formula_spec fit_rrn 2 assms(1,2))
next
case 3 show ?case by (rule *)
next
case 4 show ?case by (rule *)
}
qed
subsection \<open>Building blocks\<close>
definition for_rrn where
"for_rrn tas = fold (\<lambda>A B. relabel_reg (reg_union A B)) tas (Reg {||} (TA {||} {||}))"
lemma for_rrn:
assumes "length tas = length fs" "\<And>i. i < length fs \<Longrightarrow> formula_spec \<F> Rs vs (tas ! i) (fs ! i)"
and vs: "sorted vs" "distinct vs"
shows "formula_spec \<F> Rs vs (for_rrn tas) (FOr fs)"
using assms(1,2) unfolding for_rrn_def
proof (induct fs arbitrary: tas rule: rev_induct)
case Nil then show ?case using vs false_formula_spec[of vs \<F> Rs] by (auto simp: FFalse_def)
next
case (snoc fm fs)
have *: "Bex (set [x]) P = P x" for P x by auto
have [intro!]: "formula_spec \<F> Rs vs (reg_union A B) (FOr (fs @ [fm]))" if
"formula_spec \<F> Rs vs A fm" "formula_spec \<F> Rs vs B (FOr fs)" for A B using that
unfolding formula_spec_def
apply (intro conjI, blast, blast)
subgoal unfolding formula_relevant_def eval_formula.simps set_append bex_Un * by blast
apply (elim conjE)
subgoal premises p by (rule subst[of _ _ "RRn_spec _ _", OF _ union_automaton[OF p(6,8)]]) auto
done
show ?case using snoc(1)[of "take (length fs) tas"] snoc(2) snoc(3)[simplified, OF less_SucI] snoc(3)[of "length fs"] vs
by (cases tas rule: rev_exhaust) (auto simp: min_def nth_append intro!: relabel_formula_spec)
qed
fun fand_rrn where
"fand_rrn \<F> n [] = true_RRn \<F> n"
| "fand_rrn \<F> n (A # tas) = fold (\<lambda>A B. simplify_reg (reg_intersect A B)) tas A"
lemma fand_rrn:
assumes "\<T>\<^sub>G (fset \<F>) \<noteq> {}" "length tas = length fs" "\<And>i. i < length fs \<Longrightarrow> formula_spec (fset \<F>) Rs vs (tas ! i) (fs ! i)"
and vs: "sorted vs" "distinct vs"
shows "formula_spec (fset \<F>) Rs vs (fand_rrn \<F> (length vs) tas) (FAnd fs)"
proof (cases fs)
case Nil
have "tas = []" using assms(2) by (auto simp: Nil)
then show ?thesis using true_formula_spec[OF _ vs, of \<F> Rs] assms(1) Nil
by (simp add: FTrue_def)
next
case (Cons fm fs)
obtain ta tas' where tas: "tas = ta # tas'" using assms(2) Cons by (cases tas) auto
show ?thesis using assms(2) assms(3)[of "Suc _"] unfolding tas Cons
unfolding list.size add_Suc_right add_0_right nat.inject Suc_less_eq nth_Cons_Suc fand_rrn.simps
proof (induct fs arbitrary: tas' rule: rev_induct)
case Nil
have "formula_relevant (fset \<F>) Rs vs (FAnd [fm])" using assms(3)[of 0]
apply (simp add: tas Cons formula_spec_def)
unfolding formula_relevant_def eval_formula.simps in_set_simps by blast
then show ?case using assms(3)[of 0, unfolded tas Cons, simplified] Nil by (simp add: formula_spec_def)
next
case (snoc fm' fs)
have *: "Ball (insert x X) P \<longleftrightarrow> P x \<and> Ball X P" for P x X by auto
have [intro!]: "formula_spec (fset \<F>) Rs vs (reg_intersect A B) (FAnd (fm # fs @ [fm']))" if
"formula_spec (fset \<F>) Rs vs A fm'" "formula_spec (fset \<F>) Rs vs B (FAnd (fm # fs))" for A B using that
unfolding formula_spec_def
apply (intro conjI, blast, blast)
subgoal unfolding formula_relevant_def eval_formula.simps set_append set_simps ball_simps ball_Un in_set_simps *
by blast
apply (elim conjE)
subgoal premises p
by (rule subst[of _ _ "RRn_spec _ _", OF _ intersect_automaton[OF p(6,8)]])
(auto dest: p(5)[unfolded formula_relevant_def, rule_format])
done
show ?case using snoc(1)[of "take (length fs) tas'"] snoc(2) snoc(3)[simplified, OF less_SucI] snoc(3)[of "length fs"] vs
by (cases tas' rule: rev_exhaust) (auto simp: min_def nth_append simplify_reg_def intro!: relabel_formula_spec trim_formula_spec)
qed
qed
subsubsection \<open>IExists inference rule\<close>
lemma lift_fun_gpairD:
"map_gterm lift_fun s = gpair t u \<Longrightarrow> t = s"
"map_gterm lift_fun s = gpair t u \<Longrightarrow> u = s"
by (metis gfst_gpair gsnd_gpair map_funs_term_some_gpair)+
definition upd_bruijn :: "nat list \<Rightarrow> nat list" where
"upd_bruijn vs = tl (map (\<lambda> x. x - 1) vs)"
lemma upd_bruijn_length[simp]:
"length (upd_bruijn vs) = length vs - 1"
by (induct vs) (auto simp: upd_bruijn_def)
lemma pres_sorted_dec:
"sorted xs \<Longrightarrow> sorted (map (\<lambda>x. x - Suc 0) xs)"
by (induct xs) auto
lemma upd_bruijn_pres_sorted:
"sorted xs \<Longrightarrow> sorted (upd_bruijn xs)"
unfolding upd_bruijn_def
by (intro sorted_tl) (auto simp: pres_sorted_dec)
lemma pres_distinct_not_0_list_dec:
"distinct xs \<Longrightarrow> 0 \<notin> set xs \<Longrightarrow> distinct (map (\<lambda>x. x - Suc 0) xs)"
by (induct xs) (auto, metis Suc_pred neq0_conv)
lemma upd_bruijn_pres_distinct:
assumes "sorted xs" "distinct xs"
shows "distinct (upd_bruijn xs)"
proof -
have "sorted (ys :: nat list) \<Longrightarrow> distinct ys \<Longrightarrow> 0 \<notin> set (tl ys)" for ys
by (induct ys) auto
from this[OF assms] show ?thesis using assms(2)
using pres_distinct_not_0_list_dec[OF distinct_tl, OF assms(2)]
unfolding upd_bruijn_def
by (simp add: map_tl)
qed
lemma upd_bruijn_relevant_inv:
assumes "sorted vs" "distinct vs" "0 \<in> set vs"
and "\<And> x. x \<in> set (upd_bruijn vs) \<Longrightarrow> \<alpha> x = \<alpha>' x"
shows "\<And> x. x \<in> set vs \<Longrightarrow> (shift \<alpha> 0 z) x = (shift \<alpha>' 0 z) x"
using assms unfolding upd_bruijn_def
by (induct vs) (auto simp add: FOL_Fitting.shift_def)
lemma ExistsI_upd_brujin_0:
assumes "sorted vs" "distinct vs" "0 \<in> set vs" "formula_relevant \<F> Rs vs fm"
shows "formula_relevant \<F> Rs (upd_bruijn vs) (FExists fm)"
unfolding formula_relevant_def
proof (intro allI, intro impI)
fix \<alpha> \<alpha>' assume ass: "range \<alpha> \<subseteq> \<T>\<^sub>G \<F>" "range (\<alpha>' :: fvar \<Rightarrow> 'a gterm) \<subseteq> \<T>\<^sub>G \<F>"
"map \<alpha> (upd_bruijn vs) = map \<alpha>' (upd_bruijn vs)" "eval_formula \<F> Rs \<alpha> (FExists fm)"
from ass(4) obtain z where "z \<in> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs (\<alpha>\<langle>0 : z\<rangle>) fm"
by auto
then show "eval_formula \<F> Rs \<alpha>' (FExists fm)"
using ass(1 - 3) formula_relevantD[OF assms(4), of "\<alpha>\<langle>0:z\<rangle>" "\<alpha>'\<langle>0:z\<rangle>"]
using upd_bruijn_relevant_inv[OF assms(1 - 3), of "\<alpha>" "\<alpha>'"]
by (auto simp: shift_rangeI intro!: exI[of _ z])
qed
declare subsetI[rule del]
lemma ExistsI_upd_brujin_no_0:
assumes "0 \<notin> set vs" and "formula_relevant \<F> Rs vs fm"
shows "formula_relevant \<F> Rs (map (\<lambda>x. x - Suc 0) vs) (FExists fm)"
unfolding formula_relevant_def
proof ((intro allI)+ , (intro impI)+, unfold eval_formula.simps)
fix \<alpha> \<alpha>' assume st: "range \<alpha> \<subseteq> \<T>\<^sub>G \<F>" "range \<alpha>' \<subseteq> \<T>\<^sub>G \<F>"
"map \<alpha> (map (\<lambda>x. x - Suc 0) vs) = map \<alpha>' (map (\<lambda>x. x - Suc 0) vs)"
"\<exists> z \<in> \<T>\<^sub>G \<F>. eval_formula \<F> Rs (shift \<alpha> 0 z) fm"
then obtain z where w: "z \<in> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs (shift \<alpha> 0 z) fm" by auto
from this(1) have "eval_formula \<F> Rs (shift \<alpha>' 0 z) fm"
using st(1 - 3) assms(1) FOL_Fitting.shift_def
apply (intro formula_relevantD[OF assms(2) _ _ _ w(2), of "shift \<alpha>' 0 z"])
by auto (simp add: FOL_Fitting.shift_def)
then show "\<exists> z \<in> \<T>\<^sub>G \<F>. eval_formula \<F> Rs (shift \<alpha>' 0 z) fm" using w(1)
by blast
qed
definition shift_right where
"shift_right \<alpha> \<equiv> \<lambda> i. \<alpha> (i + 1)"
lemma shift_right_nt_0:
"i \<noteq> 0 \<Longrightarrow> \<alpha> i = shift_right \<alpha> (i - Suc 0)"
unfolding shift_right_def
by auto
lemma shift_shift_right_id [simp]:
"shift (shift_right \<alpha>) 0 (\<alpha> 0) = \<alpha>"
by (auto simp: shift_def shift_right_def)
lemma shift_right_rangeI [intro]:
"range \<alpha> \<subseteq> T \<Longrightarrow> range (shift_right \<alpha>) \<subseteq> T"
by (auto simp: shift_right_def intro: subsetI)
lemma eval_formula_shift_right_eval:
"eval_formula \<F> Rs \<alpha> fm \<Longrightarrow> eval_formula \<F> Rs (shift (shift_right \<alpha>) 0 (\<alpha> 0)) fm"
"eval_formula \<F> Rs (shift (shift_right \<alpha>) 0 (\<alpha> 0)) fm \<Longrightarrow> eval_formula \<F> Rs \<alpha> fm"
by (auto)
declare subsetI[intro!]
lemma nt_rel_0_trivial_shift:
assumes "0 \<notin> set vs"
shows "{map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> eval_formula \<F> Rs \<alpha> fm} =
{map (\<lambda>x. \<alpha> (x - Suc 0)) vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> (\<exists>z \<in> \<T>\<^sub>G \<F>. eval_formula \<F> Rs (\<alpha>\<langle>0:z\<rangle>) fm)}"
(is "?Ls = ?Rs")
proof
{fix \<alpha> assume ass: "range \<alpha> \<subseteq> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs \<alpha> fm"
then have "map \<alpha> vs = map (\<lambda>x. (shift_right \<alpha>) (x - Suc 0)) vs"
"range (shift_right \<alpha>) \<subseteq> \<T>\<^sub>G \<F>" "\<alpha> 0 \<in>\<T>\<^sub>G \<F>" "eval_formula \<F> Rs (shift (shift_right \<alpha>) 0 (\<alpha> 0)) fm"
using shift_right_rangeI[OF ass(1)] assms
by (auto intro: eval_formula_shift_right_eval(1), metis shift_right_nt_0)}
then show "?Ls \<subseteq> ?Rs"
by blast
next
show "?Rs \<subseteq> ?Ls"
by auto (metis FOL_Fitting.shift_def One_nat_def assms not_less0 shift_rangeI)
qed
lemma relevant_vars_upd_bruijn_tl:
assumes "sorted vs" "distinct vs"
shows "map (shift_right \<alpha>) (upd_bruijn vs) = tl (map \<alpha> vs)" using assms
proof (induct vs)
case (Cons a vs) then show ?case
using le_antisym
by (auto simp: upd_bruijn_def shift_right_def)
(metis One_nat_def Suc_eq_plus1 le_0_eq shift_right_def shift_right_nt_0)
qed (auto simp: upd_bruijn_def)
lemma drop_upd_bruijn_set:
assumes "sorted vs" "distinct vs"
shows "drop 1 ` {map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> eval_formula \<F> Rs \<alpha> fm} =
{map \<alpha> (upd_bruijn vs) |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> (\<exists>z\<in>\<T>\<^sub>G \<F>. eval_formula \<F> Rs (\<alpha>\<langle>0:z\<rangle>) fm)}"
(is "?Ls = ?Rs")
proof
{fix \<alpha> assume ass: "range \<alpha> \<subseteq> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs \<alpha> fm"
then have "drop 1 (map \<alpha> vs) = map (shift_right \<alpha>) (upd_bruijn vs)"
"range (shift_right \<alpha>) \<subseteq> \<T>\<^sub>G \<F>" "\<alpha> 0 \<in>\<T>\<^sub>G \<F>" "eval_formula \<F> Rs (shift (shift_right \<alpha>) 0 (\<alpha> 0)) fm"
using shift_right_rangeI[OF ass(1)]
by (auto simp: tl_drop_conv relevant_vars_upd_bruijn_tl[OF assms(1, 2)])}
then show "?Ls \<subseteq> ?Rs"
by blast
next
have [dest]: "0 \<in> set (tl vs) \<Longrightarrow> False" using assms(1, 2)
by (cases vs) auto
{fix \<alpha> z assume ass: "range \<alpha> \<subseteq> \<T>\<^sub>G \<F>" "z \<in> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs (\<alpha>\<langle>0:z\<rangle>) fm"
then have "map \<alpha> (upd_bruijn vs) = tl (map (\<alpha>\<langle>0:z\<rangle>) vs)"
"range (\<alpha>\<langle>0:z\<rangle>) \<subseteq> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs (\<alpha>\<langle>0:z\<rangle>) fm"
using shift_rangeI[OF ass(1)]
by (auto simp: upd_bruijn_def shift_def simp flip: map_tl)}
then show "?Rs \<subseteq> ?Ls"
by (auto simp: tl_drop_conv image_def) blast
qed
lemma closed_sat_form_env_dom:
assumes "formula_relevant \<F> Rs [] (FExists fm)" "range \<alpha> \<subseteq> \<T>\<^sub>G \<F>" "eval_formula \<F> Rs \<alpha> fm"
shows "{[\<alpha> 0] |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G \<F> \<and> (\<exists> z \<in> \<T>\<^sub>G \<F>. eval_formula \<F> Rs (\<alpha>\<langle>0:z\<rangle>) fm)} = {[t] | t. t \<in> \<T>\<^sub>G \<F>}"
using formula_relevantD[OF assms(1)] assms(2-)
apply auto
apply blast
by (smt rangeI shift_eq shift_rangeI shift_right_rangeI shift_shift_right_id subsetD)
(* MOVE *)
lemma find_append:
"find P (xs @ ys) = (if find P xs \<noteq> None then find P xs else find P ys)"
by (induct xs arbitrary: ys) (auto split!: if_splits)
subsection \<open>Checking inferences\<close>
derive linorder ext_step pos_step gtt_rel rr1_rel rr2_rel ftrs
derive compare ext_step pos_step gtt_rel rr1_rel rr2_rel ftrs
fun check_inference :: "(('f \<times> nat) fset \<Rightarrow> ('f, 'v) fin_trs list \<Rightarrow> ftrs rr1_rel \<Rightarrow> (nat, 'f) reg option)
\<Rightarrow> (('f \<times> nat) fset \<Rightarrow> ('f, 'v) fin_trs list \<Rightarrow> ftrs rr2_rel \<Rightarrow> (nat, 'f option \<times> 'f option) reg option)
\<Rightarrow> ('f \<times> nat) fset \<Rightarrow> ('f :: compare, 'v) fin_trs list
\<Rightarrow> (ftrs formula \<times> nat list \<times> (nat, 'f option list) reg) list
\<Rightarrow> (nat \<times> ftrs inference \<times> ftrs formula \<times> info list)
\<Rightarrow> (ftrs formula \<times> nat list \<times> (nat, 'f option list) reg) option" where
"check_inference rr1c rr2c \<F> Rs infs (l, step, fm, is) = do {
guard (l = length infs);
case step of
IRR1 s x \<Rightarrow> do {
guard (fm = FRR1 s x);
liftO1 (\<lambda>ta. (FRR1 s x, [x], fmap_funs_reg (\<lambda>f. [Some f]) ta)) (rr1c \<F> Rs s)
}
| IRR2 r x y \<Rightarrow> do {
guard (fm = FRR2 r x y);
case compare x y of
Lt \<Rightarrow> liftO1 (\<lambda>ta. (FRR2 r x y, [x, y], fmap_funs_reg (\<lambda>(f, g). [f, g]) ta)) (rr2c \<F> Rs r)
| Eq \<Rightarrow> liftO1 (\<lambda>ta. (FRR2 r x y, [x], fmap_funs_reg (\<lambda>f. [Some f]) ta))
(liftO1 (simplify_reg \<circ> proj_1_reg)
(liftO2 (\<lambda> t1 t2. simplify_reg (reg_intersect t1 t2)) (rr2c \<F> Rs r) (rr2c \<F> Rs (R2Diag R1Terms))))
| Gt \<Rightarrow> liftO1 (\<lambda>ta. (FRR2 r x y, [y, x], fmap_funs_reg (\<lambda>(f, g). [g, f]) ta)) (rr2c \<F> Rs r)
}
| IAnd ls \<Rightarrow> do {
guard (\<forall>l' \<in> set ls. l' < l);
guard (fm = FAnd (map (\<lambda>l'. fst (infs ! l')) ls));
let (vs', tas') = fit_rrns \<F> (map ((!) infs) ls) in
Some (fm, vs', fand_rrn \<F> (length vs') tas')
}
| IOr ls \<Rightarrow> do {
guard (\<forall>l' \<in> set ls. l' < l);
guard (fm = FOr (map (\<lambda>l'. fst (infs ! l')) ls));
let (vs', tas') = fit_rrns \<F> (map ((!) infs) ls) in
Some (fm, vs', for_rrn tas')
}
| INot l' \<Rightarrow> do {
guard (l' < l);
guard (fm = FNot (fst (infs ! l')));
let (vs', tas') = snd (infs ! l');
Some (fm, vs', simplify_reg (difference_reg (true_RRn \<F> (length vs')) tas'))
}
| IExists l' \<Rightarrow> do {
guard (l' < l);
guard (fm = FExists (fst (infs ! l')));
let (vs', tas') = snd (infs ! l');
if length vs' = 0 then Some (fm, [], tas') else
if reg_empty tas' then Some (fm, [], empty_reg)
else if 0 \<notin> set vs' then Some (fm, map (\<lambda> x. x - 1) vs', tas')
else if 1 = length vs' then Some (fm, [], true_RRn \<F> 0)
else Some (fm, upd_bruijn vs', rrn_drop_fst tas')
}
| IRename l' vs \<Rightarrow> guard (l' < l) \<then> None
| INNFPlus l' \<Rightarrow> do {
guard (l' < l);
let fm' = fst (infs ! l');
guard (ord_form_list_aci (nnf_to_list_aci (nnf (form_of_formula fm'))) = ord_form_list_aci (nnf_to_list_aci (nnf (form_of_formula fm))));
Some (fm, snd (infs ! l'))
}
| IRepl eq pos l' \<Rightarrow> guard (l' < l) \<then> None
}"
lemma RRn_spec_true_RRn:
"RRn_spec (Suc 0) (true_RRn \<F> (Suc 0)) {[t] |t. t \<in> \<T>\<^sub>G (fset \<F>)}"
apply (auto simp: RRn_spec_def \<T>\<^sub>G_equivalent_def fmap_funs_\<L>
image_def term_automaton[of \<F>, unfolded RR1_spec_def])
apply (metis gencode_singleton)+
done
lemma check_inference_correct:
assumes sig: "\<T>\<^sub>G (fset \<F>) \<noteq> {}" and Rs: "\<forall>R \<in> set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>"
assumes infs: "\<And>fvA. fvA \<in> set infs \<Longrightarrow> formula_spec (fset \<F>) (map fset Rs) (fst (snd fvA)) (snd (snd fvA)) (fst fvA)"
assumes inf: "check_inference rr1c rr2c \<F> Rs infs (l, step, fm, is) = Some (fm', vs, A')"
assumes rr1: "\<And>r1. \<forall>ta1. rr1c \<F> Rs r1 = Some ta1 \<longrightarrow> RR1_spec ta1 (eval_rr1_rel (fset \<F>) (map fset Rs) r1)"
assumes rr2: "\<And>r2. \<forall>ta2. rr2c \<F> Rs r2 = Some ta2 \<longrightarrow> RR2_spec ta2 (eval_rr2_rel (fset \<F>) (map fset Rs) r2)"
shows "l = length infs \<and> fm = fm' \<and> formula_spec (fset \<F>) (map fset Rs) vs A' fm'"
using inf
proof (induct step)
note [simp] = bind_eq_Some_conv guard_simps
let ?F = "fset \<F>" let ?Rs = "map fset Rs"
{
case (IRR1 s x)
then show ?case
using rr1[rule_format, of s]
subsetD[OF eval_rr12_rel_sig(1), of _ ?F ?Rs s]
by (force simp: formula_spec_def formula_relevant_def RR1_spec_def \<T>\<^sub>G_equivalent_def
intro!: RR1_to_RRn_spec[of _ "(\<lambda>\<alpha>. \<alpha> x) ` Collect P" for P, unfolded image_comp, unfolded image_Collect comp_def One_nat_def])
next
case (IRR2 r x y)
then show ?case using rr2[rule_format, of r]
subsetD[OF eval_rr12_rel_sig(2), of _ ?F ?Rs r]
two_comparisons_into_compare(1)[of x y "x = y" "x < y" "x > y"]
proof (induct "compare x y")
note [intro!] = RR1_to_RRn_spec[of _ "(\<lambda>\<alpha>. \<alpha> y) ` Collect P" for P, unfolded image_comp,
unfolded image_Collect comp_def One_nat_def prod.simps]
case Eq
then obtain A where w[simp]: "rr2c \<F> Rs r = Some A" by auto
from Eq obtain B where [simp]:"rr2c \<F> Rs (R2Diag R1Terms) = Some B" by auto
let ?B = "reg_intersect A B"
from Eq(3)[OF w] have "RR2_spec ?B (eval_rr2_rel ?F ?Rs r \<inter> Restr Id (\<T>\<^sub>G ?F))"
using rr2[rule_format, of "R2Diag R1Terms" B]
by (auto simp add: \<L>_intersect RR2_spec_def dest: lift_fun_gpairD)
then have "RR2_spec (relabel_reg (trim_reg ?B)) (eval_rr2_rel ?F ?Rs r \<inter> Restr Id (\<T>\<^sub>G ?F))" by simp
from proj_1(1)[OF this]
have "RR1_spec (proj_1_reg (relabel_reg (trim_reg ?B))) {\<alpha> y |\<alpha>. range \<alpha> \<subseteq> gterms ?F \<and> (\<alpha> y, \<alpha> y) \<in> eval_rr2_rel ?F ?Rs r}"
apply (auto simp: RR1_spec_def \<T>\<^sub>G_equivalent_def image_iff)
by (metis Eq.prems(3) IdI IntI \<T>\<^sub>G_equivalent_def fst_conv)
then show ?thesis using Eq
by (auto simp: formula_spec_def formula_relevant_def liftO1_def \<T>\<^sub>G_equivalent_def simplify_reg_def RR2_spec_def
split: if_splits intro!: exI[of _ "\<lambda>z. if z = x then _ else _"])
next
note [intro!] = RR2_to_RRn_spec[of _ "(\<lambda>\<alpha>. (\<alpha> x, \<alpha> y)) ` Collect P" for P, unfolded image_comp,
unfolded image_Collect comp_def numeral_2_eq_2 prod.simps]
case Lt then show ?thesis by (fastforce simp: formula_spec_def formula_relevant_def RR2_spec_def \<T>\<^sub>G_equivalent_def
split: if_splits intro!: exI[of _ "\<lambda>z. if z = x then _ else _"])
next
note [intro!] = RR2_to_RRn_spec[of _ "prod.swap ` (\<lambda>\<alpha>. (\<alpha> x, \<alpha> y)) ` Collect P" for P, OF swap_RR2_spec,
unfolded image_comp, unfolded image_Collect comp_def numeral_2_eq_2 prod.simps fmap_funs_reg_comp case_swap]
case Gt then show ?thesis
by (fastforce simp: formula_spec_def formula_relevant_def RR2_spec_def \<T>\<^sub>G_equivalent_def
split: if_splits intro!: exI[of _ "\<lambda>z. if z = x then _ else _"])
qed
next
case (IAnd ls)
have [simp]: "(fm, vs, ta) \<in> (!) infs ` set ls \<Longrightarrow> formula_spec ?F ?Rs vs ta fm" for fm vs ta
using infs IAnd by auto
show ?case using IAnd fit_rrns[OF assms(3), of "map ((!) infs) ls", OF _ prod.collapse]
by (force split: prod.splits intro!: fand_rrn[OF assms(1)])
next
case (IOr ls)
have [simp]: "(fm, vs, ta) \<in> (!) infs ` set ls \<Longrightarrow> formula_spec ?F ?Rs vs ta fm" for fm vs ta
using infs IOr by auto
show ?case using IOr fit_rrns[OF assms(3), of "map ((!) infs) ls", OF _ prod.collapse]
by (force split: prod.splits intro!: for_rrn)
next
case (INot l')
obtain fm vs' ta where [simp]: "infs ! l' = (fm, vs', ta)" by (cases "infs ! l'") auto
then have spec: "formula_spec ?F ?Rs vs ta fm" using infs[OF nth_mem, of l'] INot
by auto
have rel: "formula_relevant (fset \<F>) (map fset Rs) vs (FNot fm)" using spec
unfolding formula_spec_def formula_relevant_def
by (metis (no_types, lifting) eval_formula.simps(5))
have vs: "sorted vs" "distinct vs" using spec by (auto simp: formula_spec_def)
{fix xs assume ass: "\<forall>\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G (fset \<F>) \<longrightarrow> xs = map \<alpha> vs \<longrightarrow> \<not> eval_formula (fset \<F>) (map fset Rs) \<alpha> fm"
"length xs = length vs" "set xs \<subseteq> \<T>\<^sub>G (fset \<F>)"
from sig obtain s where mem: "s \<in> \<T>\<^sub>G (fset \<F>)" by blast
let ?g = "\<lambda> i. find (\<lambda> p. fst p = i) (zip vs [0 ..< length vs])"
let ?f = "\<lambda> i. if ?g i = None then s else xs ! snd (the (?g i))"
from vs(1) have *: "sorted (zip vs [0 ..< length vs])"
by (induct vs rule: rev_induct) (auto simp: sorted_append elim!: in_set_zipE intro!: sorted_append_bigger)
have "i < length vs \<Longrightarrow> ?g (vs ! i) = Some (vs ! i, i)" for i using vs(2)
by (induct vs rule: rev_induct) (auto simp: nth_append find_append find_Some_iff nth_eq_iff_index_eq split!: if_splits)
then have "map ?f vs = xs" using vs(2) ass(2)
by (intro nth_equalityI) (auto simp: find_None_iff set_zip)
moreover have "range ?f \<subseteq> \<T>\<^sub>G (fset \<F>)" using ass(2, 3) mem
using find_SomeD(2) set_zip_rightD by auto fastforce
ultimately have "\<exists>\<alpha>. xs = map \<alpha> vs \<and> range \<alpha> \<subseteq> \<T>\<^sub>G (fset \<F>) \<and> \<not> eval_formula (fset \<F>) (map fset Rs) \<alpha> fm" using ass(1)
by (intro exI[of _ ?f]) auto}
then have *: "{ts. length ts = length vs \<and> set ts \<subseteq> \<T>\<^sub>G (fset \<F>)} -
{map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G (fset \<F>) \<and> eval_formula (fset \<F>) (map fset Rs) \<alpha> fm} =
{map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G (fset \<F>) \<and> \<not> eval_formula (fset \<F>) (map fset Rs) \<alpha> fm}"
apply auto
apply force
using formula_relevantD[OF rel] unfolding eval_formula.simps
by (meson map_ext)
have "RRn_spec (length vs) (difference_reg (true_RRn \<F> (length vs)) ta)
{map \<alpha> vs |\<alpha>. range \<alpha> \<subseteq> \<T>\<^sub>G (fset \<F>) \<and> \<not> eval_formula (fset \<F>) (map fset Rs) \<alpha> fm}"
using RRn_difference[OF true_RRn_spec[of "length vs" \<F>] formula_spec_RRn_spec[OF spec]]
unfolding * by simp
then show ?case using INot spec rel
by (auto split: prod.splits simp: formula_spec_def)
next
case (IExists l')
obtain fm vs ta where [simp]: "infs ! l' = (fm, vs, ta)" by (cases "infs ! l'") auto
then have spec: "formula_spec ?F ?Rs vs ta fm" using infs[OF nth_mem, of l'] IExists
by auto
show ?case
proof (cases "length vs = 0")
case True
then show ?thesis using IExists spec
apply (auto simp: formula_spec_def)
subgoal apply (auto simp: formula_relevant_def)
apply (meson shift_rangeI)
done
subgoal apply (auto simp: RRn_spec_def image_iff)
apply (meson eval_formula_shift_right_eval(1) rangeI shift_right_rangeI subsetD)
apply (meson shift_rangeI)
done
done
next
case False note len = this
then have *[simp]: "vs \<noteq> []" by (cases vs) auto
show ?thesis
proof (cases "reg_empty ta")
case True
then show ?thesis using IExists spec formula_spec_empty[OF _ spec]
by (auto simp: \<T>\<^sub>G_equivalent_def comp_def formula_spec_def
shift_rangeI RRn_spec_def image_iff \<L>_epmty
intro!: trivial_formula_relevant)
next
case False
then have nt_empty [simp]: "\<L> ta \<noteq> {}" by auto
show ?thesis
proof (cases "0 \<notin> set vs")
case True
then have ta: "ta = A'" using spec IExists
by (auto simp: formula_spec_def)
from True have relv: "formula_relevant ?F ?Rs (map (\<lambda>x. x - Suc 0) vs) (FExists fm)"
using spec IExists
by (intro ExistsI_upd_brujin_no_0) (auto simp: formula_spec_def)
then show ?thesis using True spec IExists nt_rel_0_trivial_shift[OF True, of ?F ?Rs ]
by (auto simp: formula_spec_def \<T>\<^sub>G_equivalent_def comp_def
elim!: formula_relevantD
intro!: pres_distinct_not_0_list_dec pres_sorted_dec)
next
case False
then have rel_0: "0 \<in> set vs" by simp
show ?thesis
proof (cases "1 = length vs")
case True
then have [simp]: "vs = [0]" using rel_0 by (induct vs) auto
{fix t assume "0 |\<in>| ta_der (TA {|[] [] \<rightarrow> 0|} {||}) (term_of_gterm t)"
then have "t = GFun [] []" by (cases t) auto}
then have [simp]: "\<L> (Reg {|0|} (TA {|TA_rule [] [] 0|} {||})) = {GFun [] []}"
by (auto simp: \<L>_def gta_der_def gta_lang_def)
have [simp]: "GFun [] [] = gencode []"
by (auto simp: gencode_def gunions_def)
show ?thesis using IExists spec nt_empty
by (auto simp: formula_spec_def RRn_spec_true_RRn RRn_spec_def formula_relevant_0_FExists image_iff)
(meson eval_formula_shift_right_eval(1) in_mono rangeI shift_right_rangeI)
next
case False
from False show ?thesis using spec IExists rel_0 nt_empty
using rrn_drop_fst_lang[OF formula_spec_RRn_spec[OF spec]]
by (auto simp: formula_spec_def Suc_lessI simp flip: drop_upd_bruijn_set
split: prod.splits
intro: upd_bruijn_pres_sorted upd_bruijn_pres_distinct ExistsI_upd_brujin_0)
qed
qed
qed
qed
next
case (IRename l' vs)
then show ?case by simp
next
case (INNFPlus l')
show ?case using infs[OF nth_mem, of l'] INNFPlus
apply (auto simp: formula_spec_def formula_relevant_def eval_formula_conv)
apply (simp_all only: check_equivalence_by_nnf_sortedlist_aci[of "form_of_formula (fst (infs ! l'))" "form_of_formula fm"])
done
next
case (IRepl eq pos l')
then show ?case by simp
}
qed
end
\ No newline at end of file
diff --git a/thys/FO_Theory_Rewriting/FOR_Check_Impl.thy b/thys/FO_Theory_Rewriting/FOR_Check_Impl.thy
--- a/thys/FO_Theory_Rewriting/FOR_Check_Impl.thy
+++ b/thys/FO_Theory_Rewriting/FOR_Check_Impl.thy
@@ -1,773 +1,773 @@
theory FOR_Check_Impl
imports FOR_Check
Regular_Tree_Relations.Regular_Relation_Impl
NF_Impl
begin
section \<open>Inference checking implementation\<close>
(* we define epsilon free agtt/gtt constructions *)
definition "ftrancl_eps_free_closures \<A> = eps_free_automata (eps \<A>) \<A>"
abbreviation "ftrancl_eps_free_reg \<A> \<equiv> Reg (fin \<A>) (ftrancl_eps_free_closures (ta \<A>))"
lemma ftrancl_eps_free_ta_derI:
"(eps \<A>)|\<^sup>+| = eps \<A> \<Longrightarrow> ta_der (ftrancl_eps_free_closures \<A>) (term_of_gterm t) = ta_der \<A> (term_of_gterm t)"
using eps_free[of \<A>] ta_res_eps_free[of \<A>]
by (auto simp add: ftrancl_eps_free_closures_def)
lemma \<L>_ftrancl_eps_free_closuresI:
"(eps (ta \<A>))|\<^sup>+| = eps (ta \<A>) \<Longrightarrow> \<L> (ftrancl_eps_free_reg \<A>) = \<L> \<A>"
using ftrancl_eps_free_ta_derI[of "ta \<A>"]
unfolding \<L>_def by (auto simp: gta_lang_def gta_der_def)
definition "root_step R \<F> \<equiv> (let (TA1, TA2) = agtt_grrstep R \<F> in
(ftrancl_eps_free_closures TA1, TA2))"
definition AGTT_trancl_eps_free :: "('q, 'f) gtt \<Rightarrow> ('q + 'q, 'f) gtt" where
"AGTT_trancl_eps_free \<G> = (let (\<A>, \<B>) = AGTT_trancl \<G> in
(ftrancl_eps_free_closures \<A>, \<B>))"
definition GTT_trancl_eps_free where
"GTT_trancl_eps_free \<G> = (let (\<A>, \<B>) = GTT_trancl \<G> in
(ftrancl_eps_free_closures \<A>,
ftrancl_eps_free_closures \<B>))"
definition AGTT_comp_eps_free where
"AGTT_comp_eps_free \<G>\<^sub>1 \<G>\<^sub>2 = (let (\<A>, \<B>) = AGTT_comp' \<G>\<^sub>1 \<G>\<^sub>2 in
(ftrancl_eps_free_closures \<A>, \<B>))"
definition GTT_comp_eps_free where
"GTT_comp_eps_free \<G>\<^sub>1 \<G>\<^sub>2 =(let (\<A>, \<B>) = GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2 in
(ftrancl_eps_free_closures \<A>, ftrancl_eps_free_closures \<B>))"
(* epsilon free proves *)
lemma eps_free_relable [simp]:
"is_gtt_eps_free (relabel_gtt \<G>) = is_gtt_eps_free \<G>"
by (auto simp: is_gtt_eps_free_def relabel_gtt_def fmap_states_gtt_def fmap_states_ta_def)
lemma eps_free_prod_swap:
"is_gtt_eps_free (\<A>, \<B>) \<Longrightarrow> is_gtt_eps_free (\<B>, \<A>)"
by (auto simp: is_gtt_eps_free_def)
lemma eps_free_root_step:
"is_gtt_eps_free (root_step R \<F>)"
by (auto simp add: case_prod_beta is_gtt_eps_free_def root_step_def pair_at_to_agtt'_def ftrancl_eps_free_closures_def)
lemma eps_free_AGTT_trancl_eps_free:
"is_gtt_eps_free \<G> \<Longrightarrow> is_gtt_eps_free (AGTT_trancl_eps_free \<G>)"
by (auto simp: case_prod_beta is_gtt_eps_free_def AGTT_trancl_def Let_def
AGTT_trancl_eps_free_def ftrancl_eps_free_closures_def)
lemma eps_free_GTT_trancl_eps_free:
"is_gtt_eps_free \<G> \<Longrightarrow> is_gtt_eps_free (GTT_trancl_eps_free \<G>)"
by (auto simp: case_prod_beta is_gtt_eps_free_def GTT_trancl_eps_free_def ftrancl_eps_free_closures_def)
lemma eps_free_AGTT_comp_eps_free:
"is_gtt_eps_free \<G>\<^sub>2 \<Longrightarrow> is_gtt_eps_free (AGTT_comp_eps_free \<G>\<^sub>1 \<G>\<^sub>2)"
by (auto simp: case_prod_beta is_gtt_eps_free_def AGTT_comp_eps_free_def
ftrancl_eps_free_closures_def AGTT_comp_def fmap_states_gtt_def fmap_states_ta_def)
lemma eps_free_GTT_comp_eps_free:
"is_gtt_eps_free (GTT_comp_eps_free \<G>\<^sub>1 \<G>\<^sub>2)"
by (auto simp: case_prod_beta is_gtt_eps_free_def GTT_comp_eps_free_def ftrancl_eps_free_closures_def)
lemmas eps_free_const =
eps_free_prod_swap
eps_free_root_step
eps_free_AGTT_trancl_eps_free
eps_free_GTT_trancl_eps_free
eps_free_AGTT_comp_eps_free
eps_free_GTT_comp_eps_free
(* lang preserve proofs *)
lemma agtt_lang_derI:
assumes "\<And> t. ta_der (fst \<A>) (term_of_gterm t) = ta_der (fst \<B>) (term_of_gterm t)"
and "\<And> t. ta_der (snd \<A>) (term_of_gterm t) = ta_der (snd \<B>) (term_of_gterm t)"
shows "agtt_lang \<A> = agtt_lang \<B>" using assms
by (auto simp: agtt_lang_def gta_der_def)
lemma agtt_lang_root_step_conv:
"agtt_lang (root_step R \<F>) = agtt_lang (agtt_grrstep R \<F>)"
using ftrancl_eps_free_ta_derI[OF agtt_grrstep_eps_trancl(1), of R \<F>]
by (auto simp: case_prod_beta root_step_def intro!: agtt_lang_derI)
lemma agtt_lang_AGTT_trancl_eps_free_conv:
assumes "is_gtt_eps_free \<G>"
shows "agtt_lang (AGTT_trancl_eps_free \<G>) = agtt_lang (AGTT_trancl \<G>)"
proof -
let ?eps = "eps (fst (AGTT_trancl \<G>))"
have "?eps |O| ?eps = {||}" using assms
by (auto simp: AGTT_trancl_def is_gtt_eps_free_def Let_def fmap_states_ta_def)
from ftrancl_eps_free_ta_derI[OF frelcomp_empty_ftrancl_simp[OF this]]
show ?thesis
by (auto simp: case_prod_beta AGTT_trancl_eps_free_def intro!: agtt_lang_derI)
qed
lemma agtt_lang_GTT_trancl_eps_free_conv:
assumes "is_gtt_eps_free \<G>"
shows "agtt_lang (GTT_trancl_eps_free \<G>) = agtt_lang (GTT_trancl \<G>)"
proof -
have "(eps (fst (GTT_trancl \<G>)))|\<^sup>+| = eps (fst (GTT_trancl \<G>))"
"(eps (snd (GTT_trancl \<G>)))|\<^sup>+| = eps (snd (GTT_trancl \<G>))" using assms
by (auto simp: GTT_trancl_def Let_def is_gtt_eps_free_def \<Delta>_trancl_inv)
from ftrancl_eps_free_ta_derI[OF this(1)] ftrancl_eps_free_ta_derI[OF this(2)]
show ?thesis
by (auto simp: case_prod_beta GTT_trancl_eps_free_def intro!: agtt_lang_derI)
qed
lemma agtt_lang_AGTT_comp_eps_free_conv:
assumes "is_gtt_eps_free \<G>\<^sub>1" "is_gtt_eps_free \<G>\<^sub>2"
shows "agtt_lang (AGTT_comp_eps_free \<G>\<^sub>1 \<G>\<^sub>2) = agtt_lang (AGTT_comp' \<G>\<^sub>1 \<G>\<^sub>2)"
proof -
have "(eps (fst (AGTT_comp' \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+| = eps (fst (AGTT_comp' \<G>\<^sub>1 \<G>\<^sub>2))" using assms
by (auto simp: is_gtt_eps_free_def fmap_states_gtt_def fmap_states_ta_def
case_prod_beta AGTT_comp_def gtt_interface_def \<Q>_def intro!: frelcomp_empty_ftrancl_simp)
from ftrancl_eps_free_ta_derI[OF this] show ?thesis
by (auto simp: case_prod_beta AGTT_comp_eps_free_def intro!: agtt_lang_derI)
qed
lemma agtt_lang_GTT_comp_eps_free_conv:
assumes "is_gtt_eps_free \<G>\<^sub>1" "is_gtt_eps_free \<G>\<^sub>2"
shows "agtt_lang (GTT_comp_eps_free \<G>\<^sub>1 \<G>\<^sub>2) = agtt_lang (GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2)"
proof -
have "(eps (fst (GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+| = eps (fst (GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2))"
"(eps (snd (GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+| = eps (snd (GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2))" using assms
by (auto simp: is_gtt_eps_free_def fmap_states_gtt_def fmap_states_ta_def \<Delta>\<^sub>\<epsilon>_fmember
case_prod_beta GTT_comp_def gtt_interface_def \<Q>_def dest!: ground_ta_der_statesD
intro!: frelcomp_empty_ftrancl_simp)
from ftrancl_eps_free_ta_derI[OF this(1)] ftrancl_eps_free_ta_derI[OF this(2)]
show ?thesis
by (auto simp: case_prod_beta GTT_comp_eps_free_def intro!: agtt_lang_derI)
qed
fun gtt_of_gtt_rel_impl :: "('f \<times> nat) fset \<Rightarrow> ('f :: linorder, 'v) fin_trs list \<Rightarrow> ftrs gtt_rel \<Rightarrow> (nat, 'f) gtt option" where
"gtt_of_gtt_rel_impl \<F> Rs (ARoot is) = liftO1 (\<lambda>R. relabel_gtt (root_step R \<F>)) (is_to_trs' Rs is)"
| "gtt_of_gtt_rel_impl \<F> Rs (GInv g) = liftO1 prod.swap (gtt_of_gtt_rel_impl \<F> Rs g)"
| "gtt_of_gtt_rel_impl \<F> Rs (AUnion g1 g2) = liftO2 (\<lambda>g1 g2. relabel_gtt (AGTT_union' g1 g2)) (gtt_of_gtt_rel_impl \<F> Rs g1) (gtt_of_gtt_rel_impl \<F> Rs g2)"
| "gtt_of_gtt_rel_impl \<F> Rs (ATrancl g) = liftO1 (relabel_gtt \<circ> AGTT_trancl_eps_free) (gtt_of_gtt_rel_impl \<F> Rs g)"
| "gtt_of_gtt_rel_impl \<F> Rs (GTrancl g) = liftO1 GTT_trancl_eps_free (gtt_of_gtt_rel_impl \<F> Rs g)"
| "gtt_of_gtt_rel_impl \<F> Rs (AComp g1 g2) = liftO2 (\<lambda>g1 g2. relabel_gtt (AGTT_comp_eps_free g1 g2)) (gtt_of_gtt_rel_impl \<F> Rs g1) (gtt_of_gtt_rel_impl \<F> Rs g2)"
| "gtt_of_gtt_rel_impl \<F> Rs (GComp g1 g2) = liftO2 (\<lambda>g1 g2. relabel_gtt (GTT_comp_eps_free g1 g2)) (gtt_of_gtt_rel_impl \<F> Rs g1) (gtt_of_gtt_rel_impl \<F> Rs g2)"
lemma gtt_of_gtt_rel_impl_is_gtt_eps_free:
"gtt_of_gtt_rel_impl \<F> Rs g = Some g' \<Longrightarrow> is_gtt_eps_free g'"
proof (induct g arbitrary: g')
case (AUnion g1 g2)
then show ?case
by (auto simp: is_gtt_eps_free_def AGTT_union_def fmap_states_gtt_def fmap_states_ta_def ta_union_def relabel_gtt_def)
qed (auto simp: eps_free_const)
lemma gtt_of_gtt_rel_impl_gtt_of_gtt_rel:
"gtt_of_gtt_rel_impl \<F> Rs g \<noteq> None \<longleftrightarrow> gtt_of_gtt_rel \<F> Rs g \<noteq> None" (is "?Ls \<longleftrightarrow> ?Rs")
proof -
have "?Ls \<Longrightarrow> ?Rs" by (induct g) auto
moreover have "?Rs \<Longrightarrow> ?Ls" by (induct g) auto
ultimately show ?thesis by blast
qed
lemma gtt_of_gtt_rel_impl_sound:
"gtt_of_gtt_rel_impl \<F> Rs g = Some g' \<Longrightarrow> gtt_of_gtt_rel \<F> Rs g = Some g'' \<Longrightarrow> agtt_lang g' = agtt_lang g''"
proof (induct g arbitrary: g' g'')
case (ARoot x)
then show ?case by (simp add: agtt_lang_root_step_conv)
next
case (GInv g)
then have "agtt_lang (prod.swap g') = agtt_lang (prod.swap g'')" by auto
then show ?case
by (metis converse_agtt_lang converse_converse)
next
case (AUnion g1 g2)
then show ?case
by simp (metis AGTT_union'_sound option.sel)
next
case (ATrancl g)
then show ?case
using agtt_lang_AGTT_trancl_eps_free_conv[OF gtt_of_gtt_rel_impl_is_gtt_eps_free, of \<F> Rs g]
by simp (metis AGTT_trancl_sound option.sel)
next
case (GTrancl g)
then show ?case
using agtt_lang_GTT_trancl_eps_free_conv[OF gtt_of_gtt_rel_impl_is_gtt_eps_free, of \<F> Rs g]
by simp (metis GTT_trancl_alang option.sel)
next
case (AComp g1 g2)
then show ?case
using agtt_lang_AGTT_comp_eps_free_conv[OF gtt_of_gtt_rel_impl_is_gtt_eps_free, of \<F> Rs g
"the (gtt_of_gtt_rel_impl \<F> Rs g1)" "the (gtt_of_gtt_rel_impl \<F> Rs g2)"]
by simp (metis AGTT_comp'_sound agtt_lang_AGTT_comp_eps_free_conv gtt_of_gtt_rel_impl_is_gtt_eps_free option.sel)
next
case (GComp g1 g2)
then show ?case
using agtt_lang_GTT_comp_eps_free_conv[OF gtt_of_gtt_rel_impl_is_gtt_eps_free, of \<F> Rs g
"the (gtt_of_gtt_rel_impl \<F> Rs g1)" "the (gtt_of_gtt_rel_impl \<F> Rs g2)"]
by simp (metis agtt_lang_GTT_comp_eps_free_conv gtt_comp'_alang gtt_of_gtt_rel_impl_is_gtt_eps_free option.sel)
qed
(* eps free closure constructions *)
lemma \<L>_eps_free_nhole_ctxt_closure_reg:
assumes "is_ta_eps_free (ta \<A>)"
shows "\<L> (ftrancl_eps_free_reg (nhole_ctxt_closure_reg \<F> \<A>)) = \<L> (nhole_ctxt_closure_reg \<F> \<A>)"
proof -
have "eps (ta (nhole_ctxt_closure_reg \<F> \<A>)) |O| eps (ta (nhole_ctxt_closure_reg \<F> \<A>)) = {||}" using assms
by (auto simp: nhole_ctxt_closure_reg_def gen_nhole_ctxt_closure_reg_def
gen_nhole_ctxt_closure_automaton_def ta_union_def reflcl_over_nhole_ctxt_ta_def
fmap_states_reg_def is_ta_eps_free_def fmap_states_ta_def reg_Restr_Q\<^sub>f_def)
from frelcomp_empty_ftrancl_simp[OF this] show ?thesis
by (intro \<L>_ftrancl_eps_free_closuresI) simp
qed
lemma \<L>_eps_free_ctxt_closure_reg:
assumes "is_ta_eps_free (ta \<A>)"
shows "\<L> (ftrancl_eps_free_reg (ctxt_closure_reg \<F> \<A>)) = \<L> (ctxt_closure_reg \<F> \<A>)"
proof -
have "eps (ta (ctxt_closure_reg \<F> \<A>)) |O| eps (ta (ctxt_closure_reg \<F> \<A>)) = {||}" using assms
by (auto simp: ctxt_closure_reg_def gen_ctxt_closure_reg_def Let_def
gen_ctxt_closure_automaton_def ta_union_def reflcl_over_single_ta_def
fmap_states_reg_def is_ta_eps_free_def fmap_states_ta_def reg_Restr_Q\<^sub>f_def)
from frelcomp_empty_ftrancl_simp[OF this] show ?thesis
by (intro \<L>_ftrancl_eps_free_closuresI) simp
qed
lemma \<L>_eps_free_parallel_closure_reg:
assumes "is_ta_eps_free (ta \<A>)"
shows "\<L> (ftrancl_eps_free_reg (parallel_closure_reg \<F> \<A>)) = \<L> (parallel_closure_reg \<F> \<A>)"
proof -
have "eps (ta (parallel_closure_reg \<F> \<A>)) |O| eps (ta (parallel_closure_reg \<F> \<A>)) = {||}" using assms
by (auto simp: parallel_closure_reg_def gen_parallel_closure_automaton_def Let_def ta_union_def
refl_over_states_ta_def fmap_states_reg_def is_ta_eps_free_def fmap_states_ta_def reg_Restr_Q\<^sub>f_def)
from frelcomp_empty_ftrancl_simp[OF this] show ?thesis
by (intro \<L>_ftrancl_eps_free_closuresI) simp
qed
abbreviation "eps_free_reg' S R \<equiv> Reg (fin R) (eps_free_automata S (ta R))"
definition "eps_free_mctxt_closure_reg \<F> \<A> =
(let \<B> = mctxt_closure_reg \<F> \<A> in
eps_free_reg' ((\<lambda> p. (fst p, Inr cl_state)) |`| (eps (ta \<B>)) |\<union>| eps (ta \<B>)) \<B>)"
definition "eps_free_nhole_mctxt_reflcl_reg \<F> \<A> =
(let \<B> = nhole_mctxt_reflcl_reg \<F> \<A> in
eps_free_reg' ((\<lambda> p. (fst p, Inl (Inr cl_state))) |`| (eps (ta \<B>)) |\<union>| eps (ta \<B>)) \<B>)"
definition "eps_free_nhole_mctxt_closure_reg \<F> \<A> =
(let \<B> = nhole_mctxt_closure_reg \<F> \<A> in
eps_free_reg' ((\<lambda> p. (fst p, (Inr cl_state))) |`| (eps (ta \<B>)) |\<union>| eps (ta \<B>)) \<B>)"
lemma \<L>_eps_free_reg'I:
"(eps (ta \<A>))|\<^sup>+| = S \<Longrightarrow> \<L> (eps_free_reg' S \<A>) = \<L> \<A>"
by (auto simp: \<L>_def gta_lang_def gta_der_def ta_res_eps_free simp flip: eps_free)
lemma \<L>_eps_free_mctxt_closure_reg:
assumes "is_ta_eps_free (ta \<A>)"
shows "\<L> (eps_free_mctxt_closure_reg \<F> \<A>) = \<L> (mctxt_closure_reg \<F> \<A>)" using assms
unfolding eps_free_mctxt_closure_reg_def Let_def
apply (intro \<L>_eps_free_reg'I)
apply (auto simp: comp_def mctxt_closure_reg_def is_ta_eps_free_def Let_def
gen_nhole_mctxt_closure_automaton_def reflcl_over_nhole_mctxt_ta_def ta_union_def
reflcl_over_nhole_ctxt_ta_def gen_mctxt_closure_reg_def reg_Restr_Q\<^sub>f_def
fmap_states_reg_def fmap_states_ta_def dest: ftranclD ftranclD2)
by (meson fimageI finsert_iff finterI fr_into_trancl ftrancl_into_trancl)
lemma \<L>_eps_free_nhole_mctxt_reflcl_reg:
assumes "is_ta_eps_free (ta \<A>)"
shows "\<L> (eps_free_nhole_mctxt_reflcl_reg \<F> \<A>) = \<L> (nhole_mctxt_reflcl_reg \<F> \<A>)" using assms
unfolding eps_free_nhole_mctxt_reflcl_reg_def Let_def
apply (intro \<L>_eps_free_reg'I)
apply (auto simp: comp_def nhole_mctxt_reflcl_reg_def is_ta_eps_free_def Let_def
nhole_mctxt_closure_reg_def gen_nhole_mctxt_closure_reg_def reg_union_def ta_union_def
gen_nhole_mctxt_closure_automaton_def reflcl_over_nhole_mctxt_ta_def
reflcl_over_nhole_ctxt_ta_def reg_Restr_Q\<^sub>f_def fmap_states_reg_def fmap_states_ta_def dest: ftranclD ftranclD2)
by (meson fimageI finsert_iff finterI fr_into_trancl ftrancl_into_trancl)
lemma \<L>_eps_free_nhole_mctxt_closure_reg:
assumes "is_ta_eps_free (ta \<A>)"
shows "\<L> (eps_free_nhole_mctxt_closure_reg \<F> \<A>) = \<L> (nhole_mctxt_closure_reg \<F> \<A>)" using assms
unfolding eps_free_nhole_mctxt_closure_reg_def Let_def
apply (intro \<L>_eps_free_reg'I)
apply (auto simp: comp_def nhole_mctxt_closure_reg_def is_ta_eps_free_def Let_def
gen_nhole_mctxt_closure_reg_def reg_Restr_Q\<^sub>f_def fmap_states_reg_def fmap_states_ta_def
gen_nhole_mctxt_closure_automaton_def reflcl_over_nhole_mctxt_ta_def ta_union_def
reflcl_over_nhole_ctxt_ta_def dest: ftranclD ftranclD2)
by (meson fimageI finsert_iff finterI fr_into_trancl ftrancl_into_trancl)
fun rr1_of_rr1_rel_impl :: "('f \<times> nat) fset \<Rightarrow> ('f :: linorder, 'v) fin_trs list \<Rightarrow> ftrs rr1_rel \<Rightarrow> (nat, 'f) reg option"
and rr2_of_rr2_rel_impl :: "('f \<times> nat) fset \<Rightarrow> ('f, 'v) fin_trs list \<Rightarrow> ftrs rr2_rel \<Rightarrow> (nat, 'f option \<times> 'f option) reg option" where
"rr1_of_rr1_rel_impl \<F> Rs R1Terms = Some (relabel_reg (term_reg \<F>))"
| "rr1_of_rr1_rel_impl \<F> Rs (R1NF is) = liftO1 (\<lambda>R. (simplify_reg (nf_reg (fst |`| R) \<F>))) (is_to_trs' Rs is)"
| "rr1_of_rr1_rel_impl \<F> Rs (R1Inf r) = liftO1 (\<lambda>R.
let \<A> = trim_reg R in
simplify_reg (proj_1_reg (Inf_reg_impl \<A>))
) (rr2_of_rr2_rel_impl \<F> Rs r)"
| "rr1_of_rr1_rel_impl \<F> Rs (R1Proj i r) = (case i of 0 \<Rightarrow>
liftO1 (trim_reg \<circ> proj_1_reg) (rr2_of_rr2_rel_impl \<F> Rs r)
| _ \<Rightarrow> liftO1 (trim_reg \<circ> proj_2_reg) (rr2_of_rr2_rel_impl \<F> Rs r))"
| "rr1_of_rr1_rel_impl \<F> Rs (R1Union s1 s2) =
liftO2 (\<lambda> x y. relabel_reg (reg_union x y)) (rr1_of_rr1_rel_impl \<F> Rs s1) (rr1_of_rr1_rel_impl \<F> Rs s2)"
| "rr1_of_rr1_rel_impl \<F> Rs (R1Inter s1 s2) =
liftO2 (\<lambda> x y. simplify_reg (reg_intersect x y)) (rr1_of_rr1_rel_impl \<F> Rs s1) (rr1_of_rr1_rel_impl \<F> Rs s2)"
| "rr1_of_rr1_rel_impl \<F> Rs (R1Diff s1 s2) = liftO2 (\<lambda> x y. relabel_reg (trim_reg (difference_reg x y))) (rr1_of_rr1_rel_impl \<F> Rs s1) (rr1_of_rr1_rel_impl \<F> Rs s2)"
| "rr2_of_rr2_rel_impl \<F> Rs (R2GTT_Rel g w x) =
(case w of PRoot \<Rightarrow>
(case x of ESingle \<Rightarrow> liftO1 (simplify_reg \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g)
| EParallel \<Rightarrow> liftO1 (simplify_reg \<circ> reflcl_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g)
| EStrictParallel \<Rightarrow> liftO1 (simplify_reg \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g))
| PNonRoot \<Rightarrow>
(case x of ESingle \<Rightarrow> liftO1 (simplify_reg \<circ> ftrancl_eps_free_reg \<circ> nhole_ctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g)
| EParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_nhole_mctxt_reflcl_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g)
| EStrictParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_nhole_mctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g))
| PAny \<Rightarrow>
(case x of ESingle \<Rightarrow> liftO1 (simplify_reg \<circ> ftrancl_eps_free_reg \<circ> ctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g)
| EParallel \<Rightarrow> liftO1 (simplify_reg \<circ> ftrancl_eps_free_reg \<circ> parallel_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g)
| EStrictParallel \<Rightarrow> liftO1 (simplify_reg \<circ> eps_free_mctxt_closure_reg (lift_sig_RR2 |`| \<F>) \<circ> GTT_to_RR2_root_reg) (gtt_of_gtt_rel_impl \<F> Rs g)))"
| "rr2_of_rr2_rel_impl \<F> Rs (R2Diag s) =
liftO1 (\<lambda> x. fmap_funs_reg (\<lambda>f. (Some f, Some f)) x) (rr1_of_rr1_rel_impl \<F> Rs s)"
| "rr2_of_rr2_rel_impl \<F> Rs (R2Prod s1 s2) =
liftO2 (\<lambda> x y. simplify_reg (pair_automaton_reg x y)) (rr1_of_rr1_rel_impl \<F> Rs s1) (rr1_of_rr1_rel_impl \<F> Rs s2)"
| "rr2_of_rr2_rel_impl \<F> Rs (R2Inv r) = liftO1 (fmap_funs_reg prod.swap) (rr2_of_rr2_rel_impl \<F> Rs r)"
| "rr2_of_rr2_rel_impl \<F> Rs (R2Union r1 r2) =
liftO2 (\<lambda> x y. relabel_reg (reg_union x y)) (rr2_of_rr2_rel_impl \<F> Rs r1) (rr2_of_rr2_rel_impl \<F> Rs r2)"
| "rr2_of_rr2_rel_impl \<F> Rs (R2Inter r1 r2) =
liftO2 (\<lambda> x y. simplify_reg (reg_intersect x y)) (rr2_of_rr2_rel_impl \<F> Rs r1) (rr2_of_rr2_rel_impl \<F> Rs r2)"
| "rr2_of_rr2_rel_impl \<F> Rs (R2Diff r1 r2) = liftO2 (\<lambda> x y. simplify_reg (difference_reg x y)) (rr2_of_rr2_rel_impl \<F> Rs r1) (rr2_of_rr2_rel_impl \<F> Rs r2)"
| "rr2_of_rr2_rel_impl \<F> Rs (R2Comp r1 r2) = liftO2 (\<lambda> x y. simplify_reg (rr2_compositon \<F> x y))
(rr2_of_rr2_rel_impl \<F> Rs r1) (rr2_of_rr2_rel_impl \<F> Rs r2)"
lemmas ta_simp_unfold = simplify_reg_def relabel_reg_def trim_reg_def relabel_ta_def term_reg_def
lemma is_ta_eps_free_trim_reg [intro!]:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta (trim_reg R))"
by (simp add: is_ta_eps_free_def trim_reg_def trim_ta_def ta_restrict_def)
lemma is_ta_eps_free_relabel_reg [intro!]:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta (relabel_reg R))"
by (simp add: is_ta_eps_free_def relabel_reg_def relabel_ta_def fmap_states_ta_def)
lemma is_ta_eps_free_simplify_reg [intro!]:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta (simplify_reg R))"
by (simp add: is_ta_eps_free_def ta_simp_unfold fmap_states_ta_def trim_ta_def ta_restrict_def)
lemma is_ta_emptyI [simp]:
"is_ta_eps_free (TA R {||}) \<longleftrightarrow> True"
by (simp add: is_ta_eps_free_def)
lemma is_ta_empty_trim_reg:
"is_ta_eps_free (ta A) \<Longrightarrow> eps (ta (trim_reg A)) = {||}"
by (auto simp: is_ta_eps_free_def trim_reg_def trim_ta_def ta_restrict_def)
lemma is_proj_ta_eps_empty:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta (proj_1_reg R))"
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta (proj_2_reg R))"
by (auto simp: is_ta_eps_free_def proj_1_reg_def proj_2_reg_def collapse_automaton_reg_def collapse_automaton_def
fmap_funs_reg_def fmap_funs_ta_def trim_reg_def trim_ta_def ta_restrict_def)
lemma is_pod_ta_eps_empty:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta L) \<Longrightarrow> is_ta_eps_free (ta (reg_intersect R L))"
by (auto simp: reg_intersect_def prod_ta_def prod_epsRp_def prod_epsLp_def is_ta_eps_free_def)
lemma is_fmap_funs_reg_eps_empty:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta (fmap_funs_reg f R))"
by (auto simp: fmap_funs_reg_def fmap_funs_ta_def is_ta_eps_free_def)
lemma is_collapse_automaton_reg_eps_empty:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta (collapse_automaton_reg R))"
by (auto simp: collapse_automaton_reg_def collapse_automaton_def is_ta_eps_free_def)
lemma is_pair_automaton_reg_eps_empty:
"is_ta_eps_free (ta R) \<Longrightarrow> is_ta_eps_free (ta L) \<Longrightarrow> is_ta_eps_free (ta (pair_automaton_reg R L))"
by (auto simp: pair_automaton_reg_def pair_automaton_def is_ta_eps_free_def)
lemma is_reflcl_automaton_eps_free:
"is_ta_eps_free A \<Longrightarrow> is_ta_eps_free (reflcl_automaton (lift_sig_RR2 |`| \<F>) A)"
by (auto simp: is_ta_eps_free_def reflcl_automaton_def ta_union_def gen_reflcl_automaton_def Let_def fmap_states_ta_def)
lemma is_GTT_to_RR2_root_eps_empty:
"is_gtt_eps_free \<G> \<Longrightarrow> is_ta_eps_free (GTT_to_RR2_root \<G>)"
by (auto simp: is_gtt_eps_free_def GTT_to_RR2_root_def pair_automaton_def is_ta_eps_free_def)
lemma is_term_automata_eps_empty:
"is_ta_eps_free (ta (term_reg \<F>)) \<longleftrightarrow> True"
by (auto simp: is_ta_eps_free_def term_reg_def term_automaton_def)
lemma is_ta_eps_free_eps_free_automata [simp]:
" is_ta_eps_free (eps_free_automata S R) \<longleftrightarrow> True"
by (auto simp: eps_free_automata_def is_ta_eps_free_def)
lemma rr2_of_rr2_rel_impl_eps_free:
shows "\<forall> A. rr1_of_rr1_rel_impl \<F> Rs r1 = Some A \<longrightarrow> is_ta_eps_free (ta A)"
"\<forall> A. rr2_of_rr2_rel_impl \<F> Rs r2 = Some A \<longrightarrow> is_ta_eps_free (ta A)"
proof (induct r1 and r2)
case R1Terms
then show ?case
by (auto simp: is_ta_eps_free_def term_automaton_def fmap_states_ta_def ta_simp_unfold)
next
case (R1NF x)
then show ?case
by (auto simp: nf_reg_def nf_ta_def)
next
case (R1Inf x)
then show ?case
by (auto simp: Inf_reg_impl_def Let_def Inf_reg_def Inf_automata_def is_ta_empty_trim_reg intro!: is_proj_ta_eps_empty)
next
case (R1Proj n x2)
then show ?case
by (cases n) (auto intro!: is_proj_ta_eps_empty)
next
case (R1Union x1 x2)
then show ?case
by (simp add: reg_union_def ta_union_def fmap_states_ta_def is_ta_eps_free_def relabel_reg_def relabel_ta_def)
next
case (R1Inter x1 x2)
then show ?case
by (auto intro: is_pod_ta_eps_empty)
next
case (R1Diff x1 x2)
then show ?case
by (auto simp: difference_reg_def Let_def complement_reg_def ps_reg_def ps_ta_def intro!: is_pod_ta_eps_empty)
next
case (R2GTT_Rel x1 x2 x3)
then show ?case
by (cases x2; cases x3) (auto simp: GTT_to_RR2_root_reg_def ftrancl_eps_free_closures_def
eps_free_nhole_mctxt_closure_reg_def eps_free_nhole_mctxt_reflcl_reg_def Let_def
eps_free_mctxt_closure_reg_def reflcl_reg_def
dest: gtt_of_gtt_rel_impl_is_gtt_eps_free
intro!: is_GTT_to_RR2_root_eps_empty is_reflcl_automaton_eps_free)
next
case (R2Diag x)
then show ?case by (auto simp: fmap_funs_reg_def fmap_funs_ta_def is_ta_eps_free_def)
next
case (R2Prod x1 x2)
then show ?case by (auto intro: is_pair_automaton_reg_eps_empty)
next
case (R2Inv x)
then show ?case by (auto simp: fmap_funs_reg_def fmap_funs_ta_def is_ta_eps_free_def)
next
case (R2Union x1 x2)
then show ?case by (simp add: reg_union_def ta_union_def fmap_states_ta_def is_ta_eps_free_def relabel_reg_def relabel_ta_def)
next
case (R2Inter x1 x2)
then show ?case by (auto intro: is_pod_ta_eps_empty)
next
case (R2Diff x1 x2)
then show ?case by (auto simp: difference_reg_def Let_def complement_reg_def ps_reg_def ps_ta_def intro!: is_pod_ta_eps_empty)
next
case (R2Comp x1 x2)
then show ?case by (auto simp: is_term_automata_eps_empty rr2_compositon_def Let_def
intro!: is_pod_ta_eps_empty is_fmap_funs_reg_eps_empty is_collapse_automaton_reg_eps_empty is_pair_automaton_reg_eps_empty)
qed
lemma rr_of_rr_rel_impl_complete:
"rr1_of_rr1_rel_impl \<F> Rs r1 \<noteq> None \<longleftrightarrow> rr1_of_rr1_rel \<F> Rs r1 \<noteq> None"
"rr2_of_rr2_rel_impl \<F> Rs r2 \<noteq> None \<longleftrightarrow> rr2_of_rr2_rel \<F> Rs r2 \<noteq> None"
proof (induct r1 and r2)
case (R1Proj n x2)
then show ?case by (cases n) auto
next
case (R2GTT_Rel x1 n p)
then show ?case using gtt_of_gtt_rel_impl_gtt_of_gtt_rel[of \<F> Rs]
by (cases p; cases n) auto
qed auto
lemma \<Q>_fmap_funs_reg [simp]:
"\<Q>\<^sub>r (fmap_funs_reg f \<A>) = \<Q>\<^sub>r \<A>"
by (auto simp: fmap_funs_reg_def)
lemma ta_reachable_fmap_funs_reg [simp]:
"ta_reachable (ta (fmap_funs_reg f \<A>)) = ta_reachable (ta \<A>)"
by (auto simp: fmap_funs_reg_def)
lemma collapse_reg_cong:
"\<Q>\<^sub>r \<A> |\<subseteq>| ta_reachable (ta \<A>) \<Longrightarrow> \<Q>\<^sub>r \<B> |\<subseteq>| ta_reachable (ta \<B>) \<Longrightarrow> \<L> \<A> = \<L> \<B> \<Longrightarrow> \<L> (collapse_automaton_reg \<A>) = \<L> (collapse_automaton_reg \<B>)"
by (auto simp: collapse_automaton_reg_def \<L>_def collapse_automaton')
lemma \<L>_fmap_funs_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<L> (fmap_funs_reg h \<A>) = \<L> (fmap_funs_reg h \<B>)"
by (auto simp: fmap_funs_\<L>)
lemma \<L>_pair_automaton_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<L> \<C> = \<L> \<D> \<Longrightarrow> \<L> (pair_automaton_reg \<A> \<C>) = \<L> (pair_automaton_reg \<B> \<D>)"
by (auto simp: pair_automaton')
lemma \<L>_nhole_ctxt_closure_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<F> = \<G> \<Longrightarrow> \<L> (nhole_ctxt_closure_reg \<F> \<A>) = \<L> (nhole_ctxt_closure_reg \<G> \<B>)"
by (auto simp: nhole_ctxtcl_lang)
lemma \<L>_nhole_mctxt_closure_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<F> = \<G> \<Longrightarrow> \<L> (nhole_mctxt_closure_reg \<F> \<A>) = \<L> (nhole_mctxt_closure_reg \<G> \<B>)"
by (auto simp: nhole_gmctxt_closure_lang)
lemma \<L>_ctxt_closure_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<F> = \<G> \<Longrightarrow> \<L> (ctxt_closure_reg \<F> \<A>) = \<L> (ctxt_closure_reg \<G> \<B>)"
by (auto simp: gctxt_closure_lang)
lemma \<L>_parallel_closure_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<F> = \<G> \<Longrightarrow> \<L> (parallel_closure_reg \<F> \<A>) = \<L> (parallel_closure_reg \<G> \<B>)"
by (auto simp: parallelcl_gmctxt_lang)
lemma \<L>_mctxt_closure_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<F> = \<G> \<Longrightarrow> \<L> (mctxt_closure_reg \<F> \<A>) = \<L> (mctxt_closure_reg \<G> \<B>)"
by (auto simp: gmctxt_closure_lang)
lemma \<L>_nhole_mctxt_reflcl_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<F> = \<G> \<Longrightarrow> \<L> (nhole_mctxt_reflcl_reg \<F> \<A>) = \<L> (nhole_mctxt_reflcl_reg \<G> \<B>)"
unfolding nhole_mctxt_reflcl_lang
by (intro arg_cong2[where ?f = "(\<union>)"] \<L>_nhole_mctxt_closure_reg_cong) auto
declare equalityI[rule del]
declare fsubsetI[rule del]
lemma \<L>_proj_1_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<L> (proj_1_reg \<A>) = \<L> (proj_1_reg \<B>)"
by (auto simp: proj_1_reg_def \<L>_trim intro!: collapse_reg_cong \<L>_fmap_funs_reg_cong)
lemma \<L>_proj_2_reg_cong:
"\<L> \<A> = \<L> \<B> \<Longrightarrow> \<L> (proj_2_reg \<A>) = \<L> (proj_2_reg \<B>)"
by (auto simp: proj_2_reg_def \<L>_trim intro!: collapse_reg_cong \<L>_fmap_funs_reg_cong)
lemma rr2_of_rr2_rel_impl_sound:
assumes "\<forall>R \<in> set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>"
shows "\<And> A B. rr1_of_rr1_rel_impl \<F> Rs r1 = Some A \<Longrightarrow> rr1_of_rr1_rel \<F> Rs r1 = Some B \<Longrightarrow> \<L> A = \<L> B"
"\<And> A B. rr2_of_rr2_rel_impl \<F> Rs r2 = Some A \<Longrightarrow> rr2_of_rr2_rel \<F> Rs r2 = Some B \<Longrightarrow> \<L> A = \<L> B"
proof (induct r1 and r2)
case (R1Inf r)
then obtain C D where inf: "rr2_of_rr2_rel_impl \<F> Rs r = Some C" "rr2_of_rr2_rel \<F> Rs r = Some D"
"\<L> C = \<L> D" by auto
have spec: "RR2_spec C (eval_rr2_rel (fset \<F>) (map fset Rs) r)" "RR2_spec D (eval_rr2_rel (fset \<F>) (map fset Rs) r)"
using rr12_of_rr12_rel_correct(2)[OF assms, rule_format, OF inf(2)] inf(3)
by (auto simp: RR2_spec_def)
then have trim_spec: "RR2_spec (trim_reg C) (eval_rr2_rel (fset \<F>) (map fset Rs) r)"
"RR2_spec (trim_reg D) (eval_rr2_rel (fset \<F>) (map fset Rs) r)"
by (auto simp: RR2_spec_def \<L>_trim)
let ?C = "Inf_reg (trim_reg C) (Q_infty (ta (trim_reg C)) \<F>)" let ?D = "Inf_reg (trim_reg D) (Q_infty (ta (trim_reg D)) \<F>)"
from spec have *: "\<L> (Inf_reg_impl (trim_reg C)) = \<L> ?C"
using eval_rr12_rel_sig(2)[of "fset \<F>" "map fset Rs" r]
by (intro Inf_reg_impl_sound) (auto simp: RR2_spec_def \<L>_trim \<T>\<^sub>G_equivalent_def)
from spec have **: "\<L> (Inf_reg_impl (trim_reg D)) = \<L> ?D"
using eval_rr12_rel_sig(2)[of "fset \<F>" "map fset Rs" r]
by (intro Inf_reg_impl_sound) (auto simp: RR2_spec_def \<L>_trim \<T>\<^sub>G_equivalent_def)
then have C: "RR2_spec ?C {(s, t) | s t. gpair s t \<in> \<L> ?C}" and
D: "RR2_spec ?D {(s, t) | s t. gpair s t \<in> \<L> ?D}"
using subset_trans[OF Inf_automata_subseteq[of "trim_reg C" \<F>], of "\<L> C"] spec
using subset_trans[OF Inf_automata_subseteq[of "trim_reg D" \<F>], of "\<L> D"]
using eval_rr12_rel_sig(2)[of "fset \<F>" "map fset Rs" r]
by (auto simp: RR2_spec_def \<L>_trim \<T>\<^sub>G_equivalent_def intro!: equalityI fsubsetI)
from * ** have r: "\<L> (proj_1_reg (Inf_reg_impl (trim_reg C))) = \<L> (proj_1_reg ?C)"
"\<L> (proj_1_reg (Inf_reg_impl (trim_reg D))) = \<L> (proj_1_reg ?D)"
by (auto intro: \<L>_proj_1_reg_cong)
from \<L>_Inf_reg[OF trim_spec(1), of \<F>] \<L>_Inf_reg[OF trim_spec(2), of \<F>]
show ?case using R1Inf eval_rr12_rel_sig(2)[of "fset \<F>" "map fset Rs" r]
by (auto simp: liftO1_def r inf \<T>\<^sub>G_equivalent_def \<L>_proj(1)[OF C] \<L>_proj(1)[OF D])
next
case (R1Proj n x2)
then show ?case by (cases n)
(auto simp: liftO1_def \<L>_trim proj_1_reg_def proj_2_reg_def intro!: fsubsetI \<L>_fmap_funs_reg_cong collapse_reg_cong, (meson fin_mono trim_reg_reach)+)
next
case (R2GTT_Rel g p n) note IH = this
note ass = R2GTT_Rel
consider (a) "\<exists> A. gtt_of_gtt_rel_impl \<F> Rs g = Some A" | (b) "gtt_of_gtt_rel_impl \<F> Rs g = None" by blast
then show ?case
proof cases
case a then obtain C D where gtt [simp]: "gtt_of_gtt_rel_impl \<F> Rs g = Some C"
"gtt_of_gtt_rel \<F> Rs g = Some D" using gtt_of_gtt_rel_impl_gtt_of_gtt_rel by blast
from gtt_of_gtt_rel_impl_sound[OF this]
have spec [simp]: "agtt_lang C = agtt_lang D" by auto
have eps [simp]: "is_ta_eps_free (ta (GTT_to_RR2_root_reg C))"
using gtt_of_gtt_rel_impl_is_gtt_eps_free[OF gtt(1)]
by (auto simp: GTT_to_RR2_root_reg_def GTT_to_RR2_root_def pair_automaton_def is_ta_eps_free_def is_gtt_eps_free_def)
have lang: "\<L> (GTT_to_RR2_root_reg C) = \<L> (GTT_to_RR2_root_reg D)"
by (metis (no_types, lifting) GTT_to_RR2_root RR2_spec_def spec)
show ?thesis
proof (cases p)
case PRoot
then show ?thesis using IH spec lang
by (cases n) (auto simp: \<L>_eps_free \<L>_reflcl_reg)
next
case PNonRoot
then show ?thesis using IH
by (cases n) (auto simp: \<L>_eps_free \<L>_eps_free_nhole_ctxt_closure_reg[OF eps]
\<L>_eps_free_nhole_mctxt_reflcl_reg[OF eps] \<L>_eps_free_nhole_mctxt_closure_reg[OF eps]
lang intro: \<L>_nhole_ctxt_closure_reg_cong \<L>_nhole_mctxt_reflcl_reg_cong \<L>_nhole_mctxt_closure_reg_cong)
next
case PAny
then show ?thesis using IH
by (cases n) (auto simp: \<L>_eps_free \<L>_eps_free_ctxt_closure_reg[OF eps]
\<L>_eps_free_parallel_closure_reg[OF eps] \<L>_eps_free_mctxt_closure_reg[OF eps] lang
intro!: \<L>_ctxt_closure_reg_cong \<L>_parallel_closure_reg_cong \<L>_mctxt_closure_reg_cong)
qed
next
case b then show ?thesis using IH
by (cases p; cases n) auto
qed
next
case (R2Comp x1 x2)
then show ?case
by (auto simp: liftO1_def rr2_compositon_def \<L>_trim \<L>_intersect Let_def
intro!: \<L>_pair_automaton_reg_cong \<L>_fmap_funs_reg_cong collapse_reg_cong arg_cong2[where ?f = "(\<inter>)"])
qed (auto simp: liftO1_def \<L>_intersect \<L>_union \<L>_trim \<L>_difference_reg intro!: \<L>_fmap_funs_reg_cong \<L>_pair_automaton_reg_cong)
declare equalityI[intro!]
declare fsubsetI[intro!]
lemma rr12_of_rr12_rel_impl_correct:
assumes "\<forall>R \<in> set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>"
shows "\<forall>ta1. rr1_of_rr1_rel_impl \<F> Rs r1 = Some ta1 \<longrightarrow> RR1_spec ta1 (eval_rr1_rel (fset \<F>) (map fset Rs) r1)"
"\<forall>ta2. rr2_of_rr2_rel_impl \<F> Rs r2 = Some ta2 \<longrightarrow> RR2_spec ta2 (eval_rr2_rel (fset \<F>) (map fset Rs) r2)"
using rr12_of_rr12_rel_correct(1)[OF assms, of r1]
using rr12_of_rr12_rel_correct(2)[OF assms, of r2]
using rr2_of_rr2_rel_impl_sound(1)[OF assms, of r1]
using rr2_of_rr2_rel_impl_sound(2)[OF assms, of r2]
using rr_of_rr_rel_impl_complete(1)[of \<F> Rs r1]
using rr_of_rr_rel_impl_complete(2)[of \<F> Rs r2]
by (force simp: RR1_spec_def RR2_spec_def)+
lemma check_inference_rrn_impl_correct:
assumes sig: "\<T>\<^sub>G (fset \<F>) \<noteq> {}" and Rs: "\<forall>R \<in> set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>"
assumes infs: "\<And>fvA. fvA \<in> set infs \<Longrightarrow> formula_spec (fset \<F>) (map fset Rs) (fst (snd fvA)) (snd (snd fvA)) (fst fvA)"
assumes inf: "check_inference rr1_of_rr1_rel_impl rr2_of_rr2_rel_impl \<F> Rs infs (l, step, fm, is) = Some (fm', vs, A')"
shows "l = length infs \<and> fm = fm' \<and> formula_spec (fset \<F>) (map fset Rs) vs A' fm'"
using check_inference_correct[where ?rr1c = rr1_of_rr1_rel_impl and ?rr2c = rr2_of_rr2_rel_impl, OF assms]
using rr12_of_rr12_rel_impl_correct[OF Rs]
by auto
definition check_sig_nempty where
"check_sig_nempty \<F> = (0 |\<in>| snd |`| \<F>)"
definition check_trss where
"check_trss \<R> \<F> = list_all (\<lambda> R. lv_trs (fset R) \<and> funas_trs (fset R) \<subseteq> fset \<F>) \<R>"
lemma check_sig_nempty:
"check_sig_nempty \<F> \<longleftrightarrow> \<T>\<^sub>G (fset \<F>) \<noteq> {}" (is "?Ls \<longleftrightarrow> ?Rs")
proof -
{assume ?Ls then obtain a where "(a, 0) |\<in>| \<F>" by (auto simp: check_sig_nempty_def)
then have "GFun a [] \<in> \<T>\<^sub>G (fset \<F>)"
- by (intro const) (simp add: fmember_iff_member_fset)
+ by (intro const) simp
then have ?Rs by blast}
moreover
{assume ?Rs then obtain s where "s \<in> \<T>\<^sub>G (fset \<F>)" by blast
- then obtain a where "(a, 0) |\<in>| \<F>" unfolding fmember_iff_member_fset
+ then obtain a where "(a, 0) |\<in>| \<F>"
by (induct s) (auto, force)
then have ?Ls unfolding check_sig_nempty_def
- by (auto simp: fimage_iff fBex_def)}
+ by (auto simp: image_iff Bex_def)}
ultimately show ?thesis by blast
qed
lemma check_trss:
"check_trss \<R> \<F> \<longleftrightarrow> (\<forall> R \<in> set \<R>. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>)"
unfolding check_trss_def list_all_iff
- by (auto simp: fmember_iff_member_fset ffunas_trs.rep_eq less_eq_fset.rep_eq)
+ by (auto simp: ffunas_trs.rep_eq less_eq_fset.rep_eq)
fun check_inference_list :: "('f \<times> nat) fset \<Rightarrow> ('f :: {compare,linorder}, 'v) fin_trs list
\<Rightarrow> (nat \<times> ftrs inference \<times> ftrs formula \<times> info list) list
\<Rightarrow> (ftrs formula \<times> nat list \<times> (nat, 'f option list) reg) list option" where
"check_inference_list \<F> Rs infs = do {
guard (check_sig_nempty \<F>);
guard (check_trss Rs \<F>);
foldl (\<lambda> tas inf. do {
tas' \<leftarrow> tas;
r \<leftarrow> check_inference rr1_of_rr1_rel_impl rr2_of_rr2_rel_impl \<F> Rs tas' inf;
Some (tas' @ [r])
})
(Some []) infs
}"
lemma check_inference_list_correct:
assumes "check_inference_list \<F> Rs infs = Some fvAs"
shows "length infs = length fvAs \<and> (\<forall> i < length fvAs. fst (snd (snd (infs ! i))) = fst (fvAs ! i)) \<and>
(\<forall> i < length fvAs. formula_spec (fset \<F>) (map fset Rs) (fst (snd (fvAs ! i))) (snd (snd (fvAs ! i))) (fst (fvAs ! i)))"
using assms
proof (induct infs arbitrary: fvAs rule: rev_induct)
note [simp] = bind_eq_Some_conv guard_simps
{case Nil
then show ?case by auto
next
case (snoc a infs)
have inv: "\<T>\<^sub>G (fset \<F>) \<noteq> {}" "\<forall>R\<in>set Rs. lv_trs (fset R) \<and> ffunas_trs R |\<subseteq>| \<F>"
using snoc(2) by (auto simp: check_sig_nempty check_trss)
from snoc(2) obtain fvAs' l steps fm fm' is' vs A' where
ch: "check_inference_list \<F> Rs infs = Some fvAs'" "a = (l, steps, fm, is')"
"check_inference rr1_of_rr1_rel_impl rr2_of_rr2_rel_impl \<F> Rs fvAs' (l, steps, fm, is') = Some (fm', vs, A')" "fvAs = fvAs' @ [(fm', vs, A')]"
by (auto simp del: check_inference.simps) (metis prod_cases4)
from snoc(1)[OF ch(1)] have "fvA \<in> set fvAs' \<Longrightarrow> formula_spec (fset \<F>) (map fset Rs) (fst (snd fvA)) (snd (snd fvA)) (fst fvA)" for fvA
by (auto dest: in_set_idx)
from check_inference_rrn_impl_correct[OF inv this, of fvAs'] this
show ?case using snoc(1)[OF ch(1)] ch
by (auto simp del: check_inference.simps simp: nth_append)
}
qed
fun check_certificate where
"check_certificate \<F> Rs A fm (Certificate infs claim n) = do {
guard (n < length infs);
guard (A \<longleftrightarrow> claim = Nonempty);
guard (fm = fst (snd (snd (infs ! n))));
fvA \<leftarrow> check_inference_list \<F> Rs (take (Suc n) infs);
(let E = reg_empty (snd (snd (last fvA))) in
case claim of Empty \<Rightarrow> Some E
| _ \<Rightarrow> Some (\<not> E))
}"
definition formula_unsatisfiable where
"formula_unsatisfiable \<F> Rs fm \<longleftrightarrow> (formula_satisfiable \<F> Rs fm = False)"
definition correct_certificate where
"correct_certificate \<F> Rs claim infs n \<equiv>
(claim = Empty \<longleftrightarrow> (formula_unsatisfiable (fset \<F>) (map fset Rs) (fst (snd (snd (infs ! n))))) \<and>
claim = Nonempty \<longleftrightarrow> formula_satisfiable (fset \<F>) (map fset Rs) (fst (snd (snd (infs ! n)))))"
lemma check_certificate_sound:
assumes "check_certificate \<F> Rs A fm (Certificate infs claim n) = Some B"
shows "fm = fst (snd (snd (infs ! n)))" "A \<longleftrightarrow> claim = Nonempty"
using assms by (auto simp: bind_eq_Some_conv guard_simps)
lemma check_certificate_correct:
assumes "check_certificate \<F> Rs A fm (Certificate infs claim n) = Some B"
shows "(B = True \<longrightarrow> correct_certificate \<F> Rs claim infs n) \<and>
(B = False \<longrightarrow> correct_certificate \<F> Rs (case_claim Nonempty Empty claim) infs n)"
proof -
note [simp] = bind_eq_Some_conv guard_simps
from assms obtain fvAs where inf: "check_inference_list \<F> Rs (take (Suc n) infs) = Some fvAs"
by auto
from assms have len: "n < length infs" by auto
from check_inference_list_correct[OF inf] have
inv: "length fvAs = n + 1"
"fst (snd (snd (infs ! n))) = fst (fvAs ! n)"
"formula_spec (fset \<F>) (map fset Rs) (fst (snd (last fvAs))) (snd (snd (last fvAs))) (fst (last fvAs))"
using len last_conv_nth[of fvAs] by force+
have nth: "fst (last fvAs) = fst (fvAs ! n)" using inv(1)
using len last_conv_nth[of fvAs] by force
note spec = formula_spec_empty[OF _ inv(3)] formula_spec_nt_empty_form_sat[OF _ inv(3)]
consider (a) "claim = Empty" | (b) "claim = Nonempty" using claim.exhaust by blast
then show ?thesis
proof cases
case a
then have *: "B = reg_empty (snd (snd (last fvAs)))" using inv
using assms len last_conv_nth[of fvAs]
by (auto simp: inf simp del: check_inference_list.simps)
show ?thesis using a inv spec unfolding *
by (auto simp: formula_satisfiable_def nth correct_certificate_def formula_unsatisfiable_def simp del: reg_empty)
next
case b
then have *: "B \<longleftrightarrow> \<not> (reg_empty (snd (snd (last fvAs))))" using inv
using assms len last_conv_nth[of fvAs]
by (auto simp: inf simp del: check_inference_list.simps)
show ?thesis using b inv spec unfolding *
by (auto simp: formula_satisfiable_def nth formula_unsatisfiable_def correct_certificate_def simp del: reg_empty)
qed
qed
definition check_certificate_string ::
"(integer list \<times> fvar) fset \<Rightarrow>
((integer list, integer list) Term.term \<times> (integer list, integer list) Term.term) fset list \<Rightarrow>
bool \<Rightarrow> ftrs formula \<Rightarrow> ftrs certificate \<Rightarrow> bool option"
where "check_certificate_string = check_certificate"
(***********************************)
export_code check_certificate_string Var Fun fset_of_list nat_of_integer Certificate
R2GTT_Rel R2Eq R2Reflc R2Step R2StepEq R2Steps R2StepsEq R2StepsNF R2ParStep R2RootStep
R2RootStepEq R2RootSteps R2RootStepsEq R2NonRootStep R2NonRootStepEq R2NonRootSteps
R2NonRootStepsEq R2Meet R2Join
ARoot GSteps PRoot ESingle Empty Size EDistribAndOr
R1Terms R1Fin
FRR1 FRestrict FTrue FFalse
IRR1 Fwd in Haskell module_name FOR
end
\ No newline at end of file
diff --git a/thys/FO_Theory_Rewriting/Primitives/LV_to_GTT.thy b/thys/FO_Theory_Rewriting/Primitives/LV_to_GTT.thy
--- a/thys/FO_Theory_Rewriting/Primitives/LV_to_GTT.thy
+++ b/thys/FO_Theory_Rewriting/Primitives/LV_to_GTT.thy
@@ -1,330 +1,330 @@
section \<open>Primitive constructions\<close>
theory LV_to_GTT
imports Regular_Tree_Relations.Pair_Automaton
Bot_Terms
Rewriting
begin
subsection \<open>Recognizing subterms of linear terms\<close>
(* Pattern recognizer automaton *)
abbreviation ffunas_terms where
"ffunas_terms R \<equiv> |\<Union>| (ffunas_term |`| R)"
definition "states R \<equiv> {t\<^sup>\<bottom> | s t. s \<in> R \<and> s \<unrhd> t}"
lemma states_conv:
"states R = term_to_bot_term ` (\<Union> s \<in> R. subterms s)"
unfolding states_def set_all_subteq_subterms
by auto
lemma finite_states:
assumes "finite R" shows "finite (states R)"
proof -
have conv: "states R = term_to_bot_term ` (\<Union> s \<in> R. {t | t. s \<unrhd> t})"
unfolding states_def by auto
from assms have "finite (\<Union> s \<in> R. {t | t. s \<unrhd> t})"
by (intro finite_UN_I2 finite_imageI) (simp add: finite_subterms)+
then show ?thesis using conv by auto
qed
lemma root_bot_diff:
"root_bot ` (R - {Bot}) = (root_bot ` R) - {None}"
using root_bot.elims by auto
lemma root_bot_states_root_subterms:
"the ` (root_bot ` (states R - {Bot})) = the ` (root ` (\<Union> s \<in> R. subterms s) - {None})"
unfolding states_conv root_bot_diff
unfolding image_comp
by simp
context
includes fset.lifting
begin
lift_definition fstates :: "('f, 'v) term fset \<Rightarrow> 'f bot_term fset" is states
by (simp add: finite_states)
lift_definition fsubterms :: "('f, 'v) term \<Rightarrow> ('f, 'v) term fset" is subterms
by (simp add: finite_subterms_fun)
lemmas fsubterms [code] = subterms.simps[Transfer.transferred]
lift_definition ffunas_trs :: "(('f, 'v) term \<times> ('f, 'v) term) fset \<Rightarrow> ('f \<times> nat) fset" is funas_trs
by (simp add: finite_funas_trs)
lemma fstates_def':
"t |\<in>| fstates R \<longleftrightarrow> (\<exists> s u. s |\<in>| R \<and> s \<unrhd> u \<and> u\<^sup>\<bottom> = t)"
by transfer (auto simp: states_def)
lemma fstates_fmemberE [elim!]:
assumes "t |\<in>| fstates R"
obtains s u where "s |\<in>| R \<and> s \<unrhd> u \<and> u\<^sup>\<bottom> = t"
using assms unfolding fstates_def'
by blast
lemma fstates_fmemberI [intro]:
"s |\<in>| R \<Longrightarrow> s \<unrhd> u \<Longrightarrow> u\<^sup>\<bottom> |\<in>| fstates R"
unfolding fstates_def' by blast
lemmas froot_bot_states_root_subterms = root_bot_states_root_subterms[Transfer.transferred]
lemmas root_fsubsterms_ffunas_term_fset = root_substerms_funas_term_set[Transfer.transferred]
lemma fstates[code]:
"fstates R = term_to_bot_term |`| ( |\<Union>| (fsubterms |`| R))"
by transfer (auto simp: states_conv)
end
definition ta_rule_sig where
"ta_rule_sig = (\<lambda> r. (r_root r, length (r_lhs_states r)))"
primrec term_to_ta_rule where
"term_to_ta_rule (BFun f ts) = TA_rule f ts (BFun f ts)"
lemma ta_rule_sig_term_to_ta_rule_root:
"t \<noteq> Bot \<Longrightarrow> ta_rule_sig (term_to_ta_rule t) = the (root_bot t)"
by (cases t) (auto simp: ta_rule_sig_def)
lemma ta_rule_sig_term_to_ta_rule_root_set:
assumes "Bot |\<notin>| R"
shows "ta_rule_sig |`| (term_to_ta_rule |`| R) = the |`| (root_bot |`| R)"
proof -
{fix x assume "x |\<in>| R" then have "ta_rule_sig (term_to_ta_rule x) = the (root_bot x)"
using ta_rule_sig_term_to_ta_rule_root[of x] assms
by auto}
then show ?thesis
by (force simp: fimage_iff)
qed
definition pattern_automaton_rules where
"pattern_automaton_rules \<F> R =
(let states = (fstates R) - {|Bot|} in
term_to_ta_rule |`| states |\<union>| (\<lambda> (f, n). TA_rule f (replicate n Bot) Bot) |`| \<F>)"
lemma pattern_automaton_rules_BotD:
assumes "TA_rule f ss Bot |\<in>| pattern_automaton_rules \<F> R"
shows "TA_rule f ss Bot |\<in>| (\<lambda> (f, n). TA_rule f (replicate n Bot) Bot) |`| \<F>" using assms
by (auto simp: pattern_automaton_rules_def)
(metis ta_rule.inject term_to_bot_term.elims term_to_ta_rule.simps)
lemma pattern_automaton_rules_FunD:
assumes "TA_rule f ss (BFun g ts) |\<in>| pattern_automaton_rules \<F> R"
shows "g = f \<and> ts = ss \<and>
TA_rule f ss (BFun g ts) |\<in>| term_to_ta_rule |`| ((fstates R) - {|Bot|})" using assms
apply (auto simp: pattern_automaton_rules_def)
apply (metis bot_term.exhaust ta_rule.inject term_to_ta_rule.simps)
by (metis (no_types, lifting) ta_rule.inject term_to_bot_term.elims term_to_ta_rule.simps)
definition pattern_automaton where
"pattern_automaton \<F> R = TA (pattern_automaton_rules \<F> R) {||}"
lemma ta_sig_pattern_automaton [simp]:
"ta_sig (pattern_automaton \<F> R) = \<F> |\<union>| ffunas_terms R"
proof -
let ?r = "ta_rule_sig"
have *:"Bot |\<notin>| (fstates R) - {|Bot|}" by simp
have f: "\<F> = ?r |`| ((\<lambda> (f, n). TA_rule f (replicate n Bot) Bot) |`| \<F>)"
- by (auto simp: fimage_iff fBex_def ta_rule_sig_def split!: prod.splits)
+ by (auto simp: image_iff Bex_def ta_rule_sig_def split!: prod.splits)
moreover have "ffunas_terms R = ?r |`| (term_to_ta_rule |`| ((fstates R) - {|Bot|}))"
unfolding ta_rule_sig_term_to_ta_rule_root_set[OF *]
unfolding froot_bot_states_root_subterms root_fsubsterms_ffunas_term_fset
by simp
ultimately show ?thesis unfolding pattern_automaton_def ta_sig_def
unfolding ta_rule_sig_def pattern_automaton_rules_def
- by (auto simp add: Let_def comp_def fimage_funion)
+ by (simp add: fimage_funion comp_def) blast
qed
lemma terms_reach_Bot:
assumes "ffunas_gterm t |\<subseteq>| \<F>"
shows "Bot |\<in>| ta_der (pattern_automaton \<F> R) (term_of_gterm t)" using assms
proof (induct t)
case (GFun f ts)
have [simp]: "s \<in> set ts \<Longrightarrow> ffunas_gterm s |\<subseteq>| \<F>" for s using GFun(2)
using in_set_idx by fastforce
from GFun show ?case
by (auto simp: pattern_automaton_def pattern_automaton_rules_def rev_fimage_eqI
intro!: exI[of _ "replicate (length ts) Bot"])
qed
lemma pattern_automaton_reach_smaller_term:
assumes "l |\<in>| R" "l \<unrhd> s" "s\<^sup>\<bottom> \<le>\<^sub>b (term_of_gterm t)\<^sup>\<bottom>" "ffunas_gterm t |\<subseteq>| \<F>"
shows "s\<^sup>\<bottom> |\<in>| ta_der (pattern_automaton \<F> R) (term_of_gterm t)" using assms(2-)
proof (induct t arbitrary: s)
case (GFun f ts) show ?case
proof (cases s)
case (Var x)
then show ?thesis using terms_reach_Bot[OF GFun(4)]
by (auto simp del: ta_der_Fun)
next
case [simp]: (Fun g ss)
let ?ss = "map term_to_bot_term ss"
have [simp]: "s \<in> set ts \<Longrightarrow> ffunas_gterm s |\<subseteq>| \<F>" for s using GFun(4)
using in_set_idx by fastforce
from GFun(3) have s: "g = f" "length ss = length ts" by auto
from GFun(2) s(2) assms(1) have rule: "TA_rule f ?ss (BFun f ?ss) |\<in>| pattern_automaton_rules \<F> R"
- by (auto simp: s(1) pattern_automaton_rules_def fimage_iff fBex_def)
+ by (auto simp: s(1) pattern_automaton_rules_def image_iff Bex_def)
{fix i assume bound: "i < length ts"
then have sub: "l \<unrhd> ss ! i" using GFun(2) arg_subteq[OF nth_mem, of i ss f]
unfolding Fun s(1) using s(2) by (metis subterm.order.trans)
have "ss ! i\<^sup>\<bottom> \<le>\<^sub>b (term_of_gterm (ts ! i):: ('a, 'c) term)\<^sup>\<bottom>" using GFun(3) bound s(2)
by (auto simp: s elim!: bless_eq.cases)
from GFun(1)[OF nth_mem sub this] bound
have "ss ! i\<^sup>\<bottom> |\<in>| ta_der (pattern_automaton \<F> R) (term_of_gterm (ts ! i))"
by auto}
then show ?thesis using GFun(2-) s(2) rule
by (auto simp: s(1) pattern_automaton_def intro!: exI[of _ ?ss] exI[of _ "BFun f ?ss"])
qed
qed
lemma bot_term_of_gterm_conv:
"term_of_gterm s\<^sup>\<bottom> = term_of_gterm s\<^sup>\<bottom>"
by (induct s) auto
lemma pattern_automaton_ground_instance_reach:
assumes "l |\<in>| R" "l \<cdot> \<sigma> = (term_of_gterm t)" "ffunas_gterm t |\<subseteq>| \<F>"
shows "l\<^sup>\<bottom> |\<in>| ta_der (pattern_automaton \<F> R) (term_of_gterm t)"
proof -
let ?t = "(term_of_gterm t) :: ('a, 'a bot_term) term"
from instance_to_bless_eq[OF assms(2)] have sm: "l\<^sup>\<bottom> \<le>\<^sub>b ?t\<^sup>\<bottom>"
using bot_term_of_gterm_conv by metis
show ?thesis using pattern_automaton_reach_smaller_term[OF assms(1) _ sm] assms(3-)
by auto
qed
lemma pattern_automaton_reach_smallet_term:
assumes "l\<^sup>\<bottom> |\<in>| ta_der (pattern_automaton \<F> R) t" "ground t"
shows "l\<^sup>\<bottom> \<le>\<^sub>b t\<^sup>\<bottom>" using assms
proof (induct t arbitrary: l)
case (Fun f ts) note IH = this show ?case
proof (cases l)
case (Fun g ss)
let ?ss = "map term_to_bot_term ss"
from IH(2) have rule: "g = f" "length ss = length ts"
"TA_rule f ?ss (BFun f ?ss) |\<in>| rules (pattern_automaton \<F> R)"
by (auto simp: Fun pattern_automaton_def dest: pattern_automaton_rules_FunD)
{fix i assume "i < length ts"
then have "ss ! i\<^sup>\<bottom> \<le>\<^sub>b ts ! i\<^sup>\<bottom>" using IH(2, 3) rule(2)
by (intro IH(1)) (auto simp: Fun pattern_automaton_def dest: pattern_automaton_rules_FunD)}
then show ?thesis using rule(2)
by (auto simp: Fun rule(1))
qed auto
qed auto
subsection \<open>Recognizing root step relation of LV-TRSs\<close>
definition lv_trs :: "('f, 'v) trs \<Rightarrow> bool" where
"lv_trs R \<equiv> \<forall>(l, r) \<in> R. linear_term l \<and> linear_term r \<and> (vars_term l \<inter> vars_term r = {})"
lemma subst_unification:
assumes "vars_term s \<inter> vars_term t = {}"
obtains \<mu> where "s \<cdot> \<sigma> = s \<cdot> \<mu>" "t \<cdot> \<tau> = t \<cdot> \<mu>"
using assms
by (auto intro!: that[of "\<lambda>x. if x \<in> vars_term s then \<sigma> x else \<tau> x"] simp: term_subst_eq_conv)
lemma lv_trs_subst_unification:
assumes "lv_trs R" "(l, r) \<in> R" "s = l \<cdot> \<sigma>" "t = r \<cdot> \<tau>"
obtains \<mu> where "s = l \<cdot> \<mu> \<and> t = r \<cdot> \<mu>"
using assms subst_unification[of l r \<sigma> \<tau>]
unfolding lv_trs_def
by (force split!: prod.splits)
definition Rel\<^sub>f where
"Rel\<^sub>f R = map_both term_to_bot_term |`| R"
definition root_pair_automaton where
"root_pair_automaton \<F> R = (pattern_automaton \<F> (fst |`| R),
pattern_automaton \<F> (snd |`| R))"
definition agtt_grrstep where
"agtt_grrstep \<R> \<F> = pair_at_to_agtt' (root_pair_automaton \<F> \<R>) (Rel\<^sub>f \<R>)"
lemma agtt_grrstep_eps_trancl [simp]:
"(eps (fst (agtt_grrstep \<R> \<F>)))|\<^sup>+| = eps (fst (agtt_grrstep \<R> \<F>))"
"(eps (snd (agtt_grrstep \<R> \<F>))) = {||}"
by (auto simp add: agtt_grrstep_def pair_at_to_agtt'_def
pair_at_to_agtt_def Let_def root_pair_automaton_def pattern_automaton_def
fmap_states_ta_def intro!: frelcomp_empty_ftrancl_simp)
lemma root_pair_automaton_grrstep:
fixes R :: "('f, 'v) rule fset"
assumes "lv_trs (fset R)" "ffunas_trs R |\<subseteq>| \<F>"
shows "pair_at_lang (root_pair_automaton \<F> R) (Rel\<^sub>f R) = Restr (grrstep (fset R)) (\<T>\<^sub>G (fset \<F>))" (is "?Ls = ?Rs")
proof
let ?t_o_g = "term_of_gterm :: 'f gterm \<Rightarrow> ('f, 'v) Term.term"
have [simp]: "\<F> |\<union>| |\<Union>| ((ffunas_term \<circ> fst) |`| R) = \<F>"
"\<F> |\<union>| |\<Union>| ((ffunas_term \<circ> snd) |`| R) = \<F>" using assms(2)
- by (force simp: less_eq_fset.rep_eq ffunas_trs.rep_eq funas_trs_def ffunas_term.rep_eq fmember_iff_member_fset ffUnion.rep_eq)+
+ by (force simp: less_eq_fset.rep_eq ffunas_trs.rep_eq funas_trs_def ffunas_term.rep_eq ffUnion.rep_eq)+
{fix s t assume "(s, t) \<in> ?Ls"
from pair_at_langE[OF this] obtain p q where st: "(q, p) |\<in>| Rel\<^sub>f R"
"q |\<in>| gta_der (fst (root_pair_automaton \<F> R)) s" "p |\<in>| gta_der (snd (root_pair_automaton \<F> R)) t"
by blast
from st(1) obtain l r where tm: "q = l\<^sup>\<bottom>" "p = r\<^sup>\<bottom>" "(l, r) |\<in>| R" unfolding Rel\<^sub>f_def
- using assms(1) by (auto simp: fmember.abs_eq)
+ using assms(1) by auto
have sm: "l\<^sup>\<bottom> \<le>\<^sub>b (?t_o_g s)\<^sup>\<bottom>" "r\<^sup>\<bottom> \<le>\<^sub>b (?t_o_g t)\<^sup>\<bottom>"
using pattern_automaton_reach_smallet_term[of l \<F> "fst |`| R" "term_of_gterm s"]
using pattern_automaton_reach_smallet_term[of r \<F> "snd |`| R" "term_of_gterm t"]
using st(2, 3) tm(3) unfolding tm
by (auto simp: gta_der_def root_pair_automaton_def) (smt bot_term_of_gterm_conv)+
have "linear_term l" "linear_term r" using tm(3) assms(1)
- by (auto simp: lv_trs_def fmember_iff_member_fset)
+ by (auto simp: lv_trs_def)
then obtain \<sigma> \<tau> where "l \<cdot> \<sigma> = ?t_o_g s" "r \<cdot> \<tau> = ?t_o_g t" using sm
by (auto dest!: bless_eq_to_instance)
then obtain \<mu> where subst: "l \<cdot> \<mu> = ?t_o_g s" "r \<cdot> \<mu> = ?t_o_g t"
- using lv_trs_subst_unification[OF assms(1) tm(3)[unfolded fmember_iff_member_fset], of "?t_o_g s" \<sigma> "?t_o_g t" \<tau>]
+ using lv_trs_subst_unification[OF assms(1) tm(3), of "?t_o_g s" \<sigma> "?t_o_g t" \<tau>]
by metis
moreover have "s \<in> \<T>\<^sub>G (fset \<F>)" "t \<in> \<T>\<^sub>G (fset \<F>)" using st(2-) assms
using ta_der_gterm_sig[of q "pattern_automaton \<F> (fst |`| R)" s]
using ta_der_gterm_sig[of p "pattern_automaton \<F> (snd |`| R)" t]
by (auto simp: gta_der_def root_pair_automaton_def \<T>\<^sub>G_equivalent_def less_eq_fset.rep_eq ffunas_gterm.rep_eq)
ultimately have "(s, t) \<in> ?Rs" using tm(3)
- by (auto simp: grrstep_def rrstep_def' fmember_iff_member_fset) metis}
+ by (auto simp: grrstep_def rrstep_def') metis}
then show "?Ls \<subseteq> ?Rs" by auto
next
let ?t_o_g = "term_of_gterm :: 'f gterm \<Rightarrow> ('f, 'v) Term.term"
{fix s t assume "(s, t) \<in> ?Rs"
then obtain \<sigma> l r where st: "(l, r) |\<in>| R" "l \<cdot> \<sigma> = ?t_o_g s" "r \<cdot> \<sigma> = ?t_o_g t" "s \<in> \<T>\<^sub>G (fset \<F>)" "t \<in> \<T>\<^sub>G (fset \<F>)"
- by (auto simp: grrstep_def rrstep_def' fmember_iff_member_fset)
+ by (auto simp: grrstep_def rrstep_def')
have funas: "ffunas_gterm s |\<subseteq>| \<F>" "ffunas_gterm t |\<subseteq>| \<F>" using st(4, 5)
by (auto simp: \<T>\<^sub>G_equivalent_def)
- (metis ffunas_gterm.rep_eq notin_fset subsetD)+
+ (metis ffunas_gterm.rep_eq subsetD)+
from st(1) have "(l\<^sup>\<bottom>, r\<^sup>\<bottom>) |\<in>| Rel\<^sub>f R" unfolding Rel\<^sub>f_def using assms(1)
by (auto simp: fimage_iff fBex_def)
then have "(s, t) \<in> ?Ls" using st
using pattern_automaton_ground_instance_reach[of l "fst |`| R" \<sigma>, OF _ _ funas(1)]
using pattern_automaton_ground_instance_reach[of r "snd |`| R" \<sigma>, OF _ _ funas(2)]
- by (auto simp: \<T>\<^sub>G_equivalent_def fimage_iff fBex_def fmember.abs_eq root_pair_automaton_def gta_der_def pair_at_lang_def)}
+ by (auto simp: \<T>\<^sub>G_equivalent_def image_iff Bex_def root_pair_automaton_def gta_der_def pair_at_lang_def)}
then show "?Rs \<subseteq> ?Ls" by auto
qed
lemma agtt_grrstep:
fixes R :: "('f, 'v) rule fset"
assumes "lv_trs (fset R)" "ffunas_trs R |\<subseteq>| \<F>"
shows "agtt_lang (agtt_grrstep R \<F>) = Restr (grrstep (fset R)) (\<T>\<^sub>G (fset \<F>))"
using root_pair_automaton_grrstep[OF assms] unfolding pair_at_agtt_cost agtt_grrstep_def
by simp
(* Results for set as input *)
lemma root_pair_automaton_grrstep_set:
fixes R :: "('f, 'v) rule set"
assumes "finite R" "finite \<F>" "lv_trs R" "funas_trs R \<subseteq> \<F>"
shows "pair_at_lang (root_pair_automaton (Abs_fset \<F>) (Abs_fset R)) (Rel\<^sub>f (Abs_fset R)) = Restr (grrstep R) (\<T>\<^sub>G \<F>)"
proof -
from assms(1, 2, 4) have "ffunas_trs (Abs_fset R) |\<subseteq>| Abs_fset \<F>"
- by (auto simp add: Abs_fset_inverse ffunas_trs.rep_eq fmember_iff_member_fset subset_eq)
+ by (auto simp add: Abs_fset_inverse ffunas_trs.rep_eq subset_eq)
from root_pair_automaton_grrstep[OF _ this] assms
show ?thesis
by (auto simp: Abs_fset_inverse)
qed
lemma agtt_grrstep_set:
fixes R :: "('f, 'v) rule set"
assumes "finite R" "finite \<F>" "lv_trs R" "funas_trs R \<subseteq> \<F>"
shows "agtt_lang (agtt_grrstep (Abs_fset R) (Abs_fset \<F>)) = Restr (grrstep R) (\<T>\<^sub>G \<F>)"
using root_pair_automaton_grrstep_set[OF assms] unfolding pair_at_agtt_cost agtt_grrstep_def
by simp
end
diff --git a/thys/FO_Theory_Rewriting/Primitives/NF.thy b/thys/FO_Theory_Rewriting/Primitives/NF.thy
--- a/thys/FO_Theory_Rewriting/Primitives/NF.thy
+++ b/thys/FO_Theory_Rewriting/Primitives/NF.thy
@@ -1,295 +1,296 @@
theory NF
imports
Saturation
Bot_Terms
Regular_Tree_Relations.Tree_Automata
begin
subsection \<open>Recognizing normal forms of left linear TRSs\<close>
interpretation lift_total: semilattice_closure_partial_operator "\<lambda> x y. (x, y) \<in> mergeP" "(\<up>)" "\<lambda> x y. x \<le>\<^sub>b y" Bot
apply unfold_locales apply (auto simp: merge_refl merge_symmetric merge_terms_assoc merge_terms_idem merge_bot_args_bless_eq_merge)
using merge_dist apply blast
using megeP_ass apply blast
using merge_terms_commutative apply blast
apply (metis bless_eq_mergeP bless_eq_trans merge_bot_args_bless_eq_merge merge_dist merge_symmetric merge_terms_commutative)
apply (metis merge_bot_args_bless_eq_merge merge_symmetric merge_terms_commutative)
using bless_eq_closued_under_supremum bless_eq_trans bless_eq_anti_sym
by blast+
abbreviation "psubt_lhs_bot R \<equiv> {t\<^sup>\<bottom> | s t. s \<in> R \<and> s \<rhd> t}"
abbreviation "closure S \<equiv> lift_total.cl.pred_closure S"
definition states where
"states R = insert Bot (closure (psubt_lhs_bot R))"
lemma psubt_mono:
"R \<subseteq> S \<Longrightarrow> psubt_lhs_bot R \<subseteq> psubt_lhs_bot S" by auto
lemma states_mono:
"R \<subseteq> S \<Longrightarrow> states R \<subseteq> states S"
unfolding states_def using lift_total.cl.closure_mono[OF psubt_mono[of R S]]
by auto
lemma finite_lhs_subt [simp, intro]:
assumes "finite R"
shows "finite (psubt_lhs_bot R)"
proof -
have conv: "psubt_lhs_bot R = term_to_bot_term ` {t | s t . s \<in> R \<and> s \<rhd> t}" by auto
from assms have "finite {t | s t . s \<in> R \<and> s \<rhd> t}"
by (simp add: finite_strict_subterms)
then show ?thesis using conv by auto
qed
lemma states_ref_closure:
"states R \<subseteq> insert Bot (closure (psubt_lhs_bot R))"
unfolding states_def by auto
lemma finite_R_finite_states [simp, intro]:
"finite R \<Longrightarrow> finite (states R)"
using finite_lhs_subt states_ref_closure
using lift_total.cl.finite_S_finite_closure finite_subset
by fastforce
abbreviation "lift_sup_small s S \<equiv> lift_total.supremum (lift_total.smaller_subset (Some s) (Some ` S))"
abbreviation "bound_max s S \<equiv> the (lift_sup_small s S)"
lemma bound_max_state_set:
assumes "finite R"
shows "bound_max t (psubt_lhs_bot R) \<in> states R"
using lift_total.supremum_neut_or_in_closure[OF finite_lhs_subt[OF assms], of t]
unfolding states_def by auto
context
includes fset.lifting
begin
lift_definition fstates :: "('a, 'b) term fset \<Rightarrow> 'a bot_term fset" is states
by simp
lemma bound_max_state_fset:
"bound_max t (psubt_lhs_bot (fset R)) |\<in>| fstates R"
using bound_max_state_set[of "fset R" t]
- using fstates.rep_eq notin_fset by fastforce
+ using fstates.rep_eq by fastforce
end
definition nf_rules where
"nf_rules R \<F> = {|TA_rule f qs q | f qs q. (f, length qs) |\<in>| \<F> \<and> fset_of_list qs |\<subseteq>| fstates R \<and>
\<not>(\<exists> l |\<in>| R. l\<^sup>\<bottom> \<le>\<^sub>b BFun f qs) \<and> q = bound_max (BFun f qs) (psubt_lhs_bot (fset R))|}"
lemma nf_rules_fmember:
"TA_rule f qs q |\<in>| nf_rules R \<F> \<longleftrightarrow> (f, length qs) |\<in>| \<F> \<and> fset_of_list qs |\<subseteq>| fstates R \<and>
\<not>(\<exists> l |\<in>| R. l\<^sup>\<bottom> \<le>\<^sub>b BFun f qs) \<and> q = bound_max (BFun f qs) (psubt_lhs_bot (fset R))"
proof -
let ?subP = "\<lambda> n qs. fset_of_list qs |\<subseteq>| fstates R \<and> length qs = n"
let ?sub = "\<lambda> n. Collect (?subP n)"
have *: "finite (?sub n)" for n
using finite_lists_length_eq[of "fset (fstates R)" n]
by (simp add: less_eq_fset.rep_eq fset_of_list.rep_eq)
{fix f n assume mem: "(f, n) \<in> fset \<F>"
have **: "{f} \<times> (?sub n) = {(f, qs) |qs. ?subP n qs}" by auto
from mem have "finite {(f, qs) |qs. ?subP n qs}" using *
using finite_cartesian_product[OF _ *[of n], of "{f}"] unfolding ** by simp}
then have *: "finite (\<Union> (f, n) \<in> fset \<F> . {(f, qs) | qs. ?subP n qs})" by auto
have **: "(\<Union> (f, n) \<in> fset \<F> . {(f, qs) | qs. ?subP n qs}) = {(f, qs) | f qs. (f, length qs) |\<in>| \<F> \<and> ?subP (length qs) qs}"
- by (auto simp: fmember_iff_member_fset)
+ by auto
have *: "finite ({(f, qs) | f qs. (f, length qs) |\<in>| \<F> \<and> ?subP (length qs) qs} \<times> fset (fstates R))"
using * unfolding ** by (intro finite_cartesian_product) auto
have **: "{TA_rule f qs q | f qs q. (f, length qs) |\<in>| \<F> \<and> fset_of_list qs |\<subseteq>| fstates R \<and> q |\<in>| fstates R} =
(\<lambda> ((f, qs), q). TA_rule f qs q) ` ({(f, qs) | f qs. (f, length qs) |\<in>| \<F> \<and> ?subP (length qs) qs} \<times> fset (fstates R))"
- by (auto simp: image_def fmember_iff_member_fset split!: prod.splits)
+ by (auto simp: image_def split!: prod.splits)
have f: "finite {TA_rule f qs q | f qs q. (f, length qs) |\<in>| \<F> \<and> fset_of_list qs |\<subseteq>| fstates R \<and> q |\<in>| fstates R}"
unfolding ** using * by auto
show ?thesis
by (auto simp: nf_rules_def bound_max_state_fset intro!: finite_subset[OF _ f])
qed
definition nf_ta where
"nf_ta R \<F> = TA (nf_rules R \<F>) {||}"
definition nf_reg where
"nf_reg R \<F> = Reg (fstates R) (nf_ta R \<F>)"
lemma bound_max_sound:
assumes "finite R"
shows "bound_max t (psubt_lhs_bot R) \<le>\<^sub>b t"
using assms lift_total.lift_ord.supremum_smaller_subset[of "Some ` psubt_lhs_bot R" "Some t"]
by auto (metis (no_types, lifting) lift_less_eq_total.elims(2) option.sel option.simps(3))
lemma Bot_in_filter:
"Bot \<in> Set.filter (\<lambda>s. s \<le>\<^sub>b t) (states R)"
by (auto simp: Set.filter_def states_def)
lemma bound_max_exists:
"\<exists> p. p = bound_max t (psubt_lhs_bot R)"
by blast
lemma bound_max_unique:
assumes "p = bound_max t (psubt_lhs_bot R)" and "q = bound_max t (psubt_lhs_bot R)"
shows "p = q" using assms by force
lemma nf_rule_to_bound_max:
"f qs \<rightarrow> q |\<in>| nf_rules R \<F> \<Longrightarrow> q = bound_max (BFun f qs) (psubt_lhs_bot (fset R))"
by (auto simp: nf_rules_fmember)
lemma nf_rules_unique:
assumes "f qs \<rightarrow> q |\<in>| nf_rules R \<F>" and "f qs \<rightarrow> q' |\<in>| nf_rules R \<F>"
shows "q = q'" using assms unfolding nf_rules_def
using nf_rule_to_bound_max[OF assms(1)] nf_rule_to_bound_max[OF assms(2)]
using bound_max_unique by blast
lemma nf_ta_det:
shows "ta_det (nf_ta R \<F>)"
by (auto simp add: ta_det_def nf_ta_def nf_rules_unique)
lemma term_instance_of_reach_state:
assumes "q |\<in>| ta_der (nf_ta R \<F>) (adapt_vars t)" and "ground t"
shows "q \<le>\<^sub>b t\<^sup>\<bottom>" using assms(1, 2)
proof(induct t arbitrary: q)
case (Fun f ts)
from Fun(2) obtain qs where wit: "f qs \<rightarrow> q |\<in>| nf_rules R \<F>" "length qs = length ts"
"\<forall> i < length ts. qs ! i |\<in>| ta_der (nf_ta R \<F>) (adapt_vars (ts ! i))"
by (auto simp add: nf_ta_def)
then have "BFun f qs \<le>\<^sub>b Fun f ts\<^sup>\<bottom>" using Fun(1)[OF nth_mem, of i "qs !i" for i] using Fun(3)
by auto
then show ?case using bless_eq_trans wit(1) bound_max_sound[of "fset R"]
by (auto simp: nf_rules_fmember)
qed auto
lemma [simp]: "i < length ss \<Longrightarrow> l \<rhd> Fun f ss \<Longrightarrow> l \<rhd> ss ! i"
by (meson nth_mem subterm.dual_order.strict_trans supt.arg)
lemma subt_less_eq_res_less_eq:
assumes ground: "ground t" and "l |\<in>| R" and "l \<rhd> s" and "s\<^sup>\<bottom> \<le>\<^sub>b t\<^sup>\<bottom>"
and "q |\<in>| ta_der (nf_ta R \<F>) (adapt_vars t)"
shows "s\<^sup>\<bottom> \<le>\<^sub>b q" using assms(2-)
proof (induction t arbitrary: q s)
case (Var x)
then show ?case using lift_total.anti_sym by fastforce
next
case (Fun f ts) note IN = this
from IN obtain qs where rule: "f qs \<rightarrow> q |\<in>| nf_rules R \<F>" and
reach: "length qs = length ts" "\<forall> i < length ts. qs ! i |\<in>| ta_der (nf_ta R \<F>) (adapt_vars (ts ! i))"
by (auto simp add: nf_ta_def)
have q: "lift_sup_small (BFun f qs) (psubt_lhs_bot (fset R)) = Some q"
using nf_rule_to_bound_max[OF rule]
using lift_total.supremum_smaller_exists_unique[OF finite_lhs_subt, of "fset R" "BFun f qs"]
by simp (metis option.collapse option.distinct(1))
have subst: "s\<^sup>\<bottom> \<le>\<^sub>b BFun f qs" using IN(1)[OF nth_mem, of i "term.args s ! i" "qs ! i" for i] IN(2-) reach
by (cases s) (auto elim!: bless_eq.cases)
have "s\<^sup>\<bottom> \<in> psubt_lhs_bot (fset R)" using Fun(2 - 4)
- by auto (metis notin_fset)
+ by auto
then have "lift_total.lifted_less_eq (Some (s\<^sup>\<bottom>)) (lift_sup_small (BFun f qs) (psubt_lhs_bot (fset R)))"
using subst
by (intro lift_total.lift_ord.supremum_sound)
(auto simp: lift_total.lift_ord.smaller_subset_def)
then show ?case using subst q finite_lhs_subt
by auto
qed
lemma ta_nf_sound1:
assumes ground: "ground t" and lhs: "l |\<in>| R" and inst: "l\<^sup>\<bottom> \<le>\<^sub>b t\<^sup>\<bottom>"
shows "ta_der (nf_ta R \<F>) (adapt_vars t) = {||}"
proof (rule ccontr)
assume ass: "ta_der (nf_ta R \<F>) (adapt_vars t) \<noteq> {||}"
show False proof (cases t)
case [simp]: (Fun f ts) from ass
obtain q qs where fin: "q |\<in>| ta_der (nf_ta R \<F>) (adapt_vars (Fun f ts))" and
rule: "(f qs \<rightarrow> q) |\<in>| rules (nf_ta R \<F>)" "length qs = length ts" and
reach: "\<forall> i < length ts. qs ! i |\<in>| ta_der (nf_ta R \<F>) (adapt_vars (ts ! i))"
by (auto simp add: nf_ta_def) blast
have "l\<^sup>\<bottom> \<le>\<^sub>b BFun f qs" using reach assms(1) inst rule(2)
using subt_less_eq_res_less_eq[OF _ lhs, of "ts ! i" "term.args l ! i" "qs ! i" \<F> for i]
by (cases l) (auto elim!: bless_eq.cases intro!: bless_eq.step)
then show ?thesis using lhs rule by (auto simp: nf_ta_def nf_rules_def)
qed (metis ground ground.simps(1))
qed
lemma ta_nf_tr_to_state:
assumes "ground t" and "q |\<in>| ta_der (nf_ta R \<F>) (adapt_vars t)"
shows "q |\<in>| fstates R" using assms bound_max_state_fset
by (cases t) (auto simp: states_def nf_ta_def nf_rules_def)
lemma ta_nf_sound2:
assumes linear: "\<forall> l |\<in>| R. linear_term l"
and "ground (t :: ('f, 'v) term)" and "funas_term t \<subseteq> fset \<F>"
and NF: "\<And> l s. l |\<in>| R \<Longrightarrow> t \<unrhd> s \<Longrightarrow> \<not> l\<^sup>\<bottom> \<le>\<^sub>b s\<^sup>\<bottom>"
shows "\<exists> q. q |\<in>| ta_der (nf_ta R \<F>) (adapt_vars t)" using assms(2 - 4)
proof (induct t)
case (Fun f ts)
have sub: "\<And> i. i < length ts \<Longrightarrow> (\<And>l s. l |\<in>| R \<Longrightarrow> ts ! i \<unrhd> s \<Longrightarrow> \<not> l\<^sup>\<bottom> \<le>\<^sub>b s\<^sup>\<bottom>) " using Fun(4) nth_mem by blast
from Fun(1)[OF nth_mem] this Fun(2, 3) obtain qs where
reach: "(\<forall> i < length ts. qs ! i |\<in>| ta_der (nf_ta R \<F>) (adapt_vars (ts ! i)))" and len: "length qs = length ts"
using Ex_list_of_length_P[of "length ts" "\<lambda> x i. x |\<in>| (ta_der (nf_ta R \<F>) (adapt_vars (ts ! i)))"]
by auto (meson UN_subset_iff nth_mem)
have nt_inst: "\<not> (\<exists> s |\<in>| R. s\<^sup>\<bottom> \<le>\<^sub>b BFun f qs)"
proof (rule ccontr, simp)
assume ass: "\<exists> s |\<in>| R. s\<^sup>\<bottom> \<le>\<^sub>b BFun f qs"
from term_instance_of_reach_state[of "qs ! i" R \<F> "ts ! i" for i] reach Fun(2) len
have "BFun f qs \<le>\<^sub>b Fun f ts\<^sup>\<bottom>" by auto
then show False using ass Fun(4) bless_eq_trans by blast
qed
obtain q where "q = bound_max (BFun f qs) (psubt_lhs_bot (fset R))" by blast
then have "f qs \<rightarrow> q |\<in>| rules (nf_ta R \<F>)" using Fun(2 - 4)
using ta_nf_tr_to_state[of "ts ! i" "qs ! i" R \<F> for i] len nt_inst reach
- by (auto simp: nf_ta_def nf_rules_fmember, simp add: fmember_iff_member_fset)
+ by (auto simp: nf_ta_def nf_rules_fmember)
(metis (no_types, lifting) in_fset_idx nth_mem)
then show ?case using reach len by auto
qed auto
lemma ta_nf_lang_sound:
assumes "l |\<in>| R"
shows "C\<langle>l \<cdot> \<sigma>\<rangle> \<notin> ta_lang (fstates R) (nf_ta R \<F>)"
proof (rule ccontr, simp del: ta_lang_to_gta_lang)
assume *: "C\<langle>l \<cdot> \<sigma>\<rangle> \<in> ta_lang (fstates R) (nf_ta R \<F>)"
then have cgr:"ground (C\<langle>l\<cdot>\<sigma>\<rangle>)" unfolding ta_lang_def by force
then have gr: "ground (l \<cdot> \<sigma>)" by simp
then have "l\<^sup>\<bottom> \<le>\<^sub>b (l \<cdot> \<sigma>)\<^sup>\<bottom>" using instance_to_bless_eq by blast
from ta_nf_sound1[OF gr assms(1) this] have res: "ta_der (nf_ta R \<F>) (adapt_vars (l \<cdot> \<sigma>)) = {||}" .
from ta_langE * obtain q where "q |\<in>| ta_der (nf_ta R \<F>) (adapt_vars (C\<langle>l\<cdot>\<sigma>\<rangle>))"
by (metis adapt_vars_adapt_vars)
with ta_der_ctxt_decompose[OF this[unfolded adapt_vars_ctxt]] res
show False by blast
qed
lemma ta_nf_lang_complete:
assumes linear: "\<forall> l |\<in>| R. linear_term l"
and ground: "ground (t :: ('f, 'v) term)" and sig: "funas_term t \<subseteq> fset \<F>"
and nf: "\<And>C \<sigma> l. l |\<in>| R \<Longrightarrow> C\<langle>l\<cdot>\<sigma>\<rangle> \<noteq> t"
shows "t \<in> ta_lang (fstates R) (nf_ta R \<F>)"
proof -
from nf have "\<And> l s. l |\<in>| R \<Longrightarrow> t \<unrhd> s \<Longrightarrow> \<not> l\<^sup>\<bottom> \<le>\<^sub>b s\<^sup>\<bottom>"
using bless_eq_to_instance linear by blast
from ta_nf_sound2[OF linear ground sig] this
obtain q where "q |\<in>| ta_der (nf_ta R \<F>) (adapt_vars t)" by blast
from this ta_nf_tr_to_state[OF ground this] ground show ?thesis
by (intro ta_langI) (auto simp add: nf_ta_def)
qed
lemma ta_nf_\<L>_complete:
assumes linear: "\<forall> l |\<in>| R. linear_term l"
and sig: "funas_gterm t \<subseteq> fset \<F>"
and nf: "\<And>C \<sigma> l. l |\<in>| R \<Longrightarrow> C\<langle>l\<cdot>\<sigma>\<rangle> \<noteq> (term_of_gterm t)"
shows "t \<in> \<L> (nf_reg R \<F>)"
using ta_nf_lang_complete[of R "term_of_gterm t" \<F>] assms
by (force simp: \<L>_def nf_reg_def funas_term_of_gterm_conv)
lemma nf_ta_funas:
assumes "ground t" "q |\<in>| ta_der (nf_ta R \<F>) t"
shows "funas_term t \<subseteq> fset \<F>" using assms
proof (induct t arbitrary: q)
case (Fun f ts)
from Fun(2-) have "(f, length ts) |\<in>| \<F>"
by (auto simp: nf_ta_def nf_rules_def)
then show ?case using Fun
- by (auto simp: fmember_iff_member_fset) (metis Fun.hyps Fun.prems(2) in_set_idx subsetD ta_der_Fun)
+ apply simp
+ by (smt (verit) Union_least image_iff in_set_idx)
qed auto
lemma gta_lang_nf_ta_funas:
assumes "t \<in> \<L> (nf_reg R \<F>)"
shows "funas_gterm t \<subseteq> fset \<F>" using assms nf_ta_funas[of "term_of_gterm t" _ R \<F>]
unfolding nf_reg_def \<L>_def
by (auto simp: funas_term_of_gterm_conv)
end
diff --git a/thys/FO_Theory_Rewriting/Primitives/NF_Impl.thy b/thys/FO_Theory_Rewriting/Primitives/NF_Impl.thy
--- a/thys/FO_Theory_Rewriting/Primitives/NF_Impl.thy
+++ b/thys/FO_Theory_Rewriting/Primitives/NF_Impl.thy
@@ -1,198 +1,198 @@
theory NF_Impl
imports NF
Type_Instances_Impl
begin
subsubsection \<open>Implementation of normal form construction\<close>
(* Implementation *)
fun supteq_list :: "('f, 'v) Term.term \<Rightarrow> ('f, 'v) Term.term list"
where
"supteq_list (Var x) = [Var x]" |
"supteq_list (Fun f ts) = Fun f ts # concat (map supteq_list ts)"
fun supt_list :: "('f, 'v) Term.term \<Rightarrow> ('f, 'v) Term.term list"
where
"supt_list (Var x) = []" |
"supt_list (Fun f ts) = concat (map supteq_list ts)"
lemma supteq_list [simp]:
"set (supteq_list t) = {s. t \<unrhd> s}"
proof (rule set_eqI, simp)
fix s
show "s \<in> set(supteq_list t) = (t \<unrhd> s)"
proof (induct t, simp add: supteq_var_imp_eq)
case (Fun f ss)
show ?case
proof (cases "Fun f ss = s", (auto)[1])
case False
show ?thesis
proof
assume "Fun f ss \<unrhd> s"
with False have sup: "Fun f ss \<rhd> s" using supteq_supt_conv by auto
obtain C where "C \<noteq> \<box>" and "Fun f ss = C\<langle>s\<rangle>" using sup by auto
then obtain b D a where "Fun f ss = Fun f (b @ D\<langle>s\<rangle> # a)" by (cases C, auto)
then have D: "D\<langle>s\<rangle> \<in> set ss" by auto
with Fun[OF D] ctxt_imp_supteq[of D s] obtain t where "t \<in> set ss" and "s \<in> set (supteq_list t)" by auto
then show "s \<in> set (supteq_list (Fun f ss))" by auto
next
assume "s \<in> set (supteq_list (Fun f ss))"
with False obtain t where t: "t \<in> set ss" and "s \<in> set (supteq_list t)" by auto
with Fun[OF t] have "t \<unrhd> s" by auto
with t show "Fun f ss \<unrhd> s" by auto
qed
qed
qed
qed
lemma supt_list_sound [simp]:
"set (supt_list t) = {s. t \<rhd> s}"
by (cases t) auto
fun mergeP_impl where
"mergeP_impl Bot t = True"
| "mergeP_impl t Bot = True"
| "mergeP_impl (BFun f ss) (BFun g ts) =
(if f = g \<and> length ss = length ts then list_all (\<lambda> (x, y). mergeP_impl x y) (zip ss ts) else False)"
lemma [simp]: "mergeP_impl s Bot = True" by (cases s) auto
lemma [simp]: "mergeP_impl s t \<longleftrightarrow> (s, t) \<in> mergeP" (is "?LS = ?RS")
proof
show "?LS \<Longrightarrow> ?RS"
by (induct rule: mergeP_impl.induct, auto split: if_splits intro!: step)
(smt length_zip list_all_length mergeP.step min_less_iff_conj nth_mem nth_zip old.prod.case)
next
show "?RS \<Longrightarrow> ?LS" by (induct rule: mergeP.induct, auto simp add: list_all_length)
qed
fun bless_eq_impl where
"bless_eq_impl Bot t = True"
| "bless_eq_impl (BFun f ss) (BFun g ts) =
(if f = g \<and> length ss = length ts then list_all (\<lambda> (x, y). bless_eq_impl x y) (zip ss ts) else False)"
| "bless_eq_impl _ _ = False"
lemma [simp]: "bless_eq_impl s t \<longleftrightarrow> (s, t) \<in> bless_eq" (is "?RS = ?LS")
proof
show "?LS \<Longrightarrow> ?RS" by (induct rule: bless_eq.induct, auto simp add: list_all_length)
next
show "?RS \<Longrightarrow> ?LS"
by (induct rule: bless_eq_impl.induct, auto split: if_splits intro!: bless_eq.step)
(metis (full_types) length_zip list_all_length min_less_iff_conj nth_mem nth_zip old.prod.case)
qed
definition "psubt_bot_impl R \<equiv> remdups (map term_to_bot_term (concat (map supt_list R)))"
lemma psubt_bot_impl[simp]: "set (psubt_bot_impl R) = psubt_lhs_bot (set R)"
by (induct R, auto simp: psubt_bot_impl_def)
definition "states_impl R = List.insert Bot (map the (removeAll None
(closure_impl (lift_f_total mergeP_impl (\<up>)) (map Some (psubt_bot_impl R)))))"
lemma states_impl [simp]: "set (states_impl R) = states (set R)"
proof -
have [simp]: "lift_f_total mergeP_impl (\<up>) = lift_f_total (\<lambda> x y. mergeP_impl x y) (\<up>)" by blast
show ?thesis unfolding states_impl_def states_def
using lift_total.cl.closure_impl
by (simp add: lift_total.cl.pred_closure_lift_closure)
qed
abbreviation check_intance_lhs where
"check_intance_lhs qs f R \<equiv> list_all (\<lambda> u. \<not> bless_eq_impl u (BFun f qs)) R"
definition min_elem where
"min_elem s ss = (let ts = filter (\<lambda> x. bless_eq_impl x s) ss in
foldr (\<up>) ts Bot)"
lemma bound_impl [simp, code]:
"bound_max s (set ss) = min_elem s ss"
proof -
have [simp]: "{y. lift_total.lifted_less_eq y (Some s) \<and> y \<in> Some ` set ss} = Some ` {x \<in> set ss. x \<le>\<^sub>b s}"
by auto
then show ?thesis
using lift_total.supremum_impl[of "filter (\<lambda> x. bless_eq_impl x s) ss"]
using lift_total.supremum_smaller_exists_unique[of "set ss" s]
by (auto simp: min_elem_def Let_def lift_total.lift_ord.smaller_subset_def)
qed
definition nf_rule_impl where
"nf_rule_impl S R SR h = (let (f, n) = h in
let states = List.n_lists n S in
let nlhs_inst = filter (\<lambda> qs. check_intance_lhs qs f R) states in
map (\<lambda> qs. TA_rule f qs (min_elem (BFun f qs) SR)) nlhs_inst)"
abbreviation nf_rules_impl where
"nf_rules_impl R \<F> \<equiv> concat (map (nf_rule_impl (states_impl R) (map term_to_bot_term R) (psubt_bot_impl R)) \<F>)"
(* Section proves that the implementation constructs the same rule set *)
lemma nf_rules_in_impl:
assumes "TA_rule f qs q |\<in>| nf_rules (fset_of_list R) (fset_of_list \<F>)"
shows "TA_rule f qs q |\<in>| fset_of_list (nf_rules_impl R \<F>)"
proof -
have funas: "(f, length qs) \<in> set \<F>" and st: "fset_of_list qs |\<subseteq>| fstates (fset_of_list R)"
and nlhs: "\<not>(\<exists> s \<in> (set R). s\<^sup>\<bottom> \<le>\<^sub>b BFun f qs)"
and min: "q = bound_max (BFun f qs) (psubt_lhs_bot (set R))"
- using assms by (auto simp add: nf_rules_fmember simp flip: fset_of_list_elem fmember_iff_member_fset)
+ using assms by (auto simp add: nf_rules_fmember simp flip: fset_of_list_elem)
then have st_impl: "qs |\<in>| fset_of_list (List.n_lists (length qs) (states_impl R))"
by (auto simp add: fset_of_list_elem subset_code(1) set_n_lists
fset_of_list.rep_eq less_eq_fset.rep_eq fstates.rep_eq)
from nlhs have nlhs_impl: "check_intance_lhs qs f (map term_to_bot_term R)"
by (auto simp: list.pred_set)
have min_impl: "q = min_elem (BFun f qs) (psubt_bot_impl R)"
using bound_impl min
by (auto simp flip: psubt_bot_impl)
then show ?thesis using funas nlhs_impl funas st_impl unfolding nf_rule_impl_def
by (auto simp: fset_of_list_elem)
qed
lemma nf_rules_impl_in_rules:
assumes "TA_rule f qs q |\<in>| fset_of_list (nf_rules_impl R \<F>)"
shows "TA_rule f qs q |\<in>| nf_rules (fset_of_list R) (fset_of_list \<F>)"
proof -
have funas: "(f, length qs) \<in> set \<F>"
and st_impl: "qs |\<in>| fset_of_list (List.n_lists (length qs) (states_impl R))"
and nlhs_impl: "check_intance_lhs qs f (map term_to_bot_term R)"
and min: "q = min_elem (BFun f qs) (psubt_bot_impl R)" using assms
by (auto simp add: set_n_lists nf_rule_impl_def fset_of_list_elem)
from st_impl have st: "fset_of_list qs |\<subseteq>| fstates (fset_of_list R)"
- by (force simp: set_n_lists fset_of_list_elem fstates.rep_eq fmember_iff_member_fset fset_of_list.rep_eq)
+ by (force simp: set_n_lists fset_of_list_elem fstates.rep_eq fset_of_list.rep_eq)
from nlhs_impl have nlhs: "\<not>(\<exists> l \<in> (set R). l\<^sup>\<bottom> \<le>\<^sub>b BFun f qs)"
by auto (metis (no_types, lifting) Ball_set_list_all in_set_idx length_map nth_map nth_mem)
have "q = bound_max (BFun f qs) (psubt_lhs_bot (set R))"
using bound_impl min
by (auto simp flip: psubt_bot_impl)
then show ?thesis using funas st nlhs
by (auto simp add: nf_rules_fmember fset_of_list_elem fset_of_list.rep_eq)
qed
lemma rule_set_eq:
shows "nf_rules (fset_of_list R) (fset_of_list \<F>) = fset_of_list (nf_rules_impl R \<F>)" (is "?Ls = ?Rs")
proof -
{fix r assume "r |\<in>| ?Ls" then have "r |\<in>| ?Rs"
using nf_rules_in_impl[where ?R = R and ?\<F> = \<F>]
by (cases r) auto}
moreover
{fix r assume "r |\<in>| ?Rs" then have "r |\<in>| ?Ls"
using nf_rules_impl_in_rules[where ?R = R and ?\<F> = \<F>]
by (cases r) auto}
ultimately show ?thesis by blast
qed
(* Code equation for normal form TA *)
lemma fstates_code[code]:
"fstates R = fset_of_list (states_impl (sorted_list_of_fset R))"
- by (auto simp: fmember_iff_member_fset fstates.rep_eq fset_of_list.rep_eq)
+ by (auto simp: fstates.rep_eq fset_of_list.rep_eq)
lemma nf_ta_code [code]:
"nf_ta R \<F> = TA (fset_of_list (nf_rules_impl (sorted_list_of_fset R) (sorted_list_of_fset \<F>))) {||}"
unfolding nf_ta_def using rule_set_eq[of "sorted_list_of_fset R" "sorted_list_of_fset \<F>"]
by (intro TA_equalityI) auto
(*
export_code nf_ta in Haskell
*)
end
\ No newline at end of file
diff --git a/thys/FO_Theory_Rewriting/Util/Tree_Automata_Derivation_Split.thy b/thys/FO_Theory_Rewriting/Util/Tree_Automata_Derivation_Split.thy
--- a/thys/FO_Theory_Rewriting/Util/Tree_Automata_Derivation_Split.thy
+++ b/thys/FO_Theory_Rewriting/Util/Tree_Automata_Derivation_Split.thy
@@ -1,406 +1,406 @@
theory Tree_Automata_Derivation_Split
imports Regular_Tree_Relations.Tree_Automata
Ground_MCtxt
begin
lemma ta_der'_inf_mctxt:
assumes "t |\<in>| ta_der' \<A> s"
shows "fst (split_vars t) \<le> (mctxt_of_term s)" using assms
proof (induct s arbitrary: t)
case (Fun f ts) then show ?case
by (cases t) (auto simp: comp_def less_eq_mctxt_prime intro: less_eq_mctxt'.intros)
qed (auto simp: ta_der'.simps)
lemma ta_der'_poss_subt_at_ta_der':
assumes "t |\<in>| ta_der' \<A> s" and "p \<in> poss t"
shows "t |_ p |\<in>| ta_der' \<A> (s |_ p)" using assms
by (induct s arbitrary: t p) (auto simp: ta_der'.simps, blast+)
lemma ta_der'_varposs_to_ta_der:
assumes "t |\<in>| ta_der' \<A> s" and "p \<in> varposs t"
shows "the_Var (t |_ p) |\<in>| ta_der \<A> (s |_ p)" using assms
by (induct s arbitrary: t p) (auto simp: ta_der'.simps, blast+)
definition "ta_der'_target_mctxt t \<equiv> fst (split_vars t)"
definition "ta_der'_target_args t \<equiv> snd (split_vars t)"
definition "ta_der'_source_args t s \<equiv> unfill_holes (fst (split_vars t)) s"
lemmas ta_der'_mctxt_simps = ta_der'_target_mctxt_def ta_der'_target_args_def ta_der'_source_args_def
lemma ta_der'_target_mctxt_funas [simp]:
"funas_mctxt (ta_der'_target_mctxt u) = funas_term u"
by (auto simp: ta_der'_target_mctxt_def)
lemma ta_der'_target_mctxt_ground [simp]:
"ground_mctxt (ta_der'_target_mctxt t)"
by (auto simp: ta_der'_target_mctxt_def)
lemma ta_der'_source_args_ground:
"t |\<in>| ta_der' \<A> s \<Longrightarrow> ground s \<Longrightarrow> \<forall> u \<in> set (ta_der'_source_args t s). ground u"
by (metis fill_unfill_holes ground_fill_holes length_unfill_holes ta_der'_inf_mctxt ta_der'_mctxt_simps)
lemma ta_der'_source_args_term_of_gterm:
"t |\<in>| ta_der' \<A> (term_of_gterm s) \<Longrightarrow> \<forall> u \<in> set (ta_der'_source_args t (term_of_gterm s)). ground u"
by (intro ta_der'_source_args_ground) auto
lemma ta_der'_source_args_length:
"t |\<in>| ta_der' \<A> s \<Longrightarrow> num_holes (ta_der'_target_mctxt t) = length (ta_der'_source_args t s)"
by (auto simp: ta_der'_mctxt_simps ta_der'_inf_mctxt)
lemma ta_der'_target_args_length:
"num_holes (ta_der'_target_mctxt t) = length (ta_der'_target_args t)"
by (auto simp: ta_der'_mctxt_simps split_vars_num_holes)
lemma ta_der'_target_args_vars_term_conv:
"vars_term t = set (ta_der'_target_args t)"
by (auto simp: ta_der'_target_args_def split_vars_vars_term_list)
lemma ta_der'_target_args_vars_term_list_conv:
"ta_der'_target_args t = vars_term_list t"
by (auto simp: ta_der'_target_args_def split_vars_vars_term_list)
lemma mctxt_args_ta_der':
assumes "num_holes C = length qs" "num_holes C = length ss"
and "\<forall> i < length ss. qs ! i |\<in>| ta_der \<A> (ss ! i)"
shows "(fill_holes C (map Var qs)) |\<in>| ta_der' \<A> (fill_holes C ss)" using assms
proof (induct rule: fill_holes_induct2)
case MHole then show ?case
by (cases ss; cases qs) (auto simp: ta_der_to_ta_der')
next
case (MFun f ts) then show ?case
by (simp add: partition_by_nth_nth(1, 2))
qed auto
\<comment> \<open>Splitting derivation into multihole context containing the remaining function symbols and
the states, where each state is reached via the automata\<close>
lemma ta_der'_mctxt_structure:
assumes "t |\<in>| ta_der' \<A> s"
shows "t = fill_holes (ta_der'_target_mctxt t) (map Var (ta_der'_target_args t))" (is "?G1")
"s = fill_holes (ta_der'_target_mctxt t) (ta_der'_source_args t s)" (is "?G2")
"num_holes (ta_der'_target_mctxt t) = length (ta_der'_source_args t s) \<and>
length (ta_der'_source_args t s) = length (ta_der'_target_args t)" (is "?G3")
"i < length (ta_der'_source_args t s) \<Longrightarrow> ta_der'_target_args t ! i |\<in>| ta_der \<A> (ta_der'_source_args t s ! i)"
proof -
let ?C = "ta_der'_target_mctxt t" let ?ss = "ta_der'_source_args t s"
let ?qs = "ta_der'_target_args t"
have t_split: "?G1" by (auto simp: ta_der'_mctxt_simps split_vars_fill_holes)
have s_split: "?G2" by (auto simp: ta_der'_mctxt_simps ta_der'_inf_mctxt[OF assms]
intro!: fill_unfill_holes[symmetric])
have len: "num_holes ?C = length ?ss" "length ?ss = length ?qs" using assms
by (auto simp: ta_der'_mctxt_simps split_vars_num_holes ta_der'_inf_mctxt)
have "i < length (ta_der'_target_args t) \<Longrightarrow>
ta_der'_target_args t ! i |\<in>| ta_der \<A> (ta_der'_source_args t s ! i)" for i
using ta_der'_poss_subt_at_ta_der'[OF assms, of "varposs_list t ! i"]
unfolding ta_der'_mctxt_simps split_vars_vars_term_list length_map
by (auto simp: unfill_holes_to_subst_at_hole_poss[OF ta_der'_inf_mctxt[OF assms]]
simp flip: varposs_list_to_var_term_list[of i t, unfolded varposs_list_var_terms_length])
(metis assms hole_poss_split_vars_varposs_list nth_map nth_mem
ta_der'_varposs_to_ta_der ta_der_to_ta_der' varposs_eq_varposs_list varposs_list_var_terms_length)
then show ?G1 ?G2 ?G3 "i < length (ta_der'_source_args t s) \<Longrightarrow>
ta_der'_target_args t ! i |\<in>| ta_der \<A> (ta_der'_source_args t s ! i)" using len t_split s_split
by (simp_all add: ta_der'_mctxt_simps)
qed
lemma ta_der'_ground_mctxt_structure:
assumes "t |\<in>| ta_der' \<A> (term_of_gterm s)"
shows "t = fill_holes (ta_der'_target_mctxt t) (map Var (ta_der'_target_args t))"
"term_of_gterm s = fill_holes (ta_der'_target_mctxt t) (ta_der'_source_args t (term_of_gterm s))"
"num_holes (ta_der'_target_mctxt t) = length (ta_der'_source_args t (term_of_gterm s)) \<and>
length (ta_der'_source_args t (term_of_gterm s)) = length (ta_der'_target_args t)"
"i < length (ta_der'_target_args t) \<Longrightarrow> ta_der'_target_args t ! i |\<in>| ta_der \<A> (ta_der'_source_args t (term_of_gterm s) ! i)"
using ta_der'_mctxt_structure[OF assms]
by force+
\<comment> \<open>Splitting derivation into context containing the remaining function symbols and state\<close>
definition "ta_der'_gctxt t \<equiv> gctxt_of_gmctxt (gmctxt_of_mctxt (fst (split_vars t)))"
abbreviation "ta_der'_ctxt t \<equiv> ctxt_of_gctxt (ta_der'_gctxt t)"
definition "ta_der'_source_ctxt_arg t s \<equiv> hd (unfill_holes (fst (split_vars t)) s)"
abbreviation "ta_der'_source_gctxt_arg t s \<equiv> gterm_of_term (ta_der'_source_ctxt_arg t (term_of_gterm s))"
lemma ta_der'_ctxt_structure:
assumes "t |\<in>| ta_der' \<A> s" "vars_term_list t = [q]"
shows "t = (ta_der'_ctxt t)\<langle>Var q\<rangle>" (is "?G1")
"s = (ta_der'_ctxt t)\<langle>ta_der'_source_ctxt_arg t s\<rangle>" (is "?G2")
"ground_ctxt (ta_der'_ctxt t)" (is "?G3")
"q |\<in>| ta_der \<A> (ta_der'_source_ctxt_arg t s)" (is "?G4")
proof -
have *: "length xs = Suc 0 \<Longrightarrow> xs = [hd xs]" for xs
by (metis length_0_conv length_Suc_conv list.sel(1))
have [simp]: "length (snd (split_vars t)) = Suc 0" using assms(2) ta_der'_inf_mctxt[OF assms(1)]
by (auto simp: split_vars_vars_term_list)
have [simp]: "num_gholes (gmctxt_of_mctxt (fst (split_vars t))) = Suc 0" using assms(2)
by (simp add: split_vars_num_holes split_vars_vars_term_list)
have [simp]: "ta_der'_source_args t s = [ta_der'_source_ctxt_arg t s]"
using assms(2) ta_der'_inf_mctxt[OF assms(1)]
by (auto simp: ta_der'_source_args_def ta_der'_source_ctxt_arg_def split_vars_num_holes intro!: *)
have t_split: ?G1 using assms(2)
by (auto simp: ta_der'_gctxt_def split_vars_fill_holes
split_vars_vars_term_list simp flip: ctxt_of_gctxt_gctxt_of_gmctxt_apply)
have s_split: ?G2 using ta_der'_mctxt_structure[OF assms(1)] assms(2)
by (auto simp: ta_der'_gctxt_def ta_der'_target_mctxt_def
simp flip: ctxt_of_gctxt_gctxt_of_gmctxt_apply)
from ta_der'_mctxt_structure[OF assms(1)] have ?G4
by (auto simp: ta_der'_target_args_def assms(2) split_vars_vars_term_list)
moreover have ?G3 unfolding ta_der'_gctxt_def by auto
ultimately show ?G1 ?G2 ?G3 ?G4 using t_split s_split
by force+
qed
lemma ta_der'_ground_ctxt_structure:
assumes "t |\<in>| ta_der' \<A> (term_of_gterm s)" "vars_term_list t = [q]"
shows "t = (ta_der'_ctxt t)\<langle>Var q\<rangle>"
"s = (ta_der'_gctxt t)\<langle>ta_der'_source_gctxt_arg t s\<rangle>\<^sub>G"
"ground (ta_der'_source_ctxt_arg t (term_of_gterm s))"
"q |\<in>| ta_der \<A> (ta_der'_source_ctxt_arg t (term_of_gterm s))"
using ta_der'_ctxt_structure[OF assms] term_of_gterm_ctxt_apply
by force+
subsection \<open>Sufficient condition for splitting the reachability relation induced by a tree automaton\<close>
locale derivation_split =
fixes A :: "('q, 'f) ta" and \<A> and \<B>
assumes rule_split: "rules A = rules \<A> |\<union>| rules \<B>"
and eps_split: "eps A = eps \<A> |\<union>| eps \<B>"
and B_target_states: "rule_target_states (rules \<B>) |\<union>| (snd |`| (eps \<B>)) |\<inter>|
(rule_arg_states (rules \<A>) |\<union>| (fst |`| (eps \<A>))) = {||}"
begin
abbreviation "\<Delta>\<^sub>A \<equiv> rules \<A>"
abbreviation "\<Delta>\<^sub>\<E>\<^sub>A \<equiv> eps \<A>"
abbreviation "\<Delta>\<^sub>B \<equiv> rules \<B>"
abbreviation "\<Delta>\<^sub>\<E>\<^sub>B \<equiv> eps \<B>"
abbreviation "\<Q>\<^sub>A \<equiv> \<Q> \<A>"
definition "\<Q>\<^sub>B \<equiv> rule_target_states \<Delta>\<^sub>B |\<union>| (snd |`| \<Delta>\<^sub>\<E>\<^sub>B)"
lemmas B_target_states' = B_target_states[folded \<Q>\<^sub>B_def]
lemma states_split [simp]: "\<Q> A = \<Q> \<A> |\<union>| \<Q> \<B>"
by (auto simp add: \<Q>_def rule_split eps_split)
lemma A_args_states_not_B:
"TA_rule f qs q |\<in>| \<Delta>\<^sub>A \<Longrightarrow> p |\<in>| fset_of_list qs \<Longrightarrow> p |\<notin>| \<Q>\<^sub>B"
using B_target_states
by (force simp add: \<Q>\<^sub>B_def)
lemma rule_statesD:
"r |\<in>| \<Delta>\<^sub>A \<Longrightarrow> r_rhs r |\<in>| \<Q>\<^sub>A"
"r |\<in>| \<Delta>\<^sub>B \<Longrightarrow> r_rhs r |\<in>| \<Q>\<^sub>B"
"r |\<in>| \<Delta>\<^sub>A \<Longrightarrow> p |\<in>| fset_of_list (r_lhs_states r) \<Longrightarrow> p |\<in>| \<Q>\<^sub>A"
"TA_rule f qs q |\<in>| \<Delta>\<^sub>A \<Longrightarrow> q |\<in>| \<Q>\<^sub>A"
"TA_rule f qs q |\<in>| \<Delta>\<^sub>B \<Longrightarrow> q |\<in>| \<Q>\<^sub>B"
"TA_rule f qs q |\<in>| \<Delta>\<^sub>A \<Longrightarrow> p |\<in>| fset_of_list qs \<Longrightarrow> p |\<in>| \<Q>\<^sub>A"
- by (auto simp: rule_statesD \<Q>\<^sub>B_def rev_fimage_eqI)
+ by (auto simp: rule_statesD \<Q>\<^sub>B_def rev_image_eqI)
lemma eps_states_dest:
"(p, q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>A \<Longrightarrow> p |\<in>| \<Q>\<^sub>A"
"(p, q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>A \<Longrightarrow> q |\<in>| \<Q>\<^sub>A"
"(p, q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>A|\<^sup>+| \<Longrightarrow> p |\<in>| \<Q>\<^sub>A"
"(p, q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>A|\<^sup>+| \<Longrightarrow> q |\<in>| \<Q>\<^sub>A"
"(p, q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>B \<Longrightarrow> q |\<in>| \<Q>\<^sub>B"
"(p, q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>B|\<^sup>+| \<Longrightarrow> q |\<in>| \<Q>\<^sub>B"
- by (auto simp: eps_dest_all \<Q>\<^sub>B_def rev_fimage_eqI elim: ftranclE)
+ by (auto simp: eps_dest_all \<Q>\<^sub>B_def rev_image_eqI elim: ftranclE)
lemma transcl_eps_simp:
"(eps A)|\<^sup>+| = \<Delta>\<^sub>\<E>\<^sub>A|\<^sup>+| |\<union>| \<Delta>\<^sub>\<E>\<^sub>B|\<^sup>+| |\<union>| (\<Delta>\<^sub>\<E>\<^sub>A|\<^sup>+| |O| \<Delta>\<^sub>\<E>\<^sub>B|\<^sup>+|)"
proof -
have "\<Delta>\<^sub>\<E>\<^sub>B |O| \<Delta>\<^sub>\<E>\<^sub>A = {||}" using B_target_states'
by (metis eps_states_dest(5) ex_fin_conv fimageI finterI frelcompE fst_conv inf_sup_distrib1 sup_eq_bot_iff)
from ftrancl_Un2_separatorE[OF this] show ?thesis
unfolding eps_split by auto
qed
lemma B_rule_eps_A_False:
"f qs \<rightarrow> q |\<in>| \<Delta>\<^sub>B \<Longrightarrow> (q, p) |\<in>| \<Delta>\<^sub>\<E>\<^sub>A|\<^sup>+| \<Longrightarrow> False"
using B_target_states unfolding \<Q>\<^sub>B_def
by (metis B_target_states' equalsffemptyD fimage_eqI finter_iff fst_conv ftranclD funion_iff local.rule_statesD(5))
lemma to_A_rule_set:
assumes "TA_rule f qs q |\<in>| rules A" and "q = p \<or> (q, p) |\<in>| (eps A)|\<^sup>+|" and "p |\<notin>| \<Q>\<^sub>B"
shows "TA_rule f qs q |\<in>| \<Delta>\<^sub>A" "q = p \<or> (q, p) |\<in>| \<Delta>\<^sub>\<E>\<^sub>A|\<^sup>+|" using assms
unfolding transcl_eps_simp rule_split
by (auto dest: rule_statesD eps_states_dest dest: B_rule_eps_A_False)
lemma to_B_rule_set:
assumes "TA_rule f qs q |\<in>| rules A" and "q |\<notin>| \<Q>\<^sub>A"
shows "TA_rule f qs q |\<in>| \<Delta>\<^sub>B" using assms
unfolding transcl_eps_simp rule_split
by (auto dest: rule_statesD eps_states_dest)
declare fsubsetI[rule del]
lemma ta_der_monos:
"ta_der \<A> t |\<subseteq>| ta_der A t" "ta_der \<B> t |\<subseteq>| ta_der A t"
by (auto simp: sup.coboundedI1 rule_split eps_split intro!: ta_der_mono)
declare fsubsetI[intro!]
lemma ta_der_from_\<Delta>\<^sub>A:
assumes "q |\<in>| ta_der A (term_of_gterm t)" and "q |\<notin>| \<Q>\<^sub>B"
shows "q |\<in>| ta_der \<A> (term_of_gterm t)" using assms
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
have "i < length ts \<Longrightarrow> ps ! i |\<notin>| \<Q>\<^sub>B" for i using GFun A_args_states_not_B
by (metis fnth_mem to_A_rule_set(1))
then show ?case using GFun(2, 5) to_A_rule_set[OF GFun(1, 3, 6)]
by (auto simp: transcl_eps_simp)
qed
lemma ta_state:
assumes "q |\<in>| ta_der A (term_of_gterm s)"
shows "q |\<in>| \<Q>\<^sub>A \<or> q |\<in>| \<Q>\<^sub>B" using assms
by (cases s) (auto simp: rule_split transcl_eps_simp dest: rule_statesD eps_states_dest)
(* Main lemmas *)
lemma ta_der_split:
assumes "q |\<in>| ta_der A (term_of_gterm s)" and "q |\<in>| \<Q>\<^sub>B"
shows "\<exists> t. t |\<in>| ta_der' \<A> (term_of_gterm s) \<and> q |\<in>| ta_der \<B> t"
(is "\<exists>t . ?P s q t") using assms
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
{fix i assume ass: "i < length ts"
then have "\<exists> t. t |\<in>| ta_der' \<A> (term_of_gterm (ts ! i)) \<and> ps ! i |\<in>| ta_der \<B> t"
proof (cases "ps ! i |\<notin>| \<Q>\<^sub>B")
case True then show ?thesis
using ta_state GFun(2, 4) ta_der_from_\<Delta>\<^sub>A[of "ps ! i" "ts ! i"] ass
by (intro exI[of _ "Var (ps ! i)"]) (auto simp: ta_der_to_ta_der' \<Q>\<^sub>B_def)
next
case False
then have "ps ! i |\<in>| \<Q>\<^sub>B" using ta_state[OF GFun(4)[OF ass]]
by auto
from GFun(5)[OF ass this] show ?thesis .
qed}
then obtain h where IH:
"\<forall> i < length ts. h i |\<in>| ta_der' \<A> (term_of_gterm (ts ! i))"
"\<forall> i < length ts. ps ! i |\<in>| ta_der \<B> (h i)"
using GFun(1 - 4) choice_nat[of "length ts" "\<lambda> t i. ?P (ts ! i) (ps ! i) t"]
by blast
from GFun(1) consider (A) "f ps \<rightarrow> p |\<in>| \<Delta>\<^sub>A" | (B) "f ps \<rightarrow> p |\<in>| \<Delta>\<^sub>B" by (auto simp: rule_split)
then show ?case
proof cases
case A then obtain q' where eps_sp: "p = q' \<or> (p, q') |\<in>| \<Delta>\<^sub>\<E>\<^sub>A|\<^sup>+|"
"q' = q \<or> (q', q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>B|\<^sup>+|" using GFun(3, 6)
by (auto simp: transcl_eps_simp dest: eps_states_dest)
from GFun(4)[THEN ta_der_from_\<Delta>\<^sub>A] A GFun(2, 4)
have reach_fst: "p |\<in>| ta_der \<A> (term_of_gterm (GFun f ts))"
using A_args_states_not_B by auto
then have "q' |\<in>| ta_der \<A> (term_of_gterm (GFun f ts))" using eps_sp
by (meson ta_der_trancl_eps)
then show ?thesis using eps_sp(2)
by (intro exI[of _ "Var q'"]) (auto simp flip: ta_der_to_ta_der' simp del: ta_der'_simps)
next
case B
then have "p = q \<or> (p, q) |\<in>| \<Delta>\<^sub>\<E>\<^sub>B|\<^sup>+|" using GFun(3)
by (auto simp: transcl_eps_simp dest: B_rule_eps_A_False)
then show ?thesis using GFun(2, 4, 6) IH B
by (auto intro!: exI[of _ "Fun f (map h [0 ..< length ts])"] exI[of _ ps])
qed
qed
lemma ta_der'_split:
assumes "t |\<in>| ta_der' A (term_of_gterm s)"
shows "\<exists> u. u |\<in>| ta_der' \<A> (term_of_gterm s) \<and> t |\<in>| ta_der' \<B> u"
(is "\<exists> u. ?P s t u") using assms
proof (induct s arbitrary: t)
case (GFun f ts) show ?case
proof (cases t)
case [simp]: (Var q)
have "q |\<in>| ta_der A (term_of_gterm (GFun f ts))" using GFun(2)
by (auto simp flip: ta_der_to_ta_der')
from ta_der_split[OF this] ta_der_from_\<Delta>\<^sub>A[OF this] ta_state[OF this]
show ?thesis unfolding Var
by (metis ta_der'_refl ta_der_to_ta_der')
next
case [simp]: (Fun g ss)
obtain h where IH:
"\<forall> i < length ts. h i |\<in>| ta_der' \<A> (term_of_gterm (ts ! i))"
"\<forall> i < length ts. ss ! i |\<in>| ta_der' \<B> (h i)"
using GFun choice_nat[of "length ts" "\<lambda> t i. ?P (ts ! i) (ss ! i) t"]
by auto
then show ?thesis using GFun(2)
by (auto intro!: exI[of _ "Fun f (map h [0..<length ts])"])
qed
qed
(* TODO rewrite using ta_der'_mctxt_structure *)
lemma ta_der_to_mcxtx:
assumes "q |\<in>| ta_der A (term_of_gterm s)" and "q |\<in>| \<Q>\<^sub>B"
shows "\<exists> C ss qs. length qs = length ss \<and> num_holes C = length ss \<and>
(\<forall> i < length ss. qs ! i |\<in>| ta_der \<A> (term_of_gterm (ss ! i))) \<and>
q |\<in>| ta_der \<B> (fill_holes C (map Var qs)) \<and>
ground_mctxt C \<and> fill_holes C (map term_of_gterm ss) = term_of_gterm s"
(is "\<exists>C ss qs. ?P s q C ss qs")
proof -
from ta_der_split[OF assms] obtain t where
wit: "t |\<in>| ta_der' \<A> (term_of_gterm s)" "q |\<in>| ta_der \<B> t" by auto
let ?C = "fst (split_vars t)" let ?ss = "map (gsubt_at s) (varposs_list t)"
let ?qs = "snd (split_vars t)"
have poss [simp]:"i < length (varposs_list t) \<Longrightarrow> varposs_list t ! i \<in> gposs s" for i
by (metis nth_mem ta_der'_poss[OF wit(1)] poss_gposs_conv subset_eq varposs_eq_varposs_list
varposs_imp_poss varposs_list_var_terms_length)
have len: "num_holes ?C = length ?ss" "length ?ss = length ?qs"
by (simp_all add: split_vars_num_holes split_vars_vars_term_list varposs_list_var_terms_length)
from unfill_holes_to_subst_at_hole_poss[OF ta_der'_inf_mctxt[OF wit(1)]]
have "unfill_holes (fst (split_vars t)) (term_of_gterm s) = map (term_of_gterm \<circ> gsubt_at s) (varposs_list t)"
by (auto simp: comp_def hole_poss_split_vars_varposs_list
dest: in_set_idx intro!: nth_equalityI term_of_gterm_gsubt)
from fill_unfill_holes[OF ta_der'_inf_mctxt[OF wit(1)]] this
have rep: "fill_holes ?C (map term_of_gterm ?ss) = term_of_gterm s"
by simp
have reach_int: "i < length ?ss \<Longrightarrow> ?qs ! i |\<in>| ta_der \<A> (term_of_gterm (?ss ! i))" for i
using wit(1) ta_der'_varposs_to_ta_der
unfolding split_vars_vars_term_list length_map
unfolding varposs_list_to_var_term_list[symmetric]
by (metis nth_map nth_mem poss term_of_gterm_gsubt varposs_eq_varposs_list)
have reach_end: "q |\<in>| ta_der \<B> (fill_holes ?C (map Var ?qs))" using wit
using split_vars_fill_holes[of ?C t "map Var ?qs"]
by auto
show ?thesis using len rep reach_end reach_int
by (metis split_vars_ground')
qed
lemma ta_der_to_gmcxtx:
assumes "q |\<in>| ta_der A (term_of_gterm s)" and "q |\<in>| \<Q>\<^sub>B"
shows "\<exists> C ss qs qs'. length qs' = length qs \<and> length qs = length ss \<and> num_gholes C = length ss \<and>
(\<forall> i < length ss. qs ! i |\<in>| ta_der \<A> (term_of_gterm (ss ! i))) \<and>
q |\<in>| ta_der \<B> (fill_holes (mctxt_of_gmctxt C) (map Var qs')) \<and>
fill_gholes C ss = s"
using ta_der_to_mcxtx[OF assms]
by (metis gmctxt_of_mctxt_inv ground_gmctxt_of_gterm_of_term num_gholes_gmctxt_of_mctxt term_of_gterm_inv)
(* Reconstuction *)
lemma mctxt_const_to_ta_der:
assumes "num_holes C = length ss" "length ss = length qs"
and "\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> (ss ! i)"
and "q |\<in>| ta_der \<B> (fill_holes C (map Var qs))"
shows "q |\<in>| ta_der A (fill_holes C ss)"
proof -
have mid: "fill_holes C (map Var qs) |\<in>| ta_der' A (fill_holes C ss)"
using assms(1 - 3) ta_der_monos(1)
by (intro mctxt_args_ta_der') auto
then show ?thesis using assms(1, 2) ta_der_monos(2)[THEN fsubsetD, OF assms(4)]
using ta_der'_trans
by (simp add: ta_der'_ta_der)
qed
lemma ctxt_const_to_ta_der:
assumes "q |\<in>| ta_der \<A> s"
and "p |\<in>| ta_der \<B> C\<langle>Var q\<rangle>"
shows "p |\<in>| ta_der A C\<langle>s\<rangle>" using assms
by (meson fin_mono ta_der_ctxt ta_der_monos(1) ta_der_monos(2))
lemma gctxt_const_to_ta_der:
assumes "q |\<in>| ta_der \<A> (term_of_gterm s)"
and "p |\<in>| ta_der \<B> (ctxt_of_gctxt C)\<langle>Var q\<rangle>"
shows "p |\<in>| ta_der A (term_of_gterm C\<langle>s\<rangle>\<^sub>G)" using assms
by (metis ctxt_const_to_ta_der ctxt_of_gctxt_inv ground_ctxt_of_gctxt ground_gctxt_of_ctxt_apply_gterm)
end
end
\ No newline at end of file
diff --git a/thys/FSM_Tests/EquivalenceTesting/Simple_Convergence_Graph.thy b/thys/FSM_Tests/EquivalenceTesting/Simple_Convergence_Graph.thy
--- a/thys/FSM_Tests/EquivalenceTesting/Simple_Convergence_Graph.thy
+++ b/thys/FSM_Tests/EquivalenceTesting/Simple_Convergence_Graph.thy
@@ -1,2003 +1,2003 @@
section \<open>Simple Convergence Graphs\<close>
text \<open>This theory introduces a very simple implementation of convergence graphs that
consists of a list of convergent classes represented as sets of traces.\<close>
theory Simple_Convergence_Graph
imports Convergence_Graph
begin
subsection \<open>Basic Definitions\<close>
type_synonym 'a simple_cg = "'a list fset list"
definition simple_cg_empty :: "'a simple_cg" where
"simple_cg_empty = []"
(* collects all traces in the same convergent class set as ys *)
fun simple_cg_lookup :: "('a::linorder) simple_cg \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
"simple_cg_lookup xs ys = sorted_list_of_fset (finsert ys (foldl (|\<union>|) fempty (filter (\<lambda>x . ys |\<in>| x) xs)))"
(* collects all traces (zs@ys'') such that there exists a prefix ys' of ys with (ys=ys'@ys'')
and zs is in the same convergent class set as ys' *)
fun simple_cg_lookup_with_conv :: "('a::linorder) simple_cg \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
"simple_cg_lookup_with_conv g ys = (let
lookup_for_prefix = (\<lambda>i . let
pref = take i ys;
suff = drop i ys;
pref_conv = (foldl (|\<union>|) fempty (filter (\<lambda>x . pref |\<in>| x) g))
in fimage (\<lambda> pref' . pref'@suff) pref_conv)
in sorted_list_of_fset (finsert ys (foldl (\<lambda> cs i . lookup_for_prefix i |\<union>| cs) fempty [0..<Suc (length ys)])))"
fun simple_cg_insert' :: "('a::linorder) simple_cg \<Rightarrow> 'a list \<Rightarrow> 'a simple_cg" where
"simple_cg_insert' xs ys = (case find (\<lambda>x . ys |\<in>| x) xs
of Some x \<Rightarrow> xs |
None \<Rightarrow> {|ys|}#xs)"
fun simple_cg_insert :: "('a::linorder) simple_cg \<Rightarrow> 'a list \<Rightarrow> 'a simple_cg" where
"simple_cg_insert xs ys = foldl (\<lambda> xs' ys' . simple_cg_insert' xs' ys') xs (prefixes ys)"
fun simple_cg_initial :: "('a,'b::linorder,'c::linorder) fsm \<Rightarrow> ('b\<times>'c) prefix_tree \<Rightarrow> ('b\<times>'c) simple_cg" where
"simple_cg_initial M1 T = foldl (\<lambda> xs' ys' . simple_cg_insert' xs' ys') simple_cg_empty (filter (is_in_language M1 (initial M1)) (sorted_list_of_sequences_in_tree T))"
subsection \<open>Merging by Closure\<close>
text \<open>The following implementation of the merge operation follows the closure operation described
by Simão et al. in Simão, A., Petrenko, A. and Yevtushenko, N. (2012), On reducing test length
for FSMs with extra states. Softw. Test. Verif. Reliab., 22: 435-454. https://doi.org/10.1002/stvr.452.
That is, two traces u and v are merged by adding {u,v} to the list of convergent classes
followed by computing the closure of the graph based on two operations:
(1) classes A and B can be merged if there exists some class C such that C contains some w1, w2
and there exists some w such that A contains w1.w and B contains w2.w.
(2) classes A and B can be merged if one is a subset of the other.\<close>
(* classes x1 and x2 can be merged via class x if there exist \<alpha>, \<beta> in x and some suffix \<gamma>
such that x1 contains \<alpha>@\<gamma> and x2 contains \<beta>@\<gamma> *)
fun can_merge_by_suffix :: "'a list fset \<Rightarrow> 'a list fset \<Rightarrow> 'a list fset \<Rightarrow> bool" where
"can_merge_by_suffix x x1 x2 = (\<exists> \<alpha> \<beta> \<gamma> . \<alpha> |\<in>| x \<and> \<beta> |\<in>| x \<and> \<alpha>@\<gamma> |\<in>| x1 \<and> \<beta>@\<gamma> |\<in>| x2)"
lemma can_merge_by_suffix_code[code] :
"can_merge_by_suffix x x1 x2 =
(\<exists> ys \<in> fset x .
\<exists> ys1 \<in> fset x1 .
is_prefix ys ys1 \<and>
(\<exists> ys' \<in> fset x . ys'@(drop (length ys) ys1) |\<in>| x2))"
(is "?P1 = ?P2")
proof
show "?P1 \<Longrightarrow> ?P2"
- by (metis append_eq_conv_conj can_merge_by_suffix.elims(2) is_prefix_prefix notin_fset)
+ by (metis append_eq_conv_conj can_merge_by_suffix.elims(2) is_prefix_prefix)
show "?P2 \<Longrightarrow> ?P1"
- by (metis append_eq_conv_conj can_merge_by_suffix.elims(3) is_prefix_prefix notin_fset)
+ by (metis append_eq_conv_conj can_merge_by_suffix.elims(3) is_prefix_prefix)
qed
fun prefixes_in_list_helper :: "'a \<Rightarrow> 'a list list \<Rightarrow> (bool \<times> 'a list list) \<Rightarrow> bool \<times> 'a list list" where
"prefixes_in_list_helper x [] res = res" |
"prefixes_in_list_helper x ([]#yss) res = prefixes_in_list_helper x yss (True, snd res)" |
"prefixes_in_list_helper x ((y#ys)#yss) res =
(if x = y then prefixes_in_list_helper x yss (fst res, ys # snd res)
else prefixes_in_list_helper x yss res)"
fun prefixes_in_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list list \<Rightarrow> 'a list list \<Rightarrow> 'a list list" where
"prefixes_in_list [] prev yss res = (if List.member yss [] then prev#res else res)" |
"prefixes_in_list (x#xs) prev yss res = (let
(b,yss') = prefixes_in_list_helper x yss (False,[])
in if b then prefixes_in_list xs (prev@[x]) yss' (prev # res)
else prefixes_in_list xs (prev@[x]) yss' res)"
fun prefixes_in_set :: "('a::linorder) list \<Rightarrow> 'a list fset \<Rightarrow> 'a list list" where
"prefixes_in_set xs yss = prefixes_in_list xs [] (sorted_list_of_fset yss) []"
value "prefixes_in_list [1::nat,2,3,4,5] []
[ [1,2,3], [1,2,4], [1,3], [], [1], [1,5,3], [2,5] ] []"
value "prefixes_in_list_helper (1::nat)
[ [1,2,3], [1,2,4], [1,3], [], [1], [1,5,3], [2,5] ]
(False,[])"
lemma prefixes_in_list_helper_prop :
shows "fst (prefixes_in_list_helper x yss res) = (fst res \<or> [] \<in> list.set yss)" (is ?P1)
and "list.set (snd (prefixes_in_list_helper x yss res)) = list.set (snd res) \<union> {ys . x#ys \<in> list.set yss}" (is ?P2)
proof -
have "?P1 \<and> ?P2"
proof (induction yss arbitrary: res)
case Nil
then show ?case by auto
next
case (Cons ys yss)
show ?case proof (cases ys)
case Nil
then show ?thesis
using Cons.IH by auto
next
case (Cons y ys')
show ?thesis proof (cases "x = y")
case True
have *: "prefixes_in_list_helper x (ys # yss) res = prefixes_in_list_helper y yss (fst res, ys' # snd res)"
unfolding Cons True by auto
show ?thesis
using Cons.IH[of "(fst res, ys' # snd res)"]
unfolding *
unfolding Cons
unfolding True
by auto
next
case False
then have *: "prefixes_in_list_helper x (ys # yss) res = prefixes_in_list_helper x yss res"
unfolding Cons by auto
show ?thesis
unfolding *
unfolding Cons
using Cons.IH[of res] False
by force
qed
qed
qed
then show ?P1 and ?P2 by blast+
qed
lemma prefixes_in_list_prop :
shows "list.set (prefixes_in_list xs prev yss res) = list.set res \<union> {prev@ys | ys . ys \<in> list.set (prefixes xs) \<and> ys \<in> list.set yss}"
proof (induction xs arbitrary: prev yss res)
case Nil
show ?case
unfolding prefixes_in_list.simps List.member_def prefixes_set by auto
next
case (Cons x xs)
obtain b yss' where "prefixes_in_list_helper x yss (False,[]) = (b,yss')"
using prod.exhaust by metis
then have "b = ([] \<in> list.set yss)"
and "list.set yss' = {ys . x#ys \<in> list.set yss}"
using prefixes_in_list_helper_prop[of x yss "(False,[])"]
by auto
show ?case proof (cases b)
case True
then have *: "prefixes_in_list (x#xs) prev yss res = prefixes_in_list xs (prev@[x]) yss' (prev # res)"
using \<open>prefixes_in_list_helper x yss (False,[]) = (b,yss')\<close> by auto
show ?thesis
unfolding *
unfolding Cons \<open>list.set yss' = {ys . x#ys \<in> list.set yss}\<close>
using True unfolding \<open>b = ([] \<in> list.set yss)\<close>
by auto
next
case False
then have *: "prefixes_in_list (x#xs) prev yss res = prefixes_in_list xs (prev@[x]) yss' res"
using \<open>prefixes_in_list_helper x yss (False,[]) = (b,yss')\<close> by auto
show ?thesis
unfolding *
unfolding Cons \<open>list.set yss' = {ys . x#ys \<in> list.set yss}\<close>
using False unfolding \<open>b = ([] \<in> list.set yss)\<close>
by auto
qed
qed
lemma prefixes_in_set_prop :
"list.set (prefixes_in_set xs yss) = list.set (prefixes xs) \<inter> fset yss"
unfolding prefixes_in_set.simps
unfolding prefixes_in_list_prop
by auto
(* alternative implementation of merging *)
(*
lemma can_merge_by_suffix_code[code] :
"can_merge_by_suffix x x1 x2 =
(\<exists> ys1 \<in> fset x1 . list_ex (\<lambda>ys . ys |\<in>| x \<and> (\<exists> ys' \<in> fset x . ys'@(drop (length ys) ys1) |\<in>| x2))
(prefixes ys1))"
(is "?P1 = ?P2")
proof
show "?P1 \<Longrightarrow> ?P2"
proof -
assume "?P1"
then obtain \<alpha> \<beta> \<gamma> where "\<alpha> |\<in>| x" and "\<beta> |\<in>| x" and "\<alpha>@\<gamma> |\<in>| x1" and "\<beta>@\<gamma> |\<in>| x2"
by auto
have "\<alpha>@\<gamma> \<in> fset x1" using \<open>\<alpha>@\<gamma> |\<in>| x1\<close> notin_fset by metis
moreover have "\<alpha> \<in> list.set (prefixes (\<alpha>@\<gamma>))" by (simp add: prefixes_take_iff)
moreover note \<open>\<alpha> |\<in>| x\<close>
moreover have "\<exists> ys'' \<in> fset x . ys''@(drop (length \<alpha>) (\<alpha>@\<gamma>)) |\<in>| x2"
using \<open>\<beta>@\<gamma> |\<in>| x2\<close> \<open>\<beta> |\<in>| x\<close>
by (metis append_eq_conv_conj notin_fset)
ultimately show "?P2"
unfolding list_ex_iff by blast
qed
show "?P2 \<Longrightarrow> ?P1"
proof -
assume "?P2"
then obtain ys1 ys ys' where "ys1 \<in> fset x1"
and "ys \<in> list.set (prefixes ys1)"
and "ys |\<in>| x"
and "ys' \<in> fset x"
and "ys'@(drop (length ys) ys1) |\<in>| x2"
unfolding list_ex_iff by blast
then show "?P1"
by (metis append_take_drop_id can_merge_by_suffix.elims(3) notin_fset prefixes_take_iff)
qed
qed
*)
lemma can_merge_by_suffix_validity :
assumes "observable M1" and "observable M2"
and "\<And> u v . u |\<in>| x \<Longrightarrow> v |\<in>| x \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> u v . u |\<in>| x1 \<Longrightarrow> v |\<in>| x1 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> u v . u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "can_merge_by_suffix x x1 x2"
and "u |\<in>| (x1 |\<union>| x2)"
and "v |\<in>| (x1 |\<union>| x2)"
and "u \<in> L M1" and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
obtain \<alpha> \<beta> \<gamma> where "\<alpha> |\<in>| x" and "\<beta> |\<in>| x" and "\<alpha>@\<gamma> |\<in>| x1" and "\<beta>@\<gamma> |\<in>| x2"
using \<open>can_merge_by_suffix x x1 x2\<close> by auto
consider "u |\<in>| x1" | "u |\<in>| x2"
using \<open>u |\<in>| (x1 |\<union>| x2)\<close> by blast
then show ?thesis proof cases
case 1
then have "converge M1 u (\<alpha>@\<gamma>)" and "converge M2 u (\<alpha>@\<gamma>)"
using \<open>u |\<in>| (x1 |\<union>| x2)\<close> assms(4)[OF _ \<open>\<alpha>@\<gamma> |\<in>| x1\<close> assms(9,10)]
by blast+
then have "(\<alpha>@\<gamma>) \<in> L M1" and "(\<alpha>@\<gamma>) \<in> L M2"
by auto
then have "\<alpha> \<in> L M1" and "\<alpha> \<in> L M2"
using language_prefix by metis+
then have "converge M1 \<alpha> \<beta>" and "converge M2 \<alpha> \<beta>"
using assms(3) \<open>\<alpha> |\<in>| x\<close> \<open>\<beta> |\<in>| x\<close>
by blast+
have "converge M1 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)"
using \<open>converge M1 \<alpha> \<beta>\<close>
by (meson \<open>\<alpha> @ \<gamma> \<in> L M1\<close> assms(1) converge.simps converge_append)
then have "\<beta>@\<gamma> \<in> L M1"
by auto
have "converge M2 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)"
using \<open>converge M2 \<alpha> \<beta>\<close>
by (meson \<open>\<alpha> @ \<gamma> \<in> L M2\<close> assms(2) converge.simps converge_append)
then have "\<beta>@\<gamma> \<in> L M2"
by auto
consider (11) "v |\<in>| x1" | (12) "v |\<in>| x2"
using \<open>v |\<in>| (x1 |\<union>| x2)\<close> by blast
then show ?thesis proof cases
case 11
show ?thesis
using "1" "11" assms(10) assms(4) assms(9) by blast
next
case 12
then have "converge M1 v (\<beta>@\<gamma>)" and "converge M2 v (\<beta>@\<gamma>)"
using assms(5)[OF \<open>\<beta>@\<gamma> |\<in>| x2\<close> _ \<open>\<beta>@\<gamma> \<in> L M1\<close> \<open>\<beta>@\<gamma> \<in> L M2\<close>]
by auto
then show ?thesis
using \<open>converge M1 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)\<close> \<open>converge M2 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)\<close> \<open>converge M1 u (\<alpha>@\<gamma>)\<close> \<open>converge M2 u (\<alpha>@\<gamma>)\<close>
by auto
qed
next
case 2
then have "converge M1 u (\<beta>@\<gamma>)" and "converge M2 u (\<beta>@\<gamma>)"
using \<open>u |\<in>| (x1 |\<union>| x2)\<close> assms(5)[OF _ \<open>\<beta>@\<gamma> |\<in>| x2\<close> assms(9,10)]
by blast+
then have "(\<beta>@\<gamma>) \<in> L M1" and "(\<beta>@\<gamma>) \<in> L M2"
by auto
then have "\<beta> \<in> L M1" and "\<beta> \<in> L M2"
using language_prefix by metis+
then have "converge M1 \<alpha> \<beta>" and "converge M2 \<alpha> \<beta>"
using assms(3)[OF \<open>\<beta> |\<in>| x\<close> \<open>\<alpha> |\<in>| x\<close>]
by auto
have "converge M1 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)"
using \<open>converge M1 \<alpha> \<beta>\<close>
using \<open>\<beta> @ \<gamma> \<in> L M1\<close> \<open>\<beta> \<in> L M1\<close> assms(1) converge_append converge_append_language_iff by blast
then have "\<alpha>@\<gamma> \<in> L M1"
by auto
have "converge M2 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)"
using \<open>converge M2 \<alpha> \<beta>\<close>
using \<open>\<beta> @ \<gamma> \<in> L M2\<close> \<open>\<beta> \<in> L M2\<close> assms(2) converge_append converge_append_language_iff by blast
then have "\<alpha>@\<gamma> \<in> L M2"
by auto
consider (21) "v |\<in>| x1" | (22) "v |\<in>| x2"
using \<open>v |\<in>| (x1 |\<union>| x2)\<close> by blast
then show ?thesis proof cases
case 22
show ?thesis
using "2" "22" assms(10) assms(5) assms(9) by blast
next
case 21
then have "converge M1 v (\<alpha>@\<gamma>)" and "converge M2 v (\<alpha>@\<gamma>)"
using assms(4)[OF \<open>\<alpha>@\<gamma> |\<in>| x1\<close> _ \<open>\<alpha>@\<gamma> \<in> L M1\<close> \<open>\<alpha>@\<gamma> \<in> L M2\<close>]
by auto
then show ?thesis
using \<open>converge M1 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)\<close> \<open>converge M2 (\<alpha>@\<gamma>) (\<beta>@\<gamma>)\<close> \<open>converge M1 u (\<beta>@\<gamma>)\<close> \<open>converge M2 u (\<beta>@\<gamma>)\<close>
by auto
qed
qed
qed
fun simple_cg_closure_phase_1_helper' :: "'a list fset \<Rightarrow> 'a list fset \<Rightarrow> 'a simple_cg \<Rightarrow> (bool \<times> 'a list fset \<times> 'a simple_cg)" where
"simple_cg_closure_phase_1_helper' x x1 xs =
(let (x2s,others) = separate_by (can_merge_by_suffix x x1) xs;
x1Union = foldl (|\<union>|) x1 x2s
in (x2s \<noteq> [],x1Union,others))"
lemma simple_cg_closure_phase_1_helper'_False :
"\<not>fst (simple_cg_closure_phase_1_helper' x x1 xs) \<Longrightarrow> simple_cg_closure_phase_1_helper' x x1 xs = (False,x1,xs)"
unfolding simple_cg_closure_phase_1_helper'.simps Let_def separate_by.simps
by (simp add: filter_empty_conv)
lemma simple_cg_closure_phase_1_helper'_True :
assumes "fst (simple_cg_closure_phase_1_helper' x x1 xs)"
shows "length (snd (snd (simple_cg_closure_phase_1_helper' x x1 xs))) < length xs"
proof -
have "snd (snd (simple_cg_closure_phase_1_helper' x x1 xs)) = filter (\<lambda>x2 . \<not> (can_merge_by_suffix x x1 x2)) xs"
by auto
moreover have "filter (\<lambda>x2 . (can_merge_by_suffix x x1 x2)) xs \<noteq> []"
using assms unfolding simple_cg_closure_phase_1_helper'.simps Let_def separate_by.simps
by fastforce
ultimately show ?thesis
using filter_not_all_length[of "can_merge_by_suffix x x1" xs]
by metis
qed
lemma simple_cg_closure_phase_1_helper'_length :
"length (snd (snd (simple_cg_closure_phase_1_helper' x x1 xs))) \<le> length xs"
by auto
lemma simple_cg_closure_phase_1_helper'_validity_fst :
assumes "observable M1" and "observable M2"
and "\<And> u v . u |\<in>| x \<Longrightarrow> v |\<in>| x \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> u v . u |\<in>| x1 \<Longrightarrow> v |\<in>| x1 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> x2 u v . x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "u |\<in>| fst (snd (simple_cg_closure_phase_1_helper' x x1 xs))"
and "v |\<in>| fst (snd (simple_cg_closure_phase_1_helper' x x1 xs))"
and "u \<in> L M1" and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
have *:"\<And> w . w |\<in>| fst (snd (simple_cg_closure_phase_1_helper' x x1 xs)) \<Longrightarrow> w |\<in>| x1 \<or> (\<exists> x2 . x2 \<in> list.set xs \<and> w |\<in>| x2 \<and> can_merge_by_suffix x x1 x2)"
proof -
fix w assume "w |\<in>| fst (snd (simple_cg_closure_phase_1_helper' x x1 xs))"
then have "w |\<in>| ffUnion (fset_of_list (x1#(filter (can_merge_by_suffix x x1) xs)))"
using foldl_funion_fsingleton[where xs="(filter (can_merge_by_suffix x x1) xs)"]
by auto
then obtain x2 where "w |\<in>| x2"
and "x2 |\<in>| fset_of_list (x1#(filter (can_merge_by_suffix x x1) xs))"
using ffUnion_fmember_ob
by metis
then consider "x2=x1" | "x2 \<in> list.set (filter (can_merge_by_suffix x x1) xs)"
by (meson fset_of_list_elem set_ConsD)
then show "w |\<in>| x1 \<or> (\<exists> x2 . x2 \<in> list.set xs \<and> w |\<in>| x2 \<and> can_merge_by_suffix x x1 x2)"
using \<open>w |\<in>| x2\<close> by (cases; auto)
qed
consider "u |\<in>| x1" | "(\<exists> x2 . x2 \<in> list.set xs \<and> u |\<in>| x2 \<and> can_merge_by_suffix x x1 x2)"
using *[OF assms(6)] by blast
then show ?thesis proof cases
case 1
consider (a) "v |\<in>| x1" | (b) "(\<exists> x2 . x2 \<in> list.set xs \<and> v |\<in>| x2 \<and> can_merge_by_suffix x x1 x2)"
using *[OF assms(7)] by blast
then show ?thesis proof cases
case a
then show ?thesis using assms(4)[OF 1 _ assms(8,9)] by auto
next
case b
then obtain x2v where "x2v \<in> list.set xs" and "v |\<in>| x2v" and "can_merge_by_suffix x x1 x2v"
using *[OF assms(6)]
by blast
then have "u |\<in>| x1 |\<union>| x2v" and "v |\<in>| x1 |\<union>| x2v"
using 1 by auto
show ?thesis
using can_merge_by_suffix_validity[OF assms(1,2), of x x1 x2v, OF assms(3,4) assms(5)[OF \<open>x2v \<in> list.set xs\<close>] \<open>can_merge_by_suffix x x1 x2v\<close> \<open>u |\<in>| x1 |\<union>| x2v\<close> \<open>v |\<in>| x1 |\<union>| x2v\<close> assms(8,9)]
by blast
qed
next
case 2
then obtain x2u where "x2u \<in> list.set xs" and "u |\<in>| x2u" and "can_merge_by_suffix x x1 x2u"
using *[OF assms(6)]
by blast
then have "u |\<in>| x1 |\<union>| x2u"
by auto
consider (a) "v |\<in>| x1" | (b) "(\<exists> x2 . x2 \<in> list.set xs \<and> v |\<in>| x2 \<and> can_merge_by_suffix x x1 x2)"
using *[OF assms(7)] by blast
then show ?thesis proof cases
case a
then have "v |\<in>| x1 |\<union>| x2u"
by auto
show ?thesis
using can_merge_by_suffix_validity[OF assms(1,2), of x x1 x2u, OF assms(3,4) assms(5)[OF \<open>x2u \<in> list.set xs\<close>] \<open>can_merge_by_suffix x x1 x2u\<close> \<open>u |\<in>| x1 |\<union>| x2u\<close> \<open>v |\<in>| x1 |\<union>| x2u\<close> assms(8,9)]
by blast
next
case b
then obtain x2v where "x2v \<in> list.set xs" and "v |\<in>| x2v" and "can_merge_by_suffix x x1 x2v"
using *[OF assms(6)]
by blast
then have "v |\<in>| x1 |\<union>| x2v"
by auto
have "\<And> v . v |\<in>| x1 |\<union>| x2u \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using can_merge_by_suffix_validity[OF assms(1,2), of x x1 x2u, OF assms(3,4) assms(5)[OF \<open>x2u \<in> list.set xs\<close>] \<open>can_merge_by_suffix x x1 x2u\<close> \<open>u |\<in>| x1 |\<union>| x2u\<close> _ assms(8,9)]
by blast
have "\<And> u . u |\<in>| x1 |\<union>| x2v \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using can_merge_by_suffix_validity[OF assms(1,2), of x x1 x2v, OF assms(3,4) assms(5)[OF \<open>x2v \<in> list.set xs\<close>] \<open>can_merge_by_suffix x x1 x2v\<close> _ \<open>v |\<in>| x1 |\<union>| x2v\<close>]
by blast
obtain \<alpha>v \<beta>v \<gamma>v where "\<alpha>v |\<in>| x" and "\<beta>v |\<in>| x" and "\<alpha>v@\<gamma>v |\<in>| x1" and "\<beta>v@\<gamma>v |\<in>| x2v"
using \<open>can_merge_by_suffix x x1 x2v\<close> by auto
show ?thesis
using \<open>\<And>u. \<lbrakk>u |\<in>| x1 |\<union>| x2v; u \<in> L M1; u \<in> L M2\<rbrakk> \<Longrightarrow> converge M1 u v \<and> converge M2 u v\<close> \<open>\<And>v. v |\<in>| x1 |\<union>| x2u \<Longrightarrow> converge M1 u v \<and> converge M2 u v\<close> \<open>\<alpha>v @ \<gamma>v |\<in>| x1\<close> by fastforce
qed
qed
qed
lemma simple_cg_closure_phase_1_helper'_validity_snd :
assumes "\<And> x2 u v . x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (snd (snd (simple_cg_closure_phase_1_helper' x x1 xs)))"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1" and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
have "list.set (snd (snd (simple_cg_closure_phase_1_helper' x x1 xs))) \<subseteq> list.set xs"
by auto
then show ?thesis
using assms by blast
qed
fun simple_cg_closure_phase_1_helper :: "'a list fset \<Rightarrow> 'a simple_cg \<Rightarrow> (bool \<times> 'a simple_cg) \<Rightarrow> (bool \<times> 'a simple_cg)" where
"simple_cg_closure_phase_1_helper x [] (b,done) = (b,done)" |
"simple_cg_closure_phase_1_helper x (x1#xs) (b,done) = (let (hasChanged,x1',xs') = simple_cg_closure_phase_1_helper' x x1 xs
in simple_cg_closure_phase_1_helper x xs' (b \<or> hasChanged, x1' # done))"
lemma simple_cg_closure_phase_1_helper_validity :
assumes "observable M1" and "observable M2"
and "\<And> u v . u |\<in>| x \<Longrightarrow> v |\<in>| x \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> x2 u v . x2 \<in> list.set don \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> x2 u v . x2 \<in> list.set xss \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (snd (simple_cg_closure_phase_1_helper x xss (b,don)))"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1" and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
using assms(4,5,6)
proof (induction "length xss" arbitrary: xss don b rule: less_induct)
case less
show ?case proof (cases xss)
case Nil
then have "x2 \<in> list.set don"
using less.prems(3) by auto
then show ?thesis
using less.prems(1) assms(7,8,9,10)
by blast
next
case (Cons x1 xs)
obtain b' x1' xs' where "simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')"
using prod.exhaust by metis
then have "simple_cg_closure_phase_1_helper x xss (b,don) = simple_cg_closure_phase_1_helper x xs' (b \<or> b', x1' # don)"
unfolding Cons by auto
have *:"\<And> u v . u |\<in>| x1 \<Longrightarrow> v |\<in>| x1 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using less.prems(2)[of x1] unfolding Cons
by (meson list.set_intros(1))
have **:"\<And> x2 u v . x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using less.prems(2) unfolding Cons
by (meson list.set_intros(2))
have ***:"\<And> u v. u |\<in>| x1' \<Longrightarrow> v |\<in>| x1' \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using simple_cg_closure_phase_1_helper'_validity_fst[of M1 M2 x x1 xs _ _, OF assms(1,2,3) * **, of "\<lambda> a b c . a"]
unfolding \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> fst_conv snd_conv
by blast
have "length xs' < length xss"
using simple_cg_closure_phase_1_helper'_length[of x x1 xs]
unfolding \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> Cons by auto
have "(\<And>x2 u v. x2 \<in> list.set (x1' # don) \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using *** less.prems(1)
by (metis set_ConsD)
have "xs' = snd (snd (simple_cg_closure_phase_1_helper' x x1 xs))"
using \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> by auto
have "(\<And>x2 u v. x2 \<in> list.set xs' \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using simple_cg_closure_phase_1_helper'_validity_snd[of xs' M1]
unfolding \<open>xs' = snd (snd (simple_cg_closure_phase_1_helper' x x1 xs))\<close>
using ** simple_cg_closure_phase_1_helper'_validity_snd by blast
have "x2 \<in> list.set (snd (simple_cg_closure_phase_1_helper x xs' (b \<or> b', x1' # don)))"
using less.prems(3) unfolding \<open>simple_cg_closure_phase_1_helper x xss (b,don) = simple_cg_closure_phase_1_helper x xs' (b \<or> b', x1' # don)\<close> .
then show ?thesis
using less.hyps[OF \<open>length xs' < length xss\<close> \<open>(\<And>x2 u v. x2 \<in> list.set (x1' # don) \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)\<close> \<open>(\<And>x2 u v. x2 \<in> list.set xs' \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)\<close>, of "x1'#don" "\<lambda> a b c . a" "\<lambda> a b c . a"]
by force
qed
qed
lemma simple_cg_closure_phase_1_helper_length :
"length (snd (simple_cg_closure_phase_1_helper x xss (b,don))) \<le> length xss + length don"
proof (induction "length xss" arbitrary: xss b don rule: less_induct)
case less
show ?case proof (cases xss)
case Nil
then show ?thesis by auto
next
case (Cons x1 xs)
obtain b' x1' xs' where "simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')"
using prod.exhaust by metis
then have "simple_cg_closure_phase_1_helper x xss (b,don) = simple_cg_closure_phase_1_helper x xs' (b \<or> b', x1' # don)"
unfolding Cons by auto
have "length xs' < length xss"
using simple_cg_closure_phase_1_helper'_length[of x x1 xs]
unfolding \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> Cons by auto
then have "length (snd (simple_cg_closure_phase_1_helper x xs' (b \<or> b', x1'#don))) \<le> length xs' + length (x1'#don)"
using less[of xs'] unfolding Cons by blast
moreover have "length xs' + length (x1'#don) \<le> length xss + length don"
using simple_cg_closure_phase_1_helper'_length[of x x1 xs]
unfolding \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> snd_conv Cons by auto
ultimately show ?thesis
unfolding \<open>simple_cg_closure_phase_1_helper x xss (b,don) = simple_cg_closure_phase_1_helper x xs' (b \<or> b', x1' # don)\<close>
by presburger
qed
qed
lemma simple_cg_closure_phase_1_helper_True :
assumes "fst (simple_cg_closure_phase_1_helper x xss (False,don))"
and "xss \<noteq> []"
shows "length (snd (simple_cg_closure_phase_1_helper x xss (False,don))) < length xss + length don"
using assms
proof (induction "length xss" arbitrary: xss don rule: less_induct)
case less
show ?case proof (cases xss)
case Nil
then show ?thesis using less.prems(2) by auto
next
case (Cons x1 xs)
obtain b' x1' xs' where "simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')"
using prod.exhaust by metis
then have "simple_cg_closure_phase_1_helper x xss (False,don) = simple_cg_closure_phase_1_helper x xs' (b', x1' # don)"
unfolding Cons by auto
show ?thesis proof (cases b')
case True
then have "length xs' < length xs"
using simple_cg_closure_phase_1_helper'_True[of x x1 xs]
unfolding \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> fst_conv snd_conv
by blast
then have "length (snd (simple_cg_closure_phase_1_helper x xs' (b', x1' # don))) < length xss + length don"
using simple_cg_closure_phase_1_helper_length[of x xs' b' "x1'#don"]
unfolding Cons
by auto
then show ?thesis
unfolding \<open>simple_cg_closure_phase_1_helper x xss (False,don) = simple_cg_closure_phase_1_helper x xs' (b', x1' # don)\<close> .
next
case False
then have "simple_cg_closure_phase_1_helper x xss (False,don) = simple_cg_closure_phase_1_helper x xs' (False, x1' # don)"
using \<open>simple_cg_closure_phase_1_helper x xss (False,don) = simple_cg_closure_phase_1_helper x xs' (b', x1' # don)\<close>
by auto
then have "fst (simple_cg_closure_phase_1_helper x xs' (False, x1' # don))"
using less.prems(1) by auto
have "length xs' < length xss"
using simple_cg_closure_phase_1_helper'_length[of x x1 xs]
unfolding \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> Cons by auto
have "xs' \<noteq> []"
using \<open>simple_cg_closure_phase_1_helper' x x1 xs = (b',x1',xs')\<close> False
by (metis \<open>fst (simple_cg_closure_phase_1_helper x xs' (False, x1' # don))\<close> simple_cg_closure_phase_1_helper.simps(1) fst_eqD)
show ?thesis
using less.hyps[OF \<open>length xs' < length xss\<close> \<open>fst (simple_cg_closure_phase_1_helper x xs' (False, x1' # don))\<close> \<open>xs' \<noteq> []\<close>] \<open>length xs' < length xss\<close>
unfolding \<open>simple_cg_closure_phase_1_helper x xss (False,don) = simple_cg_closure_phase_1_helper x xs' (False, x1' # don)\<close>
unfolding Cons
by auto
qed
qed
qed
(* closure operation (1) *)
fun simple_cg_closure_phase_1 :: "'a simple_cg \<Rightarrow> (bool \<times> 'a simple_cg)" where
"simple_cg_closure_phase_1 xs = foldl (\<lambda> (b,xs) x. let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xs"
lemma simple_cg_closure_phase_1_validity :
assumes "observable M1" and "observable M2"
and "\<And> x2 u v . x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (snd (simple_cg_closure_phase_1 xs))"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1" and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
have "\<And> xss x2 u v . (\<And> x2 u v . x2 \<in> list.set xss \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v) \<Longrightarrow> x2 \<in> list.set (snd (foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss)) \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
proof -
fix xss x2 u v
assume "\<And> x2 u v . x2 \<in> list.set xss \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (snd (foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss))"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1"
and "u \<in> L M2"
then show "converge M1 u v \<and> converge M2 u v"
proof (induction xss arbitrary: x2 u v rule: rev_induct)
case Nil
then have "x2 \<in> list.set xs"
by auto
then show ?case
using Nil.prems(3,4,5,6) assms(3) by blast
next
case (snoc x xss)
obtain b xss' where "(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss) = (b,xss')"
using prod.exhaust by metis
moreover obtain b' xss'' where "simple_cg_closure_phase_1_helper x xss' (False,[]) = (b',xss'')"
using prod.exhaust by metis
ultimately have *:"(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) (xss@[x])) = (b\<or>b',xss'')"
by auto
have "(\<And>u v. u |\<in>| x \<Longrightarrow> v |\<in>| x \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using snoc.prems(1)
by (metis append_Cons list.set_intros(1) list_set_sym)
moreover have "(\<And>x2 u v. x2 \<in> list.set [] \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
by auto
moreover have "(\<And>x2 u v. x2 \<in> list.set xss' \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
proof -
have "(\<And>x2 u v. x2 \<in> list.set xss \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using snoc.prems(1)
by (metis (no_types, lifting) append_Cons append_Nil2 insertCI list.simps(15) list_set_sym)
then show "(\<And>x2 u v. x2 \<in> list.set xss' \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using snoc.IH
unfolding \<open>(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss) = (b,xss')\<close> snd_conv
by blast
qed
ultimately have "(\<And>x2 u v. x2 \<in> list.set xss'' \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using simple_cg_closure_phase_1_helper_validity[OF assms(1,2), of x "[]" xss' _ False]
unfolding \<open>simple_cg_closure_phase_1_helper x xss' (False,[]) = (b',xss'')\<close> snd_conv
by blast
then show ?case
using snoc.prems(2,3,4,5,6)
unfolding * snd_conv
by blast
qed
qed
then show ?thesis
using assms(3,4,5,6,7,8)
unfolding simple_cg_closure_phase_1.simps
by blast
qed
lemma simple_cg_closure_phase_1_length_helper :
"length (snd (foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss)) \<le> length xs"
proof (induction xss rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc x xss)
obtain b xss' where "(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss) = (b,xss')"
using prod.exhaust by metis
moreover obtain b' xss'' where "simple_cg_closure_phase_1_helper x xss' (False,[]) = (b',xss'')"
using prod.exhaust by metis
ultimately have *:"(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) (xss@[x])) = (b\<or>b',xss'')"
by auto
have "length xss' \<le> length xs"
using snoc.IH
unfolding \<open>(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss) = (b,xss')\<close>
by auto
moreover have "length xss'' \<le> length xss'"
using simple_cg_closure_phase_1_helper_length[of x xss' False "[]"]
unfolding \<open>simple_cg_closure_phase_1_helper x xss' (False,[]) = (b',xss'')\<close>
by auto
ultimately show ?case
unfolding * snd_conv
by simp
qed
lemma simple_cg_closure_phase_1_length :
"length (snd (simple_cg_closure_phase_1 xs)) \<le> length xs"
using simple_cg_closure_phase_1_length_helper by auto
lemma simple_cg_closure_phase_1_True :
assumes "fst (simple_cg_closure_phase_1 xs)"
shows "length (snd (simple_cg_closure_phase_1 xs)) < length xs"
proof -
have "\<And> xss . fst (foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss) \<Longrightarrow> length (snd (foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss)) < length xs"
proof -
fix xss
assume "fst (foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss)"
then show "length (snd (foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss)) < length xs"
proof (induction xss rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc x xss)
obtain b xss' where "(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss) = (b,xss')"
using prod.exhaust by metis
moreover obtain b' xss'' where "simple_cg_closure_phase_1_helper x xss' (False,[]) = (b',xss'')"
using prod.exhaust by metis
ultimately have "(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) (xss@[x])) = (b\<or>b',xss'')"
by auto
consider b | b'
using snoc.prems
unfolding \<open>(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) (xss@[x])) = (b\<or>b',xss'')\<close> fst_conv
by blast
then show ?case proof cases
case 1
then have "length xss' < length xs"
using snoc.IH
unfolding \<open>(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) xss) = (b,xss')\<close> fst_conv snd_conv
by auto
moreover have "length xss'' \<le> length xss'"
using simple_cg_closure_phase_1_helper_length[of x xss' False "[]"]
unfolding \<open>simple_cg_closure_phase_1_helper x xss' (False,[]) = (b',xss'')\<close>
by auto
ultimately show ?thesis
unfolding \<open>(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) (xss@[x])) = (b\<or>b',xss'')\<close> snd_conv
by simp
next
case 2
have "length xss' \<le> length xs"
using simple_cg_closure_phase_1_length_helper[of xss xs]
by (metis \<open>foldl (\<lambda>(b, xs) x. let (b', xs') = simple_cg_closure_phase_1_helper x xs (False, []) in (b \<or> b', xs')) (False, xs) xss = (b, xss')\<close> simple_cg_closure_phase_1_length_helper snd_conv)
moreover have "length xss'' < length xss'"
proof -
have "xss' \<noteq> []"
using "2" \<open>simple_cg_closure_phase_1_helper x xss' (False, []) = (b', xss'')\<close> by auto
then show ?thesis
using simple_cg_closure_phase_1_helper_True[of x xss' "[]"] 2
unfolding \<open>simple_cg_closure_phase_1_helper x xss' (False,[]) = (b',xss'')\<close> fst_conv snd_conv
by auto
qed
ultimately show ?thesis
unfolding \<open>(foldl (\<lambda> (b,xs) x . let (b',xs') = simple_cg_closure_phase_1_helper x xs (False,[]) in (b\<or>b',xs')) (False,xs) (xss@[x])) = (b\<or>b',xss'')\<close> snd_conv
by simp
qed
qed
qed
then show ?thesis
using assms by auto
qed
fun can_merge_by_intersection :: "'a list fset \<Rightarrow> 'a list fset \<Rightarrow> bool" where
"can_merge_by_intersection x1 x2 = (\<exists> \<alpha> . \<alpha> |\<in>| x1 \<and> \<alpha> |\<in>| x2)"
lemma can_merge_by_intersection_code[code] :
"can_merge_by_intersection x1 x2 = (\<exists> \<alpha> \<in> fset x1 . \<alpha> |\<in>| x2)"
unfolding can_merge_by_intersection.simps
- by (meson notin_fset)
+ by metis
lemma can_merge_by_intersection_validity :
assumes "\<And> u v . u |\<in>| x1 \<Longrightarrow> v |\<in>| x1 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> u v . u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "can_merge_by_intersection x1 x2"
and "u |\<in>| (x1 |\<union>| x2)"
and "v |\<in>| (x1 |\<union>| x2)"
and "u \<in> L M1"
and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
obtain \<alpha> where "\<alpha> |\<in>| x1" and "\<alpha> |\<in>| x2"
using assms(3) by auto
have "converge M1 u \<alpha> \<and> converge M2 u \<alpha>"
using \<open>\<alpha> |\<in>| x1\<close> \<open>\<alpha> |\<in>| x2\<close> assms(1,2,4,6,7) by blast
moreover have "converge M1 v \<alpha> \<and> converge M2 v \<alpha>"
- by (metis \<open>\<alpha> |\<in>| x1\<close> \<open>\<alpha> |\<in>| x2\<close> assms(1) assms(2) assms(5) calculation converge.elims(2) funion_iff)
+ by (metis (no_types, opaque_lifting) \<open>\<alpha> |\<in>| x1\<close> \<open>\<alpha> |\<in>| x2\<close> assms(1) assms(2) assms(5) calculation converge.simps funion_iff)
ultimately show ?thesis
by simp
qed
fun simple_cg_closure_phase_2_helper :: "'a list fset \<Rightarrow> 'a simple_cg \<Rightarrow> (bool \<times> 'a list fset \<times> 'a simple_cg)" where
"simple_cg_closure_phase_2_helper x1 xs =
(let (x2s,others) = separate_by (can_merge_by_intersection x1) xs;
x1Union = foldl (|\<union>|) x1 x2s
in (x2s \<noteq> [],x1Union,others))"
lemma simple_cg_closure_phase_2_helper_length :
"length (snd (snd (simple_cg_closure_phase_2_helper x1 xs))) \<le> length xs"
by auto
lemma simple_cg_closure_phase_2_helper_validity_fst :
assumes "\<And> u v . u |\<in>| x1 \<Longrightarrow> v |\<in>| x1 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> x2 u v . x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "u |\<in>| fst (snd (simple_cg_closure_phase_2_helper x1 xs))"
and "v |\<in>| fst (snd (simple_cg_closure_phase_2_helper x1 xs))"
and "u \<in> L M1"
and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
have *:"\<And> w . w |\<in>| fst (snd (simple_cg_closure_phase_2_helper x1 xs)) \<Longrightarrow> w |\<in>| x1 \<or> (\<exists> x2 . x2 \<in> list.set xs \<and> w |\<in>| x2 \<and> can_merge_by_intersection x1 x2)"
proof -
fix w assume "w |\<in>| fst (snd (simple_cg_closure_phase_2_helper x1 xs))"
then have "w |\<in>| ffUnion (fset_of_list (x1#(filter (can_merge_by_intersection x1) xs)))"
using foldl_funion_fsingleton[where xs="(filter (can_merge_by_intersection x1) xs)"]
by auto
then obtain x2 where "w |\<in>| x2"
and "x2 |\<in>| fset_of_list (x1#(filter (can_merge_by_intersection x1) xs))"
using ffUnion_fmember_ob
by metis
then consider "x2=x1" | "x2 \<in> list.set (filter (can_merge_by_intersection x1) xs)"
by (meson fset_of_list_elem set_ConsD)
then show "w |\<in>| x1 \<or> (\<exists> x2 . x2 \<in> list.set xs \<and> w |\<in>| x2 \<and> can_merge_by_intersection x1 x2)"
using \<open>w |\<in>| x2\<close> by (cases; auto)
qed
consider "u |\<in>| x1" | "(\<exists> x2 . x2 \<in> list.set xs \<and> u |\<in>| x2 \<and> can_merge_by_intersection x1 x2)"
using *[OF assms(3)] by blast
then show ?thesis proof cases
case 1
consider (a) "v |\<in>| x1" | (b) "(\<exists> x2 . x2 \<in> list.set xs \<and> v |\<in>| x2 \<and> can_merge_by_intersection x1 x2)"
using *[OF assms(4)] by blast
then show ?thesis proof cases
case a
then show ?thesis using assms(1)[OF 1 _ assms(5,6)] by auto
next
case b
then obtain x2v where "x2v \<in> list.set xs" and "v |\<in>| x2v" and "can_merge_by_intersection x1 x2v"
using *[OF assms(3)]
by blast
show ?thesis
using can_merge_by_intersection_validity[of x1 M1 M2 x2v, OF assms(1) assms(2)[OF \<open>x2v \<in> list.set xs\<close>] \<open>can_merge_by_intersection x1 x2v\<close>]
using 1 \<open>v |\<in>| x2v\<close> assms(5,6)
by blast
qed
next
case 2
then obtain x2u where "x2u \<in> list.set xs" and "u |\<in>| x2u" and "can_merge_by_intersection x1 x2u"
using *[OF assms(3)]
by blast
obtain \<alpha>u where "\<alpha>u |\<in>| x1" and "\<alpha>u |\<in>| x2u"
using \<open>can_merge_by_intersection x1 x2u\<close> by auto
consider (a) "v |\<in>| x1" | (b) "(\<exists> x2 . x2 \<in> list.set xs \<and> v |\<in>| x2 \<and> can_merge_by_intersection x1 x2)"
using *[OF assms(4)] by blast
then show ?thesis proof cases
case a
show ?thesis
using can_merge_by_intersection_validity[of x1 M1 M2 x2u, OF assms(1) assms(2)[OF \<open>x2u \<in> list.set xs\<close>] \<open>can_merge_by_intersection x1 x2u\<close>]
using \<open>u |\<in>| x2u\<close> a assms(5,6)
by blast
next
case b
then obtain x2v where "x2v \<in> list.set xs" and "v |\<in>| x2v" and "can_merge_by_intersection x1 x2v"
using *[OF assms(4)]
by blast
obtain \<alpha>v where "\<alpha>v |\<in>| x1" and "\<alpha>v |\<in>| x2v"
using \<open>can_merge_by_intersection x1 x2v\<close> by auto
have "\<And> v . v |\<in>| x1 |\<union>| x2u \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using can_merge_by_intersection_validity[of x1 M1 M2 x2u, OF assms(1) assms(2)[OF \<open>x2u \<in> list.set xs\<close>] \<open>can_merge_by_intersection x1 x2u\<close> _ _ assms(5,6)] \<open>u |\<in>| x2u\<close>
by blast
have "\<And> u . u |\<in>| x1 |\<union>| x2v \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using can_merge_by_intersection_validity[of x1 M1 M2 x2v, OF assms(1) assms(2)[OF \<open>x2v \<in> list.set xs\<close>] \<open>can_merge_by_intersection x1 x2v\<close> ] \<open>v |\<in>| x2v\<close>
by blast
show ?thesis
using \<open>\<And>u. \<lbrakk>u |\<in>| x1 |\<union>| x2v; u \<in> L M1; u \<in> L M2\<rbrakk> \<Longrightarrow> converge M1 u v \<and> converge M2 u v\<close> \<open>\<And>v. v |\<in>| x1 |\<union>| x2u \<Longrightarrow> converge M1 u v \<and> converge M2 u v\<close> \<open>\<alpha>u |\<in>| x1\<close> by fastforce
qed
qed
qed
lemma simple_cg_closure_phase_2_helper_validity_snd :
assumes "\<And> x2 u v . x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (snd (snd (simple_cg_closure_phase_2_helper x1 xs)))"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1"
and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
have "list.set (snd (snd (simple_cg_closure_phase_2_helper x1 xs))) \<subseteq> list.set xs"
by auto
then show ?thesis
using assms by blast
qed
lemma simple_cg_closure_phase_2_helper_True :
assumes "fst (simple_cg_closure_phase_2_helper x xs)"
shows "length (snd (snd (simple_cg_closure_phase_2_helper x xs))) < length xs"
proof -
have "snd (snd (simple_cg_closure_phase_2_helper x xs)) = filter (\<lambda>x2 . \<not> (can_merge_by_intersection x x2)) xs"
by auto
moreover have "filter (\<lambda>x2 . (can_merge_by_intersection x x2)) xs \<noteq> []"
using assms unfolding simple_cg_closure_phase_1_helper'.simps Let_def separate_by.simps
by fastforce
ultimately show ?thesis
using filter_not_all_length[of "can_merge_by_intersection x" xs]
by metis
qed
function simple_cg_closure_phase_2' :: "'a simple_cg \<Rightarrow> (bool \<times> 'a simple_cg) \<Rightarrow> (bool \<times> 'a simple_cg)" where
"simple_cg_closure_phase_2' [] (b,done) = (b,done)" |
"simple_cg_closure_phase_2' (x#xs) (b,done) = (let (hasChanged,x',xs') = simple_cg_closure_phase_2_helper x xs
in if hasChanged then simple_cg_closure_phase_2' xs' (True,x'#done)
else simple_cg_closure_phase_2' xs (b,x#done))"
by pat_completeness auto
termination
proof -
{
fix xa :: "(bool \<times> 'a list fset \<times> 'a simple_cg)"
fix x xs b don xb y xaa ya
assume "xa = simple_cg_closure_phase_2_helper x xs"
and "(xb, y) = xa"
and "(xaa, ya) = y"
and xb
have "length ya < Suc (length xs)"
using simple_cg_closure_phase_2_helper_True[of x xs] \<open>xb\<close>
unfolding \<open>xa = simple_cg_closure_phase_2_helper x xs\<close>[symmetric]
unfolding \<open>(xb, y) = xa\<close>[symmetric] \<open>(xaa, ya) = y\<close>[symmetric]
unfolding fst_conv snd_conv
by auto
then have "((ya, True, xaa # don), x # xs, b, don) \<in> measure (\<lambda>(xs, bd). length xs)"
by auto
}
then show ?thesis
apply (relation "measure (\<lambda> (xs,bd) . length xs)")
by force+
qed
lemma simple_cg_closure_phase_2'_validity :
assumes "\<And> x2 u v . x2 \<in> list.set don \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "\<And> x2 u v . x2 \<in> list.set xss \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (snd (simple_cg_closure_phase_2' xss (b,don)))"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1"
and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
using assms(1,2,3)
proof (induction "length xss" arbitrary: xss b don rule: less_induct)
case less
show ?case proof (cases xss)
case Nil
show ?thesis using less.prems(3) less.prems(1)[OF _ assms(4,5,6,7)] unfolding Nil
by auto
next
case (Cons x xs)
obtain hasChanged x' xs' where "simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')"
using prod.exhaust by metis
show ?thesis proof (cases hasChanged)
case True
then have "simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs' (True,x'#don)"
using \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close>
unfolding Cons
by auto
have *:"(\<And>u v. u |\<in>| x \<Longrightarrow> v |\<in>| x \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)" and
**:"(\<And>x2 u v. x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using less.prems(2) unfolding Cons
by (meson list.set_intros)+
have "length xs' < length xss"
unfolding Cons
using simple_cg_closure_phase_2_helper_True[of x xs] True
unfolding \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close> fst_conv snd_conv
by auto
moreover have "(\<And>x2 u v. x2 \<in> list.set (x' # don) \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using simple_cg_closure_phase_2_helper_validity_fst[of x M1 M2 xs, OF * **, of "\<lambda> a b c . a"]
using less.prems(1)
unfolding \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close> fst_conv snd_conv
using set_ConsD[of _ x' don]
by blast
moreover have "(\<And>x2 u v. x2 \<in> list.set xs' \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using simple_cg_closure_phase_2_helper_validity_snd[of xs M1 M2 _ x, OF **, of "\<lambda> a b c . a"]
unfolding \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close> fst_conv snd_conv
by blast
moreover have "x2 \<in> list.set (snd (simple_cg_closure_phase_2' xs' (True, x' # don)))"
using less.prems(3) unfolding \<open>simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs' (True,x'#don)\<close> .
ultimately show ?thesis
using less.hyps[of xs' "x'#don"]
by blast
next
case False
then have "simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs (b,x#don)"
using \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close>
unfolding Cons
by auto
have "length xs < length xss"
unfolding Cons by auto
moreover have "(\<And>x2 u v. x2 \<in> list.set (x # don) \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using less.prems(1,2) unfolding Cons
by (metis list.set_intros(1) set_ConsD)
moreover have "(\<And>x2 u v. x2 \<in> list.set xs \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using less.prems(2) unfolding Cons
by (metis list.set_intros(2))
moreover have "x2 \<in> list.set (snd (simple_cg_closure_phase_2' xs (b, x # don)))"
using less.prems(3)
unfolding \<open>simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs (b,x#don)\<close>
unfolding Cons .
ultimately show ?thesis
using less.hyps[of xs "x#don" b]
by blast
qed
qed
qed
lemma simple_cg_closure_phase_2'_length :
"length (snd (simple_cg_closure_phase_2' xss (b,don))) \<le> length xss + length don"
proof (induction "length xss" arbitrary: xss b don rule: less_induct)
case less
show ?case proof (cases xss)
case Nil
then show ?thesis by auto
next
case (Cons x xs)
obtain hasChanged x' xs' where "simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')"
using prod.exhaust by metis
show ?thesis proof (cases hasChanged)
case True
then have "simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs' (True,x'#don)"
using \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close>
unfolding Cons
by auto
have "length xs' < length xss"
using simple_cg_closure_phase_2_helper_True[of x xs] True
unfolding \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close> snd_conv fst_conv
unfolding Cons
by auto
then show ?thesis
using less.hyps[of xs' True "x'#don"]
unfolding \<open>simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs' (True,x'#don)\<close>
unfolding Cons by auto
next
case False
then have "simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs (b,x#don)"
using \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close>
unfolding Cons
by auto
show ?thesis
using less.hyps[of xs b "x#don"]
unfolding \<open>simple_cg_closure_phase_2' xss (b,don) = simple_cg_closure_phase_2' xs (b,x#don)\<close>
unfolding Cons
by auto
qed
qed
qed
lemma simple_cg_closure_phase_2'_True :
assumes "fst (simple_cg_closure_phase_2' xss (False,don))"
and "xss \<noteq> []"
shows "length (snd (simple_cg_closure_phase_2' xss (False,don))) < length xss + length don"
using assms
proof (induction "length xss" arbitrary: xss don rule: less_induct)
case less
show ?case proof (cases xss)
case Nil
then show ?thesis
using less.prems(2) by auto
next
case (Cons x xs)
obtain hasChanged x' xs' where "simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')"
using prod.exhaust by metis
show ?thesis proof (cases hasChanged)
case True
then have "simple_cg_closure_phase_2' xss (False,don) = simple_cg_closure_phase_2' xs' (True,x'#don)"
using \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close>
unfolding Cons
by auto
have "length xs' < length xs"
using simple_cg_closure_phase_2_helper_True[of x xs] True
unfolding \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close> snd_conv fst_conv
unfolding Cons
by auto
moreover have "length (snd (simple_cg_closure_phase_2' xs' (True,x'#don))) \<le> length xs' + length (x'#don)"
using simple_cg_closure_phase_2'_length by metis
ultimately show ?thesis
unfolding \<open>simple_cg_closure_phase_2' xss (False,don) = simple_cg_closure_phase_2' xs' (True,x'#don)\<close>
unfolding Cons
by auto
next
case False
then have "simple_cg_closure_phase_2' xss (False,don) = simple_cg_closure_phase_2' xs (False,x#don)"
using \<open>simple_cg_closure_phase_2_helper x xs = (hasChanged,x',xs')\<close>
unfolding Cons
by auto
have "xs \<noteq> []"
using \<open>simple_cg_closure_phase_2' xss (False, don) = simple_cg_closure_phase_2' xs (False, x # don)\<close> less.prems(1) by auto
show ?thesis
using less.hyps[of xs "x#don", OF _ _ \<open>xs \<noteq> []\<close>]
using less.prems(1)
unfolding \<open>simple_cg_closure_phase_2' xss (False,don) = simple_cg_closure_phase_2' xs (False,x#don)\<close>
unfolding Cons
by auto
qed
qed
qed
(* closure operation (2) *)
fun simple_cg_closure_phase_2 :: "'a simple_cg \<Rightarrow> (bool \<times> 'a simple_cg)" where
"simple_cg_closure_phase_2 xs = simple_cg_closure_phase_2' xs (False,[])"
lemma simple_cg_closure_phase_2_validity :
assumes "\<And> x2 u v . x2 \<in> list.set xss \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (snd (simple_cg_closure_phase_2 xss))"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1"
and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
using assms(2)
unfolding simple_cg_closure_phase_2.simps
using simple_cg_closure_phase_2'_validity[OF _ assms(1) _ assms(3,4,5,6), of "[]" xss "\<lambda> a b c . a" False]
by auto
lemma simple_cg_closure_phase_2_length :
"length (snd (simple_cg_closure_phase_2 xss)) \<le> length xss"
unfolding simple_cg_closure_phase_2.simps
using simple_cg_closure_phase_2'_length[of xss False "[]"]
by auto
lemma simple_cg_closure_phase_2_True :
assumes "fst (simple_cg_closure_phase_2 xss)"
shows "length (snd (simple_cg_closure_phase_2 xss)) < length xss"
proof -
have "xss \<noteq> []"
using assms by auto
then show ?thesis
using simple_cg_closure_phase_2'_True[of xss "[]"] assms by auto
qed
function simple_cg_closure :: "'a simple_cg \<Rightarrow> 'a simple_cg" where
"simple_cg_closure g = (let (hasChanged1,g1) = simple_cg_closure_phase_1 g;
(hasChanged2,g2) = simple_cg_closure_phase_2 g1
in if hasChanged1 \<or> hasChanged2
then simple_cg_closure g2
else g2)"
by pat_completeness auto
termination
proof -
{
fix g :: "'a simple_cg"
fix x hasChanged1 g1 xb hasChanged2 g2
assume "x = simple_cg_closure_phase_1 g"
"(hasChanged1, g1) = x"
"xb = simple_cg_closure_phase_2 g1"
"(hasChanged2, g2) = xb"
"hasChanged1 \<or> hasChanged2"
then have "simple_cg_closure_phase_1 g = (hasChanged1, g1)"
and "simple_cg_closure_phase_2 g1 = (hasChanged2, g2)"
by auto
have "length g1 \<le> length g"
using \<open>simple_cg_closure_phase_1 g = (hasChanged1, g1)\<close>
using simple_cg_closure_phase_1_length[of g]
by auto
have "length g2 \<le> length g1"
using \<open>simple_cg_closure_phase_2 g1 = (hasChanged2, g2)\<close>
using simple_cg_closure_phase_2_length[of g1]
by auto
consider hasChanged1 | hasChanged2
using \<open>hasChanged1 \<or> hasChanged2\<close> by blast
then have "length g2 < length g"
proof cases
case 1
then have "length g1 < length g"
using \<open>simple_cg_closure_phase_1 g = (hasChanged1, g1)\<close>
using simple_cg_closure_phase_1_True[of g]
by auto
then show ?thesis
using \<open>length g2 \<le> length g1\<close>
by linarith
next
case 2
then have "length g2 < length g1"
using \<open>simple_cg_closure_phase_2 g1 = (hasChanged2, g2)\<close>
using simple_cg_closure_phase_2_True[of g1]
by auto
then show ?thesis
using \<open>length g1 \<le> length g\<close>
by linarith
qed
then have "(g2, g) \<in> measure length"
by auto
}
then show ?thesis by (relation "measure length"; force)
qed
lemma simple_cg_closure_validity :
assumes "observable M1" and "observable M2"
and "\<And> x2 u v . x2 \<in> list.set g \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (simple_cg_closure g)"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1"
and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
using assms(3,4)
proof (induction "length g" arbitrary: g rule: less_induct)
case less
obtain hasChanged1 hasChanged2 g1 g2 where "simple_cg_closure_phase_1 g = (hasChanged1, g1)"
and "simple_cg_closure_phase_2 g1 = (hasChanged2, g2)"
using prod.exhaust by metis
have "length g1 \<le> length g"
using \<open>simple_cg_closure_phase_1 g = (hasChanged1, g1)\<close>
using simple_cg_closure_phase_1_length[of g]
by auto
have "length g2 \<le> length g1"
using \<open>simple_cg_closure_phase_2 g1 = (hasChanged2, g2)\<close>
using simple_cg_closure_phase_2_length[of g1]
by auto
have "(\<And>x2 u v. x2 \<in> list.set g2 \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
proof -
have "(\<And>x2 u v. x2 \<in> list.set g1 \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using simple_cg_closure_phase_1_validity[OF assms(1,2), of g]
using less.prems(1)
unfolding \<open>simple_cg_closure_phase_1 g = (hasChanged1, g1)\<close> snd_conv
by blast
then show "(\<And>x2 u v. x2 \<in> list.set g2 \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using simple_cg_closure_phase_2_validity[of g1]
unfolding \<open>simple_cg_closure_phase_2 g1 = (hasChanged2, g2)\<close> snd_conv
by blast
qed
show ?thesis proof (cases "hasChanged1 \<or> hasChanged2")
case True
then consider hasChanged1 | hasChanged2
by blast
then have "length g2 < length g"
proof cases
case 1
then have "length g1 < length g"
using \<open>simple_cg_closure_phase_1 g = (hasChanged1, g1)\<close>
using simple_cg_closure_phase_1_True[of g]
by auto
then show ?thesis
using \<open>length g2 \<le> length g1\<close>
by linarith
next
case 2
then have "length g2 < length g1"
using \<open>simple_cg_closure_phase_2 g1 = (hasChanged2, g2)\<close>
using simple_cg_closure_phase_2_True[of g1]
by auto
then show ?thesis
using \<open>length g1 \<le> length g\<close>
by linarith
qed
moreover have "x2 \<in> list.set (simple_cg_closure g2)"
using less.prems(2)
using \<open>simple_cg_closure_phase_1 g = (hasChanged1, g1)\<close> \<open>simple_cg_closure_phase_2 g1 = (hasChanged2, g2)\<close> True
by auto
moreover note \<open>(\<And>x2 u v. x2 \<in> list.set g2 \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)\<close>
ultimately show ?thesis
using less.hyps[of g2]
by blast
next
case False
then have "(simple_cg_closure g) = g2"
using \<open>simple_cg_closure_phase_1 g = (hasChanged1, g1)\<close> \<open>simple_cg_closure_phase_2 g1 = (hasChanged2, g2)\<close>
by auto
show ?thesis
using less.prems(2)
using \<open>(\<And>x2 u v. x2 \<in> list.set g2 \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)\<close> assms(5,6,7,8)
unfolding \<open>(simple_cg_closure g) = g2\<close>
by blast
qed
qed
(* when inserting \<alpha> this also for all \<alpha>1@\<alpha>2 = \<alpha> and \<beta> in [\<alpha>1] inserts \<beta>@\<alpha>2 -- extremely inefficient *)
fun simple_cg_insert_with_conv :: "('a::linorder) simple_cg \<Rightarrow> 'a list \<Rightarrow> 'a simple_cg" where
"simple_cg_insert_with_conv g ys = (let
insert_for_prefix = (\<lambda> g i . let
pref = take i ys;
suff = drop i ys;
pref_conv = simple_cg_lookup g pref
in foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) g pref_conv);
g' = simple_cg_insert g ys;
g'' = foldl insert_for_prefix g' [0..<length ys]
in simple_cg_closure g'')"
fun simple_cg_merge :: "'a simple_cg \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a simple_cg" where
"simple_cg_merge g ys1 ys2 = simple_cg_closure ({|ys1,ys2|}#g)"
lemma simple_cg_merge_validity :
assumes "observable M1" and "observable M2"
and "converge M1 u' v' \<and> converge M2 u' v'"
and "\<And> x2 u v . x2 \<in> list.set g \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
and "x2 \<in> list.set (simple_cg_merge g u' v')"
and "u |\<in>| x2"
and "v |\<in>| x2"
and "u \<in> L M1"
and "u \<in> L M2"
shows "converge M1 u v \<and> converge M2 u v"
proof -
have "(\<And>x2 u v. x2 \<in> list.set ({|u',v'|}#g) \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
proof -
fix x2 u v assume "x2 \<in> list.set ({|u',v'|}#g)" and "u |\<in>| x2" and "v |\<in>| x2" and "u \<in> L M1" and "u \<in> L M2"
then consider "x2 = {|u',v'|}" | "x2 \<in> list.set g"
by auto
then show "converge M1 u v \<and> converge M2 u v" proof cases
case 1
then have "u \<in> {u',v'}" and "v \<in> {u',v'}"
using \<open>u |\<in>| x2\<close> \<open>v |\<in>| x2\<close> by auto
then show ?thesis
using assms(3)
by (cases "u = u'"; cases "v = v'"; auto)
next
case 2
then show ?thesis
using assms(4) \<open>u |\<in>| x2\<close> \<open>v |\<in>| x2\<close> \<open>u \<in> L M1\<close> \<open>u \<in> L M2\<close>
by blast
qed
qed
moreover have "x2 \<in> list.set (simple_cg_closure ({|u',v'|}#g))"
using assms(5) by auto
ultimately show ?thesis
using simple_cg_closure_validity[OF assms(1,2) _ _ assms(6,7,8,9)]
by blast
qed
subsection \<open>Invariants\<close>
lemma simple_cg_lookup_iff :
"\<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<longleftrightarrow> (\<beta> = \<alpha> \<or> (\<exists> x . x \<in> list.set G \<and> \<alpha> |\<in>| x \<and> \<beta> |\<in>| x))"
proof (induction G rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc x G)
show ?case proof (cases "\<alpha> |\<in>| x \<and> \<beta> |\<in>| x")
case True
then have "\<beta> \<in> list.set (simple_cg_lookup (G@[x]) \<alpha>)"
unfolding simple_cg_lookup.simps
unfolding sorted_list_of_set_set
- using notin_fset by force
+ by simp
then show ?thesis
using True by auto
next
case False
have "\<beta> \<in> list.set (simple_cg_lookup (G@[x]) \<alpha>) = (\<beta> = \<alpha> \<or> (\<beta> \<in> list.set (simple_cg_lookup G \<alpha>)))"
proof -
consider "\<alpha> |\<notin>| x" | "\<beta> |\<notin>| x"
using False by blast
then show "\<beta> \<in> list.set (simple_cg_lookup (G@[x]) \<alpha>) = (\<beta> = \<alpha> \<or> (\<beta> \<in> list.set (simple_cg_lookup G \<alpha>)))"
proof cases
case 1
then show ?thesis
unfolding simple_cg_lookup.simps
unfolding sorted_list_of_set_set
by auto
next
case 2
then have "\<beta> \<notin> list.set (sorted_list_of_fset x)"
- using fmember_iff_member_fset by fastforce
+ by simp
then have "(\<beta> \<in> list.set (simple_cg_lookup (G@[x]) \<alpha>)) = (\<beta> \<in> Set.insert \<alpha> (list.set (simple_cg_lookup G \<alpha>)))"
unfolding simple_cg_lookup.simps
unfolding sorted_list_of_set_set
by auto
then show ?thesis
by (induction G; auto)
qed
qed
moreover have "(\<exists> x' . x' \<in> list.set (G@[x]) \<and> \<alpha> |\<in>| x' \<and> \<beta> |\<in>| x') = (\<exists> x . x \<in> list.set G \<and> \<alpha> |\<in>| x \<and> \<beta> |\<in>| x)"
using False by auto
ultimately show ?thesis
using snoc.IH
by blast
qed
qed
lemma simple_cg_insert'_invar :
"convergence_graph_insert_invar M1 M2 simple_cg_lookup simple_cg_insert'"
proof -
have "\<And> G \<gamma> \<alpha> \<beta> . \<gamma> \<in> L M1 \<Longrightarrow>
\<gamma> \<in> L M2 \<Longrightarrow>
(\<And>\<alpha> . \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup G \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)) \<Longrightarrow>
\<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup (simple_cg_insert' G \<gamma>) \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert' G \<gamma>) \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof
fix G \<gamma> \<alpha>
assume "\<gamma> \<in> L M1"
and "\<gamma> \<in> L M2"
and *:"(\<And>\<alpha> . \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup G \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>))"
and "\<alpha> \<in> L M1"
and "\<alpha> \<in> L M2"
show "\<alpha> \<in> list.set (simple_cg_lookup (simple_cg_insert' G \<gamma>) \<alpha>)"
unfolding simple_cg_lookup.simps
unfolding sorted_list_of_set_set
by auto
have "\<And> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert' G \<gamma>) \<alpha>) \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>"
proof -
fix \<beta>
assume **: "\<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert' G \<gamma>) \<alpha>)"
show "converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>"
proof (cases "\<beta> \<in> list.set (simple_cg_lookup G \<alpha>)")
case True
then show ?thesis
using *[OF \<open>\<alpha> \<in> L M1\<close> \<open>\<alpha> \<in> L M2\<close>]
by presburger
next
case False
show ?thesis proof (cases "find ((|\<in>|) \<gamma>) G")
case None
then have "(simple_cg_insert' G \<gamma>) = {|\<gamma>|}#G"
by auto
have "\<alpha> = \<gamma> \<and> \<beta> = \<gamma>"
using False \<open>\<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert' G \<gamma>) \<alpha>)\<close>
unfolding \<open>(simple_cg_insert' G \<gamma>) = {|\<gamma>|}#G\<close>
by (metis fsingleton_iff set_ConsD simple_cg_lookup_iff)
then show ?thesis
using \<open>\<gamma> \<in> L M1\<close> \<open>\<gamma> \<in> L M2\<close> by auto
next
case (Some x)
then have "(simple_cg_insert' G \<gamma>) = G"
by auto
then show ?thesis
using *[OF \<open>\<alpha> \<in> L M1\<close> \<open>\<alpha> \<in> L M2\<close>] **
by presburger
qed
qed
qed
then show "(\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert' G \<gamma>) \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
by blast
qed
then show ?thesis
unfolding convergence_graph_insert_invar_def convergence_graph_lookup_invar_def
by blast
qed
lemma simple_cg_insert'_foldl_helper:
assumes "list.set xss \<subseteq> L M1 \<inter> L M2"
and "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
shows "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (foldl (\<lambda> xs' ys' . simple_cg_insert' xs' ys') G xss) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using \<open>list.set xss \<subseteq> L M1 \<inter> L M2\<close>
proof (induction xss rule: rev_induct)
case Nil
then show ?case
using \<open>(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)\<close>
by auto
next
case (snoc x xs)
have "x \<in> L M1" and "x \<in> L M2"
using snoc.prems by auto
have "list.set xs \<subseteq> L M1 \<inter> L M2"
using snoc.prems by auto
then have *:"(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (foldl (\<lambda> xs' ys'. simple_cg_insert' xs' ys') G xs) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using snoc.IH
by blast
have **:"(foldl (\<lambda> xs' ys'. simple_cg_insert' xs' ys') G (xs@[x])) = simple_cg_insert' (foldl (\<lambda> xs' ys' . simple_cg_insert' xs' ys') G xs) x"
by auto
show ?case
using snoc.prems(1,2,3) * \<open>x \<in> L M1\<close> \<open>x \<in> L M2\<close>
unfolding **
using simple_cg_insert'_invar[of M1 M2]
unfolding convergence_graph_insert_invar_def convergence_graph_lookup_invar_def
using simple_cg_lookup_iff
by blast
qed
lemma simple_cg_insert_invar :
"convergence_graph_insert_invar M1 M2 simple_cg_lookup simple_cg_insert"
proof -
have "\<And> G \<gamma> \<alpha> \<beta> . \<gamma> \<in> L M1 \<Longrightarrow>
\<gamma> \<in> L M2 \<Longrightarrow>
(\<And>\<alpha> . \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup G \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)) \<Longrightarrow>
\<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup (simple_cg_insert G \<gamma>) \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert G \<gamma>) \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof
fix G \<gamma> \<alpha>
assume "\<gamma> \<in> L M1"
and "\<gamma> \<in> L M2"
and *:"(\<And>\<alpha> . \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup G \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>))"
and "\<alpha> \<in> L M1"
and "\<alpha> \<in> L M2"
show "\<alpha> \<in> list.set (simple_cg_lookup (simple_cg_insert G \<gamma>) \<alpha>)"
unfolding simple_cg_lookup.simps
unfolding sorted_list_of_set_set
by auto
note simple_cg_insert'_foldl_helper[of "prefixes \<gamma>" M1 M2]
moreover have "list.set (prefixes \<gamma>) \<subseteq> L M1 \<inter> L M2"
by (metis (no_types, lifting) IntI \<open>\<gamma> \<in> L M1\<close> \<open>\<gamma> \<in> L M2\<close> language_prefix prefixes_set_ob subsetI)
ultimately show "(\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert G \<gamma>) \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using \<open>\<alpha> \<in> L M1\<close> \<open>\<alpha> \<in> L M2\<close>
by (metis "*" simple_cg_insert.simps)
qed
then show ?thesis
unfolding convergence_graph_insert_invar_def convergence_graph_lookup_invar_def
by blast
qed
lemma simple_cg_closure_invar_helper :
assumes "observable M1" and "observable M2"
and "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
and "\<beta> \<in> list.set (simple_cg_lookup (simple_cg_closure G) \<alpha>)"
and "\<alpha> \<in> L M1" and "\<alpha> \<in> L M2"
shows "converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>"
proof (cases "\<beta> = \<alpha>")
case True
then show ?thesis using assms(5,6) by auto
next
case False
show ?thesis
proof
obtain x where "x \<in> list.set (simple_cg_closure G)" and "\<alpha> |\<in>| x" and "\<beta> |\<in>| x"
using False \<open>\<beta> \<in> list.set (simple_cg_lookup (simple_cg_closure G) \<alpha>)\<close> unfolding simple_cg_lookup_iff
by blast
have "\<And> x2 u v . x2 \<in> list.set G \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v"
using \<open>(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)\<close>
unfolding simple_cg_lookup_iff
by blast
have "(\<And>x2 u v. x2 \<in> list.set G \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using \<open>(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)\<close>
unfolding simple_cg_lookup_iff by blast
then show "converge M1 \<alpha> \<beta>"
using \<open>\<alpha> |\<in>| x\<close> \<open>\<beta> |\<in>| x\<close> \<open>x \<in> list.set (simple_cg_closure G)\<close> assms(1) assms(2) assms(5) assms(6) simple_cg_closure_validity by blast
have "(\<And>x2 u v. x2 \<in> list.set G \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using \<open>(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)\<close>
unfolding simple_cg_lookup_iff by blast
then show "converge M2 \<alpha> \<beta>"
using \<open>\<alpha> |\<in>| x\<close> \<open>\<beta> |\<in>| x\<close> \<open>x \<in> list.set (simple_cg_closure G)\<close> assms(1) assms(2) assms(5) assms(6) simple_cg_closure_validity by blast
qed
qed
lemma simple_cg_merge_invar :
assumes "observable M1" and "observable M2"
shows "convergence_graph_merge_invar M1 M2 simple_cg_lookup simple_cg_merge"
proof -
have "\<And> G \<gamma> \<gamma>' \<alpha> \<beta>.
converge M1 \<gamma> \<gamma>' \<Longrightarrow>
converge M2 \<gamma> \<gamma>' \<Longrightarrow>
(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>) \<Longrightarrow>
\<beta> \<in> list.set (simple_cg_lookup (simple_cg_merge G \<gamma> \<gamma>') \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>"
proof -
fix G \<gamma> \<gamma>' \<alpha> \<beta>
assume "converge M1 \<gamma> \<gamma>'"
"converge M2 \<gamma> \<gamma>'"
"(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
"\<beta> \<in> list.set (simple_cg_lookup (simple_cg_merge G \<gamma> \<gamma>') \<alpha>)"
"\<alpha> \<in> L M1"
"\<alpha> \<in> L M2"
show "converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>"
proof (cases "\<beta> = \<alpha>")
case True
then show ?thesis using \<open>\<alpha> \<in> L M1\<close> \<open>\<alpha> \<in> L M2\<close> by auto
next
case False
then obtain x where "x \<in> list.set (simple_cg_merge G \<gamma> \<gamma>')" and "\<alpha> |\<in>| x" and "\<beta> |\<in>| x"
using \<open>\<beta> \<in> list.set (simple_cg_lookup (simple_cg_merge G \<gamma> \<gamma>') \<alpha>)\<close> unfolding simple_cg_lookup_iff
by blast
have "(\<And>x2 u v. x2 \<in> list.set G \<Longrightarrow> u |\<in>| x2 \<Longrightarrow> v |\<in>| x2 \<Longrightarrow> u \<in> L M1 \<Longrightarrow> u \<in> L M2 \<Longrightarrow> converge M1 u v \<and> converge M2 u v)"
using \<open>(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)\<close>
unfolding simple_cg_lookup_iff by blast
then show ?thesis
using simple_cg_merge_validity[OF assms(1,2) _ _ \<open>x \<in> list.set (simple_cg_merge G \<gamma> \<gamma>')\<close> \<open>\<alpha> |\<in>| x\<close> \<open>\<beta> |\<in>| x\<close> \<open>\<alpha> \<in> L M1\<close> \<open>\<alpha> \<in> L M2\<close>]
\<open>converge M1 \<gamma> \<gamma>'\<close> \<open>converge M2 \<gamma> \<gamma>'\<close>
by blast
qed
qed
then show ?thesis
unfolding convergence_graph_merge_invar_def convergence_graph_lookup_invar_def
unfolding simple_cg_lookup_iff
by metis
qed
lemma simple_cg_empty_invar :
"convergence_graph_lookup_invar M1 M2 simple_cg_lookup simple_cg_empty"
unfolding convergence_graph_lookup_invar_def simple_cg_empty_def
by auto
lemma simple_cg_initial_invar :
assumes "observable M1"
shows "convergence_graph_initial_invar M1 M2 simple_cg_lookup simple_cg_initial"
proof -
have "\<And> T . (L M1 \<inter> set T = (L M2 \<inter> set T)) \<Longrightarrow> finite_tree T \<Longrightarrow> (\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (simple_cg_initial M1 T) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof -
fix T assume "(L M1 \<inter> set T = (L M2 \<inter> set T))" and "finite_tree T"
then have "list.set (filter (is_in_language M1 (initial M1)) (sorted_list_of_sequences_in_tree T)) \<subseteq> L M1 \<inter> L M2"
unfolding is_in_language_iff[OF assms fsm_initial]
using sorted_list_of_sequences_in_tree_set[OF \<open>finite_tree T\<close>]
by auto
moreover have "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup simple_cg_empty \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using simple_cg_empty_invar
unfolding convergence_graph_lookup_invar_def
by blast
ultimately show "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (simple_cg_initial M1 T) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using simple_cg_insert'_foldl_helper[of "(filter (is_in_language M1 (initial M1)) (sorted_list_of_sequences_in_tree T))" M1 M2]
unfolding simple_cg_initial.simps
by blast
qed
then show ?thesis
unfolding convergence_graph_initial_invar_def convergence_graph_lookup_invar_def
using simple_cg_lookup_iff by blast
qed
lemma simple_cg_insert_with_conv_invar :
assumes "observable M1"
assumes "observable M2"
shows "convergence_graph_insert_invar M1 M2 simple_cg_lookup simple_cg_insert_with_conv"
proof -
have "\<And> G \<gamma> \<alpha> \<beta> . \<gamma> \<in> L M1 \<Longrightarrow>
\<gamma> \<in> L M2 \<Longrightarrow>
(\<And>\<alpha> . \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup G \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)) \<Longrightarrow>
\<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup (simple_cg_insert_with_conv G \<gamma>) \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert_with_conv G \<gamma>) \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof
fix G ys \<alpha>
assume "ys \<in> L M1"
and "ys \<in> L M2"
and *:"(\<And>\<alpha> . \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> \<alpha> \<in> list.set (simple_cg_lookup G \<alpha>) \<and> (\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>))"
and "\<alpha> \<in> L M1"
and "\<alpha> \<in> L M2"
show "\<alpha> \<in> list.set (simple_cg_lookup (simple_cg_insert_with_conv G ys) \<alpha>)"
using simple_cg_lookup_iff by blast
have "\<And> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert_with_conv G ys) \<alpha>) \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>"
proof -
fix \<beta>
assume "\<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert_with_conv G ys) \<alpha>)"
define insert_for_prefix where insert_for_prefix:
"insert_for_prefix = (\<lambda> g i . let
pref = take i ys;
suff = drop i ys;
pref_conv = simple_cg_lookup g pref
in foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) g pref_conv)"
define g' where g': "g' = simple_cg_insert G ys"
define g'' where g'': "g'' = foldl insert_for_prefix g' [0..<length ys]"
have "simple_cg_insert_with_conv G ys = simple_cg_closure g''"
unfolding simple_cg_insert_with_conv.simps g'' g' insert_for_prefix Let_def by force
have g'_invar: "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup g' \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using g' *
using simple_cg_insert_invar \<open>ys \<in> L M1\<close> \<open>ys \<in> L M2\<close>
unfolding convergence_graph_insert_invar_def convergence_graph_lookup_invar_def
by blast
have insert_for_prefix_invar: "\<And> i g . (\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup g \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>) \<Longrightarrow> (\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (insert_for_prefix g i) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof -
fix i g assume "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup g \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
define pref where pref: "pref = take i ys"
define suff where suff: "suff = drop i ys"
let ?pref_conv = "simple_cg_lookup g pref"
have "insert_for_prefix g i = foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) g ?pref_conv"
unfolding insert_for_prefix pref suff Let_def by force
have "ys = pref @ suff"
unfolding pref suff by auto
then have "pref \<in> L M1" and "pref \<in> L M2"
using \<open>ys \<in> L M1\<close> \<open>ys \<in> L M2\<close> language_prefix by metis+
have insert_step_invar: "\<And> ys' pc G . list.set pc \<subseteq> list.set (simple_cg_lookup g pref) \<Longrightarrow> ys' \<in> list.set pc \<Longrightarrow>
(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>) \<Longrightarrow>
(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert' G (ys'@suff)) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof -
fix ys' pc G
assume "list.set pc \<subseteq> list.set (simple_cg_lookup g pref)"
and "ys' \<in> list.set pc"
and "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
then have "converge M1 pref ys'" and "converge M2 pref ys'"
using \<open>\<And>\<beta> \<alpha>. \<beta> \<in> list.set (simple_cg_lookup g \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>\<close>
using \<open>pref \<in> L M1\<close> \<open>pref \<in> L M2\<close>
by blast+
have "(ys'@suff) \<in> L M1"
using \<open>converge M1 pref ys'\<close>
using \<open>ys = pref @ suff\<close> \<open>ys \<in> L M1\<close> assms(1) converge_append_language_iff by blast
moreover have "(ys'@suff) \<in> L M2"
using \<open>converge M2 pref ys'\<close>
using \<open>ys = pref @ suff\<close> \<open>ys \<in> L M2\<close> assms(2) converge_append_language_iff by blast
ultimately show "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert' G (ys'@suff)) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using \<open>(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)\<close>
using simple_cg_insert'_invar[of M1 M2]
unfolding convergence_graph_insert_invar_def convergence_graph_lookup_invar_def
using simple_cg_lookup_iff by blast
qed
have insert_foldl_invar: "\<And> pc G . list.set pc \<subseteq> list.set (simple_cg_lookup g pref) \<Longrightarrow>
(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>) \<Longrightarrow>
(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) G pc) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof -
fix pc G assume "list.set pc \<subseteq> list.set (simple_cg_lookup g pref)"
and "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
then show "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) G pc) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof (induction pc rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc a pc)
have **:"(foldl (\<lambda>g' ys'. simple_cg_insert' g' (ys' @ suff)) G (pc @ [a]))
= simple_cg_insert' (foldl (\<lambda>g' ys'. simple_cg_insert' g' (ys' @ suff)) G pc) (a@suff)"
unfolding foldl_append by auto
have "list.set pc \<subseteq> list.set (simple_cg_lookup g pref)"
using snoc.prems(4) by auto
then have *: "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) G pc) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using snoc.IH
using snoc.prems(5) by blast
have "a \<in> list.set (pc @ [a])" by auto
then show ?case
using snoc.prems(1,2,3)
unfolding **
using insert_step_invar[OF snoc.prems(4), of a "(foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) G pc)", OF _ *]
by blast
qed
qed
show "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (insert_for_prefix g i) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
using insert_foldl_invar[of ?pref_conv g, OF _ \<open>(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup g \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)\<close>]
unfolding \<open>insert_for_prefix g i = foldl (\<lambda> g' ys' . simple_cg_insert' g' (ys'@suff)) g ?pref_conv\<close>
by blast
qed
have insert_for_prefix_foldl_invar: "\<And> ns . (\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (foldl insert_for_prefix g' ns) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof -
fix ns show "(\<And>\<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup (foldl insert_for_prefix g' ns) \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof (induction ns rule: rev_induct)
case Nil
then show ?case using g'_invar by auto
next
case (snoc a ns)
show ?case
using snoc.prems
using insert_for_prefix_invar [OF snoc.IH]
by auto
qed
qed
show \<open>converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>\<close>
using \<open>\<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert_with_conv G ys) \<alpha>)\<close>
unfolding \<open>simple_cg_insert_with_conv G ys = simple_cg_closure g''\<close> g''
using insert_for_prefix_foldl_invar[of _ "[0..<length ys]" _]
using simple_cg_closure_invar_helper[OF assms, of "(foldl insert_for_prefix g' [0..<length ys])", OF insert_for_prefix_foldl_invar[of _ "[0..<length ys]" _]]
using \<open>\<alpha> \<in> L M1\<close> \<open>\<alpha> \<in> L M2\<close> by blast
qed
then show "(\<forall> \<beta> . \<beta> \<in> list.set (simple_cg_lookup (simple_cg_insert_with_conv G ys) \<alpha>) \<longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
by blast
qed
then show ?thesis
unfolding convergence_graph_insert_invar_def convergence_graph_lookup_invar_def
by blast
qed
lemma simple_cg_lookup_with_conv_from_lookup_invar:
assumes "observable M1" and "observable M2"
and "convergence_graph_lookup_invar M1 M2 simple_cg_lookup G"
shows "convergence_graph_lookup_invar M1 M2 simple_cg_lookup_with_conv G"
proof -
have "(\<And> \<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup_with_conv G \<alpha>) \<Longrightarrow> \<alpha> \<in> L M1 \<Longrightarrow> \<alpha> \<in> L M2 \<Longrightarrow> converge M1 \<alpha> \<beta> \<and> converge M2 \<alpha> \<beta>)"
proof -
fix ys \<beta> assume "\<beta> \<in> list.set (simple_cg_lookup_with_conv G ys)" and "ys \<in> L M1" and "ys \<in> L M2"
define lookup_for_prefix where lookup_for_prefix:
"lookup_for_prefix = (\<lambda>i . let
pref = take i ys;
suff = drop i ys;
pref_conv = (foldl (|\<union>|) fempty (filter (\<lambda>x . pref |\<in>| x) G))
in fimage (\<lambda> pref' . pref'@suff) pref_conv)"
have "\<And> ns . \<beta> \<in> list.set (sorted_list_of_fset (finsert ys (foldl (\<lambda> cs i . lookup_for_prefix i |\<union>| cs) fempty ns))) \<Longrightarrow> converge M1 ys \<beta> \<and> converge M2 ys \<beta>"
proof -
fix ns assume "\<beta> \<in> list.set (sorted_list_of_fset (finsert ys (foldl (\<lambda> cs i . lookup_for_prefix i |\<union>| cs) fempty ns)))"
then show "converge M1 ys \<beta> \<and> converge M2 ys \<beta>"
proof (induction ns rule: rev_induct)
case Nil
then show ?case using \<open>ys \<in> L M1\<close> \<open>ys \<in> L M2\<close> by auto
next
case (snoc a ns)
have "list.set (sorted_list_of_fset (finsert ys (foldl (\<lambda> cs i . lookup_for_prefix i |\<union>| cs) fempty (ns@[a])))) =
(fset (lookup_for_prefix a) \<union> list.set (sorted_list_of_fset (finsert ys (foldl (\<lambda> cs i . lookup_for_prefix i |\<union>| cs) fempty ns))))"
by auto
then consider "\<beta> \<in> fset (lookup_for_prefix a)" | "\<beta> \<in> list.set (sorted_list_of_fset (finsert ys (foldl (\<lambda> cs i . lookup_for_prefix i |\<union>| cs) fempty ns)))"
using snoc.prems by auto
then show ?case proof cases
case 1
define pref where pref: "pref = take a ys"
define suff where suff: "suff = drop a ys"
define pref_conv where pref_conv: "pref_conv = (foldl (|\<union>|) fempty (filter (\<lambda>x . pref |\<in>| x) G))"
have "lookup_for_prefix a = fimage (\<lambda> pref' . pref'@suff) pref_conv"
unfolding lookup_for_prefix pref suff pref_conv
by metis
then have "\<beta> \<in> list.set (map (\<lambda> pref' . pref'@suff) (sorted_list_of_fset (finsert pref (foldl (|\<union>|) {||} (filter ((|\<in>|) pref) G)))))"
using 1 unfolding pref_conv by auto
then obtain \<gamma> where "\<gamma> \<in> list.set (simple_cg_lookup G pref)"
and "\<beta> = \<gamma>@suff"
unfolding simple_cg_lookup.simps
by (meson set_map_elem)
then have "converge M1 \<gamma> pref" and "converge M2 \<gamma> pref"
using \<open>convergence_graph_lookup_invar M1 M2 simple_cg_lookup G\<close>
unfolding convergence_graph_lookup_invar_def
by (metis \<open>ys \<in> L M1\<close> \<open>ys \<in> L M2\<close> append_take_drop_id converge_sym language_prefix pref)+
then show ?thesis
by (metis \<open>\<And>thesis. (\<And>\<gamma>. \<lbrakk>\<gamma> \<in> list.set (simple_cg_lookup G pref); \<beta> = \<gamma> @ suff\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> \<open>ys \<in> L M1\<close> \<open>ys \<in> L M2\<close> append_take_drop_id assms(1) assms(2) assms(3) converge_append converge_append_language_iff convergence_graph_lookup_invar_def language_prefix pref suff)
next
case 2
then show ?thesis using snoc.IH by blast
qed
qed
qed
then show "converge M1 ys \<beta> \<and> converge M2 ys \<beta>"
using \<open>\<beta> \<in> list.set (simple_cg_lookup_with_conv G ys)\<close>
unfolding simple_cg_lookup_with_conv.simps Let_def lookup_for_prefix sorted_list_of_set_set
by blast
qed
moreover have "\<And> \<alpha> . \<alpha> \<in> list.set (simple_cg_lookup_with_conv G \<alpha>)"
unfolding simple_cg_lookup_with_conv.simps by auto
ultimately show ?thesis
unfolding convergence_graph_lookup_invar_def
by blast
qed
lemma simple_cg_lookup_from_lookup_invar_with_conv:
assumes "convergence_graph_lookup_invar M1 M2 simple_cg_lookup_with_conv G"
shows "convergence_graph_lookup_invar M1 M2 simple_cg_lookup G"
proof -
have "\<And> \<alpha> \<beta>. \<beta> \<in> list.set (simple_cg_lookup G \<alpha>) \<Longrightarrow> \<beta> \<in> list.set (simple_cg_lookup_with_conv G \<alpha>)"
proof -
fix \<alpha> \<beta> assume "\<beta> \<in> list.set (simple_cg_lookup G \<alpha>)"
define lookup_for_prefix where lookup_for_prefix:
"lookup_for_prefix = (\<lambda>i . let
pref = take i \<alpha>;
suff = drop i \<alpha>;
pref_conv = simple_cg_lookup G pref
in map (\<lambda> pref' . pref'@suff) pref_conv)"
have "lookup_for_prefix (length \<alpha>) = simple_cg_lookup G \<alpha>"
unfolding lookup_for_prefix by auto
moreover have "list.set (lookup_for_prefix (length \<alpha>)) \<subseteq> list.set (simple_cg_lookup_with_conv G \<alpha>)"
unfolding simple_cg_lookup_with_conv.simps lookup_for_prefix Let_def sorted_list_of_set_set by auto
ultimately show "\<beta> \<in> list.set (simple_cg_lookup_with_conv G \<alpha>)"
using \<open>\<beta> \<in> list.set (simple_cg_lookup G \<alpha>)\<close>
by (metis subsetD)
qed
then show ?thesis
using assms
unfolding convergence_graph_lookup_invar_def
using simple_cg_lookup_iff by blast
qed
lemma simple_cg_lookup_invar_with_conv_eq :
assumes "observable M1" and "observable M2"
shows "convergence_graph_lookup_invar M1 M2 simple_cg_lookup_with_conv G = convergence_graph_lookup_invar M1 M2 simple_cg_lookup G"
using simple_cg_lookup_with_conv_from_lookup_invar[OF assms] simple_cg_lookup_from_lookup_invar_with_conv[of M1 M2]
by blast
lemma simple_cg_insert_invar_with_conv :
assumes "observable M1" and "observable M2"
shows "convergence_graph_insert_invar M1 M2 simple_cg_lookup_with_conv simple_cg_insert"
using simple_cg_insert_invar[of M1 M2]
unfolding convergence_graph_insert_invar_def
unfolding simple_cg_lookup_invar_with_conv_eq[OF assms]
.
lemma simple_cg_merge_invar_with_conv :
assumes "observable M1" and "observable M2"
shows "convergence_graph_merge_invar M1 M2 simple_cg_lookup_with_conv simple_cg_merge"
using simple_cg_merge_invar[OF assms]
unfolding convergence_graph_merge_invar_def
unfolding simple_cg_lookup_invar_with_conv_eq[OF assms]
.
lemma simple_cg_initial_invar_with_conv :
assumes "observable M1" and "observable M2"
shows "convergence_graph_initial_invar M1 M2 simple_cg_lookup_with_conv simple_cg_initial"
using simple_cg_initial_invar[OF assms(1), of M2]
unfolding convergence_graph_initial_invar_def
unfolding simple_cg_lookup_invar_with_conv_eq[OF assms]
.
end
diff --git a/thys/FSM_Tests/FSM.thy b/thys/FSM_Tests/FSM.thy
--- a/thys/FSM_Tests/FSM.thy
+++ b/thys/FSM_Tests/FSM.thy
@@ -1,6428 +1,6428 @@
section \<open>Finite State Machines\<close>
text \<open>This theory defines well-formed finite state machines and introduces various closely related
notions, as well as a selection of basic properties and definitions.\<close>
theory FSM
imports FSM_Impl "HOL-Library.Quotient_Type" "HOL-Library.Product_Lexorder"
begin
subsection \<open>Well-formed Finite State Machines\<close>
text \<open>A value of type @{text "fsm_impl"} constitutes a well-formed FSM if its contained sets are
finite and the initial state and the components of each transition are contained in their
respective sets.\<close>
abbreviation(input) "well_formed_fsm (M :: ('state, 'input, 'output) fsm_impl)
\<equiv> (initial M \<in> states M
\<and> finite (states M)
\<and> finite (inputs M)
\<and> finite (outputs M)
\<and> finite (transitions M)
\<and> (\<forall> t \<in> transitions M . t_source t \<in> states M \<and>
t_input t \<in> inputs M \<and>
t_target t \<in> states M \<and>
t_output t \<in> outputs M)) "
typedef ('state, 'input, 'output) fsm =
"{ M :: ('state, 'input, 'output) fsm_impl . well_formed_fsm M}"
morphisms fsm_impl_of_fsm Abs_fsm
proof -
obtain q :: 'state where "True" by blast
have "well_formed_fsm (FSMI q {q} {} {} {})" by auto
then show ?thesis by blast
qed
setup_lifting type_definition_fsm
lift_definition initial :: "('state, 'input, 'output) fsm \<Rightarrow> 'state" is FSM_Impl.initial done
lift_definition states :: "('state, 'input, 'output) fsm \<Rightarrow> 'state set" is FSM_Impl.states done
lift_definition inputs :: "('state, 'input, 'output) fsm \<Rightarrow> 'input set" is FSM_Impl.inputs done
lift_definition outputs :: "('state, 'input, 'output) fsm \<Rightarrow> 'output set" is FSM_Impl.outputs done
lift_definition transitions ::
"('state, 'input, 'output) fsm \<Rightarrow> ('state \<times> 'input \<times> 'output \<times> 'state) set"
is FSM_Impl.transitions done
lift_definition fsm_from_list :: "'a \<Rightarrow> ('a,'b,'c) transition list \<Rightarrow> ('a, 'b, 'c) fsm"
is FSM_Impl.fsm_impl_from_list
proof -
fix q :: 'a
fix ts :: "('a,'b,'c) transition list"
show "well_formed_fsm (fsm_impl_from_list q ts)"
by (induction ts; auto)
qed
lemma fsm_initial[intro]: "initial M \<in> states M"
by (transfer; blast)
lemma fsm_states_finite: "finite (states M)"
by (transfer; blast)
lemma fsm_inputs_finite: "finite (inputs M)"
by (transfer; blast)
lemma fsm_outputs_finite: "finite (outputs M)"
by (transfer; blast)
lemma fsm_transitions_finite: "finite (transitions M)"
by (transfer; blast)
lemma fsm_transition_source[intro]: "\<And> t . t \<in> (transitions M) \<Longrightarrow> t_source t \<in> states M"
by (transfer; blast)
lemma fsm_transition_target[intro]: "\<And> t . t \<in> (transitions M) \<Longrightarrow> t_target t \<in> states M"
by (transfer; blast)
lemma fsm_transition_input[intro]: "\<And> t . t \<in> (transitions M) \<Longrightarrow> t_input t \<in> inputs M"
by (transfer; blast)
lemma fsm_transition_output[intro]: "\<And> t . t \<in> (transitions M) \<Longrightarrow> t_output t \<in> outputs M"
by (transfer; blast)
instantiation fsm :: (type,type,type) equal
begin
definition equal_fsm :: "('a, 'b, 'c) fsm \<Rightarrow> ('a, 'b, 'c) fsm \<Rightarrow> bool" where
"equal_fsm x y = (initial x = initial y \<and> states x = states y \<and> inputs x = inputs y \<and> outputs x = outputs y \<and> transitions x = transitions y)"
instance
apply (intro_classes)
unfolding equal_fsm_def
apply transfer
using fsm_impl.expand by auto
end
subsubsection \<open>Example FSMs\<close>
definition m_ex_H :: "(integer,integer,integer) fsm" where
"m_ex_H = fsm_from_list 1 [ (1,0,0,2),
(1,0,1,4),
(1,1,1,4),
(2,0,0,2),
(2,1,1,4),
(3,0,1,4),
(3,1,0,1),
(3,1,1,3),
(4,0,0,3),
(4,1,0,1)]"
definition m_ex_9 :: "(integer,integer,integer) fsm" where
"m_ex_9 = fsm_from_list 0 [ (0,0,2,2),
(0,0,3,2),
(0,1,0,3),
(0,1,1,3),
(1,0,3,2),
(1,1,1,3),
(2,0,2,2),
(2,1,3,3),
(3,0,2,2),
(3,1,0,2),
(3,1,1,1)]"
definition m_ex_DR :: "(integer,integer,integer) fsm" where
"m_ex_DR = fsm_from_list 0 [(0,0,0,100),
(100,0,0,101),
(100,0,1,101),
(101,0,0,102),
(101,0,1,102),
(102,0,0,103),
(102,0,1,103),
(103,0,0,104),
(103,0,1,104),
(104,0,0,100),
(104,0,1,100),
(104,1,0,400),
(0,0,2,200),
(200,0,2,201),
(201,0,2,202),
(202,0,2,203),
(203,0,2,200),
(203,1,0,400),
(0,1,0,300),
(100,1,0,300),
(101,1,0,300),
(102,1,0,300),
(103,1,0,300),
(200,1,0,300),
(201,1,0,300),
(202,1,0,300),
(300,0,0,300),
(300,1,0,300),
(400,0,0,300),
(400,1,0,300)]"
subsection \<open>Transition Function h and related functions\<close>
lift_definition h :: "('state, 'input, 'output) fsm \<Rightarrow> ('state \<times> 'input) \<Rightarrow> ('output \<times> 'state) set"
is FSM_Impl.h .
lemma h_simps[simp]: "FSM.h M (q,x) = { (y,q') . (q,x,y,q') \<in> transitions M }"
by (transfer; auto)
lift_definition h_obs :: "('state, 'input, 'output) fsm \<Rightarrow> 'state \<Rightarrow> 'input \<Rightarrow> 'output \<Rightarrow> 'state option"
is FSM_Impl.h_obs .
lemma h_obs_simps[simp]: "FSM.h_obs M q x y = (let
tgts = snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))
in if card tgts = 1
then Some (the_elem tgts)
else None)"
by (transfer; auto)
fun defined_inputs' :: "(('a \<times>'b) \<Rightarrow> ('c\<times>'a) set) \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'b set" where
"defined_inputs' hM iM q = {x \<in> iM . hM (q,x) \<noteq> {}}"
fun defined_inputs :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> 'b set" where
"defined_inputs M q = defined_inputs' (h M) (inputs M) q"
lemma defined_inputs_set : "defined_inputs M q = {x \<in> inputs M . h M (q,x) \<noteq> {} }"
by auto
fun transitions_from' :: "(('a \<times>'b) \<Rightarrow> ('c\<times>'a) set) \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> ('a,'b,'c) transition set" where
"transitions_from' hM iM q = \<Union>(image (\<lambda>x . image (\<lambda>(y,q') . (q,x,y,q')) (hM (q,x))) iM)"
fun transitions_from :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('a,'b,'c) transition set" where
"transitions_from M q = transitions_from' (h M) (inputs M) q"
lemma transitions_from_set :
assumes "q \<in> states M"
shows "transitions_from M q = {t \<in> transitions M . t_source t = q}"
proof -
have "\<And> t . t \<in> transitions_from M q \<Longrightarrow> t \<in> transitions M \<and> t_source t = q" by auto
moreover have "\<And> t . t \<in> transitions M \<Longrightarrow> t_source t = q \<Longrightarrow> t \<in> transitions_from M q"
proof -
fix t assume "t \<in> transitions M" and "t_source t = q"
then have "(t_output t, t_target t) \<in> h M (q,t_input t)" and "t_input t \<in> inputs M" by auto
then have "t_input t \<in> defined_inputs' (h M) (inputs M) q"
unfolding defined_inputs'.simps \<open>t_source t = q\<close> by blast
have "(q, t_input t, t_output t, t_target t) \<in> transitions M"
using \<open>t_source t = q\<close> \<open>t \<in> transitions M\<close> by auto
then have "(q, t_input t, t_output t, t_target t) \<in> (\<lambda>(y, q'). (q, t_input t, y, q')) ` h M (q, t_input t)"
using \<open>(t_output t, t_target t) \<in> h M (q,t_input t)\<close>
unfolding h.simps
by (metis (no_types, lifting) image_iff prod.case_eq_if surjective_pairing)
then have "t \<in> (\<lambda>(y, q'). (q, t_input t, y, q')) ` h M (q, t_input t)"
using \<open>t_source t = q\<close> by (metis prod.collapse)
then show "t \<in> transitions_from M q"
unfolding transitions_from.simps transitions_from'.simps
using \<open>t_input t \<in> defined_inputs' (h M) (inputs M) q\<close>
using \<open>t_input t \<in> FSM.inputs M\<close> by blast
qed
ultimately show ?thesis by blast
qed
fun h_from :: "('state, 'input, 'output) fsm \<Rightarrow> 'state \<Rightarrow> ('input \<times> 'output \<times> 'state) set" where
"h_from M q = { (x,y,q') . (q,x,y,q') \<in> transitions M }"
lemma h_from[code] : "h_from M q = (let m = set_as_map (transitions M)
in (case m q of Some yqs \<Rightarrow> yqs | None \<Rightarrow> {}))"
unfolding set_as_map_def by force
fun h_out :: "('a,'b,'c) fsm \<Rightarrow> ('a \<times> 'b) \<Rightarrow> 'c set" where
"h_out M (q,x) = {y . \<exists> q' . (q,x,y,q') \<in> transitions M}"
lemma h_out_code[code]:
"h_out M = (\<lambda>qx . (case (set_as_map (image (\<lambda>(q,x,y,q') . ((q,x),y)) (transitions M))) qx of
Some yqs \<Rightarrow> yqs |
None \<Rightarrow> {}))"
proof -
let ?f = "(\<lambda>qx . (case (set_as_map (image (\<lambda>(q,x,y,q') . ((q,x),y)) (transitions M))) qx of Some yqs \<Rightarrow> yqs | None \<Rightarrow> {}))"
have "\<And> qx . (\<lambda>qx . (case (set_as_map (image (\<lambda>(q,x,y,q') . ((q,x),y)) (transitions M))) qx of Some yqs \<Rightarrow> yqs | None \<Rightarrow> {})) qx = (\<lambda> qx . {z. (qx, z) \<in> (\<lambda>(q, x, y, q'). ((q, x), y)) ` (transitions M)}) qx"
unfolding set_as_map_def by auto
moreover have "\<And> qx . (\<lambda> qx . {z. (qx, z) \<in> (\<lambda>(q, x, y, q'). ((q, x), y)) ` (transitions M)}) qx = (\<lambda> qx . {y | y . \<exists> q' . (fst qx, snd qx, y, q') \<in> (transitions M)}) qx"
by force
ultimately have "?f = (\<lambda> qx . {y | y . \<exists> q' . (fst qx, snd qx, y, q') \<in> (transitions M)})"
by blast
then have "?f = (\<lambda> (q,x) . {y | y . \<exists> q' . (q, x, y, q') \<in> (transitions M)})" by force
then show ?thesis by force
qed
lemma h_out_alt_def :
"h_out M (q,x) = {t_output t | t . t \<in> transitions M \<and> t_source t = q \<and> t_input t = x}"
unfolding h_out.simps
by auto
subsection \<open>Size\<close>
instantiation fsm :: (type,type,type) size
begin
definition size where [simp, code]: "size (m::('a, 'b, 'c) fsm) = card (states m)"
instance ..
end
lemma fsm_size_Suc :
"size M > 0"
unfolding FSM.size_def
using fsm_states_finite[of M] fsm_initial[of M]
using card_gt_0_iff by blast
subsection \<open>Paths\<close>
inductive path :: "('state, 'input, 'output) fsm \<Rightarrow> 'state \<Rightarrow> ('state, 'input, 'output) path \<Rightarrow> bool"
where
nil[intro!] : "q \<in> states M \<Longrightarrow> path M q []" |
cons[intro!] : "t \<in> transitions M \<Longrightarrow> path M (t_target t) ts \<Longrightarrow> path M (t_source t) (t#ts)"
inductive_cases path_nil_elim[elim!]: "path M q []"
inductive_cases path_cons_elim[elim!]: "path M q (t#ts)"
fun visited_states :: "'state \<Rightarrow> ('state, 'input, 'output) path \<Rightarrow> 'state list" where
"visited_states q p = (q # map t_target p)"
fun target :: "'state \<Rightarrow> ('state, 'input, 'output) path \<Rightarrow> 'state" where
"target q p = last (visited_states q p)"
lemma target_nil [simp] : "target q [] = q" by auto
lemma target_snoc [simp] : "target q (p@[t]) = t_target t" by auto
lemma path_begin_state :
assumes "path M q p"
shows "q \<in> states M"
using assms by (cases; auto)
lemma path_append[intro!] :
assumes "path M q p1"
and "path M (target q p1) p2"
shows "path M q (p1@p2)"
using assms by (induct p1 arbitrary: p2; auto)
lemma path_target_is_state :
assumes "path M q p"
shows "target q p \<in> states M"
using assms by (induct p; auto)
lemma path_suffix :
assumes "path M q (p1@p2)"
shows "path M (target q p1) p2"
using assms by (induction p1 arbitrary: q; auto)
lemma path_prefix :
assumes "path M q (p1@p2)"
shows "path M q p1"
using assms by (induction p1 arbitrary: q; auto; (metis path_begin_state))
lemma path_append_elim[elim!] :
assumes "path M q (p1@p2)"
obtains "path M q p1"
and "path M (target q p1) p2"
by (meson assms path_prefix path_suffix)
lemma path_append_target:
"target q (p1@p2) = target (target q p1) p2"
by (induction p1) (simp+)
lemma path_append_target_hd :
assumes "length p > 0"
shows "target q p = target (t_target (hd p)) (tl p)"
using assms by (induction p) (simp+)
lemma path_transitions :
assumes "path M q p"
shows "set p \<subseteq> transitions M"
using assms by (induct p arbitrary: q; fastforce)
lemma path_append_transition[intro!] :
assumes "path M q p"
and "t \<in> transitions M"
and "t_source t = target q p"
shows "path M q (p@[t])"
by (metis assms(1) assms(2) assms(3) cons fsm_transition_target nil path_append)
lemma path_append_transition_elim[elim!] :
assumes "path M q (p@[t])"
shows "path M q p"
and "t \<in> transitions M"
and "t_source t = target q p"
using assms by auto
lemma path_prepend_t : "path M q' p \<Longrightarrow> (q,x,y,q') \<in> transitions M \<Longrightarrow> path M q ((q,x,y,q')#p)"
by (metis (mono_tags, lifting) fst_conv path.intros(2) prod.sel(2))
lemma path_target_append : "target q1 p1 = q2 \<Longrightarrow> target q2 p2 = q3 \<Longrightarrow> target q1 (p1@p2) = q3"
by auto
lemma single_transition_path : "t \<in> transitions M \<Longrightarrow> path M (t_source t) [t]" by auto
lemma path_source_target_index :
assumes "Suc i < length p"
and "path M q p"
shows "t_target (p ! i) = t_source (p ! (Suc i))"
using assms proof (induction p rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc t ps)
then have "path M q ps" and "t_source t = target q ps" and "t \<in> transitions M" by auto
show ?case proof (cases "Suc i < length ps")
case True
then have "t_target (ps ! i) = t_source (ps ! Suc i)"
using snoc.IH \<open>path M q ps\<close> by auto
then show ?thesis
by (simp add: Suc_lessD True nth_append)
next
case False
then have "Suc i = length ps"
using snoc.prems(1) by auto
then have "(ps @ [t]) ! Suc i = t"
by auto
show ?thesis proof (cases "ps = []")
case True
then show ?thesis using \<open>Suc i = length ps\<close> by auto
next
case False
then have "target q ps = t_target (last ps)"
unfolding target.simps visited_states.simps
by (simp add: last_map)
then have "target q ps = t_target (ps ! i)"
using \<open>Suc i = length ps\<close>
by (metis False diff_Suc_1 last_conv_nth)
then show ?thesis
using \<open>t_source t = target q ps\<close>
by (metis \<open>(ps @ [t]) ! Suc i = t\<close> \<open>Suc i = length ps\<close> lessI nth_append)
qed
qed
qed
lemma paths_finite : "finite { p . path M q p \<and> length p \<le> k }"
proof -
have "{ p . path M q p \<and> length p \<le> k } \<subseteq> {xs . set xs \<subseteq> transitions M \<and> length xs \<le> k}"
by (metis (no_types, lifting) Collect_mono path_transitions)
then show "finite { p . path M q p \<and> length p \<le> k }"
using finite_lists_length_le[OF fsm_transitions_finite[of M], of k]
by (metis (mono_tags) finite_subset)
qed
lemma visited_states_prefix :
assumes "q' \<in> set (visited_states q p)"
shows "\<exists> p1 p2 . p = p1@p2 \<and> target q p1 = q'"
using assms proof (induction p arbitrary: q)
case Nil
then show ?case by auto
next
case (Cons a p)
then show ?case
proof (cases "q' \<in> set (visited_states (t_target a) p)")
case True
then obtain p1 p2 where "p = p1 @ p2 \<and> target (t_target a) p1 = q'"
using Cons.IH by blast
then have "(a#p) = (a#p1)@p2 \<and> target q (a#p1) = q'"
by auto
then show ?thesis by blast
next
case False
then have "q' = q"
using Cons.prems by auto
then show ?thesis
by auto
qed
qed
lemma visited_states_are_states :
assumes "path M q1 p"
shows "set (visited_states q1 p) \<subseteq> states M"
by (metis assms path_prefix path_target_is_state subsetI visited_states_prefix)
lemma transition_subset_path :
assumes "transitions A \<subseteq> transitions B"
and "path A q p"
and "q \<in> states B"
shows "path B q p"
using assms(2) proof (induction p rule: rev_induct)
case Nil
show ?case using assms(3) by auto
next
case (snoc t p)
then show ?case using assms(1) path_suffix
by fastforce
qed
subsubsection \<open>Paths of fixed length\<close>
fun paths_of_length' :: "('a,'b,'c) path \<Rightarrow> 'a \<Rightarrow> (('a \<times>'b) \<Rightarrow> ('c\<times>'a) set) \<Rightarrow> 'b set \<Rightarrow> nat \<Rightarrow> ('a,'b,'c) path set"
where
"paths_of_length' prev q hM iM 0 = {prev}" |
"paths_of_length' prev q hM iM (Suc k) =
(let hF = transitions_from' hM iM q
in \<Union> (image (\<lambda> t . paths_of_length' (prev@[t]) (t_target t) hM iM k) hF))"
fun paths_of_length :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> ('a,'b,'c) path set" where
"paths_of_length M q k = paths_of_length' [] q (h M) (inputs M) k"
subsubsection \<open>Paths up to fixed length\<close>
fun paths_up_to_length' :: "('a,'b,'c) path \<Rightarrow> 'a \<Rightarrow> (('a \<times>'b) \<Rightarrow> (('c\<times>'a) set)) \<Rightarrow> 'b set \<Rightarrow> nat \<Rightarrow> ('a,'b,'c) path set"
where
"paths_up_to_length' prev q hM iM 0 = {prev}" |
"paths_up_to_length' prev q hM iM (Suc k) =
(let hF = transitions_from' hM iM q
in insert prev (\<Union> (image (\<lambda> t . paths_up_to_length' (prev@[t]) (t_target t) hM iM k) hF)))"
fun paths_up_to_length :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> ('a,'b,'c) path set" where
"paths_up_to_length M q k = paths_up_to_length' [] q (h M) (inputs M) k"
lemma paths_up_to_length'_set :
assumes "q \<in> states M"
and "path M q prev"
shows "paths_up_to_length' prev (target q prev) (h M) (inputs M) k
= {(prev@p) | p . path M (target q prev) p \<and> length p \<le> k}"
using assms(2) proof (induction k arbitrary: prev)
case 0
show ?case unfolding paths_up_to_length'.simps using path_target_is_state[OF "0.prems"(1)] by auto
next
case (Suc k)
have "\<And> p . p \<in> paths_up_to_length' prev (target q prev) (h M) (inputs M) (Suc k)
\<Longrightarrow> p \<in> {(prev@p) | p . path M (target q prev) p \<and> length p \<le> Suc k}"
proof -
fix p assume "p \<in> paths_up_to_length' prev (target q prev) (h M) (inputs M) (Suc k)"
then show "p \<in> {(prev@p) | p . path M (target q prev) p \<and> length p \<le> Suc k}"
proof (cases "p = prev")
case True
show ?thesis using path_target_is_state[OF Suc.prems(1)] unfolding True by (simp add: nil)
next
case False
then have "p \<in> (\<Union> (image (\<lambda>t. paths_up_to_length' (prev@[t]) (t_target t) (h M) (inputs M) k)
(transitions_from' (h M) (inputs M) (target q prev))))"
using \<open>p \<in> paths_up_to_length' prev (target q prev) (h M) (inputs M) (Suc k)\<close>
unfolding paths_up_to_length'.simps Let_def by blast
then obtain t where "t \<in> \<Union>(image (\<lambda>x . image (\<lambda>(y,q') . ((target q prev),x,y,q'))
(h M ((target q prev),x))) (inputs M))"
and "p \<in> paths_up_to_length' (prev@[t]) (t_target t) (h M) (inputs M) k"
unfolding transitions_from'.simps by blast
have "t \<in> transitions M" and "t_source t = (target q prev)"
using \<open>t \<in> \<Union>(image (\<lambda>x . image (\<lambda>(y,q') . ((target q prev),x,y,q'))
(h M ((target q prev),x))) (inputs M))\<close> by auto
then have "path M q (prev@[t])"
using Suc.prems(1) using path_append_transition by simp
have "(target q (prev @ [t])) = t_target t" by auto
show ?thesis
using \<open>p \<in> paths_up_to_length' (prev@[t]) (t_target t) (h M) (inputs M) k\<close>
using Suc.IH[OF \<open>path M q (prev@[t])\<close>]
unfolding \<open>(target q (prev @ [t])) = t_target t\<close>
using \<open>path M q (prev @ [t])\<close> by auto
qed
qed
moreover have "\<And> p . p \<in> {(prev@p) | p . path M (target q prev) p \<and> length p \<le> Suc k}
\<Longrightarrow> p \<in> paths_up_to_length' prev (target q prev) (h M) (inputs M) (Suc k)"
proof -
fix p assume "p \<in> {(prev@p) | p . path M (target q prev) p \<and> length p \<le> Suc k}"
then obtain p' where "p = prev@p'"
and "path M (target q prev) p'"
and "length p' \<le> Suc k"
by blast
have "prev@p' \<in> paths_up_to_length' prev (target q prev) (h M) (inputs M) (Suc k)"
proof (cases p')
case Nil
then show ?thesis by auto
next
case (Cons t p'')
then have "t \<in> transitions M" and "t_source t = (target q prev)"
using \<open>path M (target q prev) p'\<close> by auto
then have "path M q (prev@[t])"
using Suc.prems(1) using path_append_transition by simp
have "(target q (prev @ [t])) = t_target t" by auto
have "length p'' \<le> k" using \<open>length p' \<le> Suc k\<close> Cons by auto
moreover have "path M (target q (prev@[t])) p''"
using \<open>path M (target q prev) p'\<close> unfolding Cons
by auto
ultimately have "p \<in> paths_up_to_length' (prev @ [t]) (t_target t) (h M) (FSM.inputs M) k"
using Suc.IH[OF \<open>path M q (prev@[t])\<close>]
unfolding \<open>(target q (prev @ [t])) = t_target t\<close> \<open>p = prev@p'\<close> Cons by simp
then have "prev@t#p'' \<in> paths_up_to_length' (prev @ [t]) (t_target t) (h M) (FSM.inputs M) k"
unfolding \<open>p = prev@p'\<close> Cons by auto
have "t \<in> (\<lambda>(y, q'). (t_source t, t_input t, y, q')) `
{(y, q'). (t_source t, t_input t, y, q') \<in> FSM.transitions M}"
using \<open>t \<in> transitions M\<close>
by (metis (no_types, lifting) case_prodI mem_Collect_eq pair_imageI surjective_pairing)
then have "t \<in> transitions_from' (h M) (inputs M) (target q prev)"
unfolding transitions_from'.simps
using fsm_transition_input[OF \<open>t \<in> transitions M\<close>]
unfolding \<open>t_source t = (target q prev)\<close>[symmetric] h_simps
by blast
then show ?thesis
using \<open>prev @ t # p'' \<in> paths_up_to_length' (prev@[t]) (t_target t) (h M) (FSM.inputs M) k\<close>
unfolding \<open>p = prev@p'\<close> Cons paths_up_to_length'.simps Let_def by blast
qed
then show "p \<in> paths_up_to_length' prev (target q prev) (h M) (inputs M) (Suc k)"
unfolding \<open>p = prev@p'\<close> by assumption
qed
ultimately show ?case by blast
qed
lemma paths_up_to_length_set :
assumes "q \<in> states M"
shows "paths_up_to_length M q k = {p . path M q p \<and> length p \<le> k}"
unfolding paths_up_to_length.simps
using paths_up_to_length'_set[OF assms nil[OF assms], of k] by auto
subsubsection \<open>Calculating Acyclic Paths\<close>
fun acyclic_paths_up_to_length' :: "('a,'b,'c) path \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> (('b\<times>'c\<times>'a) set)) \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> ('a,'b,'c) path set"
where
"acyclic_paths_up_to_length' prev q hF visitedStates 0 = {prev}" |
"acyclic_paths_up_to_length' prev q hF visitedStates (Suc k) =
(let tF = Set.filter (\<lambda> (x,y,q') . q' \<notin> visitedStates) (hF q)
in (insert prev (\<Union> (image (\<lambda> (x,y,q') . acyclic_paths_up_to_length' (prev@[(q,x,y,q')]) q' hF (insert q' visitedStates) k) tF))))"
fun p_source :: "'a \<Rightarrow> ('a,'b,'c) path \<Rightarrow> 'a"
where "p_source q p = hd (visited_states q p)"
lemma acyclic_paths_up_to_length'_prev :
"p' \<in> acyclic_paths_up_to_length' (prev@prev') q hF visitedStates k \<Longrightarrow> \<exists> p'' . p' = prev@p''"
by (induction k arbitrary: p' q visitedStates prev'; auto)
lemma acyclic_paths_up_to_length'_set :
assumes "path M (p_source q prev) prev"
and "\<And> q' . hF q' = {(x,y,q'') | x y q'' . (q',x,y,q'') \<in> transitions M}"
and "distinct (visited_states (p_source q prev) prev)"
and "visitedStates = set (visited_states (p_source q prev) prev)"
shows "acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates k
= { prev@p | p . path M (p_source q prev) (prev@p)
\<and> length p \<le> k
\<and> distinct (visited_states (p_source q prev) (prev@p)) }"
using assms proof (induction k arbitrary: q hF prev visitedStates)
case 0
then show ?case by auto
next
case (Suc k)
let ?tgt = "(target (p_source q prev) prev)"
have "\<And> p . (prev@p) \<in> acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates (Suc k)
\<Longrightarrow> path M (p_source q prev) (prev@p)
\<and> length p \<le> Suc k
\<and> distinct (visited_states (p_source q prev) (prev@p))"
proof -
fix p assume "(prev@p) \<in> acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates (Suc k)"
then consider (a) "(prev@p) = prev" |
(b) "(prev@p) \<in> (\<Union> (image (\<lambda> (x,y,q') . acyclic_paths_up_to_length' (prev@[(?tgt,x,y,q')]) q' hF (insert q' visitedStates) k)
(Set.filter (\<lambda> (x,y,q') . q' \<notin> visitedStates) (hF (target (p_source q prev) prev)))))"
by auto
then show "path M (p_source q prev) (prev@p) \<and> length p \<le> Suc k \<and> distinct (visited_states (p_source q prev) (prev@p))"
proof (cases)
case a
then show ?thesis using Suc.prems(1,3) by auto
next
case b
then obtain x y q' where *: "(x,y,q') \<in> Set.filter (\<lambda> (x,y,q') . q' \<notin> visitedStates) (hF ?tgt)"
and **:"(prev@p) \<in> acyclic_paths_up_to_length' (prev@[(?tgt,x,y,q')]) q' hF (insert q' visitedStates) k"
by auto
let ?t = "(?tgt,x,y,q')"
from * have "?t \<in> transitions M" and "q' \<notin> visitedStates"
using Suc.prems(2)[of ?tgt] by simp+
moreover have "t_source ?t = target (p_source q prev) prev"
by simp
moreover have "p_source (p_source q prev) (prev@[?t]) = p_source q prev"
by auto
ultimately have p1: "path M (p_source (p_source q prev) (prev@[?t])) (prev@[?t])"
using Suc.prems(1)
by (simp add: path_append_transition)
have "q' \<notin> set (visited_states (p_source q prev) prev)"
using \<open>q' \<notin> visitedStates\<close> Suc.prems(4) by auto
then have p2: "distinct (visited_states (p_source (p_source q prev) (prev@[?t])) (prev@[?t]))"
using Suc.prems(3) by auto
have p3: "(insert q' visitedStates)
= set (visited_states (p_source (p_source q prev) (prev@[?t])) (prev@[?t]))"
using Suc.prems(4) by auto
have ***: "(target (p_source (p_source q prev) (prev @ [(target (p_source q prev) prev, x, y, q')]))
(prev @ [(target (p_source q prev) prev, x, y, q')]))
= q'"
by auto
show ?thesis
using Suc.IH[OF p1 Suc.prems(2) p2 p3] **
unfolding ***
unfolding \<open>p_source (p_source q prev) (prev@[?t]) = p_source q prev\<close>
proof -
assume "acyclic_paths_up_to_length' (prev @ [(target (p_source q prev) prev, x, y, q')]) q' hF (insert q' visitedStates) k
= {(prev @ [(target (p_source q prev) prev, x, y, q')]) @ p |p.
path M (p_source q prev) ((prev @ [(target (p_source q prev) prev, x, y, q')]) @ p)
\<and> length p \<le> k
\<and> distinct (visited_states (p_source q prev) ((prev @ [(target (p_source q prev) prev, x, y, q')]) @ p))}"
then have "\<exists>ps. prev @ p = (prev @ [(target (p_source q prev) prev, x, y, q')]) @ ps
\<and> path M (p_source q prev) ((prev @ [(target (p_source q prev) prev, x, y, q')]) @ ps)
\<and> length ps \<le> k
\<and> distinct (visited_states (p_source q prev) ((prev @ [(target (p_source q prev) prev, x, y, q')]) @ ps))"
using \<open>prev @ p \<in> acyclic_paths_up_to_length' (prev @ [(target (p_source q prev) prev, x, y, q')]) q' hF (insert q' visitedStates) k\<close>
by blast
then show ?thesis
by (metis (no_types) Suc_le_mono append.assoc append.right_neutral append_Cons length_Cons same_append_eq)
qed
qed
qed
moreover have "\<And> p' . p' \<in> acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates (Suc k)
\<Longrightarrow> \<exists> p'' . p' = prev@p''"
using acyclic_paths_up_to_length'_prev[of _ prev "[]" "target (p_source q prev) prev" hF visitedStates "Suc k"]
by force
ultimately have fwd: "\<And> p' . p' \<in> acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates (Suc k)
\<Longrightarrow> p' \<in> { prev@p | p . path M (p_source q prev) (prev@p)
\<and> length p \<le> Suc k
\<and> distinct (visited_states (p_source q prev) (prev@p)) }"
by blast
have "\<And> p . path M (p_source q prev) (prev@p)
\<Longrightarrow> length p \<le> Suc k
\<Longrightarrow> distinct (visited_states (p_source q prev) (prev@p))
\<Longrightarrow> (prev@p) \<in> acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates (Suc k)"
proof -
fix p assume "path M (p_source q prev) (prev@p)"
and "length p \<le> Suc k"
and "distinct (visited_states (p_source q prev) (prev@p))"
show "(prev@p) \<in> acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates (Suc k)"
proof (cases p)
case Nil
then show ?thesis by auto
next
case (Cons t p')
then have "t_source t = target (p_source q (prev)) (prev)" and "t \<in> transitions M"
using \<open>path M (p_source q prev) (prev@p)\<close> by auto
have "path M (p_source q (prev@[t])) ((prev@[t])@p')"
and "path M (p_source q (prev@[t])) ((prev@[t]))"
using Cons \<open>path M (p_source q prev) (prev@p)\<close> by auto
have "length p' \<le> k"
using Cons \<open>length p \<le> Suc k\<close> by auto
have "distinct (visited_states (p_source q (prev@[t])) ((prev@[t])@p'))"
and "distinct (visited_states (p_source q (prev@[t])) ((prev@[t])))"
using Cons \<open>distinct (visited_states (p_source q prev) (prev@p))\<close> by auto
then have "t_target t \<notin> visitedStates"
using Suc.prems(4) by auto
let ?vN = "insert (t_target t) visitedStates"
have "?vN = set (visited_states (p_source q (prev @ [t])) (prev @ [t]))"
using Suc.prems(4) by auto
have "prev@p = prev@([t]@p')"
using Cons by auto
have "(prev@[t])@p' \<in> acyclic_paths_up_to_length' (prev @ [t]) (target (p_source q (prev @ [t])) (prev @ [t])) hF (insert (t_target t) visitedStates) k"
using Suc.IH[of q "prev@[t]", OF \<open>path M (p_source q (prev@[t])) ((prev@[t]))\<close> Suc.prems(2)
\<open>distinct (visited_states (p_source q (prev@[t])) ((prev@[t])))\<close>
\<open>?vN = set (visited_states (p_source q (prev @ [t])) (prev @ [t]))\<close> ]
using \<open>path M (p_source q (prev@[t])) ((prev@[t])@p')\<close>
\<open>length p' \<le> k\<close>
\<open>distinct (visited_states (p_source q (prev@[t])) ((prev@[t])@p'))\<close>
by force
then have "(prev@[t])@p' \<in> acyclic_paths_up_to_length' (prev@[t]) (t_target t) hF ?vN k"
by auto
moreover have "(t_input t,t_output t, t_target t) \<in> Set.filter (\<lambda> (x,y,q') . q' \<notin> visitedStates) (hF (t_source t))"
using Suc.prems(2)[of "t_source t"] \<open>t \<in> transitions M\<close> \<open>t_target t \<notin> visitedStates\<close>
proof -
have "\<exists>b c a. snd t = (b, c, a) \<and> (t_source t, b, c, a) \<in> FSM.transitions M"
by (metis (no_types) \<open>t \<in> FSM.transitions M\<close> prod.collapse)
then show ?thesis
using \<open>hF (t_source t) = {(x, y, q'') |x y q''. (t_source t, x, y, q'') \<in> FSM.transitions M}\<close>
\<open>t_target t \<notin> visitedStates\<close>
by fastforce
qed
ultimately have "\<exists> (x,y,q') \<in> (Set.filter (\<lambda> (x,y,q') . q' \<notin> visitedStates) (hF (target (p_source q prev) prev))) .
(prev@[t])@p' \<in> (acyclic_paths_up_to_length' (prev@[((target (p_source q prev) prev),x,y,q')]) q' hF (insert q' visitedStates) k)"
unfolding \<open>t_source t = target (p_source q (prev)) (prev)\<close>
by (metis (no_types, lifting) \<open>t_source t = target (p_source q prev) prev\<close> case_prodI prod.collapse)
then show ?thesis unfolding \<open>prev@p = prev@[t]@p'\<close>
unfolding acyclic_paths_up_to_length'.simps Let_def by force
qed
qed
then have rev: "\<And> p' . p' \<in> {prev@p | p . path M (p_source q prev) (prev@p)
\<and> length p \<le> Suc k
\<and> distinct (visited_states (p_source q prev) (prev@p))}
\<Longrightarrow> p' \<in> acyclic_paths_up_to_length' prev (target (p_source q prev) prev) hF visitedStates (Suc k)"
by blast
show ?case
using fwd rev by blast
qed
fun acyclic_paths_up_to_length :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> ('a,'b,'c) path set" where
"acyclic_paths_up_to_length M q k = {p. path M q p \<and> length p \<le> k \<and> distinct (visited_states q p)}"
lemma acyclic_paths_up_to_length_code[code] :
"acyclic_paths_up_to_length M q k = (if q \<in> states M
then acyclic_paths_up_to_length' [] q (m2f (set_as_map (transitions M))) {q} k
else {})"
proof (cases "q \<in> states M")
case False
then have "acyclic_paths_up_to_length M q k = {}"
using path_begin_state by fastforce
then show ?thesis using False by auto
next
case True
then have *: "path M (p_source q []) []" by auto
have **: "(\<And>q'. (m2f (set_as_map (transitions M))) q' = {(x, y, q'') |x y q''. (q', x, y, q'') \<in> FSM.transitions M})"
unfolding set_as_map_def by auto
have ***: "distinct (visited_states (p_source q []) [])"
by auto
have ****: "{q} = set (visited_states (p_source q []) [])"
by auto
show ?thesis
using acyclic_paths_up_to_length'_set[OF * ** *** ****, of k ]
using True by auto
qed
lemma path_map_target : "target (f4 q) (map (\<lambda> t . (f1 (t_source t), f2 (t_input t), f3 (t_output t), f4 (t_target t))) p) = f4 (target q p)"
by (induction p; auto)
lemma path_length_sum :
assumes "path M q p"
shows "length p = (\<Sum> q \<in> states M . length (filter (\<lambda>t. t_target t = q) p))"
using assms
proof (induction p rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc x xs)
then have "length xs = (\<Sum>q\<in>states M. length (filter (\<lambda>t. t_target t = q) xs))"
by auto
have *: "t_target x \<in> states M"
using \<open>path M q (xs @ [x])\<close> by auto
then have **: "length (filter (\<lambda>t. t_target t = t_target x) (xs @ [x]))
= Suc (length (filter (\<lambda>t. t_target t = t_target x) xs))"
by auto
have "\<And> q . q \<in> states M \<Longrightarrow> q \<noteq> t_target x
\<Longrightarrow> length (filter (\<lambda>t. t_target t = q) (xs @ [x])) = length (filter (\<lambda>t. t_target t = q) xs)"
by simp
then have ***: "(\<Sum>q\<in>states M - {t_target x}. length (filter (\<lambda>t. t_target t = q) (xs @ [x])))
= (\<Sum>q\<in>states M - {t_target x}. length (filter (\<lambda>t. t_target t = q) xs))"
using fsm_states_finite[of M]
by (metis (no_types, lifting) DiffE insertCI sum.cong)
have "(\<Sum>q\<in>states M. length (filter (\<lambda>t. t_target t = q) (xs @ [x])))
= (\<Sum>q\<in>states M - {t_target x}. length (filter (\<lambda>t. t_target t = q) (xs @ [x])))
+ (length (filter (\<lambda>t. t_target t = t_target x) (xs @ [x])))"
using * fsm_states_finite[of M]
proof -
have "(\<Sum>a\<in>insert (t_target x) (states M). length (filter (\<lambda>p. t_target p = a) (xs @ [x])))
= (\<Sum>a\<in>states M. length (filter (\<lambda>p. t_target p = a) (xs @ [x])))"
by (simp add: \<open>t_target x \<in> states M\<close> insert_absorb)
then show ?thesis
by (simp add: \<open>finite (states M)\<close> sum.insert_remove)
qed
moreover have "(\<Sum>q\<in>states M. length (filter (\<lambda>t. t_target t = q) xs))
= (\<Sum>q\<in>states M - {t_target x}. length (filter (\<lambda>t. t_target t = q) xs))
+ (length (filter (\<lambda>t. t_target t = t_target x) xs))"
using * fsm_states_finite[of M]
proof -
have "(\<Sum>a\<in>insert (t_target x) (states M). length (filter (\<lambda>p. t_target p = a) xs))
= (\<Sum>a\<in>states M. length (filter (\<lambda>p. t_target p = a) xs))"
by (simp add: \<open>t_target x \<in> states M\<close> insert_absorb)
then show ?thesis
by (simp add: \<open>finite (states M)\<close> sum.insert_remove)
qed
ultimately have "(\<Sum>q\<in>states M. length (filter (\<lambda>t. t_target t = q) (xs @ [x])))
= Suc (\<Sum>q\<in>states M. length (filter (\<lambda>t. t_target t = q) xs))"
using ** *** by auto
then show ?case
by (simp add: \<open>length xs = (\<Sum>q\<in>states M. length (filter (\<lambda>t. t_target t = q) xs))\<close>)
qed
lemma path_loop_cut :
assumes "path M q p"
and "t_target (p ! i) = t_target (p ! j)"
and "i < j"
and "j < length p"
shows "path M q ((take (Suc i) p) @ (drop (Suc j) p))"
and "target q ((take (Suc i) p) @ (drop (Suc j) p)) = target q p"
and "length ((take (Suc i) p) @ (drop (Suc j) p)) < length p"
and "path M (target q (take (Suc i) p)) (drop (Suc i) (take (Suc j) p))"
and "target (target q (take (Suc i) p)) (drop (Suc i) (take (Suc j) p)) = (target q (take (Suc i) p))"
proof -
have "p = (take (Suc j) p) @ (drop (Suc j) p)"
by auto
also have "\<dots> = ((take (Suc i) (take (Suc j) p)) @ (drop (Suc i) (take (Suc j) p))) @ (drop (Suc j) p)"
by (metis append_take_drop_id)
also have "\<dots> = ((take (Suc i) p) @ (drop (Suc i) (take (Suc j) p))) @ (drop (Suc j) p)"
using \<open>i < j\<close> by simp
finally have "p = (take (Suc i) p) @ (drop (Suc i) (take (Suc j) p)) @ (drop (Suc j) p)"
by simp
then have "path M q ((take (Suc i) p) @ (drop (Suc i) (take (Suc j) p)) @ (drop (Suc j) p))"
and "path M q (((take (Suc i) p) @ (drop (Suc i) (take (Suc j) p))) @ (drop (Suc j) p))"
using \<open>path M q p\<close> by auto
have "path M q (take (Suc i) p)" and "path M (target q (take (Suc i) p)) (drop (Suc i) (take (Suc j) p) @ drop (Suc j) p)"
using path_append_elim[OF \<open>path M q ((take (Suc i) p) @ (drop (Suc i) (take (Suc j) p)) @ (drop (Suc j) p))\<close>]
by blast+
have *: "(take (Suc i) p @ drop (Suc i) (take (Suc j) p)) = (take (Suc j) p)"
using \<open>i < j\<close> append_take_drop_id
by (metis \<open>(take (Suc i) (take (Suc j) p) @ drop (Suc i) (take (Suc j) p)) @ drop (Suc j) p = (take (Suc i) p @ drop (Suc i) (take (Suc j) p)) @ drop (Suc j) p\<close> append_same_eq)
have "path M q (take (Suc j) p)" and "path M (target q (take (Suc j) p)) (drop (Suc j) p)"
using path_append_elim[OF \<open>path M q (((take (Suc i) p) @ (drop (Suc i) (take (Suc j) p))) @ (drop (Suc j) p))\<close>]
unfolding *
by blast+
have **: "(target q (take (Suc j) p)) = (target q (take (Suc i) p))"
proof -
have "p ! i = last (take (Suc i) p)"
by (metis Suc_lessD assms(3) assms(4) less_trans_Suc take_last_index)
moreover have "p ! j = last (take (Suc j) p)"
by (simp add: assms(4) take_last_index)
ultimately show ?thesis
using assms(2) unfolding * target.simps visited_states.simps
by (simp add: last_map)
qed
show "path M q ((take (Suc i) p) @ (drop (Suc j) p))"
using \<open>path M q (take (Suc i) p)\<close> \<open>path M (target q (take (Suc j) p)) (drop (Suc j) p)\<close> unfolding ** by auto
show "target q ((take (Suc i) p) @ (drop (Suc j) p)) = target q p"
by (metis "**" append_take_drop_id path_append_target)
show "length ((take (Suc i) p) @ (drop (Suc j) p)) < length p"
proof -
have ***: "length p = length ((take (Suc j) p) @ (drop (Suc j) p))"
by auto
have "length (take (Suc i) p) < length (take (Suc j) p)"
using assms(3,4)
by (simp add: min_absorb2)
have scheme: "\<And> a b c . length a < length b \<Longrightarrow> length (a@c) < length (b@c)"
by auto
show ?thesis
unfolding *** using scheme[OF \<open>length (take (Suc i) p) < length (take (Suc j) p)\<close>, of "(drop (Suc j) p)"]
by assumption
qed
show "path M (target q (take (Suc i) p)) (drop (Suc i) (take (Suc j) p))"
using \<open>path M (target q (take (Suc i) p)) (drop (Suc i) (take (Suc j) p) @ drop (Suc j) p)\<close> by blast
show "target (target q (take (Suc i) p)) (drop (Suc i) (take (Suc j) p)) = (target q (take (Suc i) p))"
by (metis "*" "**" path_append_target)
qed
lemma path_prefix_take :
assumes "path M q p"
shows "path M q (take i p)"
proof -
have "p = (take i p)@(drop i p)" by auto
then have "path M q ((take i p)@(drop i p))" using assms by auto
then show ?thesis
by blast
qed
subsection \<open>Acyclic Paths\<close>
lemma cyclic_path_loop :
assumes "path M q p"
and "\<not> distinct (visited_states q p)"
shows "\<exists> p1 p2 p3 . p = p1@p2@p3 \<and> p2 \<noteq> [] \<and> target q p1 = target q (p1@p2)"
using assms proof (induction p arbitrary: q)
case (nil M q)
then show ?case by auto
next
case (cons t M ts)
then show ?case
proof (cases "distinct (visited_states (t_target t) ts)")
case True
then have "q \<in> set (visited_states (t_target t) ts)"
using cons.prems by simp
then obtain p2 p3 where "ts = p2@p3" and "target (t_target t) p2 = q"
using visited_states_prefix[of q "t_target t" ts] by blast
then have "(t#ts) = []@(t#p2)@p3 \<and> (t#p2) \<noteq> [] \<and> target q [] = target q ([]@(t#p2))"
using cons.hyps by auto
then show ?thesis by blast
next
case False
then obtain p1 p2 p3 where "ts = p1@p2@p3" and "p2 \<noteq> []"
and "target (t_target t) p1 = target (t_target t) (p1@p2)"
using cons.IH by blast
then have "t#ts = (t#p1)@p2@p3 \<and> p2 \<noteq> [] \<and> target q (t#p1) = target q ((t#p1)@p2)"
by simp
then show ?thesis by blast
qed
qed
lemma cyclic_path_pumping :
assumes "path M (initial M) p"
and "\<not> distinct (visited_states (initial M) p)"
shows "\<exists> p . path M (initial M) p \<and> length p \<ge> n"
proof -
from assms obtain p1 p2 p3 where "p = p1 @ p2 @ p3" and "p2 \<noteq> []"
and "target (initial M) p1 = target (initial M) (p1 @ p2)"
using cyclic_path_loop[of M "initial M" p] by blast
then have "path M (target (initial M) p1) p3"
using path_suffix[of M "initial M" "p1@p2" p3] \<open>path M (initial M) p\<close> by auto
have "path M (initial M) p1"
using path_prefix[of M "initial M" p1 "p2@p3"] \<open>path M (initial M) p\<close> \<open>p = p1 @ p2 @ p3\<close>
by auto
have "path M (initial M) ((p1@p2)@p3)"
using \<open>path M (initial M) p\<close> \<open>p = p1 @ p2 @ p3\<close>
by auto
have "path M (target (initial M) p1) p2"
using path_suffix[of M "initial M" p1 p2, OF path_prefix[of M "initial M" "p1@p2" p3, OF \<open>path M (initial M) ((p1@p2)@p3)\<close>]]
by assumption
have "target (target (initial M) p1) p2 = (target (initial M) p1)"
using path_append_target \<open>target (initial M) p1 = target (initial M) (p1 @ p2)\<close>
by auto
have "path M (initial M) (p1 @ (concat (replicate n p2)) @ p3)"
proof (induction n)
case 0
then show ?case
using path_append[OF \<open>path M (initial M) p1\<close> \<open>path M (target (initial M) p1) p3\<close>]
by auto
next
case (Suc n)
then show ?case
using \<open>path M (target (initial M) p1) p2\<close> \<open>target (target (initial M) p1) p2 = target (initial M) p1\<close>
by auto
qed
moreover have "length (p1 @ (concat (replicate n p2)) @ p3) \<ge> n"
proof -
have "length (concat (replicate n p2)) = n * (length p2)"
using concat_replicate_length by metis
moreover have "length p2 > 0"
using \<open>p2 \<noteq> []\<close> by auto
ultimately have "length (concat (replicate n p2)) \<ge> n"
by (simp add: Suc_leI)
then show ?thesis by auto
qed
ultimately show "\<exists> p . path M (initial M) p \<and> length p \<ge> n" by blast
qed
lemma cyclic_path_shortening :
assumes "path M q p"
and "\<not> distinct (visited_states q p)"
shows "\<exists> p' . path M q p' \<and> target q p' = target q p \<and> length p' < length p"
proof -
obtain p1 p2 p3 where *: "p = p1@p2@p3 \<and> p2 \<noteq> [] \<and> target q p1 = target q (p1@p2)"
using cyclic_path_loop[OF assms] by blast
then have "path M q (p1@p3)"
using assms(1) by force
moreover have "target q (p1@p3) = target q p"
by (metis (full_types) * path_append_target)
moreover have "length (p1@p3) < length p"
using * by auto
ultimately show ?thesis by blast
qed
lemma acyclic_path_from_cyclic_path :
assumes "path M q p"
and "\<not> distinct (visited_states q p)"
obtains p' where "path M q p'" and "target q p = target q p'" and "distinct (visited_states q p')"
proof -
let ?paths = "{p' . (path M q p' \<and> target q p' = target q p \<and> length p' \<le> length p)}"
let ?minPath = "arg_min length (\<lambda> io . io \<in> ?paths)"
have "?paths \<noteq> empty"
using assms(1) by auto
moreover have "finite ?paths"
using paths_finite[of M q "length p"]
by (metis (no_types, lifting) Collect_mono rev_finite_subset)
ultimately have minPath_def : "?minPath \<in> ?paths \<and> (\<forall> p' \<in> ?paths . length ?minPath \<le> length p')"
by (meson arg_min_nat_lemma equals0I)
then have "path M q ?minPath" and "target q ?minPath = target q p"
by auto
moreover have "distinct (visited_states q ?minPath)"
proof (rule ccontr)
assume "\<not> distinct (visited_states q ?minPath)"
have "\<exists> p' . path M q p' \<and> target q p' = target q p \<and> length p' < length ?minPath"
using cyclic_path_shortening[OF \<open>path M q ?minPath\<close> \<open>\<not> distinct (visited_states q ?minPath)\<close>] minPath_def
\<open>target q ?minPath= target q p\<close> by auto
then show "False"
using minPath_def using arg_min_nat_le dual_order.strict_trans1 by auto
qed
ultimately show ?thesis
by (simp add: that)
qed
lemma acyclic_path_length_limit :
assumes "path M q p"
and "distinct (visited_states q p)"
shows "length p < size M"
proof (rule ccontr)
assume *: "\<not> length p < size M"
then have "length p \<ge> card (states M)"
using size_def by auto
then have "length (visited_states q p) > card (states M)"
by auto
moreover have "set (visited_states q p) \<subseteq> states M"
by (metis assms(1) path_prefix path_target_is_state subsetI visited_states_prefix)
ultimately have "\<not> distinct (visited_states q p)"
using distinct_card[OF assms(2)]
using List.finite_set[of "visited_states q p"]
by (metis card_mono fsm_states_finite leD)
then show "False" using assms(2) by blast
qed
subsection \<open>Reachable States\<close>
definition reachable :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> bool" where
"reachable M q = (\<exists> p . path M (initial M) p \<and> target (initial M) p = q)"
definition reachable_states :: "('a,'b,'c) fsm \<Rightarrow> 'a set" where
"reachable_states M = {target (initial M) p | p . path M (initial M) p }"
abbreviation "size_r M \<equiv> card (reachable_states M)"
lemma acyclic_paths_set :
"acyclic_paths_up_to_length M q (size M - 1) = {p . path M q p \<and> distinct (visited_states q p)}"
unfolding acyclic_paths_up_to_length.simps using acyclic_path_length_limit[of M q]
by (metis (no_types, lifting) One_nat_def Suc_pred cyclic_path_shortening leD list.size(3)
not_less_eq_eq not_less_zero path.intros(1) path_begin_state)
(* inefficient calculation, as a state may be target of a large number of acyclic paths *)
lemma reachable_states_code[code] :
"reachable_states M = image (target (initial M)) (acyclic_paths_up_to_length M (initial M) (size M - 1))"
proof -
have "\<And> q' . q' \<in> reachable_states M
\<Longrightarrow> q' \<in> image (target (initial M)) (acyclic_paths_up_to_length M (initial M) (size M - 1))"
proof -
fix q' assume "q' \<in> reachable_states M"
then obtain p where "path M (initial M) p" and "target (initial M) p = q'"
unfolding reachable_states_def by blast
obtain p' where "path M (initial M) p'" and "target (initial M) p' = q'"
and "distinct (visited_states (initial M) p')"
proof (cases "distinct (visited_states (initial M) p)")
case True
then show ?thesis using \<open>path M (initial M) p\<close> \<open>target (initial M) p = q'\<close> that by auto
next
case False
then show ?thesis
using acyclic_path_from_cyclic_path[OF \<open>path M (initial M) p\<close>]
unfolding \<open>target (initial M) p = q'\<close> using that by blast
qed
then show "q' \<in> image (target (initial M)) (acyclic_paths_up_to_length M (initial M) (size M - 1))"
unfolding acyclic_paths_set by force
qed
moreover have "\<And> q' . q' \<in> image (target (initial M)) (acyclic_paths_up_to_length M (initial M) (size M - 1))
\<Longrightarrow> q' \<in> reachable_states M"
unfolding reachable_states_def acyclic_paths_set by blast
ultimately show ?thesis by blast
qed
lemma reachable_states_intro[intro!] :
assumes "path M (initial M) p"
shows "target (initial M) p \<in> reachable_states M"
using assms unfolding reachable_states_def by auto
lemma reachable_states_initial :
"initial M \<in> reachable_states M"
unfolding reachable_states_def by auto
lemma reachable_states_next :
assumes "q \<in> reachable_states M" and "t \<in> transitions M" and "t_source t = q"
shows "t_target t \<in> reachable_states M"
proof -
from \<open>q \<in> reachable_states M\<close> obtain p where * :"path M (initial M) p"
and **:"target (initial M) p = q"
unfolding reachable_states_def by auto
then have "path M (initial M) (p@[t])" using assms(2,3) path_append_transition by metis
moreover have "target (initial M) (p@[t]) = t_target t" by auto
ultimately show ?thesis
unfolding reachable_states_def
by (metis (mono_tags, lifting) mem_Collect_eq)
qed
lemma reachable_states_path :
assumes "q \<in> reachable_states M"
and "path M q p"
and "t \<in> set p"
shows "t_source t \<in> reachable_states M"
using assms unfolding reachable_states_def proof (induction p arbitrary: q)
case Nil
then show ?case by auto
next
case (Cons t' p')
then show ?case proof (cases "t = t'")
case True
then show ?thesis using Cons.prems(1,2) by force
next
case False then show ?thesis using Cons
by (metis (mono_tags, lifting) path_cons_elim reachable_states_def reachable_states_next
set_ConsD)
qed
qed
lemma reachable_states_initial_or_target :
assumes "q \<in> reachable_states M"
shows "q = initial M \<or> (\<exists> t \<in> transitions M . t_source t \<in> reachable_states M \<and> t_target t = q)"
proof -
obtain p where "path M (initial M) p" and "target (initial M) p = q"
using assms unfolding reachable_states_def by auto
show ?thesis proof (cases p rule: rev_cases)
case Nil
then show ?thesis using \<open>path M (initial M) p\<close> \<open>target (initial M) p = q\<close> by auto
next
case (snoc p' t)
have "t \<in> transitions M"
using \<open>path M (initial M) p\<close> unfolding snoc by auto
moreover have "t_target t = q"
using \<open>target (initial M) p = q\<close> unfolding snoc by auto
moreover have "t_source t \<in> reachable_states M"
using \<open>path M (initial M) p\<close> unfolding snoc
by (metis append_is_Nil_conv last_in_set last_snoc not_Cons_self2 reachable_states_initial reachable_states_path)
ultimately show ?thesis
by blast
qed
qed
lemma reachable_state_is_state :
"q \<in> reachable_states M \<Longrightarrow> q \<in> states M"
unfolding reachable_states_def using path_target_is_state by fastforce
lemma reachable_states_finite : "finite (reachable_states M)"
using fsm_states_finite[of M] reachable_state_is_state[of _ M]
by (meson finite_subset subset_eq)
subsection \<open>Language\<close>
abbreviation "p_io (p :: ('state,'input,'output) path) \<equiv> map (\<lambda> t . (t_input t, t_output t)) p"
fun language_state_for_input :: "('state,'input,'output) fsm \<Rightarrow> 'state \<Rightarrow> 'input list \<Rightarrow> ('input \<times> 'output) list set" where
"language_state_for_input M q xs = {p_io p | p . path M q p \<and> map fst (p_io p) = xs}"
fun LS\<^sub>i\<^sub>n :: "('state,'input,'output) fsm \<Rightarrow> 'state \<Rightarrow> 'input list set \<Rightarrow> ('input \<times> 'output) list set" where
"LS\<^sub>i\<^sub>n M q xss = {p_io p | p . path M q p \<and> map fst (p_io p) \<in> xss}"
abbreviation(input) "L\<^sub>i\<^sub>n M \<equiv> LS\<^sub>i\<^sub>n M (initial M)"
lemma language_state_for_input_inputs :
assumes "io \<in> language_state_for_input M q xs"
shows "map fst io = xs"
using assms by auto
lemma language_state_for_inputs_inputs :
assumes "io \<in> LS\<^sub>i\<^sub>n M q xss"
shows "map fst io \<in> xss" using assms by auto
fun LS :: "('state,'input,'output) fsm \<Rightarrow> 'state \<Rightarrow> ('input \<times> 'output) list set" where
"LS M q = { p_io p | p . path M q p }"
abbreviation "L M \<equiv> LS M (initial M)"
lemma language_state_containment :
assumes "path M q p"
and "p_io p = io"
shows "io \<in> LS M q"
using assms by auto
lemma language_prefix :
assumes "io1@io2 \<in> LS M q"
shows "io1 \<in> LS M q"
proof -
obtain p where "path M q p" and "p_io p = io1@io2"
using assms by auto
let ?tp = "take (length io1) p"
have "path M q ?tp"
by (metis (no_types) \<open>path M q p\<close> append_take_drop_id path_prefix)
moreover have "p_io ?tp = io1"
using \<open>p_io p = io1@io2\<close> by (metis append_eq_conv_conj take_map)
ultimately show ?thesis
by force
qed
lemma language_contains_empty_sequence : "[] \<in> L M"
by auto
lemma language_state_split :
assumes "io1 @ io2 \<in> LS M q1"
obtains p1 p2 where "path M q1 p1"
and "path M (target q1 p1) p2"
and "p_io p1 = io1"
and "p_io p2 = io2"
proof -
obtain p12 where "path M q1 p12" and "p_io p12 = io1 @ io2"
using assms unfolding LS.simps by auto
let ?p1 = "take (length io1) p12"
let ?p2 = "drop (length io1) p12"
have "p12 = ?p1 @ ?p2"
by auto
then have "path M q1 (?p1 @ ?p2)"
using \<open>path M q1 p12\<close> by auto
have "path M q1 ?p1" and "path M (target q1 ?p1) ?p2"
using path_append_elim[OF \<open>path M q1 (?p1 @ ?p2)\<close>] by blast+
moreover have "p_io ?p1 = io1"
using \<open>p12 = ?p1 @ ?p2\<close> \<open>p_io p12 = io1 @ io2\<close>
by (metis append_eq_conv_conj take_map)
moreover have "p_io ?p2 = io2"
using \<open>p12 = ?p1 @ ?p2\<close> \<open>p_io p12 = io1 @ io2\<close>
by (metis (no_types) \<open>p_io p12 = io1 @ io2\<close> append_eq_conv_conj drop_map)
ultimately show ?thesis using that by blast
qed
lemma language_initial_path_append_transition :
assumes "ios @ [io] \<in> L M"
obtains p t where "path M (initial M) (p@[t])" and "p_io (p@[t]) = ios @ [io]"
proof -
obtain pt where "path M (initial M) pt" and "p_io pt = ios @ [io]"
using assms unfolding LS.simps by auto
then have "pt \<noteq> []"
by auto
then obtain p t where "pt = p @ [t]"
using rev_exhaust by blast
then have "path M (initial M) (p@[t])" and "p_io (p@[t]) = ios @ [io]"
using \<open>path M (initial M) pt\<close> \<open>p_io pt = ios @ [io]\<close> by auto
then show ?thesis using that by simp
qed
lemma language_path_append_transition :
assumes "ios @ [io] \<in> LS M q"
obtains p t where "path M q (p@[t])" and "p_io (p@[t]) = ios @ [io]"
proof -
obtain pt where "path M q pt" and "p_io pt = ios @ [io]"
using assms unfolding LS.simps by auto
then have "pt \<noteq> []"
by auto
then obtain p t where "pt = p @ [t]"
using rev_exhaust by blast
then have "path M q (p@[t])" and "p_io (p@[t]) = ios @ [io]"
using \<open>path M q pt\<close> \<open>p_io pt = ios @ [io]\<close> by auto
then show ?thesis using that by simp
qed
lemma language_split :
assumes "io1@io2 \<in> L M"
obtains p1 p2 where "path M (initial M) (p1@p2)" and "p_io p1 = io1" and "p_io p2 = io2"
proof -
from assms obtain p where "path M (initial M) p" and "p_io p = io1 @ io2"
by auto
let ?p1 = "take (length io1) p"
let ?p2 = "drop (length io1) p"
have "path M (initial M) (?p1@?p2)"
using \<open>path M (initial M) p\<close> by simp
moreover have "p_io ?p1 = io1"
using \<open>p_io p = io1 @ io2\<close>
by (metis append_eq_conv_conj take_map)
moreover have "p_io ?p2 = io2"
using \<open>p_io p = io1 @ io2\<close>
by (metis append_eq_conv_conj drop_map)
ultimately show ?thesis using that by blast
qed
lemma language_io :
assumes "io \<in> LS M q"
and "(x,y) \<in> set io"
shows "x \<in> (inputs M)"
and "y \<in> outputs M"
proof -
obtain p where "path M q p" and "p_io p = io"
using \<open>io \<in> LS M q\<close> by auto
then obtain t where "t \<in> set p" and "t_input t = x" and "t_output t = y"
using \<open>(x,y) \<in> set io\<close> by auto
have "t \<in> transitions M"
using \<open>path M q p\<close> \<open>t \<in> set p\<close>
by (induction p; auto)
show "x \<in> (inputs M)"
using \<open>t \<in> transitions M\<close> \<open>t_input t = x\<close> by auto
show "y \<in> outputs M"
using \<open>t \<in> transitions M\<close> \<open>t_output t = y\<close> by auto
qed
lemma path_io_split :
assumes "path M q p"
and "p_io p = io1@io2"
shows "path M q (take (length io1) p)"
and "p_io (take (length io1) p) = io1"
and "path M (target q (take (length io1) p)) (drop (length io1) p)"
and "p_io (drop (length io1) p) = io2"
proof -
have "length io1 \<le> length p"
using \<open>p_io p = io1@io2\<close>
unfolding length_map[of "(\<lambda> t . (t_input t, t_output t))", symmetric]
by auto
have "p = (take (length io1) p)@(drop (length io1) p)"
by simp
then have *: "path M q ((take (length io1) p)@(drop (length io1) p))"
using \<open>path M q p\<close> by auto
show "path M q (take (length io1) p)"
and "path M (target q (take (length io1) p)) (drop (length io1) p)"
using path_append_elim[OF *] by blast+
show "p_io (take (length io1) p) = io1"
using \<open>p = (take (length io1) p)@(drop (length io1) p)\<close> \<open>p_io p = io1@io2\<close>
by (metis append_eq_conv_conj take_map)
show "p_io (drop (length io1) p) = io2"
using \<open>p = (take (length io1) p)@(drop (length io1) p)\<close> \<open>p_io p = io1@io2\<close>
by (metis append_eq_conv_conj drop_map)
qed
lemma language_intro :
assumes "path M q p"
shows "p_io p \<in> LS M q"
using assms unfolding LS.simps by auto
lemma language_prefix_append :
assumes "io1 @ (p_io p) \<in> L M"
shows "io1 @ p_io (take i p) \<in> L M"
proof -
fix i
have "p_io p = (p_io (take i p)) @ (p_io (drop i p))"
by (metis append_take_drop_id map_append)
then have "(io1 @ (p_io (take i p))) @ (p_io (drop i p)) \<in> L M"
using \<open>io1 @ p_io p \<in> L M\<close> by auto
show "io1 @ p_io (take i p) \<in> L M"
using language_prefix[OF \<open>(io1 @ (p_io (take i p))) @ (p_io (drop i p)) \<in> L M\<close>]
by assumption
qed
lemma language_finite: "finite {io . io \<in> L M \<and> length io \<le> k}"
proof -
have "{io . io \<in> L M \<and> length io \<le> k} \<subseteq> p_io ` {p. path M (FSM.initial M) p \<and> length p \<le> k}"
by auto
then show ?thesis
using paths_finite[of M "initial M" k]
using finite_surj by auto
qed
lemma LS_prepend_transition :
assumes "t \<in> transitions M"
and "io \<in> LS M (t_target t)"
shows "(t_input t, t_output t) # io \<in> LS M (t_source t)"
proof -
obtain p where "path M (t_target t) p" and "p_io p = io"
using assms(2) by auto
then have "path M (t_source t) (t#p)" and "p_io (t#p) = (t_input t, t_output t) # io"
using assms(1) by auto
then show ?thesis
unfolding LS.simps
by (metis (mono_tags, lifting) mem_Collect_eq)
qed
lemma language_empty_IO :
assumes "inputs M = {} \<or> outputs M = {}"
shows "L M = {[]}"
proof -
consider "inputs M = {}" | "outputs M = {}" using assms by blast
then show ?thesis proof cases
case 1
show "L M = {[]}"
using language_io(1)[of _ M "initial M"] unfolding 1
by (metis (no_types, opaque_lifting) ex_in_conv is_singletonI' is_singleton_the_elem language_contains_empty_sequence set_empty2 singleton_iff surj_pair)
next
case 2
show "L M = {[]}"
using language_io(2)[of _ M "initial M"] unfolding 2
by (metis (no_types, opaque_lifting) ex_in_conv is_singletonI' is_singleton_the_elem language_contains_empty_sequence set_empty2 singleton_iff surj_pair)
qed
qed
lemma language_equivalence_from_isomorphism_helper :
assumes "bij_betw f (states M1) (states M2)"
and "f (initial M1) = initial M2"
and "\<And> q x y q' . q \<in> states M1 \<Longrightarrow> q' \<in> states M1 \<Longrightarrow> (q,x,y,q') \<in> transitions M1 \<longleftrightarrow> (f q,x,y,f q') \<in> transitions M2"
and "q \<in> states M1"
shows "LS M1 q \<subseteq> LS M2 (f q)"
proof
fix io assume "io \<in> LS M1 q"
then obtain p where "path M1 q p" and "p_io p = io"
by auto
let ?f = "\<lambda>(q,x,y,q') . (f q,x,y,f q')"
let ?p = "map ?f p"
have "f q \<in> states M2"
using assms(1,4)
using bij_betwE by auto
have "path M2 (f q) ?p"
using \<open>path M1 q p\<close> proof (induction p rule: rev_induct)
case Nil
show ?case using \<open>f q \<in> states M2\<close> by auto
next
case (snoc a p)
then have "path M2 (f q) (map ?f p)"
by auto
have "target (f q) (map ?f p) = f (target q p)"
using \<open>f (initial M1) = initial M2\<close> assms(2)
by (induction p; auto)
then have "t_source (?f a) = target (f q) (map ?f p)"
by (metis (no_types, lifting) case_prod_beta' fst_conv path_append_transition_elim(3) snoc.prems)
have "a \<in> transitions M1"
using snoc.prems by auto
then have "?f a \<in> transitions M2"
by (metis (mono_tags, lifting) assms(3) case_prod_beta fsm_transition_source fsm_transition_target surjective_pairing)
have "map ?f (p@[a]) = (map ?f p)@[?f a]"
by auto
show ?case
unfolding \<open>map ?f (p@[a]) = (map ?f p)@[?f a]\<close>
using path_append_transition[OF \<open>path M2 (f q) (map ?f p)\<close> \<open>?f a \<in> transitions M2\<close> \<open>t_source (?f a) = target (f q) (map ?f p)\<close>]
by assumption
qed
moreover have "p_io ?p = io"
using \<open>p_io p = io\<close>
by (induction p; auto)
ultimately show "io \<in> LS M2 (f q)"
using language_state_containment by fastforce
qed
lemma language_equivalence_from_isomorphism :
assumes "bij_betw f (states M1) (states M2)"
and "f (initial M1) = initial M2"
and "\<And> q x y q' . q \<in> states M1 \<Longrightarrow> q' \<in> states M1 \<Longrightarrow> (q,x,y,q') \<in> transitions M1 \<longleftrightarrow> (f q,x,y,f q') \<in> transitions M2"
and "q \<in> states M1"
shows "LS M1 q = LS M2 (f q)"
proof
show "LS M1 q \<subseteq> LS M2 (f q)"
using language_equivalence_from_isomorphism_helper[OF assms] .
have "f q \<in> states M2"
using assms(1,4)
using bij_betwE by auto
have "(inv_into (FSM.states M1) f (f q)) = q"
by (meson assms(1) assms(4) bij_betw_imp_inj_on inv_into_f_f)
have "bij_betw (inv_into (states M1) f) (states M2) (states M1)"
using bij_betw_inv_into[OF assms(1)] .
moreover have "(inv_into (states M1) f) (initial M2) = (initial M1)"
using assms(1,2)
by (metis bij_betw_inv_into_left fsm_initial)
moreover have "\<And> q x y q' . q \<in> states M2 \<Longrightarrow> q' \<in> states M2 \<Longrightarrow> (q,x,y,q') \<in> transitions M2 \<longleftrightarrow> ((inv_into (states M1) f) q,x,y,(inv_into (states M1) f) q') \<in> transitions M1"
proof
fix q x y q' assume "q \<in> states M2" and "q' \<in> states M2"
show "(q,x,y,q') \<in> transitions M2 \<Longrightarrow> ((inv_into (states M1) f) q,x,y,(inv_into (states M1) f) q') \<in> transitions M1"
proof -
assume a1: "(q, x, y, q') \<in> FSM.transitions M2"
have f2: "\<forall>f B A. \<not> bij_betw f B A \<or> (\<forall>b. (b::'b) \<notin> B \<or> (f b::'a) \<in> A)"
using bij_betwE by blast
then have f3: "inv_into (states M1) f q \<in> states M1"
using \<open>q \<in> states M2\<close> calculation(1) by blast
have "inv_into (states M1) f q' \<in> states M1"
using f2 \<open>q' \<in> states M2\<close> calculation(1) by blast
then show ?thesis
using f3 a1 \<open>q \<in> states M2\<close> \<open>q' \<in> states M2\<close> assms(1) assms(3) bij_betw_inv_into_right by fastforce
qed
show "((inv_into (states M1) f) q,x,y,(inv_into (states M1) f) q') \<in> transitions M1 \<Longrightarrow> (q,x,y,q') \<in> transitions M2"
proof -
assume a1: "(inv_into (states M1) f q, x, y, inv_into (states M1) f q') \<in> FSM.transitions M1"
have f2: "\<forall>f B A. \<not> bij_betw f B A \<or> (\<forall>b. (b::'b) \<notin> B \<or> (f b::'a) \<in> A)"
by (metis (full_types) bij_betwE)
then have f3: "inv_into (states M1) f q' \<in> states M1"
using \<open>q' \<in> states M2\<close> calculation(1) by blast
have "inv_into (states M1) f q \<in> states M1"
using f2 \<open>q \<in> states M2\<close> calculation(1) by blast
then show ?thesis
using f3 a1 \<open>q \<in> states M2\<close> \<open>q' \<in> states M2\<close> assms(1) assms(3) bij_betw_inv_into_right by fastforce
qed
qed
ultimately show "LS M2 (f q) \<subseteq> LS M1 q"
using language_equivalence_from_isomorphism_helper[of "(inv_into (states M1) f)" M2 M1, OF _ _ _ \<open>f q \<in> states M2\<close>]
unfolding \<open>(inv_into (FSM.states M1) f (f q)) = q\<close>
by blast
qed
lemma language_equivalence_from_isomorphism_helper_reachable :
assumes "bij_betw f (reachable_states M1) (reachable_states M2)"
and "f (initial M1) = initial M2"
and "\<And> q x y q' . q \<in> reachable_states M1 \<Longrightarrow> q' \<in> reachable_states M1 \<Longrightarrow> (q,x,y,q') \<in> transitions M1 \<longleftrightarrow> (f q,x,y,f q') \<in> transitions M2"
shows "L M1 \<subseteq> L M2"
proof
fix io assume "io \<in> L M1"
then obtain p where "path M1 (initial M1) p" and "p_io p = io"
by auto
let ?f = "\<lambda>(q,x,y,q') . (f q,x,y,f q')"
let ?p = "map ?f p"
have "path M2 (initial M2) ?p"
using \<open>path M1 (initial M1) p\<close> proof (induction p rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc a p)
then have "path M2 (initial M2) (map ?f p)"
by auto
have "target (initial M2) (map ?f p) = f (target (initial M1) p)"
using \<open>f (initial M1) = initial M2\<close> assms(2)
by (induction p; auto)
then have "t_source (?f a) = target (initial M2) (map ?f p)"
by (metis (no_types, lifting) case_prod_beta' fst_conv path_append_transition_elim(3) snoc.prems)
have "t_source a \<in> reachable_states M1"
using \<open>path M1 (FSM.initial M1) (p @ [a])\<close>
by (metis path_append_transition_elim(3) path_prefix reachable_states_intro)
have "t_target a \<in> reachable_states M1"
using \<open>path M1 (FSM.initial M1) (p @ [a])\<close>
by (meson \<open>t_source a \<in> reachable_states M1\<close> path_append_transition_elim(2) reachable_states_next)
have "a \<in> transitions M1"
using snoc.prems by auto
then have "?f a \<in> transitions M2"
using assms(3)[OF \<open>t_source a \<in> reachable_states M1\<close> \<open>t_target a \<in> reachable_states M1\<close>]
by (metis (mono_tags, lifting) prod.case_eq_if prod.collapse)
have "map ?f (p@[a]) = (map ?f p)@[?f a]"
by auto
show ?case
unfolding \<open>map ?f (p@[a]) = (map ?f p)@[?f a]\<close>
using path_append_transition[OF \<open>path M2 (initial M2) (map ?f p)\<close> \<open>?f a \<in> transitions M2\<close> \<open>t_source (?f a) = target (initial M2) (map ?f p)\<close>]
by assumption
qed
moreover have "p_io ?p = io"
using \<open>p_io p = io\<close>
by (induction p; auto)
ultimately show "io \<in> L M2"
using language_state_containment by fastforce
qed
lemma language_equivalence_from_isomorphism_reachable :
assumes "bij_betw f (reachable_states M1) (reachable_states M2)"
and "f (initial M1) = initial M2"
and "\<And> q x y q' . q \<in> reachable_states M1 \<Longrightarrow> q' \<in> reachable_states M1 \<Longrightarrow> (q,x,y,q') \<in> transitions M1 \<longleftrightarrow> (f q,x,y,f q') \<in> transitions M2"
shows "L M1 = L M2"
proof
show "L M1 \<subseteq> L M2"
using language_equivalence_from_isomorphism_helper_reachable[OF assms] .
have "bij_betw (inv_into (reachable_states M1) f) (reachable_states M2) (reachable_states M1)"
using bij_betw_inv_into[OF assms(1)] .
moreover have "(inv_into (reachable_states M1) f) (initial M2) = (initial M1)"
using assms(1,2) reachable_states_initial
by (metis bij_betw_inv_into_left)
moreover have "\<And> q x y q' . q \<in> reachable_states M2 \<Longrightarrow> q' \<in> reachable_states M2 \<Longrightarrow> (q,x,y,q') \<in> transitions M2 \<longleftrightarrow> ((inv_into (reachable_states M1) f) q,x,y,(inv_into (reachable_states M1) f) q') \<in> transitions M1"
proof
fix q x y q' assume "q \<in> reachable_states M2" and "q' \<in> reachable_states M2"
show "(q,x,y,q') \<in> transitions M2 \<Longrightarrow> ((inv_into (reachable_states M1) f) q,x,y,(inv_into (reachable_states M1) f) q') \<in> transitions M1"
proof -
assume a1: "(q, x, y, q') \<in> FSM.transitions M2"
have f2: "\<forall>f B A. \<not> bij_betw f B A \<or> (\<forall>b. (b::'b) \<notin> B \<or> (f b::'a) \<in> A)"
using bij_betwE by blast
then have f3: "inv_into (FSM.reachable_states M1) f q \<in> FSM.reachable_states M1"
using \<open>q \<in> FSM.reachable_states M2\<close> calculation(1) by blast
have "inv_into (FSM.reachable_states M1) f q' \<in> FSM.reachable_states M1"
using f2 \<open>q' \<in> FSM.reachable_states M2\<close> calculation(1) by blast
then show ?thesis
using f3 a1 \<open>q \<in> FSM.reachable_states M2\<close> \<open>q' \<in> FSM.reachable_states M2\<close> assms(1) assms(3) bij_betw_inv_into_right by fastforce
qed
show "((inv_into (reachable_states M1) f) q,x,y,(inv_into (reachable_states M1) f) q') \<in> transitions M1 \<Longrightarrow> (q,x,y,q') \<in> transitions M2"
proof -
assume a1: "(inv_into (FSM.reachable_states M1) f q, x, y, inv_into (FSM.reachable_states M1) f q') \<in> FSM.transitions M1"
have f2: "\<forall>f B A. \<not> bij_betw f B A \<or> (\<forall>b. (b::'b) \<notin> B \<or> (f b::'a) \<in> A)"
by (metis (full_types) bij_betwE)
then have f3: "inv_into (FSM.reachable_states M1) f q' \<in> FSM.reachable_states M1"
using \<open>q' \<in> FSM.reachable_states M2\<close> calculation(1) by blast
have "inv_into (FSM.reachable_states M1) f q \<in> FSM.reachable_states M1"
using f2 \<open>q \<in> FSM.reachable_states M2\<close> calculation(1) by blast
then show ?thesis
using f3 a1 \<open>q \<in> FSM.reachable_states M2\<close> \<open>q' \<in> FSM.reachable_states M2\<close> assms(1) assms(3) bij_betw_inv_into_right by fastforce
qed
qed
ultimately show "L M2 \<subseteq> L M1"
using language_equivalence_from_isomorphism_helper_reachable[of "(inv_into (reachable_states M1) f)" M2 M1]
by blast
qed
lemma language_empty_io :
assumes "inputs M = {} \<or> outputs M = {}"
shows "L M = {[]}"
proof -
have "transitions M = {}"
using assms fsm_transition_input fsm_transition_output
by auto
then have "\<And> p . path M (initial M) p \<Longrightarrow> p = []"
by (metis empty_iff path.cases)
then show ?thesis
unfolding LS.simps
by blast
qed
subsection \<open>Basic FSM Properties\<close>
subsubsection \<open>Completely Specified\<close>
fun completely_specified :: "('a,'b,'c) fsm \<Rightarrow> bool" where
"completely_specified M = (\<forall> q \<in> states M . \<forall> x \<in> inputs M . \<exists> t \<in> transitions M . t_source t = q \<and> t_input t = x)"
lemma completely_specified_alt_def :
"completely_specified M = (\<forall> q \<in> states M . \<forall> x \<in> inputs M . \<exists> q' y . (q,x,y,q') \<in> transitions M)"
by force
lemma completely_specified_alt_def_h :
"completely_specified M = (\<forall> q \<in> states M . \<forall> x \<in> inputs M . h M (q,x) \<noteq> {})"
by force
fun completely_specified_state :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> bool" where
"completely_specified_state M q = (\<forall> x \<in> inputs M . \<exists> t \<in> transitions M . t_source t = q \<and> t_input t = x)"
lemma completely_specified_states :
"completely_specified M = (\<forall> q \<in> states M . completely_specified_state M q)"
unfolding completely_specified.simps completely_specified_state.simps by force
lemma completely_specified_state_alt_def_h :
"completely_specified_state M q = (\<forall> x \<in> inputs M . h M (q,x) \<noteq> {})"
by force
lemma completely_specified_path_extension :
assumes "completely_specified M"
and "q \<in> states M"
and "path M q p"
and "x \<in> (inputs M)"
obtains t where "t \<in> transitions M" and "t_input t = x" and "t_source t = target q p"
proof -
have "target q p \<in> states M"
using path_target_is_state \<open>path M q p\<close> by metis
then obtain t where "t \<in> transitions M" and "t_input t = x" and "t_source t = target q p"
using \<open>completely_specified M\<close> \<open>x \<in> (inputs M)\<close>
unfolding completely_specified.simps by blast
then show ?thesis using that by blast
qed
lemma completely_specified_language_extension :
assumes "completely_specified M"
and "q \<in> states M"
and "io \<in> LS M q"
and "x \<in> (inputs M)"
obtains y where "io@[(x,y)] \<in> LS M q"
proof -
obtain p where "path M q p" and "p_io p = io"
using \<open>io \<in> LS M q\<close> by auto
moreover obtain t where "t \<in> transitions M" and "t_input t = x" and "t_source t = target q p"
using completely_specified_path_extension[OF assms(1,2) \<open>path M q p\<close> assms(4)] by blast
ultimately have "path M q (p@[t])" and "p_io (p@[t]) = io@[(x,t_output t)]"
by (simp add: path_append_transition)+
then have "io@[(x,t_output t)] \<in> LS M q"
using language_state_containment[of M q "p@[t]" "io@[(x,t_output t)]"] by auto
then show ?thesis using that by blast
qed
lemma path_of_length_ex :
assumes "completely_specified M"
and "q \<in> states M"
and "inputs M \<noteq> {}"
shows "\<exists> p . path M q p \<and> length p = k"
using assms(2) proof (induction k arbitrary: q)
case 0
then show ?case by auto
next
case (Suc k)
obtain t where "t_source t = q" and "t \<in> transitions M"
by (meson Suc.prems assms(1) assms(3) completely_specified.simps equals0I)
then have "t_target t \<in> states M"
using fsm_transition_target by blast
then obtain p where "path M (t_target t) p \<and> length p = k"
using Suc.IH by blast
then show ?case
using \<open>t_source t = q\<close> \<open>t \<in> transitions M\<close>
by auto
qed
subsubsection \<open>Deterministic\<close>
fun deterministic :: "('a,'b,'c) fsm \<Rightarrow> bool" where
"deterministic M = (\<forall> t1 \<in> transitions M .
\<forall> t2 \<in> transitions M .
(t_source t1 = t_source t2 \<and> t_input t1 = t_input t2)
\<longrightarrow> (t_output t1 = t_output t2 \<and> t_target t1 = t_target t2))"
lemma deterministic_alt_def :
"deterministic M = (\<forall> q1 x y' y'' q1' q1'' . (q1,x,y',q1') \<in> transitions M \<and> (q1,x,y'',q1'') \<in> transitions M \<longrightarrow> y' = y'' \<and> q1' = q1'')"
by auto
lemma deterministic_alt_def_h :
"deterministic M = (\<forall> q1 x yq yq' . (yq \<in> h M (q1,x) \<and> yq' \<in> h M (q1,x)) \<longrightarrow> yq = yq')"
by auto
subsubsection \<open>Observable\<close>
fun observable :: "('a,'b,'c) fsm \<Rightarrow> bool" where
"observable M = (\<forall> t1 \<in> transitions M .
\<forall> t2 \<in> transitions M .
(t_source t1 = t_source t2 \<and> t_input t1 = t_input t2 \<and> t_output t1 = t_output t2)
\<longrightarrow> t_target t1 = t_target t2)"
lemma observable_alt_def :
"observable M = (\<forall> q1 x y q1' q1'' . (q1,x,y,q1') \<in> transitions M \<and> (q1,x,y,q1'') \<in> transitions M \<longrightarrow> q1' = q1'')"
by auto
lemma observable_alt_def_h :
"observable M = (\<forall> q1 x yq yq' . (yq \<in> h M (q1,x) \<and> yq' \<in> h M (q1,x)) \<longrightarrow> fst yq = fst yq' \<longrightarrow> snd yq = snd yq')"
by auto
lemma language_append_path_ob :
assumes "io@[(x,y)] \<in> L M"
obtains p t where "path M (initial M) (p@[t])" and "p_io p = io" and "t_input t = x" and "t_output t = y"
proof -
obtain p p2 where "path M (initial M) p" and "path M (target (initial M) p) p2" and "p_io p = io" and "p_io p2 = [(x,y)]"
using language_state_split[OF assms] by blast
obtain t where "p2 = [t]" and "t_input t = x" and "t_output t = y"
using \<open>p_io p2 = [(x,y)]\<close> by auto
have "path M (initial M) (p@[t])"
using \<open>path M (initial M) p\<close> \<open>path M (target (initial M) p) p2\<close> unfolding \<open>p2 = [t]\<close> by auto
then show ?thesis using that[OF _ \<open>p_io p = io\<close> \<open>t_input t = x\<close> \<open>t_output t = y\<close>]
by simp
qed
subsubsection \<open>Single Input\<close>
(* each state has at most one input, but may have none *)
fun single_input :: "('a,'b,'c) fsm \<Rightarrow> bool" where
"single_input M = (\<forall> t1 \<in> transitions M .
\<forall> t2 \<in> transitions M .
t_source t1 = t_source t2 \<longrightarrow> t_input t1 = t_input t2)"
lemma single_input_alt_def :
"single_input M = (\<forall> q1 x x' y y' q1' q1'' . (q1,x,y,q1') \<in> transitions M \<and> (q1,x',y',q1'') \<in> transitions M \<longrightarrow> x = x')"
by fastforce
lemma single_input_alt_def_h :
"single_input M = (\<forall> q x x' . (h M (q,x) \<noteq> {} \<and> h M (q,x') \<noteq> {}) \<longrightarrow> x = x')"
by force
subsubsection \<open>Output Complete\<close>
fun output_complete :: "('a,'b,'c) fsm \<Rightarrow> bool" where
"output_complete M = (\<forall> t \<in> transitions M .
\<forall> y \<in> outputs M .
\<exists> t' \<in> transitions M . t_source t = t_source t' \<and>
t_input t = t_input t' \<and>
t_output t' = y)"
lemma output_complete_alt_def :
"output_complete M = (\<forall> q x . (\<exists> y q' . (q,x,y,q') \<in> transitions M) \<longrightarrow> (\<forall> y \<in> (outputs M) . \<exists> q' . (q,x,y,q') \<in> transitions M))"
by force
lemma output_complete_alt_def_h :
"output_complete M = (\<forall> q x . h M (q,x) \<noteq> {} \<longrightarrow> (\<forall> y \<in> outputs M . \<exists> q' . (y,q') \<in> h M (q,x)))"
by force
subsubsection \<open>Acyclic\<close>
fun acyclic :: "('a,'b,'c) fsm \<Rightarrow> bool" where
"acyclic M = (\<forall> p . path M (initial M) p \<longrightarrow> distinct (visited_states (initial M) p))"
lemma visited_states_length : "length (visited_states q p) = Suc (length p)" by auto
lemma visited_states_take :
"(take (Suc n) (visited_states q p)) = (visited_states q (take n p))"
proof (induction p rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc x xs)
then show ?case by (cases "n \<le> length xs"; auto)
qed
(* very inefficient calculation *)
lemma acyclic_code[code] :
"acyclic M = (\<not>(\<exists> p \<in> (acyclic_paths_up_to_length M (initial M) (size M - 1)) .
\<exists> t \<in> transitions M . t_source t = target (initial M) p \<and>
t_target t \<in> set (visited_states (initial M) p)))"
proof -
have "(\<exists> p \<in> (acyclic_paths_up_to_length M (initial M) (size M - 1)) .
\<exists> t \<in> transitions M . t_source t = target (initial M) p \<and>
t_target t \<in> set (visited_states (initial M) p))
\<Longrightarrow> \<not> FSM.acyclic M"
proof -
assume "(\<exists> p \<in> (acyclic_paths_up_to_length M (initial M) (size M - 1)) .
\<exists> t \<in> transitions M . t_source t = target (initial M) p \<and>
t_target t \<in> set (visited_states (initial M) p))"
then obtain p t where "path M (initial M) p"
and "distinct (visited_states (initial M) p)"
and "t \<in> transitions M"
and "t_source t = target (initial M) p"
and "t_target t \<in> set (visited_states (initial M) p)"
unfolding acyclic_paths_set by blast
then have "path M (initial M) (p@[t])"
by (simp add: path_append_transition)
moreover have "\<not> (distinct (visited_states (initial M) (p@[t])))"
using \<open>t_target t \<in> set (visited_states (initial M) p)\<close> by auto
ultimately show "\<not> FSM.acyclic M"
by (meson acyclic.elims(2))
qed
moreover have "\<not> FSM.acyclic M \<Longrightarrow>
(\<exists> p \<in> (acyclic_paths_up_to_length M (initial M) (size M - 1)) .
\<exists> t \<in> transitions M . t_source t = target (initial M) p \<and>
t_target t \<in> set (visited_states (initial M) p))"
proof -
assume "\<not> FSM.acyclic M"
then obtain p where "path M (initial M) p"
and "\<not> distinct (visited_states (initial M) p)"
by auto
then obtain n where "distinct (take (Suc n) (visited_states (initial M) p))"
and "\<not> distinct (take (Suc (Suc n)) (visited_states (initial M) p))"
using maximal_distinct_prefix by blast
then have "distinct (visited_states (initial M) (take n p))"
and "\<not> distinct (visited_states (initial M)(take (Suc n) p))"
unfolding visited_states_take by simp+
then obtain p' t' where *: "take n p = p'"
and **: "take (Suc n) p = p' @ [t']"
by (metis Suc_less_eq \<open>\<not> distinct (visited_states (FSM.initial M) p)\<close>
le_imp_less_Suc not_less_eq_eq take_all take_hd_drop)
have ***: "visited_states (FSM.initial M) (p' @ [t']) = (visited_states (FSM.initial M) p')@[t_target t']"
by auto
have "path M (initial M) p'"
using * \<open>path M (initial M) p\<close>
by (metis append_take_drop_id path_prefix)
then have "p' \<in> (acyclic_paths_up_to_length M (initial M) (size M - 1))"
using \<open>distinct (visited_states (initial M) (take n p))\<close>
unfolding * acyclic_paths_set by blast
moreover have "t' \<in> transitions M \<and> t_source t' = target (initial M) p'"
using * ** \<open>path M (initial M) p\<close>
by (metis append_take_drop_id path_append_elim path_cons_elim)
moreover have "t_target t' \<in> set (visited_states (initial M) p')"
using \<open>distinct (visited_states (initial M) (take n p))\<close>
\<open>\<not> distinct (visited_states (initial M)(take (Suc n) p))\<close>
unfolding * ** *** by auto
ultimately show "(\<exists> p \<in> (acyclic_paths_up_to_length M (initial M) (size M - 1)) .
\<exists> t \<in> transitions M . t_source t = target (initial M) p \<and>
t_target t \<in> set (visited_states (initial M) p))"
by blast
qed
ultimately show ?thesis by blast
qed
lemma acyclic_alt_def : "acyclic M = finite (L M)"
proof
show "acyclic M \<Longrightarrow> finite (L M)"
proof -
assume "acyclic M"
then have "{ p . path M (initial M) p} \<subseteq> (acyclic_paths_up_to_length M (initial M) (size M - 1))"
unfolding acyclic_paths_set by auto
moreover have "finite (acyclic_paths_up_to_length M (initial M) (size M - 1))"
unfolding acyclic_paths_up_to_length.simps using paths_finite[of M "initial M" "size M - 1"]
by (metis (mono_tags, lifting) Collect_cong \<open>FSM.acyclic M\<close> acyclic.elims(2))
ultimately have "finite { p . path M (initial M) p}"
using finite_subset by blast
then show "finite (L M)"
unfolding LS.simps by auto
qed
show "finite (L M) \<Longrightarrow> acyclic M"
proof (rule ccontr)
assume "finite (L M)"
assume "\<not> acyclic M"
obtain max_io_len where "\<forall>io \<in> L M . length io < max_io_len"
using finite_maxlen[OF \<open>finite (L M)\<close>] by blast
then have "\<And> p . path M (initial M) p \<Longrightarrow> length p < max_io_len"
proof -
fix p assume "path M (initial M) p"
show "length p < max_io_len"
proof (rule ccontr)
assume "\<not> length p < max_io_len"
then have "\<not> length (p_io p) < max_io_len" by auto
moreover have "p_io p \<in> L M"
unfolding LS.simps using \<open>path M (initial M) p\<close> by blast
ultimately show "False"
using \<open>\<forall>io \<in> L M . length io < max_io_len\<close> by blast
qed
qed
obtain p where "path M (initial M) p" and "\<not> distinct (visited_states (initial M) p)"
using \<open>\<not> acyclic M\<close> unfolding acyclic.simps by blast
then obtain pL where "path M (initial M) pL" and "max_io_len \<le> length pL"
using cyclic_path_pumping[of M p max_io_len] by blast
then show "False"
using \<open>\<And> p . path M (initial M) p \<Longrightarrow> length p < max_io_len\<close>
using not_le by blast
qed
qed
lemma acyclic_finite_paths_from_reachable_state :
assumes "acyclic M"
and "path M (initial M) p"
and "target (initial M) p = q"
shows "finite {p . path M q p}"
proof -
from assms have "{ p . path M (initial M) p} \<subseteq> (acyclic_paths_up_to_length M (initial M) (size M - 1))"
unfolding acyclic_paths_set by auto
moreover have "finite (acyclic_paths_up_to_length M (initial M) (size M - 1))"
unfolding acyclic_paths_up_to_length.simps using paths_finite[of M "initial M" "size M - 1"]
by (metis (mono_tags, lifting) Collect_cong \<open>FSM.acyclic M\<close> acyclic.elims(2))
ultimately have "finite { p . path M (initial M) p}"
using finite_subset by blast
show "finite {p . path M q p}"
proof (cases "q \<in> states M")
case True
have "image (\<lambda>p' . p@p') {p' . path M q p'} \<subseteq> {p' . path M (initial M) p'}"
proof
fix x assume "x \<in> image (\<lambda>p' . p@p') {p' . path M q p'}"
then obtain p' where "x = p@p'" and "p' \<in> {p' . path M q p'}"
by blast
then have "path M q p'" by auto
then have "path M (initial M) (p@p')"
using path_append[OF \<open>path M (initial M) p\<close>] \<open>target (initial M) p = q\<close> by auto
then show "x \<in> {p' . path M (initial M) p'}" using \<open>x = p@p'\<close> by blast
qed
then have "finite (image (\<lambda>p' . p@p') {p' . path M q p'})"
using \<open>finite { p . path M (initial M) p}\<close> finite_subset by auto
show ?thesis using finite_imageD[OF \<open>finite (image (\<lambda>p' . p@p') {p' . path M q p'})\<close>]
by (meson inj_onI same_append_eq)
next
case False
then show ?thesis
by (meson not_finite_existsD path_begin_state)
qed
qed
lemma acyclic_paths_from_reachable_states :
assumes "acyclic M"
and "path M (initial M) p'"
and "target (initial M) p' = q"
and "path M q p"
shows "distinct (visited_states q p)"
proof -
have "path M (initial M) (p'@p)"
using assms(2,3,4) path_append by metis
then have "distinct (visited_states (initial M) (p'@p))"
using assms(1) unfolding acyclic.simps by blast
then have "distinct (initial M # (map t_target p') @ map t_target p)"
by auto
moreover have "initial M # (map t_target p') @ map t_target p
= (butlast (initial M # map t_target p')) @ ((last (initial M # map t_target p')) # map t_target p)"
by auto
ultimately have "distinct ((last (initial M # map t_target p')) # map t_target p)"
by auto
then show ?thesis
using \<open>target (initial M) p' = q\<close> unfolding visited_states.simps target.simps by simp
qed
definition LS_acyclic :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('b \<times> 'c) list set" where
"LS_acyclic M q = {p_io p | p . path M q p \<and> distinct (visited_states q p)}"
lemma LS_acyclic_code[code] :
"LS_acyclic M q = image p_io (acyclic_paths_up_to_length M q (size M - 1))"
unfolding acyclic_paths_set LS_acyclic_def by blast
lemma LS_from_LS_acyclic :
assumes "acyclic M"
shows "L M = LS_acyclic M (initial M)"
proof -
obtain pps :: "(('b \<times> 'c) list \<Rightarrow> bool) \<Rightarrow> (('b \<times> 'c) list \<Rightarrow> bool) \<Rightarrow> ('b \<times> 'c) list" where
f1: "\<forall>p pa. (\<not> p (pps pa p)) = pa (pps pa p) \<or> Collect p = Collect pa"
by (metis (no_types) Collect_cong)
have "\<forall>ps. \<not> path M (FSM.initial M) ps \<or> distinct (visited_states (FSM.initial M) ps)"
using acyclic.simps assms by blast
then have "(\<nexists>ps. pps (\<lambda>ps. \<exists>psa. ps = p_io psa \<and> path M (FSM.initial M) psa)
(\<lambda>ps. \<exists>psa. ps = p_io psa \<and> path M (FSM.initial M) psa
\<and> distinct (visited_states (FSM.initial M) psa))
= p_io ps \<and> path M (FSM.initial M) ps \<and> distinct (visited_states (FSM.initial M) ps))
\<noteq> (\<exists>ps. pps (\<lambda>ps. \<exists>psa. ps = p_io psa \<and> path M (FSM.initial M) psa)
(\<lambda>ps. \<exists>psa. ps = p_io psa \<and> path M (FSM.initial M) psa
\<and> distinct (visited_states (FSM.initial M) psa))
= p_io ps \<and> path M (FSM.initial M) ps)"
by blast
then have "{p_io ps |ps. path M (FSM.initial M) ps \<and> distinct (visited_states (FSM.initial M) ps)}
= {p_io ps |ps. path M (FSM.initial M) ps}"
using f1
by (meson \<open>\<forall>ps. \<not> path M (FSM.initial M) ps \<or> distinct (visited_states (FSM.initial M) ps)\<close>)
then show ?thesis
by (simp add: LS_acyclic_def)
qed
lemma cyclic_cycle :
assumes "\<not> acyclic M"
shows "\<exists> q p . path M q p \<and> p \<noteq> [] \<and> target q p = q"
proof -
from \<open>\<not> acyclic M\<close> obtain p t where "path M (initial M) (p@[t])"
and "\<not>distinct (visited_states (initial M) (p@[t]))"
by (metis (no_types, opaque_lifting) Nil_is_append_conv acyclic.simps append_take_drop_id
maximal_distinct_prefix rev_exhaust visited_states_take)
show ?thesis
proof (cases "initial M \<in> set (map t_target (p@[t]))")
case True
then obtain i where "last (take i (map t_target (p@[t]))) = initial M"
and "i \<le> length (map t_target (p@[t]))" and "0 < i"
using list_contains_last_take by metis
let ?p = "take i (p@[t])"
have "path M (initial M) (?p@(drop i (p@[t])))"
using \<open>path M (initial M) (p@[t])\<close>
by (metis append_take_drop_id)
then have "path M (initial M) ?p" by auto
moreover have "?p \<noteq> []" using \<open>0 < i\<close> by auto
moreover have "target (initial M) ?p = initial M"
using \<open>last (take i (map t_target (p@[t]))) = initial M\<close>
unfolding target.simps visited_states.simps
by (metis (no_types, lifting) calculation(2) last_ConsR list.map_disc_iff take_map)
ultimately show ?thesis by blast
next
case False
then have "\<not> distinct (map t_target (p@[t]))"
using \<open>\<not>distinct (visited_states (initial M) (p@[t]))\<close>
unfolding visited_states.simps
by auto
then obtain i j where "i < j" and "j < length (map t_target (p@[t]))"
and "(map t_target (p@[t])) ! i = (map t_target (p@[t])) ! j"
using non_distinct_repetition_indices by blast
let ?pre_i = "take (Suc i) (p@[t])"
let ?p = "take ((Suc j)-(Suc i)) (drop (Suc i) (p@[t]))"
let ?post_j = "drop ((Suc j)-(Suc i)) (drop (Suc i) (p@[t]))"
have "p@[t] = ?pre_i @ ?p @ ?post_j"
using \<open>i < j\<close> \<open>j < length (map t_target (p@[t]))\<close>
by (metis append_take_drop_id)
then have "path M (target (initial M) ?pre_i) ?p"
using \<open>path M (initial M) (p@[t])\<close>
by (metis path_prefix path_suffix)
have "?p \<noteq> []"
using \<open>i < j\<close> \<open>j < length (map t_target (p@[t]))\<close> by auto
have "i < length (map t_target (p@[t]))"
using \<open>i < j\<close> \<open>j < length (map t_target (p@[t]))\<close> by auto
have "(target (initial M) ?pre_i) = (map t_target (p@[t])) ! i"
unfolding target.simps visited_states.simps
using take_last_index[OF \<open>i < length (map t_target (p@[t]))\<close>]
by (metis (no_types, lifting) \<open>i < length (map t_target (p @ [t]))\<close>
last_ConsR snoc_eq_iff_butlast take_Suc_conv_app_nth take_map)
have "?pre_i@?p = take (Suc j) (p@[t])"
by (metis (no_types) \<open>i < j\<close> add_Suc add_diff_cancel_left' less_SucI less_imp_Suc_add take_add)
moreover have "(target (initial M) (take (Suc j) (p@[t]))) = (map t_target (p@[t])) ! j"
unfolding target.simps visited_states.simps
using take_last_index[OF \<open>j < length (map t_target (p@[t]))\<close>]
by (metis (no_types, lifting) \<open>j < length (map t_target (p @ [t]))\<close>
last_ConsR snoc_eq_iff_butlast take_Suc_conv_app_nth take_map)
ultimately have "(target (initial M) (?pre_i@?p)) = (map t_target (p@[t])) ! j"
by auto
then have "(target (initial M) (?pre_i@?p)) = (map t_target (p@[t])) ! i"
using \<open>(map t_target (p@[t])) ! i = (map t_target (p@[t])) ! j\<close> by simp
moreover have "(target (initial M) (?pre_i@?p)) = (target (target (initial M) ?pre_i) ?p)"
unfolding target.simps visited_states.simps last.simps by auto
ultimately have "(target (target (initial M) ?pre_i) ?p) = (map t_target (p@[t])) ! i"
by auto
then have "(target (target (initial M) ?pre_i) ?p) = (target (initial M) ?pre_i)"
using \<open>(target (initial M) ?pre_i) = (map t_target (p@[t])) ! i\<close> by auto
show ?thesis
using \<open>path M (target (initial M) ?pre_i) ?p\<close> \<open>?p \<noteq> []\<close>
\<open>(target (target (initial M) ?pre_i) ?p) = (target (initial M) ?pre_i)\<close>
by blast
qed
qed
lemma cyclic_cycle_rev :
fixes M :: "('a,'b,'c) fsm"
assumes "path M (initial M) p'"
and "target (initial M) p' = q"
and "path M q p"
and "p \<noteq> []"
and "target q p = q"
shows "\<not> acyclic M"
using assms unfolding acyclic.simps target.simps visited_states.simps
using distinct.simps(2) by fastforce
lemma acyclic_initial :
assumes "acyclic M"
shows "\<not> (\<exists> t \<in> transitions M . t_target t = initial M \<and>
(\<exists> p . path M (initial M) p \<and> target (initial M) p = t_source t))"
by (metis append_Cons assms cyclic_cycle_rev list.distinct(1) path.simps
path_append path_append_transition_elim(3) single_transition_path)
lemma cyclic_path_shift :
assumes "path M q p"
and "target q p = q"
shows "path M (target q (take i p)) ((drop i p) @ (take i p))"
and "target (target q (take i p)) ((drop i p) @ (take i p)) = (target q (take i p))"
proof -
show "path M (target q (take i p)) ((drop i p) @ (take i p))"
by (metis append_take_drop_id assms(1) assms(2) path_append path_append_elim path_append_target)
show "target (target q (take i p)) ((drop i p) @ (take i p)) = (target q (take i p))"
by (metis append_take_drop_id assms(2) path_append_target)
qed
lemma cyclic_path_transition_states_property :
assumes "\<exists> t \<in> set p . P (t_source t)"
and "\<forall> t \<in> set p . P (t_source t) \<longrightarrow> P (t_target t)"
and "path M q p"
and "target q p = q"
shows "\<forall> t \<in> set p . P (t_source t)"
and "\<forall> t \<in> set p . P (t_target t)"
proof -
obtain t0 where "t0 \<in> set p" and "P (t_source t0)"
using assms(1) by blast
then obtain i where "i < length p" and "p ! i = t0"
by (meson in_set_conv_nth)
let ?p = "(drop i p @ take i p)"
have "path M (target q (take i p)) ?p"
using cyclic_path_shift(1)[OF assms(3,4), of i] by assumption
have "set ?p = set p"
proof -
have "set ?p = set (take i p @ drop i p)"
using list_set_sym by metis
then show ?thesis by auto
qed
then have "\<And> t . t \<in> set ?p \<Longrightarrow> P (t_source t) \<Longrightarrow> P (t_target t)"
using assms(2) by blast
have "\<And> j . j < length ?p \<Longrightarrow> P (t_source (?p ! j))"
proof -
fix j assume "j < length ?p"
then show "P (t_source (?p ! j))"
proof (induction j)
case 0
then show ?case
using \<open>p ! i = t0\<close> \<open>P (t_source t0)\<close>
by (metis \<open>i < length p\<close> drop_eq_Nil hd_append2 hd_conv_nth hd_drop_conv_nth leD
length_greater_0_conv)
next
case (Suc j)
then have "P (t_source (?p ! j))"
by auto
then have "P (t_target (?p ! j))"
using Suc.prems \<open>\<And> t . t \<in> set ?p \<Longrightarrow> P (t_source t) \<Longrightarrow> P (t_target t)\<close>[of "?p ! j"]
using Suc_lessD nth_mem by blast
moreover have "t_target (?p ! j) = t_source (?p ! (Suc j))"
using path_source_target_index[OF Suc.prems \<open>path M (target q (take i p)) ?p\<close>]
by assumption
ultimately show ?case
using \<open>\<And> t . t \<in> set ?p \<Longrightarrow> P (t_source t) \<Longrightarrow> P (t_target t)\<close>[of "?p ! j"]
by simp
qed
qed
then have "\<forall> t \<in> set ?p . P (t_source t)"
by (metis in_set_conv_nth)
then show "\<forall> t \<in> set p . P (t_source t)"
using \<open>set ?p = set p\<close> by blast
then show "\<forall> t \<in> set p . P (t_target t)"
using assms(2) by blast
qed
lemma cycle_incoming_transition_ex :
assumes "path M q p"
and "p \<noteq> []"
and "target q p = q"
and "t \<in> set p"
shows "\<exists> tI \<in> set p . t_target tI = t_source t"
proof -
obtain i where "i < length p" and "p ! i = t"
using assms(4) by (meson in_set_conv_nth)
let ?p = "(drop i p @ take i p)"
have "path M (target q (take i p)) ?p"
and "target (target q (take i p)) ?p = target q (take i p)"
using cyclic_path_shift[OF assms(1,3), of i] by linarith+
have "p = (take i p @ drop i p)" by auto
then have "path M (target q (take i p)) (drop i p)"
using path_suffix assms(1) by metis
moreover have "t = hd (drop i p)"
using \<open>i < length p\<close> \<open>p ! i = t\<close>
by (simp add: hd_drop_conv_nth)
ultimately have "path M (target q (take i p)) [t]"
by (metis \<open>i < length p\<close> append_take_drop_id assms(1) path_append_elim take_hd_drop)
then have "t_source t = (target q (take i p))"
by auto
moreover have "t_target (last ?p) = (target q (take i p))"
using \<open>path M (target q (take i p)) ?p\<close> \<open>target (target q (take i p)) ?p = target q (take i p)\<close>
assms(2)
unfolding target.simps visited_states.simps last.simps
by (metis (no_types, lifting) \<open>p = take i p @ drop i p\<close> append_is_Nil_conv last_map
list.map_disc_iff)
moreover have "set ?p = set p"
proof -
have "set ?p = set (take i p @ drop i p)"
using list_set_sym by metis
then show ?thesis by auto
qed
ultimately show ?thesis
by (metis \<open>i < length p\<close> append_is_Nil_conv drop_eq_Nil last_in_set leD)
qed
lemma acyclic_paths_finite :
"finite {p . path M q p \<and> distinct (visited_states q p) }"
proof -
have "\<And> p . path M q p \<Longrightarrow> distinct (visited_states q p) \<Longrightarrow> distinct p"
proof -
fix p assume "path M q p" and "distinct (visited_states q p)"
then have "distinct (map t_target p)" by auto
then show "distinct p" by (simp add: distinct_map)
qed
then show ?thesis
using distinct_lists_finite[OF fsm_transitions_finite, of M] path_transitions[of M q]
by (metis (no_types, lifting) infinite_super mem_Collect_eq path_transitions subsetI)
qed
lemma acyclic_no_self_loop :
assumes "acyclic M"
and "q \<in> reachable_states M"
shows "\<not> (\<exists> x y . (q,x,y,q) \<in> transitions M)"
proof
assume "\<exists>x y. (q, x, y, q) \<in> FSM.transitions M"
then obtain x y where "(q, x, y, q) \<in> FSM.transitions M" by blast
moreover obtain p where "path M (initial M) p" and "target (initial M) p = q"
using assms(2) unfolding reachable_states_def by blast
ultimately have "path M (initial M) (p@[(q,x,y,q)])"
by (simp add: path_append_transition)
moreover have "\<not> (distinct (visited_states (initial M) (p@[(q,x,y,q)])))"
using \<open>target (initial M) p = q\<close> unfolding visited_states.simps target.simps by (cases p rule: rev_cases; auto)
ultimately show "False"
using assms(1) unfolding acyclic.simps
by meson
qed
subsubsection \<open>Deadlock States\<close>
fun deadlock_state :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> bool" where
"deadlock_state M q = (\<not>(\<exists> t \<in> transitions M . t_source t = q))"
lemma deadlock_state_alt_def : "deadlock_state M q = (LS M q \<subseteq> {[]})"
proof
show "deadlock_state M q \<Longrightarrow> LS M q \<subseteq> {[]}"
proof -
assume "deadlock_state M q"
moreover have "\<And> p . deadlock_state M q \<Longrightarrow> path M q p \<Longrightarrow> p = []"
unfolding deadlock_state.simps by (metis path.cases)
ultimately show "LS M q \<subseteq> {[]}"
unfolding LS.simps by blast
qed
show "LS M q \<subseteq> {[]} \<Longrightarrow> deadlock_state M q"
unfolding LS.simps deadlock_state.simps using path.cases[of M q] by blast
qed
lemma deadlock_state_alt_def_h : "deadlock_state M q = (\<forall> x \<in> inputs M . h M (q,x) = {})"
unfolding deadlock_state.simps h.simps
using fsm_transition_input by force
lemma acyclic_deadlock_reachable :
assumes "acyclic M"
shows "\<exists> q \<in> reachable_states M . deadlock_state M q"
proof (rule ccontr)
assume "\<not> (\<exists>q\<in>reachable_states M. deadlock_state M q)"
then have *: "\<And> q . q \<in> reachable_states M \<Longrightarrow> (\<exists> t \<in> transitions M . t_source t = q)"
unfolding deadlock_state.simps by blast
let ?p = "arg_max_on length {p. path M (initial M) p}"
have "finite {p. path M (initial M) p}"
by (metis Collect_cong acyclic_finite_paths_from_reachable_state assms eq_Nil_appendI fsm_initial
nil path_append path_append_elim)
moreover have "{p. path M (initial M) p} \<noteq> {}"
by auto
ultimately obtain p where "path M (initial M) p"
and "\<And> p' . path M (initial M) p' \<Longrightarrow> length p' \<le> length p"
using max_length_elem
by (metis mem_Collect_eq not_le_imp_less)
then obtain t where "t \<in> transitions M" and "t_source t = target (initial M) p"
using *[of "target (initial M) p"] unfolding reachable_states_def
by blast
then have "path M (initial M) (p@[t])"
using \<open>path M (initial M) p\<close>
by (simp add: path_append_transition)
then show "False"
using \<open>\<And> p' . path M (initial M) p' \<Longrightarrow> length p' \<le> length p\<close>
by (metis impossible_Cons length_rotate1 rotate1.simps(2))
qed
lemma deadlock_prefix :
assumes "path M q p"
and "t \<in> set (butlast p)"
shows "\<not> (deadlock_state M (t_target t))"
using assms proof (induction p rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc t' p')
show ?case proof (cases "t \<in> set (butlast p')")
case True
show ?thesis
using snoc.IH[OF _ True] snoc.prems(1)
by blast
next
case False
then have "p' = (butlast p')@[t]"
using snoc.prems(2) by (metis append_butlast_last_id append_self_conv2 butlast_snoc
in_set_butlast_appendI list_prefix_elem set_ConsD)
then have "path M q ((butlast p'@[t])@[t'])"
using snoc.prems(1)
by auto
have "t' \<in> transitions M" and "t_source t' = target q (butlast p'@[t])"
using path_suffix[OF \<open>path M q ((butlast p'@[t])@[t'])\<close>]
by auto
then have "t' \<in> transitions M \<and> t_source t' = t_target t"
unfolding target.simps visited_states.simps by auto
then show ?thesis
unfolding deadlock_state.simps using \<open>t' \<in> transitions M\<close> by blast
qed
qed
lemma states_initial_deadlock :
assumes "deadlock_state M (initial M)"
shows "reachable_states M = {initial M}"
proof -
have "\<And> q . q \<in> reachable_states M \<Longrightarrow> q = initial M"
proof -
fix q assume "q \<in> reachable_states M"
then obtain p where "path M (initial M) p" and "target (initial M) p = q"
unfolding reachable_states_def by auto
show "q = initial M" proof (cases p)
case Nil
then show ?thesis using \<open>target (initial M) p = q\<close> by auto
next
case (Cons t p')
then have "False" using assms \<open>path M (initial M) p\<close> unfolding deadlock_state.simps
by auto
then show ?thesis by simp
qed
qed
then show ?thesis
using reachable_states_initial[of M] by blast
qed
subsubsection \<open>Other\<close>
fun completed_path :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('a,'b,'c) path \<Rightarrow> bool" where
"completed_path M q p = deadlock_state M (target q p)"
fun minimal :: "('a,'b,'c) fsm \<Rightarrow> bool" where
"minimal M = (\<forall> q \<in> states M . \<forall> q' \<in> states M . q \<noteq> q' \<longrightarrow> LS M q \<noteq> LS M q')"
lemma minimal_alt_def : "minimal M = (\<forall> q q' . q \<in> states M \<longrightarrow> q' \<in> states M \<longrightarrow> LS M q = LS M q' \<longrightarrow> q = q')"
by auto
definition retains_outputs_for_states_and_inputs :: "('a,'b,'c) fsm \<Rightarrow> ('a,'b,'c) fsm \<Rightarrow> bool" where
"retains_outputs_for_states_and_inputs M S
= (\<forall> tS \<in> transitions S .
\<forall> tM \<in> transitions M .
(t_source tS = t_source tM \<and> t_input tS = t_input tM) \<longrightarrow> tM \<in> transitions S)"
subsection \<open>IO Targets and Observability\<close>
fun paths_for_io' :: "(('a \<times> 'b) \<Rightarrow> ('c \<times> 'a) set) \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> 'a \<Rightarrow> ('a,'b,'c) path \<Rightarrow> ('a,'b,'c) path set" where
"paths_for_io' f [] q prev = {prev}" |
"paths_for_io' f ((x,y)#io) q prev = \<Union>(image (\<lambda>yq' . paths_for_io' f io (snd yq') (prev@[(q,x,y,(snd yq'))])) (Set.filter (\<lambda>yq' . fst yq' = y) (f (q,x))))"
lemma paths_for_io'_set :
assumes "q \<in> states M"
shows "paths_for_io' (h M) io q prev = {prev@p | p . path M q p \<and> p_io p = io}"
using assms proof (induction io arbitrary: q prev)
case Nil
then show ?case by auto
next
case (Cons xy io)
obtain x y where "xy = (x,y)"
by (meson surj_pair)
let ?UN = "\<Union>(image (\<lambda>yq' . paths_for_io' (h M) io (snd yq') (prev@[(q,x,y,(snd yq'))]))
(Set.filter (\<lambda>yq' . fst yq' = y) (h M (q,x))))"
have "?UN = {prev@p | p . path M q p \<and> p_io p = (x,y)#io}"
proof
have "\<And> p . p \<in> ?UN \<Longrightarrow> p \<in> {prev@p | p . path M q p \<and> p_io p = (x,y)#io}"
proof -
fix p assume "p \<in> ?UN"
then obtain q' where "(y,q') \<in> (Set.filter (\<lambda>yq' . fst yq' = y) (h M (q,x)))"
and "p \<in> paths_for_io' (h M) io q' (prev@[(q,x,y,q')])"
by auto
from \<open>(y,q') \<in> (Set.filter (\<lambda>yq' . fst yq' = y) (h M (q,x)))\<close> have "q' \<in> states M"
and "(q,x,y,q') \<in> transitions M"
using fsm_transition_target unfolding h.simps by auto
have "p \<in> {(prev @ [(q, x, y, q')]) @ p |p. path M q' p \<and> p_io p = io}"
using \<open>p \<in> paths_for_io' (h M) io q' (prev@[(q,x,y,q')])\<close>
unfolding Cons.IH[OF \<open>q' \<in> states M\<close>] by assumption
moreover have "{(prev @ [(q, x, y, q')]) @ p |p. path M q' p \<and> p_io p = io}
\<subseteq> {prev@p | p . path M q p \<and> p_io p = (x,y)#io}"
using \<open>(q,x,y,q') \<in> transitions M\<close>
using cons by force
ultimately show "p \<in> {prev@p | p . path M q p \<and> p_io p = (x,y)#io}"
by blast
qed
then show "?UN \<subseteq> {prev@p | p . path M q p \<and> p_io p = (x,y)#io}"
by blast
have "\<And> p . p \<in> {prev@p | p . path M q p \<and> p_io p = (x,y)#io} \<Longrightarrow> p \<in> ?UN"
proof -
fix pp assume "pp \<in> {prev@p | p . path M q p \<and> p_io p = (x,y)#io}"
then obtain p where "pp = prev@p" and "path M q p" and "p_io p = (x,y)#io"
by fastforce
then obtain t p' where "p = t#p'" and "path M q (t#p')" and "p_io (t#p') = (x,y)#io"
and "p_io p' = io"
by (metis (no_types, lifting) map_eq_Cons_D)
then have "path M (t_target t) p'" and "t_source t = q" and "t_input t = x"
and "t_output t = y" and "t_target t \<in> states M"
and "t \<in> transitions M"
by auto
have "(y,t_target t) \<in> Set.filter (\<lambda>yq'. fst yq' = y) (h M (q, x))"
using \<open>t \<in> transitions M\<close> \<open>t_output t = y\<close> \<open>t_input t = x\<close> \<open>t_source t = q\<close>
unfolding h.simps
by auto
moreover have "(prev@p) \<in> paths_for_io' (h M) io (snd (y,t_target t)) (prev @ [(q, x, y, snd (y,t_target t))])"
using Cons.IH[OF \<open>t_target t \<in> states M\<close>, of "prev@[(q, x, y, t_target t)]"]
using \<open>p = t # p'\<close> \<open>p_io p' = io\<close> \<open>path M (t_target t) p'\<close> \<open>t_input t = x\<close>
\<open>t_output t = y\<close> \<open>t_source t = q\<close>
by auto
ultimately show "pp \<in> ?UN" unfolding \<open>pp = prev@p\<close>
by blast
qed
then show "{prev@p | p . path M q p \<and> p_io p = (x,y)#io} \<subseteq> ?UN"
by (meson subsetI)
qed
then show ?case
by (simp add: \<open>xy = (x, y)\<close>)
qed
definition paths_for_io :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> ('a,'b,'c) path set" where
"paths_for_io M q io = {p . path M q p \<and> p_io p = io}"
lemma paths_for_io_set_code[code] :
"paths_for_io M q io = (if q \<in> states M then paths_for_io' (h M) io q [] else {})"
using paths_for_io'_set[of q M io "[]"]
unfolding paths_for_io_def
proof -
have "{[] @ ps |ps. path M q ps \<and> p_io ps = io} = (if q \<in> FSM.states M then paths_for_io' (h M) io q [] else {})
\<longrightarrow> {ps. path M q ps \<and> p_io ps = io} = (if q \<in> FSM.states M then paths_for_io' (h M) io q [] else {})"
by auto
moreover
{ assume "{[] @ ps |ps. path M q ps \<and> p_io ps = io} \<noteq> (if q \<in> FSM.states M then paths_for_io' (h M) io q [] else {})"
then have "q \<notin> FSM.states M"
using \<open>q \<in> FSM.states M \<Longrightarrow> paths_for_io' (h M) io q [] = {[] @ p |p. path M q p \<and> p_io p = io}\<close> by force
then have "{ps. path M q ps \<and> p_io ps = io} = (if q \<in> FSM.states M then paths_for_io' (h M) io q [] else {})"
using path_begin_state by force }
ultimately show "{ps. path M q ps \<and> p_io ps = io} = (if q \<in> FSM.states M then paths_for_io' (h M) io q [] else {})"
by linarith
qed
fun io_targets :: "('a,'b,'c) fsm \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> 'a \<Rightarrow> 'a set" where
"io_targets M io q = {target q p | p . path M q p \<and> p_io p = io}"
lemma io_targets_code[code] : "io_targets M io q = image (target q) (paths_for_io M q io)"
unfolding io_targets.simps paths_for_io_def by blast
lemma io_targets_states :
"io_targets M io q \<subseteq> states M"
using path_target_is_state by fastforce
lemma observable_transition_unique :
assumes "observable M"
and "t \<in> transitions M"
shows "\<exists>! t' \<in> transitions M . t_source t' = t_source t \<and>
t_input t' = t_input t \<and>
t_output t' = t_output t"
by (metis assms observable.elims(2) prod.expand)
lemma observable_path_unique :
assumes "observable M"
and "path M q p"
and "path M q p'"
and "p_io p = p_io p'"
shows "p = p'"
proof -
have "length p = length p'"
using assms(4) map_eq_imp_length_eq by blast
then show ?thesis
using \<open>p_io p = p_io p'\<close> \<open>path M q p\<close> \<open>path M q p'\<close>
proof (induction p p' arbitrary: q rule: list_induct2)
case Nil
then show ?case by auto
next
case (Cons x xs y ys)
then have *: "x \<in> transitions M \<and> y \<in> transitions M \<and> t_source x = t_source y
\<and> t_input x = t_input y \<and> t_output x = t_output y"
by auto
then have "t_target x = t_target y"
using assms(1) observable.elims(2) by blast
then have "x = y"
by (simp add: "*" prod.expand)
have "p_io xs = p_io ys"
using Cons by auto
moreover have "path M (t_target x) xs"
using Cons by auto
moreover have "path M (t_target x) ys"
using Cons \<open>t_target x = t_target y\<close> by auto
ultimately have "xs = ys"
using Cons by auto
then show ?case
using \<open>x = y\<close> by simp
qed
qed
lemma observable_io_targets :
assumes "observable M"
and "io \<in> LS M q"
obtains q'
where "io_targets M io q = {q'}"
proof -
obtain p where "path M q p" and "p_io p = io"
using assms(2) by auto
then have "target q p \<in> io_targets M io q"
by auto
have "\<exists> q' . io_targets M io q = {q'}"
proof (rule ccontr)
assume "\<not>(\<exists>q'. io_targets M io q = {q'})"
then have "\<exists> q' . q' \<noteq> target q p \<and> q' \<in> io_targets M io q"
proof -
have "\<not> io_targets M io q \<subseteq> {target q p}"
using \<open>\<not>(\<exists>q'. io_targets M io q = {q'})\<close> \<open>target q p \<in> io_targets M io q\<close> by blast
then show ?thesis
by blast
qed
then obtain q' where "q' \<noteq> target q p" and "q' \<in> io_targets M io q"
by blast
then obtain p' where "path M q p'" and "target q p' = q'" and "p_io p' = io"
by auto
then have "p_io p = p_io p'"
using \<open>p_io p = io\<close> by simp
then have "p = p'"
using observable_path_unique[OF assms(1) \<open>path M q p\<close> \<open>path M q p'\<close>] by simp
then show "False"
using \<open>q' \<noteq> target q p\<close> \<open>target q p' = q'\<close> by auto
qed
then show ?thesis using that by blast
qed
lemma observable_path_io_target :
assumes "observable M"
and "path M q p"
shows "io_targets M (p_io p) q = {target q p}"
using observable_io_targets[OF assms(1) language_state_containment[OF assms(2)], of "p_io p"]
singletonD[of "target q p"]
unfolding io_targets.simps
proof -
assume a1: "\<And>a. target q p \<in> {a} \<Longrightarrow> target q p = a"
assume "\<And>thesis. \<lbrakk>p_io p = p_io p; \<And>q'. {target q pa |pa. path M q pa \<and> p_io pa = p_io p} = {q'} \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis"
then obtain aa :: 'a where "\<And>b. {target q ps |ps. path M q ps \<and> p_io ps = p_io p} = {aa} \<or> b"
by meson
then show "{target q ps |ps. path M q ps \<and> p_io ps = p_io p} = {target q p}"
using a1 assms(2) by blast
qed
lemma completely_specified_io_targets :
assumes "completely_specified M"
shows "\<forall> q \<in> io_targets M io (initial M) . \<forall> x \<in> (inputs M) . \<exists> t \<in> transitions M . t_source t = q \<and> t_input t = x"
by (meson assms completely_specified.elims(2) io_targets_states subsetD)
lemma observable_path_language_step :
assumes "observable M"
and "path M q p"
and "\<not> (\<exists>t\<in>transitions M.
t_source t = target q p \<and>
t_input t = x \<and> t_output t = y)"
shows "(p_io p)@[(x,y)] \<notin> LS M q"
using assms proof (induction p rule: rev_induct)
case Nil
show ?case proof
assume "p_io [] @ [(x, y)] \<in> LS M q"
then obtain p' where "path M q p'" and "p_io p' = [(x,y)]" unfolding LS.simps
by force
then obtain t where "p' = [t]" by blast
have "t\<in>transitions M" and "t_source t = target q []"
using \<open>path M q p'\<close> \<open>p' = [t]\<close> by auto
moreover have "t_input t = x \<and> t_output t = y"
using \<open>p_io p' = [(x,y)]\<close> \<open>p' = [t]\<close> by auto
ultimately show "False"
using Nil.prems(3) by blast
qed
next
case (snoc t p)
from \<open>path M q (p @ [t])\<close> have "path M q p" and "t_source t = target q p"
and "t \<in> transitions M"
by auto
show ?case proof
assume "p_io (p @ [t]) @ [(x, y)] \<in> LS M q"
then obtain p' where "path M q p'" and "p_io p' = p_io (p @ [t]) @ [(x, y)]"
by auto
then obtain p'' t' t'' where "p' = p''@[t']@[t'']"
by (metis (no_types, lifting) append.assoc map_butlast map_is_Nil_conv snoc_eq_iff_butlast)
then have "path M q p''"
using \<open>path M q p'\<close> by blast
have "p_io p'' = p_io p"
using \<open>p' = p''@[t']@[t'']\<close> \<open>p_io p' = p_io (p @ [t]) @ [(x, y)]\<close> by auto
then have "p'' = p"
using observable_path_unique[OF assms(1) \<open>path M q p''\<close> \<open>path M q p\<close>] by blast
have "t_source t' = target q p''" and "t' \<in> transitions M"
using \<open>path M q p'\<close> \<open>p' = p''@[t']@[t'']\<close> by auto
then have "t_source t' = t_source t"
using \<open>p'' = p\<close> \<open>t_source t = target q p\<close> by auto
moreover have "t_input t' = t_input t \<and> t_output t' = t_output t"
using \<open>p_io p' = p_io (p @ [t]) @ [(x, y)]\<close> \<open>p' = p''@[t']@[t'']\<close> \<open>p'' = p\<close> by auto
ultimately have "t' = t"
using \<open>t \<in> transitions M\<close> \<open>t' \<in> transitions M\<close> assms(1) unfolding observable.simps
by (meson prod.expand)
have "t'' \<in> transitions M" and "t_source t'' = target q (p@[t])"
using \<open>path M q p'\<close> \<open>p' = p''@[t']@[t'']\<close> \<open>p'' = p\<close> \<open>t' = t\<close> by auto
moreover have "t_input t'' = x \<and> t_output t'' = y"
using \<open>p_io p' = p_io (p @ [t]) @ [(x, y)]\<close> \<open>p' = p''@[t']@[t'']\<close> by auto
ultimately show "False"
using snoc.prems(3) by blast
qed
qed
lemma observable_io_targets_language :
assumes "io1 @ io2 \<in> LS M q1"
and "observable M"
and "q2 \<in> io_targets M io1 q1"
shows "io2 \<in> LS M q2"
proof -
obtain p1 p2 where "path M q1 p1" and "path M (target q1 p1) p2"
and "p_io p1 = io1" and "p_io p2 = io2"
using language_state_split[OF assms(1)] by blast
then have "io1 \<in> LS M q1" and "io2 \<in> LS M (target q1 p1)"
by auto
have "target q1 p1 \<in> io_targets M io1 q1"
using \<open>path M q1 p1\<close> \<open>p_io p1 = io1\<close>
unfolding io_targets.simps by blast
then have "target q1 p1 = q2"
using observable_io_targets[OF assms(2) \<open>io1 \<in> LS M q1\<close>]
by (metis assms(3) singletonD)
then show ?thesis
using \<open>io2 \<in> LS M (target q1 p1)\<close> by auto
qed
lemma io_targets_language_append :
assumes "q1 \<in> io_targets M io1 q"
and "io2 \<in> LS M q1"
shows "io1@io2 \<in> LS M q"
proof -
obtain p1 where "path M q p1" and "p_io p1 = io1" and "target q p1 = q1"
using assms(1) by auto
moreover obtain p2 where "path M q1 p2" and "p_io p2 = io2"
using assms(2) by auto
ultimately have "path M q (p1@p2)" and "p_io (p1@p2) = io1@io2"
by auto
then show ?thesis
using language_state_containment[of M q "p1@p2" "io1@io2"] by simp
qed
lemma io_targets_next :
assumes "t \<in> transitions M"
shows "io_targets M io (t_target t) \<subseteq> io_targets M (p_io [t] @ io) (t_source t)"
unfolding io_targets.simps
proof
fix q assume "q \<in> {target (t_target t) p |p. path M (t_target t) p \<and> p_io p = io}"
then obtain p where "path M (t_target t) p \<and> p_io p = io \<and> target (t_target t) p = q"
by auto
then have "path M (t_source t) (t#p) \<and> p_io (t#p) = p_io [t] @ io \<and> target (t_source t) (t#p) = q"
using FSM.path.cons[OF assms] by auto
then show "q \<in> {target (t_source t) p |p. path M (t_source t) p \<and> p_io p = p_io [t] @ io}"
by blast
qed
lemma observable_io_targets_next :
assumes "observable M"
and "t \<in> transitions M"
shows "io_targets M (p_io [t] @ io) (t_source t) = io_targets M io (t_target t)"
proof
show "io_targets M (p_io [t] @ io) (t_source t) \<subseteq> io_targets M io (t_target t)"
proof
fix q assume "q \<in> io_targets M (p_io [t] @ io) (t_source t)"
then obtain p where "q = target (t_source t) p"
and "path M (t_source t) p"
and "p_io p = p_io [t] @ io"
unfolding io_targets.simps by blast
then have "q = t_target (last p)" unfolding target.simps visited_states.simps
using last_map by auto
obtain t' p' where "p = t' # p'"
using \<open>p_io p = p_io [t] @ io\<close> by auto
then have "t' \<in> transitions M" and "t_source t' = t_source t"
using \<open>path M (t_source t) p\<close> by auto
moreover have "t_input t' = t_input t" and "t_output t' = t_output t"
using \<open>p = t' # p'\<close> \<open>p_io p = p_io [t] @ io\<close> by auto
ultimately have "t' = t"
using \<open>t \<in> transitions M\<close> \<open>observable M\<close> unfolding observable.simps
by (meson prod.expand)
then have "path M (t_target t) p'"
using \<open>path M (t_source t) p\<close> \<open>p = t' # p'\<close> by auto
moreover have "p_io p' = io"
using \<open>p_io p = p_io [t] @ io\<close> \<open>p = t' # p'\<close> by auto
moreover have "q = target (t_target t) p'"
using \<open>q = target (t_source t) p\<close> \<open>p = t' # p'\<close> \<open>t' = t\<close> by auto
ultimately show "q \<in> io_targets M io (t_target t)"
by auto
qed
show "io_targets M io (t_target t) \<subseteq> io_targets M (p_io [t] @ io) (t_source t)"
using io_targets_next[OF assms(2)] by assumption
qed
lemma observable_language_target :
assumes "observable M"
and "q \<in> io_targets M io1 (initial M)"
and "t \<in> io_targets T io1 (initial T)"
and "L T \<subseteq> L M"
shows "LS T t \<subseteq> LS M q"
proof
fix io2 assume "io2 \<in> LS T t"
then obtain pT2 where "path T t pT2" and "p_io pT2 = io2"
by auto
obtain pT1 where "path T (initial T) pT1" and "p_io pT1 = io1" and "target (initial T) pT1 = t"
using \<open>t \<in> io_targets T io1 (initial T)\<close> by auto
then have "path T (initial T) (pT1@pT2)"
using \<open>path T t pT2\<close> using path_append by metis
moreover have "p_io (pT1@pT2) = io1@io2"
using \<open>p_io pT1 = io1\<close> \<open>p_io pT2 = io2\<close> by auto
ultimately have "io1@io2 \<in> L T"
using language_state_containment[of T] by auto
then have "io1@io2 \<in> L M"
using \<open>L T \<subseteq> L M\<close> by blast
then obtain pM where "path M (initial M) pM" and "p_io pM = io1@io2"
by auto
let ?pM1 = "take (length io1) pM"
let ?pM2 = "drop (length io1) pM"
have "path M (initial M) (?pM1@?pM2)"
using \<open>path M (initial M) pM\<close> by auto
then have "path M (initial M) ?pM1" and "path M (target (initial M) ?pM1) ?pM2"
by blast+
have "p_io ?pM1 = io1"
using \<open>p_io pM = io1@io2\<close>
by (metis append_eq_conv_conj take_map)
have "p_io ?pM2 = io2"
using \<open>p_io pM = io1@io2\<close>
by (metis append_eq_conv_conj drop_map)
obtain pM1 where "path M (initial M) pM1" and "p_io pM1 = io1" and "target (initial M) pM1 = q"
using \<open>q \<in> io_targets M io1 (initial M)\<close> by auto
have "pM1 = ?pM1"
using observable_path_unique[OF \<open>observable M\<close> \<open>path M (initial M) pM1\<close> \<open>path M (initial M) ?pM1\<close>]
unfolding \<open>p_io pM1 = io1\<close> \<open>p_io ?pM1 = io1\<close> by simp
then have "path M q ?pM2"
using \<open>path M (target (initial M) ?pM1) ?pM2\<close> \<open>target (initial M) pM1 = q\<close> by auto
then show "io2 \<in> LS M q"
using language_state_containment[OF _ \<open>p_io ?pM2 = io2\<close>, of M] by auto
qed
lemma observable_language_target_failure :
assumes "observable M"
and "q \<in> io_targets M io1 (initial M)"
and "t \<in> io_targets T io1 (initial T)"
and "\<not> LS T t \<subseteq> LS M q"
shows "\<not> L T \<subseteq> L M"
using observable_language_target[OF assms(1,2,3)] assms(4) by blast
lemma language_path_append_transition_observable :
assumes "(p_io p) @ [(x,y)] \<in> LS M q"
and "path M q p"
and "observable M"
obtains t where "path M q (p@[t])" and "t_input t = x" and "t_output t = y"
proof -
obtain p' t where "path M q (p'@[t])" and "p_io (p'@[t]) = (p_io p) @ [(x,y)]"
using language_path_append_transition[OF assms(1)] by blast
then have "path M q p'" and "p_io p' = p_io p" and "t_input t = x" and "t_output t = y"
by auto
have "p' = p"
using observable_path_unique[OF assms(3) \<open>path M q p'\<close> \<open>path M q p\<close> \<open>p_io p' = p_io p\<close>] by assumption
then have "path M q (p@[t])"
using \<open>path M q (p'@[t])\<close> by auto
then show ?thesis using that \<open>t_input t = x\<close> \<open>t_output t = y\<close> by metis
qed
lemma language_io_target_append :
assumes "q' \<in> io_targets M io1 q"
and "io2 \<in> LS M q'"
shows "(io1@io2) \<in> LS M q"
proof -
obtain p2 where "path M q' p2" and "p_io p2 = io2"
using assms(2) by auto
moreover obtain p1 where "q' = target q p1" and "path M q p1" and "p_io p1 = io1"
using assms(1) by auto
ultimately show ?thesis unfolding LS.simps
by (metis (mono_tags, lifting) map_append mem_Collect_eq path_append)
qed
lemma observable_path_suffix :
assumes "(p_io p)@io \<in> LS M q"
and "path M q p"
and "observable M"
obtains p' where "path M (target q p) p'" and "p_io p' = io"
proof -
obtain p1 p2 where "path M q p1" and "path M (target q p1) p2" and "p_io p1 = p_io p" and "p_io p2 = io"
using language_state_split[OF assms(1)] by blast
have "p1 = p"
using observable_path_unique[OF assms(3,2) \<open>path M q p1\<close> \<open>p_io p1 = p_io p\<close>[symmetric]]
by simp
show ?thesis using that[of p2] \<open>path M (target q p1) p2\<close> \<open>p_io p2 = io\<close> unfolding \<open>p1 = p\<close>
by blast
qed
lemma io_targets_finite :
"finite (io_targets M io q)"
proof -
have "(io_targets M io q) \<subseteq> {target q p | p . path M q p \<and> length p \<le> length io}"
unfolding io_targets.simps length_map[of "(\<lambda> t . (t_input t, t_output t))", symmetric] by force
moreover have "finite {target q p | p . path M q p \<and> length p \<le> length io}"
using paths_finite[of M q "length io"]
by simp
ultimately show ?thesis
using rev_finite_subset by blast
qed
lemma language_next_transition_ob :
assumes "(x,y)#ios \<in> LS M q"
obtains t where "t_source t = q"
and "t \<in> transitions M"
and "t_input t = x"
and "t_output t = y"
and "ios \<in> LS M (t_target t)"
proof -
obtain p where "path M q p" and "p_io p = (x,y)#ios"
using assms unfolding LS.simps mem_Collect_eq
by (metis (no_types, lifting))
then obtain t p' where "p = t#p'"
by blast
have "t_source t = q"
and "t \<in> transitions M"
and "path M (t_target t) p'"
using \<open>path M q p\<close> unfolding \<open>p = t#p'\<close> by auto
moreover have "t_input t = x"
and "t_output t = y"
and "p_io p' = ios"
using \<open>p_io p = (x,y)#ios\<close> unfolding \<open>p = t#p'\<close> by auto
ultimately show ?thesis using that[of t] by auto
qed
lemma h_observable_card :
assumes "observable M"
shows "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) \<le> 1"
and "finite (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)))"
proof -
have "snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)) = {q' . (q,x,y,q') \<in> transitions M}"
unfolding h.simps by force
moreover have "{q' . (q,x,y,q') \<in> transitions M} = {} \<or> (\<exists> q' . {q' . (q,x,y,q') \<in> transitions M} = {q'})"
using assms unfolding observable_alt_def by blast
ultimately show "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) \<le> 1"
and "finite (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)))"
by auto
qed
lemma h_obs_None :
assumes "observable M"
shows "(h_obs M q x y = None) = (\<nexists>q' . (q,x,y,q') \<in> transitions M)"
proof
show "(h_obs M q x y = None) \<Longrightarrow> (\<nexists>q' . (q,x,y,q') \<in> transitions M)"
proof -
assume "h_obs M q x y = None"
then have "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) \<noteq> 1"
by auto
then have "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) = 0"
using h_observable_card(1)[OF assms, of y q x] by presburger
then have "(snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) = {}"
using h_observable_card(2)[OF assms, of y q x] card_0_eq[of "(snd ` Set.filter (\<lambda>(y', q'). y' = y) (h M (q, x)))"] by blast
then show ?thesis
unfolding h.simps by force
qed
show "(\<nexists>q' . (q,x,y,q') \<in> transitions M) \<Longrightarrow> (h_obs M q x y = None)"
proof -
assume "(\<nexists>q' . (q,x,y,q') \<in> transitions M)"
then have "snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)) = {}"
unfolding h.simps by force
then have "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) = 0"
by simp
then show ?thesis
unfolding h_obs_simps Let_def \<open>snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)) = {}\<close>
by auto
qed
qed
lemma h_obs_Some :
assumes "observable M"
shows "(h_obs M q x y = Some q') = ({q' . (q,x,y,q') \<in> transitions M} = {q'})"
proof
have *: "snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)) = {q' . (q,x,y,q') \<in> transitions M}"
unfolding h.simps by force
show "h_obs M q x y = Some q' \<Longrightarrow> ({q' . (q,x,y,q') \<in> transitions M} = {q'})"
proof -
assume "h_obs M q x y = Some q'"
then have "(snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) \<noteq> {}"
by force
then have "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) > 0"
unfolding h_simps using fsm_transitions_finite[of M]
by (metis assms card_0_eq h_observable_card(2) h_simps neq0_conv)
moreover have "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) \<le> 1"
using assms unfolding observable_alt_def h_simps
by (metis assms h_observable_card(1) h_simps)
ultimately have "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) = 1"
by auto
then have "(snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) = {q'}"
using \<open>h_obs M q x y = Some q'\<close> unfolding h_obs_simps Let_def
by (metis card_1_singletonE option.inject the_elem_eq)
then show ?thesis
using * unfolding h.simps by blast
qed
show "({q' . (q,x,y,q') \<in> transitions M} = {q'}) \<Longrightarrow> (h_obs M q x y = Some q')"
proof -
assume "({q' . (q,x,y,q') \<in> transitions M} = {q'})"
then have "snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)) = {q'}"
unfolding h.simps by force
then show ?thesis
unfolding Let_def
by simp
qed
qed
lemma h_obs_state :
assumes "h_obs M q x y = Some q'"
shows "q' \<in> states M"
proof (cases "card (snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) = 1")
case True
then have "(snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x))) = {q'}"
using \<open>h_obs M q x y = Some q'\<close> unfolding h_obs_simps Let_def
by (metis card_1_singletonE option.inject the_elem_eq)
then have "(q,x,y,q') \<in> transitions M"
unfolding h_simps by auto
then show ?thesis
by (metis fsm_transition_target snd_conv)
next
case False
then have "h_obs M q x y = None"
using False unfolding h_obs_simps Let_def by auto
then show ?thesis using assms by auto
qed
fun after :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> 'a" where
"after M q [] = q" |
"after M q ((x,y)#io) = after M (the (h_obs M q x y)) io"
(*abbreviation(input) "after_initial M io \<equiv> after M (initial M) io" *)
abbreviation "after_initial M io \<equiv> after M (initial M) io"
lemma after_path :
assumes "observable M"
and "path M q p"
shows "after M q (p_io p) = target q p"
using assms(2) proof (induction p arbitrary: q rule: list.induct)
case Nil
then show ?case by auto
next
case (Cons t p)
then have "t \<in> transitions M" and "path M (t_target t) p" and "t_source t = q"
by auto
have "\<And> q' . (q, t_input t, t_output t, q') \<in> FSM.transitions M \<Longrightarrow> q' = t_target t"
using observable_transition_unique[OF assms(1) \<open>t \<in> transitions M\<close>] \<open>t \<in> transitions M\<close>
using \<open>t_source t = q\<close> assms(1) by auto
then have "({q'. (q, t_input t, t_output t, q') \<in> FSM.transitions M} = {t_target t})"
using \<open>t \<in> transitions M\<close> \<open>t_source t = q\<close> by auto
then have "(h_obs M q (t_input t) (t_output t)) = Some (t_target t)"
using h_obs_Some[OF assms(1), of q "t_input t" "t_output t" "t_target t"]
by blast
then have "after M q (p_io (t#p)) = after M (t_target t) (p_io p)"
by auto
moreover have "target (t_target t) p = target q (t#p)"
using \<open>t_source t = q\<close> by auto
ultimately show ?case
using Cons.IH[OF \<open>path M (t_target t) p\<close>]
by simp
qed
lemma observable_after_path :
assumes "observable M"
and "io \<in> LS M q"
obtains p where "path M q p"
and "p_io p = io"
and "target q p = after M q io"
using after_path[OF assms(1)]
using assms(2) by auto
lemma h_obs_from_LS :
assumes "observable M"
and "[(x,y)] \<in> LS M q"
obtains q' where "h_obs M q x y = Some q'"
using assms(2) h_obs_None[OF assms(1), of q x y] by force
lemma after_h_obs :
assumes "observable M"
and "h_obs M q x y = Some q'"
shows "after M q [(x,y)] = q'"
proof -
have "path M q [(q,x,y,q')]"
using assms(2) unfolding h_obs_Some[OF assms(1)]
using single_transition_path by fastforce
then show ?thesis
using assms(2) after_path[OF assms(1), of q "[(q,x,y,q')]"] by auto
qed
lemma after_h_obs_prepend :
assumes "observable M"
and "h_obs M q x y = Some q'"
and "io \<in> LS M q'"
shows "after M q ((x,y)#io) = after M q' io"
proof -
obtain p where "path M q' p" and "p_io p = io"
using assms(3) by auto
then have "after M q' io = target q' p"
using after_path[OF assms(1)]
by blast
have "path M q ((q,x,y,q')#p)"
using assms(2) path_prepend_t[OF \<open>path M q' p\<close>, of q x y] unfolding h_obs_Some[OF assms(1)] by auto
moreover have "p_io ((q,x,y,q')#p) = (x,y)#io"
using \<open>p_io p = io\<close> by auto
ultimately have "after M q ((x,y)#io) = target q ((q,x,y,q')#p)"
using after_path[OF assms(1), of q "(q,x,y,q')#p"] by simp
moreover have "target q ((q,x,y,q')#p) = target q' p"
by auto
ultimately show ?thesis
using \<open>after M q' io = target q' p\<close> by simp
qed
lemma after_split :
assumes "observable M"
and "\<alpha>@\<gamma> \<in> LS M q"
shows "after M (after M q \<alpha>) \<gamma> = after M q (\<alpha> @ \<gamma>)"
proof -
obtain p1 p2 where "path M q p1" and "path M (target q p1) p2" and "p_io p1 = \<alpha>" and "p_io p2 = \<gamma>"
using language_state_split[OF assms(2)]
by blast
then have "path M q (p1@p2)" and "p_io (p1@p2) = (\<alpha> @ \<gamma>)"
by auto
then have "after M q (\<alpha> @ \<gamma>) = target q (p1@p2)"
using assms(1)
by (metis (mono_tags, lifting) after_path)
moreover have "after M q \<alpha> = target q p1"
using \<open>path M q p1\<close> \<open>p_io p1 = \<alpha>\<close> assms(1)
by (metis (mono_tags, lifting) after_path)
moreover have "after M (target q p1) \<gamma> = target (target q p1) p2"
using \<open>path M (target q p1) p2\<close> \<open>p_io p2 = \<gamma>\<close> assms(1)
by (metis (mono_tags, lifting) after_path)
moreover have "target (target q p1) p2 = target q (p1@p2)"
by auto
ultimately show ?thesis
by auto
qed
lemma after_io_targets :
assumes "observable M"
and "io \<in> LS M q"
shows "after M q io = the_elem (io_targets M io q)"
proof -
have "after M q io \<in> io_targets M io q"
using after_path[OF assms(1)] assms(2)
unfolding io_targets.simps LS.simps
by blast
then show ?thesis
using observable_io_targets[OF assms]
by (metis singletonD the_elem_eq)
qed
lemma after_language_subset :
assumes "observable M"
and "\<alpha>@\<gamma> \<in> L M"
and "\<beta> \<in> LS M (after_initial M (\<alpha>@\<gamma>))"
shows "\<gamma>@\<beta> \<in> LS M (after_initial M \<alpha>)"
by (metis after_io_targets after_split assms(1) assms(2) assms(3) language_io_target_append language_prefix observable_io_targets observable_io_targets_language singletonI the_elem_eq)
lemma after_language_append_iff :
assumes "observable M"
and "\<alpha>@\<gamma> \<in> L M"
shows "\<beta> \<in> LS M (after_initial M (\<alpha>@\<gamma>)) = (\<gamma>@\<beta> \<in> LS M (after_initial M \<alpha>))"
by (metis after_io_targets after_language_subset after_split assms(1) assms(2) language_prefix observable_io_targets observable_io_targets_language singletonI the_elem_eq)
lemma h_obs_language_iff :
assumes "observable M"
shows "(x,y)#io \<in> LS M q = (\<exists> q' . h_obs M q x y = Some q' \<and> io \<in> LS M q')"
(is "?P1 = ?P2")
proof
show "?P1 \<Longrightarrow> ?P2"
proof -
assume ?P1
then obtain t p where "t \<in> transitions M"
and "path M (t_target t) p"
and "t_input t = x"
and "t_output t = y"
and "t_source t = q"
and "p_io p = io"
by auto
then have "(q,x,y,t_target t) \<in> transitions M"
by auto
then have "h_obs M q x y = Some (t_target t)"
unfolding h_obs_Some[OF assms]
using assms by auto
moreover have "io \<in> LS M (t_target t)"
using \<open>path M (t_target t) p\<close> \<open>p_io p = io\<close>
by auto
ultimately show ?P2
by blast
qed
show "?P2 \<Longrightarrow> ?P1"
unfolding h_obs_Some[OF assms] using LS_prepend_transition[where io=io and M=M]
by (metis fst_conv mem_Collect_eq singletonI snd_conv)
qed
lemma after_language_iff :
assumes "observable M"
and "\<alpha> \<in> LS M q"
shows "(\<gamma> \<in> LS M (after M q \<alpha>)) = (\<alpha>@\<gamma> \<in> LS M q)"
by (metis after_io_targets assms(1) assms(2) language_io_target_append observable_io_targets observable_io_targets_language singletonI the_elem_eq)
(* TODO: generalise to non-observable FSMs *)
lemma language_maximal_contained_prefix_ob :
assumes "io \<notin> LS M q"
and "q \<in> states M"
and "observable M"
obtains io' x y io'' where "io = io'@[(x,y)]@io''"
and "io' \<in> LS M q"
and "io'@[(x,y)] \<notin> LS M q"
proof -
have "\<exists> io' x y io'' . io = io'@[(x,y)]@io'' \<and> io' \<in> LS M q \<and> io'@[(x,y)] \<notin> LS M q"
using assms(1,2) proof (induction io arbitrary: q)
case Nil
then show ?case by auto
next
case (Cons xy io)
obtain x y where "xy = (x,y)"
by fastforce
show ?case proof (cases "h_obs M q x y")
case None
then have "[]@[(x,y)] \<notin> LS M q"
unfolding h_obs_None[OF assms(3)] by auto
moreover have "[] \<in> LS M q"
using Cons.prems by auto
moreover have "(x,y)#io = []@[(x,y)]@io"
using Cons.prems
unfolding \<open>xy = (x,y)\<close> by auto
ultimately show ?thesis
unfolding \<open>xy = (x,y)\<close> by blast
next
case (Some q')
then have "io \<notin> LS M q'"
using h_obs_language_iff[OF assms(3), of x y io q] Cons.prems(1)
unfolding \<open>xy = (x,y)\<close>
by auto
then obtain io' x' y' io'' where "io = io'@[(x',y')]@io''"
and "io' \<in> LS M q'"
and "io'@[(x',y')] \<notin> LS M q'"
using Cons.IH[OF _ h_obs_state[OF Some]]
by blast
have "xy#io = (xy#io')@[(x',y')]@io''"
using \<open>io = io'@[(x',y')]@io''\<close> by auto
moreover have "(xy#io') \<in> LS M q"
using \<open>io' \<in> LS M q'\<close> Some
unfolding \<open>xy = (x,y)\<close> h_obs_language_iff[OF assms(3)]
by blast
moreover have "(xy#io')@[(x',y')] \<notin> LS M q"
using \<open>io'@[(x',y')] \<notin> LS M q'\<close> Some h_obs_language_iff[OF assms(3), of x y "io'@[(x',y')]" q]
unfolding \<open>xy = (x,y)\<close>
by auto
ultimately show ?thesis
by blast
qed
qed
then show ?thesis
using that by blast
qed
lemma after_is_state :
assumes "observable M"
assumes "io \<in> LS M q"
shows "FSM.after M q io \<in> states M"
using assms
by (metis observable_after_path path_target_is_state)
lemma after_reachable_initial :
assumes "observable M"
and "io \<in> L M"
shows "after_initial M io \<in> reachable_states M"
proof -
obtain p where "path M (initial M) p" and "p_io p = io"
using assms(2) by auto
then have "after_initial M io = target (initial M) p"
using after_path[OF assms(1)]
by blast
then show ?thesis
unfolding reachable_states_def using \<open>path M (initial M) p\<close> by blast
qed
lemma after_transition :
assumes "observable M"
and "(q,x,y,q') \<in> transitions M"
shows "after M q [(x,y)] = q'"
using after_path[OF assms(1) single_transition_path[OF assms(2)]]
by auto
lemma after_transition_exhaust :
assumes "observable M"
and "t \<in> transitions M"
shows "t_target t = after M (t_source t) [(t_input t, t_output t)]"
using after_transition[OF assms(1)] assms(2)
by (metis surjective_pairing)
lemma after_reachable :
assumes "observable M"
and "io \<in> LS M q"
and "q \<in> reachable_states M"
shows "after M q io \<in> reachable_states M"
proof -
obtain p where "path M q p" and "p_io p = io"
using assms(2) by auto
then have "after M q io = target q p"
using after_path[OF assms(1)] by force
obtain p' where "path M (initial M) p'" and "target (initial M) p' = q"
using assms(3) unfolding reachable_states_def by blast
then have "path M (initial M) (p'@p)"
using \<open>path M q p\<close> by auto
moreover have "after M q io = target (initial M) (p'@p)"
using \<open>target (initial M) p' = q\<close>
unfolding \<open>after M q io = target q p\<close>
by auto
ultimately show ?thesis
unfolding reachable_states_def by blast
qed
lemma observable_after_language_append :
assumes "observable M"
and "io1 \<in> LS M q"
and "io2 \<in> LS M (after M q io1)"
shows "io1@io2 \<in> LS M q"
using observable_after_path[OF assms(1,2)] assms(3)
proof -
assume a1: "\<And>thesis. (\<And>p. \<lbrakk>path M q p; p_io p = io1; target q p = after M q io1\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis"
have "\<exists>ps. io2 = p_io ps \<and> path M (after M q io1) ps"
using \<open>io2 \<in> LS M (after M q io1)\<close> by auto
moreover
{ assume "(\<exists>ps. io2 = p_io ps \<and> path M (after M q io1) ps) \<and> (\<forall>ps. io1 @ io2 \<noteq> p_io ps \<or> \<not> path M q ps)"
then have "io1 @ io2 \<in> {p_io ps |ps. path M q ps}"
using a1 by (metis (lifting) map_append path_append) }
ultimately show ?thesis
by auto
qed
lemma observable_after_language_none :
assumes "observable M"
and "io1 \<in> LS M q"
and "io2 \<notin> LS M (after M q io1)"
shows "io1@io2 \<notin> LS M q"
using after_path[OF assms(1)] language_state_split[of io1 io2 M q]
by (metis (mono_tags, lifting) assms(3) language_intro)
lemma observable_after_eq :
assumes "observable M"
and "after M q io1 = after M q io2"
and "io1 \<in> LS M q"
and "io2 \<in> LS M q"
shows "io1@io \<in> LS M q \<longleftrightarrow> io2@io \<in> LS M q"
using observable_after_language_append[OF assms(1,3), of io]
observable_after_language_append[OF assms(1,4), of io]
assms(2)
by (metis assms(1) language_prefix observable_after_language_none)
lemma observable_after_target :
assumes "observable M"
and "io @ io' \<in> LS M q"
and "path M (FSM.after M q io) p"
and "p_io p = io'"
shows "target (FSM.after M q io) p = (FSM.after M q (io @ io'))"
proof -
obtain p' where "path M q p'" and "p_io p' = io @ io'"
using \<open>io @ io' \<in> LS M q\<close> by auto
then have "path M q (take (length io) p')"
and "p_io (take (length io) p') = io"
and "path M (target q (take (length io) p')) (drop (length io) p')"
and "p_io (drop (length io) p') = io'"
using path_io_split[of M q p' io io']
by auto
then have "FSM.after M q io = target q (take (length io) p')"
using after_path assms(1) by fastforce
then have "p = (drop (length io) p')"
using \<open>path M (target q (take (length io) p')) (drop (length io) p')\<close> \<open>p_io (drop (length io) p') = io'\<close>
assms(3,4)
observable_path_unique[OF \<open>observable M\<close>]
by force
have "(FSM.after M q (io @ io')) = target q p'"
using after_path[OF \<open>observable M\<close> \<open>path M q p'\<close>] unfolding \<open>p_io p' = io @ io'\<close> .
moreover have "target (FSM.after M q io) p = target q p'"
using \<open>FSM.after M q io = target q (take (length io) p')\<close>
by (metis \<open>p = drop (length io) p'\<close> append_take_drop_id path_append_target)
ultimately show ?thesis
by simp
qed
fun is_in_language :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('b \<times>'c) list \<Rightarrow> bool" where
"is_in_language M q [] = True" |
"is_in_language M q ((x,y)#io) = (case h_obs M q x y of
None \<Rightarrow> False |
Some q' \<Rightarrow> is_in_language M q' io)"
lemma is_in_language_iff :
assumes "observable M"
and "q \<in> states M"
shows "is_in_language M q io \<longleftrightarrow> io \<in> LS M q"
using assms(2) proof (induction io arbitrary: q)
case Nil
then show ?case
by auto
next
case (Cons xy io)
obtain x y where "xy = (x,y)"
using prod.exhaust by metis
show ?case
unfolding \<open>xy = (x,y)\<close>
unfolding h_obs_language_iff[OF assms(1), of x y io q]
unfolding is_in_language.simps
apply (cases "h_obs M q x y")
apply auto[1]
by (metis Cons.IH h_obs_state option.simps(5))
qed
lemma observable_paths_for_io :
assumes "observable M"
and "io \<in> LS M q"
obtains p where "paths_for_io M q io = {p}"
proof -
obtain p where "path M q p" and "p_io p = io"
using assms(2) by auto
then have "p \<in> paths_for_io M q io"
unfolding paths_for_io_def
by blast
then show ?thesis
using that[of p]
using observable_path_unique[OF assms(1) \<open>path M q p\<close>] \<open>p_io p = io\<close>
unfolding paths_for_io_def
by force
qed
lemma io_targets_language :
assumes "q' \<in> io_targets M io q"
shows "io \<in> LS M q"
using assms by auto
lemma observable_after_reachable_surj :
assumes "observable M"
shows "(after_initial M) ` (L M) = reachable_states M"
proof
show "after_initial M ` L M \<subseteq> reachable_states M"
using after_reachable[OF assms _ reachable_states_initial]
by blast
show "reachable_states M \<subseteq> after_initial M ` L M"
unfolding reachable_states_def
using after_path[OF assms]
using image_iff by fastforce
qed
lemma observable_minimal_size_r_language_distinct :
assumes "minimal M1"
and "minimal M2"
and "observable M1"
and "observable M2"
and "size_r M1 < size_r M2"
shows "L M1 \<noteq> L M2"
proof
assume "L M1 = L M2"
define V where "V = (\<lambda> q . SOME io . io \<in> L M1 \<and> after_initial M2 io = q)"
have "\<And> q . q \<in> reachable_states M2 \<Longrightarrow> V q \<in> L M1 \<and> after_initial M2 (V q) = q"
proof -
fix q assume "q \<in> reachable_states M2"
then have "\<exists> io . io \<in> L M1 \<and> after_initial M2 io = q"
unfolding \<open>L M1 = L M2\<close>
by (metis assms(4) imageE observable_after_reachable_surj)
then show "V q \<in> L M1 \<and> after_initial M2 (V q) = q"
unfolding V_def
using someI_ex[of "\<lambda> io . io \<in> L M1 \<and> after_initial M2 io = q"] by blast
qed
then have "(after_initial M1) ` V ` reachable_states M2 \<subseteq> reachable_states M1"
by (metis assms(3) image_mono image_subsetI observable_after_reachable_surj)
then have "card (after_initial M1 ` V ` reachable_states M2) \<le> size_r M1"
using reachable_states_finite[of M1]
by (meson card_mono)
have "(after_initial M2) ` V ` reachable_states M2 = reachable_states M2"
proof
show "after_initial M2 ` V ` reachable_states M2 \<subseteq> reachable_states M2"
using \<open>\<And> q . q \<in> reachable_states M2 \<Longrightarrow> V q \<in> L M1 \<and> after_initial M2 (V q) = q\<close> by auto
show "reachable_states M2 \<subseteq> after_initial M2 ` V ` reachable_states M2"
using \<open>\<And> q . q \<in> reachable_states M2 \<Longrightarrow> V q \<in> L M1 \<and> after_initial M2 (V q) = q\<close> observable_after_reachable_surj[OF assms(4)] unfolding \<open>L M1 = L M2\<close>
using image_iff by fastforce
qed
then have "card ((after_initial M2) ` V ` reachable_states M2) = size_r M2"
by auto
have *: "finite (V ` reachable_states M2)"
by (simp add: reachable_states_finite)
have **: "card ((after_initial M1) ` V ` reachable_states M2) < card ((after_initial M2) ` V ` reachable_states M2)"
using assms(5) \<open>card (after_initial M1 ` V ` reachable_states M2) \<le> size_r M1\<close>
unfolding \<open>card ((after_initial M2) ` V ` reachable_states M2) = size_r M2\<close>
by linarith
obtain io1 io2 where "io1 \<in> V ` reachable_states M2"
"io2 \<in> V ` reachable_states M2"
"after_initial M2 io1 \<noteq> after_initial M2 io2"
"after_initial M1 io1 = after_initial M1 io2"
using finite_card_less_witnesses[OF * **]
by blast
then have "io1 \<in> L M1" and "io2 \<in> L M1" and "io1 \<in> L M2" and "io2 \<in> L M2"
using \<open>\<And> q . q \<in> reachable_states M2 \<Longrightarrow> V q \<in> L M1 \<and> after_initial M2 (V q) = q\<close> unfolding \<open>L M1 = L M2\<close>
by auto
then have "after_initial M1 io1 \<in> reachable_states M1"
"after_initial M1 io2 \<in> reachable_states M1"
"after_initial M2 io1 \<in> reachable_states M2"
"after_initial M2 io2 \<in> reachable_states M2"
using after_reachable[OF assms(3) _ reachable_states_initial] after_reachable[OF assms(4) _ reachable_states_initial]
by blast+
obtain io3 where "io3 \<in> LS M2 (after_initial M2 io1) = (io3 \<notin> LS M2 (after_initial M2 io2))"
using reachable_state_is_state[OF \<open>after_initial M2 io1 \<in> reachable_states M2\<close>]
reachable_state_is_state[OF \<open>after_initial M2 io2 \<in> reachable_states M2\<close>]
\<open>after_initial M2 io1 \<noteq> after_initial M2 io2\<close> assms(2)
unfolding minimal.simps by blast
then have "io1@io3 \<in> L M2 = (io2@io3 \<notin> L M2)"
using observable_after_language_append[OF assms(4) \<open>io1 \<in> L M2\<close>]
observable_after_language_append[OF assms(4) \<open>io2 \<in> L M2\<close>]
observable_after_language_none[OF assms(4) \<open>io1 \<in> L M2\<close>]
observable_after_language_none[OF assms(4) \<open>io2 \<in> L M2\<close>]
by blast
moreover have "io1@io3 \<in> L M1 = (io2@io3 \<in> L M1)"
by (meson \<open>after_initial M1 io1 = after_initial M1 io2\<close> \<open>io1 \<in> L M1\<close> \<open>io2 \<in> L M1\<close> assms(3) observable_after_eq)
ultimately show False
using \<open>L M1 = L M2\<close> by blast
qed
(* language equivalent minimal FSMs have the same number of reachable states *)
lemma minimal_equivalence_size_r :
assumes "minimal M1"
and "minimal M2"
and "observable M1"
and "observable M2"
and "L M1 = L M2"
shows "size_r M1 = size_r M2"
using observable_minimal_size_r_language_distinct[OF assms(1-4)]
observable_minimal_size_r_language_distinct[OF assms(2,1,4,3)]
assms(5)
using nat_neq_iff by auto
subsection \<open>Conformity Relations\<close>
fun is_io_reduction_state :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('d,'b,'c) fsm \<Rightarrow> 'd \<Rightarrow> bool" where
"is_io_reduction_state A a B b = (LS A a \<subseteq> LS B b)"
abbreviation(input) "is_io_reduction A B \<equiv> is_io_reduction_state A (initial A) B (initial B)"
notation
is_io_reduction ("_ \<preceq> _")
fun is_io_reduction_state_on_inputs :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> 'b list set \<Rightarrow> ('d,'b,'c) fsm \<Rightarrow> 'd \<Rightarrow> bool" where
"is_io_reduction_state_on_inputs A a U B b = (LS\<^sub>i\<^sub>n A a U \<subseteq> LS\<^sub>i\<^sub>n B b U)"
abbreviation(input) "is_io_reduction_on_inputs A U B \<equiv> is_io_reduction_state_on_inputs A (initial A) U B (initial B)"
notation
is_io_reduction_on_inputs ("_ \<preceq>\<lbrakk>_\<rbrakk> _")
subsection \<open>A Pass Relation for Reduction and Test Represented as Sets of Input-Output Sequences\<close>
definition pass_io_set :: "('a,'b,'c) fsm \<Rightarrow> ('b \<times> 'c) list set \<Rightarrow> bool" where
"pass_io_set M ios = (\<forall> io x y . io@[(x,y)] \<in> ios \<longrightarrow> (\<forall> y' . io@[(x,y')] \<in> L M \<longrightarrow> io@[(x,y')] \<in> ios))"
definition pass_io_set_maximal :: "('a,'b,'c) fsm \<Rightarrow> ('b \<times> 'c) list set \<Rightarrow> bool" where
"pass_io_set_maximal M ios = (\<forall> io x y io' . io@[(x,y)]@io' \<in> ios \<longrightarrow> (\<forall> y' . io@[(x,y')] \<in> L M \<longrightarrow> (\<exists> io''. io@[(x,y')]@io'' \<in> ios)))"
lemma pass_io_set_from_pass_io_set_maximal :
"pass_io_set_maximal M ios = pass_io_set M {io' . \<exists> io io'' . io = io'@io'' \<and> io \<in> ios}"
proof -
have "\<And> io x y io' . io@[(x,y)]@io' \<in> ios \<Longrightarrow> io@[(x,y)] \<in> {io' . \<exists> io io'' . io = io'@io'' \<and> io \<in> ios}"
by auto
moreover have "\<And> io x y . io@[(x,y)] \<in> {io' . \<exists> io io'' . io = io'@io'' \<and> io \<in> ios} \<Longrightarrow> \<exists> io' . io@[(x,y)]@io' \<in> ios"
by auto
ultimately show ?thesis
unfolding pass_io_set_def pass_io_set_maximal_def
by meson
qed
lemma pass_io_set_maximal_from_pass_io_set :
assumes "finite ios"
and "\<And> io' io'' . io'@io'' \<in> ios \<Longrightarrow> io' \<in> ios"
shows "pass_io_set M ios = pass_io_set_maximal M {io' \<in> ios . \<not> (\<exists> io'' . io'' \<noteq> [] \<and> io'@io'' \<in> ios)}"
proof -
have "\<And> io x y . io@[(x,y)] \<in> ios \<Longrightarrow> \<exists> io' . io@[(x,y)]@io' \<in> {io'' \<in> ios . \<not> (\<exists> io''' . io''' \<noteq> [] \<and> io''@io''' \<in> ios)}"
proof -
fix io x y assume "io@[(x,y)] \<in> ios"
show "\<exists> io' . io@[(x,y)]@io' \<in> {io'' \<in> ios . \<not> (\<exists> io''' . io''' \<noteq> [] \<and> io''@io''' \<in> ios)}"
using finite_set_elem_maximal_extension_ex[OF \<open>io@[(x,y)] \<in> ios\<close> assms(1)] by force
qed
moreover have "\<And> io x y io' . io@[(x,y)]@io' \<in> {io'' \<in> ios . \<not> (\<exists> io''' . io''' \<noteq> [] \<and> io''@io''' \<in> ios)} \<Longrightarrow> io@[(x,y)] \<in> ios"
using \<open>\<And> io' io'' . io'@io'' \<in> ios \<Longrightarrow> io' \<in> ios\<close> by force
ultimately show ?thesis
unfolding pass_io_set_def pass_io_set_maximal_def
by meson
qed
subsection \<open>Relaxation of IO based test suites to sets of input sequences\<close>
abbreviation(input) "input_portion xs \<equiv> map fst xs"
lemma equivalence_io_relaxation :
assumes "(L M1 = L M2) \<longleftrightarrow> (L M1 \<inter> T = L M2 \<inter> T)"
shows "(L M1 = L M2) \<longleftrightarrow> ({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} = {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')})"
proof
show "(L M1 = L M2) \<Longrightarrow> ({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} = {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')})"
by blast
show "({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} = {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')}) \<Longrightarrow> L M1 = L M2"
proof -
have *:"\<And> M . {io . io \<in> L M \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} = L M \<inter> {io . \<exists> io' \<in> T . input_portion io = input_portion io'}"
by blast
have "({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} = {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')}) \<Longrightarrow> (L M1 \<inter> T = L M2 \<inter> T)"
unfolding * by blast
then show "({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} = {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')}) \<Longrightarrow> L M1 = L M2"
using assms by blast
qed
qed
lemma reduction_io_relaxation :
assumes "(L M1 \<subseteq> L M2) \<longleftrightarrow> (L M1 \<inter> T \<subseteq> L M2 \<inter> T)"
shows "(L M1 \<subseteq> L M2) \<longleftrightarrow> ({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} \<subseteq> {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')})"
proof
show "(L M1 \<subseteq> L M2) \<Longrightarrow> ({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} \<subseteq> {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')})"
by blast
show "({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} \<subseteq> {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')}) \<Longrightarrow> L M1 \<subseteq> L M2"
proof -
have *:"\<And> M . {io . io \<in> L M \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} \<subseteq> L M \<inter> {io . \<exists> io' \<in> T . input_portion io = input_portion io'}"
by blast
have "({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} \<subseteq> {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')}) \<Longrightarrow> (L M1 \<inter> T \<subseteq> L M2 \<inter> T)"
unfolding * by blast
then show "({io . io \<in> L M1 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')} \<subseteq> {io . io \<in> L M2 \<and> (\<exists> io' \<in> T . input_portion io = input_portion io')}) \<Longrightarrow> L M1 \<subseteq> L M2"
using assms by blast
qed
qed
subsection \<open>Submachines\<close>
fun is_submachine :: "('a,'b,'c) fsm \<Rightarrow> ('a,'b,'c) fsm \<Rightarrow> bool" where
"is_submachine A B = (initial A = initial B \<and> transitions A \<subseteq> transitions B \<and> inputs A = inputs B \<and> outputs A = outputs B \<and> states A \<subseteq> states B)"
lemma submachine_path_initial :
assumes "is_submachine A B"
and "path A (initial A) p"
shows "path B (initial B) p"
using assms proof (induction p rule: rev_induct)
case Nil
then show ?case by auto
next
case (snoc a p)
then show ?case
by fastforce
qed
lemma submachine_path :
assumes "is_submachine A B"
and "path A q p"
shows "path B q p"
by (meson assms(1) assms(2) is_submachine.elims(2) path_begin_state subsetD transition_subset_path)
lemma submachine_reduction :
assumes "is_submachine A B"
shows "is_io_reduction A B"
using submachine_path[OF assms] assms by auto
lemma complete_submachine_initial :
assumes "is_submachine A B"
and "completely_specified A"
shows "completely_specified_state B (initial B)"
using assms(1) assms(2) fsm_initial subset_iff by fastforce
lemma submachine_language :
assumes "is_submachine S M"
shows "L S \<subseteq> L M"
by (meson assms is_io_reduction_state.elims(2) submachine_reduction)
lemma submachine_observable :
assumes "is_submachine S M"
and "observable M"
shows "observable S"
using assms unfolding is_submachine.simps observable.simps by blast
lemma submachine_transitive :
assumes "is_submachine S M"
and "is_submachine S' S"
shows "is_submachine S' M"
using assms unfolding is_submachine.simps by force
lemma transitions_subset_path :
assumes "set p \<subseteq> transitions M"
and "p \<noteq> []"
and "path S q p"
shows "path M q p"
using assms by (induction p arbitrary: q; auto)
lemma transition_subset_paths :
assumes "transitions S \<subseteq> transitions M"
and "initial S \<in> states M"
and "inputs S = inputs M"
and "outputs S = outputs M"
and "path S (initial S) p"
shows "path M (initial S) p"
using assms(5) proof (induction p rule: rev_induct)
case Nil
then show ?case using assms(2) by auto
next
case (snoc t p)
then have "path S (initial S) p"
and "t \<in> transitions S"
and "t_source t = target (initial S) p"
and "path M (initial S) p"
by auto
have "t \<in> transitions M"
using assms(1) \<open>t \<in> transitions S\<close> by auto
moreover have "t_source t \<in> states M"
using \<open>t_source t = target (initial S) p\<close> \<open>path M (initial S) p\<close>
using path_target_is_state by fastforce
ultimately have "t \<in> transitions M"
using \<open>t \<in> transitions S\<close> assms(3,4) by auto
then show ?case
using \<open>path M (initial S) p\<close>
using snoc.prems by auto
qed
lemma submachine_reachable_subset :
assumes "is_submachine A B"
shows "reachable_states A \<subseteq> reachable_states B"
using assms submachine_path_initial[OF assms]
unfolding is_submachine.simps reachable_states_def by force
lemma submachine_simps :
assumes "is_submachine A B"
shows "initial A = initial B"
and "states A \<subseteq> states B"
and "inputs A = inputs B"
and "outputs A = outputs B"
and "transitions A \<subseteq> transitions B"
using assms unfolding is_submachine.simps by blast+
lemma submachine_deadlock :
assumes "is_submachine A B"
and "deadlock_state B q"
shows "deadlock_state A q"
using assms(1) assms(2) in_mono by auto
subsection \<open>Changing Initial States\<close>
lift_definition from_FSM :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('a,'b,'c) fsm" is FSM_Impl.from_FSMI
by simp
lemma from_FSM_simps[simp]:
assumes "q \<in> states M"
shows
"initial (from_FSM M q) = q"
"inputs (from_FSM M q) = inputs M"
"outputs (from_FSM M q) = outputs M"
"transitions (from_FSM M q) = transitions M"
"states (from_FSM M q) = states M" using assms by (transfer; simp)+
lemma from_FSM_path_initial :
assumes "q \<in> states M"
shows "path M q p = path (from_FSM M q) (initial (from_FSM M q)) p"
by (metis assms from_FSM_simps(1) from_FSM_simps(4) from_FSM_simps(5) order_refl
transition_subset_path)
lemma from_FSM_path :
assumes "q \<in> states M"
and "path (from_FSM M q) q' p"
shows "path M q' p"
using assms(1) assms(2) path_transitions transitions_subset_path by fastforce
lemma from_FSM_reachable_states :
assumes "q \<in> reachable_states M"
shows "reachable_states (from_FSM M q) \<subseteq> reachable_states M"
proof
from assms obtain p where "path M (initial M) p" and "target (initial M) p = q"
unfolding reachable_states_def by blast
then have "q \<in> states M"
by (meson path_target_is_state)
fix q' assume "q' \<in> reachable_states (from_FSM M q)"
then obtain p' where "path (from_FSM M q) q p'" and "target q p' = q'"
unfolding reachable_states_def from_FSM_simps[OF \<open>q \<in> states M\<close>] by blast
then have "path M (initial M) (p@p')" and "target (initial M) (p@p') = q'"
using from_FSM_path[OF \<open>q \<in> states M\<close> ] \<open>path M (initial M) p\<close>
using \<open>target (FSM.initial M) p = q\<close> by auto
then show "q' \<in> reachable_states M"
unfolding reachable_states_def by blast
qed
lemma submachine_from :
assumes "is_submachine S M"
and "q \<in> states S"
shows "is_submachine (from_FSM S q) (from_FSM M q)"
proof -
have "path S q []"
using assms(2) by blast
then have "path M q []"
by (meson assms(1) submachine_path)
then show ?thesis
using assms(1) assms(2) by force
qed
lemma from_FSM_path_rev_initial :
assumes "path M q p"
shows "path (from_FSM M q) q p"
by (metis (no_types) assms from_FSM_path_initial from_FSM_simps(1) path_begin_state)
lemma from_from[simp] :
assumes "q1 \<in> states M"
and "q1' \<in> states M"
shows "from_FSM (from_FSM M q1) q1' = from_FSM M q1'" (is "?M = ?M'")
proof -
have *: "q1' \<in> states (from_FSM M q1)"
using assms(2) unfolding from_FSM_simps(5)[OF assms(1)] by assumption
have "initial ?M = initial ?M'"
and "states ?M = states ?M'"
and "inputs ?M = inputs ?M'"
and "outputs ?M = outputs ?M'"
and "transitions ?M = transitions ?M'"
unfolding from_FSM_simps[OF *] from_FSM_simps[OF assms(1)] from_FSM_simps[OF assms(2)] by simp+
then show ?thesis by (transfer; force)
qed
lemma from_FSM_completely_specified :
assumes "completely_specified M"
shows "completely_specified (from_FSM M q)" proof (cases "q \<in> states M")
case True
then show ?thesis
using assms by auto
next
case False
then have "from_FSM M q = M" by (transfer; auto)
then show ?thesis using assms by auto
qed
lemma from_FSM_single_input :
assumes "single_input M"
shows "single_input (from_FSM M q)" proof (cases "q \<in> states M")
case True
then show ?thesis
using assms
by (metis from_FSM_simps(4) single_input.elims(1))
next
case False
then have "from_FSM M q = M" by (transfer; auto)
then show ?thesis using assms
by presburger
qed
lemma from_FSM_acyclic :
assumes "q \<in> reachable_states M"
and "acyclic M"
shows "acyclic (from_FSM M q)"
using assms(1)
acyclic_paths_from_reachable_states[OF assms(2), of _ q]
from_FSM_path[of q M q]
path_target_is_state
reachable_state_is_state[OF assms(1)]
from_FSM_simps(1)
unfolding acyclic.simps
reachable_states_def
by force
lemma from_FSM_observable :
assumes "observable M"
shows "observable (from_FSM M q)"
proof (cases "q \<in> states M")
case True
then show ?thesis
using assms
proof -
have f1: "\<forall>f. observable f = (\<forall>a b c aa ab. ((a::'a, b::'b, c::'c, aa) \<notin> FSM.transitions f \<or> (a, b, c, ab) \<notin> FSM.transitions f) \<or> aa = ab)"
by force
have "\<forall>a f. a \<notin> FSM.states (f::('a, 'b, 'c) fsm) \<or> FSM.transitions (FSM.from_FSM f a) = FSM.transitions f"
by (meson from_FSM_simps(4))
then show ?thesis
using f1 True assms by presburger
qed
next
case False
then have "from_FSM M q = M" by (transfer; auto)
then show ?thesis using assms by presburger
qed
lemma observable_language_next :
assumes "io#ios \<in> LS M (t_source t)"
and "observable M"
and "t \<in> transitions M"
and "t_input t = fst io"
and "t_output t = snd io"
shows "ios \<in> L (from_FSM M (t_target t))"
proof -
obtain p where "path M (t_source t) p" and "p_io p = io#ios"
using assms(1)
proof -
assume a1: "\<And>p. \<lbrakk>path M (t_source t) p; p_io p = io # ios\<rbrakk> \<Longrightarrow> thesis"
obtain pps :: "('a \<times> 'b) list \<Rightarrow> 'c \<Rightarrow> ('c, 'a, 'b) fsm \<Rightarrow> ('c \<times> 'a \<times> 'b \<times> 'c) list" where
"\<forall>x0 x1 x2. (\<exists>v3. x0 = p_io v3 \<and> path x2 x1 v3) = (x0 = p_io (pps x0 x1 x2) \<and> path x2 x1 (pps x0 x1 x2))"
by moura
then have "\<exists>ps. path M (t_source t) ps \<and> p_io ps = io # ios"
using assms(1) by auto
then show ?thesis
using a1 by meson
qed
then obtain t' p' where "p = t' # p'"
by auto
then have "t' \<in> transitions M" and "t_source t' = t_source t" and "t_input t' = fst io" and "t_output t' = snd io"
using \<open>path M (t_source t) p\<close> \<open>p_io p = io#ios\<close> by auto
then have "t = t'"
using assms(2,3,4,5) unfolding observable.simps
by (metis (no_types, opaque_lifting) prod.expand)
then have "path M (t_target t) p'" and "p_io p' = ios"
using \<open>p = t' # p'\<close> \<open>path M (t_source t) p\<close> \<open>p_io p = io#ios\<close> by auto
then have "path (from_FSM M (t_target t)) (initial (from_FSM M (t_target t))) p'"
by (meson assms(3) from_FSM_path_initial fsm_transition_target)
then show ?thesis using \<open>p_io p' = ios\<close> by auto
qed
lemma from_FSM_language :
assumes "q \<in> states M"
shows "L (from_FSM M q) = LS M q"
using assms unfolding LS.simps by (meson from_FSM_path_initial)
lemma observable_transition_target_language_subset :
assumes "LS M (t_source t1) \<subseteq> LS M (t_source t2)"
and "t1 \<in> transitions M"
and "t2 \<in> transitions M"
and "t_input t1 = t_input t2"
and "t_output t1 = t_output t2"
and "observable M"
shows "LS M (t_target t1) \<subseteq> LS M (t_target t2)"
proof (rule ccontr)
assume "\<not> LS M (t_target t1) \<subseteq> LS M (t_target t2)"
then obtain ioF where "ioF \<in> LS M (t_target t1)" and "ioF \<notin> LS M (t_target t2)"
by blast
then have "(t_input t1, t_output t1)#ioF \<in> LS M (t_source t1)"
using LS_prepend_transition assms(2) by blast
then have *: "(t_input t1, t_output t1)#ioF \<in> LS M (t_source t2)"
using assms(1) by blast
have "ioF \<in> LS M (t_target t2)"
using observable_language_next[OF * \<open>observable M\<close> \<open>t2 \<in> transitions M\<close> ] unfolding assms(4,5) fst_conv snd_conv
by (metis assms(3) from_FSM_language fsm_transition_target)
then show False
using \<open>ioF \<notin> LS M (t_target t2)\<close> by blast
qed
lemma observable_transition_target_language_eq :
assumes "LS M (t_source t1) = LS M (t_source t2)"
and "t1 \<in> transitions M"
and "t2 \<in> transitions M"
and "t_input t1 = t_input t2"
and "t_output t1 = t_output t2"
and "observable M"
shows "LS M (t_target t1) = LS M (t_target t2)"
using observable_transition_target_language_subset[OF _ assms(2,3,4,5,6)]
observable_transition_target_language_subset[OF _ assms(3,2) assms(4,5)[symmetric] assms(6)]
assms(1)
by blast
lemma language_state_prepend_transition :
assumes "io \<in> LS (from_FSM A (t_target t)) (initial (from_FSM A (t_target t)))"
and "t \<in> transitions A"
shows "p_io [t] @ io \<in> LS A (t_source t)"
proof -
obtain p where "path (from_FSM A (t_target t)) (initial (from_FSM A (t_target t))) p"
and "p_io p = io"
using assms(1) unfolding LS.simps by blast
then have "path A (t_target t) p"
by (meson assms(2) from_FSM_path_initial fsm_transition_target)
then have "path A (t_source t) (t # p)"
using assms(2) by auto
then show ?thesis
using \<open>p_io p = io\<close> unfolding LS.simps
by force
qed
lemma observable_language_transition_target :
assumes "observable M"
and "t \<in> transitions M"
and "(t_input t, t_output t) # io \<in> LS M (t_source t)"
shows "io \<in> LS M (t_target t)"
by (metis (no_types) assms(1) assms(2) assms(3) from_FSM_language fsm_transition_target fst_conv observable_language_next snd_conv)
lemma LS_single_transition :
"[(x,y)] \<in> LS M q \<longleftrightarrow> (\<exists> t \<in> transitions M . t_source t = q \<and> t_input t = x \<and> t_output t = y)"
proof
show "[(x, y)] \<in> LS M q \<Longrightarrow> \<exists>t\<in>FSM.transitions M. t_source t = q \<and> t_input t = x \<and> t_output t = y"
by auto
show "\<exists>t\<in>FSM.transitions M. t_source t = q \<and> t_input t = x \<and> t_output t = y \<Longrightarrow> [(x, y)] \<in> LS M q"
by (metis LS_prepend_transition from_FSM_language fsm_transition_target language_contains_empty_sequence)
qed
lemma h_obs_language_append :
assumes "observable M"
and "u \<in> L M"
and "h_obs M (after_initial M u) x y \<noteq> None"
shows "u@[(x,y)] \<in> L M"
using after_language_iff[OF assms(1,2), of "[(x,y)]"]
using h_obs_None[OF assms(1)] assms(3)
unfolding LS_single_transition
by (metis old.prod.inject prod.collapse)
lemma h_obs_language_single_transition_iff :
assumes "observable M"
shows "[(x,y)] \<in> LS M q \<longleftrightarrow> h_obs M q x y \<noteq> None"
using h_obs_None[OF assms(1), of q x y]
unfolding LS_single_transition
by (metis fst_conv prod.exhaust_sel snd_conv)
(* TODO: generalise to non-observable FSMs *)
lemma minimal_failure_prefix_ob :
assumes "observable M"
and "observable I"
and "qM \<in> states M"
and "qI \<in> states I"
and "io \<in> LS I qI - LS M qM"
obtains io' xy io'' where "io = io'@[xy]@io''"
and "io' \<in> LS I qI \<inter> LS M qM"
and "io'@[xy] \<in> LS I qI - LS M qM"
proof -
have "\<exists> io' xy io'' . io = io'@[xy]@io'' \<and> io' \<in> LS I qI \<inter> LS M qM \<and> io'@[xy] \<in> LS I qI - LS M qM"
using assms(3,4,5) proof (induction io arbitrary: qM qI)
case Nil
then show ?case by auto
next
case (Cons xy io)
show ?case proof (cases "[xy] \<in> LS I qI - LS M qM")
case True
have "xy # io = []@[xy]@io"
by auto
moreover have "[] \<in> LS I qI \<inter> LS M qM"
using \<open>qM \<in> states M\<close> \<open>qI \<in> states I\<close> by auto
moreover have "[]@[xy] \<in> LS I qI - LS M qM"
using True by auto
ultimately show ?thesis
by blast
next
case False
obtain x y where "xy = (x,y)"
by (meson surj_pair)
have "[(x,y)] \<in> LS M qM"
using \<open>xy = (x,y)\<close> False \<open>xy # io \<in> LS I qI - LS M qM\<close>
by (metis DiffD1 DiffI append_Cons append_Nil language_prefix)
then obtain qM' where "(qM,x,y,qM') \<in> transitions M"
by auto
then have "io \<notin> LS M qM'"
using observable_language_transition_target[OF \<open>observable M\<close>]
\<open>xy = (x,y)\<close> \<open>xy # io \<in> LS I qI - LS M qM\<close>
by (metis DiffD2 LS_prepend_transition fst_conv snd_conv)
have "[(x,y)] \<in> LS I qI"
using \<open>xy = (x,y)\<close> \<open>xy # io \<in> LS I qI - LS M qM\<close>
by (metis DiffD1 append_Cons append_Nil language_prefix)
then obtain qI' where "(qI,x,y,qI') \<in> transitions I"
by auto
then have "io \<in> LS I qI'"
using observable_language_next[of xy io I "(qI,x,y,qI')", OF _ \<open>observable I\<close>]
\<open>xy # io \<in> LS I qI - LS M qM\<close> fsm_transition_target[OF \<open>(qI,x,y,qI') \<in> transitions I\<close>]
unfolding \<open>xy = (x,y)\<close> fst_conv snd_conv
by (metis DiffD1 from_FSM_language)
obtain io' xy' io'' where "io = io'@[xy']@io''" and "io' \<in> LS I qI' \<inter> LS M qM'" and "io'@[xy'] \<in> LS I qI' - LS M qM'"
using \<open>io \<in> LS I qI'\<close> \<open>io \<notin> LS M qM'\<close>
Cons.IH[OF fsm_transition_target[OF \<open>(qM,x,y,qM') \<in> transitions M\<close>]
fsm_transition_target[OF \<open>(qI,x,y,qI') \<in> transitions I\<close>] ]
unfolding fst_conv snd_conv
by blast
have "xy#io = (xy#io')@[xy']@io''"
using \<open>io = io'@[xy']@io''\<close> \<open>xy = (x,y)\<close> by auto
moreover have "xy#io' \<in> LS I qI \<inter> LS M qM"
using LS_prepend_transition[OF \<open>(qI,x,y,qI') \<in> transitions I\<close>, of io']
using LS_prepend_transition[OF \<open>(qM,x,y,qM') \<in> transitions M\<close>, of io']
using \<open>io' \<in> LS I qI' \<inter> LS M qM'\<close>
unfolding \<open>xy = (x,y)\<close> fst_conv snd_conv
by auto
moreover have "(xy#io')@[xy'] \<in> LS I qI - LS M qM"
using LS_prepend_transition[OF \<open>(qI,x,y,qI') \<in> transitions I\<close>, of "io'@[xy']"]
using observable_language_transition_target[OF \<open>observable M\<close> \<open>(qM,x,y,qM') \<in> transitions M\<close>, of "io'@[xy']"]
using \<open>io'@[xy'] \<in> LS I qI' - LS M qM'\<close>
unfolding \<open>xy = (x,y)\<close> fst_conv snd_conv
by fastforce
ultimately show ?thesis
by blast
qed
qed
then show ?thesis
using that by blast
qed
subsection \<open>Language and Defined Inputs\<close>
lemma defined_inputs_code : "defined_inputs M q = t_input ` Set.filter (\<lambda>t . t_source t = q) (transitions M)"
unfolding defined_inputs_set by force
lemma defined_inputs_alt_def : "defined_inputs M q = {t_input t | t . t \<in> transitions M \<and> t_source t = q}"
unfolding defined_inputs_code by force
lemma defined_inputs_language_diff :
assumes "x \<in> defined_inputs M1 q1"
and "x \<notin> defined_inputs M2 q2"
obtains y where "[(x,y)] \<in> LS M1 q1 - LS M2 q2"
using assms unfolding defined_inputs_alt_def
proof -
assume a1: "x \<notin> {t_input t |t. t \<in> FSM.transitions M2 \<and> t_source t = q2}"
assume a2: "x \<in> {t_input t |t. t \<in> FSM.transitions M1 \<and> t_source t = q1}"
assume a3: "\<And>y. [(x, y)] \<in> LS M1 q1 - LS M2 q2 \<Longrightarrow> thesis"
have f4: "\<nexists>p. x = t_input p \<and> p \<in> FSM.transitions M2 \<and> t_source p = q2"
using a1 by blast
obtain pp :: "'a \<Rightarrow> 'b \<times> 'a \<times> 'c \<times> 'b" where
"\<forall>a. ((\<nexists>p. a = t_input p \<and> p \<in> FSM.transitions M1 \<and> t_source p = q1) \<or> a = t_input (pp a) \<and> pp a \<in> FSM.transitions M1 \<and> t_source (pp a) = q1) \<and> ((\<exists>p. a = t_input p \<and> p \<in> FSM.transitions M1 \<and> t_source p = q1) \<or> (\<forall>p. a \<noteq> t_input p \<or> p \<notin> FSM.transitions M1 \<or> t_source p \<noteq> q1))"
by moura
then have "x = t_input (pp x) \<and> pp x \<in> FSM.transitions M1 \<and> t_source (pp x) = q1"
using a2 by blast
then show ?thesis
using f4 a3 by (metis (no_types) DiffI LS_single_transition)
qed
lemma language_path_append :
assumes "path M1 q1 p1"
and "io \<in> LS M1 (target q1 p1)"
shows "(p_io p1 @ io) \<in> LS M1 q1"
proof -
obtain p2 where "path M1 (target q1 p1) p2" and "p_io p2 = io"
using assms(2) by auto
then have "path M1 q1 (p1@p2)"
using assms(1) by auto
moreover have "p_io (p1@p2) = (p_io p1 @ io)"
using \<open>p_io p2 = io\<close> by auto
ultimately show ?thesis
by (metis (mono_tags, lifting) language_intro)
qed
lemma observable_defined_inputs_diff_ob :
assumes "observable M1"
and "observable M2"
and "path M1 q1 p1"
and "path M2 q2 p2"
and "p_io p1 = p_io p2"
and "x \<in> defined_inputs M1 (target q1 p1)"
and "x \<notin> defined_inputs M2 (target q2 p2)"
obtains y where "(p_io p1)@[(x,y)] \<in> LS M1 q1 - LS M2 q2"
proof -
obtain y where "[(x,y)] \<in> LS M1 (target q1 p1) - LS M2 (target q2 p2)"
using defined_inputs_language_diff[OF assms(6,7)] by blast
then have "(p_io p1)@[(x,y)] \<in> LS M1 q1"
using language_path_append[OF assms(3)]
by blast
moreover have "(p_io p1)@[(x,y)] \<notin> LS M2 q2"
by (metis (mono_tags, lifting) DiffD2 \<open>[(x, y)] \<in> LS M1 (target q1 p1) - LS M2 (target q2 p2)\<close> assms(2) assms(4) assms(5) language_state_containment observable_path_suffix)
ultimately show ?thesis
using that[of y] by blast
qed
lemma observable_defined_inputs_diff_language :
assumes "observable M1"
and "observable M2"
and "path M1 q1 p1"
and "path M2 q2 p2"
and "p_io p1 = p_io p2"
and "defined_inputs M1 (target q1 p1) \<noteq> defined_inputs M2 (target q2 p2)"
shows "LS M1 q1 \<noteq> LS M2 q2"
proof -
obtain x where "(x \<in> defined_inputs M1 (target q1 p1) - defined_inputs M2 (target q2 p2))
\<or> (x \<in> defined_inputs M2 (target q2 p2) - defined_inputs M1 (target q1 p1))"
using assms by blast
then consider "(x \<in> defined_inputs M1 (target q1 p1) - defined_inputs M2 (target q2 p2))" |
"(x \<in> defined_inputs M2 (target q2 p2) - defined_inputs M1 (target q1 p1))"
by blast
then show ?thesis
proof cases
case 1
then show ?thesis
using observable_defined_inputs_diff_ob[OF assms(1-5), of x] by blast
next
case 2
then show ?thesis
using observable_defined_inputs_diff_ob[OF assms(2,1,4,3) assms(5)[symmetric], of x] by blast
qed
qed
fun maximal_prefix_in_language :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('b \<times>'c) list \<Rightarrow> ('b \<times>'c) list" where
"maximal_prefix_in_language M q [] = []" |
"maximal_prefix_in_language M q ((x,y)#io) = (case h_obs M q x y of
None \<Rightarrow> [] |
Some q' \<Rightarrow> (x,y)#maximal_prefix_in_language M q' io)"
lemma maximal_prefix_in_language_properties :
assumes "observable M"
and "q \<in> states M"
shows "maximal_prefix_in_language M q io \<in> LS M q"
and "maximal_prefix_in_language M q io \<in> list.set (prefixes io)"
proof -
have "maximal_prefix_in_language M q io \<in> LS M q \<and> maximal_prefix_in_language M q io \<in> list.set (prefixes io)"
using assms(2) proof (induction io arbitrary: q)
case Nil
then show ?case by auto
next
case (Cons xy io)
obtain x y where "xy = (x,y)"
using prod.exhaust by metis
show ?case proof (cases "h_obs M q x y")
case None
then have "maximal_prefix_in_language M q (xy#io) = []"
unfolding \<open>xy = (x,y)\<close>
by auto
then show ?thesis
by (metis (mono_tags, lifting) Cons.prems append_self_conv2 from_FSM_language language_contains_empty_sequence mem_Collect_eq prefixes_set)
next
case (Some q')
then have *: "maximal_prefix_in_language M q (xy#io) = (x,y)#maximal_prefix_in_language M q' io"
unfolding \<open>xy = (x,y)\<close>
by auto
have "q' \<in> states M"
using h_obs_state[OF Some] by auto
then have "maximal_prefix_in_language M q' io \<in> LS M q'"
and "maximal_prefix_in_language M q' io \<in> list.set (prefixes io)"
using Cons.IH by auto
have "maximal_prefix_in_language M q (xy # io) \<in> LS M q"
unfolding *
using Some \<open>maximal_prefix_in_language M q' io \<in> LS M q'\<close>
by (meson assms(1) h_obs_language_iff)
moreover have "maximal_prefix_in_language M q (xy # io) \<in> list.set (prefixes (xy # io))"
unfolding *
unfolding \<open>xy = (x,y)\<close>
using \<open>maximal_prefix_in_language M q' io \<in> list.set (prefixes io)\<close> append_Cons
unfolding prefixes_set
by auto
ultimately show ?thesis
by blast
qed
qed
then show "maximal_prefix_in_language M q io \<in> LS M q"
and "maximal_prefix_in_language M q io \<in> list.set (prefixes io)"
by auto
qed
subsection \<open>Further Reachability Formalisations\<close>
(* states that are reachable in at most k transitions *)
fun reachable_k :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a set" where
"reachable_k M q n = {target q p | p . path M q p \<and> length p \<le> n}"
lemma reachable_k_0_initial : "reachable_k M (initial M) 0 = {initial M}"
by auto
lemma reachable_k_states : "reachable_states M = reachable_k M (initial M) ( size M - 1)"
proof -
have "\<And>q. q \<in> reachable_states M \<Longrightarrow> q \<in> reachable_k M (initial M) ( size M - 1)"
proof -
fix q assume "q \<in> reachable_states M"
then obtain p where "path M (initial M) p" and "target (initial M) p = q"
unfolding reachable_states_def by blast
then obtain p' where "path M (initial M) p'"
and "target (initial M) p' = target (initial M) p"
and "length p' < size M"
by (metis acyclic_path_from_cyclic_path acyclic_path_length_limit)
then show "q \<in> reachable_k M (initial M) ( size M - 1)"
using \<open>target (FSM.initial M) p = q\<close> less_trans by auto
qed
moreover have "\<And>x. x \<in> reachable_k M (initial M) ( size M - 1) \<Longrightarrow> x \<in> reachable_states M"
unfolding reachable_states_def reachable_k.simps by blast
ultimately show ?thesis by blast
qed
subsubsection \<open>Induction Schemes\<close>
lemma acyclic_induction [consumes 1, case_names reachable_state]:
assumes "acyclic M"
and "\<And> q . q \<in> reachable_states M \<Longrightarrow> (\<And> t . t \<in> transitions M \<Longrightarrow> ((t_source t = q) \<Longrightarrow> P (t_target t))) \<Longrightarrow> P q"
shows "\<forall> q \<in> reachable_states M . P q"
proof
fix q assume "q \<in> reachable_states M"
let ?k = "Max (image length {p . path M q p})"
have "finite {p . path M q p}" using acyclic_finite_paths_from_reachable_state[OF assms(1)]
using \<open>q \<in> reachable_states M\<close> unfolding reachable_states_def by force
then have k_prop: "(\<forall> p . path M q p \<longrightarrow> length p \<le> ?k)" by auto
moreover have "\<And> q k . q \<in> reachable_states M \<Longrightarrow> (\<forall> p . path M q p \<longrightarrow> length p \<le> k) \<Longrightarrow> P q"
proof -
fix q k assume "q \<in> reachable_states M" and "(\<forall> p . path M q p \<longrightarrow> length p \<le> k)"
then show "P q"
proof (induction k arbitrary: q)
case 0
then have "{p . path M q p} = {[]}" using reachable_state_is_state[OF \<open>q \<in> reachable_states M\<close>]
by blast
then have "LS M q \<subseteq> {[]}" unfolding LS.simps by blast
then have "deadlock_state M q" using deadlock_state_alt_def by metis
then show ?case using assms(2)[OF \<open>q \<in> reachable_states M\<close>] unfolding deadlock_state.simps by blast
next
case (Suc k)
have "\<And> t . t \<in> transitions M \<Longrightarrow> (t_source t = q) \<Longrightarrow> P (t_target t)"
proof -
fix t assume "t \<in> transitions M" and "t_source t = q"
then have "t_target t \<in> reachable_states M"
using \<open>q \<in> reachable_states M\<close> using reachable_states_next by metis
moreover have "\<forall>p. path M (t_target t) p \<longrightarrow> length p \<le> k"
using Suc.prems(2) \<open>t \<in> transitions M\<close> \<open>t_source t = q\<close> by auto
ultimately show "P (t_target t)"
using Suc.IH unfolding reachable_states_def by blast
qed
then show ?case using assms(2)[OF Suc.prems(1)] by blast
qed
qed
ultimately show "P q" using \<open>q \<in> reachable_states M\<close> by blast
qed
lemma reachable_states_induct [consumes 1, case_names init transition] :
assumes "q \<in> reachable_states M"
and "P (initial M)"
and "\<And> t . t \<in> transitions M \<Longrightarrow> t_source t \<in> reachable_states M \<Longrightarrow> P (t_source t) \<Longrightarrow> P (t_target t)"
shows "P q"
proof -
from assms(1) obtain p where "path M (initial M) p" and "target (initial M) p = q"
unfolding reachable_states_def by auto
then show "P q"
proof (induction p arbitrary: q rule: rev_induct)
case Nil
then show ?case using assms(2) by auto
next
case (snoc t p)
then have "target (initial M) p = t_source t"
by auto
then have "P (t_source t)"
using snoc.IH snoc.prems by auto
moreover have "t \<in> transitions M"
using snoc.prems by auto
moreover have "t_source t \<in> reachable_states M"
by (metis \<open>target (FSM.initial M) p = t_source t\<close> path_prefix reachable_states_intro snoc.prems(1))
moreover have "t_target t = q"
using snoc.prems by auto
ultimately show ?case
using assms(3) by blast
qed
qed
lemma reachable_states_cases [consumes 1, case_names init transition] :
assumes "q \<in> reachable_states M"
and "P (initial M)"
and "\<And> t . t \<in> transitions M \<Longrightarrow> t_source t \<in> reachable_states M \<Longrightarrow> P (t_target t)"
shows "P q"
by (metis assms(1) assms(2) assms(3) reachable_states_induct)
subsection \<open>Further Path Enumeration Algorithms\<close>
fun paths_for_input' :: "('a \<Rightarrow> ('b \<times> 'c \<times> 'a) set) \<Rightarrow> 'b list \<Rightarrow> 'a \<Rightarrow> ('a,'b,'c) path \<Rightarrow> ('a,'b,'c) path set" where
"paths_for_input' f [] q prev = {prev}" |
"paths_for_input' f (x#xs) q prev = \<Union>(image (\<lambda>(x',y',q') . paths_for_input' f xs q' (prev@[(q,x,y',q')])) (Set.filter (\<lambda>(x',y',q') . x' = x) (f q)))"
lemma paths_for_input'_set :
assumes "q \<in> states M"
shows "paths_for_input' (h_from M) xs q prev = {prev@p | p . path M q p \<and> map fst (p_io p) = xs}"
using assms proof (induction xs arbitrary: q prev)
case Nil
then show ?case by auto
next
case (Cons x xs)
let ?UN = "\<Union>(image (\<lambda>(x',y',q') . paths_for_input' (h_from M) xs q' (prev@[(q,x,y',q')])) (Set.filter (\<lambda>(x',y',q') . x' = x) (h_from M q)))"
have "?UN = {prev@p | p . path M q p \<and> map fst (p_io p) = x#xs}"
proof
have "\<And> p . p \<in> ?UN \<Longrightarrow> p \<in> {prev@p | p . path M q p \<and> map fst (p_io p) = x#xs}"
proof -
fix p assume "p \<in> ?UN"
then obtain y' q' where "(x,y',q') \<in> (Set.filter (\<lambda>(x',y',q') . x' = x) (h_from M q))"
and "p \<in> paths_for_input' (h_from M) xs q' (prev@[(q,x,y',q')])"
by auto
from \<open>(x,y',q') \<in> (Set.filter (\<lambda>(x',y',q') . x' = x) (h_from M q))\<close> have "q' \<in> states M" and "(q,x,y',q') \<in> transitions M"
using fsm_transition_target unfolding h.simps by auto
have "p \<in> {(prev @ [(q, x, y', q')]) @ p |p. path M q' p \<and> map fst (p_io p) = xs}"
using \<open>p \<in> paths_for_input' (h_from M) xs q' (prev@[(q,x,y',q')])\<close>
unfolding Cons.IH[OF \<open>q' \<in> states M\<close>] by assumption
moreover have "{(prev @ [(q, x, y', q')]) @ p |p. path M q' p \<and> map fst (p_io p) = xs}
\<subseteq> {prev@p | p . path M q p \<and> map fst (p_io p) = x#xs}"
using \<open>(q,x,y',q') \<in> transitions M\<close>
using cons by force
ultimately show "p \<in> {prev@p | p . path M q p \<and> map fst (p_io p) = x#xs}"
by blast
qed
then show "?UN \<subseteq> {prev@p | p . path M q p \<and> map fst (p_io p) = x#xs}"
by blast
have "\<And> p . p \<in> {prev@p | p . path M q p \<and> map fst (p_io p) = x#xs} \<Longrightarrow> p \<in> ?UN"
proof -
fix pp assume "pp \<in> {prev@p | p . path M q p \<and> map fst (p_io p) = x#xs}"
then obtain p where "pp = prev@p" and "path M q p" and "map fst (p_io p) = x#xs"
by fastforce
then obtain t p' where "p = t#p'" and "path M q (t#p')" and "map fst (p_io (t#p')) = x#xs" and "map fst (p_io p') = xs"
by (metis (no_types, lifting) map_eq_Cons_D)
then have "path M (t_target t) p'" and "t_source t = q" and "t_input t = x" and "t_target t \<in> states M" and "t \<in> transitions M"
by auto
have "(x,t_output t,t_target t) \<in> (Set.filter (\<lambda>(x',y',q') . x' = x) (h_from M q))"
using \<open>t \<in> transitions M\<close> \<open>t_input t = x\<close> \<open>t_source t = q\<close>
unfolding h.simps by auto
moreover have "(prev@p) \<in> paths_for_input' (h_from M) xs (t_target t) (prev@[(q,x,t_output t,t_target t)])"
using Cons.IH[OF \<open>t_target t \<in> states M\<close>, of "prev@[(q, x, t_output t, t_target t)]"]
using \<open>\<And>thesis. (\<And>t p'. \<lbrakk>p = t # p'; path M q (t # p'); map fst (p_io (t # p')) = x # xs; map fst (p_io p') = xs\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
\<open>p = t # p'\<close>
\<open>paths_for_input' (h_from M) xs (t_target t) (prev @ [(q, x, t_output t, t_target t)])
= {(prev @ [(q, x, t_output t, t_target t)]) @ p |p. path M (t_target t) p \<and> map fst (p_io p) = xs}\<close>
\<open>t_input t = x\<close>
\<open>t_source t = q\<close>
by fastforce
ultimately show "pp \<in> ?UN" unfolding \<open>pp = prev@p\<close>
by blast
qed
then show "{prev@p | p . path M q p \<and> map fst (p_io p) = x#xs} \<subseteq> ?UN"
by (meson subsetI)
qed
then show ?case
by (metis paths_for_input'.simps(2))
qed
definition paths_for_input :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> ('a,'b,'c) path set" where
"paths_for_input M q xs = {p . path M q p \<and> map fst (p_io p) = xs}"
lemma paths_for_input_set_code[code] :
"paths_for_input M q xs = (if q \<in> states M then paths_for_input' (h_from M) xs q [] else {})"
using paths_for_input'_set[of q M xs "[]"]
unfolding paths_for_input_def
by (cases "q \<in> states M"; auto; simp add: path_begin_state)
fun paths_up_to_length_or_condition_with_witness' ::
"('a \<Rightarrow> ('b \<times> 'c \<times> 'a) set) \<Rightarrow> (('a,'b,'c) path \<Rightarrow> 'd option) \<Rightarrow> ('a,'b,'c) path \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> (('a,'b,'c) path \<times> 'd) set"
where
"paths_up_to_length_or_condition_with_witness' f P prev 0 q = (case P prev of Some w \<Rightarrow> {(prev,w)} | None \<Rightarrow> {})" |
"paths_up_to_length_or_condition_with_witness' f P prev (Suc k) q = (case P prev of
Some w \<Rightarrow> {(prev,w)} |
None \<Rightarrow> (\<Union>(image (\<lambda>(x,y,q') . paths_up_to_length_or_condition_with_witness' f P (prev@[(q,x,y,q')]) k q') (f q))))"
lemma paths_up_to_length_or_condition_with_witness'_set :
assumes "q \<in> states M"
shows "paths_up_to_length_or_condition_with_witness' (h_from M) P prev k q
= {(prev@p,x) | p x . path M q p
\<and> length p \<le> k
\<and> P (prev@p) = Some x
\<and> (\<forall> p' p'' . (p = p'@p'' \<and> p'' \<noteq> []) \<longrightarrow> P (prev@p') = None)}"
using assms proof (induction k arbitrary: q prev)
case 0
then show ?case proof (cases "P prev")
case None then show ?thesis by auto
next
case (Some w)
then show ?thesis by (simp add: "0.prems" nil)
qed
next
case (Suc k)
then show ?case proof (cases "P prev")
case (Some w)
then have "(prev,w) \<in> {(prev@p,x) | p x . path M q p
\<and> length p \<le> Suc k
\<and> P (prev@p) = Some x
\<and> (\<forall> p' p'' . (p = p'@p'' \<and> p'' \<noteq> []) \<longrightarrow> P (prev@p') = None)}"
by (simp add: Suc.prems nil)
then have "{(prev@p,x) | p x . path M q p
\<and> length p \<le> Suc k
\<and> P (prev@p) = Some x
\<and> (\<forall> p' p'' . (p = p'@p'' \<and> p'' \<noteq> []) \<longrightarrow> P (prev@p') = None)}
= {(prev,w)}"
using Some by fastforce
then show ?thesis using Some by auto
next
case None
have "(\<Union>(image (\<lambda>(x,y,q') . paths_up_to_length_or_condition_with_witness' (h_from M) P (prev@[(q,x,y,q')]) k q') (h_from M q)))
= {(prev@p,x) | p x . path M q p
\<and> length p \<le> Suc k
\<and> P (prev@p) = Some x
\<and> (\<forall> p' p'' . (p = p'@p'' \<and> p'' \<noteq> []) \<longrightarrow> P (prev@p') = None)}"
(is "?UN = ?PX")
proof -
have *: "\<And> pp . pp \<in> ?UN \<Longrightarrow> pp \<in> ?PX"
proof -
fix pp assume "pp \<in> ?UN"
then obtain x y q' where "(x,y,q') \<in> h_from M q"
and "pp \<in> paths_up_to_length_or_condition_with_witness' (h_from M) P (prev@[(q,x,y,q')]) k q'"
by blast
then have "(q,x,y,q') \<in> transitions M" by auto
then have "q' \<in> states M" using fsm_transition_target by auto
obtain p w where "pp = ((prev@[(q,x,y,q')])@p,w)"
and "path M q' p"
and "length p \<le> k"
and "P ((prev @ [(q, x, y, q')]) @ p) = Some w"
and "\<And> p' p''. p = p' @ p'' \<Longrightarrow> p'' \<noteq> [] \<Longrightarrow> P ((prev @ [(q, x, y, q')]) @ p') = None"
using \<open>pp \<in> paths_up_to_length_or_condition_with_witness' (h_from M) P (prev@[(q,x,y,q')]) k q'\<close>
unfolding Suc.IH[OF \<open>q' \<in> states M\<close>, of "prev@[(q,x,y,q')]"]
by blast
have "path M q ((q,x,y,q')#p)"
using \<open>path M q' p\<close> \<open>(q,x,y,q') \<in> transitions M\<close> by (simp add: path_prepend_t)
moreover have "length ((q,x,y,q')#p) \<le> Suc k"
using \<open>length p \<le> k\<close> by auto
moreover have "P (prev @ ([(q, x, y, q')] @ p)) = Some w"
using \<open>P ((prev @ [(q, x, y, q')]) @ p) = Some w\<close> by auto
moreover have "\<And> p' p''. ((q,x,y,q')#p) = p' @ p'' \<Longrightarrow> p'' \<noteq> [] \<Longrightarrow> P (prev @ p') = None"
using \<open>\<And> p' p''. p = p' @ p'' \<Longrightarrow> p'' \<noteq> [] \<Longrightarrow> P ((prev @ [(q, x, y, q')]) @ p') = None\<close>
using None
by (metis (no_types, opaque_lifting) append.simps(1) append_Cons append_Nil2 append_assoc
list.inject neq_Nil_conv)
ultimately show "pp \<in> ?PX"
unfolding \<open>pp = ((prev@[(q,x,y,q')])@p,w)\<close> by auto
qed
have **: "\<And> pp . pp \<in> ?PX \<Longrightarrow> pp \<in> ?UN"
proof -
fix pp assume "pp \<in> ?PX"
then obtain p' w where "pp = (prev @ p', w)"
and "path M q p'"
and "length p' \<le> Suc k"
and "P (prev @ p') = Some w"
and "\<And> p' p''. p' = p' @ p'' \<Longrightarrow> p'' \<noteq> [] \<Longrightarrow> P (prev @ p') = None"
by blast
moreover obtain t p where "p' = t#p" using \<open>P (prev @ p') = Some w\<close> using None
by (metis append_Nil2 list.exhaust option.distinct(1))
have "pp = ((prev @ [t])@p, w)"
using \<open>pp = (prev @ p', w)\<close> unfolding \<open>p' = t#p\<close> by auto
have "path M q (t#p)"
using \<open>path M q p'\<close> unfolding \<open>p' = t#p\<close> by auto
have p2: "length (t#p) \<le> Suc k"
using \<open>length p' \<le> Suc k\<close> unfolding \<open>p' = t#p\<close> by auto
have p3: "P ((prev @ [t])@p) = Some w"
using \<open>P (prev @ p') = Some w\<close> unfolding \<open>p' = t#p\<close> by auto
have p4: "\<And> p' p''. p = p' @ p'' \<Longrightarrow> p'' \<noteq> [] \<Longrightarrow> P ((prev@[t]) @ p') = None"
using \<open>\<And> p' p''. p' = p' @ p'' \<Longrightarrow> p'' \<noteq> [] \<Longrightarrow> P (prev @ p') = None\<close> \<open>pp \<in> ?PX\<close>
unfolding \<open>pp = ((prev @ [t]) @ p, w)\<close> \<open>p' = t#p\<close>
by auto
have "t \<in> transitions M" and p1: "path M (t_target t) p" and "t_source t = q"
using \<open>path M q (t#p)\<close> by auto
then have "t_target t \<in> states M"
and "(t_input t, t_output t, t_target t) \<in> h_from M q"
and "t_source t = q"
using fsm_transition_target by auto
then have "t = (q,t_input t, t_output t, t_target t)"
by auto
have "((prev @ [t])@p, w) \<in> paths_up_to_length_or_condition_with_witness' (h_from M) P (prev@[t]) k (t_target t)"
unfolding Suc.IH[OF \<open>t_target t \<in> states M\<close>, of "prev@[t]"]
using p1 p2 p3 p4 by auto
then show "pp \<in> ?UN"
unfolding \<open>pp = ((prev @ [t])@p, w)\<close>
proof -
have "paths_up_to_length_or_condition_with_witness' (h_from M) P (prev @ [t]) k (t_target t)
= paths_up_to_length_or_condition_with_witness' (h_from M) P (prev @ [(q, t_input t, t_output t, t_target t)]) k (t_target t)"
using \<open>t = (q, t_input t, t_output t, t_target t)\<close> by presburger
then show "((prev @ [t]) @ p, w) \<in> (\<Union>(b, c, a)\<in>h_from M q. paths_up_to_length_or_condition_with_witness' (h_from M) P (prev @ [(q, b, c, a)]) k a)"
using \<open>((prev @ [t]) @ p, w) \<in> paths_up_to_length_or_condition_with_witness' (h_from M) P (prev @ [t]) k (t_target t)\<close>
\<open>(t_input t, t_output t, t_target t) \<in> h_from M q\<close>
by blast
qed
qed
show ?thesis
using subsetI[of ?UN ?PX, OF *] subsetI[of ?PX ?UN, OF **] subset_antisym by blast
qed
then show ?thesis
using None unfolding paths_up_to_length_or_condition_with_witness'.simps by simp
qed
qed
definition paths_up_to_length_or_condition_with_witness ::
"('a,'b,'c) fsm \<Rightarrow> (('a,'b,'c) path \<Rightarrow> 'd option) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> (('a,'b,'c) path \<times> 'd) set"
where
"paths_up_to_length_or_condition_with_witness M P k q
= {(p,x) | p x . path M q p
\<and> length p \<le> k
\<and> P p = Some x
\<and> (\<forall> p' p'' . (p = p'@p'' \<and> p'' \<noteq> []) \<longrightarrow> P p' = None)}"
lemma paths_up_to_length_or_condition_with_witness_code[code] :
"paths_up_to_length_or_condition_with_witness M P k q
= (if q \<in> states M then paths_up_to_length_or_condition_with_witness' (h_from M) P [] k q
else {})"
proof (cases "q \<in> states M")
case True
then show ?thesis
unfolding paths_up_to_length_or_condition_with_witness_def
paths_up_to_length_or_condition_with_witness'_set[OF True]
by auto
next
case False
then show ?thesis
unfolding paths_up_to_length_or_condition_with_witness_def
using path_begin_state by fastforce
qed
lemma paths_up_to_length_or_condition_with_witness_finite :
"finite (paths_up_to_length_or_condition_with_witness M P k q)"
proof -
have "paths_up_to_length_or_condition_with_witness M P k q
\<subseteq> {(p, the (P p)) | p . path M q p \<and> length p \<le> k}"
unfolding paths_up_to_length_or_condition_with_witness_def
by auto
moreover have "finite {(p, the (P p)) | p . path M q p \<and> length p \<le> k}"
using paths_finite[of M q k]
by simp
ultimately show ?thesis
using rev_finite_subset by auto
qed
subsection \<open>More Acyclicity Properties\<close>
lemma maximal_path_target_deadlock :
assumes "path M (initial M) p"
and "\<not>(\<exists> p' . path M (initial M) p' \<and> is_prefix p p' \<and> p \<noteq> p')"
shows "deadlock_state M (target (initial M) p)"
proof -
have "\<not>(\<exists> t \<in> transitions M . t_source t = target (initial M) p)"
using assms(2) unfolding is_prefix_prefix
by (metis append_Nil2 assms(1) not_Cons_self2 path_append_transition same_append_eq)
then show ?thesis by auto
qed
lemma path_to_deadlock_is_maximal :
assumes "path M (initial M) p"
and "deadlock_state M (target (initial M) p)"
shows "\<not>(\<exists> p' . path M (initial M) p' \<and> is_prefix p p' \<and> p \<noteq> p')"
proof
assume "\<exists>p'. path M (initial M) p' \<and> is_prefix p p' \<and> p \<noteq> p'"
then obtain p' where "path M (initial M) p'" and "is_prefix p p'" and "p \<noteq> p'" by blast
then have "length p' > length p"
unfolding is_prefix_prefix by auto
then obtain t p2 where "p' = p @ [t] @ p2"
using \<open>is_prefix p p'\<close> unfolding is_prefix_prefix
by (metis \<open>p \<noteq> p'\<close> append.left_neutral append_Cons append_Nil2 non_sym_dist_pairs'.cases)
then have "path M (initial M) (p@[t])"
using \<open>path M (initial M) p'\<close> by auto
then have "t \<in> transitions M" and "t_source t = target (initial M) p"
by auto
then show "False"
using \<open>deadlock_state M (target (initial M) p)\<close> unfolding deadlock_state.simps by blast
qed
definition maximal_acyclic_paths :: "('a,'b,'c) fsm \<Rightarrow> ('a,'b,'c) path set" where
"maximal_acyclic_paths M = {p . path M (initial M) p
\<and> distinct (visited_states (initial M) p)
\<and> \<not>(\<exists> p' . p' \<noteq> [] \<and> path M (initial M) (p@p')
\<and> distinct (visited_states (initial M) (p@p')))}"
(* very inefficient construction *)
lemma maximal_acyclic_paths_code[code] :
"maximal_acyclic_paths M = (let ps = acyclic_paths_up_to_length M (initial M) (size M - 1)
in Set.filter (\<lambda>p . \<not> (\<exists> p' \<in> ps . p' \<noteq> p \<and> is_prefix p p')) ps)"
proof -
have scheme1: "\<And> P p . (\<exists> p' . p' \<noteq> [] \<and> P (p@p')) = (\<exists> p' \<in> {p . P p} . p' \<noteq> p \<and> is_prefix p p')"
unfolding is_prefix_prefix by blast
have scheme2: "\<And> p . (path M (FSM.initial M) p
\<and> length p \<le> FSM.size M - 1
\<and> distinct (visited_states (FSM.initial M) p))
= (path M (FSM.initial M) p \<and> distinct (visited_states (FSM.initial M) p))"
using acyclic_path_length_limit by fastforce
show ?thesis
unfolding maximal_acyclic_paths_def acyclic_paths_up_to_length.simps Let_def
unfolding scheme1[of "\<lambda>p . path M (initial M) p \<and> distinct (visited_states (initial M) p)"]
unfolding scheme2 by fastforce
qed
lemma maximal_acyclic_path_deadlock :
assumes "acyclic M"
and "path M (initial M) p"
shows "\<not>(\<exists> p' . p' \<noteq> [] \<and> path M (initial M) (p@p') \<and> distinct (visited_states (initial M) (p@p')))
= deadlock_state M (target (initial M) p)"
proof -
have "deadlock_state M (target (initial M) p) \<Longrightarrow> \<not>(\<exists> p' . p' \<noteq> [] \<and> path M (initial M) (p@p')
\<and> distinct (visited_states (initial M) (p@p')))"
unfolding deadlock_state.simps
using assms(2) by (metis path.cases path_suffix)
then show ?thesis
by (metis acyclic.elims(2) assms(1) assms(2) is_prefix_prefix maximal_path_target_deadlock
self_append_conv)
qed
lemma maximal_acyclic_paths_deadlock_targets :
assumes "acyclic M"
shows "maximal_acyclic_paths M
= { p . path M (initial M) p \<and> deadlock_state M (target (initial M) p)}"
using maximal_acyclic_path_deadlock[OF assms]
unfolding maximal_acyclic_paths_def
by (metis (no_types, lifting) acyclic.elims(2) assms)
lemma cycle_from_cyclic_path :
assumes "path M q p"
and "\<not> distinct (visited_states q p)"
obtains i j where
"take j (drop i p) \<noteq> []"
"target (target q (take i p)) (take j (drop i p)) = (target q (take i p))"
"path M (target q (take i p)) (take j (drop i p))"
proof -
obtain i j where "i < j" and "j < length (visited_states q p)"
and "(visited_states q p) ! i = (visited_states q p) ! j"
using assms(2) non_distinct_repetition_indices by blast
have "(target q (take i p)) = (visited_states q p) ! i"
using \<open>i < j\<close> \<open>j < length (visited_states q p)\<close>
by (metis less_trans take_last_index target.simps visited_states_take)
then have "(target q (take i p)) = (visited_states q p) ! j"
using \<open>(visited_states q p) ! i = (visited_states q p) ! j\<close> by auto
have p1: "take (j-i) (drop i p) \<noteq> []"
using \<open>i < j\<close> \<open>j < length (visited_states q p)\<close> by auto
have "target (target q (take i p)) (take (j-i) (drop i p)) = (target q (take j p))"
using \<open>i < j\<close> by (metis add_diff_inverse_nat less_asym' path_append_target take_add)
then have p2: "target (target q (take i p)) (take (j-i) (drop i p)) = (target q (take i p))"
using \<open>(target q (take i p)) = (visited_states q p) ! i\<close>
using \<open>(target q (take i p)) = (visited_states q p) ! j\<close>
by (metis \<open>j < length (visited_states q p)\<close> take_last_index target.simps visited_states_take)
have p3: "path M (target q (take i p)) (take (j-i) (drop i p))"
by (metis append_take_drop_id assms(1) path_append_elim)
show ?thesis using p1 p2 p3 that by blast
qed
lemma acyclic_single_deadlock_reachable :
assumes "acyclic M"
and "\<And> q' . q' \<in> reachable_states M \<Longrightarrow> q' = qd \<or> \<not> deadlock_state M q'"
shows "qd \<in> reachable_states M"
using acyclic_deadlock_reachable[OF assms(1)]
using assms(2) by auto
lemma acyclic_paths_to_single_deadlock :
assumes "acyclic M"
and "\<And> q' . q' \<in> reachable_states M \<Longrightarrow> q' = qd \<or> \<not> deadlock_state M q'"
and "q \<in> reachable_states M"
obtains p where "path M q p" and "target q p = qd"
proof -
have "q \<in> states M" using assms(3) reachable_state_is_state by metis
have "acyclic (from_FSM M q)"
using from_FSM_acyclic[OF assms(3,1)] by assumption
have *: "(\<And>q'. q' \<in> reachable_states (FSM.from_FSM M q)
\<Longrightarrow> q' = qd \<or> \<not> deadlock_state (FSM.from_FSM M q) q')"
using assms(2) from_FSM_reachable_states[OF assms(3)]
unfolding deadlock_state.simps from_FSM_simps[OF \<open>q \<in> states M\<close>] by blast
obtain p where "path (from_FSM M q) q p" and "target q p = qd"
using acyclic_single_deadlock_reachable[OF \<open>acyclic (from_FSM M q)\<close> *]
unfolding reachable_states_def from_FSM_simps[OF \<open>q \<in> states M\<close>]
by blast
then show ?thesis
using that by (metis \<open>q \<in> FSM.states M\<close> from_FSM_path)
qed
subsection \<open>Elements as Lists\<close>
fun states_as_list :: "('a :: linorder, 'b, 'c) fsm \<Rightarrow> 'a list" where
"states_as_list M = sorted_list_of_set (states M)"
lemma states_as_list_distinct : "distinct (states_as_list M)" by auto
lemma states_as_list_set : "set (states_as_list M) = states M"
by (simp add: fsm_states_finite)
fun reachable_states_as_list :: "('a :: linorder, 'b, 'c) fsm \<Rightarrow> 'a list" where
"reachable_states_as_list M = sorted_list_of_set (reachable_states M)"
lemma reachable_states_as_list_distinct : "distinct (reachable_states_as_list M)" by auto
lemma reachable_states_as_list_set : "set (reachable_states_as_list M) = reachable_states M"
by (metis fsm_states_finite infinite_super reachable_state_is_state reachable_states_as_list.simps
set_sorted_list_of_set subsetI)
fun inputs_as_list :: "('a, 'b :: linorder, 'c) fsm \<Rightarrow> 'b list" where
"inputs_as_list M = sorted_list_of_set (inputs M)"
lemma inputs_as_list_set : "set (inputs_as_list M) = inputs M"
by (simp add: fsm_inputs_finite)
lemma inputs_as_list_distinct : "distinct (inputs_as_list M)" by auto
fun transitions_as_list :: "('a :: linorder,'b :: linorder,'c :: linorder) fsm \<Rightarrow> ('a,'b,'c) transition list" where
"transitions_as_list M = sorted_list_of_set (transitions M)"
lemma transitions_as_list_set : "set (transitions_as_list M) = transitions M"
by (simp add: fsm_transitions_finite)
fun outputs_as_list :: "('a,'b,'c :: linorder) fsm \<Rightarrow> 'c list" where
"outputs_as_list M = sorted_list_of_set (outputs M)"
lemma outputs_as_list_set : "set (outputs_as_list M) = outputs M"
by (simp add: fsm_outputs_finite)
fun ftransitions :: "('a :: linorder,'b :: linorder,'c :: linorder) fsm \<Rightarrow> ('a,'b,'c) transition fset" where
"ftransitions M = fset_of_list (transitions_as_list M)"
fun fstates :: "('a :: linorder,'b,'c) fsm \<Rightarrow> 'a fset" where
"fstates M = fset_of_list (states_as_list M)"
fun finputs :: "('a,'b :: linorder,'c) fsm \<Rightarrow> 'b fset" where
"finputs M = fset_of_list (inputs_as_list M)"
fun foutputs :: "('a,'b,'c :: linorder) fsm \<Rightarrow> 'c fset" where
"foutputs M = fset_of_list (outputs_as_list M)"
lemma fstates_set : "fset (fstates M) = states M"
using fsm_states_finite[of M] by (simp add: fset_of_list.rep_eq)
lemma finputs_set : "fset (finputs M) = inputs M"
using fsm_inputs_finite[of M] by (simp add: fset_of_list.rep_eq)
lemma foutputs_set : "fset (foutputs M) = outputs M"
using fsm_outputs_finite[of M] by (simp add: fset_of_list.rep_eq)
lemma ftransitions_set: "fset (ftransitions M) = transitions M"
by (metis (no_types) fset_of_list.rep_eq ftransitions.simps transitions_as_list_set)
lemma ftransitions_source:
"q |\<in>| (t_source |`| ftransitions M) \<Longrightarrow> q \<in> states M"
using ftransitions_set[of M] fsm_transition_source[of _ M]
- by (metis (no_types, lifting) fimageE fmember_iff_member_fset)
+ by (metis (no_types, opaque_lifting) fimageE)
lemma ftransitions_target:
"q |\<in>| (t_target |`| ftransitions M) \<Longrightarrow> q \<in> states M"
using ftransitions_set[of M] fsm_transition_target[of _ M]
- by (metis (no_types, lifting) fimageE fmember_iff_member_fset)
+ by (metis (no_types, lifting) fimageE)
subsection \<open>Responses to Input Sequences\<close>
fun language_for_input :: "('a::linorder,'b::linorder,'c::linorder) fsm \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> ('b\<times>'c) list list" where
"language_for_input M q [] = [[]]" |
"language_for_input M q (x#xs) =
(let outs = outputs_as_list M
in concat (map (\<lambda>y . case h_obs M q x y of None \<Rightarrow> [] | Some q' \<Rightarrow> map ((#) (x,y)) (language_for_input M q' xs)) outs))"
lemma language_for_input_set :
assumes "observable M"
and "q \<in> states M"
shows "list.set (language_for_input M q xs) = {io . io \<in> LS M q \<and> map fst io = xs}"
using assms(2) proof (induction xs arbitrary: q)
case Nil
then show ?case by auto
next
case (Cons x xs)
have "list.set (language_for_input M q (x#xs)) \<subseteq> {io . io \<in> LS M q \<and> map fst io = (x#xs)}"
proof
fix io assume "io \<in> list.set (language_for_input M q (x#xs))"
then obtain y where "y \<in> outputs M"
and "io \<in> list.set (case h_obs M q x y of None \<Rightarrow> [] | Some q' \<Rightarrow> map ((#) (x,y)) (language_for_input M q' xs))"
unfolding outputs_as_list_set[symmetric]
by auto
then obtain q' where "h_obs M q x y = Some q'" and "io \<in> list.set (map ((#) (x,y)) (language_for_input M q' xs))"
by (cases "h_obs M q x y"; auto)
then obtain io' where "io = (x,y)#io'"
and "io' \<in> list.set (language_for_input M q' xs)"
by auto
then have "io' \<in> LS M q'" and "map fst io' = xs"
using Cons.IH[OF h_obs_state[OF \<open>h_obs M q x y = Some q'\<close>]]
by blast+
then have "(x,y)#io' \<in> LS M q"
using \<open>h_obs M q x y = Some q'\<close>
unfolding h_obs_language_iff[OF assms(1), of x y io' q]
by blast
then show "io \<in> {io . io \<in> LS M q \<and> map fst io = (x#xs)}"
unfolding \<open>io = (x,y)#io'\<close>
using \<open>map fst io' = xs\<close>
by auto
qed
moreover have "{io . io \<in> LS M q \<and> map fst io = (x#xs)} \<subseteq> list.set (language_for_input M q (x#xs))"
proof
have scheme : "\<And> x y f xs . y \<in> list.set (f x) \<Longrightarrow> x \<in> list.set xs \<Longrightarrow> y \<in> list.set (concat (map f xs))"
by auto
fix io assume "io \<in> {io . io \<in> LS M q \<and> map fst io = (x#xs)}"
then have "io \<in> LS M q" and "map fst io = (x#xs)"
by auto
then obtain y io' where "io = (x,y)#io'"
by fastforce
then have "(x,y)#io' \<in> LS M q"
using \<open>io \<in> LS M q\<close>
by auto
then obtain q' where "h_obs M q x y = Some q'" and "io' \<in> LS M q'"
unfolding h_obs_language_iff[OF assms(1), of x y io' q]
by blast
moreover have "io' \<in> list.set (language_for_input M q' xs)"
using Cons.IH[OF h_obs_state[OF \<open>h_obs M q x y = Some q'\<close>]] \<open>io' \<in> LS M q'\<close> \<open>map fst io = (x#xs)\<close>
unfolding \<open>io = (x,y)#io'\<close> by auto
ultimately have "io \<in> list.set ((\<lambda> y .(case h_obs M q x y of None \<Rightarrow> [] | Some q' \<Rightarrow> map ((#) (x,y)) (language_for_input M q' xs))) y)"
unfolding \<open>io = (x,y)#io'\<close>
by force
moreover have "y \<in> list.set (outputs_as_list M)"
unfolding outputs_as_list_set
using language_io(2)[OF \<open>(x,y)#io' \<in> LS M q\<close>] by auto
ultimately show "io \<in> list.set (language_for_input M q (x#xs))"
unfolding language_for_input.simps Let_def
using scheme[of io "(\<lambda> y .(case h_obs M q x y of None \<Rightarrow> [] | Some q' \<Rightarrow> map ((#) (x,y)) (language_for_input M q' xs)))" y]
by blast
qed
ultimately show ?case
by blast
qed
subsection \<open>Filtering Transitions\<close>
lift_definition filter_transitions ::
"('a,'b,'c) fsm \<Rightarrow> (('a,'b,'c) transition \<Rightarrow> bool) \<Rightarrow> ('a,'b,'c) fsm" is FSM_Impl.filter_transitions
proof -
fix M :: "('a,'b,'c) fsm_impl"
fix P :: "('a,'b,'c) transition \<Rightarrow> bool"
assume "well_formed_fsm M"
then show "well_formed_fsm (FSM_Impl.filter_transitions M P)"
unfolding FSM_Impl.filter_transitions.simps by force
qed
lemma filter_transitions_simps[simp] :
"initial (filter_transitions M P) = initial M"
"states (filter_transitions M P) = states M"
"inputs (filter_transitions M P) = inputs M"
"outputs (filter_transitions M P) = outputs M"
"transitions (filter_transitions M P) = {t \<in> transitions M . P t}"
by (transfer;auto)+
lemma filter_transitions_submachine :
"is_submachine (filter_transitions M P) M"
unfolding filter_transitions_simps by fastforce
lemma filter_transitions_path :
assumes "path (filter_transitions M P) q p"
shows "path M q p"
using path_begin_state[OF assms]
transition_subset_path[of "filter_transitions M P" M, OF _ assms]
unfolding filter_transitions_simps by blast
lemma filter_transitions_reachable_states :
assumes "q \<in> reachable_states (filter_transitions M P)"
shows "q \<in> reachable_states M"
using assms unfolding reachable_states_def filter_transitions_simps
using filter_transitions_path[of M P "initial M"]
by blast
subsection \<open>Filtering States\<close>
lift_definition filter_states :: "('a,'b,'c) fsm \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a,'b,'c) fsm"
is FSM_Impl.filter_states
proof -
fix M :: "('a,'b,'c) fsm_impl"
fix P :: "'a \<Rightarrow> bool"
assume *: "well_formed_fsm M"
then show "well_formed_fsm (FSM_Impl.filter_states M P)"
by (cases "P (FSM_Impl.initial M)"; auto)
qed
lemma filter_states_simps[simp] :
assumes "P (initial M)"
shows "initial (filter_states M P) = initial M"
"states (filter_states M P) = Set.filter P (states M)"
"inputs (filter_states M P) = inputs M"
"outputs (filter_states M P) = outputs M"
"transitions (filter_states M P) = {t \<in> transitions M . P (t_source t) \<and> P (t_target t)}"
using assms by (transfer;auto)+
lemma filter_states_submachine :
assumes "P (initial M)"
shows "is_submachine (filter_states M P) M"
using filter_states_simps[of P M, OF assms] by fastforce
fun restrict_to_reachable_states :: "('a,'b,'c) fsm \<Rightarrow> ('a,'b,'c) fsm" where
"restrict_to_reachable_states M = filter_states M (\<lambda> q . q \<in> reachable_states M)"
lemma restrict_to_reachable_states_simps[simp] :
shows "initial (restrict_to_reachable_states M) = initial M"
"states (restrict_to_reachable_states M) = reachable_states M"
"inputs (restrict_to_reachable_states M) = inputs M"
"outputs (restrict_to_reachable_states M) = outputs M"
"transitions (restrict_to_reachable_states M)
= {t \<in> transitions M . (t_source t) \<in> reachable_states M}"
proof -
show "initial (restrict_to_reachable_states M) = initial M"
"states (restrict_to_reachable_states M) = reachable_states M"
"inputs (restrict_to_reachable_states M) = inputs M"
"outputs (restrict_to_reachable_states M) = outputs M"
using filter_states_simps[of "(\<lambda> q . q \<in> reachable_states M)", OF reachable_states_initial]
using reachable_state_is_state[of _ M] by auto
have "transitions (restrict_to_reachable_states M)
= {t \<in> transitions M. (t_source t) \<in> reachable_states M \<and> (t_target t) \<in> reachable_states M}"
using filter_states_simps[of "(\<lambda> q . q \<in> reachable_states M)", OF reachable_states_initial]
by auto
then show "transitions (restrict_to_reachable_states M)
= {t \<in> transitions M . (t_source t) \<in> reachable_states M}"
using reachable_states_next[of _ M] by auto
qed
lemma restrict_to_reachable_states_path :
assumes "q \<in> reachable_states M"
shows "path M q p = path (restrict_to_reachable_states M) q p"
proof
show "path M q p \<Longrightarrow> path (restrict_to_reachable_states M) q p"
proof -
assume "path M q p"
then show "path (restrict_to_reachable_states M) q p"
using assms proof (induction p arbitrary: q rule: list.induct)
case Nil
then show ?case
using restrict_to_reachable_states_simps(2) by fastforce
next
case (Cons t' p')
then have "path M (t_target t') p'" by auto
moreover have "t_target t' \<in> reachable_states M" using Cons.prems
by (metis path_cons_elim reachable_states_next)
ultimately show ?case using Cons.IH
by (metis (no_types, lifting) Cons.prems(1) Cons.prems(2) mem_Collect_eq path.simps
path_cons_elim restrict_to_reachable_states_simps(5))
qed
qed
show "path (restrict_to_reachable_states M) q p \<Longrightarrow> path M q p"
by (metis (no_types, lifting) assms mem_Collect_eq reachable_state_is_state
restrict_to_reachable_states_simps(5) subsetI transition_subset_path)
qed
lemma restrict_to_reachable_states_language :
"L (restrict_to_reachable_states M) = L M"
unfolding LS.simps
unfolding restrict_to_reachable_states_simps
unfolding restrict_to_reachable_states_path[OF reachable_states_initial, of M]
by blast
lemma restrict_to_reachable_states_observable :
assumes "observable M"
shows "observable (restrict_to_reachable_states M)"
using assms unfolding observable.simps
unfolding restrict_to_reachable_states_simps
by blast
lemma restrict_to_reachable_states_minimal :
assumes "minimal M"
shows "minimal (restrict_to_reachable_states M)"
proof -
have "\<And> q1 q2 . q1 \<in> reachable_states M \<Longrightarrow>
q2 \<in> reachable_states M \<Longrightarrow>
q1 \<noteq> q2 \<Longrightarrow>
LS (restrict_to_reachable_states M) q1 \<noteq> LS (restrict_to_reachable_states M) q2"
proof -
fix q1 q2 assume "q1 \<in> reachable_states M" and "q2 \<in> reachable_states M" and "q1 \<noteq> q2"
then have "q1 \<in> states M" and "q2 \<in> states M"
by (simp add: reachable_state_is_state)+
then have "LS M q1 \<noteq> LS M q2"
using \<open>q1 \<noteq> q2\<close> assms by auto
then show "LS (restrict_to_reachable_states M) q1 \<noteq> LS (restrict_to_reachable_states M) q2"
unfolding LS.simps
unfolding restrict_to_reachable_states_path[OF \<open>q1 \<in> reachable_states M\<close>]
unfolding restrict_to_reachable_states_path[OF \<open>q2 \<in> reachable_states M\<close>] .
qed
then show ?thesis
unfolding minimal.simps restrict_to_reachable_states_simps
by blast
qed
lemma restrict_to_reachable_states_reachable_states :
"reachable_states (restrict_to_reachable_states M) = states (restrict_to_reachable_states M)"
proof
show "reachable_states (restrict_to_reachable_states M) \<subseteq> states (restrict_to_reachable_states M)"
by (simp add: reachable_state_is_state subsetI)
show "states (restrict_to_reachable_states M) \<subseteq> reachable_states (restrict_to_reachable_states M)"
proof
fix q assume "q \<in> states (restrict_to_reachable_states M)"
then have "q \<in> reachable_states M"
unfolding restrict_to_reachable_states_simps .
then show "q \<in> reachable_states (restrict_to_reachable_states M)"
unfolding reachable_states_def
unfolding restrict_to_reachable_states_simps
unfolding restrict_to_reachable_states_path[OF reachable_states_initial, symmetric] .
qed
qed
subsection \<open>Adding Transitions\<close>
lift_definition create_unconnected_fsm :: "'a \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> 'c set \<Rightarrow> ('a,'b,'c) fsm"
is FSM_Impl.create_unconnected_FSMI by (transfer; simp)
lemma create_unconnected_fsm_simps :
assumes "finite ns" and "finite ins" and "finite outs" and "q \<in> ns"
shows "initial (create_unconnected_fsm q ns ins outs) = q"
"states (create_unconnected_fsm q ns ins outs) = ns"
"inputs (create_unconnected_fsm q ns ins outs) = ins"
"outputs (create_unconnected_fsm q ns ins outs) = outs"
"transitions (create_unconnected_fsm q ns ins outs) = {}"
using assms by (transfer; auto)+
lift_definition create_unconnected_fsm_from_lists :: "'a \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list \<Rightarrow> ('a,'b,'c) fsm"
is FSM_Impl.create_unconnected_fsm_from_lists by (transfer; simp)
lemma create_unconnected_fsm_from_lists_simps :
assumes "q \<in> set ns"
shows "initial (create_unconnected_fsm_from_lists q ns ins outs) = q"
"states (create_unconnected_fsm_from_lists q ns ins outs) = set ns"
"inputs (create_unconnected_fsm_from_lists q ns ins outs) = set ins"
"outputs (create_unconnected_fsm_from_lists q ns ins outs) = set outs"
"transitions (create_unconnected_fsm_from_lists q ns ins outs) = {}"
using assms by (transfer; auto)+
lift_definition create_unconnected_fsm_from_fsets :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> 'c fset \<Rightarrow> ('a,'b,'c) fsm"
is FSM_Impl.create_unconnected_fsm_from_fsets by (transfer; simp)
lemma create_unconnected_fsm_from_fsets_simps :
assumes "q |\<in>| ns"
shows "initial (create_unconnected_fsm_from_fsets q ns ins outs) = q"
"states (create_unconnected_fsm_from_fsets q ns ins outs) = fset ns"
"inputs (create_unconnected_fsm_from_fsets q ns ins outs) = fset ins"
"outputs (create_unconnected_fsm_from_fsets q ns ins outs) = fset outs"
"transitions (create_unconnected_fsm_from_fsets q ns ins outs) = {}"
- using assms unfolding fmember_def by (transfer; auto)+
+ using assms by (transfer; auto)+
lift_definition add_transitions :: "('a,'b,'c) fsm \<Rightarrow> ('a,'b,'c) transition set \<Rightarrow> ('a,'b,'c) fsm"
is FSM_Impl.add_transitions
proof -
fix M :: "('a,'b,'c) fsm_impl"
fix ts :: "('a,'b,'c) transition set"
assume *: "well_formed_fsm M"
then show "well_formed_fsm (FSM_Impl.add_transitions M ts)"
proof (cases "\<forall> t \<in> ts . t_source t \<in> FSM_Impl.states M \<and> t_input t \<in> FSM_Impl.inputs M
\<and> t_output t \<in> FSM_Impl.outputs M \<and> t_target t \<in> FSM_Impl.states M")
case True
then have "ts \<subseteq> FSM_Impl.states M \<times> FSM_Impl.inputs M \<times> FSM_Impl.outputs M \<times> FSM_Impl.states M"
by fastforce
moreover have "finite (FSM_Impl.states M \<times> FSM_Impl.inputs M \<times> FSM_Impl.outputs M \<times> FSM_Impl.states M)"
using * by blast
ultimately have "finite ts"
using rev_finite_subset by auto
then show ?thesis using * by auto
next
case False
then show ?thesis using * by auto
qed
qed
lemma add_transitions_simps :
assumes "\<And> t . t \<in> ts \<Longrightarrow> t_source t \<in> states M \<and> t_input t \<in> inputs M \<and> t_output t \<in> outputs M \<and> t_target t \<in> states M"
shows "initial (add_transitions M ts) = initial M"
"states (add_transitions M ts) = states M"
"inputs (add_transitions M ts) = inputs M"
"outputs (add_transitions M ts) = outputs M"
"transitions (add_transitions M ts) = transitions M \<union> ts"
using assms by (transfer; auto)+
lift_definition create_fsm_from_sets :: "'a \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> 'c set \<Rightarrow> ('a,'b,'c) transition set \<Rightarrow> ('a,'b,'c) fsm"
is FSM_Impl.create_fsm_from_sets
proof -
fix q :: 'a
fix qs :: "'a set"
fix ins :: "'b set"
fix outs :: "'c set"
fix ts :: "('a,'b,'c) transition set"
show "well_formed_fsm (FSM_Impl.create_fsm_from_sets q qs ins outs ts)"
proof (cases "q \<in> qs \<and> finite qs \<and> finite ins \<and> finite outs")
case True
let ?M = "(FSMI q qs ins outs {})"
show ?thesis proof (cases "\<forall> t \<in> ts . t_source t \<in> FSM_Impl.states ?M \<and> t_input t \<in> FSM_Impl.inputs ?M
\<and> t_output t \<in> FSM_Impl.outputs ?M \<and> t_target t \<in> FSM_Impl.states ?M")
case True
then have "ts \<subseteq> FSM_Impl.states ?M \<times> FSM_Impl.inputs ?M \<times> FSM_Impl.outputs ?M \<times> FSM_Impl.states ?M"
by fastforce
moreover have "finite (FSM_Impl.states ?M \<times> FSM_Impl.inputs ?M \<times> FSM_Impl.outputs ?M \<times> FSM_Impl.states ?M)"
using \<open>q \<in> qs \<and> finite qs \<and> finite ins \<and> finite outs\<close> by force
ultimately have "finite ts"
using rev_finite_subset by auto
then show ?thesis by auto
next
case False
then show ?thesis by auto
qed
next
case False
then show ?thesis by auto
qed
qed
lemma create_fsm_from_sets_simps :
assumes "q \<in> qs" and "finite qs" and "finite ins" and "finite outs"
assumes "\<And> t . t \<in> ts \<Longrightarrow> t_source t \<in> qs \<and> t_input t \<in> ins \<and> t_output t \<in> outs \<and> t_target t \<in> qs"
shows "initial (create_fsm_from_sets q qs ins outs ts) = q"
"states (create_fsm_from_sets q qs ins outs ts) = qs"
"inputs (create_fsm_from_sets q qs ins outs ts) = ins"
"outputs (create_fsm_from_sets q qs ins outs ts) = outs"
"transitions (create_fsm_from_sets q qs ins outs ts) = ts"
using assms by (transfer; auto)+
lemma create_fsm_from_self :
"m = create_fsm_from_sets (initial m) (states m) (inputs m) (outputs m) (transitions m)"
proof -
have *: "\<And> t . t \<in> transitions m \<Longrightarrow> t_source t \<in> states m \<and> t_input t \<in> inputs m \<and> t_output t \<in> outputs m \<and> t_target t \<in> states m"
by auto
show ?thesis
using create_fsm_from_sets_simps[OF fsm_initial fsm_states_finite fsm_inputs_finite fsm_outputs_finite *, of "transitions m"]
apply transfer
by force
qed
lemma create_fsm_from_sets_surj :
assumes "finite (UNIV :: 'a set)"
and "finite (UNIV :: 'b set)"
and "finite (UNIV :: 'c set)"
shows "surj (\<lambda>(q::'a,Q,X::'b set,Y::'c set,T) . create_fsm_from_sets q Q X Y T)"
proof
show "range (\<lambda>(q::'a,Q,X::'b set,Y::'c set,T) . create_fsm_from_sets q Q X Y T) \<subseteq> UNIV"
by simp
show "UNIV \<subseteq> range (\<lambda>(q::'a,Q,X::'b set,Y::'c set,T) . create_fsm_from_sets q Q X Y T)"
proof
fix m assume "m \<in> (UNIV :: ('a,'b,'c) fsm set)"
then have "m = create_fsm_from_sets (initial m) (states m) (inputs m) (outputs m) (transitions m)"
using create_fsm_from_self by blast
then have "m = (\<lambda>(q::'a,Q,X::'b set,Y::'c set,T) . create_fsm_from_sets q Q X Y T) (initial m,states m,inputs m,outputs m,transitions m)"
by auto
then show "m \<in> range (\<lambda>(q::'a,Q,X::'b set,Y::'c set,T) . create_fsm_from_sets q Q X Y T)"
by blast
qed
qed
subsection \<open>Distinguishability\<close>
definition distinguishes :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('b \<times>'c) list \<Rightarrow> bool" where
"distinguishes M q1 q2 io = (io \<in> LS M q1 \<union> LS M q2 \<and> io \<notin> LS M q1 \<inter> LS M q2)"
definition minimally_distinguishes :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('b \<times>'c) list \<Rightarrow> bool" where
"minimally_distinguishes M q1 q2 io = (distinguishes M q1 q2 io
\<and> (\<forall> io' . distinguishes M q1 q2 io' \<longrightarrow> length io \<le> length io'))"
lemma minimally_distinguishes_ex :
assumes "q1 \<in> states M"
and "q2 \<in> states M"
and "LS M q1 \<noteq> LS M q2"
obtains v where "minimally_distinguishes M q1 q2 v"
proof -
let ?vs = "{v . distinguishes M q1 q2 v}"
define vMin where vMin: "vMin = arg_min length (\<lambda>v . v \<in> ?vs)"
obtain v' where "distinguishes M q1 q2 v'"
using assms unfolding distinguishes_def by blast
then have "vMin \<in> ?vs \<and> (\<forall> v'' . distinguishes M q1 q2 v'' \<longrightarrow> length vMin \<le> length v'')"
unfolding vMin using arg_min_nat_lemma[of "\<lambda>v . distinguishes M q1 q2 v" v' length]
by simp
then show ?thesis
using that[of vMin] unfolding minimally_distinguishes_def by blast
qed
lemma distinguish_prepend :
assumes "observable M"
and "distinguishes M (FSM.after M q1 io) (FSM.after M q2 io) w"
and "q1 \<in> states M"
and "q2 \<in> states M"
and "io \<in> LS M q1"
and "io \<in> LS M q2"
shows "distinguishes M q1 q2 (io@w)"
proof -
have "(io@w \<in> LS M q1) = (w \<in> LS M (after M q1 io))"
using assms(1,3,5)
by (metis after_language_iff)
moreover have "(io@w \<in> LS M q2) = (w \<in> LS M (after M q2 io))"
using assms(1,4,6)
by (metis after_language_iff)
ultimately show ?thesis
using assms(2) unfolding distinguishes_def by blast
qed
lemma distinguish_prepend_initial :
assumes "observable M"
and "distinguishes M (after_initial M (io1@io)) (after_initial M (io2@io)) w"
and "io1@io \<in> L M"
and "io2@io \<in> L M"
shows "distinguishes M (after_initial M io1) (after_initial M io2) (io@w)"
proof -
have f1: "\<forall>ps psa f a. (ps::('b \<times> 'c) list) @ psa \<notin> LS f (a::'a) \<or> ps \<in> LS f a"
by (meson language_prefix)
then have f2: "io1 \<in> L M"
by (meson assms(3))
have f3: "io2 \<in> L M"
using f1 by (metis assms(4))
have "io1 \<in> L M"
using f1 by (metis assms(3))
then show ?thesis
by (metis after_is_state after_language_iff after_split assms(1) assms(2) assms(3) assms(4) distinguish_prepend f3)
qed
lemma minimally_distinguishes_no_prefix :
assumes "observable M"
and "u@w \<in> L M"
and "v@w \<in> L M"
and "minimally_distinguishes M (after_initial M u) (after_initial M v) (w@w'@w'')"
and "w' \<noteq> []"
shows "\<not>distinguishes M (after_initial M (u@w)) (after_initial M (v@w)) w''"
proof
assume "distinguishes M (after_initial M (u @ w)) (after_initial M (v @ w)) w''"
then have "distinguishes M (after_initial M u) (after_initial M v) (w@w'')"
using assms(1-3) distinguish_prepend_initial by blast
moreover have "length (w@w'') < length (w@w'@w'')"
using assms(5) by auto
ultimately show False
using assms(4) unfolding minimally_distinguishes_def
using leD by blast
qed
lemma minimally_distinguishes_after_append :
assumes "observable M"
and "minimal M"
and "q1 \<in> states M"
and "q2 \<in> states M"
and "minimally_distinguishes M q1 q2 (w@w')"
and "w' \<noteq> []"
shows "minimally_distinguishes M (after M q1 w) (after M q2 w) w'"
proof -
have "\<not> distinguishes M q1 q2 w"
using assms(5,6)
by (metis add.right_neutral add_le_cancel_left length_append length_greater_0_conv linorder_not_le minimally_distinguishes_def)
then have "w \<in> LS M q1 = (w \<in> LS M q2)"
unfolding distinguishes_def
by blast
moreover have "(w@w') \<in> LS M q1 \<union> LS M q2"
using assms(5) unfolding minimally_distinguishes_def distinguishes_def
by blast
ultimately have "w \<in> LS M q1" and "w \<in> LS M q2"
by (meson Un_iff language_prefix)+
have "(w@w') \<in> LS M q1 = (w' \<in> LS M (after M q1 w))"
by (meson \<open>w \<in> LS M q1\<close> after_language_iff assms(1))
moreover have "(w@w') \<in> LS M q2 = (w' \<in> LS M (after M q2 w))"
by (meson \<open>w \<in> LS M q2\<close> after_language_iff assms(1))
ultimately have "distinguishes M (after M q1 w) (after M q2 w) w'"
using assms(5) unfolding minimally_distinguishes_def distinguishes_def
by blast
moreover have "\<And> w'' . distinguishes M (after M q1 w) (after M q2 w) w'' \<Longrightarrow> length w' \<le> length w''"
proof -
fix w'' assume "distinguishes M (after M q1 w) (after M q2 w) w''"
then have "distinguishes M q1 q2 (w@w'')"
by (metis \<open>w \<in> LS M q1\<close> \<open>w \<in> LS M q2\<close> assms(1) assms(3) assms(4) distinguish_prepend)
then have "length (w@w') \<le> length (w@w'')"
using assms(5) unfolding minimally_distinguishes_def distinguishes_def
by blast
then show "length w' \<le> length w''"
by auto
qed
ultimately show ?thesis
unfolding minimally_distinguishes_def distinguishes_def
by blast
qed
lemma minimally_distinguishes_after_append_initial :
assumes "observable M"
and "minimal M"
and "u \<in> L M"
and "v \<in> L M"
and "minimally_distinguishes M (after_initial M u) (after_initial M v) (w@w')"
and "w' \<noteq> []"
shows "minimally_distinguishes M (after_initial M (u@w)) (after_initial M (v@w)) w'"
proof -
have "\<not> distinguishes M (after_initial M u) (after_initial M v) w"
using assms(5,6)
by (metis add.right_neutral add_le_cancel_left length_append length_greater_0_conv linorder_not_le minimally_distinguishes_def)
then have "w \<in> LS M (after_initial M u) = (w \<in> LS M (after_initial M v))"
unfolding distinguishes_def
by blast
moreover have "(w@w') \<in> LS M (after_initial M u) \<union> LS M (after_initial M v)"
using assms(5) unfolding minimally_distinguishes_def distinguishes_def
by blast
ultimately have "w \<in> LS M (after_initial M u)" and "w \<in> LS M (after_initial M v)"
by (meson Un_iff language_prefix)+
have "(w@w') \<in> LS M (after_initial M u) = (w' \<in> LS M (after_initial M (u@w)))"
by (meson \<open>w \<in> LS M (after_initial M u)\<close> after_language_append_iff after_language_iff assms(1) assms(3))
moreover have "(w@w') \<in> LS M (after_initial M v) = (w' \<in> LS M (after_initial M (v@w)))"
by (meson \<open>w \<in> LS M (after_initial M v)\<close> after_language_append_iff after_language_iff assms(1) assms(4))
ultimately have "distinguishes M (after_initial M (u@w)) (after_initial M (v@w)) w'"
using assms(5) unfolding minimally_distinguishes_def distinguishes_def
by blast
moreover have "\<And> w'' . distinguishes M (after_initial M (u@w)) (after_initial M (v@w)) w'' \<Longrightarrow> length w' \<le> length w''"
proof -
fix w'' assume "distinguishes M (after_initial M (u@w)) (after_initial M (v@w)) w''"
then have "distinguishes M (after_initial M u) (after_initial M v) (w@w'')"
by (meson \<open>w \<in> LS M (after_initial M u)\<close> \<open>w \<in> LS M (after_initial M v)\<close> after_language_iff assms(1) assms(3) assms(4) distinguish_prepend_initial)
then have "length (w@w') \<le> length (w@w'')"
using assms(5) unfolding minimally_distinguishes_def distinguishes_def
by blast
then show "length w' \<le> length w''"
by auto
qed
ultimately show ?thesis
unfolding minimally_distinguishes_def distinguishes_def
by blast
qed
lemma minimally_distinguishes_proper_prefixes_card :
assumes "observable M"
and "minimal M"
and "q1 \<in> states M"
and "q2 \<in> states M"
and "minimally_distinguishes M q1 q2 w"
and "S \<subseteq> states M"
shows "card {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S} \<le> card S - 1"
(is "?P S")
proof -
define k where "k = card S"
then show ?thesis
using assms(6)
proof (induction k arbitrary: S rule: less_induct)
case (less k)
then have "finite S"
by (metis fsm_states_finite rev_finite_subset)
show ?case proof (cases k)
case 0
then have "S = {}"
using less.prems \<open>finite S\<close> by auto
then show ?thesis
by fastforce
next
case (Suc k')
show ?thesis proof (cases "{w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S} = {}")
case True
then show ?thesis
by (metis bot.extremum dual_order.eq_iff obtain_subset_with_card_n)
next
case False
define wk where wk: "wk = arg_max length (\<lambda>wk . wk \<in> {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S})"
obtain wk' where *:"wk' \<in> {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S}"
using False by blast
have "finite {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S}"
by (metis (no_types) Collect_mem_eq List.finite_set finite_Collect_conjI)
then have "wk \<in> {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S}"
and "\<And> wk' . wk' \<in> {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S} \<Longrightarrow> length wk' \<le> length wk"
using False unfolding wk
using arg_max_nat_lemma[of "(\<lambda>wk . wk \<in> {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S})", OF *]
by (meson finite_maxlen)+
then have "wk \<in> set (prefixes w)" and "wk \<noteq> w" and "after M q1 wk \<in> S" and "after M q2 wk \<in> S"
by blast+
obtain wk_suffix where "w = wk@wk_suffix" and "wk_suffix \<noteq> []"
using \<open>wk \<in> set (prefixes w)\<close>
using prefixes_set_ob \<open>wk \<noteq> w\<close>
by blast
have "distinguishes M (after M q1 []) (after M q2 []) w"
using \<open>minimally_distinguishes M q1 q2 w\<close>
by (metis after.simps(1) minimally_distinguishes_def)
have "minimally_distinguishes M (after M q1 wk) (after M q2 wk) wk_suffix"
using \<open>minimally_distinguishes M q1 q2 w\<close> \<open>wk_suffix \<noteq> []\<close>
unfolding \<open>w = wk@wk_suffix\<close>
using minimally_distinguishes_after_append[OF assms(1,2,3,4), of wk wk_suffix]
by blast
then have "distinguishes M (after M q1 wk) (after M q2 wk) wk_suffix"
unfolding minimally_distinguishes_def
by auto
then have "wk_suffix \<in> LS M (after M q1 wk) = (wk_suffix \<notin> LS M (after M q2 wk))"
unfolding distinguishes_def by blast
define S1 where S1: "S1 = Set.filter (\<lambda>q . wk_suffix \<in> LS M q) S"
define S2 where S2: "S2 = Set.filter (\<lambda>q . wk_suffix \<notin> LS M q) S"
have "S = S1 \<union> S2"
unfolding S1 S2 by auto
moreover have "S1 \<inter> S2 = {}"
unfolding S1 S2 by auto
ultimately have "card S = card S1 + card S2"
using \<open>finite S\<close> card_Un_disjoint by blast
have "S1 \<subseteq> states M" and "S2 \<subseteq> states M"
using \<open>S = S1 \<union> S2\<close> less.prems(2) by blast+
have "S1 \<noteq> {}" and "S2 \<noteq> {}"
using \<open>wk_suffix \<in> LS M (after M q1 wk) = (wk_suffix \<notin> LS M (after M q2 wk))\<close> \<open>after M q1 wk \<in> S\<close> \<open>after M q2 wk \<in> S\<close>
unfolding S1 S2
by (metis empty_iff member_filter)+
then have "card S1 > 0" and "card S2 > 0"
using \<open>S = S1 \<union> S2\<close> \<open>finite S\<close>
by (meson card_0_eq finite_Un neq0_conv)+
then have "card S1 < k" and "card S2 < k"
using \<open>card S = card S1 + card S2\<close> unfolding less.prems
by auto
define W where W: "W = (\<lambda> S1 S2 . {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S1 \<and> after M q2 w' \<in> S2})"
then have "\<And> S' S'' . W S' S'' \<subseteq> set (prefixes w)"
by auto
then have W_finite: "\<And> S' S'' . finite (W S' S'')"
using List.finite_set[of "prefixes w"]
by (meson finite_subset)
have "\<And> w' . w' \<in> W S S \<Longrightarrow> w' \<noteq> wk \<Longrightarrow> after M q1 w' \<in> S1 = (after M q2 w' \<in> S1)"
proof -
fix w' assume *:"w' \<in> W S S" and "w' \<noteq> wk"
then have "w' \<in> set (prefixes w)" and "w' \<noteq> w" and "after M q1 w' \<in> S" and "after M q2 w' \<in> S"
unfolding W
by blast+
then have "w' \<in> LS M q1"
by (metis IntE UnCI UnE append_self_conv assms(5) distinguishes_def language_prefix leD length_append length_greater_0_conv less_add_same_cancel1 minimally_distinguishes_def prefixes_set_ob)
have "w' \<in> LS M q2"
by (metis IntE UnCI \<open>w' \<in> LS M q1\<close> \<open>w' \<in> set (prefixes w)\<close> \<open>w' \<noteq> w\<close> append_Nil2 assms(5) distinguishes_def leD length_append length_greater_0_conv less_add_same_cancel1 minimally_distinguishes_def prefixes_set_ob)
have "length w' < length wk"
using \<open>w' \<noteq> wk\<close> *
\<open>\<And> wk' . wk' \<in> {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S} \<Longrightarrow> length wk' \<le> length wk\<close>
unfolding W
by (metis (no_types, lifting) \<open>w = wk @ wk_suffix\<close> \<open>w' \<in> set (prefixes w)\<close> append_eq_append_conv le_neq_implies_less prefixes_set_ob)
show "after M q1 w' \<in> S1 = (after M q2 w' \<in> S1)"
proof (rule ccontr)
assume "(after M q1 w' \<in> S1) \<noteq> (after M q2 w' \<in> S1)"
then have "(after M q1 w' \<in> S1 \<and> (after M q2 w' \<in> S2)) \<or> (after M q1 w' \<in> S2 \<and> (after M q2 w' \<in> S1))"
using \<open>after M q1 w' \<in> S\<close> \<open>after M q2 w' \<in> S\<close>
unfolding \<open>S = S1 \<union> S2\<close>
by blast
then have "wk_suffix \<in> LS M (after M q1 w') = (wk_suffix \<notin> LS M (after M q2 w'))"
unfolding S1 S2
by (metis member_filter)
then have "distinguishes M (after M q1 w') (after M q2 w') wk_suffix"
unfolding distinguishes_def by blast
then have "distinguishes M q1 q2 (w'@wk_suffix)"
using distinguish_prepend[OF assms(1) _ \<open>q1 \<in> states M\<close> \<open>q2 \<in> states M\<close> \<open>w' \<in> LS M q1\<close> \<open>w' \<in> LS M q2\<close>]
by blast
moreover have "length (w'@wk_suffix) < length (wk@wk_suffix)"
using \<open>length w' < length wk\<close>
by auto
ultimately show False
using \<open>minimally_distinguishes M q1 q2 w\<close>
unfolding \<open>w = wk@wk_suffix\<close> minimally_distinguishes_def
by auto
qed
qed
have "\<And> x . x \<in> W S1 S2 \<union> W S2 S1 \<Longrightarrow> x = wk"
proof -
fix x assume "x \<in> W S1 S2 \<union> W S2 S1"
then have "x \<in> W S S"
unfolding W \<open>S = S1 \<union> S2\<close> by blast
show "x = wk"
using \<open>x \<in> W S1 S2 \<union> W S2 S1\<close>
using \<open>\<And> w' . w' \<in> W S S \<Longrightarrow> w' \<noteq> wk \<Longrightarrow> after M q1 w' \<in> S1 = (after M q2 w' \<in> S1)\<close>[OF \<open>x \<in> W S S\<close>]
unfolding W
using \<open>S1 \<inter> S2 = {}\<close>
by blast
qed
moreover have "wk \<in> W S1 S2 \<union> W S2 S1"
unfolding W
using \<open>wk \<in> {w' . w' \<in> set (prefixes w) \<and> w' \<noteq> w \<and> after M q1 w' \<in> S \<and> after M q2 w' \<in> S}\<close>
\<open>wk_suffix \<in> LS M (after M q1 wk) = (wk_suffix \<notin> LS M (after M q2 wk))\<close>
by (metis (no_types, lifting) S1 Un_iff \<open>S = S1 \<union> S2\<close> mem_Collect_eq member_filter)
ultimately have "W S1 S2 \<union> W S2 S1 = {wk}"
by blast
have "W S S = (W S1 S1 \<union> W S2 S2 \<union> (W S1 S2 \<union> W S2 S1))"
unfolding W \<open>S = S1 \<union> S2\<close> by blast
moreover have "W S1 S1 \<inter> W S2 S2 = {}"
using \<open>S1 \<inter> S2 = {}\<close> unfolding W
by blast
moreover have "W S1 S1 \<inter> (W S1 S2 \<union> W S2 S1) = {}"
unfolding W
using \<open>S1 \<inter> S2 = {}\<close>
by blast
moreover have "W S2 S2 \<inter> (W S1 S2 \<union> W S2 S1) = {}"
unfolding W
using \<open>S1 \<inter> S2 = {}\<close>
by blast
moreover have "finite (W S1 S1)" and "finite (W S2 S2)" and "finite {wk}"
using W_finite by auto
ultimately have "card (W S S) = card (W S1 S1) + card (W S2 S2) + 1"
unfolding \<open>W S1 S2 \<union> W S2 S1 = {wk}\<close>
by (metis card_Un_disjoint finite_UnI inf_sup_distrib2 is_singletonI is_singleton_altdef sup_idem)
moreover have "card (W S1 S1) \<le> card S1 - 1"
using less.IH[OF \<open>card S1 < k\<close> _ \<open>S1 \<subseteq> states M\<close>]
unfolding W by blast
moreover have "card (W S2 S2) \<le> card S2 - 1"
using less.IH[OF \<open>card S2 < k\<close> _ \<open>S2 \<subseteq> states M\<close>]
unfolding W by blast
ultimately have "card (W S S) \<le> card S - 1"
using \<open>card S = card S1 + card S2\<close>
using \<open>card S1 < k\<close> \<open>card S2 < k\<close> less.prems(1) by linarith
then show ?thesis
unfolding W .
qed
qed
qed
qed
lemma minimally_distinguishes_proper_prefix_in_language :
assumes "minimally_distinguishes M q1 q2 io"
and "io' \<in> set (prefixes io)"
and "io' \<noteq> io"
shows "io' \<in> LS M q1 \<inter> LS M q2"
proof -
have "io \<in> LS M q1 \<or> io \<in> LS M q2"
using assms(1) unfolding minimally_distinguishes_def distinguishes_def by blast
then have "io' \<in> LS M q1 \<or> io' \<in> LS M q2"
by (metis assms(2) prefixes_set_ob language_prefix)
have "length io' < length io"
using assms(2,3) unfolding prefixes_set by auto
then have "io' \<in> LS M q1 \<longleftrightarrow> io' \<in> LS M q2"
using assms(1) unfolding minimally_distinguishes_def distinguishes_def
by (metis Int_iff Un_Int_eq(1) Un_Int_eq(2) leD)
then show ?thesis
using \<open>io' \<in> LS M q1 \<or> io' \<in> LS M q2\<close>
by blast
qed
lemma distinguishes_not_Nil:
assumes "distinguishes M q1 q2 io"
and "q1 \<in> states M"
and "q2 \<in> states M"
shows "io \<noteq> []"
using assms unfolding distinguishes_def by auto
fun does_distinguish :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> bool" where
"does_distinguish M q1 q2 io = (is_in_language M q1 io \<noteq> is_in_language M q2 io)"
lemma does_distinguish_correctness :
assumes "observable M"
and "q1 \<in> states M"
and "q2 \<in> states M"
shows "does_distinguish M q1 q2 io = distinguishes M q1 q2 io"
unfolding does_distinguish.simps
is_in_language_iff[OF assms(1,2)]
is_in_language_iff[OF assms(1,3)]
distinguishes_def
by blast
lemma h_obs_distinguishes :
assumes "observable M"
and "h_obs M q1 x y = Some q1'"
and "h_obs M q2 x y = None"
shows "distinguishes M q1 q2 [(x,y)]"
using assms(2,3) LS_single_transition[of x y M] unfolding distinguishes_def h_obs_Some[OF assms(1)] h_obs_None[OF assms(1)]
by (metis Int_iff UnI1 \<open>\<And>y x q. (h_obs M q x y = None) = (\<nexists>q'. (q, x, y, q') \<in> FSM.transitions M)\<close> assms(1) assms(2) fst_conv h_obs_language_iff option.distinct(1) snd_conv)
lemma distinguishes_sym :
assumes "distinguishes M q1 q2 io"
shows "distinguishes M q2 q1 io"
using assms unfolding distinguishes_def by blast
lemma distinguishes_after_prepend :
assumes "observable M"
and "h_obs M q1 x y \<noteq> None"
and "h_obs M q2 x y \<noteq> None"
and "distinguishes M (FSM.after M q1 [(x,y)]) (FSM.after M q2 [(x,y)]) \<gamma>"
shows "distinguishes M q1 q2 ((x,y)#\<gamma>)"
proof -
have "[(x,y)] \<in> LS M q1"
using assms(2) h_obs_language_single_transition_iff[OF assms(1)] by auto
have "[(x,y)] \<in> LS M q2"
using assms(3) h_obs_language_single_transition_iff[OF assms(1)] by auto
show ?thesis
using after_language_iff[OF assms(1) \<open>[(x,y)] \<in> LS M q1\<close>, of \<gamma>]
using after_language_iff[OF assms(1) \<open>[(x,y)] \<in> LS M q2\<close>, of \<gamma>]
using assms(4)
unfolding distinguishes_def
by simp
qed
lemma distinguishes_after_initial_prepend :
assumes "observable M"
and "io1 \<in> L M"
and "io2 \<in> L M"
and "h_obs M (after_initial M io1) x y \<noteq> None"
and "h_obs M (after_initial M io2) x y \<noteq> None"
and "distinguishes M (after_initial M (io1@[(x,y)])) (after_initial M (io2@[(x,y)])) \<gamma>"
shows "distinguishes M (after_initial M io1) (after_initial M io2) ((x,y)#\<gamma>)"
by (metis after_split assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) distinguishes_after_prepend h_obs_language_append)
subsection \<open>Extending FSMs by single elements\<close>
lemma fsm_from_list_simps[simp] :
"initial (fsm_from_list q ts) = (case ts of [] \<Rightarrow> q | (t#ts) \<Rightarrow> t_source t)"
"states (fsm_from_list q ts) = (case ts of [] \<Rightarrow> {q} | (t#ts') \<Rightarrow> ((image t_source (set ts)) \<union> (image t_target (set ts))))"
"inputs (fsm_from_list q ts) = image t_input (set ts)"
"outputs (fsm_from_list q ts) = image t_output (set ts)"
"transitions (fsm_from_list q ts) = set ts"
by (cases ts; transfer; simp)+
lift_definition add_transition :: "('a,'b,'c) fsm \<Rightarrow> ('a,'b,'c) transition \<Rightarrow> ('a,'b,'c) fsm" is FSM_Impl.add_transition
by simp
lemma add_transition_simps[simp]:
assumes "t_source t \<in> states M" and "t_input t \<in> inputs M" and "t_output t \<in> outputs M" and "t_target t \<in> states M"
shows
"initial (add_transition M t) = initial M"
"inputs (add_transition M t) = inputs M"
"outputs (add_transition M t) = outputs M"
"transitions (add_transition M t) = insert t (transitions M)"
"states (add_transition M t) = states M" using assms by (transfer; simp)+
lift_definition add_state :: "('a,'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> ('a,'b,'c) fsm" is FSM_Impl.add_state
by simp
lemma add_state_simps[simp]:
"initial (add_state M q) = initial M"
"inputs (add_state M q) = inputs M"
"outputs (add_state M q) = outputs M"
"transitions (add_state M q) = transitions M"
"states (add_state M q) = insert q (states M)" by (transfer; simp)+
lift_definition add_input :: "('a,'b,'c) fsm \<Rightarrow> 'b \<Rightarrow> ('a,'b,'c) fsm" is FSM_Impl.add_input
by simp
lemma add_input_simps[simp]:
"initial (add_input M x) = initial M"
"inputs (add_input M x) = insert x (inputs M)"
"outputs (add_input M x) = outputs M"
"transitions (add_input M x) = transitions M"
"states (add_input M x) = states M" by (transfer; simp)+
lift_definition add_output :: "('a,'b,'c) fsm \<Rightarrow> 'c \<Rightarrow> ('a,'b,'c) fsm" is FSM_Impl.add_output
by simp
lemma add_output_simps[simp]:
"initial (add_output M y) = initial M"
"inputs (add_output M y) = inputs M"
"outputs (add_output M y) = insert y (outputs M)"
"transitions (add_output M y) = transitions M"
"states (add_output M y) = states M" by (transfer; simp)+
lift_definition add_transition_with_components :: "('a,'b,'c) fsm \<Rightarrow> ('a,'b,'c) transition \<Rightarrow> ('a,'b,'c) fsm" is FSM_Impl.add_transition_with_components
by simp
lemma add_transition_with_components_simps[simp]:
"initial (add_transition_with_components M t) = initial M"
"inputs (add_transition_with_components M t) = insert (t_input t) (inputs M)"
"outputs (add_transition_with_components M t) = insert (t_output t) (outputs M)"
"transitions (add_transition_with_components M t) = insert t (transitions M)"
"states (add_transition_with_components M t) = insert (t_target t) (insert (t_source t) (states M))"
by (transfer; simp)+
subsection \<open>Renaming Elements\<close>
lift_definition rename_states :: "('a,'b,'c) fsm \<Rightarrow> ('a \<Rightarrow> 'd) \<Rightarrow> ('d,'b,'c) fsm" is FSM_Impl.rename_states
by simp
lemma rename_states_simps[simp]:
"initial (rename_states M f) = f (initial M)"
"states (rename_states M f) = f ` (states M)"
"inputs (rename_states M f) = inputs M"
"outputs (rename_states M f) = outputs M"
"transitions (rename_states M f) = (\<lambda>t . (f (t_source t), t_input t, t_output t, f (t_target t))) ` transitions M"
by (transfer; simp)+
lemma rename_states_isomorphism_language_state :
assumes "bij_betw f (states M) (f ` states M)"
and "q \<in> states M"
shows "LS (rename_states M f) (f q) = LS M q"
proof -
have *: "bij_betw f (FSM.states M) (FSM.states (FSM.rename_states M f))"
using assms rename_states_simps by auto
have **: "f (initial M) = initial (rename_states M f)"
using rename_states_simps by auto
have ***: "(\<And>q x y q'.
q \<in> states M \<Longrightarrow>
q' \<in> states M \<Longrightarrow> ((q, x, y, q') \<in> transitions M) = ((f q, x, y, f q') \<in> transitions (rename_states M f)))"
proof
fix q x y q' assume "q \<in> states M" and "q' \<in> states M"
show "(q, x, y, q') \<in> transitions M \<Longrightarrow> (f q, x, y, f q') \<in> transitions (rename_states M f)"
unfolding assms rename_states_simps by force
show "(f q, x, y, f q') \<in> transitions (rename_states M f) \<Longrightarrow> (q, x, y, q') \<in> transitions M"
proof -
assume "(f q, x, y, f q') \<in> transitions (rename_states M f)"
then obtain t where "(f q, x, y, f q') = (f (t_source t), t_input t, t_output t, f (t_target t))"
and "t \<in> transitions M"
unfolding assms rename_states_simps
by blast
then have "t_source t \<in> states M" and "t_target t \<in> states M" and "f (t_source t) = f q" and "f (t_target t) = f q'" and "t_input t = x" and "t_output t = y"
by auto
have "f q \<in> states (rename_states M f)" and "f q' \<in> states (rename_states M f)"
using \<open>(f q, x, y, f q') \<in> transitions (rename_states M f)\<close>
by auto
have "t_source t = q"
using \<open>f (t_source t) = f q\<close> \<open>q \<in> states M\<close> \<open>t_source t \<in> states M\<close>
using assms unfolding bij_betw_def inj_on_def
by blast
moreover have "t_target t = q'"
using \<open>f (t_target t) = f q'\<close> \<open>q' \<in> states M\<close> \<open>t_target t \<in> states M\<close>
using assms unfolding bij_betw_def inj_on_def
by blast
ultimately show "(q, x, y, q') \<in> transitions M"
using \<open>t_input t = x\<close> \<open>t_output t = y\<close> \<open>t \<in> transitions M\<close>
by auto
qed
qed
show ?thesis
using language_equivalence_from_isomorphism[OF * ** *** assms(2)]
by blast
qed
lemma rename_states_isomorphism_language :
assumes "bij_betw f (states M) (f ` states M)"
shows "L (rename_states M f) = L M"
using rename_states_isomorphism_language_state[OF assms fsm_initial]
unfolding rename_states_simps .
lemma rename_states_observable :
assumes "bij_betw f (states M) (f ` states M)"
and "observable M"
shows "observable (rename_states M f)"
proof -
have "\<And> q1 x y q1' q1'' . (q1,x,y,q1') \<in> transitions (rename_states M f) \<Longrightarrow> (q1,x,y,q1'') \<in> transitions (rename_states M f) \<Longrightarrow> q1' = q1''"
proof -
fix q1 x y q1' q1''
assume "(q1,x,y,q1') \<in> transitions (rename_states M f)" and "(q1,x,y,q1'') \<in> transitions (rename_states M f)"
then obtain t' t'' where "t' \<in> transitions M"
and "t'' \<in> transitions M"
and "(f (t_source t'), t_input t', t_output t', f (t_target t')) = (q1,x,y,q1')"
and "(f (t_source t''), t_input t'', t_output t'', f (t_target t'')) = (q1,x,y,q1'')"
unfolding rename_states_simps
by force
then have "f (t_source t') = f (t_source t'')"
by auto
moreover have "t_source t' \<in> states M" and "t_source t'' \<in> states M"
using \<open>t' \<in> transitions M\<close> \<open>t'' \<in> transitions M\<close>
by auto
ultimately have "t_source t' = t_source t''"
using assms(1)
unfolding bij_betw_def inj_on_def by blast
then have "t_target t' = t_target t''"
using assms(2) unfolding observable.simps
by (metis Pair_inject \<open>(f (t_source t''), t_input t'', t_output t'', f (t_target t'')) = (q1, x, y, q1'')\<close> \<open>(f (t_source t'), t_input t', t_output t', f (t_target t')) = (q1, x, y, q1')\<close> \<open>t' \<in> FSM.transitions M\<close> \<open>t'' \<in> FSM.transitions M\<close>)
then show "q1' = q1''"
using \<open>(f (t_source t''), t_input t'', t_output t'', f (t_target t'')) = (q1, x, y, q1'')\<close> \<open>(f (t_source t'), t_input t', t_output t', f (t_target t')) = (q1, x, y, q1')\<close> by auto
qed
then show ?thesis
unfolding observable_alt_def by blast
qed
lemma rename_states_minimal :
assumes "bij_betw f (states M) (f ` states M)"
and "minimal M"
shows "minimal (rename_states M f)"
proof -
have "\<And> q q' . q \<in> f ` FSM.states M \<Longrightarrow> q' \<in> f ` FSM.states M \<Longrightarrow> q \<noteq> q' \<Longrightarrow> LS (rename_states M f) q \<noteq> LS (rename_states M f) q'"
proof -
fix q q' assume "q \<in> f ` FSM.states M" and "q' \<in> f ` FSM.states M" and "q \<noteq> q'"
then obtain fq fq' where "fq \<in> states M" and "fq' \<in> states M" and "q = f fq" and "q' = f fq'"
by auto
then have "fq \<noteq> fq'"
using \<open>q \<noteq> q'\<close> by auto
then have "LS M fq \<noteq> LS M fq'"
by (meson \<open>fq \<in> FSM.states M\<close> \<open>fq' \<in> FSM.states M\<close> assms(2) minimal.elims(2))
then show "LS (rename_states M f) q \<noteq> LS (rename_states M f) q'"
using rename_states_isomorphism_language_state[OF assms(1)]
by (simp add: \<open>fq \<in> FSM.states M\<close> \<open>fq' \<in> FSM.states M\<close> \<open>q = f fq\<close> \<open>q' = f fq'\<close>)
qed
then show ?thesis
by auto
qed
fun index_states :: "('a::linorder,'b,'c) fsm \<Rightarrow> (nat,'b,'c) fsm" where
"index_states M = rename_states M (assign_indices (states M))"
lemma assign_indices_bij_betw: "bij_betw (assign_indices (FSM.states M)) (FSM.states M) (assign_indices (FSM.states M) ` FSM.states M)"
using assign_indices_bij[OF fsm_states_finite[of M]]
by (simp add: bij_betw_def)
lemma index_states_language :
"L (index_states M) = L M"
using rename_states_isomorphism_language[of "assign_indices (states M)" M, OF assign_indices_bij_betw]
by auto
lemma index_states_observable :
assumes "observable M"
shows "observable (index_states M)"
using rename_states_observable[of "assign_indices (states M)", OF assign_indices_bij_betw assms]
unfolding index_states.simps .
lemma index_states_minimal :
assumes "minimal M"
shows "minimal (index_states M)"
using rename_states_minimal[of "assign_indices (states M)", OF assign_indices_bij_betw assms]
unfolding index_states.simps .
fun index_states_integer :: "('a::linorder,'b,'c) fsm \<Rightarrow> (integer,'b,'c) fsm" where
"index_states_integer M = rename_states M (integer_of_nat \<circ> assign_indices (states M))"
lemma assign_indices_integer_bij_betw: "bij_betw (integer_of_nat \<circ> assign_indices (states M)) (FSM.states M) ((integer_of_nat \<circ> assign_indices (states M)) ` FSM.states M)"
proof -
have *: "inj_on (assign_indices (FSM.states M)) (FSM.states M)"
using assign_indices_bij[OF fsm_states_finite[of M]]
unfolding bij_betw_def
by auto
then have "inj_on (integer_of_nat \<circ> assign_indices (states M)) (FSM.states M)"
unfolding inj_on_def
by (metis comp_apply nat_of_integer_integer_of_nat)
then show ?thesis
unfolding bij_betw_def
by auto
qed
lemma index_states_integer_language :
"L (index_states_integer M) = L M"
using rename_states_isomorphism_language[of "integer_of_nat \<circ> assign_indices (states M)" M, OF assign_indices_integer_bij_betw]
by auto
lemma index_states_integer_observable :
assumes "observable M"
shows "observable (index_states_integer M)"
using rename_states_observable[of "integer_of_nat \<circ> assign_indices (states M)" M, OF assign_indices_integer_bij_betw assms]
unfolding index_states_integer.simps .
lemma index_states_integer_minimal :
assumes "minimal M"
shows "minimal (index_states_integer M)"
using rename_states_minimal[of "integer_of_nat \<circ> assign_indices (states M)" M, OF assign_indices_integer_bij_betw assms]
unfolding index_states_integer.simps .
subsection \<open>Canonical Separators\<close>
lift_definition canonical_separator' :: "('a,'b,'c) fsm \<Rightarrow> (('a \<times> 'a),'b,'c) fsm \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> (('a \<times> 'a) + 'a,'b,'c) fsm" is FSM_Impl.canonical_separator'
proof -
fix A :: "('a,'b,'c) fsm_impl"
fix B :: "('a \<times> 'a,'b,'c) fsm_impl"
fix q1 :: 'a
fix q2 :: 'a
assume "well_formed_fsm A" and "well_formed_fsm B"
then have p1a: "fsm_impl.initial A \<in> fsm_impl.states A"
and p2a: "finite (fsm_impl.states A)"
and p3a: "finite (fsm_impl.inputs A)"
and p4a: "finite (fsm_impl.outputs A)"
and p5a: "finite (fsm_impl.transitions A)"
and p6a: "(\<forall>t\<in>fsm_impl.transitions A.
t_source t \<in> fsm_impl.states A \<and>
t_input t \<in> fsm_impl.inputs A \<and> t_target t \<in> fsm_impl.states A \<and>
t_output t \<in> fsm_impl.outputs A)"
and p1b: "fsm_impl.initial B \<in> fsm_impl.states B"
and p2b: "finite (fsm_impl.states B)"
and p3b: "finite (fsm_impl.inputs B)"
and p4b: "finite (fsm_impl.outputs B)"
and p5b: "finite (fsm_impl.transitions B)"
and p6b: "(\<forall>t\<in>fsm_impl.transitions B.
t_source t \<in> fsm_impl.states B \<and>
t_input t \<in> fsm_impl.inputs B \<and> t_target t \<in> fsm_impl.states B \<and>
t_output t \<in> fsm_impl.outputs B)"
by simp+
let ?P = "FSM_Impl.canonical_separator' A B q1 q2"
show "well_formed_fsm ?P" proof (cases "fsm_impl.initial B = (q1,q2)")
case False
then show ?thesis by auto
next
case True
let ?f = "(\<lambda>qx . (case (set_as_map (image (\<lambda>(q,x,y,q') . ((q,x),y)) (fsm_impl.transitions A))) qx of Some yqs \<Rightarrow> yqs | None \<Rightarrow> {}))"
have "\<And> qx . (\<lambda>qx . (case (set_as_map (image (\<lambda>(q,x,y,q') . ((q,x),y)) (fsm_impl.transitions A))) qx of Some yqs \<Rightarrow> yqs | None \<Rightarrow> {})) qx = (\<lambda> qx . {z. (qx, z) \<in> (\<lambda>(q, x, y, q'). ((q, x), y)) ` fsm_impl.transitions A}) qx"
proof -
fix qx
show "\<And> qx . (\<lambda>qx . (case (set_as_map (image (\<lambda>(q,x,y,q') . ((q,x),y)) (fsm_impl.transitions A))) qx of Some yqs \<Rightarrow> yqs | None \<Rightarrow> {})) qx = (\<lambda> qx . {z. (qx, z) \<in> (\<lambda>(q, x, y, q'). ((q, x), y)) ` fsm_impl.transitions A}) qx"
unfolding set_as_map_def by (cases "\<exists>z. (qx, z) \<in> (\<lambda>(q, x, y, q'). ((q, x), y)) ` fsm_impl.transitions A"; auto)
qed
moreover have "\<And> qx . (\<lambda> qx . {z. (qx, z) \<in> (\<lambda>(q, x, y, q'). ((q, x), y)) ` fsm_impl.transitions A}) qx = (\<lambda> qx . {y | y . \<exists> q' . (fst qx, snd qx, y, q') \<in> fsm_impl.transitions A}) qx"
proof -
fix qx
show "(\<lambda> qx . {z. (qx, z) \<in> (\<lambda>(q, x, y, q'). ((q, x), y)) ` fsm_impl.transitions A}) qx = (\<lambda> qx . {y | y . \<exists> q' . (fst qx, snd qx, y, q') \<in> fsm_impl.transitions A}) qx"
by force
qed
ultimately have *:" ?f = (\<lambda> qx . {y | y . \<exists> q' . (fst qx, snd qx, y, q') \<in> fsm_impl.transitions A})"
by blast
let ?shifted_transitions' = "shifted_transitions (fsm_impl.transitions B)"
let ?distinguishing_transitions_lr = "distinguishing_transitions ?f q1 q2 (fsm_impl.states B) (fsm_impl.inputs B)"
let ?ts = "?shifted_transitions' \<union> ?distinguishing_transitions_lr"
have "FSM_Impl.states ?P = (image Inl (FSM_Impl.states B)) \<union> {Inr q1, Inr q2}"
and "FSM_Impl.transitions ?P = ?ts"
unfolding FSM_Impl.canonical_separator'.simps Let_def True by simp+
have p2: "finite (fsm_impl.states ?P)"
unfolding \<open>FSM_Impl.states ?P = (image Inl (FSM_Impl.states B)) \<union> {Inr q1, Inr q2}\<close> using p2b by blast
have "fsm_impl.initial ?P = Inl (q1,q2)" by auto
then have p1: "fsm_impl.initial ?P \<in> fsm_impl.states ?P"
using p1a p1b unfolding canonical_separator'.simps True by auto
have p3: "finite (fsm_impl.inputs ?P)"
using p3a p3b by auto
have p4: "finite (fsm_impl.outputs ?P)"
using p4a p4b by auto
have "finite (fsm_impl.states B \<times> fsm_impl.inputs B)"
using p2b p3b by blast
moreover have **: "\<And> x q1 . finite ({y |y. \<exists>q'. (fst (q1, x), snd (q1, x), y, q') \<in> fsm_impl.transitions A})"
proof -
fix x q1
have "{y |y. \<exists>q'. (fst (q1, x), snd (q1, x), y, q') \<in> fsm_impl.transitions A} = {t_output t | t . t \<in> fsm_impl.transitions A \<and> t_source t = q1 \<and> t_input t = x}"
by auto
then have "{y |y. \<exists>q'. (fst (q1, x), snd (q1, x), y, q') \<in> fsm_impl.transitions A} \<subseteq> image t_output (fsm_impl.transitions A)"
unfolding fst_conv snd_conv by blast
moreover have "finite (image t_output (fsm_impl.transitions A))"
using p5a by auto
ultimately show "finite ({y |y. \<exists>q'. (fst (q1, x), snd (q1, x), y, q') \<in> fsm_impl.transitions A})"
by (simp add: finite_subset)
qed
ultimately have "finite ?distinguishing_transitions_lr"
unfolding * distinguishing_transitions_def by force
moreover have "finite ?shifted_transitions'"
unfolding shifted_transitions_def using p5b by auto
ultimately have "finite ?ts" by blast
then have p5: "finite (fsm_impl.transitions ?P)"
by simp
have "fsm_impl.inputs ?P = fsm_impl.inputs A \<union> fsm_impl.inputs B"
using True by auto
have "fsm_impl.outputs ?P = fsm_impl.outputs A \<union> fsm_impl.outputs B"
using True by auto
have "\<And> t . t \<in> ?shifted_transitions' \<Longrightarrow> t_source t \<in> fsm_impl.states ?P \<and> t_target t \<in> fsm_impl.states ?P"
unfolding \<open>FSM_Impl.states ?P = (image Inl (FSM_Impl.states B)) \<union> {Inr q1, Inr q2}\<close> shifted_transitions_def
using p6b by force
moreover have "\<And> t . t \<in> ?distinguishing_transitions_lr \<Longrightarrow> t_source t \<in> fsm_impl.states ?P \<and> t_target t \<in> fsm_impl.states ?P"
unfolding \<open>FSM_Impl.states ?P = (image Inl (FSM_Impl.states B)) \<union> {Inr q1, Inr q2}\<close> distinguishing_transitions_def * by force
ultimately have "\<And> t . t \<in> ?ts \<Longrightarrow> t_source t \<in> fsm_impl.states ?P \<and> t_target t \<in> fsm_impl.states ?P"
by blast
moreover have "\<And> t . t \<in> ?shifted_transitions' \<Longrightarrow> t_input t \<in> fsm_impl.inputs ?P \<and> t_output t \<in> fsm_impl.outputs ?P"
proof -
have "\<And> t . t \<in> ?shifted_transitions' \<Longrightarrow> t_input t \<in> fsm_impl.inputs B \<and> t_output t \<in> fsm_impl.outputs B"
unfolding shifted_transitions_def using p6b by auto
then show "\<And> t . t \<in> ?shifted_transitions' \<Longrightarrow> t_input t \<in> fsm_impl.inputs ?P \<and> t_output t \<in> fsm_impl.outputs ?P"
unfolding \<open>fsm_impl.inputs ?P = fsm_impl.inputs A \<union> fsm_impl.inputs B\<close>
\<open>fsm_impl.outputs ?P = fsm_impl.outputs A \<union> fsm_impl.outputs B\<close> by blast
qed
moreover have "\<And> t . t \<in> ?distinguishing_transitions_lr \<Longrightarrow> t_input t \<in> fsm_impl.inputs ?P \<and> t_output t \<in> fsm_impl.outputs ?P"
unfolding * distinguishing_transitions_def using p6a p6b True by auto
ultimately have p6: "(\<forall>t\<in>fsm_impl.transitions ?P.
t_source t \<in> fsm_impl.states ?P \<and>
t_input t \<in> fsm_impl.inputs ?P \<and> t_target t \<in> fsm_impl.states ?P \<and>
t_output t \<in> fsm_impl.outputs ?P)"
unfolding \<open>FSM_Impl.transitions ?P = ?ts\<close> by blast
show "well_formed_fsm ?P"
using p1 p2 p3 p4 p5 p6 by linarith
qed
qed
lemma canonical_separator'_simps :
assumes "initial P = (q1,q2)"
shows "initial (canonical_separator' M P q1 q2) = Inl (q1,q2)"
"states (canonical_separator' M P q1 q2) = (image Inl (states P)) \<union> {Inr q1, Inr q2}"
"inputs (canonical_separator' M P q1 q2) = inputs M \<union> inputs P"
"outputs (canonical_separator' M P q1 q2) = outputs M \<union> outputs P"
"transitions (canonical_separator' M P q1 q2)
= shifted_transitions (transitions P)
\<union> distinguishing_transitions (h_out M) q1 q2 (states P) (inputs P)"
using assms unfolding h_out_code by (transfer; auto)+
lemma canonical_separator'_simps_without_assm :
"initial (canonical_separator' M P q1 q2) = Inl (q1,q2)"
"states (canonical_separator' M P q1 q2) = (if initial P = (q1,q2) then (image Inl (states P)) \<union> {Inr q1, Inr q2} else {Inl (q1,q2)})"
"inputs (canonical_separator' M P q1 q2) = (if initial P = (q1,q2) then inputs M \<union> inputs P else {})"
"outputs (canonical_separator' M P q1 q2) = (if initial P = (q1,q2) then outputs M \<union> outputs P else {})"
"transitions (canonical_separator' M P q1 q2) = (if initial P = (q1,q2) then shifted_transitions (transitions P) \<union> distinguishing_transitions (h_out M) q1 q2 (states P) (inputs P) else {})"
unfolding h_out_code by (transfer; simp add: Let_def)+
end
\ No newline at end of file
diff --git a/thys/FSM_Tests/Minimisation.thy b/thys/FSM_Tests/Minimisation.thy
--- a/thys/FSM_Tests/Minimisation.thy
+++ b/thys/FSM_Tests/Minimisation.thy
@@ -1,1846 +1,1846 @@
section \<open>Minimisation by OFSM Tables\<close>
text \<open>This theory presents the classical algorithm for transforming observable FSMs into
language-equivalent observable and minimal FSMs in analogy to the minimisation of
finite automata.\<close>
theory Minimisation
imports FSM
begin
subsection \<open>OFSM Tables\<close>
text \<open>OFSM tables partition the states of an FSM based on an initial partition and an iteration
counter.
States are in the same element of the 0th table iff they are in the same element of the
initial partition.
States q1, q2 are in the same element of the (k+1)-th table if they are in the same element of
the k-th table and furthermore for each IO pair (x,y) either (x,y) is not in the language of
both q1 and q2 or it is in the language of both states and the states q1', q2' reached via
(x,y) from q1 and q2, respectively, are in the same element of the k-th table.\<close>
fun ofsm_table :: "('a,'b,'c) fsm \<Rightarrow> ('a \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a set" where
"ofsm_table M f 0 q = (if q \<in> states M then f q else {})" |
"ofsm_table M f (Suc k) q = (let
prev_table = ofsm_table M f k
in {q' \<in> prev_table q . \<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> prev_table qT = prev_table qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None) })"
lemma ofsm_table_non_state :
assumes "q \<notin> states M"
shows "ofsm_table M f k q = {}"
using assms by (induction k; auto)
lemma ofsm_table_subset:
assumes "i \<le> j"
shows "ofsm_table M f j q \<subseteq> ofsm_table M f i q"
proof -
have *: "\<And> k . ofsm_table M f (Suc k) q \<subseteq> ofsm_table M f k q"
proof -
fix k show "ofsm_table M f (Suc k) q \<subseteq> ofsm_table M f k q"
proof (cases k)
case 0
show ?thesis unfolding 0 ofsm_table.simps Let_def by blast
next
case (Suc k')
show ?thesis
unfolding Suc ofsm_table.simps Let_def by force
qed
qed
show ?thesis
using assms
proof (induction j)
case 0
then show ?case by auto
next
case (Suc x)
then show ?case using *[of x]
using le_SucE by blast
qed
qed
lemma ofsm_table_case_helper :
"(case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)
= ((\<exists> qT qT' . h_obs M q x y = Some qT \<and> h_obs M q' x y = Some qT' \<and> ofsm_table M f k qT = ofsm_table M f k qT') \<or> (h_obs M q x y = None \<and> h_obs M q' x y = None))"
proof -
have *: "\<And> a b P . (case a of Some a' \<Rightarrow> (case b of Some b' \<Rightarrow> P a' b' | None \<Rightarrow> False) | None \<Rightarrow> b = None)
= ((\<exists> a' b' . a = Some a' \<and> b = Some b' \<and> P a' b') \<or> (a = None \<and> b = None))"
(is "\<And> a b P . ?P1 a b P = ?P2 a b P")
proof
fix a b P
show "?P1 a b P \<Longrightarrow> ?P2 a b P" using case_optionE[of "b = None" "\<lambda>a' . (case b of Some b' \<Rightarrow> P a' b' | None \<Rightarrow> False)" a]
by (metis case_optionE)
show "?P2 a b P \<Longrightarrow> ?P1 a b P" by auto
qed
show ?thesis
using *[of "h_obs M q' x y" "\<lambda>qT qT' . ofsm_table M f k qT = ofsm_table M f k qT'" "h_obs M q x y"] .
qed
lemma ofsm_table_case_helper_neg :
"(\<not> (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None))
= ((\<exists> qT qT' . h_obs M q x y = Some qT \<and> h_obs M q' x y = Some qT' \<and> ofsm_table M f k qT \<noteq> ofsm_table M f k qT') \<or> (h_obs M q x y = None \<longleftrightarrow> h_obs M q' x y \<noteq> None))"
unfolding ofsm_table_case_helper by force
lemma ofsm_table_fixpoint :
assumes "i \<le> j"
and "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f (Suc i) q = ofsm_table M f i q"
and "q \<in> states M"
shows "ofsm_table M f j q = ofsm_table M f i q"
proof -
have *: "\<And> k . k \<ge> i \<Longrightarrow> (\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f (Suc k) q = ofsm_table M f k q)"
proof -
fix k :: nat assume "k \<ge> i"
then show "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f (Suc k) q = ofsm_table M f k q"
proof (induction k)
case 0
then show ?case using assms(2) by auto
next
case (Suc k)
show "ofsm_table M f (Suc (Suc k)) q = ofsm_table M f (Suc k) q"
proof (cases "i = Suc k")
case True
then show ?thesis using assms(2)[OF \<open>q \<in> states M\<close>] by simp
next
case False
then have "i \<le> k"
using \<open>i \<le> Suc k\<close> by auto
have h_obs_state: "\<And> q x y qT . h_obs M q x y = Some qT \<Longrightarrow> qT \<in> states M"
using h_obs_state by fastforce
show ?thesis
proof (rule ccontr)
assume "ofsm_table M f (Suc (Suc k)) q \<noteq> ofsm_table M f (Suc k) q"
moreover have "ofsm_table M f (Suc (Suc k)) q \<subseteq> ofsm_table M f (Suc k) q"
using ofsm_table_subset
by (metis (full_types) Suc_n_not_le_n nat_le_linear)
ultimately obtain q' where "q' \<notin> {q' \<in> ofsm_table M f (Suc k) q . \<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f (Suc k) qT = ofsm_table M f (Suc k) qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None) }"
and "q' \<in> ofsm_table M f (Suc k) q"
using ofsm_table.simps(2)[of M f "Suc k" q] unfolding Let_def by blast
then have "\<not>(\<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f (Suc k) qT = ofsm_table M f (Suc k) qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None))"
by blast
then obtain x y where "x \<in> inputs M" and "y \<in> outputs M" and "\<not> (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f (Suc k) qT = ofsm_table M f (Suc k) qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
by blast
then consider "\<exists> qT qT' . h_obs M q x y = Some qT \<and> h_obs M q' x y = Some qT' \<and> ofsm_table M f (Suc k) qT \<noteq> ofsm_table M f (Suc k) qT'" |
"(h_obs M q x y = None \<longleftrightarrow> h_obs M q' x y \<noteq> None)"
unfolding ofsm_table_case_helper_neg by blast
then show False proof cases
case 1
then obtain qT qT' where "h_obs M q x y = Some qT" and "h_obs M q' x y = Some qT'" and "ofsm_table M f (Suc k) qT \<noteq> ofsm_table M f (Suc k) qT'"
by blast
then have "ofsm_table M f k qT \<noteq> ofsm_table M f k qT'"
using Suc.IH[OF h_obs_state[OF \<open>h_obs M q x y = Some qT\<close>] \<open>i \<le> k\<close>]
Suc.IH[OF h_obs_state[OF \<open>h_obs M q' x y = Some qT'\<close>] \<open>i \<le> k\<close>]
by fast
moreover have "q' \<in> ofsm_table M f k q"
using ofsm_table_subset[of k "Suc k"] \<open>q' \<in> ofsm_table M f (Suc k) q\<close> by force
ultimately have "ofsm_table M f (Suc k) q \<noteq> ofsm_table M f k q"
using \<open>x \<in> inputs M\<close> \<open>y \<in> outputs M\<close> \<open>h_obs M q x y = Some qT\<close> \<open>h_obs M q' x y = Some qT'\<close>
unfolding ofsm_table.simps(2) Let_def by force
then show ?thesis
using Suc.IH[OF Suc.prems(1) \<open>i \<le> k\<close>] by simp
next
case 2
then have "\<not> (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
unfolding ofsm_table_case_helper_neg by blast
moreover have "q' \<in> ofsm_table M f k q"
using ofsm_table_subset[of k "Suc k"] \<open>q' \<in> ofsm_table M f (Suc k) q\<close> by force
ultimately have "ofsm_table M f (Suc k) q \<noteq> ofsm_table M f k q"
using \<open>x \<in> inputs M\<close> \<open>y \<in> outputs M\<close>
unfolding ofsm_table.simps(2) Let_def by force
then show ?thesis
using Suc.IH[OF Suc.prems(1) \<open>i \<le> k\<close>] by simp
qed
qed
qed
qed
qed
show ?thesis
using assms(1) proof (induction "j")
case 0
then show ?case by auto
next
case (Suc j)
show ?case proof (cases "i = Suc j")
case True
then show ?thesis by simp
next
case False
then have "i \<le> j"
using Suc.prems(1) by auto
then have "ofsm_table M f j q = ofsm_table M f i q"
using Suc.IH by auto
moreover have "ofsm_table M f (Suc j) q = ofsm_table M f j q"
using *[OF \<open>i\<le>j\<close> \<open>q\<in>states M\<close>] by assumption
ultimately show ?thesis
by blast
qed
qed
qed
(* restricts the range of the supplied function to the states of the FSM - required for (easy) termination *)
function ofsm_table_fix :: "('a,'b,'c) fsm \<Rightarrow> ('a \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a set" where
"ofsm_table_fix M f k = (let
cur_table = ofsm_table M (\<lambda>q. f q \<inter> states M) k;
next_table = ofsm_table M (\<lambda>q. f q \<inter> states M) (Suc k)
in if (\<forall> q \<in> states M . cur_table q = next_table q)
then cur_table
else ofsm_table_fix M f (Suc k))"
by pat_completeness auto
termination
proof -
{
fix M :: "('a,'b,'c) fsm"
and f :: "('a \<Rightarrow> 'a set)"
and k :: nat
define f' where f': "f' = (\<lambda>q. f q \<inter> states M)"
assume "\<exists>q\<in>FSM.states M. ofsm_table M (\<lambda>q. f q \<inter> states M) k q \<noteq> ofsm_table M (\<lambda>q. f q \<inter> states M) (Suc k) q"
then obtain q where "q \<in> states M"
and "ofsm_table M f' k q \<noteq> ofsm_table M f' (Suc k) q"
unfolding f' by blast
have *: "\<And> k . (\<Sum>q\<in>FSM.states M. card (ofsm_table M f' k q)) = card (ofsm_table M f' k q) + (\<Sum>q\<in>FSM.states M - {q}. card (ofsm_table M f' k q))"
using \<open>q \<in> states M\<close> by (meson fsm_states_finite sum.remove)
have "\<And> q . ofsm_table M f' (Suc k) q \<subseteq> ofsm_table M f' k q"
using ofsm_table_subset[of k "Suc k" M ] by auto
moreover have "\<And> q . finite (ofsm_table M f' k q)"
proof -
fix q
have "ofsm_table M (\<lambda>q. f q \<inter> states M) k q \<subseteq> ofsm_table M (\<lambda>q. f q \<inter> states M) 0 q"
using ofsm_table_subset[of 0 k M "(\<lambda>q. f q \<inter> FSM.states M)" q] by auto
then have "ofsm_table M f' k q \<subseteq> states M"
unfolding f'
using ofsm_table_non_state[of q M "(\<lambda>q. f q \<inter> FSM.states M)" k]
by force
then show "finite (ofsm_table M f' k q)"
using fsm_states_finite finite_subset by auto
qed
ultimately have "\<And> q . card (ofsm_table M f' (Suc k) q) \<le> card (ofsm_table M f' k q)"
by (simp add: card_mono)
then have "(\<Sum>q\<in>FSM.states M - {q}. card (ofsm_table M f' (Suc k) q)) \<le> (\<Sum>q\<in>FSM.states M - {q}. card (ofsm_table M f' k q))"
by (simp add: sum_mono)
moreover have "card (ofsm_table M f' (Suc k) q) < card (ofsm_table M f' k q)"
using \<open>ofsm_table M f' k q \<noteq> ofsm_table M f' (Suc k) q\<close> \<open>ofsm_table M f' (Suc k) q \<subseteq> ofsm_table M f' k q\<close> \<open>finite (ofsm_table M f' k q)\<close>
by (metis psubsetI psubset_card_mono)
ultimately have "(\<Sum>q\<in>FSM.states M. card (ofsm_table M (\<lambda>q. f q \<inter> states M) (Suc k) q)) < (\<Sum>q\<in>FSM.states M. card (ofsm_table M (\<lambda>q. f q \<inter> states M) k q))"
unfolding f'[symmetric] *
by linarith
} note t = this
show ?thesis
apply (relation "measure (\<lambda> (M, f, k) . \<Sum> q \<in> states M . card (ofsm_table M (\<lambda>q. f q \<inter> states M) k q))")
apply (simp del: h_obs.simps ofsm_table.simps)+
by (erule t)
qed
lemma ofsm_table_restriction_to_states :
assumes "\<And> q . q \<in> states M \<Longrightarrow> f q \<subseteq> states M"
and "q \<in> states M"
shows "ofsm_table M f k q = ofsm_table M (\<lambda>q . f q \<inter> states M) k q"
using assms(2) proof (induction k arbitrary: q)
case 0
then show ?case using assms(1) by auto
next
case (Suc k)
have "\<And> x y q q' . (case h_obs M q x y of None \<Rightarrow> h_obs M q' x y = None | Some qT \<Rightarrow> (case h_obs M q' x y of None \<Rightarrow> False | Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT'))
= (case h_obs M q x y of None \<Rightarrow> h_obs M q' x y = None | Some qT \<Rightarrow> (case h_obs M q' x y of None \<Rightarrow> False | Some qT' \<Rightarrow> ofsm_table M (\<lambda>q . f q \<inter> states M) k qT = ofsm_table M (\<lambda>q . f q \<inter> states M) k qT'))"
(is "\<And> x y q q' . ?C1 x y q q' = ?C2 x y q q' ")
proof -
fix x y q q'
show "?C1 x y q q' = ?C2 x y q q'"
using Suc.IH[OF h_obs_state, of q x y]
using Suc.IH[OF h_obs_state, of q' x y]
by (cases "h_obs M q x y"; cases "h_obs M q' x y"; auto)
qed
then show ?case
unfolding ofsm_table.simps Let_def Suc.IH[OF Suc.prems]
by blast
qed
lemma ofsm_table_fix_length :
assumes "\<And> q . q \<in> states M \<Longrightarrow> f q \<subseteq> states M"
obtains k where "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table_fix M f 0 q = ofsm_table M f k q" and "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> k \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f k q"
proof -
have "\<exists> k . \<forall> q \<in> states M . \<forall> k' \<ge> k . ofsm_table M f k' q = ofsm_table M f k q"
proof -
have "\<exists> fp . \<forall> q k' . q \<in> states M \<longrightarrow> k' \<ge> (fp q) \<longrightarrow> ofsm_table M f k' q = ofsm_table M f (fp q) q"
proof
fix q
let ?assignK = "\<lambda> q . SOME k . \<forall> k' \<ge> k . ofsm_table M f k' q = ofsm_table M f k q"
have "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> ?assignK q \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f (?assignK q) q"
proof -
fix q k' assume "q \<in> states M" and "k' \<ge> ?assignK q"
then have p1: "finite (ofsm_table M f 0 q)"
using fsm_states_finite assms(1)
using infinite_super by fastforce
have "\<exists> k . \<forall> k' \<ge> k . ofsm_table M f k' q = ofsm_table M f k q"
using finite_subset_mapping_limit[of "\<lambda> k . ofsm_table M f k q", OF p1 ofsm_table_subset] by metis
have "\<forall> k' \<ge> (?assignK q) . ofsm_table M f k' q = ofsm_table M f (?assignK q) q"
using someI_ex[of "\<lambda> k . \<forall> k' \<ge> k . ofsm_table M f k' q = ofsm_table M f k q", OF \<open>\<exists> k . \<forall> k' \<ge> k . ofsm_table M f k' q = ofsm_table M f k q\<close>] by assumption
then show "ofsm_table M f k' q = ofsm_table M f (?assignK q) q"
using \<open>k' \<ge> ?assignK q\<close> by blast
qed
then show "\<forall>q k'. q \<in> states M \<longrightarrow> ?assignK q \<le> k' \<longrightarrow> ofsm_table M f k' q = ofsm_table M f (?assignK q) q"
by blast
qed
then obtain assignK where assignK_prop: "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> assignK q \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f (assignK q) q"
by blast
have "finite (assignK ` states M)"
by (simp add: fsm_states_finite)
moreover have "assignK ` FSM.states M \<noteq> {}"
using fsm_initial by auto
ultimately obtain k where "k \<in> (assignK ` states M)" and "\<And> k' . k' \<in> (assignK ` states M) \<Longrightarrow> k' \<le> k"
using Max_elem[OF \<open>finite (assignK ` states M)\<close> \<open>assignK ` FSM.states M \<noteq> {}\<close>] by (meson eq_Max_iff)
have "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> k \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f k q"
proof -
fix q k' assume "k' \<ge> k" and "q \<in> states M"
then have "k' \<ge> assignK q"
using \<open>\<And> k' . k' \<in> (assignK ` states M) \<Longrightarrow> k' \<le> k\<close>
using dual_order.trans by auto
then show "ofsm_table M f k' q = ofsm_table M f k q"
using assignK_prop \<open>\<And>k'. k' \<in> assignK ` FSM.states M \<Longrightarrow> k' \<le> k\<close> \<open>q \<in> FSM.states M\<close> by blast
qed
then show ?thesis
by blast
qed
then obtain k where k_prop: "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> k \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f k q"
by blast
then have "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f k q = ofsm_table M f (Suc k) q"
by (metis (full_types) le_SucI order_refl)
let ?ks = "(Set.filter (\<lambda> k . \<forall> q \<in> states M . ofsm_table M f k q = ofsm_table M f (Suc k) q) {..k})"
have f1: "finite ?ks"
by simp
moreover have f2: "?ks \<noteq> {}"
using \<open>\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f k q = ofsm_table M f (Suc k) q\<close> unfolding Set.filter_def by blast
ultimately obtain kMin where "kMin \<in> ?ks" and "\<And> k' . k' \<in> ?ks \<Longrightarrow> k' \<ge> kMin"
using Min_elem[OF f1 f2] by (meson eq_Min_iff)
have k1: "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f (Suc kMin) q = ofsm_table M f kMin q"
using \<open>kMin \<in> ?ks\<close>
by (metis (mono_tags, lifting) member_filter)
have k2: "\<And> k' . (\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f (Suc k') q) \<Longrightarrow> k' \<ge> kMin"
proof -
fix k' assume "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f (Suc k') q"
show "k' \<ge> kMin" proof (cases "k' \<in> ?ks")
case True
then show ?thesis using \<open>\<And> k' . k' \<in> ?ks \<Longrightarrow> k' \<ge> kMin\<close> by blast
next
case False
then have "k' > k"
using \<open>\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f (Suc k') q\<close>
unfolding member_filter atMost_iff
by (meson not_less)
moreover have "kMin \<le> k"
using \<open>kMin \<in> ?ks\<close> by auto
ultimately show ?thesis
by auto
qed
qed
have "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table_fix M f 0 q = ofsm_table M (\<lambda> q . f q \<inter> states M) kMin q"
proof -
fix q assume "q \<in> states M"
show "ofsm_table_fix M f 0 q = ofsm_table M (\<lambda> q . f q \<inter> states M) kMin q"
proof (cases kMin)
case 0
have "\<forall>q\<in>FSM.states M. ofsm_table M (\<lambda>q. f q \<inter> FSM.states M) 0 q = ofsm_table M (\<lambda>q. f q \<inter> FSM.states M) (Suc 0) q"
using k1
using ofsm_table_restriction_to_states[of M f _, OF assms(1) _ ]
using "0" by blast
then show ?thesis
apply (subst ofsm_table_fix.simps)
unfolding "0" Let_def by force
next
case (Suc kMin')
have *: "\<And> i . i < kMin \<Longrightarrow> \<not>(\<forall> q \<in> states M . ofsm_table M f i q = ofsm_table M f (Suc i) q)"
using k2
by (meson leD)
have "\<And> i . i < kMin \<Longrightarrow> ofsm_table_fix M f 0 = ofsm_table_fix M f (Suc i)"
proof -
fix i assume "i < kMin"
then show "ofsm_table_fix M f 0 = ofsm_table_fix M f (Suc i)"
proof (induction i)
case 0
show ?case
using *[OF 0] ofsm_table_restriction_to_states[of _ f, OF assms(1) _ ] unfolding ofsm_table_fix.simps[of M f 0] Let_def
by (metis (no_types, lifting))
next
case (Suc i)
then have "i < kMin" by auto
have "ofsm_table_fix M f (Suc i) = ofsm_table_fix M f (Suc (Suc i))"
using *[OF \<open>Suc i < kMin\<close>] ofsm_table_restriction_to_states[of _ f, OF assms(1) _ ] unfolding ofsm_table_fix.simps[of M f "Suc i"] Let_def by metis
then show ?case using Suc.IH[OF \<open>i < kMin\<close>]
by presburger
qed
qed
then have "ofsm_table_fix M f 0 = ofsm_table_fix M f kMin"
using Suc by blast
moreover have "ofsm_table_fix M f kMin q = ofsm_table M f kMin q"
proof -
have "\<forall>q\<in>FSM.states M. ofsm_table M (\<lambda>q. f q \<inter> FSM.states M) kMin q = ofsm_table M (\<lambda>q. f q \<inter> FSM.states M) (Suc kMin) q"
using ofsm_table_restriction_to_states[of _ f, OF assms(1) _ ]
using k1 by blast
then show ?thesis
using ofsm_table_restriction_to_states[of _ f, OF assms(1) _ ] \<open>q \<in> states M\<close>
unfolding ofsm_table_fix.simps[of M f kMin] Let_def
by presburger
qed
ultimately show ?thesis
using ofsm_table_restriction_to_states[of _ f, OF assms(1) \<open>q \<in> states M\<close>]
by presburger
qed
qed
moreover have "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> kMin \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f kMin q"
using ofsm_table_fixpoint[OF _ k1 ] by blast
ultimately show ?thesis
using that[of kMin]
using ofsm_table_restriction_to_states[of M f, OF assms(1) _ ]
by blast
qed
lemma ofsm_table_containment :
assumes "q \<in> states M"
and "\<And> q . q \<in> states M \<Longrightarrow> q \<in> f q"
shows "q \<in> ofsm_table M f k q"
proof (induction k)
case 0
then show ?case using assms by auto
next
case (Suc k)
then show ?case
unfolding ofsm_table.simps Let_def option.case_eq_if
by auto
qed
lemma ofsm_table_states :
assumes "\<And> q . q \<in> states M \<Longrightarrow> f q \<subseteq> states M"
and "q \<in> states M"
shows "ofsm_table M f k q \<subseteq> states M"
proof -
have "ofsm_table M f k q \<subseteq> ofsm_table M f 0 q"
using ofsm_table_subset[OF le0] by metis
moreover have "ofsm_table M f 0 q \<subseteq> states M"
using assms
unfolding ofsm_table.simps(1) by (metis (full_types))
ultimately show ?thesis
by blast
qed
subsubsection \<open>Properties of Initial Partitions\<close>
definition equivalence_relation_on_states :: "('a,'b,'c) fsm \<Rightarrow> ('a \<Rightarrow> 'a set) \<Rightarrow> bool" where
"equivalence_relation_on_states M f =
(equiv (states M) {(q1,q2) | q1 q2 . q1 \<in> states M \<and> q2 \<in> f q1}
\<and> (\<forall> q \<in> states M . f q \<subseteq> states M))"
lemma equivalence_relation_on_states_refl :
assumes "equivalence_relation_on_states M f"
and "q \<in> states M"
shows "q \<in> f q"
using assms unfolding equivalence_relation_on_states_def equiv_def refl_on_def by blast
lemma equivalence_relation_on_states_sym :
assumes "equivalence_relation_on_states M f"
and "q1 \<in> states M"
and "q2 \<in> f q1"
shows "q1 \<in> f q2"
using assms unfolding equivalence_relation_on_states_def equiv_def sym_def by blast
lemma equivalence_relation_on_states_trans :
assumes "equivalence_relation_on_states M f"
and "q1 \<in> states M"
and "q2 \<in> f q1"
and "q3 \<in> f q2"
shows "q3 \<in> f q1"
proof -
have "(q1,q2) \<in> {(q1,q2) | q1 q2 . q1 \<in> states M \<and> q2 \<in> f q1}"
using assms(2,3) by blast
then have "q2 \<in> states M"
using assms(1) unfolding equivalence_relation_on_states_def
by auto
then have "(q2,q3) \<in> {(q1,q2) | q1 q2 . q1 \<in> states M \<and> q2 \<in> f q1}"
using assms(4) by blast
moreover have "trans {(q1,q2) | q1 q2 . q1 \<in> states M \<and> q2 \<in> f q1}"
using assms(1) unfolding equivalence_relation_on_states_def equiv_def by auto
ultimately show ?thesis
using \<open>(q1,q2) \<in> {(q1,q2) | q1 q2 . q1 \<in> states M \<and> q2 \<in> f q1}\<close>
unfolding trans_def by blast
qed
lemma equivalence_relation_on_states_ran :
assumes "equivalence_relation_on_states M f"
and "q \<in> states M"
shows "f q \<subseteq> states M"
using assms unfolding equivalence_relation_on_states_def by blast
subsubsection \<open>Properties of OFSM tables for initial partitions based on equivalence relations\<close>
lemma h_obs_io :
assumes "h_obs M q x y = Some q'"
shows "x \<in> inputs M" and "y \<in> outputs M"
proof -
have "snd ` Set.filter (\<lambda> (y',q') . y' = y) (h M (q,x)) \<noteq> {}"
using assms unfolding h_obs_simps Let_def by auto
then show "x \<in> inputs M" and "y \<in> outputs M"
unfolding h_simps
using fsm_transition_input fsm_transition_output
by fastforce+
qed
lemma ofsm_table_language :
assumes "q' \<in> ofsm_table M f k q"
and "length io \<le> k"
and "q \<in> states M"
and "equivalence_relation_on_states M f"
shows "is_in_language M q io \<longleftrightarrow> is_in_language M q' io"
and "is_in_language M q io \<Longrightarrow> (after M q' io) \<in> f (after M q io)"
proof -
have "(is_in_language M q io \<longleftrightarrow> is_in_language M q' io) \<and> (is_in_language M q io \<longrightarrow> (after M q' io) \<in> f (after M q io))"
using assms(1,2,3)
proof (induction k arbitrary: q q' io)
case 0
then have "io = []" by auto
then show ?case
using "0.prems"(1,3) by auto
next
case (Suc k)
show ?case proof (cases "length io \<le> k")
case True
have *: "q' \<in> ofsm_table M f k q"
using \<open>q' \<in> ofsm_table M f (Suc k) q\<close> ofsm_table_subset
by (metis (full_types) le_SucI order_refl subsetD)
show ?thesis using Suc.IH[OF * True \<open>q \<in> states M\<close>] by assumption
next
case False
then have "length io = Suc k"
using \<open>length io \<le> Suc k\<close> by auto
then obtain ioT ioP where "io = ioT#ioP"
by (meson length_Suc_conv)
then have "length ioP \<le> k"
using \<open>length io \<le> Suc k\<close> by auto
obtain x y where "io = (x,y)#ioP"
using \<open>io = ioT#ioP\<close> prod.exhaust_sel
by fastforce
have "ofsm_table M f (Suc k) q = {q' \<in> ofsm_table M f k q . \<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None) }"
unfolding ofsm_table.simps Let_def by blast
then have "q' \<in> ofsm_table M f k q"
and *: "\<And> x y . x \<in> inputs M \<Longrightarrow> y \<in> outputs M \<Longrightarrow> (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
using \<open>q' \<in> ofsm_table M f (Suc k) q\<close> by blast+
show ?thesis
unfolding \<open>io = (x,y)#ioP\<close>
proof -
have "is_in_language M q ((x,y)#ioP) \<Longrightarrow> is_in_language M q' ((x,y)#ioP) \<and> after M q' ((x,y)#ioP) \<in> f (after M q ((x,y)#ioP))"
proof -
assume "is_in_language M q ((x,y)#ioP)"
then obtain qT where "h_obs M q x y = Some qT" and "is_in_language M qT ioP"
by (metis case_optionE is_in_language.simps(2))
moreover have "(case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
using *[of x y, OF h_obs_io[OF \<open>h_obs M q x y = Some qT\<close>]] .
ultimately obtain qT' where "h_obs M q' x y = Some qT'" and "ofsm_table M f k qT = ofsm_table M f k qT'"
using ofsm_table_case_helper[of M q' x y f k q]
unfolding ofsm_table.simps by force
then have "qT' \<in> ofsm_table M f k qT"
using ofsm_table_containment[OF h_obs_state equivalence_relation_on_states_refl[OF \<open>equivalence_relation_on_states M f\<close>]]
by metis
have "(is_in_language M qT ioP) = (is_in_language M qT' ioP)"
"(is_in_language M qT ioP \<longrightarrow> after M qT' ioP \<in> f (after M qT ioP))"
using Suc.IH[OF \<open>qT' \<in> ofsm_table M f k qT\<close> \<open>length ioP \<le> k\<close> h_obs_state[OF \<open>h_obs M q x y = Some qT\<close>]]
by blast+
have "(is_in_language M qT' ioP)"
using \<open>(is_in_language M qT ioP) = (is_in_language M qT' ioP)\<close> \<open>is_in_language M qT ioP\<close>
by auto
then have "is_in_language M q' ((x,y)#ioP)"
unfolding is_in_language.simps \<open>h_obs M q' x y = Some qT'\<close> by auto
moreover have "after M q' ((x,y)#ioP) \<in> f (after M q ((x,y)#ioP))"
unfolding after.simps \<open>h_obs M q' x y = Some qT'\<close> \<open>h_obs M q x y = Some qT\<close>
using \<open>(is_in_language M qT ioP \<longrightarrow> after M qT' ioP \<in> f (after M qT ioP))\<close> \<open>is_in_language M qT ioP\<close>
by auto
ultimately show "is_in_language M q' ((x,y)#ioP) \<and> after M q' ((x,y)#ioP) \<in> f (after M q ((x,y)#ioP))"
by blast
qed
moreover have "is_in_language M q' ((x,y)#ioP) \<Longrightarrow> is_in_language M q ((x,y)#ioP)"
proof -
assume "is_in_language M q' ((x,y)#ioP)"
then obtain qT' where "h_obs M q' x y = Some qT'" and "is_in_language M qT' ioP"
by (metis case_optionE is_in_language.simps(2))
moreover have "(case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
using *[of x y, OF h_obs_io[OF \<open>h_obs M q' x y = Some qT'\<close>]] .
ultimately obtain qT where "h_obs M q x y = Some qT" and "ofsm_table M f k qT = ofsm_table M f k qT'"
using ofsm_table_case_helper[of M q' x y f k q]
unfolding ofsm_table.simps by force
then have "qT \<in> ofsm_table M f k qT'"
using ofsm_table_containment[OF h_obs_state equivalence_relation_on_states_refl[OF \<open>equivalence_relation_on_states M f\<close>]]
by metis
have "(is_in_language M qT ioP) = (is_in_language M qT' ioP)"
using Suc.IH[OF \<open>qT \<in> ofsm_table M f k qT'\<close> \<open>length ioP \<le> k\<close> h_obs_state[OF \<open>h_obs M q' x y = Some qT'\<close>]]
by blast
then have "is_in_language M qT ioP"
using \<open>is_in_language M qT' ioP\<close>
by auto
then show "is_in_language M q ((x,y)#ioP)"
unfolding is_in_language.simps \<open>h_obs M q x y = Some qT\<close> by auto
qed
ultimately show "is_in_language M q ((x, y) # ioP) = is_in_language M q' ((x, y) # ioP) \<and> (is_in_language M q ((x, y) # ioP) \<longrightarrow> after M q' ((x, y) # ioP) \<in> f (after M q ((x, y) # ioP)))"
by blast
qed
qed
qed
then show "is_in_language M q io = is_in_language M q' io" and "(is_in_language M q io \<Longrightarrow> after M q' io \<in> f (after M q io))"
by blast+
qed
lemma after_is_state_is_in_language :
assumes "q \<in> states M"
and "is_in_language M q io"
shows "FSM.after M q io \<in> states M"
using assms
proof (induction io arbitrary: q)
case Nil
then show ?case by auto
next
case (Cons a io)
then obtain x y where "a = (x,y)" using prod.exhaust by metis
show ?case
using \<open>is_in_language M q (a # io)\<close> Cons.IH[OF h_obs_state[of M q x y]]
unfolding \<open>a = (x,y)\<close>
unfolding after.simps is_in_language.simps
by (metis option.case_eq_if option.exhaust_sel)
qed
lemma ofsm_table_elem :
assumes "q \<in> states M"
and "q' \<in> states M"
and "equivalence_relation_on_states M f"
and "\<And> io . length io \<le> k \<Longrightarrow> is_in_language M q io \<longleftrightarrow> is_in_language M q' io"
and "\<And> io . length io \<le> k \<Longrightarrow> is_in_language M q io \<Longrightarrow> (after M q' io) \<in> f (after M q io)"
shows "q' \<in> ofsm_table M f k q"
using assms(1,2,4,5) proof (induction k arbitrary: q q')
case 0
then show ?case by auto
next
case (Suc k)
have "q' \<in> ofsm_table M f k q"
using Suc.IH[OF Suc.prems(1,2)] Suc.prems(3,4) by auto
moreover have "\<And> x y . x \<in> inputs M \<Longrightarrow> y \<in> outputs M \<Longrightarrow> (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
proof -
fix x y assume "x \<in> inputs M" and "y \<in> outputs M"
show "(case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
proof (cases "\<exists> qT qT' . h_obs M q x y = Some qT \<and> h_obs M q' x y = Some qT'")
case True
then obtain qT qT' where "h_obs M q x y = Some qT" and "h_obs M q' x y = Some qT'"
by blast
have *: "\<And> io . length io \<le> k \<Longrightarrow> is_in_language M qT io = is_in_language M qT' io"
proof -
fix io :: "('b \<times> 'c) list "
assume "length io \<le> k"
have "is_in_language M qT io = is_in_language M q ([(x,y)]@io)"
using \<open>h_obs M q x y = Some qT\<close> by auto
moreover have "is_in_language M qT' io = is_in_language M q' ([(x,y)]@io)"
using \<open>h_obs M q' x y = Some qT'\<close> by auto
ultimately show "is_in_language M qT io = is_in_language M qT' io"
using Suc.prems(3) \<open>length io \<le> k\<close>
by (metis append.left_neutral append_Cons length_Cons not_less_eq_eq)
qed
have "ofsm_table M f k qT = ofsm_table M f k qT'"
proof
have "qT \<in> states M"
using h_obs_state[OF \<open>h_obs M q x y = Some qT\<close>] .
have "qT' \<in> states M"
using h_obs_state[OF \<open>h_obs M q' x y = Some qT'\<close>] .
show "ofsm_table M f k qT \<subseteq> ofsm_table M f k qT'"
proof
fix s assume "s \<in> ofsm_table M f k qT"
then have "s \<in> states M"
using ofsm_table_subset[of 0 k M f qT] equivalence_relation_on_states_ran[OF assms(3) \<open>qT \<in> states M\<close>] \<open>qT \<in> states M\<close> by auto
have **: "(\<And>io. length io \<le> k \<Longrightarrow> is_in_language M qT' io = is_in_language M s io)"
using ofsm_table_language(1)[OF \<open>s \<in> ofsm_table M f k qT\<close> _ \<open>qT\<in> states M\<close> assms(3)] * by blast
have ***: "(\<And>io. length io \<le> k \<Longrightarrow> is_in_language M qT' io \<Longrightarrow> after M s io \<in> f (after M qT' io))"
proof -
fix io assume "length io \<le> k" and "is_in_language M qT' io"
then have "is_in_language M qT io"
using * by blast
then have "after M s io \<in> f (after M qT io)"
using ofsm_table_language(2)[OF \<open>s \<in> ofsm_table M f k qT\<close> \<open>length io \<le> k\<close> \<open>qT\<in> states M\<close> assms(3)]
by blast
have "after M qT io = after M q ((x,y)#io)"
using \<open>h_obs M q x y = Some qT\<close> by auto
moreover have "after M qT' io = after M q' ((x,y)#io)"
using \<open>h_obs M q' x y = Some qT'\<close> by auto
moreover have "is_in_language M q ((x,y)#io)"
using \<open>h_obs M q x y = Some qT\<close> \<open>is_in_language M qT io\<close> by auto
ultimately have "after M qT' io \<in> f (after M qT io)"
using Suc.prems(4) \<open>length io \<le> k\<close>
by (metis Suc_le_mono length_Cons)
show "after M s io \<in> f (after M qT' io)"
using equivalence_relation_on_states_trans[OF \<open>equivalence_relation_on_states M f\<close> after_is_state_is_in_language[OF \<open>qT' \<in> states M\<close> \<open>is_in_language M qT' io\<close>]
equivalence_relation_on_states_sym[OF \<open>equivalence_relation_on_states M f\<close> after_is_state_is_in_language[OF \<open>qT \<in> states M\<close> \<open>is_in_language M qT io\<close>]
\<open>after M qT' io \<in> f (after M qT io)\<close>] \<open>after M s io \<in> f (after M qT io)\<close>] .
qed
show "s \<in> ofsm_table M f k qT'"
using Suc.IH[OF \<open>qT' \<in> states M\<close> \<open>s \<in> states M\<close> ** ***] by blast
qed
show "ofsm_table M f k qT' \<subseteq> ofsm_table M f k qT"
proof
fix s assume "s \<in> ofsm_table M f k qT'"
then have "s \<in> states M"
using ofsm_table_subset[of 0 k M f qT'] equivalence_relation_on_states_ran[OF assms(3) \<open>qT' \<in> states M\<close>] \<open>qT' \<in> states M\<close> by auto
have **: "(\<And>io. length io \<le> k \<Longrightarrow> is_in_language M qT io = is_in_language M s io)"
using ofsm_table_language(1)[OF \<open>s \<in> ofsm_table M f k qT'\<close> _ \<open>qT'\<in> states M\<close> assms(3)] * by blast
have ***: "(\<And>io. length io \<le> k \<Longrightarrow> is_in_language M qT io \<Longrightarrow> after M s io \<in> f (after M qT io))"
proof -
fix io assume "length io \<le> k" and "is_in_language M qT io"
then have "is_in_language M qT' io"
using * by blast
then have "after M s io \<in> f (after M qT' io)"
using ofsm_table_language(2)[OF \<open>s \<in> ofsm_table M f k qT'\<close> \<open>length io \<le> k\<close> \<open>qT'\<in> states M\<close> assms(3)]
by blast
have "after M qT' io = after M q' ((x,y)#io)"
using \<open>h_obs M q' x y = Some qT'\<close> by auto
moreover have "after M qT io = after M q ((x,y)#io)"
using \<open>h_obs M q x y = Some qT\<close> by auto
moreover have "is_in_language M q' ((x,y)#io)"
using \<open>h_obs M q' x y = Some qT'\<close> \<open>is_in_language M qT' io\<close> by auto
ultimately have "after M qT io \<in> f (after M qT' io)"
using Suc.prems(4) \<open>length io \<le> k\<close>
by (metis Suc.prems(3) Suc_le_mono \<open>is_in_language M qT io\<close> \<open>qT \<in> FSM.states M\<close> after_is_state_is_in_language assms(3) equivalence_relation_on_states_sym length_Cons)
show "after M s io \<in> f (after M qT io)"
using equivalence_relation_on_states_trans[OF \<open>equivalence_relation_on_states M f\<close> after_is_state_is_in_language[OF \<open>qT \<in> states M\<close> \<open>is_in_language M qT io\<close>]
equivalence_relation_on_states_sym[OF \<open>equivalence_relation_on_states M f\<close> after_is_state_is_in_language[OF \<open>qT' \<in> states M\<close> \<open>is_in_language M qT' io\<close>]
\<open>after M qT io \<in> f (after M qT' io)\<close>] \<open>after M s io \<in> f (after M qT' io)\<close>] .
qed
show "s \<in> ofsm_table M f k qT"
using Suc.IH[OF \<open>qT \<in> states M\<close> \<open>s \<in> states M\<close> ** ***] by blast
qed
qed
then show ?thesis
unfolding \<open>h_obs M q x y = Some qT\<close> \<open>h_obs M q' x y = Some qT'\<close>
by auto
next
case False
have "h_obs M q x y = None \<and> h_obs M q' x y = None"
proof (rule ccontr)
assume "\<not> (h_obs M q x y = None \<and> h_obs M q' x y = None)"
then have "is_in_language M q [(x,y)] \<or> is_in_language M q' [(x,y)]"
unfolding is_in_language.simps
using option.disc_eq_case(2) by blast
moreover have "is_in_language M q [(x,y)] \<noteq> is_in_language M q' [(x,y)]"
using False
by (metis calculation case_optionE is_in_language.simps(2))
moreover have "length [(x,y)] \<le> Suc k"
by auto
ultimately show False
using Suc.prems(3) by blast
qed
then show ?thesis
unfolding ofsm_table_case_helper
by blast
qed
qed
ultimately show ?case
unfolding Suc ofsm_table.simps Let_def by force
qed
lemma ofsm_table_set :
assumes "q \<in> states M"
and "equivalence_relation_on_states M f"
shows "ofsm_table M f k q = {q' . q' \<in> states M \<and> (\<forall> io . length io \<le> k \<longrightarrow> (is_in_language M q io \<longleftrightarrow> is_in_language M q' io) \<and> (is_in_language M q io \<longrightarrow> after M q' io \<in> f (after M q io)))}"
using ofsm_table_language[OF _ _ assms(1,2) ]
ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(2)] assms(1)]
ofsm_table_elem[OF assms(1) _ assms(2)]
by blast
lemma ofsm_table_set_observable :
assumes "observable M" and "q \<in> states M"
and "equivalence_relation_on_states M f"
shows "ofsm_table M f k q = {q' . q' \<in> states M \<and> (\<forall> io . length io \<le> k \<longrightarrow> (io \<in> LS M q \<longleftrightarrow> io \<in> LS M q') \<and> (io \<in> LS M q \<longrightarrow> after M q' io \<in> f (after M q io)))}"
unfolding ofsm_table_set[OF assms(2,3)]
unfolding is_in_language_iff[OF assms(1,2)]
using is_in_language_iff[OF assms(1)]
by blast
lemma ofsm_table_eq_if_elem :
assumes "q1 \<in> states M" and "q2 \<in> states M" and "equivalence_relation_on_states M f"
shows "(ofsm_table M f k q1 = ofsm_table M f k q2) = (q2 \<in> ofsm_table M f k q1)"
proof
show "ofsm_table M f k q1 = ofsm_table M f k q2 \<Longrightarrow> q2 \<in> ofsm_table M f k q1"
using ofsm_table_containment[OF assms(2) equivalence_relation_on_states_refl[OF \<open>equivalence_relation_on_states M f\<close>]]
by blast
show "q2 \<in> ofsm_table M f k q1 \<Longrightarrow> ofsm_table M f k q1 = ofsm_table M f k q2"
proof -
assume *: "q2 \<in> ofsm_table M f k q1"
have "ofsm_table M f k q1 = {q' \<in> FSM.states M. \<forall>io. length io \<le> k \<longrightarrow> (is_in_language M q1 io) = (is_in_language M q' io) \<and> (is_in_language M q1 io \<longrightarrow> after M q' io \<in> f (after M q1 io))}"
using ofsm_table_set[OF assms(1,3)] by auto
moreover have "ofsm_table M f k q2 = {q' \<in> FSM.states M. \<forall>io. length io \<le> k \<longrightarrow> (is_in_language M q1 io) = (is_in_language M q' io) \<and> (is_in_language M q1 io \<longrightarrow> after M q' io \<in> f (after M q1 io))}"
proof -
have "ofsm_table M f k q2 = {q' \<in> FSM.states M. \<forall>io. length io \<le> k \<longrightarrow> (is_in_language M q2 io) = (is_in_language M q' io) \<and> (is_in_language M q2 io \<longrightarrow> after M q' io \<in> f (after M q2 io))}"
using ofsm_table_set[OF assms(2,3)] by auto
moreover have "\<And> io . length io \<le> k \<Longrightarrow> (is_in_language M q1 io) = (is_in_language M q2 io)"
using ofsm_table_language(1)[OF * _ assms(1,3)] by blast
moreover have "\<And> io q' . q' \<in> states M \<Longrightarrow> length io \<le> k \<Longrightarrow> (is_in_language M q2 io \<longrightarrow> after M q' io \<in> f (after M q2 io)) = (is_in_language M q1 io \<longrightarrow> after M q' io \<in> f (after M q1 io))"
using ofsm_table_language(2)[OF * _ assms(1,3)]
by (meson after_is_state_is_in_language assms(1) assms(2) assms(3) calculation(2) equivalence_relation_on_states_sym equivalence_relation_on_states_trans)
ultimately show ?thesis
by blast
qed
ultimately show ?thesis
by blast
qed
qed
lemma ofsm_table_fix_language :
fixes M :: "('a,'b,'c) fsm"
assumes "q' \<in> ofsm_table_fix M f 0 q"
and "q \<in> states M"
and "observable M"
and "equivalence_relation_on_states M f"
shows "LS M q = LS M q'"
and "io \<in> LS M q \<Longrightarrow> after M q' io \<in> f (after M q io)"
proof -
obtain k where *:"\<And> q . q \<in> states M \<Longrightarrow> ofsm_table_fix M f 0 q = ofsm_table M f k q"
and **: "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> k \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f k q"
using ofsm_table_fix_length[of M f,OF equivalence_relation_on_states_ran[OF assms(4)]]
by blast
have "q' \<in> ofsm_table M f k q"
using * assms(1,2) by blast
then have "q' \<in> states M"
by (metis assms(2) assms(4) equivalence_relation_on_states_ran le0 ofsm_table.simps(1) ofsm_table_subset subset_iff)
have "\<And> k' . q' \<in> ofsm_table M f k' q"
proof -
fix k' show "q' \<in> ofsm_table M f k' q"
proof (cases "k' \<le> k")
case True
show ?thesis using ofsm_table_subset[OF True, of M f q] \<open>q' \<in> ofsm_table M f k q\<close> by blast
next
case False
then have "k \<le> k'"
by auto
show ?thesis
unfolding **[OF assms(2) \<open>k \<le> k'\<close>]
using \<open>q' \<in> ofsm_table M f k q\<close> by assumption
qed
qed
have "\<And> io . io \<in> LS M q \<longleftrightarrow> io \<in> LS M q'"
proof -
fix io :: "('b \<times> 'c) list"
show "io \<in> LS M q \<longleftrightarrow> io \<in> LS M q'"
using ofsm_table_language(1)[OF \<open>q' \<in> ofsm_table M f (length io) q\<close> _ assms(2,4), of io]
using is_in_language_iff[OF assms(3,2)] is_in_language_iff[OF assms(3) \<open>q' \<in> states M\<close>]
by blast
qed
then show "LS M q = LS M q'"
by blast
show "io \<in> LS M q \<Longrightarrow> after M q' io \<in> f (after M q io)"
using ofsm_table_language(2)[OF \<open>q' \<in> ofsm_table M f (length io) q\<close> _ assms(2,4), of io]
using is_in_language_iff[OF assms(3,2)] is_in_language_iff[OF assms(3) \<open>q' \<in> states M\<close>]
by blast
qed
lemma ofsm_table_same_language :
assumes "LS M q = LS M q'"
and "\<And> io . io \<in> LS M q \<Longrightarrow> after M q' io \<in> f (after M q io)"
and "observable M"
and "q' \<in> states M"
and "q \<in> states M"
and "equivalence_relation_on_states M f"
shows "ofsm_table M f k q = ofsm_table M f k q'"
using assms(1,2,4,5)
proof (induction k arbitrary: q q')
case 0
then show ?case
by (metis after.simps(1) assms(6) from_FSM_language language_contains_empty_sequence ofsm_table.simps(1) ofsm_table_eq_if_elem)
next
case (Suc k)
have "ofsm_table M f (Suc k) q = {q'' \<in> ofsm_table M f k q' . \<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None) }"
using Suc.IH[OF Suc.prems] unfolding ofsm_table.simps Suc Let_def Suc by simp
moreover have "ofsm_table M f (Suc k) q' = {q'' \<in> ofsm_table M f k q' . \<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q' x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None) }"
unfolding ofsm_table.simps Suc Let_def
by auto
moreover have "{q'' \<in> ofsm_table M f k q' . \<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None) }
= {q'' \<in> ofsm_table M f k q' . \<forall> x \<in> inputs M . \<forall> y \<in> outputs M . (case h_obs M q' x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None) }"
proof -
have "\<And> q'' x y . q'' \<in> ofsm_table M f k q' \<Longrightarrow> x \<in> inputs M \<Longrightarrow> y \<in> outputs M \<Longrightarrow>
(case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None)
= (case h_obs M q' x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None)"
proof -
fix q'' x y assume "q'' \<in> ofsm_table M f k q'" and "x \<in> inputs M" and "y \<in> outputs M"
have *:"(\<exists> qT . h_obs M q x y = Some qT) = (\<exists> qT' . h_obs M q' x y = Some qT')"
proof -
have "([(x,y)] \<in> LS M q) = ([(x,y)] \<in> LS M q')"
using \<open>LS M q = LS M q'\<close> by auto
then have "(\<exists> qT . (q, x, y, qT) \<in> FSM.transitions M) = (\<exists> qT' . (q', x, y, qT') \<in> FSM.transitions M)"
unfolding LS_single_transition by force
then show "(\<exists> qT . h_obs M q x y = Some qT) = (\<exists> qT' . h_obs M q' x y = Some qT')"
unfolding h_obs_Some[OF \<open>observable M\<close>] using \<open>observable M\<close> unfolding observable_alt_def by blast
qed
have **: "(case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q' x y = None)"
proof (cases "h_obs M q x y")
case None
then show ?thesis using * by auto
next
case (Some qT)
show ?thesis proof (cases "h_obs M q' x y")
case None
then show ?thesis using * by auto
next
case (Some qT')
have "(q,x,y,qT) \<in> transitions M"
using \<open>h_obs M q x y = Some qT\<close> unfolding h_obs_Some[OF \<open>observable M\<close>] by blast
have "(q',x,y,qT') \<in> transitions M"
using \<open>h_obs M q' x y = Some qT'\<close> unfolding h_obs_Some[OF \<open>observable M\<close>] by blast
have "LS M qT = LS M qT'"
using observable_transition_target_language_eq[OF _ \<open>(q,x,y,qT) \<in> transitions M\<close> \<open>(q',x,y,qT') \<in> transitions M\<close> _ _ \<open>observable M\<close>]
\<open>LS M q = LS M q'\<close>
by auto
moreover have "(\<And>io. io \<in> LS M qT \<Longrightarrow> after M qT' io \<in> f (after M qT io))"
proof -
fix io assume "io \<in> LS M qT"
have "io \<in> LS M qT'"
using \<open>io \<in> LS M qT\<close> calculation by auto
have "after M qT io = after M q ((x,y)#io)"
using after_h_obs_prepend[OF \<open>observable M\<close> \<open>h_obs M q x y = Some qT\<close> \<open>io \<in> LS M qT\<close>]
by simp
moreover have "after M qT' io = after M q' ((x,y)#io)"
using after_h_obs_prepend[OF \<open>observable M\<close> \<open>h_obs M q' x y = Some qT'\<close> \<open>io \<in> LS M qT'\<close>]
by simp
moreover have "(x,y)#io \<in> LS M q"
using \<open>h_obs M q x y = Some qT\<close> \<open>io \<in> LS M qT\<close> unfolding h_obs_language_iff[OF \<open>observable M\<close>]
by blast
ultimately show "after M qT' io \<in> f (after M qT io)"
using Suc.prems(2) by presburger
qed
ultimately have "ofsm_table M f k qT = ofsm_table M f k qT'"
using Suc.IH[OF _ _ fsm_transition_target[OF \<open>(q',x,y,qT') \<in> transitions M\<close>] fsm_transition_target[OF \<open>(q,x,y,qT) \<in> transitions M\<close>]]
unfolding snd_conv
by blast
then show ?thesis
using \<open>h_obs M q x y = Some qT\<close> \<open>h_obs M q' x y = Some qT'\<close> by auto
qed
qed
show "(case h_obs M q x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None)
= (case h_obs M q' x y of Some qT \<Rightarrow> (case h_obs M q'' x y of Some qT' \<Rightarrow> ofsm_table M f k qT = ofsm_table M f k qT' | None \<Rightarrow> False) | None \<Rightarrow> h_obs M q'' x y = None)" (is "?P")
proof (cases "h_obs M q x y")
case None
then have "h_obs M q' x y = None"
using * by auto
show ?thesis unfolding None \<open>h_obs M q' x y = None\<close> by auto
next
case (Some qT)
then obtain qT' where "h_obs M q' x y = Some qT'"
using \<open>(\<exists> qT . h_obs M q x y = Some qT) = (\<exists> qT' . h_obs M q' x y = Some qT')\<close> by auto
show ?thesis
proof (cases "h_obs M q'' x y")
case None
then show ?thesis using *
by (metis Some option.case_eq_if option.simps(5))
next
case (Some qT'')
show ?thesis
using **
unfolding Some \<open>h_obs M q x y = Some qT\<close> \<open>h_obs M q' x y = Some qT'\<close> by auto
qed
qed
qed
then show ?thesis
by blast
qed
ultimately show ?case by blast
qed
lemma ofsm_table_fix_set :
assumes "q \<in> states M"
and "observable M"
and "equivalence_relation_on_states M f"
shows "ofsm_table_fix M f 0 q = {q' \<in> states M . LS M q' = LS M q \<and> (\<forall> io \<in> LS M q . after M q' io \<in> f (after M q io))}"
proof
have "ofsm_table_fix M f 0 q \<subseteq> ofsm_table M f 0 q"
using ofsm_table_fix_length[of M f]
ofsm_table_subset[OF zero_le, of M f _ q]
by (metis assms(1) assms(3) equivalence_relation_on_states_ran)
then have "ofsm_table_fix M f 0 q \<subseteq> states M"
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(3)] assms(1)] by blast
then show "ofsm_table_fix M f 0 q \<subseteq> {q' \<in> states M . LS M q' = LS M q \<and> (\<forall> io \<in> LS M q . after M q' io \<in> f (after M q io))}"
using ofsm_table_fix_language[OF _ assms] by blast
show "{q' \<in> states M . LS M q' = LS M q \<and> (\<forall> io \<in> LS M q . after M q' io \<in> f (after M q io))} \<subseteq> ofsm_table_fix M f 0 q"
proof
fix q' assume "q' \<in> {q' \<in> states M . LS M q' = LS M q \<and> (\<forall> io \<in> LS M q . after M q' io \<in> f (after M q io))}"
then have "q' \<in> states M" and "LS M q' = LS M q" and "\<And> io . io \<in> LS M q \<Longrightarrow> after M q' io \<in> f (after M q io)"
by blast+
then have "\<And> io . io \<in> LS M q' \<Longrightarrow> after M q io \<in> f (after M q' io)"
by (metis after_is_state assms(2) assms(3) equivalence_relation_on_states_sym)
obtain k where "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table_fix M f 0 q = ofsm_table M f k q"
and "\<And> q k' . q \<in> states M \<Longrightarrow> k' \<ge> k \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f k q"
using ofsm_table_fix_length[of M f, OF equivalence_relation_on_states_ran[OF assms(3)]] by blast
have "ofsm_table M f k q' = ofsm_table M f k q"
using ofsm_table_same_language[OF \<open>LS M q' = LS M q\<close> \<open>\<And> io . io \<in> LS M q' \<Longrightarrow> after M q io \<in> f (after M q' io)\<close> assms(2,1) \<open>q' \<in> states M\<close> assms(3)]
by blast
then show "q' \<in> ofsm_table_fix M f 0 q"
using ofsm_table_containment[OF \<open>q' \<in> states M\<close>, of f k]
using \<open>\<And> q . q \<in> states M \<Longrightarrow> ofsm_table_fix M f 0 q = ofsm_table M f k q\<close>
by (metis assms(1) assms(3) equivalence_relation_on_states_refl)
qed
qed
lemma ofsm_table_fix_eq_if_elem :
assumes "q1 \<in> states M" and "q2 \<in> states M"
and "equivalence_relation_on_states M f"
shows "(ofsm_table_fix M f 0 q1 = ofsm_table_fix M f 0 q2) = (q2 \<in> ofsm_table_fix M f 0 q1)"
proof
have "(\<And>q. q \<in> FSM.states M \<Longrightarrow> q \<in> f q)"
using assms(3)
by (meson equivalence_relation_on_states_refl)
show "ofsm_table_fix M f 0 q1 = ofsm_table_fix M f 0 q2 \<Longrightarrow> q2 \<in> ofsm_table_fix M f 0 q1"
using ofsm_table_containment[of _ M f, OF assms(2) \<open>(\<And>q. q \<in> FSM.states M \<Longrightarrow> q \<in> f q)\<close>]
using ofsm_table_fix_length[of M f]
by (metis assms(2) assms(3) equivalence_relation_on_states_ran)
show "q2 \<in> ofsm_table_fix M f 0 q1 \<Longrightarrow> ofsm_table_fix M f 0 q1 = ofsm_table_fix M f 0 q2"
using ofsm_table_eq_if_elem[OF assms(1,2,3)]
using ofsm_table_fix_length[of M f]
by (metis assms(1) assms(2) assms(3) equivalence_relation_on_states_ran)
qed
lemma ofsm_table_refinement_disjoint :
assumes "q1 \<in> states M" and "q2 \<in> states M"
and "equivalence_relation_on_states M f"
and "ofsm_table M f k q1 \<noteq> ofsm_table M f k q2"
shows "ofsm_table M f (Suc k) q1 \<noteq> ofsm_table M f (Suc k) q2"
proof -
have "ofsm_table M f (Suc k) q1 \<subseteq> ofsm_table M f k q1"
and "ofsm_table M f (Suc k) q2 \<subseteq> ofsm_table M f k q2"
using ofsm_table_subset[of k "Suc k" M f]
by fastforce+
moreover have "ofsm_table M f k q1 \<inter> ofsm_table M f k q2 = {}"
proof (rule ccontr)
assume "ofsm_table M f k q1 \<inter> ofsm_table M f k q2 \<noteq> {}"
then obtain q where "q \<in> ofsm_table M f k q1"
and "q \<in> ofsm_table M f k q2"
by blast
then have "q \<in> states M"
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(3)] assms(1)]
by blast
have "ofsm_table M f k q1 = ofsm_table M f k q2"
using \<open>q \<in> ofsm_table M f k q1\<close> \<open>q \<in> ofsm_table M f k q2\<close>
unfolding ofsm_table_eq_if_elem[OF assms(1) \<open>q \<in> states M\<close> assms(3), symmetric]
unfolding ofsm_table_eq_if_elem[OF assms(2) \<open>q \<in> states M\<close> assms(3), symmetric]
by blast
then show False
using assms(4) by simp
qed
ultimately show ?thesis
by (metis Int_subset_iff all_not_in_conv assms(2) assms(3) ofsm_table_eq_if_elem subset_empty)
qed
lemma ofsm_table_partition_finite :
assumes "equivalence_relation_on_states M f"
shows "finite (ofsm_table M f k ` states M)"
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms]]
fsm_states_finite[of M]
unfolding finite_Pow_iff[of "states M", symmetric]
by simp
lemma ofsm_table_refinement_card :
assumes "equivalence_relation_on_states M f"
and "A \<subseteq> states M"
and "i \<le> j"
shows "card (ofsm_table M f j ` A) \<ge> card (ofsm_table M f i ` A)"
proof -
have "\<And> k . card (ofsm_table M f (Suc k) ` A) \<ge> card (ofsm_table M f k ` A)"
proof -
fix k show "card (ofsm_table M f (Suc k) ` A) \<ge> card (ofsm_table M f k ` A)"
proof (rule ccontr)
have "finite A"
using fsm_states_finite[of M] assms(2)
using finite_subset by blast
assume "\<not> card (ofsm_table M f k ` A) \<le> card (ofsm_table M f (Suc k) ` A)"
then have "card (ofsm_table M f (Suc k) ` A) < card (ofsm_table M f k ` A)"
by simp
then obtain q1 q2 where "q1 \<in> A"
and "q2 \<in> A"
and "ofsm_table M f k q1 \<noteq> ofsm_table M f k q2"
and "ofsm_table M f (Suc k) q1 = ofsm_table M f (Suc k) q2"
using finite_card_less_witnesses[OF \<open>finite A\<close>] by blast
then show False
using ofsm_table_refinement_disjoint[OF _ _ assms(1), of q1 q2 k]
using assms(2)
by blast
qed
qed
then show ?thesis
using lift_Suc_mono_le[OF _ assms(3), where f="\<lambda> k . card (ofsm_table M f k ` A)"]
by blast
qed
lemma ofsm_table_refinement_card_fix_Suc :
assumes "equivalence_relation_on_states M f"
and "card (ofsm_table M f (Suc k) ` states M) = card (ofsm_table M f k ` states M)"
and "q \<in> states M"
shows "ofsm_table M f (Suc k) q = ofsm_table M f k q"
proof (rule ccontr)
assume "ofsm_table M f (Suc k) q \<noteq> ofsm_table M f k q"
then have "ofsm_table M f (Suc k) q \<subset> ofsm_table M f k q"
using ofsm_table_subset
by (metis Suc_leD order_refl psubsetI)
then obtain q' where "q' \<in> ofsm_table M f k q"
and "q' \<notin> ofsm_table M f (Suc k) q"
by blast
then have "q' \<in> states M"
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(1)] assms(3)] by blast
have card_qq: "\<And> k . card (ofsm_table M f k ` states M)
= card (ofsm_table M f k ` (states M - \<Union>(ofsm_table M f k ` {q,q'}))) + card (ofsm_table M f k ` (\<Union>(ofsm_table M f k ` {q,q'})))"
proof -
fix k
have "states M = (states M - \<Union>(ofsm_table M f k ` {q,q'})) \<union> \<Union>(ofsm_table M f k ` {q,q'})"
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(1)] \<open>q \<in> states M\<close>]
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(1)] \<open>q' \<in> states M\<close>]
by blast
then have "finite (states M - \<Union>(ofsm_table M f k ` {q,q'}))"
and "finite (\<Union>(ofsm_table M f k ` {q,q'}))"
using fsm_states_finite[of M] finite_Un[of "(states M - \<Union>(ofsm_table M f k ` {q,q'}))" "\<Union>(ofsm_table M f k ` {q,q'})"]
by force+
then have *:"finite (ofsm_table M f k ` (states M - \<Union>(ofsm_table M f k ` {q,q'})))"
and **:"finite (ofsm_table M f k ` \<Union>(ofsm_table M f k ` {q,q'}))"
by blast+
have ***:"(ofsm_table M f k ` (states M - \<Union>(ofsm_table M f k ` {q,q'}))) \<inter> (ofsm_table M f k ` \<Union>(ofsm_table M f k ` {q,q'})) = {}"
proof (rule ccontr)
assume "ofsm_table M f k ` (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'})) \<inter> ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'}) \<noteq> {}"
then obtain Q where "Q \<in> ofsm_table M f k ` (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'}))"
and "Q \<in> ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'})"
by blast
obtain q1 where "q1 \<in> (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'}))"
and "Q = ofsm_table M f k q1"
using \<open>Q \<in> ofsm_table M f k ` (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'}))\<close> by blast
moreover obtain q2 where "q2 \<in> \<Union> (ofsm_table M f k ` {q, q'})"
and "Q = ofsm_table M f k q2"
using \<open>Q \<in> ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'})\<close> by blast
ultimately have "ofsm_table M f k q1 = ofsm_table M f k q2"
by auto
have "q1 \<in> states M" and "q1 \<notin> \<Union> (ofsm_table M f k ` {q, q'})"
using \<open>q1 \<in> (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'}))\<close>
by blast+
have "q2 \<in> states M"
using \<open>q2 \<in> \<Union> (ofsm_table M f k ` {q, q'})\<close> \<open>states M = (states M - \<Union>(ofsm_table M f k ` {q,q'})) \<union> \<Union>(ofsm_table M f k ` {q,q'})\<close>
by blast
have "q1 \<in> ofsm_table M f k q2"
using \<open>ofsm_table M f k q1 = ofsm_table M f k q2\<close>
using ofsm_table_eq_if_elem[OF \<open>q2 \<in> states M\<close> \<open>q1 \<in> states M\<close> assms(1)]
by blast
moreover have "q2 \<in> ofsm_table M f k q \<or> q2 \<in> ofsm_table M f k q'"
using \<open>q2 \<in> \<Union> (ofsm_table M f k ` {q, q'})\<close>
by blast
ultimately have "q1 \<in> \<Union> (ofsm_table M f k ` {q, q'})"
unfolding ofsm_table_eq_if_elem[OF \<open>q \<in> states M\<close> \<open>q2 \<in> states M\<close> assms(1), symmetric]
unfolding ofsm_table_eq_if_elem[OF \<open>q' \<in> states M\<close> \<open>q2 \<in> states M\<close> assms(1), symmetric]
by blast
then show False
using \<open>q1 \<notin> \<Union> (ofsm_table M f k ` {q, q'})\<close>
by blast
qed
show "card (ofsm_table M f k ` states M)
= card (ofsm_table M f k ` (states M - \<Union>(ofsm_table M f k ` {q,q'}))) + card (ofsm_table M f k ` (\<Union>(ofsm_table M f k ` {q,q'})))"
using card_Un_disjoint[OF * ** ***]
using \<open>states M = (states M - \<Union>(ofsm_table M f k ` {q,q'})) \<union> \<Union>(ofsm_table M f k ` {q,q'})\<close>
by (metis image_Un)
qed
have s1: "\<And> k . (states M - \<Union>(ofsm_table M f k ` {q,q'})) \<subseteq> states M"
and s2: "\<And> k . (\<Union>(ofsm_table M f k ` {q,q'})) \<subseteq> states M"
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(1)] \<open>q \<in> states M\<close>]
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(1)] \<open>q' \<in> states M\<close>]
by blast+
have "card (ofsm_table M f (Suc k) ` states M) > card (ofsm_table M f k ` states M)"
proof -
have *: "\<Union> (ofsm_table M f (Suc k) ` {q, q'}) \<subseteq> \<Union> (ofsm_table M f k ` {q, q'})"
using ofsm_table_subset
by (metis SUP_mono' lessI less_imp_le_nat)
have "card (ofsm_table M f k ` (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'}))) \<le> card (ofsm_table M f (Suc k) ` (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'})))"
using ofsm_table_refinement_card[OF assms(1), where i=k and j="Suc k", OF s1]
using le_SucI by blast
moreover have "card (ofsm_table M f (Suc k) ` (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'}))) \<le> card (ofsm_table M f (Suc k) ` (FSM.states M - \<Union> (ofsm_table M f (Suc k) ` {q, q'})))"
using *
using fsm_states_finite[of M]
by (meson Diff_mono card_mono finite_Diff finite_imageI image_mono subset_refl)
ultimately have "card (ofsm_table M f k ` (FSM.states M - \<Union> (ofsm_table M f k ` {q, q'}))) \<le> card (ofsm_table M f (Suc k) ` (FSM.states M - \<Union> (ofsm_table M f (Suc k) ` {q, q'})))"
by presburger
moreover have "card (ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'})) < card (ofsm_table M f (Suc k) ` \<Union> (ofsm_table M f (Suc k) ` {q, q'}))"
proof -
have *: "\<And> k . ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'}) = {ofsm_table M f k q, ofsm_table M f k q'}"
proof -
fix k show "ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'}) = {ofsm_table M f k q, ofsm_table M f k q'}"
proof
show "ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'}) \<subseteq> {ofsm_table M f k q, ofsm_table M f k q'}"
proof
fix Q assume "Q \<in> ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'})"
then obtain qq where "Q = ofsm_table M f k qq"
and "qq \<in> \<Union> (ofsm_table M f k ` {q, q'})"
by blast
then have "qq \<in> ofsm_table M f k q \<or> qq \<in> ofsm_table M f k q'"
by blast
then have "qq \<in> states M"
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(1)]] \<open>q \<in> states M\<close> \<open>q' \<in> states M\<close>
by blast
have "ofsm_table M f k qq = ofsm_table M f k q \<or> ofsm_table M f k qq = ofsm_table M f k q'"
using \<open>qq \<in> ofsm_table M f k q \<or> qq \<in> ofsm_table M f k q'\<close>
using ofsm_table_eq_if_elem[OF _ \<open>qq \<in> states M\<close> assms(1)] \<open>q \<in> states M\<close> \<open>q' \<in> states M\<close>
by blast
then show "Q \<in> {ofsm_table M f k q, ofsm_table M f k q'}"
using \<open>Q = ofsm_table M f k qq\<close>
by blast
qed
show "{ofsm_table M f k q, ofsm_table M f k q'} \<subseteq> ofsm_table M f k ` \<Union> (ofsm_table M f k ` {q, q'})"
using ofsm_table_containment[of _ M f, OF _ equivalence_relation_on_states_refl[OF assms(1)]] \<open>q \<in> states M\<close> \<open>q' \<in> states M\<close>
by blast
qed
qed
have "ofsm_table M f k q = ofsm_table M f k q'"
using \<open>q' \<in> ofsm_table M f k q\<close>
using ofsm_table_eq_if_elem[OF \<open>q \<in> states M\<close> \<open>q' \<in> states M\<close> assms(1)]
by blast
moreover have "ofsm_table M f (Suc k) q \<noteq> ofsm_table M f (Suc k) q'"
using \<open>q' \<notin> ofsm_table M f (Suc k) q\<close>
using ofsm_table_eq_if_elem[OF \<open>q \<in> states M\<close> \<open>q' \<in> states M\<close> assms(1)]
by blast
ultimately show ?thesis
unfolding *
by (metis card_insert_if finite.emptyI finite.insertI insert_absorb insert_absorb2 insert_not_empty lessI singleton_insert_inj_eq)
qed
ultimately show ?thesis
unfolding card_qq by presburger
qed
then show False
using assms(2) by linarith
qed
lemma ofsm_table_refinement_card_fix :
assumes "equivalence_relation_on_states M f"
and "card (ofsm_table M f j ` states M) = card (ofsm_table M f i ` states M)"
and "q \<in> states M"
and "i \<le> j"
shows "ofsm_table M f j q = ofsm_table M f i q"
using assms (2,4) proof (induction "j-i" arbitrary: i j)
case 0
then have "i = j" by auto
then show ?case by auto
next
case (Suc k)
then have "j \<ge> Suc i" and "k = j - Suc i"
by auto
have *:"card (ofsm_table M f j ` FSM.states M) = card (ofsm_table M f (Suc i) ` FSM.states M)"
and **:"card (ofsm_table M f (Suc i) ` FSM.states M) = card (ofsm_table M f i ` FSM.states M)"
using ofsm_table_refinement_card[OF assms(1), where A="states M"]
by (metis Suc.prems(1) \<open>Suc i \<le> j\<close> eq_iff le_SucI)+
show ?case
using Suc.hyps(1)[OF \<open>k = j - Suc i\<close> * \<open>Suc i \<le> j\<close>]
using ofsm_table_refinement_card_fix_Suc[OF assms(1) ** assms(3)]
by blast
qed
lemma ofsm_table_partition_fixpoint_Suc :
assumes "equivalence_relation_on_states M f"
and "q \<in> states M"
shows "ofsm_table M f (size M - card (f ` states M)) q = ofsm_table M f (Suc (size M - card (f ` states M))) q"
proof -
have "\<And> q . q \<in> states M \<Longrightarrow> f q = ofsm_table M f 0 q"
unfolding ofsm_table.simps by auto
define n where n: "n = (\<lambda> i . card (ofsm_table M f i ` states M))"
have "\<And> i j . i \<le> j \<Longrightarrow> n i \<le> n j"
unfolding n
using ofsm_table_refinement_card[OF assms(1), where A="states M"]
by blast
moreover have "\<And> i j m . i < j \<Longrightarrow> n i = n j \<Longrightarrow> j \<le> m \<Longrightarrow> n i = n m"
proof -
fix i j m assume "i < j" and "n i = n j" and "j \<le> m"
then have "Suc i \<le> j" and "i \<le> Suc i" and "i \<le> m"
by auto
have "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f j q = ofsm_table M f i q"
using \<open>i < j\<close> \<open>n i = n j\<close> ofsm_table_refinement_card_fix[OF assms(1) _]
unfolding n
using less_imp_le_nat by presburger
then have "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f (Suc i) q = ofsm_table M f i q"
using ofsm_table_subset[OF \<open>Suc i \<le> j\<close>, of M f]
using ofsm_table_subset[OF \<open>i \<le> Suc i\<close>, of M f]
by blast
then have "\<And> q . q \<in> states M \<Longrightarrow> ofsm_table M f m q = ofsm_table M f i q"
using ofsm_table_fixpoint[OF \<open>i \<le> m\<close>]
by metis
then show "n i = n m"
unfolding n
by auto
qed
moreover have "\<And> i . n i \<le> size M"
unfolding n
using ofsm_table_states[of M f, OF equivalence_relation_on_states_ran[OF assms(1)]]
using fsm_states_finite[of M]
by (simp add: card_image_le)
ultimately obtain k where "n (Suc k) = n k"
and "k \<le> size M - n 0"
using monotone_function_with_limit_witness_helper[where f=n and k="size M"]
by blast
then show ?thesis
unfolding n
using \<open>\<And> q . q \<in> states M \<Longrightarrow> f q = ofsm_table M f 0 q\<close>[symmetric]
using ofsm_table_refinement_card_fix_Suc[OF assms(1) _ ]
using ofsm_table_fixpoint[OF _ _ assms(2)]
by (metis (mono_tags, lifting) image_cong nat_le_linear not_less_eq_eq)
qed
lemma ofsm_table_partition_fixpoint :
assumes "equivalence_relation_on_states M f"
and "size M \<le> m"
and "q \<in> states M"
shows "ofsm_table M f (m - card (f ` states M)) q = ofsm_table M f (Suc (m - card (f ` states M))) q"
proof -
have *: "size M - card (f ` states M) \<le> m - card (f ` states M)"
using assms(2) by simp
have **: "(size M - card (f ` states M)) \<le> Suc (m - card (f ` states M))"
using assms(2) by simp
have ***: "\<And> q . q \<in> FSM.states M \<Longrightarrow> ofsm_table M f (FSM.size M - card (f ` FSM.states M)) q = ofsm_table M f (Suc (FSM.size M - card (f ` FSM.states M))) q"
using ofsm_table_partition_fixpoint_Suc[OF assms(1)] .
have "ofsm_table M f (m - card (f ` states M)) q = ofsm_table M f (FSM.size M - card (f ` FSM.states M)) q"
using ofsm_table_fixpoint[OF * _ assms(3)] ***
by blast
moreover have "ofsm_table M f (Suc (m - card (f ` states M))) q = ofsm_table M f (FSM.size M - card (f ` FSM.states M)) q"
using ofsm_table_fixpoint[OF ** _ assms(3), of f] ***
by blast
ultimately show ?thesis
by simp
qed
lemma ofsm_table_fix_partition_fixpoint :
assumes "equivalence_relation_on_states M f"
and "size M \<le> m"
and "q \<in> states M"
shows "ofsm_table M f (m - card (f ` states M)) q = ofsm_table_fix M f 0 q"
proof -
obtain k where k1: "ofsm_table_fix M f 0 q = ofsm_table M f k q"
and k2: "\<And> k' . k' \<ge> k \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f k q"
using ofsm_table_fix_length[of M f, OF equivalence_relation_on_states_ran[OF assms(1)]]
assms(3)
by metis
have m1: "\<And> k' . k' \<ge> m - card (f ` states M) \<Longrightarrow> ofsm_table M f k' q = ofsm_table M f (m - card (f ` states M)) q"
using ofsm_table_partition_fixpoint[OF assms(1,2)]
using ofsm_table_fixpoint[OF _ _ assms(3)]
by presburger
show ?thesis proof (cases "k \<le> m - card (f ` states M)")
case True
show ?thesis
using k1 k2[OF True] by simp
next
case False
then have "k \<ge> m - card (f ` states M)"
by auto
then have "ofsm_table M f k q = ofsm_table M f (m - card (f ` states M)) q"
using ofsm_table_partition_fixpoint[OF assms(1,2)]
using ofsm_table_fixpoint[OF _ _ assms(3)]
by presburger
then show ?thesis
using k1 by simp
qed
qed
subsection \<open>A minimisation function based on OFSM-tables\<close>
lemma language_equivalence_classes_preserve_observability:
assumes "transitions M' = (\<lambda> t . ({q \<in> states M . LS M q = LS M (t_source t)} , t_input t, t_output t, {q \<in> states M . LS M q = LS M (t_target t)})) ` transitions M"
and "observable M"
shows "observable M'"
proof -
have "\<And> t1 t2 . t1 \<in> transitions M' \<Longrightarrow>
t2 \<in> transitions M' \<Longrightarrow>
t_source t1 = t_source t2 \<Longrightarrow>
t_input t1 = t_input t2 \<Longrightarrow>
t_output t1 = t_output t2 \<Longrightarrow>
t_target t1 = t_target t2"
proof -
fix t1 t2 assume "t1 \<in> transitions M'" and "t2 \<in> transitions M'" and "t_source t1 = t_source t2" and "t_input t1 = t_input t2" and "t_output t1 = t_output t2"
obtain t1' where t1'_def: "t1 = ({q \<in> states M . LS M q = LS M (t_source t1')} , t_input t1', t_output t1', {q \<in> states M . LS M q = LS M (t_target t1')})"
and "t1' \<in> transitions M"
using \<open>t1 \<in> transitions M'\<close> assms(1) by auto
obtain t2' where t2'_def: "t2 = ({q \<in> states M . LS M q = LS M (t_source t2')} , t_input t2', t_output t2', {q \<in> states M . LS M q = LS M (t_target t2')})"
and "t2' \<in> transitions M"
using \<open>t2 \<in> transitions M'\<close> assms(1) \<open>t_input t1 = t_input t2\<close> \<open>t_output t1 = t_output t2\<close> by auto
have "{q \<in> FSM.states M. LS M q = LS M (t_source t1')} = {q \<in> FSM.states M. LS M q = LS M (t_source t2')}"
using t1'_def t2'_def \<open>t_source t1 = t_source t2\<close>
by (metis (no_types, lifting) fst_eqD)
then have "LS M (t_source t1') = LS M (t_source t2')"
using fsm_transition_source[OF \<open>t1' \<in> transitions M\<close>] fsm_transition_source[OF \<open>t2' \<in> transitions M\<close>] by blast
then have "LS M (t_target t1') = LS M (t_target t2')"
using observable_transition_target_language_eq[OF _ \<open>t1' \<in> transitions M\<close> \<open>t2' \<in> transitions M\<close> _ _ \<open>observable M\<close>]
using \<open>t_input t1 = t_input t2\<close> \<open>t_output t1 = t_output t2\<close>
unfolding t1'_def t2'_def fst_conv snd_conv by blast
then show "t_target t1 = t_target t2"
unfolding t1'_def t2'_def snd_conv by blast
qed
then show ?thesis
unfolding observable.simps by blast
qed
lemma language_equivalence_classes_retain_language_and_induce_minimality :
assumes "transitions M' = (\<lambda> t . ({q \<in> states M . LS M q = LS M (t_source t)} , t_input t, t_output t, {q \<in> states M . LS M q = LS M (t_target t)})) ` transitions M"
and "states M' = (\<lambda>q . {q' \<in> states M . LS M q = LS M q'}) ` states M"
and "initial M' = {q' \<in> states M . LS M q' = LS M (initial M)}"
and "observable M"
shows "L M = L M'"
and "minimal M'"
proof -
have "observable M'"
using assms(1,4) language_equivalence_classes_preserve_observability by blast
have ls_prop: "\<And> io q . q \<in> states M \<Longrightarrow> (io \<in> LS M q) \<longleftrightarrow> (io \<in> LS M' {q' \<in> states M . LS M q = LS M q'})"
proof -
fix io q assume "q \<in> states M"
then show "(io \<in> LS M q) \<longleftrightarrow> (io \<in> LS M' {q' \<in> states M . LS M q = LS M q'})"
proof (induction io arbitrary: q)
case Nil
then show ?case using assms(2) by auto
next
case (Cons xy io)
obtain x y where "xy = (x,y)"
using surjective_pairing by blast
have "xy#io \<in> LS M q \<Longrightarrow> xy#io \<in> LS M' {q' \<in> states M . LS M q = LS M q'}"
proof -
assume "xy#io \<in> LS M q"
then obtain p where "path M q p" and "p_io p = xy#io"
unfolding LS.simps mem_Collect_eq by (metis (no_types, lifting))
let ?t = "hd p"
let ?p = "tl p"
let ?q' = "{q' \<in> states M . LS M (t_target ?t) = LS M q'}"
have "p = ?t # ?p" and "p_io ?p = io" and "t_input ?t = x" and "t_output ?t = y"
using \<open>p_io p = xy#io\<close> unfolding \<open>xy = (x,y)\<close> by auto
moreover have "?t \<in> transitions M" and "path M (t_target ?t) ?p" and "t_source ?t = q"
using \<open>path M q p\<close> path_cons_elim[of M q ?t ?p] calculation by auto
ultimately have "[(x,y)] \<in> LS M q"
unfolding LS_single_transition[of x y M q] by auto
then have "io \<in> LS M (t_target ?t)"
using observable_language_next[OF _ \<open>observable M\<close>, of "(x,y)" io, OF _ \<open>?t \<in> transitions M\<close>]
\<open>xy#io \<in> LS M q\<close>
unfolding \<open>xy = (x,y)\<close> \<open>t_source ?t = q\<close> \<open>t_input ?t = x\<close> \<open>t_output ?t = y\<close>
by (metis \<open>?t \<in> FSM.transitions M\<close> from_FSM_language fsm_transition_target fst_conv snd_conv)
then have "io \<in> LS M' ?q'"
using Cons.IH[OF fsm_transition_target[OF \<open>?t \<in> transitions M\<close>]] by blast
then obtain p' where "path M' ?q' p'" and "p_io p' = io"
by auto
have *: "({q' \<in> states M . LS M q = LS M q'},x,y,{q' \<in> states M . LS M (t_target ?t) = LS M q'}) \<in> transitions M'"
using \<open>?t \<in> transitions M\<close> \<open>t_source ?t = q\<close> \<open>t_input ?t = x\<close> \<open>t_output ?t = y\<close>
unfolding assms(1) by auto
show "xy#io \<in> LS M' {q' \<in> states M . LS M q = LS M q'}"
using LS_prepend_transition[OF * ] unfolding snd_conv fst_conv \<open>xy = (x,y)\<close>
using \<open>io \<in> LS M' ?q'\<close> by blast
qed
moreover have "xy#io \<in> LS M' {q' \<in> states M . LS M q = LS M q'} \<Longrightarrow> xy#io \<in> LS M q"
proof -
let ?q = "{q' \<in> states M . LS M q = LS M q'}"
assume "xy#io \<in> LS M' ?q"
then obtain p where "path M' ?q p" and "p_io p = xy#io"
unfolding LS.simps mem_Collect_eq by (metis (no_types, lifting))
let ?t = "hd p"
let ?p = "tl p"
have "p = ?t # ?p" and "p_io ?p = io" and "t_input ?t = x" and "t_output ?t = y"
using \<open>p_io p = xy#io\<close> unfolding \<open>xy = (x,y)\<close> by auto
then have "path M' ?q (?t#?p)"
using \<open>path M' ?q p\<close> by auto
then have "?t \<in> transitions M'" and "path M' (t_target ?t) ?p" and "t_source ?t = ?q"
by force+
then have "io \<in> LS M' (t_target ?t)"
using \<open>p_io ?p = io\<close> by auto
obtain t0 where t0_def: "?t = (\<lambda> t . ({q \<in> states M . LS M q = LS M (t_source t)} , t_input t, t_output t, {q \<in> states M . LS M q = LS M (t_target t)})) t0"
and "t0 \<in> transitions M"
using \<open>?t \<in> transitions M'\<close>
unfolding assms(1)
by auto
then have "t_source t0 \<in> ?q"
using \<open>t_source ?t = ?q\<close>
by (metis (mono_tags, lifting) fsm_transition_source fst_eqD mem_Collect_eq)
then have "LS M q = LS M (t_source t0)"
by auto
moreover have "[(x,y)] \<in> LS M (t_source t0)"
using t0_def \<open>t_input ?t = x\<close> \<open>t0 \<in> transitions M\<close> \<open>t_output ?t = y\<close> \<open>t_source t0 \<in> ?q\<close> unfolding LS_single_transition by auto
ultimately obtain t where "t \<in> transitions M" and "t_source t = q" and "t_input t = x" and "t_output t = y"
by (metis LS_single_transition)
have "LS M (t_target t) = LS M (t_target t0)"
using observable_transition_target_language_eq[OF _\<open>t \<in> transitions M\<close> \<open>t0 \<in> transitions M\<close> _ _ \<open>observable M\<close>]
using \<open>LS M q = LS M (t_source t0)\<close>
unfolding \<open>t_source t = q\<close> \<open>t_input t = x\<close> \<open>t_output t = y\<close>
using t0_def \<open>t_input ?t = x\<close> \<open>t_output ?t = y\<close>
by auto
moreover have "t_target ?t = {q' \<in> FSM.states M. LS M (t_target t) = LS M q'}"
using calculation t0_def by fastforce
ultimately have "io \<in> LS M (t_target t)"
using Cons.IH[OF fsm_transition_target[OF \<open>t \<in> transitions M\<close>]]
\<open>io \<in> LS M' (t_target ?t)\<close>
by auto
then show "xy#io \<in> LS M q"
unfolding \<open>t_source t = q\<close>[symmetric] \<open>xy = (x,y)\<close>
using \<open>t_input t = x\<close> \<open>t_output t = y\<close>
using LS_prepend_transition \<open>t \<in> FSM.transitions M\<close>
by blast
qed
ultimately show ?case
by blast
qed
qed
have "L M' = LS M' {q' \<in> states M . LS M (initial M) = LS M q'}"
using assms(3)
by (metis (mono_tags, lifting) Collect_cong)
then show "L M = L M'"
using ls_prop[OF fsm_initial] by blast
show "minimal M'"
proof -
have"\<And> q q' . q \<in> states M' \<Longrightarrow> q' \<in> states M' \<Longrightarrow> LS M' q = LS M' q' \<Longrightarrow> q = q'"
proof -
fix q q' assume "q \<in> states M'" and "q' \<in> states M'" and "LS M' q = LS M' q'"
obtain qM where "q = {q \<in> states M . LS M qM = LS M q}" and "qM \<in> states M"
using \<open>q \<in> states M'\<close> assms(2) by auto
obtain qM' where "q' = {q \<in> states M . LS M qM' = LS M q}" and "qM' \<in> states M"
using \<open>q' \<in> states M'\<close> assms(2) by auto
have "LS M qM = LS M' q"
using ls_prop[OF \<open>qM \<in> states M\<close>] unfolding \<open>q = {q \<in> states M . LS M qM = LS M q}\<close> by blast
moreover have "LS M qM' = LS M' q'"
using ls_prop[OF \<open>qM' \<in> states M\<close>] unfolding \<open>q' = {q \<in> states M . LS M qM' = LS M q}\<close> by blast
ultimately have "LS M qM = LS M qM'"
using \<open>LS M' q = LS M' q'\<close> by blast
then show "q = q'"
unfolding \<open>q = {q \<in> states M . LS M qM = LS M q}\<close> \<open>q' = {q \<in> states M . LS M qM' = LS M q}\<close> by blast
qed
then show ?thesis
unfolding minimal_alt_def by blast
qed
qed
fun minimise :: "('a :: linorder,'b :: linorder,'c :: linorder) fsm \<Rightarrow> ('a set,'b,'c) fsm" where
"minimise M = (let
eq_class = ofsm_table_fix M (\<lambda>q . states M) 0;
ts = (\<lambda> t . (eq_class (t_source t), t_input t, t_output t, eq_class (t_target t))) ` (transitions M);
q0 = eq_class (initial M);
eq_states = eq_class |`| fstates M;
M' = create_unconnected_fsm_from_fsets q0 eq_states (finputs M) (foutputs M)
in add_transitions M' ts)"
lemma minimise_initial_partition :
"equivalence_relation_on_states M (\<lambda>q . states M)"
proof -
let ?r = "{(q1,q2) | q1 q2 . q1 \<in> states M \<and> q2 \<in> (\<lambda>q . states M) q1}"
have "refl_on (FSM.states M) ?r"
unfolding refl_on_def by blast
moreover have "sym ?r"
unfolding sym_def by blast
moreover have "trans ?r"
unfolding trans_def by blast
ultimately show ?thesis
unfolding equivalence_relation_on_states_def equiv_def by auto
qed
lemma minimise_props:
assumes "observable M"
shows "initial (minimise M) = {q' \<in> states M . LS M q' = LS M (initial M)}"
and "states (minimise M) = (\<lambda>q . {q' \<in> states M . LS M q = LS M q'}) ` states M"
and "inputs (minimise M) = inputs M"
and "outputs (minimise M) = outputs M"
and "transitions (minimise M) = (\<lambda> t . ({q \<in> states M . LS M q = LS M (t_source t)} , t_input t, t_output t, {q \<in> states M . LS M q = LS M (t_target t)})) ` transitions M"
proof -
let ?f = "\<lambda>q . states M"
define eq_class where "eq_class = ofsm_table_fix M (\<lambda>q . states M) 0"
moreover define M' where M'_def: "M' = create_unconnected_fsm_from_fsets (eq_class (initial M)) (eq_class |`| fstates M) (finputs M) (foutputs M)"
ultimately have *: "minimise M = add_transitions M' ((\<lambda> t . (eq_class (t_source t), t_input t, t_output t, eq_class (t_target t))) ` (transitions M))"
by auto
have **: "\<And> q . q \<in> states M \<Longrightarrow> eq_class q = {q' \<in> FSM.states M. LS M q = LS M q'}"
using ofsm_table_fix_set[OF _ assms minimise_initial_partition] \<open>eq_class = ofsm_table_fix M ?f 0\<close> after_is_state[OF \<open>observable M\<close>] by blast
then have ****: "\<And> q . q \<in> states M \<Longrightarrow> eq_class q = {q' \<in> FSM.states M. LS M q' = LS M q}"
using ofsm_table_fix_set[OF _ assms] \<open>eq_class = ofsm_table_fix M ?f 0\<close> by blast
have ***: "(eq_class (initial M)) |\<in>| (eq_class |`| fstates M)"
- using fsm_initial[of M] fstates_set fmember_iff_member_fset by fastforce
+ using fsm_initial[of M] fstates_set by fastforce
have m1:"initial M' = {q' \<in> states M . LS M q' = LS M (initial M)}"
by (metis (mono_tags) "***" "****" M'_def create_unconnected_fsm_from_fsets_simps(1) fsm_initial)
have m2: "states M' = (\<lambda>q . {q' \<in> states M . LS M q = LS M q'}) ` states M"
unfolding M'_def
proof -
have "FSM.states (FSM.create_unconnected_fsm_from_fsets (eq_class (FSM.initial M)) (eq_class |`| fstates M) (finputs M) (foutputs M)) = eq_class ` FSM.states M"
by (metis (no_types) "***" create_unconnected_fsm_from_fsets_simps(2) fset.set_map fstates_set)
then show "FSM.states (FSM.create_unconnected_fsm_from_fsets (eq_class (FSM.initial M)) (eq_class |`| fstates M) (finputs M) (foutputs M)) = (\<lambda>a. {aa \<in> FSM.states M. LS M a = LS M aa}) ` FSM.states M"
using "**" by force
qed
have m3: "inputs M' = inputs M"
- using create_unconnected_fsm_from_fsets_simps(3)[OF ***] finputs_set unfolding M'_def by force
+ using create_unconnected_fsm_from_fsets_simps(3)[OF ***] finputs_set unfolding M'_def by metis
have m4: "outputs M' = outputs M"
- using create_unconnected_fsm_from_fsets_simps(4)[OF ***] foutputs_set unfolding M'_def by force
+ using create_unconnected_fsm_from_fsets_simps(4)[OF ***] foutputs_set unfolding M'_def by metis
have m5: "transitions M' = {}"
using create_unconnected_fsm_from_fsets_simps(5)[OF ***] unfolding M'_def by force
let ?ts = "((\<lambda> t . (eq_class (t_source t), t_input t, t_output t, eq_class (t_target t))) ` (transitions M))"
have wf: "\<And> t . t \<in>?ts \<Longrightarrow> t_source t \<in> states M' \<and> t_input t \<in> inputs M' \<and> t_output t \<in> outputs M' \<and> t_target t \<in> states M'"
proof -
fix t assume "t \<in> ?ts"
then obtain tM where "tM \<in> transitions M"
and *: "t = (\<lambda> t . (eq_class (t_source t), t_input t, t_output t, eq_class (t_target t))) tM"
by blast
have "t_source t \<in> states M'"
using fsm_transition_source[OF \<open>tM \<in> transitions M\<close>]
unfolding m2 * **[OF fsm_transition_source[OF \<open>tM \<in> transitions M\<close>]] by auto
moreover have "t_input t \<in> inputs M'"
unfolding m3 * using fsm_transition_input[OF \<open>tM \<in> transitions M\<close>] by auto
moreover have "t_output t \<in> outputs M'"
unfolding m4 * using fsm_transition_output[OF \<open>tM \<in> transitions M\<close>] by auto
moreover have "t_target t \<in> states M'"
using fsm_transition_target[OF \<open>tM \<in> transitions M\<close>]
unfolding m2 * **[OF fsm_transition_target[OF \<open>tM \<in> transitions M\<close>]] by auto
ultimately show "t_source t \<in> states M' \<and> t_input t \<in> inputs M' \<and> t_output t \<in> outputs M' \<and> t_target t \<in> states M'"
by simp
qed
show "initial (minimise M) = {q' \<in> states M . LS M q' = LS M (initial M)}"
using add_transitions_simps(1)[OF wf] unfolding * m1 .
show "states (minimise M) = (\<lambda>q . {q' \<in> states M . LS M q = LS M q'}) ` states M"
using add_transitions_simps(2)[OF wf] unfolding * m2 .
show "inputs (minimise M) = inputs M"
using add_transitions_simps(3)[OF wf] unfolding * m3 .
show "outputs (minimise M) = outputs M"
using add_transitions_simps(4)[OF wf] unfolding * m4 .
show "transitions (minimise M) = (\<lambda> t . ({q \<in> states M . LS M q = LS M (t_source t)} , t_input t, t_output t, {q \<in> states M . LS M q = LS M (t_target t)})) ` transitions M"
using add_transitions_simps(5)[OF wf] ****[OF fsm_transition_source] ****[OF fsm_transition_target] unfolding * m5 by auto
qed
lemma minimise_observable:
assumes "observable M"
shows "observable (minimise M)"
using language_equivalence_classes_preserve_observability[OF minimise_props(5)[OF assms] assms]
by assumption
lemma minimise_minimal:
assumes "observable M"
shows "minimal (minimise M)"
using language_equivalence_classes_retain_language_and_induce_minimality(2)[OF minimise_props(5,2,1)[OF assms] assms]
by assumption
lemma minimise_language:
assumes "observable M"
shows "L (minimise M) = L M"
using language_equivalence_classes_retain_language_and_induce_minimality(1)[OF minimise_props(5,2,1)[OF assms] assms]
by blast
lemma minimal_observable_code :
assumes "observable M"
shows "minimal M = (\<forall> q \<in> states M . ofsm_table_fix M (\<lambda>q . states M) 0 q = {q})"
proof
show "minimal M \<Longrightarrow> (\<forall> q \<in> states M . ofsm_table_fix M (\<lambda>q . states M) 0 q = {q})"
proof
fix q assume "minimal M" and "q \<in> states M"
then show "ofsm_table_fix M (\<lambda>q . states M) 0 q = {q}"
unfolding ofsm_table_fix_set[OF \<open>q \<in> states M\<close> \<open>observable M\<close> minimise_initial_partition]
minimal_alt_def
using after_is_state[OF \<open>observable M\<close>]
by blast
qed
show "\<forall>q\<in>FSM.states M. ofsm_table_fix M (\<lambda>q . states M) 0 q = {q} \<Longrightarrow> minimal M"
using ofsm_table_fix_set[OF _ \<open>observable M\<close> minimise_initial_partition] after_is_state[OF \<open>observable M\<close>]
unfolding minimal_alt_def
by blast
qed
lemma minimise_states_subset :
assumes "observable M"
and "q \<in> states (minimise M)"
shows "q \<subseteq> states M"
using assms(2)
unfolding minimise_props[OF assms(1)]
by auto
lemma minimise_states_finite :
assumes "observable M"
and "q \<in> states (minimise M)"
shows "finite q"
using minimise_states_subset[OF assms] fsm_states_finite[of M]
using finite_subset by auto
end
\ No newline at end of file
diff --git a/thys/FSM_Tests/Observability.thy b/thys/FSM_Tests/Observability.thy
--- a/thys/FSM_Tests/Observability.thy
+++ b/thys/FSM_Tests/Observability.thy
@@ -1,1288 +1,1280 @@
section \<open>Observability\<close>
text \<open>This theory presents the classical algorithm for transforming FSMs into
language-equivalent observable FSMs in analogy to the determinisation of nondeterministic
finite automata.\<close>
theory Observability
imports FSM
begin
lemma fPow_Pow : "Pow (fset A) = fset (fset |`| fPow A)"
proof (induction A)
case empty
then show ?case by auto
next
case (insert x A)
have "Pow (fset (finsert x A)) = Pow (fset A) \<union> (insert x) ` Pow (fset A)"
by (simp add: Pow_insert)
moreover have "fset (fset |`| fPow (finsert x A)) = fset (fset |`| fPow A) \<union> (insert x) ` fset (fset |`| fPow A)"
proof -
have "fset |`| ((fPow A) |\<union>| (finsert x) |`| (fPow A)) = (fset |`| fPow A) |\<union>| (insert x) |`| (fset |`| fPow A)"
unfolding fimage_funion
by fastforce
moreover have "(fPow (finsert x A)) = (fPow A) |\<union>| (finsert x) |`| (fPow A)"
by (simp add: fPow_finsert)
ultimately show ?thesis
by auto
qed
ultimately show ?case
using insert.IH by simp
qed
lemma fcard_fsubset: "\<not> fcard (A |-| (B |\<union>| C)) < fcard (A |-| B) \<Longrightarrow> C |\<subseteq>| A \<Longrightarrow> C |\<subseteq>| B"
proof (induction C)
case empty
then show ?case by auto
next
case (insert x C)
then show ?case
unfolding finsert_fsubset funion_finsert_right not_less
proof -
assume a1: "fcard (A |-| B) \<le> fcard (A |-| finsert x (B |\<union>| C))"
assume "\<lbrakk>fcard (A |-| B) \<le> fcard (A |-| (B |\<union>| C)); C |\<subseteq>| A\<rbrakk> \<Longrightarrow> C |\<subseteq>| B"
assume a2: "x |\<in>| A \<and> C |\<subseteq>| A"
have "A |-| (C |\<union>| finsert x B) = A |-| B \<or> \<not> A |-| (C |\<union>| finsert x B) |\<subseteq>| A |-| B"
using a1 by (metis (no_types) fcard_seteq funion_commute funion_finsert_right)
then show "x |\<in>| B \<and> C |\<subseteq>| B"
using a2 by blast
qed
qed
lemma make_observable_transitions_qtrans_helper:
assumes "qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) A;
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)) nexts)"
shows "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| A \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
proof -
have "fset qtrans = { (q,x,y,q') | q x y q' . q |\<in>| nexts \<and> q' \<noteq> {||} \<and> fset q' = t_target ` {t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y}}"
proof -
have "\<And> q . fset (ffilter (\<lambda>t . t_source t |\<in>| q) A) = Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A)"
using ffilter.rep_eq assms(1) by auto
then have "\<And> q . fset (fimage (\<lambda> t . (t_input t, t_output t)) (ffilter (\<lambda>t . t_source t |\<in>| q) A)) = image (\<lambda> t . (t_input t, t_output t)) (Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A))"
by simp
then have *:"\<And> q . fset (fimage (\<lambda>(x,y) . (q,x,y, (t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))))) (fimage (\<lambda> t . (t_input t, t_output t)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))
= image (\<lambda>(x,y) . (q,x,y, (t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))))) (image (\<lambda> t . (t_input t, t_output t)) (Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A)))"
by (metis (no_types, lifting) ffilter.rep_eq fset.set_map)
have **: "\<And> f1 f2 xs ys ys' . (\<And> x . fset (f1 x ys) = (f2 x ys')) \<Longrightarrow>
fset (ffUnion (fimage (\<lambda> x . (f1 x ys)) xs)) = (\<Union> x \<in> fset xs . (f2 x ys'))"
unfolding ffUnion.rep_eq fimage.rep_eq by force
have "fset (ffUnion (fimage (\<lambda> q . (fimage (\<lambda>(x,y) . (q,x,y, (t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))))) (fimage (\<lambda> t . (t_input t, t_output t)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))) nexts))
= (\<Union> q \<in> fset nexts . image (\<lambda>(x,y) . (q,x,y, (t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))))) (image (\<lambda> t . (t_input t, t_output t)) (Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A))))"
unfolding ffUnion.rep_eq fimage.rep_eq
using "*" by force
also have "\<dots> = { (q,x,y,q') | q x y q' . q |\<in>| nexts \<and> q' \<noteq> {||} \<and> fset q' = t_target ` {t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y}}"
(is "?A = ?B") proof -
have "\<And> t . t \<in> ?A \<Longrightarrow> t \<in> ?B"
proof -
fix t assume "t \<in> ?A"
then obtain q where "q \<in> fset nexts"
and "t \<in> image (\<lambda>(x,y) . (q,x,y, (t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))))) (image (\<lambda> t . (t_input t, t_output t)) (Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A)))"
by blast
then obtain x y q' where *: "(x,y) \<in> (image (\<lambda> t . (t_input t, t_output t)) (Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A)))"
and "t = (q,x,y,q')"
and **:"q' = (t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A)))))"
by force
have "q |\<in>| nexts"
using \<open>q \<in> fset nexts\<close>
- by (meson notin_fset)
+ by simp
moreover have "q' \<noteq> {||}"
proof -
have ***:"(Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A)) = fset (ffilter (\<lambda>t . t_source t |\<in>| q) (A))"
by auto
have "\<exists> t . t |\<in>| (ffilter (\<lambda>t. t_source t |\<in>| q) A) \<and> (t_input t, t_output t) = (x,y)"
using *
- by (metis (no_types, lifting) "***" image_iff notin_fset)
+ by (metis (no_types, lifting) "***" image_iff)
then show ?thesis unfolding **
by force
qed
moreover have "fset q' = t_target ` {t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y}"
proof -
have "{t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y} = ((Set.filter (\<lambda>t . (t_input t, t_output t) = (x,y)) (fset (ffilter (\<lambda>t . t_source t |\<in>| q) (A)))))"
- using notin_fset by fastforce
+ by fastforce
also have "\<dots> = fset ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))"
by fastforce
finally have "{t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y} = fset ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))" .
then show ?thesis
unfolding **
by simp
qed
ultimately show "t \<in> ?B"
unfolding \<open>t = (q,x,y,q')\<close>
by blast
qed
moreover have "\<And> t . t \<in> ?B \<Longrightarrow> t \<in> ?A"
proof -
fix t assume "t \<in> ?B"
then obtain q x y q' where "t = (q,x,y,q')" and "(q,x,y,q') \<in> ?B" by force
then have "q |\<in>| nexts"
and "q' \<noteq> {||}"
and *: "fset q' = t_target ` {t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y}"
by force+
then have "fset q' \<noteq> {}"
by (metis bot_fset.rep_eq fset_inject)
have "(x,y) \<in> (image (\<lambda> t . (t_input t, t_output t)) (Set.filter (\<lambda>t . t_source t |\<in>| q) (fset A)))"
- proof -
- have **:"\<And> t . t |\<in>| A = (t \<in> fset A)"
- by (meson notin_fset)
- show ?thesis
- using \<open>fset q' \<noteq> {}\<close> unfolding * Set.filter_def ** by blast
- qed
+ using \<open>fset q' \<noteq> {}\<close> unfolding * Set.filter_def by blast
moreover have "q' = t_target |`| ffilter (\<lambda>t. (t_input t, t_output t) = (x, y)) (ffilter (\<lambda>t. t_source t |\<in>| q) A)"
proof -
have "{t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y} = ((Set.filter (\<lambda>t . (t_input t, t_output t) = (x,y)) (fset (ffilter (\<lambda>t . t_source t |\<in>| q) (A)))))"
- using notin_fset by fastforce
+ by fastforce
also have "\<dots> = fset ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))"
by fastforce
finally have ***:"{t' . t' |\<in>| A \<and> t_source t' |\<in>| q \<and> t_input t' = x \<and> t_output t' = y} = fset ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) (ffilter (\<lambda>t . t_source t |\<in>| q) (A))))" .
show ?thesis
using *
unfolding ***
by (metis (no_types, lifting) fimage.rep_eq fset_inject)
qed
- moreover have "q \<in> fset nexts"
+ ultimately show "t \<in> ?A"
using \<open>q |\<in>| nexts\<close>
- by (meson notin_fset)
- ultimately show "t \<in> ?A"
unfolding \<open>t = (q,x,y,q')\<close>
by force
qed
ultimately show ?thesis
by (metis (no_types, lifting) Collect_cong Sup_set_def mem_Collect_eq)
qed
finally show ?thesis
unfolding assms Let_def by blast
qed
then show "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| A \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
- unfolding fmember_iff_member_fset by force
+ by force
qed
function make_observable_transitions :: "('a,'b,'c) transition fset \<Rightarrow> 'a fset fset \<Rightarrow> 'a fset fset \<Rightarrow> ('a fset \<times> 'b \<times> 'c \<times> 'a fset) fset \<Rightarrow> ('a fset \<times> 'b \<times> 'c \<times> 'a fset) fset" where
"make_observable_transitions base_trans nexts dones ts = (let
qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) base_trans;
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| (ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts))) ios)) nexts);
dones' = dones |\<union>| nexts;
ts' = ts |\<union>| qtrans;
nexts' = (fimage t_target qtrans) |-| dones'
in if nexts' = {||}
then ts'
else make_observable_transitions base_trans nexts' dones' ts')"
by auto
termination
proof -
{
fix base_trans :: "('a,'b,'c) transition fset"
fix nexts :: "'a fset fset"
fix dones :: "'a fset fset"
fix ts :: "('a fset \<times> 'b \<times> 'c \<times> 'a fset) fset"
fix q x y q'
assume assm1: "\<not> fcard
(fPow (t_source |`| base_trans |\<union>| t_target |`| base_trans) |-|
(dones |\<union>| nexts |\<union>|
t_target |`|
ffUnion
((\<lambda>q. let qts = ffilter (\<lambda>t. t_source t |\<in>| q) base_trans
in ((\<lambda>(x, y). (q, x, y, t_target |`| ffilter (\<lambda>t. t_input t = x \<and> t_output t = y) qts)) \<circ> (\<lambda>t. (t_input t, t_output t))) |`|
qts) |`|
nexts)))
< fcard (fPow (t_source |`| base_trans |\<union>| t_target |`| base_trans) |-| (dones |\<union>| nexts))"
and assm2: "(q, x, y, q') |\<in>|
ffUnion
((\<lambda>q. let qts = ffilter (\<lambda>t. t_source t |\<in>| q) base_trans
in ((\<lambda>(x, y). (q, x, y, t_target |`| ffilter (\<lambda>t. t_input t = x \<and> t_output t = y) qts)) \<circ> (\<lambda>t. (t_input t, t_output t))) |`| qts) |`|
nexts)"
and assm3: "q' |\<notin>| nexts"
define qtrans where qtrans_def: "qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) base_trans;
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)) nexts)"
have qtrans_prop: "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' | t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using make_observable_transitions_qtrans_helper[OF qtrans_def]
by presburger
have "\<And> t . t |\<in>| qtrans \<Longrightarrow> t_target t |\<in>| fPow (t_target |`| base_trans)"
proof -
fix t assume "t |\<in>| qtrans"
then have *: "fset (t_target t) = t_target ` {t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using qtrans_prop by blast
then have "fset (t_target t) \<subseteq> t_target ` (fset base_trans)"
- by (metis (mono_tags, lifting) imageI image_Collect_subsetI notin_fset)
+ by (metis (mono_tags, lifting) imageI image_Collect_subsetI)
then show "t_target t |\<in>| fPow (t_target |`| base_trans)"
by (simp add: less_eq_fset.rep_eq)
qed
then have "t_target |`| qtrans |\<subseteq>| (fPow (t_source |`| base_trans |\<union>| t_target |`| base_trans))"
by fastforce
moreover have " \<not> fcard (fPow (t_source |`| base_trans |\<union>| t_target |`| base_trans) |-| (dones |\<union>| nexts |\<union>| t_target |`| qtrans))
< fcard (fPow (t_source |`| base_trans |\<union>| t_target |`| base_trans) |-| (dones |\<union>| nexts))"
using assm1 unfolding qtrans_def by force
ultimately have "t_target |`| qtrans |\<subseteq>| dones |\<union>| nexts"
using fcard_fsubset by fastforce
moreover have "q' |\<in>| t_target |`| qtrans"
using assm2 unfolding qtrans_def by force
ultimately have "q' |\<in>| dones"
using \<open>q' |\<notin>| nexts\<close> by blast
} note t = this
show ?thesis
apply (relation "measure (\<lambda> (base_trans, nexts, dones, ts) . fcard ((fPow (t_source |`| base_trans |\<union>| t_target |`| base_trans)) |-| (dones |\<union>| nexts)))")
apply auto
by (erule t)
qed
lemma make_observable_transitions_mono: "ts |\<subseteq>| (make_observable_transitions base_trans nexts dones ts)"
proof (induction rule: make_observable_transitions.induct[of "\<lambda> base_trans nexts dones ts . ts |\<subseteq>| (make_observable_transitions base_trans nexts dones ts)"])
case (1 base_trans nexts dones ts)
define qtrans where qtrans_def: "qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) base_trans;
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)) nexts)"
have qtrans_prop: "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' | t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using make_observable_transitions_qtrans_helper[OF qtrans_def]
by presburger
let ?dones' = "dones |\<union>| nexts"
let ?ts' = "ts |\<union>| qtrans"
let ?nexts' = "(fimage t_target qtrans) |-| ?dones'"
have res_cases: "make_observable_transitions base_trans nexts dones ts = (if ?nexts' = {||}
then ?ts'
else make_observable_transitions base_trans ?nexts' ?dones' ?ts')"
unfolding make_observable_transitions.simps[of base_trans nexts dones ts] qtrans_def Let_def by simp
show ?case proof (cases "?nexts' = {||}")
case True
then show ?thesis using res_cases by simp
next
case False
then have "make_observable_transitions base_trans nexts dones ts = make_observable_transitions base_trans ?nexts' ?dones' ?ts'"
using res_cases by simp
moreover have "ts |\<union>| qtrans |\<subseteq>| make_observable_transitions base_trans ?nexts' ?dones' ?ts'"
using "1"[OF qtrans_def _ _ _ False, of ?dones' ?ts'] by blast
ultimately show ?thesis
by blast
qed
qed
inductive pathlike :: "('state, 'input, 'output) transition fset \<Rightarrow> 'state \<Rightarrow> ('state, 'input, 'output) path \<Rightarrow> bool"
where
nil[intro!] : "pathlike ts q []" |
cons[intro!] : "t |\<in>| ts \<Longrightarrow> pathlike ts (t_target t) p \<Longrightarrow> pathlike ts (t_source t) (t#p)"
inductive_cases pathlike_nil_elim[elim!]: "pathlike ts q []"
inductive_cases pathlike_cons_elim[elim!]: "pathlike ts q (t#p)"
lemma make_observable_transitions_t_source :
assumes "\<And> t . t |\<in>| ts \<Longrightarrow> t_source t |\<in>| dones \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
and "\<And> q t' . q |\<in>| dones \<Longrightarrow> t' |\<in>| base_trans \<Longrightarrow> t_source t' |\<in>| q \<Longrightarrow> \<exists> t . t |\<in>| ts \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t'"
and "t |\<in>| make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts"
and "t_source t |\<in>| dones"
shows "t |\<in>| ts"
using assms proof (induction base_trans "(fimage t_target ts) |-| dones" dones ts rule: make_observable_transitions.induct)
case (1 base_trans dones ts)
let ?nexts = "(fimage t_target ts) |-| dones"
define qtrans where qtrans_def: "qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) base_trans;
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)) ?nexts)"
have qtrans_prop: "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| ?nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using make_observable_transitions_qtrans_helper[OF qtrans_def]
by presburger
let ?dones' = "dones |\<union>| ?nexts"
let ?ts' = "ts |\<union>| qtrans"
let ?nexts' = "(fimage t_target qtrans) |-| ?dones'"
have res_cases: "make_observable_transitions base_trans ?nexts dones ts = (if ?nexts' = {||}
then ?ts'
else make_observable_transitions base_trans ?nexts' ?dones' ?ts')"
unfolding make_observable_transitions.simps[of base_trans ?nexts dones ts] qtrans_def Let_def by simp
show ?case proof (cases "?nexts' = {||}")
case True
then have "make_observable_transitions base_trans ?nexts dones ts = ?ts'"
using res_cases by auto
then have "t |\<in>| ts |\<union>| qtrans"
using \<open>t |\<in>| make_observable_transitions base_trans ?nexts dones ts\<close> \<open>t_source t |\<in>| dones\<close> by blast
then show ?thesis
using qtrans_prop "1.prems"(3,4) by blast
next
case False
then have "make_observable_transitions base_trans ?nexts dones ts = make_observable_transitions base_trans ?nexts' ?dones' ?ts'"
using res_cases by simp
have i1: "(\<And>t. t |\<in>| ts |\<union>| qtrans \<Longrightarrow>
t_source t |\<in>| dones |\<union>| ?nexts \<and>
t_target t \<noteq> {||} \<and>
fset (t_target t) =
t_target `
{t' . t' |\<in>| base_trans \<and>
t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t})"
using "1.prems"(1) qtrans_prop by blast
have i3: "t_target |`| qtrans |-| (dones |\<union>| ?nexts) = t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| ?nexts)"
unfolding "1.prems"(3) by blast
have i2: "(\<And>q t'.
q |\<in>| dones |\<union>| ?nexts \<Longrightarrow>
t' |\<in>| base_trans \<Longrightarrow>
t_source t' |\<in>| q \<Longrightarrow>
\<exists>t. t |\<in>| ts |\<union>| qtrans \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t')"
proof -
fix q t' assume "q |\<in>| dones |\<union>| ?nexts"
and *:"t' |\<in>| base_trans"
and **:"t_source t' |\<in>| q"
then consider (a) "q |\<in>| dones" | (b) "q |\<in>| ?nexts" by blast
then show "\<exists>t. t |\<in>| ts |\<union>| qtrans \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t'"
proof cases
case a
then show ?thesis using * **
using "1.prems"(2) by blast
next
case b
let ?tgts = "{t'' . t'' |\<in>| base_trans \<and> t_source t'' |\<in>| q \<and> t_input t'' = t_input t' \<and> t_output t'' = t_output t'}"
define tgts where tgts: "tgts = Abs_fset (t_target ` ?tgts)"
have "?tgts \<subseteq> fset base_trans"
- using notin_fset by fastforce
+ by fastforce
then have "finite (t_target ` ?tgts)"
by (meson finite_fset finite_imageI finite_subset)
then have "fset tgts = (t_target ` ?tgts)"
unfolding tgts
using Abs_fset_inverse
by blast
have "?tgts \<noteq> {}"
using * ** by blast
then have "t_target ` ?tgts \<noteq> {}"
by blast
then have "tgts \<noteq> {||}"
using \<open>fset tgts = (t_target ` ?tgts)\<close>
by force
then have "(q, t_input t', t_output t', tgts) |\<in>| qtrans"
using b
unfolding qtrans_prop[of "(q,t_input t',t_output t',tgts)"]
unfolding fst_conv snd_conv
unfolding \<open>fset tgts = (t_target ` ?tgts)\<close>[symmetric]
by blast
then show ?thesis
by auto
qed
qed
have "t |\<in>| make_observable_transitions base_trans ?nexts dones ts \<Longrightarrow> t_source t |\<in>| dones |\<union>| ?nexts \<Longrightarrow> t |\<in>| ts |\<union>| qtrans"
unfolding \<open>make_observable_transitions base_trans ?nexts dones ts = make_observable_transitions base_trans ?nexts' ?dones' ?ts'\<close>
using "1.hyps"[OF qtrans_def _ _ _ _ i1 i2] False i3 by force
then have "t |\<in>| ts |\<union>| qtrans"
using \<open>t |\<in>| make_observable_transitions base_trans ?nexts dones ts\<close> \<open>t_source t |\<in>| dones\<close> by blast
then show ?thesis
using qtrans_prop "1.prems"(3,4) by blast
qed
qed
lemma make_observable_transitions_path :
assumes "\<And> t . t |\<in>| ts \<Longrightarrow> t_source t |\<in>| dones \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' \<in> transitions M . t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
and "\<And> q t' . q |\<in>| dones \<Longrightarrow> t' \<in> transitions M \<Longrightarrow> t_source t' |\<in>| q \<Longrightarrow> \<exists> t . t |\<in>| ts \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t'"
and "\<And> q . q |\<in>| (fimage t_target ts) |-| dones \<Longrightarrow> q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M)"
and "\<And> q . q |\<in>| dones \<Longrightarrow> q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M |\<union>| {|initial M|})"
and "{||} |\<notin>| dones"
and "q |\<in>| dones"
shows "(\<exists> q' p . q' |\<in>| q \<and> path M q' p \<and> p_io p = io) \<longleftrightarrow> (\<exists> p'. pathlike (make_observable_transitions (ftransitions M) ((fimage t_target ts) |-| dones) dones ts) q p' \<and> p_io p' = io)"
using assms proof (induction "ftransitions M" "(fimage t_target ts) |-| dones" dones ts arbitrary: q io rule: make_observable_transitions.induct)
case (1 dones ts q)
let ?obs = "(make_observable_transitions (ftransitions M) ((fimage t_target ts) |-| dones) dones ts)"
let ?nexts = "(fimage t_target ts) |-| dones"
show ?case proof (cases io)
case Nil
have scheme: "\<And> q q' X . q' |\<in>| q \<Longrightarrow> q |\<in>| fPow X \<Longrightarrow> q' |\<in>| X"
by (simp add: fsubsetD)
obtain q' where "q' |\<in>| q"
- using \<open>{||} |\<notin>| dones\<close> \<open>q |\<in>| dones\<close> by fastforce
+ using \<open>{||} |\<notin>| dones\<close> \<open>q |\<in>| dones\<close>
+ by (metis all_not_in_conv bot_fset.rep_eq fset_cong)
have "q' |\<in>| t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M |\<union>| {|FSM.initial M|}"
using scheme[OF \<open>q' |\<in>| q\<close> "1.prems"(4)[OF \<open>q |\<in>| dones\<close>]] .
then have "q' \<in> states M"
using ftransitions_source[of q' M]
using ftransitions_target[of q' M]
by force
then have "\<exists> q' p . q' |\<in>| q \<and> path M q' p \<and> p_io p = io"
using \<open>q' |\<in>| q\<close> Nil by auto
moreover have "(\<exists> p'. pathlike ?obs q p' \<and> p_io p' = io)"
using Nil by auto
ultimately show ?thesis
by simp
next
case (Cons ioT ioP)
define qtrans where qtrans_def: "qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) (ftransitions M);
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)) ?nexts)"
have qtrans_prop: "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| ?nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| (ftransitions M) \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using make_observable_transitions_qtrans_helper[OF qtrans_def]
by presburger
let ?dones' = "dones |\<union>| ?nexts"
let ?ts' = "ts |\<union>| qtrans"
let ?nexts' = "(fimage t_target qtrans) |-| ?dones'"
have res_cases: "make_observable_transitions (ftransitions M) ?nexts dones ts = (if ?nexts' = {||}
then ?ts'
else make_observable_transitions (ftransitions M) ?nexts' ?dones' ?ts')"
unfolding make_observable_transitions.simps[of "ftransitions M" ?nexts dones ts] qtrans_def Let_def by simp
have i1: "(\<And>t. t |\<in>| ts |\<union>| qtrans \<Longrightarrow>
t_source t |\<in>| dones |\<union>| ?nexts \<and>
t_target t \<noteq> {||} \<and>
fset (t_target t) =
t_target `
{t' \<in> FSM.transitions M .
t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t})"
using "1.prems"(1) qtrans_prop
using ftransitions_set[of M]
- by (metis (mono_tags, lifting) Collect_cong fmember_iff_member_fset funion_iff)
+ by (metis (mono_tags, lifting) Collect_cong funion_iff)
have i2: "(\<And>q t'.
q |\<in>| dones |\<union>| ?nexts \<Longrightarrow>
t' \<in> FSM.transitions M \<Longrightarrow>
t_source t' |\<in>| q \<Longrightarrow>
\<exists>t. t |\<in>| ts |\<union>| qtrans \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t')"
proof -
fix q t' assume "q |\<in>| dones |\<union>| ?nexts"
and *:"t' \<in> FSM.transitions M"
and **:"t_source t' |\<in>| q"
then consider (a) "q |\<in>| dones" | (b) "q |\<in>| ?nexts" by blast
then show "\<exists>t. t |\<in>| ts |\<union>| qtrans \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t'"
proof cases
case a
then show ?thesis using "1.prems"(2) * ** by blast
next
case b
let ?tgts = "{t'' \<in> FSM.transitions M. t_source t'' |\<in>| q \<and> t_input t'' = t_input t' \<and> t_output t'' = t_output t'}"
have "?tgts \<noteq> {}"
using * ** by blast
let ?tgts = "{t'' . t'' |\<in>| ftransitions M \<and> t_source t'' |\<in>| q \<and> t_input t'' = t_input t' \<and> t_output t'' = t_output t'}"
define tgts where tgts: "tgts = Abs_fset (t_target ` ?tgts)"
have "?tgts \<subseteq> transitions M"
using ftransitions_set[of M]
- by (metis (no_types, lifting) mem_Collect_eq notin_fset subsetI)
+ by (metis (no_types, lifting) mem_Collect_eq subsetI)
then have "finite (t_target ` ?tgts)"
by (meson finite_imageI finite_subset fsm_transitions_finite)
then have "fset tgts = (t_target ` ?tgts)"
unfolding tgts
using Abs_fset_inverse
by blast
have "?tgts \<noteq> {}"
using * **
- by (metis (mono_tags, lifting) empty_iff ftransitions_set mem_Collect_eq notin_fset)
+ by (metis (mono_tags, lifting) empty_iff ftransitions_set mem_Collect_eq)
then have "t_target ` ?tgts \<noteq> {}"
by blast
then have "tgts \<noteq> {||}"
using \<open>fset tgts = (t_target ` ?tgts)\<close>
by force
then have "(q, t_input t', t_output t', tgts) |\<in>| qtrans"
using b
unfolding qtrans_prop[of "(q,t_input t',t_output t',tgts)"]
unfolding fst_conv snd_conv
unfolding \<open>fset tgts = (t_target ` ?tgts)\<close>[symmetric]
by blast
then show ?thesis
by auto
qed
qed
have i3: "t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones)) = t_target |`| qtrans |-| (dones |\<union>| (t_target |`| ts |-| dones))"
by blast
have i4: "(\<And>q. q |\<in>| t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones)) \<Longrightarrow>
q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M))"
proof -
fix q assume "q |\<in>| t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones))"
then have "q |\<in>| t_target |`| qtrans"
by auto
then obtain t where "t |\<in>| qtrans" and "t_target t = q"
by auto
then have "fset q = t_target ` {t'. t' |\<in>| ftransitions M \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
unfolding qtrans_prop by auto
then have "fset q \<subseteq> t_target ` transitions M"
- by (metis (no_types, lifting) ftransitions_set image_Collect_subsetI image_eqI notin_fset)
+ by (metis (no_types, lifting) ftransitions_set image_Collect_subsetI image_eqI)
then show "q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M)"
by (metis (no_types, lifting) fPowI fset.set_map fset_inject ftransitions_set le_supI2 sup.orderE sup.orderI sup_fset.rep_eq)
qed
have i5: "(\<And>q. q |\<in>| dones |\<union>| ?nexts \<Longrightarrow> q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M |\<union>| {|initial M|}))"
using "1.prems"(4,3) qtrans_prop
by auto
have i7: "{||} |\<notin>| dones |\<union>| ?nexts"
using "1.prems" by fastforce
show ?thesis
proof (cases "?nexts' \<noteq> {||}")
case False
then have "?obs = ?ts'"
using res_cases by auto
have "\<And> q io . q |\<in>| ?dones' \<Longrightarrow> q \<noteq> {||} \<Longrightarrow> (\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = io) \<longleftrightarrow> (\<exists>p'. pathlike ?obs q p' \<and> p_io p' = io)"
proof -
fix q io assume "q |\<in>| ?dones'" and "q \<noteq> {||}"
then show "(\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = io) \<longleftrightarrow> (\<exists>p'. pathlike ?obs q p' \<and> p_io p' = io)"
proof (induction io arbitrary: q)
case Nil
have scheme: "\<And> q q' X . q' |\<in>| q \<Longrightarrow> q |\<in>| fPow X \<Longrightarrow> q' |\<in>| X"
by (simp add: fsubsetD)
obtain q' where "q' |\<in>| q"
using \<open>q \<noteq> {||}\<close> by fastforce
have "q' |\<in>| t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M |\<union>| {|FSM.initial M|}"
using scheme[OF \<open>q' |\<in>| q\<close> i5[OF \<open>q |\<in>| ?dones'\<close>]] .
then have "q' \<in> states M"
using ftransitions_source[of q' M]
using ftransitions_target[of q' M]
by force
then have "\<exists> q' p . q' |\<in>| q \<and> path M q' p \<and> p_io p = []"
using \<open>q' |\<in>| q\<close> by auto
moreover have "(\<exists> p'. pathlike ?obs q p' \<and> p_io p' = [])"
by auto
ultimately show ?case
by simp
next
case (Cons ioT ioP)
have "(\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = ioT # ioP) \<Longrightarrow> (\<exists>p'. pathlike ?obs q p' \<and> p_io p' = ioT # ioP)"
proof -
assume "\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = ioT # ioP"
then obtain q' p where "q' |\<in>| q" and "path M q' p" and "p_io p = ioT # ioP"
by meson
then obtain tM pM where "p = tM # pM"
by auto
then have "tM \<in> transitions M" and "t_source tM |\<in>| q"
using \<open>path M q' p\<close> \<open>q' |\<in>| q\<close> by blast+
then obtain tP where "tP |\<in>| ts |\<union>| qtrans"
and "t_source tP = q"
and "t_input tP = t_input tM"
and "t_output tP = t_output tM"
using Cons.prems i2 by blast
have "path M (t_target tM) pM" and "p_io pM = ioP"
using \<open>path M q' p\<close> \<open>p_io p = ioT # ioP\<close> unfolding \<open>p = tM # pM\<close> by auto
moreover have "t_target tM |\<in>| t_target tP"
using i1[OF \<open>tP |\<in>| ts |\<union>| qtrans\<close>]
using \<open>p = tM # pM\<close> \<open>path M q' p\<close> \<open>q' |\<in>| q\<close>
unfolding \<open>t_input tP = t_input tM\<close> \<open>t_output tP = t_output tM\<close> \<open>t_source tP = q\<close>
- using fmember_iff_member_fset by fastforce
+ by fastforce
ultimately have "\<exists>q' p. q' |\<in>| t_target tP \<and> path M q' p \<and> p_io p = ioP"
using \<open>p_io pM = ioP\<close> \<open>path M (t_target tM) pM\<close> by blast
have "t_target tP |\<in>| dones |\<union>| (t_target |`| ts |-| dones)"
using False \<open>tP |\<in>| ts |\<union>| qtrans\<close> by blast
moreover have "t_target tP \<noteq> {||}"
using i1[OF \<open>tP |\<in>| ts |\<union>| qtrans\<close>] by blast
ultimately obtain pP where "pathlike ?obs (t_target tP) pP" and "p_io pP = ioP"
using Cons.IH \<open>\<exists>q' p. q' |\<in>| t_target tP \<and> path M q' p \<and> p_io p = ioP\<close> by blast
then have "pathlike ?obs q (tP#pP)"
using \<open>t_source tP = q\<close> \<open>tP |\<in>| ts |\<union>| qtrans\<close> \<open>?obs = ?ts'\<close> by auto
moreover have "p_io (tP#pP) = ioT # ioP"
using \<open>t_input tP = t_input tM\<close> \<open>t_output tP = t_output tM\<close> \<open>p_io p = ioT # ioP\<close> \<open>p = tM # pM\<close> \<open>p_io pP = ioP\<close> by simp
ultimately show ?thesis
by auto
qed
moreover have "(\<exists>p'. pathlike ?obs q p' \<and> p_io p' = ioT # ioP) \<Longrightarrow> (\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = ioT # ioP)"
proof -
assume "\<exists>p'. pathlike ?obs q p' \<and> p_io p' = ioT # ioP"
then obtain p' where "pathlike ?ts' q p'" and "p_io p' = ioT # ioP"
unfolding \<open>?obs = ?ts'\<close> by meson
then obtain tP pP where "p' = tP#pP"
by auto
then have "t_source tP = q" and "tP |\<in>| ?ts'"
using \<open>pathlike ?ts' q p'\<close> by auto
have "pathlike ?ts' (t_target tP) pP" and "p_io pP = ioP"
using \<open>pathlike ?ts' q p'\<close> \<open>p_io p' = ioT # ioP\<close> \<open>p' = tP#pP\<close> by auto
then have "\<exists>p'. pathlike ?ts' (t_target tP) p' \<and> p_io p' = ioP"
by auto
moreover have "t_target tP |\<in>| dones |\<union>| (t_target |`| ts |-| dones)"
using False \<open>tP |\<in>| ts |\<union>| qtrans\<close> by blast
moreover have "t_target tP \<noteq> {||}"
using i1[OF \<open>tP |\<in>| ts |\<union>| qtrans\<close>] by blast
ultimately obtain q'' pM where "q'' |\<in>| t_target tP" and "path M q'' pM" and "p_io pM = ioP"
using Cons.IH unfolding \<open>?obs = ?ts'\<close> by blast
- have "q'' \<in> fset (t_target tP)"
- using \<open>q'' |\<in>| t_target tP\<close>
- by (meson notin_fset)
then obtain tM where "t_source tM |\<in>| q" and "tM \<in> transitions M" and "t_input tM = t_input tP" and "t_output tM = t_output tP" and "t_target tM = q''"
- using i1[OF \<open>tP |\<in>| ts |\<union>| qtrans\<close>]
+ using i1[OF \<open>tP |\<in>| ts |\<union>| qtrans\<close>]
+ using \<open>q'' |\<in>| t_target tP\<close>
unfolding \<open>t_source tP = q\<close> by force
have "path M (t_source tM) (tM#pM)"
using \<open>tM \<in> transitions M\<close> \<open>t_target tM = q''\<close> \<open>path M q'' pM\<close> by auto
moreover have "p_io (tM#pM) = ioT # ioP"
using \<open>p_io pM = ioP\<close> \<open>t_input tM = t_input tP\<close> \<open>t_output tM = t_output tP\<close> \<open>p_io p' = ioT # ioP\<close> \<open>p' = tP#pP\<close> by auto
ultimately show ?thesis
using \<open>t_source tM |\<in>| q\<close> by meson
qed
ultimately show ?case
by meson
qed
qed
then show ?thesis
using \<open>q |\<in>| dones\<close> \<open>{||} |\<notin>| dones\<close> by blast
next
case True
have "make_observable_transitions (ftransitions M) ?nexts' ?dones' ?ts' = make_observable_transitions (ftransitions M) ?nexts dones ts"
proof (cases "?nexts' = {||}")
case True
then have "?obs = ?ts'"
using qtrans_def by auto
moreover have "make_observable_transitions (ftransitions M) ?nexts' ?dones' ?ts' = ?ts'"
unfolding make_observable_transitions.simps[of "ftransitions M" ?nexts' ?dones' ?ts']
unfolding True Let_def by auto
ultimately show ?thesis
by blast
next
case False
then show ?thesis
unfolding make_observable_transitions.simps[of "ftransitions M" ?nexts dones ts] qtrans_def Let_def by auto
qed
then have IStep: "\<And> q io . q |\<in>| ?dones' \<Longrightarrow>
(\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = io) =
(\<exists>p'. pathlike (make_observable_transitions (ftransitions M) ?nexts dones ts) q p' \<and> p_io p' = io)"
using "1.hyps"[OF qtrans_def _ _ _ _ i1 i2 i4 i5 i7] True
unfolding i3
by presburger
show ?thesis
unfolding \<open>io = ioT # ioP\<close>
proof
show "\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = ioT # ioP \<Longrightarrow> \<exists>p'. pathlike ?obs q p' \<and> p_io p' = ioT # ioP"
proof -
assume "\<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = ioT # ioP"
then obtain q' p where "q' |\<in>| q" and "path M q' p" and "p_io p = ioT # ioP"
by meson
then obtain tM pM where "p = tM # pM"
by auto
then have "tM \<in> transitions M" and "t_source tM |\<in>| q"
using \<open>path M q' p\<close> \<open>q' |\<in>| q\<close> by blast+
then obtain tP where "tP |\<in>| ts"
and "t_source tP = q"
and "t_input tP = t_input tM"
and "t_output tP = t_output tM"
using "1.prems"(2,6) by blast
then have i9: "t_target tP |\<in>| dones |\<union>| ?nexts"
by simp
have "path M (t_target tM) pM" and "p_io pM = ioP"
using \<open>path M q' p\<close> \<open>p_io p = ioT # ioP\<close> unfolding \<open>p = tM # pM\<close> by auto
moreover have "t_target tM |\<in>| t_target tP"
using "1.prems"(1)[OF \<open>tP |\<in>| ts\<close>] \<open>p = tM # pM\<close> \<open>path M q' p\<close> \<open>q' |\<in>| q\<close>
unfolding \<open>t_input tP = t_input tM\<close> \<open>t_output tP = t_output tM\<close> \<open>t_source tP = q\<close>
- using fmember_iff_member_fset by fastforce
+ by fastforce
ultimately have "\<exists>q' p. q' |\<in>| t_target tP \<and> path M q' p \<and> p_io p = ioP"
using \<open>p_io pM = ioP\<close> \<open>path M (t_target tM) pM\<close> by blast
obtain pP where "pathlike ?obs (t_target tP) pP" and "p_io pP = ioP"
using \<open>\<exists>q' p. q' |\<in>| t_target tP \<and> path M q' p \<and> p_io p = ioP\<close> unfolding IStep[OF i9]
using that by blast
then have "pathlike ?obs q (tP#pP)"
using \<open>t_source tP = q\<close> \<open>tP |\<in>| ts\<close> make_observable_transitions_mono by blast
moreover have "p_io (tP#pP) = ioT # ioP"
using \<open>t_input tP = t_input tM\<close> \<open>t_output tP = t_output tM\<close> \<open>p_io p = ioT # ioP\<close> \<open>p = tM # pM\<close> \<open>p_io pP = ioP\<close> by simp
ultimately show ?thesis
by auto
qed
show "\<exists>p'. pathlike ?obs q p' \<and> p_io p' = ioT # ioP \<Longrightarrow> \<exists>q' p. q' |\<in>| q \<and> path M q' p \<and> p_io p = ioT # ioP"
proof -
assume "\<exists>p'. pathlike ?obs q p' \<and> p_io p' = ioT # ioP"
then obtain p' where "pathlike ?obs q p'" and "p_io p' = ioT # ioP"
by meson
then obtain tP pP where "p' = tP#pP"
by auto
have "\<And> t' . t' |\<in>| ftransitions M = (t' \<in> transitions M)"
using ftransitions_set
- by (metis notin_fset)
+ by metis
from \<open>p' = tP#pP\<close> have "t_source tP = q" and "tP |\<in>| ?obs"
using \<open>pathlike ?obs q p'\<close> by auto
then have "tP |\<in>| ts"
using "1.prems"(6) make_observable_transitions_t_source[of ts dones "ftransitions M"] "1.prems"(1,2)
unfolding \<open>\<And> t' . t' |\<in>| ftransitions M = (t' \<in> transitions M)\<close>
by blast
then have i9: "t_target tP |\<in>| dones |\<union>| ?nexts"
by simp
have "pathlike ?obs (t_target tP) pP" and "p_io pP = ioP"
using \<open>pathlike ?obs q p'\<close> \<open>p_io p' = ioT # ioP\<close> \<open>p' = tP#pP\<close> by auto
then have "\<exists>p'. pathlike ?obs (t_target tP) p' \<and> p_io p' = ioP"
by auto
then obtain q'' pM where "q'' |\<in>| t_target tP" and "path M q'' pM" and "p_io pM = ioP"
using IStep[OF i9] by blast
obtain tM where "t_source tM |\<in>| q" and "tM \<in> transitions M" and "t_input tM = t_input tP" and "t_output tM = t_output tP" and "t_target tM = q''"
using "1.prems"(1)[OF \<open>tP |\<in>| ts\<close>] \<open>q'' |\<in>| t_target tP\<close>
unfolding \<open>t_source tP = q\<close>
- unfolding fmember_iff_member_fset by force
+ by force
have "path M (t_source tM) (tM#pM)"
using \<open>tM \<in> transitions M\<close> \<open>t_target tM = q''\<close> \<open>path M q'' pM\<close> by auto
moreover have "p_io (tM#pM) = ioT # ioP"
using \<open>p_io pM = ioP\<close> \<open>t_input tM = t_input tP\<close> \<open>t_output tM = t_output tP\<close> \<open>p_io p' = ioT # ioP\<close> \<open>p' = tP#pP\<close> by auto
ultimately show ?thesis
using \<open>t_source tM |\<in>| q\<close> by meson
qed
qed
qed
qed
qed
fun observable_fset :: "('a,'b,'c) transition fset \<Rightarrow> bool" where
"observable_fset ts = (\<forall> t1 t2 . t1 |\<in>| ts \<longrightarrow> t2 |\<in>| ts \<longrightarrow>
t_source t1 = t_source t2 \<longrightarrow> t_input t1 = t_input t2 \<longrightarrow> t_output t1 = t_output t2
\<longrightarrow> t_target t1 = t_target t2)"
lemma make_observable_transitions_observable :
assumes "\<And> t . t |\<in>| ts \<Longrightarrow> t_source t |\<in>| dones \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
and "observable_fset ts"
shows "observable_fset (make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts)"
using assms proof (induction base_trans "(fimage t_target ts) |-| dones" dones ts rule: make_observable_transitions.induct)
case (1 base_trans dones ts)
let ?nexts = "(fimage t_target ts) |-| dones"
define qtrans where qtrans_def: "qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) base_trans;
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)) ?nexts)"
have qtrans_prop: "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| ?nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using make_observable_transitions_qtrans_helper[OF qtrans_def]
by presburger
let ?dones' = "dones |\<union>| ?nexts"
let ?ts' = "ts |\<union>| qtrans"
let ?nexts' = "(fimage t_target qtrans) |-| ?dones'"
have "observable_fset qtrans"
using qtrans_prop
unfolding observable_fset.simps
by (metis (mono_tags, lifting) Collect_cong fset_inject)
moreover have "t_source |`| qtrans |\<inter>| t_source |`| ts = {||}"
using "1.prems"(1) qtrans_prop by force
ultimately have "observable_fset ?ts'"
using "1.prems"(2) unfolding observable_fset.simps
by blast
have res_cases: "make_observable_transitions base_trans ?nexts dones ts = (if ?nexts' = {||}
then ?ts'
else make_observable_transitions base_trans ?nexts' ?dones' ?ts')"
unfolding make_observable_transitions.simps[of base_trans ?nexts dones ts] qtrans_def Let_def by simp
show ?case proof (cases "?nexts' = {||}")
case True
then have "make_observable_transitions base_trans ?nexts dones ts = ?ts'"
using res_cases by simp
then show ?thesis
using \<open>observable_fset ?ts'\<close> by simp
next
case False
then have *: "make_observable_transitions base_trans ?nexts dones ts = make_observable_transitions base_trans ?nexts' ?dones' ?ts'"
using res_cases by simp
have i1: "(\<And>t. t |\<in>| ts |\<union>| qtrans \<Longrightarrow>
t_source t |\<in>| dones |\<union>| ?nexts \<and>
t_target t \<noteq> {||} \<and>
fset (t_target t) =
t_target `
{t' . t' |\<in>| base_trans \<and>
t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t})"
using "1.prems"(1) qtrans_prop by blast
have i3: "t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones)) = t_target |`| qtrans |-| (dones |\<union>| (t_target |`| ts |-| dones))"
by auto
have i4: "t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones)) \<noteq> {||}"
using False by auto
show ?thesis
using "1.hyps"[OF qtrans_def _ _ i3 i4 i1 \<open>observable_fset ?ts'\<close>] unfolding * i3 by metis
qed
qed
lemma make_observable_transitions_transition_props :
assumes "\<And> t . t |\<in>| ts \<Longrightarrow> t_source t |\<in>| dones \<and> t_target t |\<in>| dones |\<union>| ((fimage t_target ts) |-| dones) \<and> t_input t |\<in>| t_input |`| base_trans \<and> t_output t |\<in>| t_output |`| base_trans"
assumes "t |\<in>| make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts"
shows "t_source t |\<in>| dones |\<union>| (t_target |`| (make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts))"
and "t_target t |\<in>| dones |\<union>| (t_target |`| (make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts))"
and "t_input t |\<in>| t_input |`| base_trans"
and "t_output t |\<in>| t_output |`| base_trans"
proof -
have "t_source t |\<in>| dones |\<union>| (t_target |`| (make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts))
\<and> t_target t |\<in>| dones |\<union>| (t_target |`| (make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts))
\<and> t_input t |\<in>| t_input |`| base_trans
\<and> t_output t |\<in>| t_output |`| base_trans"
using assms(1,2)
proof (induction base_trans "((fimage t_target ts) |-| dones)" dones ts rule: make_observable_transitions.induct)
case (1 base_trans dones ts)
let ?nexts = "((fimage t_target ts) |-| dones)"
define qtrans where qtrans_def: "qtrans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) base_trans;
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)) ?nexts)"
have qtrans_prop: "\<And> t . t |\<in>| qtrans \<longleftrightarrow> t_source t |\<in>| ?nexts \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' . t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using make_observable_transitions_qtrans_helper[OF qtrans_def]
by presburger
let ?dones' = "dones |\<union>| ?nexts"
let ?ts' = "ts |\<union>| qtrans"
let ?nexts' = "(fimage t_target qtrans) |-| ?dones'"
have res_cases: "make_observable_transitions base_trans ?nexts dones ts = (if ?nexts' = {||}
then ?ts'
else make_observable_transitions base_trans ?nexts' ?dones' ?ts')"
unfolding make_observable_transitions.simps[of base_trans ?nexts dones ts] qtrans_def Let_def by simp
have qtrans_trans_prop: "(\<And>t. t |\<in>| qtrans \<Longrightarrow>
t_source t |\<in>| dones |\<union>| (t_target |`| ts |-| dones) \<and>
t_target t |\<in>| dones |\<union>| (t_target |`| ts |-| dones) |\<union>| (t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones))) \<and>
t_input t |\<in>| t_input |`| base_trans \<and> t_output t |\<in>| t_output |`| base_trans)" (is "\<And> t . t |\<in>| qtrans \<Longrightarrow> ?P t")
proof -
fix t assume "t |\<in>| qtrans"
then have "t_source t |\<in>| dones |\<union>| (t_target |`| ts |-| dones)"
using \<open>t |\<in>| qtrans\<close> unfolding qtrans_prop[of t] by blast
moreover have "t_target t |\<in>| dones |\<union>| (t_target |`| ts |-| dones) |\<union>| (t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones)))"
using \<open>t |\<in>| qtrans\<close> "1.prems"(1) by blast
moreover have "t_input t |\<in>| t_input |`| base_trans \<and> t_output t |\<in>| t_output |`| base_trans"
proof -
obtain t' where "t' \<in> {t'. t' |\<in>| base_trans \<and> t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using \<open>t |\<in>| qtrans\<close> unfolding qtrans_prop[of t]
by (metis (mono_tags, lifting) Collect_empty_eq bot_fset.rep_eq empty_is_image fset_inject mem_Collect_eq)
then show ?thesis
by force
qed
ultimately show "?P t"
by blast
qed
show ?case proof (cases "?nexts' = {||}")
case True
then have "t |\<in>| ?ts'"
using "1.prems"(2) res_cases by force
then show ?thesis
using "1.prems"(1) qtrans_trans_prop
by (metis True fimage_funion funion_fminus_cancel funion_iff res_cases)
next
case False
then have *: "make_observable_transitions base_trans ?nexts dones ts = make_observable_transitions base_trans ?nexts' ?dones' ?ts'"
using res_cases by simp
have i1: "t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones)) = t_target |`| qtrans |-| (dones |\<union>| (t_target |`| ts |-| dones))"
by blast
have i2: "t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones)) \<noteq> {||}"
using False by blast
have i3: "(\<And>t. t |\<in>| ts |\<union>| qtrans \<Longrightarrow>
t_source t |\<in>| dones |\<union>| (t_target |`| ts |-| dones) \<and>
t_target t |\<in>| dones |\<union>| (t_target |`| ts |-| dones) |\<union>| (t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones))) \<and>
t_input t |\<in>| t_input |`| base_trans \<and> t_output t |\<in>| t_output |`| base_trans)"
using "1.prems"(1) qtrans_trans_prop by blast
have i4: "t |\<in>| make_observable_transitions base_trans (t_target |`| (ts |\<union>| qtrans) |-| (dones |\<union>| (t_target |`| ts |-| dones))) (dones |\<union>| (t_target |`| ts |-| dones)) (ts |\<union>| qtrans)"
using "1.prems"(2) unfolding * i1 by assumption
show ?thesis
using "1.hyps"[OF qtrans_def _ _ i1 i2 i3 i4] unfolding i1 unfolding *[symmetric]
using make_observable_transitions_mono[of ts base_trans ?nexts dones] by blast
qed
qed
then show "t_source t |\<in>| dones |\<union>| (t_target |`| (make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts))"
and "t_target t |\<in>| dones |\<union>| (t_target |`| (make_observable_transitions base_trans ((fimage t_target ts) |-| dones) dones ts))"
and "t_input t |\<in>| t_input |`| base_trans"
and "t_output t |\<in>| t_output |`| base_trans"
by blast+
qed
fun make_observable :: "('a :: linorder,'b :: linorder,'c :: linorder) fsm \<Rightarrow> ('a fset,'b,'c) fsm" where
"make_observable M = (let
initial_trans = (let qts = ffilter (\<lambda>t . t_source t = initial M) (ftransitions M);
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . ({|initial M|},x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios);
nexts = fimage t_target initial_trans |-| {|{|initial M|}|};
ptransitions = make_observable_transitions (ftransitions M) nexts {|{|initial M|}|} initial_trans;
pstates = finsert {|initial M|} (t_target |`| ptransitions);
M' = create_unconnected_fsm_from_fsets {|initial M|} pstates (finputs M) (foutputs M)
in add_transitions M' (fset ptransitions))"
lemma make_observable_language_observable :
shows "L (make_observable M) = L M"
and "observable (make_observable M)"
and "initial (make_observable M) = {|initial M|}"
and "inputs (make_observable M) = inputs M"
and "outputs (make_observable M) = outputs M"
proof -
define initial_trans where "initial_trans = (let qts = ffilter (\<lambda>t . t_source t = initial M) (ftransitions M);
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . ({|initial M|},x,y, t_target |`| ((ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts)))) ios)"
moreover define ptransitions where "ptransitions = make_observable_transitions (ftransitions M) (fimage t_target initial_trans |-| {|{|initial M|}|}) {|{|initial M|}|} initial_trans"
moreover define pstates where "pstates = finsert {|initial M|} (t_target |`| ptransitions)"
moreover define M' where "M' = create_unconnected_fsm_from_fsets {|initial M|} pstates (finputs M) (foutputs M)"
ultimately have "make_observable M = add_transitions M' (fset ptransitions)"
unfolding make_observable.simps Let_def by blast
have "{|initial M|} |\<in>| pstates"
unfolding pstates_def by blast
have "inputs M' = inputs M"
unfolding M'_def create_unconnected_fsm_from_fsets_simps(3)[OF \<open>{|initial M|} |\<in>| pstates\<close>, of "finputs M" "foutputs M"]
using fset_of_list.rep_eq inputs_as_list_set by fastforce
have "outputs M' = outputs M"
unfolding M'_def create_unconnected_fsm_from_fsets_simps(4)[OF \<open>{|initial M|} |\<in>| pstates\<close>, of "finputs M" "foutputs M"]
using fset_of_list.rep_eq outputs_as_list_set by fastforce
have "states M' = fset pstates" and "transitions M' = {}" and "initial M' = {|initial M|}"
unfolding M'_def create_unconnected_fsm_from_fsets_simps(1,2,5)[OF \<open>{|initial M|} |\<in>| pstates\<close>] by simp+
have initial_trans_prop: "\<And> t . t |\<in>| initial_trans \<longleftrightarrow> t_source t |\<in>| {|{|FSM.initial M|}|} \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' \<in> transitions M . t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
proof -
have *:"\<And> t' . t' |\<in>| ftransitions M = (t' \<in> transitions M)"
using ftransitions_set
- by (metis notin_fset)
+ by metis
have **: "initial_trans = ffUnion (fimage (\<lambda> q . (let qts = ffilter (\<lambda>t . t_source t |\<in>| q) (ftransitions M);
ios = fimage (\<lambda> t . (t_input t, t_output t)) qts
in fimage (\<lambda>(x,y) . (q,x,y, t_target |`| (ffilter (\<lambda>t . (t_input t, t_output t) = (x,y)) qts))) ios)) {|{|initial M|}|})"
unfolding initial_trans_def by auto
show "\<And> t . t |\<in>| initial_trans \<longleftrightarrow> t_source t |\<in>| {|{|FSM.initial M|}|} \<and> t_target t \<noteq> {||} \<and> fset (t_target t) = t_target ` {t' \<in> transitions M . t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
unfolding make_observable_transitions_qtrans_helper[OF **] *
by presburger
qed
have well_formed_transitions: "\<And> t . t \<in> (fset ptransitions) \<Longrightarrow> t_source t \<in> states M' \<and> t_input t \<in> inputs M' \<and> t_output t \<in> outputs M' \<and> t_target t \<in> states M'"
(is "\<And> t . t \<in> (fset ptransitions) \<Longrightarrow> ?P1 t \<and> ?P2 t \<and> ?P3 t \<and> ?P4 t")
proof -
fix t assume "t \<in> (fset ptransitions)"
then have i2: "t |\<in>| make_observable_transitions (ftransitions M) (fimage t_target initial_trans |-| {|{|initial M|}|}) {|{|initial M|}|} initial_trans"
using ptransitions_def
- by (meson notin_fset)
+ by metis
have i1: "(\<And>t. t |\<in>| initial_trans \<Longrightarrow>
t_source t |\<in>| {|{|FSM.initial M|}|} \<and>
t_target t |\<in>| {|{|FSM.initial M|}|} |\<union>| (t_target |`| initial_trans |-| {|{|FSM.initial M|}|}) \<and>
t_input t |\<in>| t_input |`| ftransitions M \<and> t_output t |\<in>| t_output |`| ftransitions M)" (is "\<And> t . t |\<in>| initial_trans \<Longrightarrow> ?P t")
proof -
fix t assume *: "t |\<in>| initial_trans"
then have "t_source t |\<in>| {|{|FSM.initial M|}|}"
and "t_target t \<noteq> {||}"
and "fset (t_target t) = t_target ` {t' \<in> FSM.transitions M. t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}"
using initial_trans_prop by blast+
have "t_target t |\<in>| {|{|FSM.initial M|}|} |\<union>| (t_target |`| initial_trans |-| {|{|FSM.initial M|}|})"
using "*" by blast
moreover have "t_input t |\<in>| t_input |`| ftransitions M \<and> t_output t |\<in>| t_output |`| ftransitions M"
proof -
obtain t' where "t' \<in> transitions M" and "t_input t = t_input t'" and "t_output t = t_output t'"
using \<open>t_target t \<noteq> {||}\<close> \<open>fset (t_target t) = t_target ` {t' \<in> FSM.transitions M. t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t}\<close>
by (metis (mono_tags, lifting) bot_fset.rep_eq empty_Collect_eq fset_inject image_empty)
have "fset (ftransitions M) = transitions M"
by (simp add: fset_of_list.rep_eq fsm_transitions_finite)
- then have "t' |\<in>| ftransitions M"
- using \<open>t' \<in> transitions M\<close> notin_fset by fastforce
then show ?thesis
- unfolding \<open>t_input t = t_input t'\<close> \<open>t_output t = t_output t'\<close> by auto
+ unfolding \<open>t_input t = t_input t'\<close> \<open>t_output t = t_output t'\<close>
+ using \<open>t' \<in> transitions M\<close>
+ by auto
qed
ultimately show "?P t"
using \<open>t_source t |\<in>| {|{|FSM.initial M|}|}\<close> by blast
qed
have "?P1 t"
using make_observable_transitions_transition_props(1)[OF i1 i2] unfolding pstates_def ptransitions_def \<open>states M' = fset pstates\<close>
- by (metis finsert_is_funion fmember_iff_member_fset)
+ by (metis finsert_is_funion)
moreover have "?P2 t"
proof-
have "t_input t |\<in>| t_input |`| ftransitions M"
using make_observable_transitions_transition_props(3)[OF i1 i2] by blast
then have "t_input t \<in> t_input ` transitions M"
- using ftransitions_set by (metis (mono_tags, lifting) fmember_iff_member_fset fset.set_map)
+ using ftransitions_set by (metis (mono_tags, lifting) fset.set_map)
then show ?thesis
using finputs_set fsm_transition_input \<open>inputs M' = inputs M\<close> by fastforce
qed
moreover have "?P3 t"
proof-
have "t_output t |\<in>| t_output |`| ftransitions M"
using make_observable_transitions_transition_props(4)[OF i1 i2] by blast
then have "t_output t \<in> t_output ` transitions M"
- using ftransitions_set by (metis (mono_tags, lifting) fmember_iff_member_fset fset.set_map)
+ using ftransitions_set by (metis (mono_tags, lifting) fset.set_map)
then show ?thesis
using foutputs_set fsm_transition_output \<open>outputs M' = outputs M\<close> by fastforce
qed
moreover have "?P4 t"
using make_observable_transitions_transition_props(2)[OF i1 i2] unfolding pstates_def ptransitions_def \<open>states M' = fset pstates\<close>
- by (metis finsert_is_funion fmember_iff_member_fset)
+ by (metis finsert_is_funion)
ultimately show "?P1 t \<and> ?P2 t \<and> ?P3 t \<and> ?P4 t"
by blast
qed
have "initial (make_observable M) = {|initial M|}"
and "states (make_observable M) = fset pstates"
and "inputs (make_observable M) = inputs M"
and "outputs (make_observable M) = outputs M"
and "transitions (make_observable M) = fset ptransitions"
using add_transitions_simps[OF well_formed_transitions, of "fset ptransitions"]
unfolding \<open>make_observable M = add_transitions M' (fset ptransitions)\<close>[symmetric]
\<open>inputs M' = inputs M\<close> \<open>outputs M' = outputs M\<close> \<open>initial M' = {|initial M|}\<close> \<open>states M' = fset pstates\<close> \<open>transitions M' = {}\<close>
by blast+
then show "initial (make_observable M) = {|initial M|}" and "inputs (make_observable M) = inputs M" and "outputs (make_observable M) = outputs M"
by presburger+
have i1: "(\<And>t. t |\<in>| initial_trans \<Longrightarrow>
t_source t |\<in>| {|{|FSM.initial M|}|} \<and>
t_target t \<noteq> {||} \<and>
fset (t_target t) = t_target ` {t' \<in> FSM.transitions M. t_source t' |\<in>| t_source t \<and> t_input t' = t_input t \<and> t_output t' = t_output t})"
using initial_trans_prop by blast
have i2: "(\<And>q t'.
q |\<in>| {|{|FSM.initial M|}|} \<Longrightarrow>
t' \<in> FSM.transitions M \<Longrightarrow> t_source t' |\<in>| q \<Longrightarrow> \<exists>t. t |\<in>| initial_trans \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t')"
proof -
fix q t' assume "q |\<in>| {|{|FSM.initial M|}|}" and "t' \<in> FSM.transitions M" and "t_source t' |\<in>| q"
then have "q = {|FSM.initial M|}" and "t_source t' = initial M"
by auto
define tgt where "tgt = t_target ` {t'' \<in> FSM.transitions M. t_source t'' |\<in>| {|FSM.initial M|} \<and> t_input t'' = t_input t' \<and> t_output t'' = t_output t'}"
have "t_target t' \<in> tgt"
unfolding tgt_def using \<open>t' \<in> FSM.transitions M\<close> \<open>t_source t' = initial M\<close> by auto
then have "tgt \<noteq> {}"
by auto
have "finite tgt"
using fsm_transitions_finite[of M] unfolding tgt_def by auto
then have "fset (Abs_fset tgt) = tgt"
by (simp add: Abs_fset_inverse)
then have "Abs_fset tgt \<noteq> {||}"
using \<open>tgt \<noteq> {}\<close> by auto
let ?t = "({|FSM.initial M|}, t_input t', t_output t', Abs_fset tgt)"
have "?t |\<in>| initial_trans"
unfolding initial_trans_prop fst_conv snd_conv \<open>fset (Abs_fset tgt) = tgt\<close>
unfolding \<open>tgt = t_target ` {t'' \<in> FSM.transitions M. t_source t'' |\<in>| {|FSM.initial M|} \<and> t_input t'' = t_input t' \<and> t_output t'' = t_output t'}\<close>[symmetric]
using \<open>Abs_fset tgt \<noteq> {||}\<close>
by blast
then show "\<exists>t. t |\<in>| initial_trans \<and> t_source t = q \<and> t_input t = t_input t' \<and> t_output t = t_output t'"
using \<open>q = {|FSM.initial M|}\<close> by auto
qed
have i3: "(\<And>q. q |\<in>| t_target |`| initial_trans |-| {|{|FSM.initial M|}|} \<Longrightarrow> q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M))"
proof -
fix q assume "q |\<in>| t_target |`| initial_trans |-| {|{|FSM.initial M|}|}"
then obtain t where "t |\<in>| initial_trans" and "t_target t = q"
by auto
have "fset q \<subseteq> t_target ` (transitions M)"
using \<open>t |\<in>| initial_trans\<close>
unfolding initial_trans_prop \<open>t_target t = q\<close>
by auto
then have "q |\<subseteq>| (t_target |`| ftransitions M)"
using ftransitions_set[of M]
by (simp add: less_eq_fset.rep_eq)
then show "q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M)"
by auto
qed
have i4: "(\<And>q. q |\<in>| {|{|FSM.initial M|}|} \<Longrightarrow> q |\<in>| fPow (t_source |`| ftransitions M |\<union>| t_target |`| ftransitions M |\<union>| {|FSM.initial M|}))"
and i5: "{||} |\<notin>| {|{|FSM.initial M|}|}"
and i6: "{|FSM.initial M|} |\<in>| {|{|FSM.initial M|}|}"
by blast+
show "L (make_observable M) = L M"
proof -
have *: "\<And> p . pathlike ptransitions {|initial M|} p = path (make_observable M) {|initial M|} p"
proof
have "\<And> q p . p \<noteq> [] \<Longrightarrow> pathlike ptransitions q p \<Longrightarrow> path (make_observable M) q p"
proof -
fix q p assume "p \<noteq> []" and "pathlike ptransitions q p"
then show "path (make_observable M) q p"
proof (induction p arbitrary: q)
case Nil
then show ?case by blast
next
case (Cons t p)
then have "t |\<in>| ptransitions" and "pathlike ptransitions (t_target t) p" and "t_source t = q"
by blast+
have "t \<in> transitions (make_observable M)"
using \<open>t |\<in>| ptransitions\<close> \<open>transitions (make_observable M) = fset ptransitions\<close>
- by (metis notin_fset)
+ by metis
moreover have "path (make_observable M) (t_target t) p"
using Cons.IH[OF _ \<open>pathlike ptransitions (t_target t) p\<close>] calculation by blast
ultimately show ?case
using \<open>t_source t = q\<close> by blast
qed
qed
then show "\<And> p . pathlike ptransitions {|initial M|} p \<Longrightarrow> path (make_observable M) {|initial M|} p"
by (metis \<open>FSM.initial (make_observable M) = {|FSM.initial M|}\<close> fsm_initial path.nil)
show "\<And> q p . path (make_observable M) q p \<Longrightarrow> pathlike ptransitions q p"
proof -
fix q p assume "path (make_observable M) q p"
then show "pathlike ptransitions q p"
proof (induction p arbitrary: q rule: list.induct)
case Nil
then show ?case by blast
next
case (Cons t p)
then have "t \<in> transitions (make_observable M)" and "path (make_observable M) (t_target t) p" and "t_source t = q"
by blast+
have "t |\<in>| ptransitions"
using \<open>t \<in> transitions (make_observable M)\<close> \<open>transitions (make_observable M) = fset ptransitions\<close>
- by (metis notin_fset)
+ by metis
then show ?case
using Cons.IH[OF \<open>path (make_observable M) (t_target t) p\<close>] \<open>t_source t = q\<close> by blast
qed
qed
qed
have "\<And> io . (\<exists>q' p. q' |\<in>| {|FSM.initial M|} \<and> path M q' p \<and> p_io p = io) = (\<exists>p'. pathlike ptransitions {|FSM.initial M|} p' \<and> p_io p' = io)"
using make_observable_transitions_path[OF i1 i2 i3 i4 i5 i6] unfolding ptransitions_def[symmetric] by blast
then have "\<And> io . (\<exists>p. path M (FSM.initial M) p \<and> p_io p = io) = (\<exists>p' . path (make_observable M) {|initial M|} p' \<and> p_io p' = io)"
unfolding *
by (metis (no_types, lifting) fempty_iff finsert_iff)
then show ?thesis
unfolding LS.simps \<open>initial (make_observable M) = {|initial M|}\<close> by (metis (no_types, lifting))
qed
show "observable (make_observable M)"
proof -
have i2: "observable_fset initial_trans"
unfolding observable_fset.simps
unfolding initial_trans_prop
using fset_inject
by metis
have "\<And> t' . t' |\<in>| ftransitions M = (t' \<in> transitions M)"
using ftransitions_set
- by (metis notin_fset)
+ by metis
have "observable_fset ptransitions"
using make_observable_transitions_observable[OF _ i2, of "{| {|initial M|} |}" "ftransitions M"] i1
unfolding ptransitions_def \<open>\<And> t' . t' |\<in>| ftransitions M = (t' \<in> transitions M)\<close>
by blast
then show ?thesis
unfolding observable.simps observable_fset.simps \<open>transitions (make_observable M) = fset ptransitions\<close>
- by (meson notin_fset)
+ by metis
qed
qed
end
\ No newline at end of file
diff --git a/thys/Factored_Transition_System_Bounding/FactoredSystem.thy b/thys/Factored_Transition_System_Bounding/FactoredSystem.thy
--- a/thys/Factored_Transition_System_Bounding/FactoredSystem.thy
+++ b/thys/Factored_Transition_System_Bounding/FactoredSystem.thy
@@ -1,3595 +1,3595 @@
theory FactoredSystem
imports Main "HOL-Library.Finite_Map" "HOL-Library.Sublist" FSSublist
FactoredSystemLib ListUtils HoArithUtils FmapUtils
begin
section "Factored System"
\<comment> \<open>NOTE hide the '++' operator from 'Map' to prevent warnings.\<close>
hide_const (open) Map.map_add
no_notation Map.map_add (infixl "++" 100)
subsection "Semantics of Plan Execution"
text \<open> This section aims at characterizing the semantics of executing plans---i.e. sequences
of actions---on a given initial state.
The semantics of action execution were previously introduced
via the notion of succeding state (`state\_succ`). Plan execution (`exec\_plan`) extends this notion
to sequences of actions by calculating the succeding state from the given state and action pair and
then recursively executing the remaining actions on the succeding state. [Abdulaziz et al., HOL4
Definition 3, p.9] \<close>
lemma state_succ_pair: "state_succ s (p, e) = (if (p \<subseteq>\<^sub>f s) then (e ++ s) else s)"
by (simp add: state_succ_def)
\<comment> \<open>NOTE shortened to 'exec\_plan'\<close>
\<comment> \<open>NOTE using 'fun' because of multiple definining equations.\<close>
\<comment> \<open>NOTE first argument was curried.\<close>
fun exec_plan where
"exec_plan s [] = s"
| "exec_plan s (a # as) = exec_plan (state_succ s a) as"
lemma exec_plan_Append:
fixes as_a as_b s
shows "exec_plan s (as_a @ as_b) = exec_plan (exec_plan s as_a) as_b"
by (induction as_a arbitrary: s as_b) auto
text \<open> Plan execution effectively eliminates cycles: i.e., if a given plan `as` may be partitioned
into plans `as1`, `as2` and `as3`, s.t. the sequential execution of `as1` and `as2` yields the same
state, `as2` may be skipped during plan execution. \<close>
lemma cycle_removal_lemma:
fixes as1 as2 as3
assumes "(exec_plan s (as1 @ as2) = exec_plan s as1)"
shows "(exec_plan s (as1 @ as2 @ as3) = exec_plan s (as1 @ as3))"
using assms exec_plan_Append
by metis
subsubsection "Characterization of the Set of Possible States"
text \<open> To show the construction principle of the set of possible states---in lemma
`construction\_of\_all\_possible\_states\_lemma`---the following ancillary proves of finite map
properties are required.
Most importantly, in lemma `fmupd\_fmrestrict\_subset` we show how finite mappings `s` with domain
@{term "{v} \<union> X"} and `s v = (Some x)` are constructed from their restrictions to `X` via update, i.e.
s = fmupd v x (fmrestrict\_set X s)
This is used in lemma `construction\_of\_all\_possible\_states\_lemma` to show that the set of possible
states for variables @{term "{v} \<union> X"} is constructed inductively from the set of all possible states for
variables `X` via update on point @{term "v \<notin> X"}.
\<close>
\<comment> \<open>NOTE added lemma.\<close>
lemma empty_domain_fmap_set: "{s. fmdom' s = {}} = {fmempty}"
proof -
let ?A = "{s. fmdom' s = {}}"
let ?B = "{fmempty}"
fix s
show ?thesis proof(rule ccontr)
assume C: "?A \<noteq> ?B"
then show False proof -
{
assume C1: "?A \<subset> ?B"
have "?A = {}" using C1 by force
then have False using fmdom'_empty by blast
}
moreover
{
assume C2: "\<not>(?A \<subset> ?B)"
then have "fmdom' fmempty = {}"
by auto
moreover have "fmempty \<in> ?A"
by auto
moreover have "?A \<noteq> {}"
using calculation(2) by blast
moreover have "\<forall>a\<in>?A.a\<notin>?B"
by (metis (mono_tags, lifting)
C Collect_cong calculation(1) fmrestrict_set_dom fmrestrict_set_null singleton_conv)
moreover have "fmempty \<in> ?B" by auto
moreover have "\<exists>a\<in>?A.a\<in>?B"
by simp
moreover have "\<not>(\<forall>a\<in>?A.a\<notin>?B)"
by simp
ultimately have False
by blast
}
ultimately show False
by fastforce
qed
qed
qed
\<comment> \<open>NOTE added lemma.\<close>
lemma possible_states_set_ii_a:
fixes s x v
assumes "(v \<in> fmdom' s)"
shows "(fmdom' ((\<lambda>s. fmupd v x s) s) = fmdom' s)"
using assms insert_absorb
by auto
\<comment> \<open>NOTE added lemma.\<close>
lemma possible_states_set_ii_b:
fixes s x v
assumes "(v \<notin> fmdom' s)"
shows "(fmdom' ((\<lambda>s. fmupd v x s) s) = fmdom' s \<union> {v})"
by auto
\<comment> \<open>NOTE added lemma.\<close>
lemma fmap_neq:
fixes s :: "('a, bool) fmap" and s' :: "('a, bool) fmap"
assumes "(fmdom' s = fmdom' s')"
shows "((s \<noteq> s') \<longleftrightarrow> (\<exists>v\<in>(fmdom' s). fmlookup s v \<noteq> fmlookup s' v))"
using assms fmap_ext fmdom'_notD
by metis
\<comment> \<open>NOTE added lemma.\<close>
lemma fmdom'_fmsubset_restrict_set:
fixes X1 X2 and s :: "('a, bool) fmap"
assumes "X1 \<subseteq> X2" "fmdom' s = X2"
shows "fmdom' (fmrestrict_set X1 s) = X1"
using assms
by (metis (no_types, lifting)
antisym_conv fmdom'_notD fmdom'_notI fmlookup_restrict_set rev_subsetD subsetI)
\<comment> \<open>NOTE added lemma.\<close>
lemma fmsubset_restrict_set:
fixes X1 X2 and s :: "'a state"
assumes "X1 \<subseteq> X2" "s \<in> {s. fmdom' s = X2}"
shows "fmrestrict_set X1 s \<in> {s. fmdom' s = X1}"
using assms fmdom'_fmsubset_restrict_set
by blast
\<comment> \<open>NOTE added lemma.\<close>
lemma fmupd_fmsubset_restrict_set:
fixes X v x and s :: "'a state"
assumes "s \<in> {s. fmdom' s = insert v X}" "fmlookup s v = Some x"
shows "s = fmupd v x (fmrestrict_set X s)"
proof -
\<comment> \<open>Show that domains of 's' and 'fmupd v x (fmrestrict\_set X s)' are identical.\<close>
have 1: "fmdom' s = insert v X"
using assms(1)
by simp
{
have "X \<subseteq> insert v X"
by auto
then have "fmdom' (fmrestrict_set X s) = X"
using 1 fmdom'_fmsubset_restrict_set
by metis
then have "fmdom' (fmupd v x (fmrestrict_set X s)) = insert v X"
using assms(1) fmdom'_fmupd
by auto
}
note 2 = this
moreover
{
fix w
\<comment> \<open>Show case for undefined variables (where lookup yields 'None').\<close>
{
assume "w \<notin> insert v X"
then have "w \<notin> fmdom' s" "w \<notin> fmdom' (fmupd v x (fmrestrict_set X s))"
using 1 2
by argo+
then have "fmlookup s w = fmlookup (fmupd v x (fmrestrict_set X s)) w"
using fmdom'_notD
by metis
}
\<comment> \<open>Show case for defined variables (where lookup yields 'Some y').\<close>
moreover {
assume "w \<in> insert v X"
then have "w \<in> fmdom' s" "w \<in> fmdom' (fmupd v x (fmrestrict_set X s))"
using 1 2
by argo+
then have "fmlookup s w = fmlookup (fmupd v x (fmrestrict_set X s)) w"
by (cases "w = v")
(auto simp add: assms calculation)
}
ultimately have "fmlookup s w = fmlookup (fmupd v x (fmrestrict_set X s)) w"
by blast
}
then show ?thesis
using fmap_ext
by blast
qed
lemma construction_of_all_possible_states_lemma:
fixes v X
assumes "(v \<notin> X)"
shows "({s. fmdom' s = insert v X}
= ((\<lambda>s. fmupd v True s) ` {s. fmdom' s = X})
\<union> ((\<lambda>s. fmupd v False s) ` {s. fmdom' s = X})
)"
proof -
fix v X
let ?A = "{s :: 'a state. fmdom' s = insert v X}"
let ?B = "
((\<lambda>s. fmupd v True s) ` {s :: 'a state. fmdom' s = X})
\<union> ((\<lambda>s. fmupd v False s) ` {s :: 'a state. fmdom' s = X})
"
text \<open>Show the goal by mutual inclusion. The inclusion @{term "?B \<subseteq> ?A"} is trivial and can be solved by
automation. For the complimentary proof @{term "?A \<subseteq> ?B"} however we need to do more work.
In our case we choose a proof by contradiction and show that an @{term "s \<in> ?A"} which is not also in
'?B' cannot exist.\<close>
{
have "?A \<subseteq> ?B" proof(rule ccontr)
assume C: "\<not>(?A \<subseteq> ?B)"
moreover have "\<exists>s\<in>?A. s\<notin>?B"
using C
by auto
moreover obtain s where obtain_s: "s\<in>?A \<and> s\<notin>?B"
using calculation
by auto
moreover have "s\<notin>?B"
using obtain_s
by auto
moreover have "fmdom' s = X \<union> {v}"
using obtain_s
by auto
moreover have "\<forall>s'\<in>?B. fmdom' s' = X \<union> {v}"
by auto
moreover have
"(s \<notin> ((\<lambda>s. fmupd v True s) ` {s. fmdom' s = X}))"
"(s \<notin> ((\<lambda>s. fmupd v False s) ` {s. fmdom' s = X}))"
using obtain_s
by blast+
text \<open> Show that every state @{term "s \<in> ?A"} has been constructed from another state with domain
'X'. \<close>
moreover
{
fix s :: "'a state"
assume 1: "s \<in> {s :: 'a state. fmdom' s = insert v X}"
then have "fmrestrict_set X s \<in> {s :: 'a state. fmdom' s = X}"
using subset_insertI fmsubset_restrict_set
by metis
moreover
{
assume "fmlookup s v = Some True"
then have "s = fmupd v True (fmrestrict_set X s)"
using 1 fmupd_fmsubset_restrict_set
by metis
}
moreover {
assume "fmlookup s v = Some False"
then have "s = fmupd v False (fmrestrict_set X s)"
using 1 fmupd_fmsubset_restrict_set
by fastforce
}
moreover have "fmlookup s v \<noteq> None"
using 1 fmdom'_notI
by fastforce
ultimately have "
(s \<in> ((\<lambda>s. fmupd v True s) ` {s. fmdom' s = X}))
\<or> (s \<in> ((\<lambda>s. fmupd v False s) ` {s. fmdom' s = X}))
"
by force
}
ultimately show False
by meson
qed
}
moreover have "?B \<subseteq> ?A"
by force
ultimately show "?A = ?B" by blast
qed
text \<open> Another important property of the state set is cardinality, i.e. the number of distinct
states which can be modelled using a given finite variable set.
As lemma `card\_of\_set\_of\_all\_possible\_states` shows, for a finite variable set `X`, the number of
possible states is `2 \^ card X`, i.e. the number of assigning two discrete values to `card X` slots
as known from combinatorics.
Again, some additional properties of finite maps had to be proven. Pivotally, in lemma
`updates\_disjoint`, it is shown that the image of updating a set of states with domain `X` on a
point @{term "x \<notin> X"} with either `True` or `False` yields two distinct sets of states with domain
@{term "{x} \<union> X"}. \<close>
\<comment> \<open>NOTE goal has to stay implication otherwise induction rule won't watch.\<close>
lemma FINITE_states:
fixes X :: "'a set"
shows "finite X \<Longrightarrow> finite {(s :: 'a state). fmdom' s = X}"
proof (induction rule: finite.induct)
case emptyI
then have "{s. fmdom' s = {}} = {fmempty}"
by (simp add: empty_domain_fmap_set)
then show ?case
by (simp add: \<open>{s. fmdom' s = {}} = {fmempty}\<close>)
next
case (insertI A a)
assume P1: "finite A"
and P2: "finite {s. fmdom' s = A}"
then show ?case
proof (cases "a \<in> A")
case True
then show ?thesis
using insertI.IH insert_Diff
by fastforce
next
case False
then show ?thesis
proof -
have "finite (
((\<lambda>s. fmupd a True s) ` {s. fmdom' s = A})
\<union> ((\<lambda>s. fmupd a False s) ` {s. fmdom' s = A}))"
using False construction_of_all_possible_states_lemma insertI.IH
by blast
then show ?thesis
using False construction_of_all_possible_states_lemma
by fastforce
qed
qed
qed
\<comment> \<open>NOTE added lemma.\<close>
lemma bool_update_effect:
fixes s X x v b
assumes "finite X" "s \<in> {s :: 'a state. fmdom' s = X}" "x \<in> X" "x \<noteq> v"
shows "fmlookup ((\<lambda>s :: 'a state. fmupd v b s) s) x = fmlookup s x"
using assms fmupd_lookup
by auto
\<comment> \<open>NOTE added lemma.\<close>
lemma bool_update_inj:
fixes X :: "'a set" and v b
assumes "finite X" "v \<notin> X"
shows "inj_on (\<lambda>s. fmupd v b s) {s :: 'a state. fmdom' s = X}"
proof -
let ?f = "\<lambda>s :: 'a state. fmupd v b s"
{
fix s1 s2 :: "'a state"
assume "s1 \<in> {s :: 'a state. fmdom' s = X}" "s2 \<in> {s :: 'a state. fmdom' s = X}"
"?f s1 = ?f s2"
moreover
{
have
"\<forall>x\<in>X. x \<noteq> v \<longrightarrow> fmlookup (?f s1) x = fmlookup s1 x"
"\<forall>x\<in>X. x \<noteq> v \<longrightarrow> fmlookup (?f s2) x = fmlookup s2 x"
by simp+
then have
"\<forall>x\<in>X. x \<noteq> v \<longrightarrow> fmlookup s1 x = fmlookup s2 x"
using calculation(3)
by auto
}
moreover have "fmlookup s1 v = fmlookup s2 v"
using calculation \<open>v \<notin> X\<close>
by force
ultimately have "s1 = s2"
using fmap_neq
by fastforce
}
then show "inj_on (\<lambda>s. fmupd v b s) {s :: 'a state. fmdom' s = X}"
using inj_onI
by blast
qed
\<comment> \<open>NOTE added lemma.\<close>
lemma card_update:
fixes X v b
assumes "finite (X :: 'a set)" "v \<notin> X"
shows "
card ((\<lambda>s. fmupd v b s) ` {s :: 'a state. fmdom' s = X})
= card {s :: 'a state. fmdom' s = X}
"
proof -
have "inj_on (\<lambda>s. fmupd v b s) {s :: 'a state. fmdom' s = X}"
using assms bool_update_inj
by fast
then show
"card ((\<lambda>s. fmupd v b s) ` {s :: 'a state. fmdom' s = X}) = card {s :: 'a state. fmdom' s = X}"
using card_image by blast
qed
\<comment> \<open>NOTE added lemma.\<close>
lemma updates_disjoint:
fixes X x
assumes "finite X" "x \<notin> X"
shows "
((\<lambda>s. fmupd x True s) ` {s. fmdom' s = X})
\<inter> ((\<lambda>s. fmupd x False s) ` {s. fmdom' s = X}) = {}
"
proof -
let ?A = "((\<lambda>s. fmupd x True s) ` {s. fmdom' s = X})"
let ?B = "((\<lambda>s. fmupd x False s) ` {s. fmdom' s = X})"
{
assume C: "\<not>(\<forall>a\<in>?A. \<forall>b\<in>?B. a \<noteq> b)"
then have
"\<forall>a\<in>?A. \<forall>b\<in>?B. fmlookup a x \<noteq> fmlookup b x"
by simp
then have "\<forall>a\<in>?A. \<forall>b\<in>?B. a \<noteq> b"
by blast
then have False
using C
by blast
}
then show "?A \<inter> ?B = {}"
using disjoint_iff_not_equal
by blast
qed
lemma card_of_set_of_all_possible_states:
fixes X :: "'a set"
assumes "finite X"
shows "card {(s :: 'a state). fmdom' s = X} = 2 ^ (card X)"
using assms
proof (induction X)
case empty
then have 1: "{s :: 'a state. fmdom' s = {}} = {fmempty}"
using empty_domain_fmap_set
by simp
then have "card {fmempty} = 1"
using is_singleton_altdef
by blast
then have "2^(card {}) = 1"
by auto
then show ?case
using 1
by auto
next
case (insert x F)
then show ?case
\<comment> \<open>TODO refactor and simplify proof further.\<close>
proof (cases "x \<in> F")
case True
then show ?thesis
using insert.hyps(2)
by blast
next
case False
then have "
{s :: 'a state. fmdom' s = insert x F}
= (\<lambda>s. fmupd x True s) ` {s. fmdom' s = F} \<union> (\<lambda>s. fmupd x False s) ` {s. fmdom' s = F}
"
using False construction_of_all_possible_states_lemma
by metis
then have 2: "
card ({s :: 'a state. fmdom' s = insert x F})
= card ((\<lambda>s. fmupd x True s) ` {s. fmdom' s = F} \<union> (\<lambda>s. fmupd x False s) ` {s. fmdom' s = F})
"
by argo
then have 3: "2^(card (insert x F)) = 2 * 2^(card F)"
using False insert.hyps(1)
by simp
then have
"card ((\<lambda>s. fmupd x True s) ` {s. fmdom' s = F}) = 2^(card F)"
"card ((\<lambda>s. fmupd x False s) ` {s. fmdom' s = F}) = 2^(card F)"
using False card_update insert.IH insert.hyps(1)
by metis+
moreover have "
((\<lambda>s. fmupd x True s) ` {s. fmdom' s = F})
\<inter> ((\<lambda>s. fmupd x False s) ` {s. fmdom' s = F})
= {}
"
using False insert.hyps(1) updates_disjoint
by metis
moreover have "card (
((\<lambda>s. fmupd x True s) ` {s. fmdom' s = F})
\<union> ((\<lambda>s. fmupd x False s) ` {s. fmdom' s = F})
)
= card (((\<lambda>s. fmupd x True s) ` {s. fmdom' s = F}))
+ card ((\<lambda>s. fmupd x False s) ` {s. fmdom' s = F})
"
using calculation card_Un_disjoint card.infinite
power_eq_0_iff rel_simps(76)
by metis
then have "card (
((\<lambda>s. fmupd x True s) ` {s. fmdom' s = F})
\<union> ((\<lambda>s. fmupd x False s) ` {s. fmdom' s = F})
)
= 2 * (2^(card F))"
using calculation(1, 2)
by presburger
then have "card (
((\<lambda>s. fmupd x True s) ` {s. fmdom' s = F})
\<union> ((\<lambda>s. fmupd x False s) ` {s. fmdom' s = F})
)
= 2^(card (insert x F))"
using insert.IH 3
by metis
then show ?thesis
using "2"
by argo
qed
qed
subsubsection "State Lists and State Sets"
\<comment> \<open>NOTE using fun because of two defining equations.\<close>
\<comment> \<open>NOTE paired argument replaced by currying.\<close>
fun state_list where
"state_list s [] = [s]"
| "state_list s (a # as) = s # state_list (state_succ s a) as"
lemma empty_state_list_lemma:
fixes as s
shows "\<not>([] = state_list s as)"
proof (induction as)
qed auto
lemma state_list_length_non_zero:
fixes as s
shows "\<not>(0 = length (state_list s as))"
proof (induction as)
qed auto
lemma state_list_length_lemma:
fixes as s
shows "length as = length (state_list s as) - 1"
proof (induction as arbitrary: s)
next case (Cons a as)
have "length (state_list s (Cons a as)) - 1 = length (state_list (state_succ s a) as)"
by auto
\<comment> \<open>TODO unwrap metis proof.\<close>
then show "length (Cons a as) = length (state_list s (Cons a as)) - 1"
by (metis Cons.IH Suc_diff_1 empty_state_list_lemma length_Cons length_greater_0_conv)
qed simp
lemma state_list_length_lemma_2:
fixes as s
shows "(length (state_list s as)) = (length as + 1)"
proof (induction as arbitrary: s)
qed auto
\<comment> \<open>NOTE using fun because of multiple defining equations.\<close>
\<comment> \<open>NOTE name shortened to 'state\_def'\<close>
fun state_set where
"state_set [] = {}"
| "state_set (s # ss) = insert [s] (Cons s ` (state_set ss))"
lemma state_set_thm:
fixes s1
shows "s1 \<in> state_set s2 \<longleftrightarrow> prefix s1 s2 \<and> s1 \<noteq> []"
proof -
\<comment> \<open>NOTE Show equivalence by proving both directions. Left-to-right is trivial. Right-to-Left
primarily involves exploiting the prefix premise, induction hypothesis and `state\_set`
definition.\<close>
have "s1 \<in> state_set s2 \<Longrightarrow> prefix s1 s2 \<and> s1 \<noteq> []"
by (induction s2 arbitrary: s1) auto
moreover {
assume P: "prefix s1 s2" "s1 \<noteq> []"
then have "s1 \<in> state_set s2"
proof (induction s2 arbitrary: s1)
case (Cons a s2)
obtain s1' where 1: "s1 = a # s1'" "prefix s1' s2"
using Cons.prems(1, 2) prefix_Cons
by metis
then show ?case proof (cases "s1' = []")
case True
then show ?thesis
using 1
by force
next
case False
then have "s1' \<in> state_set s2"
using 1 False Cons.IH
by blast
then show ?thesis
using 1
by fastforce
qed
qed simp
}
ultimately show "s1 \<in> state_set s2 \<longleftrightarrow> prefix s1 s2 \<and> s1 \<noteq> []"
by blast
qed
lemma state_set_finite:
fixes X
shows "finite (state_set X)"
by (induction X) auto
lemma LENGTH_state_set:
fixes X e
assumes "e \<in> state_set X"
shows "length e \<le> length X"
using assms
by (induction X arbitrary: e) auto
lemma lemma_temp:
fixes x s as h
assumes "x \<in> state_set (state_list s as)"
shows "length (h # state_list s as) > length x"
using assms LENGTH_state_set le_imp_less_Suc
by force
lemma NIL_NOTIN_stateset:
fixes X
shows "[] \<notin> state_set X"
by (induction X) auto
\<comment> \<open>NOTE added lemma.\<close>
lemma state_set_card_i:
fixes X a
shows "[a] \<notin> (Cons a ` state_set X)"
by (induction X) auto
\<comment> \<open>NOTE added lemma.\<close>
lemma state_set_card_ii:
fixes X a
shows "card (Cons a ` state_set X) = card (state_set X)"
proof -
have "inj_on (Cons a) (state_set X)"
by simp
then show ?thesis
using card_image
by blast
qed
\<comment> \<open>NOTE added lemma.\<close>
lemma state_set_card_iii:
fixes X a
shows "card (state_set (a # X)) = 1 + card (state_set X)"
proof -
have "card (state_set (a # X)) = card (insert [a] (Cons a ` state_set X))"
by auto
\<comment> \<open>TODO unwrap this metis step.\<close>
also have "\<dots> = 1 + card (Cons a ` state_set X)"
using state_set_card_i
by (metis Suc_eq_plus1_left card_insert_disjoint finite_imageI state_set_finite)
also have"\<dots> = 1 + card (state_set X)"
using state_set_card_ii
by metis
finally show "card (state_set (a # X)) = 1 + card (state_set X)"
by blast
qed
lemma state_set_card:
fixes X
shows "card (state_set X) = length X"
proof (induction X)
case (Cons a X)
then have "card (state_set (a # X)) = 1 + card (state_set X)"
using state_set_card_iii
by fast
then show ?case
using Cons
by fastforce
qed auto
subsubsection "Properties of Domain Changes During Plan Execution"
lemma FDOM_state_succ:
assumes "fmdom' (snd a) \<subseteq> fmdom' s"
shows "(fmdom' (state_succ s a) = fmdom' s)"
unfolding state_succ_def fmap_add_ltr_def
using assms
by force
lemma FDOM_state_succ_subset:
"fmdom' (state_succ s a) \<subseteq> (fmdom' s \<union> fmdom' (snd a))"
unfolding state_succ_def fmap_add_ltr_def
by simp
\<comment> \<open>NOTE definition `qispl\_then` removed (was not being used).\<close>
lemma FDOM_eff_subset_FDOM_valid_states:
fixes p e s
assumes "(p, e) \<in> PROB" "(s \<in> valid_states PROB)"
shows "(fmdom' e \<subseteq> fmdom' s)"
proof -
{
have "fmdom' e \<subseteq> action_dom p e"
unfolding action_dom_def
by blast
also have "\<dots> \<subseteq> prob_dom PROB"
unfolding action_dom_def prob_dom_def
using assms(1)
by blast
finally have "fmdom' e \<subseteq> fmdom' s"
using assms
by (auto simp: valid_states_def)
}
then show "fmdom' e \<subseteq> fmdom' s"
by simp
qed
lemma FDOM_eff_subset_FDOM_valid_states_pair:
fixes a s
assumes "a \<in> PROB" "s \<in> valid_states PROB"
shows "fmdom' (snd a) \<subseteq> fmdom' s"
proof -
{
have "fmdom' (snd a) \<subseteq> (\<lambda>(s1, s2). action_dom s1 s2) a"
unfolding action_dom_def
using case_prod_beta
by fastforce
also have "\<dots> \<subseteq> prob_dom PROB"
using assms(1) prob_dom_def Sup_upper
by fast
finally have "fmdom' (snd a) \<subseteq> fmdom' s"
using assms(2) valid_states_def
by fast
}
then show ?thesis
by simp
qed
lemma FDOM_pre_subset_FDOM_valid_states:
fixes p e s
assumes "(p, e) \<in> PROB" "s \<in> valid_states PROB"
shows "fmdom' p \<subseteq> fmdom' s"
proof -
{
have "fmdom' p \<subseteq> (\<lambda>(s1, s2). action_dom s1 s2) (p, e)"
using action_dom_def
by fast
also have "\<dots> \<subseteq> prob_dom PROB"
using assms(1)
by (simp add: Sup_upper pair_imageI prob_dom_def)
finally have "fmdom' p \<subseteq> fmdom' s"
using assms(2) valid_states_def
by fast
}
then show ?thesis
by simp
qed
lemma FDOM_pre_subset_FDOM_valid_states_pair:
fixes a s
assumes "a \<in> PROB" "s \<in> valid_states PROB"
shows "fmdom' (fst a) \<subseteq> fmdom' s"
proof -
{
have "fmdom' (fst a) \<subseteq> (\<lambda>(s1, s2). action_dom s1 s2) a"
using action_dom_def
by force
also have "\<dots> \<subseteq> prob_dom PROB"
using assms(1)
by (simp add: Sup_upper pair_imageI prob_dom_def)
finally have "fmdom' (fst a) \<subseteq> fmdom' s"
using assms(2) valid_states_def
by fast
}
then show ?thesis
by simp
qed
\<comment> \<open>TODO unwrap the simp proof.\<close>
lemma action_dom_subset_valid_states_FDOM:
fixes p e s
assumes "(p, e) \<in> PROB" "s \<in> valid_states PROB"
shows "action_dom p e \<subseteq> fmdom' s"
using assms
by (simp add: Sup_upper pair_imageI prob_dom_def valid_states_def)
\<comment> \<open>TODO unwrap the metis proof.\<close>
lemma FDOM_eff_subset_prob_dom:
fixes p e
assumes "(p, e) \<in> PROB"
shows "fmdom' e \<subseteq> prob_dom PROB"
using assms
by (metis Sup_upper Un_subset_iff action_dom_def pair_imageI prob_dom_def)
lemma FDOM_eff_subset_prob_dom_pair:
fixes a
assumes "a \<in> PROB"
shows "fmdom' (snd a) \<subseteq> prob_dom PROB"
using assms(1) FDOM_eff_subset_prob_dom surjective_pairing
by metis
\<comment> \<open>TODO unwrap metis proof.\<close>
lemma FDOM_pre_subset_prob_dom:
fixes p e
assumes "(p, e) \<in> PROB"
shows "fmdom' p \<subseteq> prob_dom PROB"
using assms
by (metis (no_types) Sup_upper Un_subset_iff action_dom_def pair_imageI prob_dom_def)
lemma FDOM_pre_subset_prob_dom_pair:
fixes a
assumes "a \<in> PROB"
shows "fmdom' (fst a) \<subseteq> prob_dom PROB"
using assms FDOM_pre_subset_prob_dom surjective_pairing
by metis
subsubsection "Properties of Valid Plans"
lemma valid_plan_valid_head:
assumes "(h # as \<in> valid_plans PROB)"
shows "h \<in> PROB"
using assms valid_plans_def
by force
lemma valid_plan_valid_tail:
assumes "(h # as \<in> valid_plans PROB)"
shows "(as \<in> valid_plans PROB)"
using assms
by (simp add: valid_plans_def)
\<comment> \<open>TODO unwrap simp proof.\<close>
lemma valid_plan_pre_subset_prob_dom_pair:
assumes "as \<in> valid_plans PROB"
shows "(\<forall>a. ListMem a as \<longrightarrow> fmdom' (fst a) \<subseteq> (prob_dom PROB))"
unfolding valid_plans_def
using assms
by (simp add: FDOM_pre_subset_prob_dom_pair ListMem_iff rev_subsetD valid_plans_def)
lemma valid_append_valid_suff:
assumes "as1 @ as2 \<in> (valid_plans PROB)"
shows "as2 \<in> (valid_plans PROB)"
using assms
by (simp add: valid_plans_def)
lemma valid_append_valid_pref:
assumes "as1 @ as2 \<in> (valid_plans PROB)"
shows "as1 \<in> (valid_plans PROB)"
using assms
by (simp add: valid_plans_def)
lemma valid_pref_suff_valid_append:
assumes "as1 \<in> (valid_plans PROB)" "as2 \<in> (valid_plans PROB)"
shows "(as1 @ as2) \<in> (valid_plans PROB)"
using assms
by (simp add: valid_plans_def)
\<comment> \<open>NOTE showcase (case split seems necessary for MP of IH but the original proof does not need it).\<close>
lemma MEM_statelist_FDOM:
fixes PROB h as s0
assumes "s0 \<in> (valid_states PROB)" "as \<in> (valid_plans PROB)" "ListMem h (state_list s0 as)"
shows "(fmdom' h = fmdom' s0)"
using assms
proof (induction as arbitrary: PROB h s0)
case Nil
have "h = s0"
using Nil.prems(3) ListMem_iff
by force
then show ?case
by simp
next
case (Cons a as)
then show ?case
\<comment> \<open>NOTE This case split seems necessary to be able to infer
'ListMem h (state\_list (state\_succ s0 a) as)'
which is required in order to apply MP to the induction hypothesis.\<close>
proof (cases "h = s0")
case False
\<comment> \<open>TODO proof steps could be refactored into auxillary lemmas.\<close>
{
have "a \<in> PROB"
using Cons.prems(2) valid_plan_valid_head
by fast
then have "fmdom' (snd a) \<subseteq> fmdom' s0"
using Cons.prems(1) FDOM_eff_subset_FDOM_valid_states_pair
by blast
then have "fmdom' (state_succ s0 a) = fmdom' s0"
using FDOM_state_succ[of _ s0] Cons.prems(1) valid_states_def
by presburger
}
note 1 = this
{
have "fmdom' s0 = prob_dom PROB"
using Cons.prems(1) valid_states_def
by fast
then have "state_succ s0 a \<in> valid_states PROB"
unfolding valid_states_def
using 1
by force
}
note 2 = this
{
have "ListMem h (state_list (state_succ s0 a) as)"
using Cons.prems(3) False
by (simp add: ListMem_iff)
}
note 3 = this
{
have "as \<in> valid_plans PROB"
using Cons.prems(2) valid_plan_valid_tail
by fast
then have "fmdom' h = fmdom' (state_succ s0 a)"
using 1 2 3 Cons.IH[of "state_succ s0 a"]
by blast
}
then show ?thesis
using 1
by argo
qed simp
qed
\<comment> \<open>TODO unwrap metis proof.\<close>
lemma MEM_statelist_valid_state:
fixes PROB h as s0
assumes "s0 \<in> valid_states PROB" "as \<in> valid_plans PROB" "ListMem h (state_list s0 as)"
shows "(h \<in> valid_states PROB)"
using assms
by (metis MEM_statelist_FDOM mem_Collect_eq valid_states_def)
\<comment> \<open>TODO refactor (characterization lemma for 'state\_succ').\<close>
\<comment> \<open>TODO unwrap metis proof.\<close>
\<comment> \<open>NOTE added lemma.\<close>
lemma lemma_1_i:
fixes s a PROB
assumes "s \<in> valid_states PROB" "a \<in> PROB"
shows "state_succ s a \<in> valid_states PROB"
using assms
by (metis FDOM_eff_subset_FDOM_valid_states_pair FDOM_state_succ mem_Collect_eq valid_states_def)
\<comment> \<open>TODO unwrap smt proof.\<close>
\<comment> \<open>NOTE added lemma.\<close>
lemma lemma_1_ii:
"last ` ((#) s ` state_set (state_list (state_succ s a) as))
= last ` state_set (state_list (state_succ s a) as)"
by (smt NIL_NOTIN_stateset image_cong image_image last_ConsR)
lemma lemma_1:
fixes as :: "(('a, 'b) fmap \<times> ('a, 'b) fmap) list" and PPROB
assumes "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
shows "((last ` (state_set (state_list s as))) \<subseteq> valid_states PROB)"
using assms
proof (induction as arbitrary: s PROB)
\<comment> \<open>NOTE Base case simplifies to @{term "{s} \<subseteq> valid_states PROB"} which itself follows directly from
1st assumption.\<close>
case (Cons a as)
text \<open> Split the 'insert' term produced by @{term "state_set (state_list s (a # as))"} and proof
inclusion in 'valid\_states PROB' for both parts. \<close>
{
\<comment> \<open>NOTE Inclusion of the first subset follows from the induction premise by simplification.
The inclusion of the second subset is shown by applying the induction hypothesis to
`state\_succ s a` and some elementary set simplifications.\<close>
have "last [s] \<in> valid_states PROB"
using Cons.prems(1)
by simp
moreover {
{
have "a \<in> PROB"
using Cons.prems(2) valid_plan_valid_head
by fast
then have "state_succ s a \<in> valid_states PROB"
using Cons.prems(1) lemma_1_i
by blast
}
moreover have "as \<in> valid_plans PROB"
using Cons.prems(2) valid_plan_valid_tail
by fast
then have "(last ` state_set (state_list (state_succ s a) as)) \<subseteq> valid_states PROB"
using calculation Cons.IH[of "state_succ s a"]
by presburger
then have "(last ` ((#) s ` state_set (state_list (state_succ s a) as))) \<subseteq> valid_states PROB"
using lemma_1_ii
by metis
}
ultimately have
"(last ` insert [s] ((#) s ` state_set (state_list (state_succ s a) as))) \<subseteq> valid_states PROB"
by simp
}
then show ?case
by fastforce
qed auto
\<comment> \<open>TODO unwrap metis proof.\<close>
lemma len_in_state_set_le_max_len:
fixes as x PROB
assumes "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)" "\<not>(as = [])"
"(x \<in> state_set (state_list s as))"
shows "(length x \<le> (Suc (length as)))"
using assms
by (metis LENGTH_state_set Suc_eq_plus1_left add.commute state_list_length_lemma_2)
lemma card_state_set_cons:
fixes as s h
shows "
(card (state_set (state_list s (h # as)))
= Suc (card (state_set (state_list (state_succ s h) as))))
"
by (metis length_Cons state_list.simps(2) state_set_card)
lemma card_state_set:
fixes as s
shows "(Suc (length as)) = card (state_set (state_list s as))"
by (simp add: state_list_length_lemma_2 state_set_card)
lemma neq_mems_state_set_neq_len:
fixes as x y s
assumes "x \<in> state_set (state_list s as)" "(y \<in> state_set (state_list s as))" "\<not>(x = y)"
shows "\<not>(length x = length y)"
proof -
have "x \<noteq> []" "prefix x (state_list s as)"
using assms(1) state_set_thm
by blast+
moreover have "y \<noteq> []" "prefix y (state_list s as)"
using assms(2) state_set_thm
by blast+
ultimately show ?thesis
using assms(3) append_eq_append_conv prefixE
by metis
qed
\<comment> \<open>NOTE added definition (imported from pred\_setScript.sml:1562).\<close>
definition inj :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool" where
"inj f A B \<equiv> (\<forall>x \<in> A. f x \<in> B) \<and> inj_on f A"
\<comment> \<open>NOTE added lemma; refactored from `not\_eq\_last\_diff\_paths`.\<close>
lemma not_eq_last_diff_paths_i:
fixes s as PROB
assumes "s \<in> valid_states PROB" "as \<in> valid_plans PROB" "x \<in> state_set (state_list s as)"
shows "last x \<in> valid_states PROB"
proof -
have "last x \<in> last ` (state_set (state_list s as))"
using assms(3)
by simp
then show ?thesis
using assms(1, 2) lemma_1
by blast
qed
lemma not_eq_last_diff_paths_ii:
assumes "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
"\<not>(inj (last) (state_set (state_list s as)) (valid_states PROB))"
shows "\<exists>l1. \<exists>l2.
l1 \<in> state_set (state_list s as)
\<and> l2 \<in> state_set (state_list s as)
\<and> last l1 = last l2
\<and> l1 \<noteq> l2
"
proof -
let ?S="state_set (state_list s as)"
have 1: "\<not>(\<forall>x\<in>?S. last x \<in> valid_states PROB) = False"
using assms(1, 2) not_eq_last_diff_paths_i
by blast
{
have
"(\<not>(inj (last) ?S (valid_states PROB))) = (\<not>((\<forall>x\<in>?S. \<forall>y\<in>?S. last x = last y \<longrightarrow> x = y)))"
unfolding inj_def inj_on_def
using 1
by blast
then have "
(\<not>(inj (last) ?S (valid_states PROB)))
= (\<exists>x. \<exists>y. x\<in>?S \<and> y\<in>?S \<and> last x = last y \<and> x \<noteq> y)
"
using assms(3)
by blast
}
then show ?thesis
using assms(3) by blast
qed
lemma not_eq_last_diff_paths:
fixes as PROB s
assumes "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
"\<not>(inj (last) (state_set (state_list s as)) (valid_states PROB))"
shows "(\<exists>slist_1 slist_2.
(slist_1 \<in> state_set (state_list s as))
\<and> (slist_2 \<in> state_set (state_list s as))
\<and> ((last slist_1) = (last slist_2))
\<and> \<not>(length slist_1 = length slist_2))
"
proof -
obtain l1 l2 where "
l1 \<in> state_set (state_list s as)
\<and> l2 \<in> state_set (state_list s as)
\<and> last l1 = last l2
\<and> l1 \<noteq> l2
"
using assms(1, 2, 3) not_eq_last_diff_paths_ii
by blast
then show ?thesis
using neq_mems_state_set_neq_len
by blast
qed
\<comment> \<open>NOTE this lemma was removed due to being redundant and being shadowed later on:
lemma empty\_list\_nin\_state\_set\<close>
lemma nempty_sl_in_state_set:
fixes sl
assumes "sl \<noteq> []"
shows "sl \<in> state_set sl"
using assms state_set_thm
by auto
lemma empty_list_nin_state_set:
fixes h slist as
assumes "(h # slist) \<in> state_set (state_list s as)"
shows "(h = s)"
using assms
by (induction as) auto
lemma cons_in_state_set_2:
fixes s slist h t
assumes "(slist \<noteq> [])" "((s # slist) \<in> state_set (state_list s (h # t)))"
shows "(slist \<in> state_set (state_list (state_succ s h) t))"
using assms
by (induction slist) auto
\<comment> \<open>TODO move up and replace 'FactoredSystem.lemma\_1\_i'?\<close>
lemma valid_action_valid_succ:
assumes "h \<in> PROB" "s \<in> valid_states PROB"
shows "(state_succ s h) \<in> valid_states PROB"
using assms lemma_1_i
by blast
lemma in_state_set_imp_eq_exec_prefix:
fixes slist as PROB s
assumes "(as \<noteq> [])" "(slist \<noteq> [])" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
"(slist \<in> state_set (state_list s as))"
shows
"(\<exists>as'. (prefix as' as) \<and> (exec_plan s as' = last slist) \<and> (length slist = Suc (length as')))"
using assms
proof (induction slist arbitrary: as s PROB)
case cons_1: (Cons a slist)
have 1: "s # slist \<in> state_set (state_list s as)"
using cons_1.prems(5) empty_list_nin_state_set
by auto
then show ?case
using cons_1
proof (cases as)
case cons_2: (Cons a' R\<^sub>a\<^sub>s)
then have a: "state_succ s a' \<in> valid_states PROB"
using cons_1.prems(3, 4) valid_action_valid_succ valid_plan_valid_head
by metis
then have b: "R\<^sub>a\<^sub>s \<in> valid_plans PROB"
using cons_1.prems(4) cons_2 valid_plan_valid_tail
by fast
then show ?thesis
proof (cases slist)
case Nil
then show ?thesis
using cons_1.prems(5) empty_list_nin_state_set
by auto
next
case cons_3: (Cons a'' R\<^sub>s\<^sub>l\<^sub>i\<^sub>s\<^sub>t)
then have i: "a'' # R\<^sub>s\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<in> state_set (state_list (state_succ s a') R\<^sub>a\<^sub>s)"
using 1 cons_2 cons_in_state_set_2
by blast
then show ?thesis
proof (cases R\<^sub>a\<^sub>s)
case Nil
then show ?thesis
using i cons_2 cons_3
by auto
next
case (Cons a''' R\<^sub>a\<^sub>s')
then obtain as' where
"prefix as' (a''' # R\<^sub>a\<^sub>s')" "exec_plan (state_succ s a') as' = last slist"
"length slist = Suc (length as')"
using cons_1.IH[of "a''' # R\<^sub>a\<^sub>s'" "state_succ s a'" PROB]
using i a b cons_3
by blast
then show ?thesis
using Cons_prefix_Cons cons_2 cons_3 exec_plan.simps(2) last.simps length_Cons
list.distinct(1) local.Cons
by metis
qed
qed
qed auto
qed auto
lemma eq_last_state_imp_append_nempty_as:
fixes as PROB slist_1 slist_2
assumes "(as \<noteq> [])" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)" "(slist_1 \<noteq> [])"
"(slist_2 \<noteq> [])" "(slist_1 \<in> state_set (state_list s as))"
"(slist_2 \<in> state_set (state_list s as))" "\<not>(length slist_1 = length slist_2)"
"(last slist_1 = last slist_2)"
shows "(\<exists>as1 as2 as3.
(as1 @ as2 @ as3 = as)
\<and> (exec_plan s (as1 @ as2) = exec_plan s as1)
\<and> \<not>(as2 = [])
)"
proof -
obtain as_1 where 1: "(prefix as_1 as)" "(exec_plan s as_1 = last slist_1)"
"length slist_1 = Suc (length as_1)"
using assms(1, 2, 3, 4, 6) in_state_set_imp_eq_exec_prefix
by blast
obtain as_2 where 2: "(prefix as_2 as)" "(exec_plan s as_2 = last slist_2)"
"(length slist_2) = Suc (length as_2)"
using assms(1, 2, 3, 5, 7) in_state_set_imp_eq_exec_prefix
by blast
then have "length as_1 \<noteq> length as_2"
using assms(8) 1(3) 2(3)
by fastforce
then consider (i) "length as_1 < length as_2" | (ii) "length as_1 > length as_2"
by force
then show ?thesis
proof (cases)
case i
then have "prefix as_1 as_2"
using 1(1) 2(1) len_gt_pref_is_pref
by blast
then obtain a where a1: "as_2 = as_1 @ a"
using prefixE
by blast
then obtain b where b1: "as = as_2 @ b"
using prefixE 2(1)
by blast
let ?as1="as_1"
let ?as2="a"
let ?as3="b"
have "as = ?as1 @ ?as2 @ ?as3"
using a1 b1
by simp
moreover have "exec_plan s (?as1 @ ?as2) = exec_plan s ?as1"
using 1(2) 2(2) a1 assms(9)
by auto
moreover have "?as2 \<noteq> []"
using i a1
by simp
ultimately show ?thesis
by blast
next
case ii
then have "prefix as_2 as_1"
using 1(1) 2(1) len_gt_pref_is_pref
by blast
then obtain a where a2: "as_1 = as_2 @ a"
using prefixE
by blast
then obtain b where b2: "as = as_1 @ b"
using prefixE 1(1)
by blast
let ?as1="as_2"
let ?as2="a"
let ?as3="b"
have "as = ?as1 @ ?as2 @ ?as3"
using a2 b2
by simp
moreover have "exec_plan s (?as1 @ ?as2) = exec_plan s ?as1"
using 1(2) 2(2) a2 assms(9)
by auto
moreover have "?as2 \<noteq> []"
using ii a2
by simp
ultimately show ?thesis
by blast
qed
qed
lemma FINITE_prob_dom:
assumes "finite PROB"
shows "finite (prob_dom PROB)"
proof -
{
fix x
assume P2: "x \<in> PROB"
then have 1: "(\<lambda>(s1, s2). action_dom s1 s2) x = fmdom' (fst x) \<union> fmdom' (snd x)"
by (simp add: action_dom_def case_prod_beta')
then have 2: "finite (fset (fmdom (fst x)))" "finite (fset (fmdom (snd x)))"
by auto
then have 3: "fset (fmdom (fst x)) = fmdom' (fst x)" "fset (fmdom (snd x)) = fmdom' (snd x)"
by (auto simp add: fmdom'_alt_def)
then have "finite (fmdom' (fst x))"
using 2 by auto
then have "finite (fmdom' (snd x))"
using 2 3 by auto
then have "finite ((\<lambda>(s1, s2). action_dom s1 s2) x)"
using 1 2 3
by simp
}
then show "finite (prob_dom PROB)"
unfolding prob_dom_def
using assms
by blast
qed
lemma CARD_valid_states:
assumes "finite (PROB :: 'a problem)"
shows "(card (valid_states PROB :: 'a state set) = 2 ^ card (prob_dom PROB))"
proof -
have 1: "finite (prob_dom PROB)"
using assms FINITE_prob_dom
by blast
have"(card (valid_states PROB :: 'a state set)) = card {s :: 'a state. fmdom' s = prob_dom PROB}"
unfolding valid_states_def
by simp
also have "... = 2 ^ (card (prob_dom PROB))"
using 1 card_of_set_of_all_possible_states
by blast
finally show ?thesis
by blast
qed
\<comment> \<open>NOTE type of 'valid\_states PROB' has to be asserted to match 'FINITE\_states' in the proof.\<close>
lemma FINITE_valid_states:
fixes PROB :: "'a problem"
shows "finite PROB \<Longrightarrow> finite ((valid_states PROB) :: 'a state set)"
proof (induction PROB rule: finite.induct)
case emptyI
then have "valid_states {} = {fmempty}"
unfolding valid_states_def prob_dom_def
using empty_domain_fmap_set
by force
then show ?case
by(subst \<open>valid_states {} = {fmempty}\<close>) auto
next
case (insertI A a)
{
then have "finite (insert a A)"
by blast
then have "finite (prob_dom (insert a A))"
using FINITE_prob_dom
by blast
then have "finite {s :: 'a state. fmdom' s = prob_dom (insert a A)}"
using FINITE_states
by blast
}
then show ?case
unfolding valid_states_def
by simp
qed
\<comment> \<open>NOTE type of 'PROB' had to be fixed for use of 'FINITE\_valid\_states'.\<close>
lemma lemma_2:
fixes PROB :: "'a problem" and as :: "('a action) list" and s :: "'a state"
assumes "finite PROB" "s \<in> (valid_states PROB)" "(as \<in> valid_plans PROB)"
"((length as) > (2 ^ (card (fmdom' s)) - 1))"
shows "(\<exists>as1 as2 as3.
(as1 @ as2 @ as3 = as)
\<and> (exec_plan s (as1 @ as2) = exec_plan s as1)
\<and> \<not>(as2 = [])
)"
proof -
have "Suc (length as) > 2^(card (fmdom' s))"
using assms(4)
by linarith
then have 1: "card (state_set (state_list s as)) > 2^card (fmdom' s)"
using card_state_set[symmetric]
by metis
{
\<comment> \<open>NOTE type of 'valid\_states PROB' had to be asserted to match 'FINITE\_valid\_states'.\<close>
have 2: "finite (prob_dom PROB)" "finite ((valid_states PROB) :: 'a state set)"
using assms(1) FINITE_prob_dom FINITE_valid_states
by blast+
have 3: "fmdom' s = prob_dom PROB"
using assms(2) valid_states_def
by fast
then have "card ((valid_states PROB) :: 'a state set) = 2^card (fmdom' s)"
using assms(1) CARD_valid_states
by auto
then have 4: "card (state_set (state_list (s :: 'a state) as)) > card ((valid_states PROB) :: 'a state set)"
unfolding valid_states_def
using 1 2(1) 3 card_of_set_of_all_possible_states[of "prob_dom PROB"]
by argo
\<comment> \<open>TODO refactor into lemma.\<close>
{
let ?S="state_set (state_list (s :: 'a state) as)"
let ?T="valid_states PROB :: 'a state set"
assume C2: "inj_on last ?S"
\<comment> \<open>TODO unwrap the metis step or refactor into lemma.\<close>
have a: "?T \<subseteq> last ` ?S"
using C2
by (metis "2"(2) "4" assms(2) assms(3) card_image card_mono lemma_1 not_less)
have "finite (state_set (state_list s as))"
using state_set_finite
by auto
then have "card (last ` ?S) = card ?S"
using C2 inj_on_iff_eq_card
by blast
also have "\<dots> > card ?T"
using 4
by blast
then have "\<exists>x. x \<in> (last ` ?S) \<and> x \<notin> ?T"
using C2 a assms(2) assms(3) calculation lemma_1
by fastforce
}
note 5 = this
moreover
{
assume C: "inj last (state_set (state_list (s :: 'a state) as)) (valid_states PROB)"
then have "inj_on last (state_set (state_list (s :: 'a state) as))"
using C inj_def
by blast
then obtain x where "x \<in> last ` (state_set (state_list s as)) \<and> x \<notin> valid_states PROB"
using 5
by presburger
then have "\<not>(\<forall>x\<in>state_set (state_list s as). last x \<in> valid_states PROB)"
by blast
then have "\<not>inj last (state_set (state_list (s :: 'a state) as)) (valid_states PROB)"
using inj_def
by metis
then have False
using C
by simp
}
ultimately have "\<not>inj last (state_set (state_list (s :: 'a state) as)) (valid_states PROB)"
unfolding inj_def
by blast
}
then obtain slist_1 slist_2 where 6:
"slist_1 \<in> state_set (state_list s as)"
"slist_2 \<in> state_set (state_list s as)"
"(last slist_1 = last slist_2)"
"length slist_1 \<noteq> length slist_2"
using assms(2, 3) not_eq_last_diff_paths
by blast
then show ?thesis
proof (cases as)
case Nil
text \<open> 4th assumption is violated in the 'Nil' case. \<close>
then have "\<not>(2 ^ card (fmdom' s) - 1 < length as)"
using Nil
by simp
then show ?thesis
using assms(4)
by blast
next
case (Cons a list)
then have "as \<noteq> []"
by simp
moreover have "slist_1 \<noteq> []" "slist_2 \<noteq> []"
using 6(1, 2) NIL_NOTIN_stateset
by blast+
ultimately show ?thesis
using assms(2, 3) 6(1, 2, 3, 4) eq_last_state_imp_append_nempty_as
by fastforce
qed
qed
lemma lemma_2_prob_dom:
fixes PROB and as :: "('a action) list" and s :: "'a state"
assumes "finite PROB" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
"(length as > (2 ^ (card (prob_dom PROB))) - 1)"
shows "(\<exists>as1 as2 as3.
(as1 @ as2 @ as3 = as)
\<and> (exec_plan s (as1 @ as2) = exec_plan s as1)
\<and> \<not>(as2 = [])
)"
proof -
have "prob_dom PROB = fmdom' s"
using assms(2) valid_states_def
by fast
then have "2 ^ card (fmdom' s) - 1 < length as"
using assms(4)
by argo
then show ?thesis
using assms(1, 2, 3) lemma_2
by blast
qed
\<comment> \<open>NOTE type for `s` had to be fixed (type mismatch in obtain statement).\<close>
\<comment> \<open>NOTE type for `as1`, `as2` and `as3` had to be fixed (due type mismatch on `as1` in
`cycle\_removal\_lemma`)\<close>
lemma lemma_3:
fixes PROB :: "'a problem" and s :: "'a state"
assumes "finite PROB" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
"(length as > (2 ^ (card (prob_dom PROB)) - 1))"
shows "(\<exists>as'.
(exec_plan s as = exec_plan s as')
\<and> (length as' < length as)
\<and> (subseq as' as)
)"
proof -
have "prob_dom PROB = fmdom' s"
using assms(2) valid_states_def
by fast
then have "2 ^ card (fmdom' s) - 1 < length as"
using assms(4)
by argo
then obtain as1 as2 as3 :: "'a action list" where 1:
"as1 @ as2 @ as3 = as" "exec_plan s (as1 @ as2) = exec_plan s as1" "as2 \<noteq> []"
using assms(1, 2, 3) lemma_2
by metis
have 2: "exec_plan s (as1 @ as3) = exec_plan s (as1 @ as2 @ as3)"
using 1 cycle_removal_lemma
by fastforce
let ?as' = "as1 @ as3"
have "exec_plan s as = exec_plan s ?as'"
using 1 2
by auto
moreover have "length ?as' < length as"
using 1 nempty_list_append_length_add
by blast
moreover have "subseq ?as' as"
using 1 subseq_append'
by blast
ultimately show "(\<exists>as'.
(exec_plan s as = exec_plan s as') \<and> (length as' < length as) \<and> (subseq as' as))"
by blast
qed
\<comment> \<open>TODO unwrap meson step.\<close>
lemma sublist_valid_is_valid:
fixes as' as PROB
assumes "(as \<in> valid_plans PROB)" "(subseq as' as)"
shows "as' \<in> valid_plans PROB"
using assms
by (simp add: valid_plans_def) (meson dual_order.trans fset_of_list_subset sublist_subset)
\<comment> \<open>NOTE type of 's' had to be fixed (type mismatch in goal).\<close>
theorem main_lemma:
fixes PROB :: "'a problem" and as s
assumes "finite PROB" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
shows "(\<exists>as'.
(exec_plan s as = exec_plan s as')
\<and> (subseq as' as)
\<and> (length as' \<le> (2 ^ (card (prob_dom PROB))) - 1)
)"
proof (cases "length as \<le> (2 ^ (card (prob_dom PROB))) - 1")
case True
then have "exec_plan s as = exec_plan s as"
by simp
then have "subseq as as"
by auto
then have "length as \<le> (2^(card (prob_dom PROB)) - 1)"
using True
by auto
then show ?thesis
by blast
next
case False
then have "length as > (2 ^ (card (prob_dom PROB))) - 1"
using False
by auto
then obtain as' where 1:
"exec_plan s as = exec_plan s as'" "length as' < length as" "subseq as' as"
using assms lemma_3
by blast
{
fix p
assume "exec_plan s as = exec_plan s p" "subseq p as"
"2 ^ card (prob_dom PROB) - 1 < length p"
then have "(\<exists>p'. (exec_plan s as = exec_plan s p' \<and> subseq p' as) \<and> length p' < length p)"
using assms(1, 2, 3) lemma_3 sublist_valid_is_valid
by fastforce
}
then have "\<forall>p. exec_plan s as = exec_plan s p \<and> subseq p as \<longrightarrow>
(\<exists>p'. (exec_plan s as = exec_plan s p' \<and> subseq p' as)
\<and> length p' \<le> 2 ^ card (prob_dom PROB) - 1)"
using general_theorem[where
P = "\<lambda>(as'' :: 'a action list). (exec_plan s as = exec_plan s as'') \<and> subseq as'' as"
and l = "(2 ^ (card (prob_dom (PROB ::'a problem)))) - 1" and f = length]
by blast
then obtain p' where
"exec_plan s as = exec_plan s p'" "subseq p' as" "length p' \<le> 2 ^ card (prob_dom PROB) - 1"
by blast
then show ?thesis
using sublist_refl
by blast
qed
subsection "Reachable States"
\<comment> \<open>NOTE shortened to 'reachable\_s'\<close>
definition reachable_s where
"reachable_s PROB s \<equiv> {exec_plan s as | as. as \<in> valid_plans PROB}"
\<comment> \<open>NOTE types for `s` and `PROB` had to be fixed (type mismatch in goal).\<close>
lemma valid_as_valid_exec:
fixes as and s :: "'a state" and PROB :: "'a problem"
assumes "(as \<in> valid_plans PROB)" "(s \<in> valid_states PROB)"
shows "(exec_plan s as \<in> valid_states PROB)"
using assms
proof (induction as arbitrary: s PROB)
case (Cons a as)
then have "a \<in> PROB"
using valid_plan_valid_head
by metis
then have "state_succ s a \<in> valid_states PROB"
using Cons.prems(2) valid_action_valid_succ
by blast
moreover have "as \<in> valid_plans PROB"
using Cons.prems(1) valid_plan_valid_tail
by fast
ultimately show ?case
using Cons.IH
by force
qed simp
lemma exec_plan_fdom_subset:
fixes as s PROB
assumes "(as \<in> valid_plans PROB)"
shows "(fmdom' (exec_plan s as) \<subseteq> (fmdom' s \<union> prob_dom PROB))"
using assms
proof (induction as arbitrary: s PROB)
case (Cons a as)
have "as \<in> valid_plans PROB"
using Cons.prems valid_plan_valid_tail
by fast
then have "fmdom' (exec_plan (state_succ s a) as) \<subseteq> fmdom' (state_succ s a) \<union> prob_dom PROB"
using Cons.IH[of _ "state_succ s a"]
by simp
\<comment> \<open>TODO unwrap metis proofs.\<close>
moreover have "fmdom' s \<union> fmdom' (snd a) \<union> prob_dom PROB = fmdom' s \<union> prob_dom PROB"
by (metis
Cons.prems FDOM_eff_subset_prob_dom_pair sup_absorb2 sup_assoc valid_plan_valid_head)
ultimately show ?case
by (metis (no_types, lifting)
FDOM_state_succ_subset exec_plan.simps(2) order_refl subset_trans sup.mono)
qed simp
\<comment> \<open>NOTE added lemma.\<close>
lemma reachable_s_finite_thm_1_a:
fixes s and PROB :: "'a problem"
assumes "(s :: 'a state) \<in> valid_states PROB"
shows "(\<forall>l\<in>reachable_s PROB s. l\<in>valid_states PROB)"
proof -
have 1: "\<forall>l\<in>reachable_s PROB s. \<exists>as. l = exec_plan s as \<and> as \<in> valid_plans PROB"
using reachable_s_def
by fastforce
{
fix l
assume P1: "l \<in> reachable_s PROB s"
\<comment> \<open>NOTE type for 's' and 'as' had to be fixed due to type mismatch in obtain statement.\<close>
then obtain as :: "'a action list" where a: "l = exec_plan s as \<and> as \<in> valid_plans PROB"
using 1
by blast
then have "exec_plan s as \<in> valid_states PROB"
using assms a valid_as_valid_exec
by blast
then have "l \<in> valid_states PROB"
using a
by simp
}
then show "\<forall>l \<in> reachable_s PROB s. l \<in> valid_states PROB"
by blast
qed
lemma reachable_s_finite_thm_1:
assumes "((s :: 'a state) \<in> valid_states PROB)"
shows "(reachable_s PROB s \<subseteq> valid_states PROB)"
using assms reachable_s_finite_thm_1_a
by blast
\<comment> \<open>NOTE second declaration skipped (this is declared twice in the source; see above)\<close>
\<comment> \<open>NOTE type for `s` had to be fixed (type mismatch in goal).\<close>
lemma reachable_s_finite_thm:
fixes s :: "'a state"
assumes "finite (PROB :: 'a problem)" "(s \<in> valid_states PROB)"
shows "finite (reachable_s PROB s)"
using assms
by (meson FINITE_valid_states reachable_s_finite_thm_1 rev_finite_subset)
lemma empty_plan_is_valid: "[] \<in> (valid_plans PROB)"
by (simp add: valid_plans_def)
lemma valid_head_and_tail_valid_plan:
assumes "(h \<in> PROB)" "(as \<in> valid_plans PROB)"
shows "((h # as) \<in> valid_plans PROB)"
using assms
by (auto simp: valid_plans_def)
\<comment> \<open>TODO refactor\<close>
\<comment> \<open>NOTE added lemma\<close>
lemma lemma_1_reachability_s_i:
fixes PROB s
assumes "s \<in> valid_states PROB"
shows "s \<in> reachable_s PROB s"
proof -
have "[] \<in> valid_plans PROB"
using empty_plan_is_valid
by blast
then show ?thesis
unfolding reachable_s_def
by force
qed
\<comment> \<open>NOTE types for 'PROB' and 's' had to be fixed (type mismatch in goal).\<close>
lemma lemma_1_reachability_s:
fixes PROB :: "'a problem" and s :: "'a state" and as
assumes "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
shows "((last ` state_set (state_list s as)) \<subseteq> (reachable_s PROB s))"
using assms
proof(induction as arbitrary: PROB s)
case Nil
then have "(last ` state_set (state_list s [])) = {s}"
by force
then show ?case
unfolding reachable_s_def
using empty_plan_is_valid
by force
next
case cons: (Cons a as)
let ?S="last ` state_set (state_list s (a # as))"
{
let ?as="[]"
have "last [s] = exec_plan s ?as"
by simp
moreover have "?as \<in> valid_plans PROB"
using empty_plan_is_valid
by auto
ultimately have "\<exists>as. (last [s] = exec_plan s as) \<and> as \<in> valid_plans PROB"
by blast
}
note 1 = this
{
fix x
assume P: "x \<in> ?S"
then consider
(a) "x = last [s]"
| (b) "x \<in> last ` ((#) s ` state_set (state_list (state_succ s a) as))"
by auto
then have "x \<in> reachable_s PROB s"
proof (cases)
case a
then have "x = s"
by simp
then show ?thesis
using cons.prems(1) P lemma_1_reachability_s_i
by blast
next
case b
then obtain x'' where i:
"x'' \<in> state_set (state_list (state_succ s a) as)"
"x = last (s # x'')"
by blast
then show ?thesis
proof (cases "x''")
case Nil
then have "x = s"
using i
by fastforce
then show ?thesis
using cons.prems(1) lemma_1_reachability_s_i
by blast
next
case (Cons a' list)
then obtain x' where a:
"last (a' # list) = last x'" "x' \<in> state_set (state_list (state_succ s a) as)"
using i(1)
by blast
{
have "state_succ s a \<in> valid_states PROB"
using cons.prems(1, 2) valid_action_valid_succ valid_plan_valid_head
by metis
moreover have "as \<in> valid_plans PROB"
using cons.prems(2) valid_plan_valid_tail
by fast
ultimately have
"last ` state_set (state_list (state_succ s a) as) \<subseteq> reachable_s PROB (state_succ s a)"
using cons.IH[of "state_succ s a"]
by auto
then have "\<exists>as'.
last (a' # list) = exec_plan (state_succ s a) as' \<and> (as' \<in> (valid_plans PROB))"
unfolding state_set.simps state_list.simps reachable_s_def
using i(1) Cons
by blast
}
then obtain as' where b:
"last (a' # list) = exec_plan (state_succ s a) as'" "(as' \<in> (valid_plans PROB))"
by blast
then have "x = exec_plan (state_succ s a) as'"
using i(2) Cons a(1)
by auto
then show ?thesis unfolding reachable_s_def
using cons.prems(2) b(2)
by (metis (mono_tags, lifting) exec_plan.simps(2) mem_Collect_eq
valid_head_and_tail_valid_plan valid_plan_valid_head)
qed
qed
}
then show ?case
by blast
qed
\<comment> \<open>NOTE types for `PROB` and `s` had to be fixed for use of `lemma\_1\_reachability\_s`.\<close>
lemma not_eq_last_diff_paths_reachability_s:
fixes PROB :: "'a problem" and s :: "'a state" and as
assumes "s \<in> valid_states PROB" "as \<in> valid_plans PROB"
"\<not>(inj last (state_set (state_list s as)) (reachable_s PROB s))"
shows "(\<exists>slist_1 slist_2.
slist_1 \<in> state_set (state_list s as)
\<and> slist_2 \<in> state_set (state_list s as)
\<and> (last slist_1 = last slist_2)
\<and> \<not>(length slist_1 = length slist_2)
)"
proof -
{
fix x
assume P1: "x \<in> state_set (state_list s as)"
have a: "last ` state_set (state_list s as) \<subseteq> reachable_s PROB s"
using assms(1, 2) lemma_1_reachability_s
by fast
then have "\<forall>as PROB. s \<in> (valid_states PROB) \<and> as \<in> (valid_plans PROB) \<longrightarrow> (last ` (state_set (state_list s as)) \<subseteq> reachable_s PROB s)"
using lemma_1_reachability_s
by fast
then have "last x \<in> valid_states PROB"
using assms(1, 2) P1 lemma_1
by fast
then have "last x \<in> reachable_s PROB s"
using P1 a
by fast
}
note 1 = this
text \<open> Show the goal by disproving the contradiction. \<close>
{
assume C: "(\<forall>slist_1 slist_2. (slist_1 \<in> state_set (state_list s as)
\<and> slist_2 \<in> state_set (state_list s as)
\<and> (last slist_1 = last slist_2)) \<longrightarrow> (length slist_1 = length slist_2))"
moreover {
fix slist_1 slist_2
assume C1: "slist_1 \<in> state_set (state_list s as)" "slist_2 \<in> state_set (state_list s as)"
"(last slist_1 = last slist_2)"
moreover have i: "(length slist_1 = length slist_2)"
using C1 C
by blast
moreover have "slist_1 = slist_2"
using C1(1, 2) i neq_mems_state_set_neq_len
by auto
ultimately have "inj_on last (state_set (state_list s as))"
unfolding inj_on_def
using C neq_mems_state_set_neq_len
by blast
then have False
using 1 inj_def assms(3)
by blast
}
ultimately have False
by (metis empty_state_list_lemma nempty_sl_in_state_set)
}
then show ?thesis
by blast
qed
\<comment> \<open>NOTE added lemma ( translation of `PHP` in pred\_setScript.sml:3155).\<close>
lemma lemma_2_reachability_s_i:
fixes f :: "'a \<Rightarrow> 'b" and s t
assumes "finite t" "card t < card s"
shows "\<not>(inj f s t)"
proof -
{
assume C: "inj f s t"
then have 1: "(\<forall>x\<in>s. f x \<in> t)" "inj_on f s"
unfolding inj_def
by blast+
moreover {
have "f ` s \<subseteq> t"
using 1
by fast
then have "card (f ` s) \<le> card t"
using assms(1) card_mono
by auto
}
moreover have "card (f ` s) = card s"
using 1 card_image
by fast
ultimately have False
using assms(2)
by linarith
}
then show ?thesis
by blast
qed
lemma lemma_2_reachability_s:
fixes PROB :: "'a problem" and as s
assumes "finite PROB" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
"(length as > card (reachable_s PROB s) - 1)"
shows "(\<exists>as1 as2 as3.
(as1 @ as2 @ as3 = as) \<and> (exec_plan s (as1 @ as2) = exec_plan s as1) \<and> \<not>(as2 = []))"
proof -
{
have "Suc (length as) > card (reachable_s PROB s)"
using assms(4)
by fastforce
then have "card (state_set (state_list s as)) > card (reachable_s PROB s)"
using card_state_set
by metis
}
note 1 = this
{
have "finite (reachable_s PROB s)"
using assms(1, 2) reachable_s_finite_thm
by blast
then have "\<not>(inj last (state_set (state_list s as)) (reachable_s PROB s))"
using assms(4) 1 lemma_2_reachability_s_i
by blast
}
note 2 = this
obtain slist_1 slist_2 where 3:
"slist_1 \<in> state_set (state_list s as)" "slist_2 \<in> state_set (state_list s as)"
"(last slist_1 = last slist_2)" "length slist_1 \<noteq> length slist_2"
using assms(2, 3) 2 not_eq_last_diff_paths_reachability_s
by blast
then show ?thesis using assms
proof(cases as)
case (Cons a list)
then show ?thesis
using assms(2, 3) 3 eq_last_state_imp_append_nempty_as state_set_thm list.distinct(1)
by metis
qed force
qed
lemma lemma_3_reachability_s:
fixes as and PROB :: "'a problem" and s
assumes "finite PROB" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
"(length as > (card (reachable_s PROB s) - 1))"
shows "(\<exists>as'.
(exec_plan s as = exec_plan s as')
\<and> (length as' < length as)
\<and> (subseq as' as)
)"
proof -
obtain as1 as2 as3 :: "'a action list" where 1:
"(as1 @ as2 @ as3 = as)" "(exec_plan s (as1 @ as2) = exec_plan s as1)" "~(as2=[])"
using assms lemma_2_reachability_s
by metis
then have "(exec_plan s (as1 @ as2) = exec_plan s as1)"
using 1
by blast
then have 2: "exec_plan s (as1 @ as3) = exec_plan s (as1 @ as2 @ as3)"
using 1 cycle_removal_lemma
by fastforce
let ?as' = "as1 @ as3"
have 3: "exec_plan s as = exec_plan s ?as'"
using 1 2
by argo
then have "as2 \<noteq> []"
using 1
by blast
then have 4: "length ?as' < length as"
using nempty_list_append_length_add 1
by blast
then have "subseq ?as' as"
using 1 subseq_append'
by blast
then show ?thesis
using 3 4
by blast
qed
\<comment> \<open>NOTE type for `as` had to be fixed (type mismatch in goal).\<close>
lemma main_lemma_reachability_s:
fixes PROB :: "'a problem" and as and s :: "'a state"
assumes "finite PROB" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
shows "(\<exists>as'.
(exec_plan s as = exec_plan s as') \<and> subseq as' as
\<and> (length as' \<le> (card (reachable_s PROB s) - 1)))"
proof (cases "length as \<le> card (reachable_s PROB s) - 1")
case False
let ?as' = "as"
have "length as > card (reachable_s PROB s) - 1"
using False
by simp
{
fix p
assume P: "exec_plan s as = exec_plan s p" "subseq p as"
"card (reachable_s PROB s) - 1 < length p"
moreover have "p \<in> valid_plans PROB"
using assms(3) P(2) sublist_valid_is_valid
by blast
ultimately obtain as' where 1:
"exec_plan s p = exec_plan s as'" "length as' < length p" "subseq as' p"
using assms lemma_3_reachability_s
by blast
then have "exec_plan s as = exec_plan s as'"
using P
by presburger
moreover have "subseq as' as"
using P 1 sublist_trans
by blast
ultimately have "(\<exists>p'. (exec_plan s as = exec_plan s p' \<and> subseq p' as) \<and> length p' < length p)"
using 1
by blast
}
then have "\<forall>p.
exec_plan s as = exec_plan s p \<and> subseq p as
\<longrightarrow> (\<exists>p'.
(exec_plan s as = exec_plan s p' \<and> subseq p' as)
\<and> length p' \<le> card (reachable_s PROB s) - 1)"
using general_theorem[of "\<lambda>as''. (exec_plan s as = exec_plan s as'') \<and> subseq as'' as"
"(card (reachable_s (PROB :: 'a problem) (s :: 'a state)) - 1)" length]
by blast
then show ?thesis
by blast
qed blast
lemma reachable_s_non_empty: "\<not>(reachable_s PROB s = {})"
using empty_plan_is_valid reachable_s_def
by blast
lemma card_reachable_s_non_zero:
fixes s
assumes "finite (PROB :: 'a problem)" "(s \<in> valid_states PROB)"
shows "(0 < card (reachable_s PROB s))"
using assms
by (simp add: card_gt_0_iff reachable_s_finite_thm reachable_s_non_empty)
lemma exec_fdom_empty_prob:
fixes s
assumes "(prob_dom PROB = {})" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
shows "(exec_plan s as = fmempty)"
proof -
have "fmdom' s = {}"
using assms(1, 2)
by (simp add: valid_states_def)
then show "exec_plan s as = fmempty"
using assms(1, 3)
by (metis
exec_plan_fdom_subset fmrestrict_set_dom fmrestrict_set_null subset_empty
sup_bot.left_neutral)
qed
\<comment> \<open>NOTE types for `PROB` and `s` had to be fixed (type mismatch in goal).\<close>
lemma reachable_s_empty_prob:
fixes PROB :: "'a problem" and s :: "'a state"
assumes "(prob_dom PROB = {})" "(s \<in> valid_states PROB)"
shows "((reachable_s PROB s) \<subseteq> {fmempty})"
proof -
{
fix x
assume P1: "x \<in> reachable_s PROB s"
then obtain as :: "'a action list" where a:
"as \<in> valid_plans PROB" "x = exec_plan s as"
using reachable_s_def
by blast
then have "as \<in> valid_plans PROB" "x = exec_plan s as"
using a
by auto
then have "x = fmempty" using assms(1, 2) exec_fdom_empty_prob
by blast
}
then show "((reachable_s PROB s) \<subseteq> {fmempty})"
by blast
qed
\<comment> \<open>NOTE this is semantically equivalent to `sublist\_valid\_is\_valid`.\<close>
\<comment> \<open>NOTE Renamed to 'sublist\_valid\_plan\_alt' because another lemma by the same name is declared
later.\<close>
lemma sublist_valid_plan__alt:
assumes "(as1 \<in> valid_plans PROB)" "(subseq as2 as1)"
shows "(as2 \<in> valid_plans PROB)"
using assms
by (auto simp add: sublist_valid_is_valid)
lemma fmsubset_eq:
assumes "s1 \<subseteq>\<^sub>f s2"
shows "(\<forall>a. a |\<in>| fmdom s1 \<longrightarrow> fmlookup s1 a = fmlookup s2 a)"
using assms
by (metis (mono_tags, lifting) domIff fmdom_notI fmsubset.rep_eq map_le_def)
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor/move into 'FmapUtils.thy'.\<close>
lemma submap_imp_state_succ_submap_a:
assumes "s1 \<subseteq>\<^sub>f s2" "s2 \<subseteq>\<^sub>f s3"
shows "s1 \<subseteq>\<^sub>f s3"
using assms fmsubset.rep_eq map_le_trans
by blast
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor into FmapUtils?\<close>
lemma submap_imp_state_succ_submap_b:
assumes "s1 \<subseteq>\<^sub>f s2"
shows "(s0 ++ s1) \<subseteq>\<^sub>f (s0 ++ s2)"
proof -
{
assume C: "\<not>((s0 ++ s1) \<subseteq>\<^sub>f (s0 ++ s2))"
then have 1: "(s0 ++ s1) = (s1 ++\<^sub>f s0)"
using fmap_add_ltr_def
by blast
then have 2: "(s0 ++ s2) = (s2 ++\<^sub>f s0)"
using fmap_add_ltr_def
by auto
then obtain a where 3:
"a |\<in>| fmdom (s1 ++\<^sub>f s0) \<and> fmlookup (s1 ++\<^sub>f s0) \<noteq> fmlookup (s2 ++\<^sub>f s0)"
using C 1 2 fmsubset.rep_eq domIff fmdom_notD map_le_def
by (metis (no_types, lifting))
then have False
using assms(1) C proof (cases "a |\<in>| fmdom s1")
case True
moreover have "fmlookup s1 a = fmlookup s2 a"
by (meson assms(1) calculation fmsubset_eq)
moreover have "fmlookup (s0 ++\<^sub>f s1) a = fmlookup s1 a"
by (simp add: True)
moreover have "a |\<in>| fmdom s2"
using True calculation(2) fmdom_notD by fastforce
moreover have "fmlookup (s0 ++\<^sub>f s2) a = fmlookup s2 a"
by (simp add: calculation(4))
moreover have "fmlookup (s0 ++\<^sub>f s1) a = fmlookup (s0 ++\<^sub>f s2) a"
using calculation(2, 3, 5)
by auto
ultimately show ?thesis
by (smt "1" "2" C assms domIff fmlookup_add fmsubset.rep_eq map_le_def)
next
case False
moreover have "fmlookup (s0 ++\<^sub>f s1) a = fmlookup s0 a"
by (auto simp add: False)
ultimately show ?thesis proof (cases "a |\<in>| fmdom s0")
case True
have "a |\<notin>| fmdom (s1 ++\<^sub>f s0)"
by (smt "1" "2" C UnE assms dom_map_add fmadd.rep_eq fmsubset.rep_eq map_add_def
map_add_dom_app_simps(1) map_le_def)
then show ?thesis
using 3 by blast
next
case False
then have "a |\<notin>| fmdom (s1 ++\<^sub>f s0)"
using \<open>fmlookup (s0 ++\<^sub>f s1) a = fmlookup s0 a\<close>
by force
then show ?thesis
using 3
by blast
qed
qed
}
then show ?thesis
by blast
qed
\<comment> \<open>NOTE type for `a` had to be fixed (type mismatch in goal).\<close>
lemma submap_imp_state_succ_submap:
fixes a :: "'a action" and s1 s2
assumes "(fst a \<subseteq>\<^sub>f s1)" "(s1 \<subseteq>\<^sub>f s2)"
shows "(state_succ s1 a \<subseteq>\<^sub>f state_succ s2 a)"
proof -
have 1: "state_succ s1 a = (snd a ++ s1)"
using assms(1)
by (simp add: state_succ_def)
then have "fst a \<subseteq>\<^sub>f s2"
using assms(1, 2) submap_imp_state_succ_submap_a
by auto
then have 2: "state_succ s2 a = (snd a ++ s2)"
using 1 state_succ_def
by metis
then have "snd a ++ s1 \<subseteq>\<^sub>f snd a ++ s2"
using assms(2) submap_imp_state_succ_submap_b
by fast
then show ?thesis
using 1 2
by argo
qed
\<comment> \<open>NOTE types for `a`, `s1` and `s2` had to be fixed (type mismatch in goal).\<close>
lemma pred_dom_subset_succ_submap:
fixes a :: "'a action" and s1 s2 :: "'a state"
assumes "(fmdom' (fst a) \<subseteq> fmdom' s1)" "(s1 \<subseteq>\<^sub>f s2)"
shows "(state_succ s1 a \<subseteq>\<^sub>f state_succ s2 a)"
using assms
unfolding state_succ_def
proof (auto)
assume P1: "fmdom' (fst a) \<subseteq> fmdom' s1" "s1 \<subseteq>\<^sub>f s2" "fst a \<subseteq>\<^sub>f s1" "fst a \<subseteq>\<^sub>f s2"
then show "snd a ++ s1 \<subseteq>\<^sub>f snd a ++ s2"
using submap_imp_state_succ_submap_b
by fast
next
assume P2: "fmdom' (fst a) \<subseteq> fmdom' s1" "s1 \<subseteq>\<^sub>f s2" "fst a \<subseteq>\<^sub>f s1" "\<not> fst a \<subseteq>\<^sub>f s2"
then show "snd a ++ s1 \<subseteq>\<^sub>f s2"
using submap_imp_state_succ_submap_a
by blast
next
assume P3: "fmdom' (fst a) \<subseteq> fmdom' s1" "s1 \<subseteq>\<^sub>f s2" "\<not> fst a \<subseteq>\<^sub>f s1" "fst a \<subseteq>\<^sub>f s2"
{
have a: "fmlookup s1 \<subseteq>\<^sub>m fmlookup s2"
using P3(2) fmsubset.rep_eq
by blast
{
have "\<not>(fmlookup (fst a) \<subseteq>\<^sub>m fmlookup s1)"
using P3(3) fmsubset.rep_eq
by blast
then have "\<exists>v \<in> dom (fmlookup (fst a)). fmlookup (fst a) v \<noteq> fmlookup s1 v"
using map_le_def
by fast
}
then obtain v where b: "v \<in> dom (fmlookup (fst a))" "fmlookup (fst a) v \<noteq> fmlookup s1 v"
by blast
then have "fmlookup (fst a) v \<noteq> fmlookup s2 v"
using assms(1) a contra_subsetD fmdom'.rep_eq map_le_def
by metis
then have "\<not>(fst a \<subseteq>\<^sub>f s2)"
using b fmsubset.rep_eq map_le_def
by metis
}
then show "s1 \<subseteq>\<^sub>f snd a ++ s2"
using P3(4)
by simp
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor.\<close>
lemma valid_as_submap_init_submap_exec_i:
fixes s a
shows "fmdom' s \<subseteq> fmdom' (state_succ s a)"
proof (cases "fst a \<subseteq>\<^sub>f s")
case True
then have "state_succ s a = s ++\<^sub>f (snd a)"
unfolding state_succ_def
using fmap_add_ltr_def
by auto
then have "fmdom' (state_succ s a) = fmdom' s \<union> fmdom' (snd a)"
using fmdom'_add
by simp
then show ?thesis
by simp
next
case False
then show ?thesis
unfolding state_succ_def
by simp
qed
\<comment> \<open>NOTE types for `s1` and `s2` had to be fixed in order to apply `pred\_dom\_subset\_succ\_submap`.\<close>
lemma valid_as_submap_init_submap_exec:
fixes s1 s2 :: "'a state"
assumes "(s1 \<subseteq>\<^sub>f s2) " "(\<forall>a. ListMem a as \<longrightarrow> (fmdom' (fst a) \<subseteq> fmdom' s1))"
shows "(exec_plan s1 as \<subseteq>\<^sub>f exec_plan s2 as)"
using assms
proof (induction as arbitrary: s1 s2)
case (Cons a as)
{
have "ListMem a (a # as)"
using elem
by fast
then have "fmdom' (fst a) \<subseteq> fmdom' s1"
using Cons.prems(2)
by blast
then have "state_succ s1 a \<subseteq>\<^sub>f state_succ s2 a"
using Cons.prems(1) pred_dom_subset_succ_submap
by fast
}
note 1 = this
{
fix b
assume "ListMem b as"
then have "ListMem b (a # as)"
using insert
by fast
then have a: "fmdom' (fst b) \<subseteq> fmdom' s1"
using Cons.prems(2)
by blast
then have "fmdom' s1 \<subseteq> fmdom' (state_succ s1 a)"
using valid_as_submap_init_submap_exec_i
by metis
then have "fmdom' (fst b) \<subseteq> fmdom' (state_succ s1 a)"
using a
by simp
}
then show ?case
using 1 Cons.IH[of "(state_succ s1 a)" "(state_succ s2 a)"]
by fastforce
qed auto
lemma valid_plan_mems:
assumes "(as \<in> valid_plans PROB)" "(ListMem a as)"
shows "a \<in> PROB"
using assms ListMem_iff in_set_conv_decomp valid_append_valid_suff valid_plan_valid_head
by (metis)
\<comment> \<open>NOTE typing moved into 'fixes' due to type mismatches when using lemma.\<close>
\<comment> \<open>NOTE showcase (this can't be used due to type problems when the type is specified within
proposition.\<close>
lemma valid_states_nempty:
fixes PROB :: "(('a, 'b) fmap \<times> ('a, 'b) fmap) set"
assumes "finite PROB"
shows "\<exists>s. s \<in> (valid_states PROB)"
unfolding valid_states_def
using fmchoice'[OF FINITE_prob_dom[OF assms], where Q = "\<lambda>_ _. True"]
by auto
lemma empty_prob_dom_single_val_state:
assumes "(prob_dom PROB = {})"
shows "(\<exists>s. valid_states PROB = {s})"
proof -
{
assume C: "\<not>(\<exists>s. valid_states PROB = {s})"
then have "valid_states PROB = {s. fmdom' s = {}}"
using assms
by (simp add: valid_states_def)
then have "\<exists>s. valid_states PROB = {s}"
using empty_domain_fmap_set
by blast
then have False
using C
by blast
}
then show ?thesis
by blast
qed
lemma empty_prob_dom_imp_empty_plan_always_good:
fixes PROB s
assumes "(prob_dom PROB = {})" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)"
shows "(exec_plan s [] = exec_plan s as)"
using assms empty_plan_is_valid exec_fdom_empty_prob
by fastforce
lemma empty_prob_dom:
fixes PROB
assumes "(prob_dom PROB = {})"
shows "(PROB = {(fmempty, fmempty)} \<or> PROB = {})"
using assms
proof (cases "PROB = {}")
case False
have "\<Union>((\<lambda>(s1, s2). fmdom' s1 \<union> fmdom' s2) ` PROB) = {}"
using assms
by (simp add: prob_dom_def action_dom_def)
then have 1:"\<forall>a\<in>PROB. (\<lambda>(s1, s2). fmdom' s1 \<union> fmdom' s2) a = {}"
using Union_empty_conv
by auto
{
fix a
assume P1: "a\<in>PROB"
then have "(\<lambda>(s1, s2). fmdom' s1 \<union> fmdom' s2) a = {}"
using 1
by simp
then have a: "fmdom' (fst a) = {}" "fmdom' (snd a) = {}"
by auto+
then have b: "fst a = fmempty"
using fmrestrict_set_dom fmrestrict_set_null
by metis
then have "snd a = fmempty"
using a(2) fmrestrict_set_dom fmrestrict_set_null
by metis
then have "a = (fmempty, fmempty)"
using b surjective_pairing
by metis
}
then have "PROB = {(fmempty, fmempty)}"
using False
by blast
then show ?thesis
by blast
qed simp
lemma empty_prob_dom_finite:
fixes PROB :: "'a problem"
assumes "prob_dom PROB = {}"
shows "finite PROB"
proof -
consider (i) "PROB = {(fmempty, fmempty)}" | (ii) "PROB = {}"
using assms empty_prob_dom
by auto
then show ?thesis by (cases) auto
qed
\<comment> \<open>NOTE type for `a` had to be fixed (type mismatch in goal).\<close>
lemma disj_imp_eq_proj_exec:
fixes a :: "('a, 'b) fmap \<times> ('a, 'b) fmap" and vs s
assumes "(fmdom' (snd a) \<inter> vs) = {}"
shows "(fmrestrict_set vs s = fmrestrict_set vs (state_succ s a))"
proof -
have "disjnt (fmdom' (snd a)) vs"
using assms disjnt_def
by fast
then show ?thesis
using disj_dom_drest_fupdate_eq state_succ_pair surjective_pairing
by metis
qed
lemma no_change_vs_eff_submap:
fixes a vs s
assumes "(fmrestrict_set vs s = fmrestrict_set vs (state_succ s a))" "(fst a \<subseteq>\<^sub>f s)"
shows "(fmrestrict_set vs (snd a) \<subseteq>\<^sub>f (fmrestrict_set vs s))"
proof -
{
fix x
assume P3: "x \<in> dom (fmlookup (fmrestrict_set vs (snd a)))"
then have "(fmlookup (fmrestrict_set vs (snd a))) x = (fmlookup (fmrestrict_set vs s)) x"
proof (cases "fmlookup (fmrestrict_set vs (snd a)) x")
case None
then show ?thesis using P3 by blast
next
case (Some y)
then have "fmrestrict_set vs s = fmrestrict_set vs (s ++\<^sub>f snd a)"
using assms
by (simp add: state_succ_def fmap_add_ltr_def)
then have "fmlookup (fmrestrict_set vs s) = fmlookup (fmrestrict_set vs (s ++\<^sub>f snd a))"
by auto
then have 1: "
fmlookup (fmrestrict_set vs s) x
= (if x \<in> vs then fmlookup (s ++\<^sub>f snd a) x else None)
"
using fmlookup_restrict_set
by metis
then show ?thesis
proof (cases "x \<in> vs")
case True
then have "fmlookup (fmrestrict_set vs s) x = fmlookup (s ++\<^sub>f snd a) x"
using True 1
by auto
then show ?thesis
using Some fmadd.rep_eq fmlookup_restrict_set map_add_Some_iff
by (metis (mono_tags, lifting))
next
case False
then have 1: "fmlookup (fmrestrict_set vs s) x = None"
using False "1"
by auto
then show ?thesis
using 1 False
by auto
qed
qed
}
then have "(fmlookup (fmrestrict_set vs (snd a)) \<subseteq>\<^sub>m fmlookup (fmrestrict_set vs s))"
using map_le_def
by blast
then show ?thesis
using fmsubset.rep_eq
by blast
qed
\<comment> \<open>NOTE type of `a` had to be fixed.\<close>
lemma sat_precond_as_proj_3:
fixes s and a :: "('a, 'b) fmap \<times> ('a, 'b) fmap" and vs
assumes "(fmdom' (fmrestrict_set vs (snd a)) = {})"
shows "((fmrestrict_set vs (state_succ s a)) = (fmrestrict_set vs s))"
proof -
have "fmdom' (fmrestrict_set vs (fmrestrict_set vs (snd a))) = {}"
using assms fmrestrict_set_dom fmrestrict_set_empty fmrestrict_set_null
by metis
{
fix x
assume C: "x \<in> fmdom' (snd a) \<and> x \<in> vs"
then have a: "x \<in> fmdom' (snd a)" "x \<in> vs"
using C
by blast+
then have "fmlookup (snd a) x \<noteq> None"
using fmdom'_notI
by metis
then have "fmlookup (fmrestrict_set vs (snd a)) x \<noteq> None"
using a(2)
by force
then have "x \<in> fmdom' (fmrestrict_set vs (snd a))"
using fmdom'_notD
by metis
then have "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}"
by blast
then have False
using assms
by blast
}
then have "\<forall>x. \<not>(x \<in> fmdom' (snd a) \<and> x \<in> vs)"
by blast
then have 1: "fmdom' (snd a) \<inter> vs = {}"
by blast
have "disjnt (fmdom' (snd a)) vs"
using 1 disjnt_def
by blast
then show ?thesis
using 1 disj_imp_eq_proj_exec
by metis
qed
\<comment> \<open>NOTE type for `a` had to be fixed (type mismatch in goal).\<close>
\<comment> \<open>TODO showcase (quick win with simp).\<close>
lemma proj_eq_proj_exec_eq:
fixes s s' vs and a :: "('a, 'b) fmap \<times> ('a, 'b) fmap" and a'
assumes "((fmrestrict_set vs s) = (fmrestrict_set vs s'))" "((fst a \<subseteq>\<^sub>f s) = (fst a' \<subseteq>\<^sub>f s'))"
"(fmrestrict_set vs (snd a) = fmrestrict_set vs (snd a'))"
shows "(fmrestrict_set vs (state_succ s a) = fmrestrict_set vs (state_succ s' a'))"
using assms
by (simp add: fmap_add_ltr_def state_succ_def)
lemma empty_eff_exec_eq:
fixes s a
assumes "(fmdom' (snd a) = {})"
shows "(state_succ s a = s)"
using assms
unfolding state_succ_def fmap_add_ltr_def
by (metis fmadd_empty(2) fmrestrict_set_dom fmrestrict_set_null)
lemma exec_as_proj_valid_2:
fixes a
assumes "a \<in> PROB"
shows "(action_dom (fst a) (snd a) \<subseteq> prob_dom PROB)"
using assms
by (simp add: FDOM_eff_subset_prob_dom_pair FDOM_pre_subset_prob_dom_pair action_dom_def)
lemma valid_filter_valid_as:
assumes "(as \<in> valid_plans PROB)"
shows "(filter P as \<in> valid_plans PROB)"
using assms
by(auto simp: valid_plans_def)
lemma sublist_valid_plan:
assumes "(subseq as' as)" "(as \<in> valid_plans PROB)"
shows "(as' \<in> valid_plans PROB)"
using assms
by (auto simp: valid_plans_def) (meson fset_mp fset_of_list_elem sublist_subset subsetCE)
lemma prob_subset_dom_subset:
assumes "PROB1 \<subseteq> PROB2"
shows "(prob_dom PROB1 \<subseteq> prob_dom PROB2)"
using assms
by (auto simp add: prob_dom_def)
lemma state_succ_valid_act_disjoint:
assumes "(a \<in> PROB)" "(vs \<inter> (prob_dom PROB) = {})"
shows "(fmrestrict_set vs (state_succ s a) = fmrestrict_set vs s)"
using assms
by (smt
FDOM_eff_subset_prob_dom_pair disj_imp_eq_proj_exec inf.absorb1
inf_bot_right inf_commute inf_left_commute
)
lemma exec_valid_as_disjoint:
fixes s
assumes "(vs \<inter> (prob_dom PROB) = {})" "(as \<in> valid_plans PROB)"
shows "(fmrestrict_set vs (exec_plan s as) = fmrestrict_set vs s)"
using assms
proof (induction as arbitrary: s vs PROB)
case (Cons a as)
then show ?case
by (metis exec_plan.simps(2) state_succ_valid_act_disjoint valid_plan_valid_head
valid_plan_valid_tail)
qed simp
definition state_successors where
"state_successors PROB s \<equiv> ((state_succ s ` PROB) - {s})"
subsection "State Spaces"
definition stateSpace where
"stateSpace ss vs \<equiv> (\<forall>s. s \<in> ss \<longrightarrow> (fmdom' s = vs))"
lemma EQ_SS_DOM:
assumes "\<not>(ss = {})" "(stateSpace ss vs1)" "(stateSpace ss vs2)"
shows "(vs1 = vs2)"
using assms
by (auto simp: stateSpace_def)
\<comment> \<open>NOTE Name 'dom' changed to 'domain' because of name clash with 'Map.dom'.\<close>
lemma FINITE_SS:
fixes ss :: "('a, bool) fmap set"
assumes "\<not>(ss = {})" "(stateSpace ss domain)"
shows "finite ss"
proof -
have 1: "stateSpace ss domain = (\<forall>s. s \<in> ss \<longrightarrow> (fmdom' s = domain))"
by (simp add: stateSpace_def)
{
fix s
assume P1: "s \<in> ss"
have "fmdom' s = domain"
using assms 1 P1
by blast
then have "s \<in> {s. fmdom' s = domain}"
by auto
}
then have 2: "ss \<subseteq> {s. fmdom' s = domain}"
by blast
\<comment> \<open>TODO add lemma (finite (fmdom' s))\<close>
then have "finite domain"
using 1 assms
by fastforce
then have "finite {s :: 'a state. fmdom' s = domain}"
using FINITE_states
by blast
then show ?thesis
using 2 finite_subset
by auto
qed
lemma disjoint_effects_no_effects:
fixes s
assumes "(\<forall>a. ListMem a as \<longrightarrow> (fmdom' (fmrestrict_set vs (snd a)) = {}))"
shows "(fmrestrict_set vs (exec_plan s as) = (fmrestrict_set vs s))"
using assms
proof (induction as arbitrary: s vs)
case (Cons a as)
then have "ListMem a (a # as)"
using elem
by fast
then have "fmdom' (fmrestrict_set vs (snd a)) = {}"
using Cons.prems(1)
by blast
then have "fmrestrict_set vs (state_succ s a) = fmrestrict_set vs s"
using sat_precond_as_proj_3
by blast
then show ?case
by (simp add: Cons.IH Cons.prems insert)
qed auto
subsection "Needed Asses"
\<comment> \<open>NOTE name shortened.\<close>
definition action_needed_vars where
"action_needed_vars a s \<equiv> {v. (v \<in> fmdom' s) \<and> (v \<in> fmdom' (fst a))
\<and> (fmlookup (fst a) v = fmlookup s v)}"
\<comment> \<open>NOTE name shortened to 'action\_needed\_asses'.\<close>
definition action_needed_asses where
"action_needed_asses a s \<equiv> fmrestrict_set (action_needed_vars a s) s"
\<comment> \<open>NOTE type for 'a' had to be fixed (type mismatch in goal).\<close>
lemma act_needed_asses_submap_succ_submap:
fixes a s1 s2
assumes "(action_needed_asses a s2 \<subseteq>\<^sub>f action_needed_asses a s1)" "(s1 \<subseteq>\<^sub>f s2)"
shows "(state_succ s1 a \<subseteq>\<^sub>f state_succ s2 a)"
using assms
unfolding state_succ_def
proof (auto)
assume P1: "action_needed_asses a s2 \<subseteq>\<^sub>f action_needed_asses a s1" "s1 \<subseteq>\<^sub>f s2" "fst a \<subseteq>\<^sub>f s1"
"fst a \<subseteq>\<^sub>f s2"
then show "snd a ++ s1 \<subseteq>\<^sub>f snd a ++ s2"
using submap_imp_state_succ_submap_b
by blast
next
assume P2: "action_needed_asses a s2 \<subseteq>\<^sub>f action_needed_asses a s1" "s1 \<subseteq>\<^sub>f s2" "fst a \<subseteq>\<^sub>f s1"
"\<not> fst a \<subseteq>\<^sub>f s2"
then show "snd a ++ s1 \<subseteq>\<^sub>f s2"
using submap_imp_state_succ_submap_a
by blast
next
assume P3: "action_needed_asses a s2 \<subseteq>\<^sub>f action_needed_asses a s1" "s1 \<subseteq>\<^sub>f s2" "\<not> fst a \<subseteq>\<^sub>f s1"
"fst a \<subseteq>\<^sub>f s2"
let ?vs1="{v \<in> fmdom' s1. v \<in> fmdom' (fst a) \<and> fmlookup (fst a) v = fmlookup s1 v}"
let ?vs2="{v \<in> fmdom' s2. v \<in> fmdom' (fst a) \<and> fmlookup (fst a) v = fmlookup s2 v}"
let ?f="fmrestrict_set ?vs1 s1"
let ?g="fmrestrict_set ?vs2 s2"
have 1: "fmdom' ?f = ?vs1" "fmdom' ?g = ?vs2"
unfolding action_needed_asses_def action_needed_vars_def fmdom'_restrict_set_precise
by blast+
have 2: "fmlookup ?g \<subseteq>\<^sub>m fmlookup ?f"
using P3(1)
unfolding action_needed_asses_def action_needed_vars_def
using fmsubset.rep_eq
by blast
{
{
fix v
assume P3_1: "v \<in> fmdom' ?g"
then have "v \<in> fmdom' s2" "v \<in> fmdom' (fst a)" "fmlookup (fst a) v = fmlookup s2 v"
using 1
by simp+
then have "fmlookup (fst a) v = fmlookup ?g v"
by simp
then have "fmlookup (fst a) v = fmlookup ?f v"
using 2
by (metis (mono_tags, lifting) P3_1 domIff fmdom'_notI map_le_def)
}
then have i: "fmlookup (fst a) \<subseteq>\<^sub>m fmlookup ?f"
using P3(4) 1(2)
by (smt domIff fmdom'_notD fmsubset.rep_eq map_le_def mem_Collect_eq)
{
fix v
assume P3_2: "v \<in> dom (fmlookup (fst a))"
then have "fmlookup (fst a) v = fmlookup ?f v"
using i
by (meson domIff fmdom'_notI map_le_def)
then have "v \<in> ?vs1"
using P3_2 1(1)
by (metis (no_types, lifting) domIff fmdom'_notD)
then have "fmlookup (fst a) v = fmlookup s1 v"
by blast
}
then have "fst a \<subseteq>\<^sub>f s1"
by (simp add: map_le_def fmsubset.rep_eq)
}
then show "s1 \<subseteq>\<^sub>f snd a ++ s2"
using P3(3)
by simp
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor.\<close>
lemma as_needed_asses_submap_exec_i:
fixes a s
assumes "v \<in> fmdom' (action_needed_asses a s)"
shows "
fmlookup (action_needed_asses a s) v = fmlookup s v
\<and> fmlookup (action_needed_asses a s) v = fmlookup (fst a) v"
using assms
unfolding action_needed_asses_def action_needed_vars_def
using fmdom'_notI fmlookup_restrict_set
by (smt mem_Collect_eq)
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor.\<close>
lemma as_needed_asses_submap_exec_ii:
fixes f g v
assumes "v \<in> fmdom' f" "f \<subseteq>\<^sub>f g"
shows "fmlookup f v = fmlookup g v"
using assms
by (meson fmdom'_notI fmdom_notD fmsubset_eq)
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor.\<close>
lemma as_needed_asses_submap_exec_iii:
fixes f g v
shows "
fmdom' (action_needed_asses a s)
= {v \<in> fmdom' s. v \<in> fmdom' (fst a) \<and> fmlookup (fst a) v = fmlookup s v}"
unfolding action_needed_asses_def action_needed_vars_def
by (simp add: Set.filter_def fmfilter_alt_defs(4))
\<comment> \<open>NOTE added lemma.\<close>
lemma as_needed_asses_submap_exec_iv:
fixes f a v
assumes "v \<in> fmdom' (action_needed_asses a s)"
shows "
fmlookup (action_needed_asses a s) v = fmlookup s v
\<and> fmlookup (action_needed_asses a s) v = fmlookup (fst a) v
\<and> fmlookup (fst a) v = fmlookup s v"
using assms
proof -
have 1: "v \<in> {v \<in> fmdom' s. v \<in> fmdom' (fst a) \<and> fmlookup (fst a) v = fmlookup s v}"
using assms as_needed_asses_submap_exec_iii
by metis
then have 2: "fmlookup (action_needed_asses a s) v = fmlookup s v"
unfolding action_needed_asses_def action_needed_vars_def
by force
moreover have 3: "fmlookup (action_needed_asses a s) v = fmlookup (fst a) v"
using 1 2
by simp
moreover have "fmlookup (fst a) v = fmlookup s v"
using 2 3
by argo
ultimately show ?thesis
by blast
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor (into Fmap\_Utils.thy).\<close>
lemma as_needed_asses_submap_exec_v:
fixes f g v
assumes "v \<in> fmdom' f" "f \<subseteq>\<^sub>f g"
shows "v \<in> fmdom' g"
proof -
obtain b where 1: "fmlookup f v = b" "b \<noteq> None"
using assms(1)
by (meson fmdom'_notI)
then have "fmlookup g v = b"
using as_needed_asses_submap_exec_ii[OF assms]
by argo
then show ?thesis
using 1 fmdom'_notD
by fastforce
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor.\<close>
lemma as_needed_asses_submap_exec_vi:
fixes a s1 s2 v
assumes "v \<in> fmdom' (action_needed_asses a s1)"
"(action_needed_asses a s1) \<subseteq>\<^sub>f (action_needed_asses a s2)"
shows
"(fmlookup (action_needed_asses a s1) v) = fmlookup (fst a) v
\<and> (fmlookup (action_needed_asses a s2) v) = fmlookup (fst a) v \<and>
fmlookup s1 v = fmlookup (fst a) v \<and> fmlookup s2 v = fmlookup (fst a) v"
using assms
proof -
have 1:
"fmlookup (action_needed_asses a s1) v = fmlookup s1 v"
"fmlookup (action_needed_asses a s1) v = fmlookup (fst a) v"
"fmlookup (fst a) v = fmlookup s1 v"
using as_needed_asses_submap_exec_iv[OF assms(1)]
by blast+
moreover {
have "fmlookup (action_needed_asses a s1) v = fmlookup (action_needed_asses a s2) v"
using as_needed_asses_submap_exec_ii[OF assms]
by simp
then have "fmlookup (action_needed_asses a s2) v = fmlookup (fst a) v"
using 1(2)
by argo
}
note 2 = this
moreover {
have "v \<in> fmdom' (action_needed_asses a s2)"
using as_needed_asses_submap_exec_v[OF assms]
by simp
then have "fmlookup s2 v = fmlookup (action_needed_asses a s2) v"
using as_needed_asses_submap_exec_i
by metis
also have "\<dots> = fmlookup (fst a) v"
using 2
by simp
finally have "fmlookup s2 v = fmlookup (fst a) v"
by simp
}
ultimately show ?thesis
by argo
qed
\<comment> \<open>TODO refactor.\<close>
\<comment> \<open>NOTE added lemma.\<close>
lemma as_needed_asses_submap_exec_vii:
fixes f g v
assumes "\<forall>v \<in> fmdom' f. fmlookup f v = fmlookup g v"
shows "f \<subseteq>\<^sub>f g"
proof -
{
fix v
assume a: "v \<in> fmdom' f"
then have "v \<in> dom (fmlookup f)"
by simp
moreover have "fmlookup f v = fmlookup g v"
using assms a
by blast
ultimately have "v \<in> dom (fmlookup f) \<longrightarrow> fmlookup f v = fmlookup g v"
by blast
}
then have "fmlookup f \<subseteq>\<^sub>m fmlookup g"
by (simp add: map_le_def)
then show ?thesis
by (simp add: fmsubset.rep_eq)
qed
\<comment> \<open>TODO refactor.\<close>
\<comment> \<open>NOTE added lemma.\<close>
lemma as_needed_asses_submap_exec_viii:
fixes f g v
assumes "f \<subseteq>\<^sub>f g"
shows "\<forall>v \<in> fmdom' f. fmlookup f v = fmlookup g v"
proof -
have 1: "fmlookup f \<subseteq>\<^sub>m fmlookup g"
using assms
by (simp add: fmsubset.rep_eq)
{
fix v
assume "v \<in> fmdom' f"
then have "v \<in> dom (fmlookup f)"
by simp
then have "fmlookup f v = fmlookup g v"
using 1 map_le_def
by metis
}
then show ?thesis
by blast
qed
\<comment> \<open>NOTE added lemma.\<close>
lemma as_needed_asses_submap_exec_viii':
fixes f g v
assumes "f \<subseteq>\<^sub>f g"
shows "fmdom' f \<subseteq> fmdom' g"
using assms as_needed_asses_submap_exec_v subsetI
by metis
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor.\<close>
lemma as_needed_asses_submap_exec_ix:
fixes f g
shows "f \<subseteq>\<^sub>f g = (\<forall>v \<in> fmdom' f. fmlookup f v = fmlookup g v)"
using as_needed_asses_submap_exec_vii as_needed_asses_submap_exec_viii
by metis
\<comment> \<open>NOTE added lemma.\<close>
lemma as_needed_asses_submap_exec_x:
fixes f a v
assumes "v \<in> fmdom' (action_needed_asses a f)"
shows "v \<in> fmdom' (fst a) \<and> v \<in> fmdom' f \<and> fmlookup (fst a) v = fmlookup f v"
using assms
unfolding action_needed_asses_def action_needed_vars_def
using as_needed_asses_submap_exec_i assms
by (metis fmdom'_notD fmdom'_notI)
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor.\<close>
lemma as_needed_asses_submap_exec_xi:
fixes v a f g
assumes "v \<in> fmdom' (action_needed_asses a (f ++ g))" "v \<in> fmdom' f"
shows "
fmlookup (action_needed_asses a (f ++ g)) v = fmlookup f v
\<and> fmlookup (action_needed_asses a (f ++ g)) v = fmlookup (fst a) v"
proof -
have 1: "v \<in> {v \<in> fmdom' (f ++ g). v \<in> fmdom' (fst a) \<and> fmlookup (fst a) v = fmlookup (f ++ g) v}"
using as_needed_asses_submap_exec_x[OF assms(1)]
by blast
{
have "v |\<in>| fmdom f"
using assms(2)
by (meson fmdom'_notI fmdom_notD)
then have "fmlookup (f ++ g) v = fmlookup f v"
unfolding fmap_add_ltr_def fmlookup_add
by simp
}
note 2 = this
{
have "fmlookup (action_needed_asses a (f ++ g)) v = fmlookup (f ++ g) v"
unfolding action_needed_asses_def action_needed_vars_def
using 1
by force
then have "fmlookup (action_needed_asses a (f ++ g)) v = fmlookup f v"
using 2
by simp
}
note 3 = this
moreover {
have "fmlookup (fst a) v = fmlookup (f ++ g) v"
using 1
by simp
also have "\<dots> = fmlookup f v"
using 2
by simp
also have "\<dots> = fmlookup (action_needed_asses a (f ++ g)) v"
using 3
by simp
finally have "fmlookup (action_needed_asses a (f ++ g)) v = fmlookup (fst a) v"
by simp
}
ultimately show ?thesis
by blast
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor (into Fmap\_Utils.thy).\<close>
lemma as_needed_asses_submap_exec_xii:
fixes f g v
assumes "v \<in> fmdom' f"
shows "fmlookup (f ++ g) v = fmlookup f v"
proof -
have "v |\<in>| fmdom f"
using assms(1) fmdom'_notI fmdom_notD
by metis
then show ?thesis
unfolding fmap_add_ltr_def
using fmlookup_add
by force
qed
\<comment> \<open>NOTE added lemma.\<close>
lemma as_needed_asses_submap_exec_xii':
fixes f g v
assumes "v \<notin> fmdom' f" "v \<in> fmdom' g"
shows "fmlookup (f ++ g) v = fmlookup g v"
proof -
have "\<not>(v |\<in>| fmdom f)"
using assms(1) fmdom'_notI fmdom_notD
by fastforce
moreover have "v |\<in>| fmdom g"
using assms(2) fmdom'_notI fmdom_notD
by metis
ultimately show ?thesis
unfolding fmap_add_ltr_def
using fmlookup_add
by simp
qed
\<comment> \<open>NOTE showcase.\<close>
lemma as_needed_asses_submap_exec:
fixes s1 s2
assumes "(s1 \<subseteq>\<^sub>f s2)"
"(\<forall>a. ListMem a as \<longrightarrow> (action_needed_asses a s2 \<subseteq>\<^sub>f action_needed_asses a s1))"
shows "(exec_plan s1 as \<subseteq>\<^sub>f exec_plan s2 as)"
using assms
proof (induction as arbitrary: s1 s2)
case (Cons a as)
\<comment> \<open>Proof the premises of the induction hypothesis for 'state\_succ s1 a' and 'state\_succ s2 a'.\<close>
{
then have "action_needed_asses a s2 \<subseteq>\<^sub>f action_needed_asses a s1"
using Cons.prems(2) elem
by metis
then have "state_succ s1 a \<subseteq>\<^sub>f state_succ s2 a"
using Cons.prems(1) act_needed_asses_submap_succ_submap
by blast
}
note 1 = this
moreover {
fix a'
assume P: "ListMem a' as"
\<comment> \<open>Show the goal by rule 'as\_needed\_asses\_submap\_exec\_ix'.\<close>
let ?f="action_needed_asses a' (state_succ s2 a)"
let ?g="action_needed_asses a' (state_succ s1 a)"
{
fix v
assume P_1: "v \<in> fmdom' ?f"
then have "fmlookup ?f v = fmlookup ?g v"
unfolding state_succ_def
text \<open> Split cases on the if-then branches introduced by the definition of 'state\_succ'.\<close>
proof (auto)
assume P_1_1: "v \<in> fmdom' (action_needed_asses a' (snd a ++ s2))" "fst a \<subseteq>\<^sub>f s2"
"fst a \<subseteq>\<^sub>f s1"
have i: "action_needed_asses a' s2 \<subseteq>\<^sub>f action_needed_asses a' s1"
using Cons.prems(2) P insert
by fast
then show "
fmlookup (action_needed_asses a' (snd a ++ s2)) v
= fmlookup (action_needed_asses a' (snd a ++ s1)) v"
proof (cases "v \<in> fmdom' ?g")
case true: True
then have A:
"v \<in> fmdom' (fst a') \<and> v \<in> fmdom' (snd a ++ s1)
\<and> fmlookup (fst a') v = fmlookup (snd a ++ s1) v"
using as_needed_asses_submap_exec_x[OF true]
unfolding state_succ_def
using P_1_1(3)
by simp
then have B:
"v \<in> fmdom' (fst a') \<and> v \<in> fmdom' (snd a ++ s2)
\<and> fmlookup (fst a') v = fmlookup (snd a ++ s2) v"
using as_needed_asses_submap_exec_x[OF P_1]
unfolding state_succ_def
using P_1_1(2)
by simp
then show ?thesis
proof (cases "v \<in> fmdom' (snd a)")
case True
then have I:
"fmlookup (snd a ++ s2) v = fmlookup (snd a) v"
"fmlookup (snd a ++ s1) v = fmlookup (snd a) v"
using as_needed_asses_submap_exec_xii
by fast+
moreover {
have "fmlookup ?f v = fmlookup (snd a ++ s2) v"
using as_needed_asses_submap_exec_iv[OF P_1]
unfolding state_succ_def
using P_1_1(2)
by presburger
then have "fmlookup ?f v = fmlookup (snd a) v"
using I(1)
by argo
}
moreover {
have "fmlookup ?g v = fmlookup (snd a ++ s1) v"
using as_needed_asses_submap_exec_iv[OF true]
unfolding state_succ_def
using P_1_1(3)
by presburger
then have "fmlookup ?g v = fmlookup (snd a) v"
using I(2)
by argo
}
ultimately show ?thesis
unfolding state_succ_def
using P_1_1(2, 3)
by presburger
next
case False
then have I: "v \<in> fmdom' s1" "v \<in> fmdom' s2"
using A B
unfolding fmap_add_ltr_def fmdom'_add
by blast+
{
have "fmlookup ?g v = fmlookup (snd a ++ s1) v"
using as_needed_asses_submap_exec_iv[OF true]
unfolding state_succ_def
using P_1_1(3)
by presburger
then have "fmlookup ?g v = fmlookup s1 v"
using as_needed_asses_submap_exec_xii'[OF False I(1)]
by simp
moreover {
have "fmlookup (snd a ++ s1) v = fmlookup s1 v"
using as_needed_asses_submap_exec_xii'[OF False I(1)]
by simp
moreover from \<open>fmlookup (snd a ++ s1) v = fmlookup s1 v\<close>
have "fmlookup (fst a') v = fmlookup s1 v"
using A(1)
by argo
ultimately have "fmlookup (action_needed_asses a' s1) v = fmlookup s1 v"
using A(1) I(1)
unfolding action_needed_asses_def action_needed_vars_def
fmlookup_restrict_set
by simp
}
ultimately have "fmlookup ?g v = fmlookup (action_needed_asses a' s1) v"
by argo
}
note II = this
{
have "fmlookup ?f v = fmlookup (snd a ++ s2) v"
using as_needed_asses_submap_exec_iv[OF P_1]
unfolding state_succ_def
using P_1_1(2)
by presburger
moreover from \<open>fmlookup ?f v = fmlookup (snd a ++ s2) v\<close>
have \<alpha>: "fmlookup ?f v = fmlookup s2 v"
using as_needed_asses_submap_exec_xii'[OF False I(2)]
by argo
ultimately have "fmlookup (snd a ++ s2) v = fmlookup s2 v"
by argo
moreover {
from \<open>fmlookup (snd a ++ s2) v = fmlookup s2 v\<close>
have "fmlookup (fst a') v = fmlookup s2 v"
using B(1)
by argo
then have "fmlookup (action_needed_asses a' s2) v = fmlookup s2 v"
using B(1) I(2)
unfolding action_needed_asses_def action_needed_vars_def
fmlookup_restrict_set
by simp
}
ultimately have "fmlookup ?f v = fmlookup (action_needed_asses a' s2) v"
using \<alpha>
by argo
}
note III = this
{
have "v \<in> fmdom' (action_needed_asses a' s2)"
proof -
have "fmlookup (fst a') v = fmlookup s1 v"
by (simp add: A False I(1) as_needed_asses_submap_exec_xii')
then show ?thesis
by (simp add: A Cons.prems(1) I(1, 2)
as_needed_asses_submap_exec_ii as_needed_asses_submap_exec_iii)
qed
then have "
fmlookup (action_needed_asses a' s2) v
= fmlookup (action_needed_asses a' s1) v"
using i as_needed_asses_submap_exec_ix[of "action_needed_asses a' s2"
"action_needed_asses a' s1"]
by blast
}
note IV = this
{
have "fmlookup ?f v = fmlookup (action_needed_asses a' s2) v"
using III
by simp
also have "\<dots> = fmlookup (action_needed_asses a' s1) v"
using IV
by simp
finally have "\<dots> = fmlookup ?g v"
using II
by simp
}
then show ?thesis
unfolding action_needed_asses_def action_needed_vars_def state_succ_def
using P_1_1 A B
by simp
qed
next
case false: False
have A:
"v \<in> fmdom' (fst a') \<and> v \<in> fmdom' (snd a ++ s2)
\<and> fmlookup (fst a') v = fmlookup (snd a ++ s2) v"
using as_needed_asses_submap_exec_x[OF P_1]
unfolding state_succ_def
using P_1_1(2)
by simp
from false have B:
"\<not>(v \<in> fmdom' (snd a ++ s1)) \<or> \<not>(fmlookup (fst a') v = fmlookup (snd a ++ s1) v)"
by (simp add: A P_1_1(3) as_needed_asses_submap_exec_iii state_succ_def)
then show ?thesis
proof (cases "v \<in> fmdom' (snd a)")
case True
then have I: "v \<in> fmdom' (snd a ++ s1)"
unfolding fmap_add_ltr_def fmdom'_add
by simp
{
from True have
"fmlookup (snd a ++ s2) v = fmlookup (snd a) v"
"fmlookup (snd a ++ s1) v = fmlookup (snd a) v"
using as_needed_asses_submap_exec_xii
by fast+
then have "fmlookup (snd a ++ s1) v = fmlookup (snd a ++ s2) v"
by auto
also have " \<dots> = fmlookup (fst a') v"
using A
by simp
finally have "fmlookup (snd a ++ s1) v = fmlookup (fst a') v"
by simp
}
then show ?thesis using B I
by presburger
next
case False
then have I: "v \<in> fmdom' s2"
using A unfolding fmap_add_ltr_def fmdom'_add
by blast
{
from P_1 have "fmlookup ?f v \<noteq> None"
by (meson fmdom'_notI)
moreover from false
have "fmlookup ?g v = None"
by (simp add: fmdom'_notD)
ultimately have "fmlookup ?f v \<noteq> fmlookup ?g v"
by simp
}
moreover
{
{
from P_1_1(2) have "state_succ s2 a = snd a ++ s2"
unfolding state_succ_def
by simp
moreover from \<open>state_succ s2 a = snd a ++ s2\<close> have
"fmlookup (state_succ s2 a) v = fmlookup s2 v"
using as_needed_asses_submap_exec_xii'[OF False I]
by simp
ultimately have "fmlookup ?f v = fmlookup (action_needed_asses a' s2) v"
unfolding action_needed_asses_def action_needed_vars_def
by (simp add: A I)
}
note I = this
moreover {
from P_1_1(3) have "state_succ s1 a = snd a ++ s1"
unfolding state_succ_def
by simp
moreover from \<open>state_succ s1 a = snd a ++ s1\<close> False
have "fmlookup (state_succ s1 a) v = fmlookup s1 v"
unfolding fmap_add_ltr_def
using fmlookup_add
- by (simp add: fmdom'_alt_def fmember_iff_member_fset)
+ by (simp add: fmdom'_alt_def)
ultimately have "fmlookup ?g v = fmlookup (action_needed_asses a' s1) v"
unfolding action_needed_asses_def action_needed_vars_def
using FDOM_state_succ_subset
by auto
}
moreover {
have "v \<in> fmdom' (action_needed_asses a' s2)"
proof -
have "v \<in> fmdom' s2 \<union> fmdom' (snd a)"
by (metis (no_types) A FDOM_state_succ_subset P_1_1(2) state_succ_def subsetCE)
then show ?thesis
by (simp add: A False as_needed_asses_submap_exec_iii as_needed_asses_submap_exec_xii')
qed
then have "
fmlookup (action_needed_asses a' s2) v
= fmlookup (action_needed_asses a' s1) v"
using i as_needed_asses_submap_exec_ix[of "action_needed_asses a' s2"
"action_needed_asses a' s1"]
by blast
}
ultimately have "fmlookup ?f v = fmlookup ?g v"
by simp
}
ultimately show ?thesis
by simp
qed
qed
next
assume P2: "v \<in> fmdom' (action_needed_asses a' (snd a ++ s2))" "fst a \<subseteq>\<^sub>f s2"
"\<not> fst a \<subseteq>\<^sub>f s1"
then show "
fmlookup (action_needed_asses a' (snd a ++ s2)) v
= fmlookup (action_needed_asses a' s1) v"
proof -
obtain aa :: "('a, 'b) fmap \<Rightarrow> ('a, 'b) fmap \<Rightarrow> 'a" where
"\<forall>x0 x1. (\<exists>v2. v2 \<in> fmdom' x1
\<and> fmlookup x1 v2 \<noteq> fmlookup x0 v2) = (aa x0 x1 \<in> fmdom' x1
\<and> fmlookup x1 (aa x0 x1) \<noteq> fmlookup x0 (aa x0 x1))"
by moura
then have f1: "\<forall>f fa. aa fa f \<in> fmdom' f
\<and> fmlookup f (aa fa f) \<noteq> fmlookup fa (aa fa f) \<or> f \<subseteq>\<^sub>f fa"
by (meson as_needed_asses_submap_exec_vii)
then have f2: "aa s1 (fst a) \<in> fmdom' (fst a)
\<and> fmlookup (fst a) (aa s1 (fst a)) \<noteq> fmlookup s1 (aa s1 (fst a))"
using P2(3) by blast
then have "aa s1 (fst a) \<in> fmdom' s2"
by (metis (full_types) P2(2) as_needed_asses_submap_exec_v)
then have "aa s1 (fst a) \<in> fmdom' (action_needed_asses a s2)"
using f2 by (simp add: P2(2) as_needed_asses_submap_exec_iii
as_needed_asses_submap_exec_viii)
then show ?thesis
using f1 by (metis (no_types) Cons.prems(2) P2(3) as_needed_asses_submap_exec_vi elem)
qed
next
assume P3: "v \<in> fmdom' (action_needed_asses a' s2)" "\<not> fst a \<subseteq>\<^sub>f s2" "fst a \<subseteq>\<^sub>f s1"
then show "
fmlookup (action_needed_asses a' s2) v
= fmlookup (action_needed_asses a' (snd a ++ s1)) v"
using Cons.prems(1) submap_imp_state_succ_submap_a
by blast
next
assume P4: "v \<in> fmdom' (action_needed_asses a' s2)" "\<not> fst a \<subseteq>\<^sub>f s2" "\<not> fst a \<subseteq>\<^sub>f s1"
then show "
fmlookup (action_needed_asses a' s2) v
= fmlookup (action_needed_asses a' s1) v"
by (simp add: Cons.prems(2) P as_needed_asses_submap_exec_ii insert)
qed
}
then have a: "?f \<subseteq>\<^sub>f ?g"
using as_needed_asses_submap_exec_ix
by blast
}
note 2 = this
then show ?case
unfolding exec_plan.simps
using Cons.IH[of "state_succ s1 a" "state_succ s2 a", OF 1]
by blast
qed simp
\<comment> \<open>NOTE name shortened.\<close>
definition system_needed_vars where
"system_needed_vars PROB s \<equiv> (\<Union>{action_needed_vars a s | a. a \<in> PROB})"
\<comment> \<open>NOTE name shortened.\<close>
definition system_needed_asses where
"system_needed_asses PROB s \<equiv> (fmrestrict_set (system_needed_vars PROB s) s)"
lemma action_needed_vars_subset_sys_needed_vars_subset:
assumes "(a \<in> PROB)"
shows "(action_needed_vars a s \<subseteq> system_needed_vars PROB s)"
using assms
by (auto simp: system_needed_vars_def) (metis surjective_pairing)
lemma action_needed_asses_submap_sys_needed_asses:
assumes "(a \<in> PROB)"
shows "(action_needed_asses a s \<subseteq>\<^sub>f system_needed_asses PROB s)"
proof -
have "action_needed_asses a s = fmrestrict_set (action_needed_vars a s) s"
unfolding action_needed_asses_def
by simp
then have "system_needed_asses PROB s = (fmrestrict_set (system_needed_vars PROB s) s)"
unfolding system_needed_asses_def
by simp
then have 1: "action_needed_vars a s \<subseteq> system_needed_vars PROB s"
unfolding action_needed_vars_subset_sys_needed_vars_subset
using assms action_needed_vars_subset_sys_needed_vars_subset
by fast
{
fix x
assume P1: "x \<in> dom (fmlookup (fmrestrict_set (action_needed_vars a s) s))"
then have a: "fmlookup (fmrestrict_set (action_needed_vars a s) s) x = fmlookup s x"
by (auto simp: fmdom'_restrict_set_precise)
then have "fmlookup (fmrestrict_set (system_needed_vars PROB s) s) x = fmlookup s x"
using 1 contra_subsetD
by fastforce
then have "
fmlookup (fmrestrict_set (action_needed_vars a s) s) x
= fmlookup (fmrestrict_set (system_needed_vars PROB s) s) x
"
using a
by argo
}
then have "
fmlookup (fmrestrict_set (action_needed_vars a s) s)
\<subseteq>\<^sub>m fmlookup (fmrestrict_set (system_needed_vars PROB s) s)
"
using map_le_def
by blast
then show "(action_needed_asses a s \<subseteq>\<^sub>f system_needed_asses PROB s)"
by (simp add: fmsubset.rep_eq action_needed_asses_def system_needed_asses_def)
qed
lemma system_needed_asses_include_action_needed_asses_1:
assumes "(a \<in> PROB)"
shows "(action_needed_vars a (fmrestrict_set (system_needed_vars PROB s) s) = action_needed_vars a s)"
proof -
let ?A="{v \<in> fmdom' (fmrestrict_set (system_needed_vars PROB s) s).
v \<in> fmdom' (fst a)
\<and> fmlookup (fst a) v = fmlookup (fmrestrict_set (system_needed_vars PROB s) s) v}"
let ?B="{v \<in> fmdom' s. v \<in> fmdom' (fst a) \<and> fmlookup (fst a) v = fmlookup s v}"
{
fix v
assume "v \<in> ?A"
then have i: "v \<in> fmdom' (fmrestrict_set (system_needed_vars PROB s) s)" "v \<in> fmdom' (fst a)"
"fmlookup (fst a) v = fmlookup (fmrestrict_set (system_needed_vars PROB s) s) v"
by blast+
then have "v \<in> fmdom' s"
by (simp add: fmdom'_restrict_set_precise)
moreover have "fmlookup (fst a) v = fmlookup s v"
using i(2, 3) fmdom'_notI
by force
ultimately have "v \<in> ?B"
using i
by blast
}
then have 1: "?A \<subseteq> ?B"
by blast
{
fix v
assume P: "v \<in> ?B"
then have ii: "v \<in> fmdom' s" "v \<in> fmdom' (fst a)" "fmlookup (fst a) v = fmlookup s v"
by blast+
moreover {
have "\<exists>s'. v \<in> s' \<and> (\<exists>a. (s' = action_needed_vars a s) \<and> a \<in> PROB)"
unfolding action_needed_vars_def
using assms P action_needed_vars_def
by metis
then obtain s' where \<alpha>: "v \<in> s'" "(\<exists>a. (s' = action_needed_vars a s) \<and> a \<in> PROB)"
by blast
moreover obtain a' where "s' = action_needed_vars a' s" "a' \<in> PROB"
using \<alpha>
by blast
ultimately have "v \<in> fmdom' (fmrestrict_set (system_needed_vars PROB s) s)"
unfolding fmdom'_restrict_set_precise
using action_needed_vars_subset_sys_needed_vars_subset ii(1) by blast
}
note iii = this
moreover have "fmlookup (fst a) v = fmlookup (fmrestrict_set (system_needed_vars PROB s) s) v"
using ii(3) iii fmdom'_notI
by force
ultimately have "v \<in> ?A"
by blast
}
then have "?B \<subseteq> ?A"
by blast
then show ?thesis
unfolding action_needed_vars_def
using 1
by blast
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor (proven elsewhere?).\<close>
lemma system_needed_asses_include_action_needed_asses_i:
fixes A B f
assumes "A \<subseteq> B"
shows "fmrestrict_set A (fmrestrict_set B f) = fmrestrict_set A f"
proof -
{
let ?f'="fmrestrict_set A f"
let ?f''="fmrestrict_set A (fmrestrict_set B f)"
assume C: "?f'' \<noteq> ?f'"
then obtain v where 1: "fmlookup ?f'' v \<noteq> fmlookup ?f' v"
by (meson fmap_ext)
then have False
proof (cases "v \<in> A")
case True
have "fmlookup ?f'' v = fmlookup (fmrestrict_set B f) v"
using True fmlookup_restrict_set
by simp
moreover have "fmlookup (fmrestrict_set B f) v = fmlookup ?f' v"
using True assms(1)
by auto
ultimately show ?thesis
using 1
by argo
next
case False
then have "fmlookup ?f' v = None" "fmlookup ?f'' v = None"
using fmlookup_restrict_set
by auto+
then show ?thesis
using 1
by argo
qed
}
then show ?thesis
by blast
qed
lemma system_needed_asses_include_action_needed_asses:
assumes "(a \<in> PROB)"
shows "(action_needed_asses a (system_needed_asses PROB s) = action_needed_asses a s)"
proof -
{
have " action_needed_vars a s \<subseteq> system_needed_vars PROB s"
using action_needed_vars_subset_sys_needed_vars_subset[OF assms]
by simp
then have "
fmrestrict_set (action_needed_vars a s) (fmrestrict_set (system_needed_vars PROB s) s) =
fmrestrict_set (action_needed_vars a s) s"
using system_needed_asses_include_action_needed_asses_i
by fast
}
moreover
{
have
"action_needed_vars a (fmrestrict_set (system_needed_vars PROB s) s) = action_needed_vars a s"
using system_needed_asses_include_action_needed_asses_1[OF assms]
by simp
then have "fmrestrict_set (action_needed_vars a (fmrestrict_set (system_needed_vars PROB s) s))
(fmrestrict_set (system_needed_vars PROB s) s) =
fmrestrict_set (action_needed_vars a s) s
\<longleftrightarrow> fmrestrict_set (action_needed_vars a s) (fmrestrict_set (system_needed_vars PROB s) s) =
fmrestrict_set (action_needed_vars a s) s"
by simp
}
ultimately show ?thesis
unfolding action_needed_asses_def system_needed_asses_def
by simp
qed
lemma system_needed_asses_submap:
"system_needed_asses PROB s \<subseteq>\<^sub>f s"
proof -
{
fix x
assume P: "x\<in> dom (fmlookup (system_needed_asses PROB s))"
then have "system_needed_asses PROB s = (fmrestrict_set (system_needed_vars PROB s) s)"
by (simp add: system_needed_asses_def)
then have "fmlookup (system_needed_asses PROB s) x = fmlookup s x"
using P
by (auto simp: fmdom'_restrict_set_precise)
}
then have "fmlookup (system_needed_asses PROB s) \<subseteq>\<^sub>m fmlookup s"
using map_le_def
by blast
then show ?thesis
using fmsubset.rep_eq
by fast
qed
lemma as_works_from_system_needed_asses:
assumes "(as \<in> valid_plans PROB)"
shows "(exec_plan (system_needed_asses PROB s) as \<subseteq>\<^sub>f exec_plan s as)"
using assms
by (metis
action_needed_asses_def
as_needed_asses_submap_exec
fmsubset_restrict_set_mono system_needed_asses_def
system_needed_asses_include_action_needed_asses
system_needed_asses_include_action_needed_asses_1
system_needed_asses_submap
valid_plan_mems
)
end
\ No newline at end of file
diff --git a/thys/Factored_Transition_System_Bounding/SystemAbstraction.thy b/thys/Factored_Transition_System_Bounding/SystemAbstraction.thy
--- a/thys/Factored_Transition_System_Bounding/SystemAbstraction.thy
+++ b/thys/Factored_Transition_System_Bounding/SystemAbstraction.thy
@@ -1,2687 +1,2687 @@
theory SystemAbstraction
imports
Main
"HOL-Library.Sublist"
"HOL-Library.Finite_Map"
FactoredSystem
FactoredSystemLib
ActionSeqProcess
Dependency
TopologicalProps
FmapUtils
ListUtils
begin
\<comment> \<open>NOTE hide 'Map.map\_add' because of conflicting notation with 'FactoredSystemLib.map\_add\_ltr'.\<close>
hide_const (open) Map.map_add
no_notation Map.map_add (infixl "++" 100)
section "System Abstraction"
text \<open> Projection of an object (state, action, sequence of action or factored representation)
to a variable set `vs` restricts the domain of the object or its components---in case of composite
objects---to `vs`. [Abdulaziz et al., p.12]
This section presents the relevant definitions (`action\_proj`, `as\_proj`, `prob\_proj` and `ss\_proj`)
as well as their characterization. \<close>
subsection "Projection of Actions, Sequences of Actions and Factored Representations."
definition action_proj where
"action_proj a vs \<equiv> (fmrestrict_set vs (fst a), fmrestrict_set vs (snd a))"
lemma action_proj_pair: "action_proj (p, e) vs = (fmrestrict_set vs p, fmrestrict_set vs e)"
unfolding action_proj_def
by simp
definition prob_proj where
"prob_proj PROB vs \<equiv> (\<lambda>a. action_proj a vs) ` PROB"
\<comment> \<open>NOTE using 'fun' due to multiple defining equations.\<close>
\<comment> \<open>NOTE name shortened.\<close>
fun as_proj where
"as_proj [] _ = []"
| "as_proj (a # as) vs = (if fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}
then action_proj a vs # as_proj as vs
else as_proj as vs
)"
\<comment> \<open>TODO the lemma might be superfluous (follows directly from 'as\_proj.simps').\<close>
lemma as_proj_pair:
"as_proj ((p, e) # as) vs = (if (fmdom' (fmrestrict_set vs e) \<noteq> {})
then action_proj (p, e) vs # as_proj as vs
else as_proj as vs
)"
"as_proj [] vs = []"
by (simp)+
lemma proj_state_succ:
fixes s a vs
assumes "(fst a \<subseteq>\<^sub>f s)"
shows "(state_succ (fmrestrict_set vs s) (action_proj a vs) = fmrestrict_set vs (state_succ s a))"
proof -
have "
fmrestrict_set vs (if fst a \<subseteq>\<^sub>f s then snd a ++ s else s)
= fmrestrict_set vs (snd a ++ s)
"
using assms
by simp
moreover
{
assume "fst (action_proj a vs) \<subseteq>\<^sub>f fmrestrict_set vs s"
then have "
(state_succ (fmrestrict_set vs s) (action_proj a vs)
= fmrestrict_set vs (snd a ++ s))
"
unfolding state_succ_def action_proj_def fmap_add_ltr_def
by force
}
moreover {
assume "\<not>(fst (action_proj a vs) \<subseteq>\<^sub>f fmrestrict_set vs s)"
then have "
(state_succ (fmrestrict_set vs s) (action_proj a vs)
= fmrestrict_set vs (snd a ++ s))
"
unfolding state_succ_def action_proj_def
using assms fmsubset_restrict_set_mono
by auto
}
ultimately show ?thesis
unfolding state_succ_def
by argo
qed
lemma graph_plan_lemma_1:
fixes s vs as
assumes "sat_precond_as s as"
shows "(exec_plan (fmrestrict_set vs s) (as_proj as vs) = (fmrestrict_set vs (exec_plan s as)))"
using assms
proof (induction as arbitrary: s vs)
case (Cons a as)
then show ?case
proof (cases "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}")
case True
then have
"state_succ (fmrestrict_set vs s) (action_proj a vs) = fmrestrict_set vs (state_succ s a)"
using Cons.prems proj_state_succ
by fastforce
then show ?thesis
unfolding exec_plan.simps sat_precond_as.simps as_proj.simps
using Cons.IH Cons.prems True
by simp
next
case False
then have "(fmdom' (snd a) \<inter> vs = {})"
using False fmdom'_restrict_set_precise[of vs "snd a"]
by argo
then have "fmrestrict_set vs s = fmrestrict_set vs (state_succ s a)"
using disj_imp_eq_proj_exec
by blast
then show ?thesis
unfolding exec_plan.simps sat_precond_as.simps as_proj.simps
using Cons.IH Cons.prems False
by simp
qed
qed simp
\<comment> \<open>TODO the proofs are inefficient (detailed proofs?).\<close>
lemma proj_action_dom_eq_inter:
shows "
action_dom (fst (action_proj a vs)) (snd (action_proj a vs))
= (action_dom (fst a) (snd a) \<inter> vs)
"
unfolding action_dom_def action_proj_def
by (auto simp: fmdom'_restrict_set_precise)
lemma graph_plan_neq_mems_state_set_neq_len:
shows "prob_dom (prob_proj PROB vs) = (prob_dom PROB \<inter> vs)"
proof -
have "
prob_dom (prob_proj PROB vs)
= (
\<Union>(s1, s2)\<in>(\<lambda>a. (fmrestrict_set vs (fst a), fmrestrict_set vs (snd a)))
` PROB. action_dom s1 s2
)
"
unfolding prob_dom_def prob_proj_def action_proj_def
by blast
moreover
{
have "
(prob_dom PROB \<inter> vs)
= (\<Union>a\<in>PROB. action_dom (fst a) (snd a) \<inter> vs)
"
unfolding prob_dom_def prob_proj_def
using SUP_cong
by auto
also have "\<dots> = (\<Union>a\<in>PROB. action_dom (fst (action_proj a vs)) (snd (action_proj a vs)))"
using proj_action_dom_eq_inter[symmetric]
by fast
finally have "
(prob_dom PROB \<inter> vs)
= (\<Union>a\<in>PROB. fmdom' (fmrestrict_set vs (fst a)) \<union> fmdom' (fmrestrict_set vs (snd a)))
"
unfolding action_dom_def action_proj_def
by simp
}
ultimately show ?thesis
by (metis (mono_tags, lifting) SUP_cong UN_simps(10) action_dom_def case_prod_beta' prod.sel(1)
snd_conv)
qed
\<comment> \<open>TODO more detailed proof.\<close>
lemma graph_plan_not_eq_last_diff_paths:
fixes PROB vs
assumes "(s \<in> valid_states PROB)"
shows "((fmrestrict_set vs s) \<in> valid_states (prob_proj PROB vs))"
unfolding valid_states_def
using graph_plan_neq_mems_state_set_neq_len
by (metis (mono_tags, lifting)
assms fmdom'.rep_eq fmlookup_fmrestrict_set_dom inf_commute mem_Collect_eq valid_states_def)
lemma dom_eff_subset_imp_dom_succ_eq_proj:
fixes h s vs
assumes "(fmdom' (snd h) \<subseteq> fmdom' s)"
shows "(fmdom' (state_succ s (action_proj h vs)) = fmdom' (state_succ s h))"
proof (cases "fst (fmrestrict_set vs (fst h), fmrestrict_set vs (snd h)) \<subseteq>\<^sub>f s")
case true: True
then show ?thesis
proof (cases "fst h \<subseteq>\<^sub>f s")
case True
then show ?thesis
unfolding state_succ_def action_proj_def
using true True
by simp (smt assms fmap_add_ltr_def fmdom'.rep_eq fmdom'_add fmlookup_fmrestrict_set_dom
inf.absorb_iff2 inf.left_commute sup.absorb_iff1)
next
case False
then show ?thesis
unfolding state_succ_def action_proj_def
using true False
by simp (metis (no_types) assms dual_order.trans fmap_add_ltr_def fmdom'.rep_eq fmdom'_add
fmlookup_fmrestrict_set_dom inf_le2 sup.absorb_iff1)
qed
next
case False
then have "fmdom' s = fmdom' (if fst h \<subseteq>\<^sub>f s then snd h ++ s else s)"
using sat_precond_as_proj_4
by auto
then show ?thesis
unfolding state_succ_def action_proj_def
using False
by presburger
qed
lemma drest_proj_succ_eq_drest_succ:
fixes h s vs
assumes "fst h \<subseteq>\<^sub>f s" "(fmdom' (snd h) \<subseteq> fmdom' s)"
shows "(fmrestrict_set vs (state_succ s (action_proj h vs)) = fmrestrict_set vs (state_succ s h))"
proof -
{
have 1: "fmrestrict_set vs (fst h) \<subseteq>\<^sub>f s"
using assms(1) submap_imp_state_succ_submap_a
by (simp add: sat_precond_as_proj_4)
then have "
fmrestrict_set vs (state_succ s (action_proj h vs))
= fmrestrict_set vs (fmrestrict_set vs (snd h) ++ s)
"
unfolding state_succ_def action_proj_def
by simp
also have "\<dots> = fmrestrict_set vs s ++\<^sub>f fmrestrict_set vs (fmrestrict_set vs (snd h))"
unfolding fmap_add_ltr_def
by simp
\<comment> \<open>TODO
refactor the step 'fmrestrict\_set ?X (fmrestrict\_set ?X ?f) = fmrestrict\_set ?X ?f' into
own lemma in 'FmapUtils.thy'.\<close>
also have "\<dots> = fmrestrict_set vs s ++\<^sub>f fmrestrict_set vs (snd h)"
using fmfilter_alt_defs(4) fmfilter_cong fmlookup_filter fmrestrict_set_dom option.simps(3)
by metis
finally have "
fmrestrict_set vs (state_succ s (action_proj h vs))
= fmrestrict_set vs (snd h ++ s)
"
unfolding fmap_add_ltr_def
by simp
}
moreover have "fmrestrict_set vs (state_succ s h) = fmrestrict_set vs ((snd h) ++ s)"
unfolding state_succ_def
using assms(1)
by simp
ultimately show ?thesis
by simp
qed
\<comment> \<open>TODO remove? This is equivalent to 'proj\_state\_succ'.\<close>
lemma drest_succ_proj_eq_drest_succ:
fixes s vs as
assumes "(fst a \<subseteq>\<^sub>f s)"
shows "(state_succ (fmrestrict_set vs s) (action_proj a vs) = fmrestrict_set vs (state_succ s a))"
using assms proj_state_succ
by blast
lemma exec_drest_cons_proj_eq_succ:
fixes as PROB vs a
assumes "fst a \<subseteq>\<^sub>f s"
shows "(
exec_plan (fmrestrict_set vs s) (action_proj a vs # as)
= exec_plan (fmrestrict_set vs (state_succ s a)) as
)"
proof -
have "exec_plan (state_succ (fmrestrict_set vs s) (action_proj a vs)) as =
exec_plan (fmrestrict_set vs (state_succ s a)) as"
using assms drest_succ_proj_eq_drest_succ
by metis
then show ?thesis
unfolding prob_proj_def
by simp
qed
lemma exec_drest:
fixes as a vs
assumes "(fst a \<subseteq>\<^sub>f s)"
shows "(
exec_plan (fmrestrict_set vs (state_succ s a)) as
= exec_plan (fmrestrict_set vs s) (action_proj a vs # as)
)"
using assms proj_state_succ
by fastforce
lemma not_empty_eff_in_as_proj:
fixes as a vs
assumes "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}"
shows "(as_proj (a # as) vs = (action_proj a vs # as_proj as vs))"
unfolding action_proj_def as_proj.simps
using assms
by argo
lemma empty_eff_not_in_as_proj:
fixes as a vs
assumes "(fmdom' (fmrestrict_set vs (snd a)) = {})"
shows "(as_proj (a # as) vs = as_proj as vs)"
unfolding action_proj_def
using assms
by simp
lemma empty_eff_drest_no_eff:
fixes s and a and vs
assumes "(fmdom' (fmrestrict_set vs (snd a)) = {})"
shows "(fmrestrict_set vs (state_succ s (action_proj a vs)) = fmrestrict_set vs s)"
proof -
have "fmdom' (snd (action_proj a vs)) = {}"
unfolding action_proj_def
using assms
by simp
then have "state_succ s (action_proj a vs) = s"
using empty_eff_exec_eq
by fast
then show ?thesis
by simp
qed
lemma sat_precond_exec_as_proj_eq_proj_exec:
fixes as vs s
assumes "(sat_precond_as s as)"
shows "(exec_plan (fmrestrict_set vs s) (as_proj as vs) = fmrestrict_set vs (exec_plan s as))"
using assms
proof (induction as)
case (Cons a as)
then show ?case
using Cons.prems graph_plan_lemma_1
by blast
qed auto
lemma action_proj_in_prob_proj:
assumes "(a \<in> PROB)"
shows "(action_proj a vs \<in> prob_proj PROB vs)"
unfolding action_proj_def prob_proj_def
using assms
by simp
lemma valid_as_valid_as_proj:
fixes PROB vs
assumes "(as \<in> valid_plans PROB)"
shows "(as_proj as vs \<in> valid_plans (prob_proj PROB vs))"
using assms
proof (induction "as" arbitrary: PROB vs)
case (Cons a as)
then show ?case
using assms Cons
proof(cases "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}")
case True
then have 1: "as_proj (a # as) vs = action_proj a vs # as_proj as vs"
using True
by simp
then have "as \<in> valid_plans PROB"
using Cons.prems valid_plan_valid_tail
by fast
then have "as_proj as vs \<in> valid_plans (prob_proj PROB vs)"
using Cons.IH 1
by simp
then have "action_proj a vs # as_proj as vs \<in> valid_plans (prob_proj PROB vs)"
using Cons.prems action_proj_in_prob_proj valid_head_and_tail_valid_plan valid_plan_valid_head
by metis
then show ?thesis
using 1
by argo
next
case False
then have "as_proj (a # as) vs = as_proj as vs"
using False
by auto
then have "as_proj (a # as) vs \<in> valid_plans (prob_proj PROB vs)"
using assms Cons valid_plan_valid_tail
by metis
then show ?thesis
using assms Cons.IH(1)
by blast
qed
qed (simp add: valid_plans_def)
lemma finite_imp_finite_prob_proj:
fixes PROB
assumes "finite PROB"
shows "(finite (prob_proj PROB vs))"
unfolding prob_proj_def
using assms
by simp
\<comment> \<open>NOTE Base 2 in 5th assumption had to be explicitely fixed to 'nat' type to be able to use the
linearity lemma for powers of natural numbers.\<close>
lemma
fixes PROB vs as and s :: "'a state"
assumes "finite PROB" "s \<in> valid_states PROB" "as \<in> (valid_plans PROB)" "finite vs"
"length (as_proj as vs) > ((2 :: nat) ^ card vs) - 1" "sat_precond_as s as"
shows "(\<exists>as1 as2 as3.
(as1 @ as2 @ as3 = as_proj as vs)
\<and> (exec_plan (fmrestrict_set vs s) (as1 @ as2) = exec_plan (fmrestrict_set vs s) as1)
\<and> (as2 \<noteq> [])
)"
proof -
{
have "card (fmdom' (fmrestrict_set vs s)) \<le> card vs"
using assms(4) graph_plan_card_state_set
by fast
then have "(2 :: nat) ^ (card (fmdom' (fmrestrict_set vs s))) - 1 \<le> 2 ^ (card vs) - 1"
using power_increasing diff_le_mono
by force
also have "... < length (as_proj as vs)"
using assms(5)
by blast
finally have "2 ^ card (fmdom' (fmrestrict_set vs s)) - 1 < length (as_proj as vs)"
by blast
}
note 1 = this
moreover have "fmrestrict_set vs s \<in> valid_states (prob_proj PROB vs)"
using assms(2) graph_plan_not_eq_last_diff_paths
by blast
moreover have "as_proj as vs \<in> valid_plans (prob_proj PROB vs)"
using assms(3) valid_as_valid_as_proj
by blast
moreover have "finite (prob_proj PROB vs)"
using assms(1) finite_imp_finite_prob_proj
by blast
ultimately show ?thesis
using lemma_2[where PROB="prob_proj PROB vs" and as="as_proj as vs" and s="fmrestrict_set vs s"]
by blast
qed
lemma as_proj_eq_filter_action_proj:
fixes as vs
shows "as_proj as vs = filter (\<lambda>a. fmdom' (snd a) \<noteq> {}) (map (\<lambda>a. action_proj a vs) as)"
by (induction as) (auto simp add: action_proj_def)
lemma append_eq_as_proj:
fixes as1 as2 as3 p vs
assumes "(as1 @ as2 @ as3 = as_proj p vs)"
shows "(\<exists>p_1 p_2 p_3.
(p_1 @ p_2 @ p_3 = p)
\<and> (as2 = as_proj p_2 vs)
\<and> (as1 = as_proj p_1 vs)
)"
using assms append_eq_as_proj_1 as_proj_eq_filter_action_proj
by (metis (no_types, lifting))
lemma succ_drest_eq_drest_succ:
fixes a s vs
shows "
state_succ (fmrestrict_set vs s) (action_proj a vs)
= fmrestrict_set vs (state_succ s (action_proj a vs))
"
proof -
let ?lhs = "state_succ (fmrestrict_set vs s) (action_proj a vs)"
let ?rhs = "fmrestrict_set vs (state_succ s (action_proj a vs))"
\<comment> \<open>NOTE Show lhs and rhs equality by splitting on the cases introduced by the if-then branching
of 'state\_succ'.\<close>
{
assume P1: "fst (fmrestrict_set vs (fst a), fmrestrict_set vs (snd a)) \<subseteq>\<^sub>f fmrestrict_set vs s"
then have a: "fst (fmrestrict_set vs (fst a), fmrestrict_set vs (snd a)) \<subseteq>\<^sub>f s"
using drest_smap_drest_smap_drest
by auto
then have "?lhs = fmrestrict_set vs (snd a) ++ fmrestrict_set vs s"
unfolding state_succ_def action_proj_def
using P1
by simp
moreover {
have rhs: "?rhs = fmrestrict_set vs (fmrestrict_set vs (snd a) ++ s)"
unfolding state_succ_def action_proj_def
using a
by auto
also have "\<dots> = (fmrestrict_set vs (fmrestrict_set vs (snd a)) ++ fmrestrict_set vs s)"
unfolding fmap_add_ltr_def
by simp
finally have "?rhs = (fmrestrict_set vs (snd a) ++ fmrestrict_set vs s)"
unfolding fmfilter_alt_defs(4)
by fastforce
}
ultimately have "?lhs = ?rhs"
by argo
}
moreover {
assume P2: "\<not>(fst (fmrestrict_set vs (fst a), fmrestrict_set vs (snd a)) \<subseteq>\<^sub>f fmrestrict_set vs s)"
then have a: "\<not>(fst (fmrestrict_set vs (fst a), fmrestrict_set vs (snd a)) \<subseteq>\<^sub>f s)"
using drest_smap_drest_smap_drest
by auto
then have "?lhs = fmrestrict_set vs s"
unfolding state_succ_def action_proj_def
using P2
by argo
moreover have "?rhs = fmrestrict_set vs s"
unfolding state_succ_def action_proj_def
using a
by presburger
ultimately have "?lhs = ?rhs"
by simp
}
ultimately show "?lhs = ?rhs"
by blast
qed
lemma proj_exec_proj_eq_exec_proj:
fixes s as vs
shows "
fmrestrict_set vs (exec_plan (fmrestrict_set vs s) (as_proj as vs))
= exec_plan (fmrestrict_set vs s) (as_proj as vs)
"
proof (induction as arbitrary: s vs)
case (Cons a as)
then show ?case
by (simp add: succ_drest_eq_drest_succ)
qed (simp add: fmfilter_alt_defs(4))
lemma proj_exec_proj_eq_exec_proj':
fixes s as vs
shows "
fmrestrict_set vs (exec_plan (fmrestrict_set vs s) (as_proj as vs))
= fmrestrict_set vs (exec_plan s (as_proj as vs))
"
proof (induction as arbitrary: s vs)
case (Cons a as)
then show ?case
by (simp add: succ_drest_eq_drest_succ)
qed (simp add: fmfilter_alt_defs(4))
lemma graph_plan_lemma_9:
fixes s as vs
shows "
fmrestrict_set vs (exec_plan s (as_proj as vs))
= exec_plan (fmrestrict_set vs s) (as_proj as vs)
"
by (metis proj_exec_proj_eq_exec_proj' proj_exec_proj_eq_exec_proj)
lemma act_dom_proj_eff_subset_act_dom_eff:
fixes a vs
shows "fmdom' (snd (action_proj a vs)) \<subseteq> fmdom' (snd a)"
proof -
have "snd (action_proj a vs) = fmrestrict_set vs (snd a)"
unfolding action_proj_def
by simp
then have "fmlookup (fmrestrict_set vs (snd a)) \<subseteq>\<^sub>m fmlookup (snd a)"
by (simp add: map_le_def fmdom'_restrict_set_precise)
then have "dom (fmlookup (fmrestrict_set vs (snd a))) \<subseteq> dom (fmlookup (snd a))"
using map_le_implies_dom_le
by blast
then have "fmdom' (fmrestrict_set vs (snd a)) \<subseteq> fmdom' (snd a)"
using fmdom'.rep_eq
by metis
then show ?thesis
unfolding action_proj_def
by simp
qed
lemma exec_as_proj_valid:
fixes as s PROB vs
assumes "s \<in> valid_states PROB" "(as \<in> valid_plans PROB)"
shows "(exec_plan s (as_proj as vs) \<in> valid_states PROB)"
using assms
proof (induction as arbitrary: s PROB vs)
case (Cons a as)
then have 1: "as \<in> valid_plans PROB"
using Cons.prems(2) valid_plan_valid_tail
by fast
then have 2: "exec_plan s (as_proj as vs) \<in> valid_states PROB"
using Cons.prems(1) Cons.IH(1)
by blast
\<comment> \<open>NOTE split on the if-then branch introduced by 'as\_proj'.\<close>
moreover {
assume P: "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}"
then have "
exec_plan s (as_proj (a # as) vs)
= exec_plan (state_succ s (action_proj a vs)) (as_proj as vs)
"
by simp
\<comment> \<open>NOTE split on the if-then branch introduced by 'state\_succ'\<close>
moreover
{
assume "fst (action_proj a vs) \<subseteq>\<^sub>f s"
then have 3: "
exec_plan (state_succ s (action_proj a vs)) (as_proj as vs)
= exec_plan (snd (action_proj a vs) ++ s) (as_proj as vs)
"
unfolding state_succ_def
using calculation
by simp
{
\<comment> \<open>TODO Unsure why this proof step is necessary at all, but it should be refactored into a
dedicated lemma @{term "s \<in> valid_states PROB \<Longrightarrow> fmdom' s = prob_dom PROB"}.\<close>
{
have "s \<in> valid_states PROB"
using Cons.prems
by simp
then have "s \<in> {s'. fmdom' s' = prob_dom PROB}"
unfolding valid_states_def
by simp
then obtain s' where "s' = s" "fmdom' s' = prob_dom PROB"
by auto
then have "fmdom' s = prob_dom PROB"
by simp
}
\<comment> \<open>TODO Refactor this step ('also ...' for subset chain; replace fact
`fmdom' s = prob\_dom PROB` in last step with MP step from lemma refactored above.\<close>
moreover {
have "(snd (action_proj a vs) ++ s) = (s ++\<^sub>f fmrestrict_set vs (snd a))"
unfolding action_proj_def fmap_add_ltr_def
by simp
then have a: "a \<in> PROB"
using Cons.prems(2) valid_plan_valid_head
by fast
then have "action_dom (fst a) (snd a) \<subseteq> prob_dom PROB"
using exec_as_proj_valid_2
by blast
then have "fmdom' (snd a) \<subseteq> action_dom (fst a) (snd a)"
unfolding action_dom_def
by simp
then have "fmdom' (fmrestrict_set vs (snd a)) \<subseteq> fmdom' (snd a)"
using action_proj_def act_dom_proj_eff_subset_act_dom_eff snd_conv
by metis
then have "fmdom' (fmrestrict_set vs (snd a)) \<subseteq> prob_dom PROB"
using FDOM_eff_subset_prob_dom_pair a
by blast
then have "fmdom' (s ++\<^sub>f fmrestrict_set vs (snd a)) = fmdom' s"
by (simp add: calculation sup.absorb_iff1)
}
ultimately have "(snd (action_proj a vs) ++ s) \<in> valid_states PROB"
unfolding action_proj_def fmap_add_ltr_def valid_states_def
by simp
}
then have "exec_plan s (as_proj (a # as) vs) \<in> valid_states PROB"
using 1 3 calculation(1) Cons.IH[where s = "snd (action_proj a vs) ++ s"]
by presburger
}
moreover {
assume "\<not>(fst (action_proj a vs) \<subseteq>\<^sub>f s)"
then have "
exec_plan (state_succ s (action_proj a vs)) (as_proj as vs)
= exec_plan s (as_proj as vs)
"
unfolding state_succ_def
by simp
then have "exec_plan s (as_proj (a # as) vs) \<in> valid_states PROB"
using 2
by force
}
ultimately have "exec_plan s (as_proj (a # as) vs) \<in> valid_states PROB"
by blast
}
moreover
{
assume "fmdom' (fmrestrict_set vs (snd a)) = {}"
then have "
exec_plan s (as_proj (a # as) vs) =
exec_plan s (as_proj as vs)
"
by simp
then have "exec_plan s (as_proj (a # as) vs) \<in> valid_states PROB"
using 2
by argo
}
ultimately show ?case
by blast
qed simp
lemma drest_exec_as_proj_eq_drest_exec:
fixes s as vs
assumes "sat_precond_as s as"
shows "(fmrestrict_set vs (exec_plan s (as_proj as vs)) = fmrestrict_set vs (exec_plan s as))"
proof -
have 1: "
(fmrestrict_set vs (exec_plan s (as_proj as vs))
= exec_plan (fmrestrict_set vs s) (as_proj as vs))
"
using graph_plan_lemma_9 by auto
then obtain s' where 2: "exec_plan (fmrestrict_set vs s) (as_proj as vs) = fmrestrict_set vs s'"
using 1
by metis
then have "fmrestrict_set vs s' = fmrestrict_set vs (exec_plan s as)"
using assms sat_precond_exec_as_proj_eq_proj_exec
by metis
then show
"fmrestrict_set vs (exec_plan s (as_proj as vs)) = fmrestrict_set vs (exec_plan s as)"
using 1 2
by argo
qed
lemma action_proj_idempot:
fixes a vs
shows "action_proj (action_proj a vs) vs = (action_proj a vs)"
unfolding action_proj_def
by (simp add: fmfilter_alt_defs(4))
lemma action_proj_idempot':
fixes a vs
assumes "(action_dom (fst a) (snd a) \<subseteq> vs)"
shows "(action_proj a vs = a)"
using assms
proof -
have 1: "action_proj a vs = (fmrestrict_set vs (fst a), fmrestrict_set vs (snd a))"
by (simp add: action_proj_def)
then have 2: "(fmdom' (fst a) \<union> fmdom' (snd a)) \<subseteq> vs"
unfolding action_dom_def
using assms
by (auto simp add: action_dom_def)
\<comment> \<open>NOTE Show that both components of 'a' remain unchanged.\<close>
{
then have "fmdom' (fst a) \<subseteq> vs"
by blast
then have "fmrestrict_set vs (fst a) = (fst a)"
using exec_drest_5
by auto
}
moreover {
have "fmdom' (snd a) \<subseteq> vs"
using 2
by auto
then have "fmrestrict_set vs (snd a) = (snd a)"
using exec_drest_5
by blast
}
ultimately show ?thesis
using 1
by simp
qed
lemma action_proj_idempot'':
fixes P vs
assumes "prob_dom P \<subseteq> vs"
shows "prob_proj P vs = P"
using assms
proof -
\<comment> \<open>TODO refactor.\<close>
{
fix a
assume "a \<in> P"
then have "action_dom (fst a) (snd a) \<subseteq> vs"
using assms exec_as_proj_valid_2
by fast
then have "action_proj a vs = a"
using action_proj_idempot'
by fast
}
then have "prob_proj P vs = P"
unfolding prob_proj_def
by force
then show ?thesis
unfolding prob_proj_def
by simp
qed
lemma sat_precond_as_proj:
fixes as s s' vs
assumes "(sat_precond_as s as)" "(fmrestrict_set vs s = fmrestrict_set vs s')"
shows "(sat_precond_as s' (as_proj as vs))"
using assms
proof (induction as arbitrary: s s' vs)
case (Cons a as)
then have 1:
"fst a \<subseteq>\<^sub>f s" "sat_precond_as (state_succ s a) as"
using Cons.prems(1)
by simp+
then have 2: "fmrestrict_set vs (fst a) \<subseteq>\<^sub>f s"
using assms(1) sat_precond_as_proj_4
by blast
moreover
{
assume "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}"
then have "
sat_precond_as s' (as_proj (a # as) vs)
= (
fst (action_proj a vs) \<subseteq>\<^sub>f s'
\<and> sat_precond_as (state_succ s' (action_proj a vs)) (as_proj as vs)
)
"
using calculation
by simp
moreover
{
have "fst (action_proj a vs) \<subseteq>\<^sub>f s' = (fmrestrict_set vs (fst a) \<subseteq>\<^sub>f s')"
unfolding action_proj_def
by simp
moreover have "(fmrestrict_set vs (fst a) \<subseteq>\<^sub>f s) = (fmrestrict_set vs (fst a) \<subseteq>\<^sub>f s')"
using Cons.prems(2) sat_precond_as_proj_1
by blast
ultimately have "fst (action_proj a vs) \<subseteq>\<^sub>f s'"
using 2
by blast
}
\<comment> \<open>TODO detailled proof for this sledgehammered step.\<close>
moreover have "sat_precond_as (state_succ s' (action_proj a vs)) (as_proj as vs)"
using 1 Cons.IH Cons.prems(2) drest_succ_proj_eq_drest_succ succ_drest_eq_drest_succ
by metis
ultimately have "(sat_precond_as s' (as_proj (a # as) vs))"
by blast
}
moreover
{
assume P1: "\<not>(fmdom' (fmrestrict_set vs (snd a)) \<noteq> {})"
then have "sat_precond_as s' (as_proj (a # as) vs)"
proof (cases "as_proj (a # as) vs")
case Cons2: (Cons a' list)
\<comment> \<open>TODO unfold the sledgehammered metis steps.\<close>
then have a: "
sat_precond_as s' (as_proj (a # as) vs)
= (fst a' \<subseteq>\<^sub>f s') \<and> sat_precond_as (state_succ s' a') list
"
using P1 Cons.IH Cons.prems(1, 2) Cons2
by (metis sat_precond_as_proj_3 empty_eff_not_in_as_proj sat_precond_as.simps(2))
then have b: "fst a' \<subseteq>\<^sub>f s'"
unfolding sat_precond_as.simps(2)
using P1 Cons.IH Cons.prems(1, 2) sat_precond_as_proj_3 empty_eff_not_in_as_proj
by (metis sat_precond_as.simps(2))
then have "sat_precond_as (state_succ s' a') list"
using a
by blast
then show ?thesis
using a b
by blast
qed fastforce
}
ultimately show ?case
by blast
qed simp
lemma sat_precond_drest_as_proj:
fixes as s s' vs
assumes "(sat_precond_as s as)" "(fmrestrict_set vs s = fmrestrict_set vs s')"
shows "(sat_precond_as (fmrestrict_set vs s') (as_proj as vs))"
using assms
proof (induction as arbitrary: s s' vs)
case (Cons a as)
then have 1: "fst a \<subseteq>\<^sub>f s" "sat_precond_as (state_succ s a) as"
using Cons.prems
by auto+
then have "fmrestrict_set vs (fst a) \<subseteq>\<^sub>f fmrestrict_set vs s"
using fmsubset_restrict_set_mono
by blast
then have "fst (action_proj a vs) \<subseteq>\<^sub>f fmrestrict_set vs s'"
unfolding action_proj_def
using Cons.prems(2) sat_precond_as_proj_1
by simp
then have "fmrestrict_set vs (snd a) = fmrestrict_set vs (snd (action_proj a vs))"
unfolding action_proj_def
by (simp add: fmfilter_alt_defs(4))
then have "fst (action_proj a vs) \<subseteq>\<^sub>f s"
unfolding action_proj_def
using 1(1) fst_conv sat_precond_as_proj_4
by auto
\<comment> \<open>TODO unfold these sledgehammered steps.\<close>
then have "
fmrestrict_set vs (state_succ s a)
= fmrestrict_set vs (state_succ (fmrestrict_set vs s') (action_proj a vs))
"
using 1(1) Cons.prems(2)
by (metis fmfilter_alt_defs(4) fmfilter_true fmlookup_restrict_set
drest_succ_proj_eq_drest_succ option.simps(3))
then show ?case
using Cons.prems(1, 2)
by (metis fmfilter_alt_defs(4) fmfilter_true fmlookup_restrict_set sat_precond_as_proj
option.simps(3))
qed simp
lemma as_proj_eq_as:
assumes "(no_effectless_act as)" "(as \<in> valid_plans PROB)" "(prob_dom PROB \<subseteq> vs)"
shows "(as_proj as vs = as)"
using assms
proof (induction as arbitrary: PROB vs)
case (Cons a as)
\<comment> \<open>NOTE We only need to look at the first branch of 'as\_proj'.\<close>
\<comment> \<open>TODO step should be refactored and proven explicitely because it's so pivotal.\<close>
then have "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}"
unfolding fmdom'_restrict_set_precise
by (metis
FDOM_eff_subset_prob_dom_pair dual_order.trans inf.orderE
no_effectless_act.simps(2) valid_plan_valid_head)
\<comment> \<open>NOTE Proof 'action\_proj a vs = a' for the first branch of 'as\_proj'.\<close>
moreover {
assume "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}"
\<comment> \<open>NOTE show 'action\_proj a vs = a'.\<close>
moreover {
have "as_proj (a # as) vs = action_proj a vs # as_proj as vs"
using calculation
by force
then have "a \<in> PROB"
using Cons.prems(2) valid_plan_valid_head
by fast
then have "action_dom (fst a) (snd a) \<subseteq> prob_dom PROB"
using exec_as_proj_valid_2
by fast
then have "action_dom (fst a) (snd a) \<subseteq> vs"
using Cons.prems(3)
by fast
then have "action_proj a vs = a"
using action_proj_idempot'
by fast
}
\<comment> \<open>NOTE show that 'as\_proj as vs = as'.\<close>
moreover {
have 1: "no_effectless_act as"
using Cons.prems(1)
by simp
then have "as \<in> valid_plans PROB"
using Cons.prems(2) valid_plan_valid_tail
by fast
then have "as_proj as vs = as"
using Cons.prems(3) Cons.IH 1
by blast
}
ultimately have "as_proj (a # as) vs = a # as"
by simp
}
ultimately show ?case
by fast
qed simp
lemma exec_rem_effless_as_proj_eq_exec_as_proj:
fixes s
shows "exec_plan s (as_proj (rem_effectless_act as) vs) = exec_plan s (as_proj as vs)"
proof (induction as arbitrary: s vs)
case (Cons a as)
\<comment> \<open>Split cases on the branching introduced by `remove\_effectless\_act` and `as\_proj`.\<close>
then show ?case
proof (cases "fmdom' (snd a) \<noteq> {}")
case true1: True
then show ?thesis
proof (cases "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}")
case False
then show ?thesis by (simp add: Cons true1)
qed (simp add: Cons true1)
next
case False
then show ?thesis
proof (cases "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}")
case true2: True
then have 1: "fmdom' (snd a) \<inter> vs = {}"
using False Int_empty_left
by force
\<comment> \<open>NOTE This step shows that the case for @{term "fmdom' (fmrestrict_set vs (snd a)) \<noteq> {}"} is
impossible.\<close>
\<comment> \<open>TODO could be refactored into a (simp) lemma (`as\_proj\_eq\_as` also uses this?).\<close>
then have "fmdom' (fmrestrict_set vs (snd a)) = {}"
by (simp add: fmdom'_restrict_set_precise)
then show ?thesis
using true2
by blast
qed (simp add: Cons)
qed
qed simp
lemma exec_as_proj_eq_exec_as:
fixes PROB as vs s
assumes "(as \<in> valid_plans PROB)" "(prob_dom PROB \<subseteq> vs)"
shows "(exec_plan s (as_proj as vs) = exec_plan s as)"
using assms as_proj_eq_as exec_rem_effless_as_proj_eq_exec_as_proj rem_effectless_works_1 rem_effectless_works_6
rem_effectless_works_9 sublist_valid_plan
by metis
lemma dom_prob_proj: "prob_dom (prob_proj PROB vs) \<subseteq> vs"
using graph_plan_neq_mems_state_set_neq_len
by fast
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor into `FmapUtils.thy`.\<close>
lemma subset_proj_absorb_1_a:
fixes f vs1 vs2
assumes "(vs1 \<subseteq> vs2)"
shows "fmrestrict_set vs1 (fmrestrict_set vs2 f) = fmrestrict_set vs1 f"
using assms
proof -
{
fix v
have "fmlookup (fmrestrict_set vs1 (fmrestrict_set vs2 f)) v = fmlookup (fmrestrict_set vs1 f) v"
using assms
proof (cases "v \<in> vs1")
case False
then show ?thesis
proof (cases "v \<in> vs2")
case False
then have "v \<notin> vs1"
using False assms
by blast
then have
"fmlookup (fmrestrict_set vs1 (fmrestrict_set vs2 f)) v = None"
"fmlookup (fmrestrict_set vs1 f) v = None"
by simp+
then show ?thesis
by argo
qed simp
qed auto
}
then show ?thesis
using fmap_ext
by blast
qed
lemma subset_proj_absorb_1:
assumes "(vs1 \<subseteq> vs2)"
shows "(action_proj (action_proj a vs2) vs1 = action_proj a vs1)"
using assms
proof -
have
"fmrestrict_set vs1 (fmrestrict_set vs2 (fst a)) = fmrestrict_set vs1 (fst a)"
"fmrestrict_set vs1 (fmrestrict_set vs2 (snd a)) = fmrestrict_set vs1 (snd a)"
using assms subset_proj_absorb_1_a
by blast+
then show ?thesis
unfolding action_proj_def
by simp
qed
lemma subset_proj_absorb:
fixes PROB vs1 vs2
assumes "vs1 \<subseteq> vs2"
shows "prob_proj (prob_proj PROB vs2) vs1 = prob_proj PROB vs1"
proof -
{
have "
prob_proj (prob_proj PROB vs2) vs1
= ((\<lambda>a. action_proj a vs1) \<circ> (\<lambda>a. action_proj a vs2)) ` PROB
"
unfolding prob_proj_def
by fastforce
also have "\<dots> = (\<lambda>a. action_proj (action_proj a vs2) vs1) ` PROB"
by fastforce
also have "\<dots> = (\<lambda>a. action_proj a vs1) ` PROB"
using assms subset_proj_absorb_1
by metis
also have "\<dots> = prob_proj PROB vs1"
unfolding prob_proj_def
by simp
finally have "prob_proj (prob_proj PROB vs2) vs1 = prob_proj PROB vs1"
by simp
}
then show ?thesis
by simp
qed
lemma union_proj_absorb:
fixes PROB vs vs'
shows "prob_proj (prob_proj PROB (vs \<union> vs')) vs = prob_proj PROB vs"
by (simp add: subset_proj_absorb)
lemma NOT_VS_IN_DOM_PROJ_PRE_EFF:
fixes ROB vs v a
assumes "\<not>(v \<in> vs)" "(a \<in> PROB)"
shows "(
((v \<in> fmdom' (fst a)) \<longrightarrow> (v \<in> fmdom' (fst (action_proj a (prob_dom PROB - vs)))))
\<and> ((v \<in> fmdom' (snd a)) \<longrightarrow> (v \<in> fmdom' (snd (action_proj a (prob_dom PROB - vs)))))
)"
unfolding action_proj_def
using assms
by (simp add: IN_FDOM_DRESTRICT_DIFF FDOM_pre_subset_prob_dom_pair
FDOM_eff_subset_prob_dom_pair)
lemma IN_DISJ_DEP_IMP_DEP_DIFF:
fixes PROB vs vs' v v'
assumes "(v \<in> vs')" "(v' \<in> vs')" "(disjnt vs vs')"
shows "(dep PROB v v' \<longrightarrow> dep (prob_proj PROB (prob_dom PROB - vs)) v v')"
using assms
proof (cases "v = v'")
case False
{
assume P: "dep PROB v v'"
then obtain a where a:
"(v \<in> fmdom' (fst a) \<and> v' \<in> fmdom' (snd a) \<or> v \<in> fmdom' (snd a) \<and> v' \<in> fmdom' (snd a))"
"a \<in> PROB"
unfolding dep_def
using False
by blast
{
have "v \<notin> vs"
using assms(1, 3)
unfolding disjnt_def
by blast
then have "(v \<in> fmdom' (fst a) \<longrightarrow> v \<in> fmdom' (fst (action_proj a (prob_dom PROB - vs))))"
"(v \<in> fmdom' (snd a) \<longrightarrow> v \<in> fmdom' (snd (action_proj a (prob_dom PROB - vs))))"
using a NOT_VS_IN_DOM_PROJ_PRE_EFF
by metis+
}
note b = this
then consider (i) "v \<in> fmdom' (fst a) \<and> v' \<in> fmdom' (snd a)"
| (ii) "v \<in> fmdom' (snd a) \<and> v' \<in> fmdom' (snd a)"
using a
by blast
then have "dep (prob_proj PROB (prob_dom PROB - vs)) v v'"
proof (cases)
case i
then show ?thesis
using assms(2, 3) a(2) b(1)
by (meson dep_def disjnt_iff action_proj_in_prob_proj NOT_VS_IN_DOM_PROJ_PRE_EFF)
next
case ii
then show ?thesis
using assms(2, 3) a(2) b(2)
by (meson dep_def disjnt_iff action_proj_in_prob_proj NOT_VS_IN_DOM_PROJ_PRE_EFF)
qed
}
then show ?thesis
by blast
qed (auto simp: dep_def prob_proj_def disjnt_def)
lemma PROB_DOM_PROJ_DIFF:
fixes P vs
shows "prob_dom (prob_proj PROB (prob_dom PROB - vs)) = (prob_dom PROB) - vs"
using graph_plan_neq_mems_state_set_neq_len
by fastforce
lemma two_children_parent_mems_le_finite:
fixes PROB vs
assumes "(vs \<subseteq> prob_dom PROB)"
shows "(prob_dom (prob_proj PROB vs) = vs)"
using assms graph_plan_neq_mems_state_set_neq_len
by fast
\<comment> \<open>TODO showcase (non-trivial proof).\<close>
\<comment> \<open>TODO find explicit proof.\<close>
lemma PROJ_DOM_PRE_EFF_SUBSET_DOM:
fixes a vs
shows "
(fmdom' (fst (action_proj a vs)) \<subseteq> fmdom' (fst a))
\<and> (fmdom' (snd (action_proj a vs)) \<subseteq> fmdom' (snd a))
"
unfolding action_proj_def
by (auto simp: fmdom'_restrict_set_precise)
lemma NOT_IN_PRE_EFF_NOT_IN_PRE_EFF_PROJ:
fixes a v vs
shows "
(\<not>(v \<in> fmdom' (fst a)) \<longrightarrow> \<not>(v \<in> fmdom' (fst (action_proj a vs))))
\<and> (\<not>(v \<in> fmdom' (snd a)) \<longrightarrow> \<not>(v \<in> fmdom' (snd (action_proj a vs))))
"
using PROJ_DOM_PRE_EFF_SUBSET_DOM rev_subsetD
by metis
lemma dep_proj_dep:
assumes "dep (prob_proj PROB vs) v v'"
shows "dep PROB v v'"
using assms
unfolding dep_def prob_proj_def action_proj_def image_def
apply (auto simp: fmdom'_restrict_set_precise)
by auto
lemma NDEP_PROJ_NDEP:
fixes PROB vs vs' vs''
assumes "(\<not>dep_var_set PROB vs vs')"
shows "(\<not>dep_var_set (prob_proj PROB vs'') vs vs')"
using assms dep_proj_dep
unfolding dep_var_set_def
by metis
lemma SUBSET_PROJ_DOM_DISJ:
fixes PROB vs vs'
assumes "(vs \<subseteq> (prob_dom (prob_proj PROB (prob_dom PROB - vs'))))"
shows "disjnt vs vs'"
using assms
by (auto simp add: PROB_DOM_PROJ_DIFF subset_iff disjnt_iff)
\<comment> \<open>TODO showcase (lemma which is solved effortlessly by automation).\<close>
lemma NOT_VS_DEP_IMP_DEP_PROJ:
fixes PROB vs v v'
assumes "\<not>(v \<in> vs)" "\<not>(v' \<in> vs)" "(dep PROB v v')"
shows "(dep (prob_proj PROB (prob_dom PROB - vs)) v v')"
using assms
by (metis Diff_disjoint Diff_iff disjnt_def insertCI IN_DISJ_DEP_IMP_DEP_DIFF)
lemma DISJ_PROJ_NDEP_IMP_NDEP:
fixes PROB vs vs' vs''
assumes
"(disjnt vs vs'')" "disjnt vs vs'"
"\<not>(dep_var_set (prob_proj PROB (prob_dom PROB - vs)) vs' vs'')"
shows "\<not>(dep_var_set PROB vs' vs'')"
proof -
{
assume C: "dep_var_set PROB vs' vs''"
then obtain v1 v2 where "v1 \<in> vs'" "v2 \<in> vs''" "disjnt vs' vs''" "dep PROB v1 v2"
unfolding dep_var_set_def
by blast
then have "\<exists>v1 v2.
v1 \<in> vs' \<and> v2 \<in> vs'' \<and> disjnt vs' vs'' \<and> dep (prob_proj PROB (prob_dom PROB - vs)) v1 v2
"
using assms(1, 2) IntI disjnt_def empty_iff NOT_VS_DEP_IMP_DEP_PROJ
by metis
then have False
using assms
unfolding dep_var_set_def
by blast
}
then show ?thesis
using assms
unfolding dep_var_set_def
by argo
qed
lemma PROJ_DOM_IDEMPOT:
fixes PROB
shows "prob_proj PROB (prob_dom PROB) = PROB"
using action_proj_idempot''
by blast
lemma prob_proj_idempot:
fixes vs vs'
assumes "(vs \<subseteq> vs')"
shows "(prob_proj PROB vs = prob_proj (prob_proj PROB vs') vs)"
using assms subset_proj_absorb
by blast
lemma prob_proj_dom_diff_eq_prob_proj_prob_proj_dom_diff:
fixes vs vs'
shows "
prob_proj PROB (prob_dom PROB - (vs \<union> vs'))
= prob_proj
(prob_proj PROB (prob_dom PROB - vs))
(prob_dom (prob_proj PROB (prob_dom PROB - vs)) - vs')
"
using PROB_DOM_PROJ_DIFF subset_proj_absorb
by (metis Compl_Diff_eq Diff_subset compl_eq_compl_iff sup_assoc)
lemma PROJ_DEP_IMP_DEP:
fixes PROB vs v v'
assumes "dep (prob_proj PROB (prob_dom PROB - vs)) v v'"
shows "dep PROB v v'"
using assms
unfolding dep_def prob_proj_def
proof (cases "v = v'")
case False
then show "(\<exists>a.
a \<in> PROB
\<and> (v \<in> fmdom' (fst a) \<and> v' \<in> fmdom' (snd a) \<or> v \<in> fmdom' (snd a) \<and> v' \<in> fmdom' (snd a)))
\<or> v = v'"
using assms
unfolding dep_def prob_proj_def
by (smt image_iff NOT_IN_PRE_EFF_NOT_IN_PRE_EFF_PROJ)
qed blast
lemma PROJ_NDEP_TC_IMP_NDEP_TC_OR:
fixes PROB vs v v'
assumes "\<not>((\<lambda>v1' v2'. dep (prob_proj PROB (prob_dom PROB - vs)) v1' v2')\<^sup>+\<^sup>+ v v')"
shows "(
(\<not>((\<lambda>v1' v2'. dep PROB v1' v2')\<^sup>+\<^sup>+ v v'))
\<or> (\<exists>v''.
v'' \<in> vs
\<and> ((\<lambda>v1' v2'. dep PROB v1' v2')\<^sup>+\<^sup>+ v v'')
\<and> ((\<lambda>v1' v2'. dep PROB v1' v2')\<^sup>+\<^sup>+ v'' v')
)
)"
using assms NOT_VS_DEP_IMP_DEP_PROJ DEP_REFL REFL_TC_CONJ[of
"\<lambda>v v'. dep PROB v v'" "\<lambda>v. \<not>(v \<in> vs)" "\<lambda>v v'. dep (prob_proj PROB (prob_dom PROB- vs)) v v'"
v v']
by fastforce
lemma every_action_proj_eq_as_proj:
fixes as vs
shows "list_all (\<lambda> a. action_proj a vs = a) (as_proj as vs)"
by (induction as) (auto simp add: action_proj_idempot)
lemma empty_eff_not_in_as_proj_2:
fixes a as vs
assumes "fmdom' (snd (action_proj a vs)) = {}"
shows "(as_proj as vs = as_proj (a # as) vs)"
using assms
by (auto simp add: action_proj_def)
declare[[smt_timeout=100]]
lemma sublist_as_proj_eq_as:
fixes as' as vs
assumes "subseq as' (as_proj as vs)"
shows "(as_proj as' vs = as')"
using assms
proof (induction as arbitrary: as' vs)
case Nil
moreover have "as' = []"
using Nil.prems sublist_NIL
by force
then show ?case
by simp
next
case cons: (Cons a as)
then show ?case
proof (cases "as'")
case (Cons aa list)
then show ?thesis
proof (cases "fmdom' (fmrestrict_set vs (snd aa)) \<noteq> {}")
case True
then have "as_proj as' vs = action_proj aa vs # as_proj list vs"
using Cons True
by auto
then show ?thesis
by (metis as_proj.simps(2) cons.IH cons.prems action_proj_idempot local.Cons
subseq_Cons2_iff)
next
case False
then have "as_proj as' vs = as_proj list vs"
using Cons False
by simp
then show ?thesis using cons False
unfolding Cons
by (smt action_proj_def action_proj_idempot as_proj.simps(2) prod.inject subseq_Cons2_neq)
qed
qed simp
qed
lemma DISJ_EFF_DISJ_PROJ_EFF:
fixes a s vs
assumes "fmdom' (snd a) \<inter> s = {}"
shows "(fmdom' (snd (action_proj a vs)) \<inter> s = {})
"
proof -
have 1: "snd (action_proj a vs) = fmrestrict_set vs (snd a)"
unfolding action_proj_def
by simp
then have "fmdom' (fmrestrict_set vs (snd a)) \<subseteq> fmdom' (snd a)"
using act_dom_proj_eff_subset_act_dom_eff
by metis
then show ?thesis
using assms 1
by auto
qed
\<comment> \<open>NOTE showcase (the step using `graph\_plan\_lemma\_5`--- labelled by '[1]'--- is non-trivial proof
due to missing premises and the last six proof steps are redundant).\<close>
lemma state_succ_proj_eq_state_succ:
fixes a s vs
assumes "(varset_action a vs)" "(fst a \<subseteq>\<^sub>f s)" "(fmdom' (snd a) \<subseteq> fmdom' s)"
shows "(state_succ s (action_proj a vs) = state_succ s a)"
proof -
have 1: "fmdom' (snd a) \<inter> (fmdom' s - vs) = {}"
using assms(1) vset_disj_eff_diff
by blast
then have 2:
"fmrestrict_set (fmdom' s - vs) s = fmrestrict_set (fmdom' s - vs) (state_succ s a)"
using disj_imp_eq_proj_exec[where vs = "fmdom' s - vs"]
by blast
then have "fmdom' (snd (action_proj a vs)) \<inter> (fmdom' s - vs) = {}"
using 1 DISJ_EFF_DISJ_PROJ_EFF[where s = "(fmdom' s - vs)"]
by blast
then have "
fmrestrict_set (fmdom' s - vs) s
= fmrestrict_set (fmdom' s - vs) (state_succ s (action_proj a vs))
"
using disj_imp_eq_proj_exec[where a = "(action_proj a vs)" and vs = "fmdom' s - vs"]
by blast
then have "fmdom' (snd (action_proj a vs)) \<inter> (fmdom' s - vs) = {}"
using 1 DISJ_EFF_DISJ_PROJ_EFF[where s = "(fmdom' s - vs)"]
by blast
then have "
fmrestrict_set (fmdom' s - vs) s =
fmrestrict_set (fmdom' s - vs) (state_succ s (action_proj a vs))
"
using disj_imp_eq_proj_exec[of "action_proj a vs" "fmdom' s - vs"]
by fast
\<comment> \<open>[1]\<close>
\<comment> \<open>TODO unwrap this step.\<close>
then show ?thesis
using 2 FDOM_state_succ graph_plan_lemma_5[where s = "state_succ s (action_proj a vs)"
and s' = "state_succ s a" and vs = vs] assms(2, 3) dom_eff_subset_imp_dom_succ_eq_proj
drest_proj_succ_eq_drest_succ
by metis
qed
\<comment> \<open>NOTE duplicate declaration of lemma `state\_succ\_proj\_eq\_state\_succ` removed.\<close>
lemma no_effectless_proj:
fixes vs as
shows "no_effectless_act (as_proj as vs)"
by (induction as arbitrary: vs) (auto simp add: action_proj_def)
\<comment> \<open>NOTE duplicate (this is identical to `valid\_as\_valid\_as\_proj`).\<close>
lemma as_proj_valid_in_prob_proj:
fixes PROB vs as
assumes "(as \<in> valid_plans PROB)"
shows "(as_proj as vs \<in> valid_plans (prob_proj PROB vs))"
using assms valid_as_valid_as_proj
by blast
\<comment> \<open>TODO Unwrap the smt proof.\<close>
lemma prob_proj_comm:
fixes PROB vs vs'
shows "prob_proj (prob_proj PROB vs) vs' = prob_proj (prob_proj PROB vs') vs"
by (smt graph_plan_neq_mems_state_set_neq_len inf_commute inf_le2 PROJ_DOM_IDEMPOT prob_proj_idempot)
\<comment> \<open>TODO Unwrap the metis proof.\<close>
lemma vset_proj_imp_vset:
fixes vs vs' a
assumes "(varset_action a vs')" "(varset_action (action_proj a vs') vs)"
shows "(varset_action a vs)"
unfolding varset_action_def action_proj_def
using assms
by (metis action_proj_def exec_drest_5 snd_conv varset_action_def)
lemma vset_imp_vset_act_proj_diff:
fixes PROB vs vs' a
assumes "(varset_action a vs)"
shows "(varset_action (action_proj a (prob_dom PROB - vs')) vs)"
proof -
have 1: "(fmdom' (snd a) \<subseteq> vs)"
using assms varset_action_def
by metis
moreover
{
\<comment> \<open>TODO refactor and put into `Fmap\_Utils`.\<close>
have "
fmdom' (snd (
fmrestrict_set (prob_dom PROB - vs') (fst a)
, fmrestrict_set (prob_dom PROB - vs') (snd a)
))
= (fmdom' (snd a) \<inter> (prob_dom PROB - vs'))
"
by (simp add: Int_def Set.filter_def fmfilter_alt_defs(4))
also have "\<dots> \<subseteq> fmdom' (snd a)"
by simp
finally have "fmdom' (snd (
fmrestrict_set (prob_dom PROB - vs') (fst a)
, fmrestrict_set (prob_dom PROB - vs') (snd a)
))
\<subseteq> vs
"
using 1 by simp
}
ultimately show ?thesis
unfolding varset_action_def dep_var_set_def dep_def action_proj_def
by blast
qed
lemma action_proj_disj_diff:
assumes "(action_dom (fst a) (snd a) \<subseteq> vs1)" "(vs2 \<inter> vs3 = {})"
shows "(action_proj (action_proj a (vs1 - vs2)) vs3 = action_proj a vs3)"
proof -
have "\<forall>f fa fb p.
action_proj (action_proj (action_proj p f) fb) fa = action_proj (action_proj p f) fb
\<or> \<not> action_dom (fst p::('a, 'b) fmap) (snd p::(_, 'c) fmap) \<inter> (f \<inter> fb) \<subseteq> fa
"
by (metis (no_types) action_proj_idempot' proj_action_dom_eq_inter inf_assoc)
then have "\<forall>f fa p.
action_proj (action_proj (p::('a, 'b) fmap \<times> (_, 'c) fmap) f) fa
= action_proj p (f \<inter> fa)
"
by (metis (no_types) inf.cobounded2 inf_commute subset_proj_absorb_1)
then show ?thesis
using assms
by (metis Diff_Int_distrib2 Diff_empty action_proj_idempot')
qed
lemma disj_proj_proj_eq_proj:
fixes PROB vs vs'
assumes "(vs \<inter> vs' = {})"
shows "prob_proj (prob_proj PROB (prob_dom PROB - vs')) vs = prob_proj PROB vs"
proof -
{
fix a
assume P: "a \<in> PROB"
moreover have "action_dom (fst a) (snd a) \<subseteq> prob_dom PROB"
using P exec_as_proj_valid_2
by blast
ultimately have "action_proj (action_proj a (prob_dom PROB - vs')) vs = action_proj a vs"
using assms action_proj_disj_diff[of a "prob_dom PROB" vs' vs]
by blast
}
then show ?thesis
unfolding prob_proj_def
by (smt image_cong image_image)
qed
lemma n_replace_proj_le_n_as_2:
fixes a vs vs'
assumes "(vs \<subseteq> vs')" "(varset_action a vs')"
shows "(varset_action (action_proj a vs') vs \<longleftrightarrow> varset_action a vs)"
unfolding varset_action_def action_proj_def
using assms
by (simp add: exec_drest_5 varset_action_def)
\<comment> \<open>NOTE type of `PROB` had to be fixed for use of `empty\_problem\_bound`.\<close>
lemma empty_problem_proj_bound:
fixes PROB :: "'a problem"
shows "problem_plan_bound (prob_proj PROB {}) = 0"
proof -
\<comment> \<open>TODO refactor?\<close>
{
have "prob_proj {} {} = {}"
unfolding prob_proj_def action_proj_def
using image_empty
by simp
moreover {
assume P: "PROB \<noteq> {}"
have "\<forall>a. (fmrestrict_set {} (fst a), fmrestrict_set {} (snd a)) = (fmempty, fmempty)"
using fmrestrict_set_null
by simp
then have "prob_proj PROB {} = {(fmempty, fmempty)}"
unfolding prob_proj_def action_proj_def
using P
by auto
}
ultimately consider
(i) "prob_proj PROB {} = {}"
| (ii) "prob_proj PROB {} = {(fmempty, fmempty)}"
by (cases "PROB = {}") force+
then have "prob_dom (prob_proj PROB {}) = {}"
unfolding prob_dom_def action_dom_def using fmdom'_empty
by (cases) force+
}
then show ?thesis
using empty_problem_bound[where PROB="prob_proj PROB {}"]
by blast
qed
lemma problem_plan_bound_works_proj:
fixes PROB :: "'a problem" and s as vs
assumes "finite PROB" "(s \<in> valid_states PROB)" "(as \<in> valid_plans PROB)" "(sat_precond_as s as)"
shows "(\<exists>as'.
(exec_plan (fmrestrict_set vs s) as' = exec_plan (fmrestrict_set vs s) (as_proj as vs))
\<and> (length as' \<le> problem_plan_bound (prob_proj PROB vs))
\<and> (subseq as' (as_proj as vs))
\<and> (sat_precond_as s as')
\<and> (no_effectless_act as')
)"
proof -
{
have "exec_plan (fmrestrict_set vs s) (as_proj as vs) = fmrestrict_set vs (exec_plan s as)"
using assms(4) sat_precond_exec_as_proj_eq_proj_exec
by blast
moreover have "fmrestrict_set vs s \<in> valid_states (prob_proj PROB vs)"
using assms(2) graph_plan_not_eq_last_diff_paths
by auto
moreover have "as_proj as vs \<in> valid_plans (prob_proj PROB vs)"
using assms(3) valid_as_valid_as_proj
by blast
moreover have "finite (prob_proj PROB vs)"
unfolding prob_proj_def
using assms(1)
by simp
ultimately have "\<exists>as'.
exec_plan (fmrestrict_set vs s) (as_proj as vs) = exec_plan (fmrestrict_set vs s) as'
\<and> subseq as' (as_proj as vs) \<and> length as' \<le> problem_plan_bound (prob_proj PROB vs)
"
using problem_plan_bound_works[of "prob_proj PROB vs"
"fmrestrict_set vs s" "as_proj as vs"]
by blast
}
then obtain as' where
"exec_plan (fmrestrict_set vs s) (as_proj as vs) = exec_plan (fmrestrict_set vs s) as'"
"subseq as' (as_proj as vs) \<and> length as' \<le> problem_plan_bound (prob_proj PROB vs)"
by fast
moreover {
have "
exec_plan (fmrestrict_set vs s) as
= exec_plan (fmrestrict_set vs s) (rem_condless_act (fmrestrict_set vs s) [] as)
"
using rem_condless_valid_1[of "fmrestrict_set vs s" as]
by blast
then have "subseq (rem_condless_act (fmrestrict_set vs s) [] as') as'"
using rem_condless_valid_8 [of "fmrestrict_set vs s" as']
by blast
}
moreover have "length (rem_condless_act (fmrestrict_set vs s) [] as') \<le> length as'"
using rem_condless_valid_3[of "fmrestrict_set vs s"]
by fast
moreover have 4:
"sat_precond_as (fmrestrict_set vs s) (rem_condless_act (fmrestrict_set vs s) [] as')"
using rem_condless_valid_2[of "fmrestrict_set vs s" as']
by blast
moreover have "
exec_plan (fmrestrict_set vs s) (rem_condless_act (fmrestrict_set vs s) [] as')
= exec_plan (fmrestrict_set vs s)
(rem_effectless_act (rem_condless_act (fmrestrict_set vs s) [] as'))
"
using rem_effectless_works_1[of "fmrestrict_set vs s"
"rem_condless_act (fmrestrict_set vs s) [] as'"]
by blast
moreover {
have "
subseq (rem_effectless_act (rem_condless_act (fmrestrict_set vs s) [] as))
(rem_condless_act (fmrestrict_set vs s) [] as)
"
using rem_effectless_works_9[of
"(rem_condless_act (fmrestrict_set vs s) [] (as :: 'a action list))"]
by blast
then have "
length (rem_effectless_act (rem_condless_act (fmrestrict_set vs s) [] as'))
\<le> length (rem_condless_act (fmrestrict_set vs s) [] as')
"
using rem_effectless_works_3[of
"(rem_condless_act (fmrestrict_set vs s) [] (as' :: 'a action list))"]
by simp
then have "
sat_precond_as (fmrestrict_set vs s)
(rem_effectless_act (rem_condless_act (fmrestrict_set vs s) [] as'))
"
using 4 rem_effectless_works_2[of "fmrestrict_set vs s"
"(rem_condless_act (fmrestrict_set vs s) [] as')"]
by blast
then have
"no_effectless_act (rem_effectless_act (rem_condless_act (fmrestrict_set vs s) [] as'))"
using rem_effectless_works_6[of "(rem_condless_act (fmrestrict_set vs s) [] (as' ::'a action list))"]
by simp
}
ultimately show ?thesis
using rem_effectless_works_13 rem_condless_valid_1 order_trans
no_effectless_proj sat_precond_drest_sat_precond subseq_order.order_trans
by (metis (no_types, lifting))
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor into `Fmap\_Utils`.\<close>
lemma action_proj_inter_i: "fmrestrict_set V (fmrestrict_set W f) = fmrestrict_set (V \<inter> W) f"
unfolding fmfilter_alt_defs(4)
by simp
lemma action_proj_inter: "action_proj (action_proj a vs1) vs2 = action_proj a (vs1 \<inter> vs2)"
proof -
have
"fmrestrict_set vs2 (fmrestrict_set vs1 (fst a)) = fmrestrict_set (vs1 \<inter> vs2) (fst a)"
"fmrestrict_set vs2 (fmrestrict_set vs1 (snd a)) = fmrestrict_set (vs1 \<inter> vs2) (snd a)"
using inf_commute action_proj_inter_i
by metis+
then show ?thesis
unfolding action_proj_def
by simp
qed
lemma prob_proj_inter: "prob_proj (prob_proj PROB vs1) vs2 = prob_proj PROB (vs1 \<inter> vs2)"
unfolding prob_proj_def
using set_eq_iff image_iff action_proj_inter
supply[[smt_timeout=100]]
by (smt image_cong image_image)
subsection "Snapshotting"
text \<open> A snapshot is an abstraction concept of the system in which the assignment of a set of
variables is fixed and actions whose preconditions or effects violate the fixed assignments are
eliminated. [Abdulaziz et al., p.28]
Formally this notion is build on the definition of agreement of states (`agree`), which
states that variables `v`, `v'`in the shared domain of two states must be assigned to the same
value. A snapshot w.r.t to a state `s` is then defined as the set of actions of a problem where the
precondition and the effect agree. [Abdulaziz et al., Definition 16, HOL4 Definition 16, p.28] \<close>
\<comment> \<open>NOTE name shortened.\<close>
definition agree where
"agree s1 s2 \<equiv> (\<forall>v. (v \<in> fmdom' s1) \<and> (v \<in> fmdom' s2) \<longrightarrow> (fmlookup s1 v = fmlookup s2 v))"
\<comment> \<open>NOTE added lemma.\<close>
lemma state_succ_fixpoint_if:
fixes a s PROB
assumes "a \<in> PROB" "(s \<in> valid_states PROB)" "fst a \<subseteq>\<^sub>f s" "agree (snd a) s"
shows "state_succ s a = s"
proof -
{
have "fmdom' (snd a) \<subseteq> fmdom' s"
using assms(1, 2) FDOM_eff_subset_FDOM_valid_states_pair
by blast
moreover have "\<forall>x. x \<in> fmdom' (snd a) \<longrightarrow> fmlookup (snd a) x = fmlookup s x"
using assms(4) calculation(1) agree_def subsetCE
by metis
moreover have "s ++\<^sub>f snd a = s"
using calculation(2)
by (metis fmap_ext fmdom'_notD fmdom_notI fmlookup_add)
}
then show ?thesis
using fmap_add_ltr_def state_succ_def
by metis
qed
lemma agree_state_succ_idempot:
assumes "(a \<in> PROB)" "(s \<in> valid_states PROB)" "(agree (snd a) s)"
shows "(state_succ s a = s)"
proof (cases "fst a \<subseteq>\<^sub>f s")
case True
then show ?thesis
using assms state_succ_fixpoint_if
by blast
next
case False
then show ?thesis
unfolding state_succ_def fmap_add_ltr_def
by simp
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor into `Fmap\_Utils`.\<close>
lemma fmdom'_fmrestrict_set:
fixes X f
shows "fmdom' (fmrestrict_set X f) = X \<inter> (fmdom' f)"
unfolding fmdom'_alt_def fmfilter_alt_defs(4)
by auto
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor into 'Fmap\_Utils'.\<close>
lemma fmdom'_fmrestrict_set_fmadd:
fixes X f g
shows "fmdom' (fmrestrict_set X (f ++\<^sub>f g)) = X \<inter> (fmdom' f \<union> fmdom' g)"
proof -
have "fmrestrict_set X (f ++\<^sub>f g) = fmrestrict_set X f ++\<^sub>f fmrestrict_set X g"
using fmrestrict_set_add_distrib
by fast
then show ?thesis
using fmdom'_fmrestrict_set fmdom'_add
by metis
qed
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor into 'Fmap\_Utils'.\<close>
lemma fmrestrict_agree:
fixes X x f g
assumes "agree (fmrestrict_set X f) (fmrestrict_set X g)" "x \<in> X \<inter> fmdom' f \<inter> fmdom' g"
shows "fmlookup (fmrestrict_set X f) x = fmlookup (fmrestrict_set X g) x"
proof -
{
fix v
assume "v \<in> X \<inter> fmdom' f \<inter> fmdom' g"
then have "v \<in> fmdom' (fmrestrict_set X f) \<and> v \<in> fmdom' (fmrestrict_set X g)"
using fmdom'_fmrestrict_set
by force
then have "fmlookup (fmrestrict_set X f) v = fmlookup (fmrestrict_set X g) v"
using assms(1)
unfolding agree_def
by blast
}
then show ?thesis
using assms
by blast
qed
lemma agree_restrict_state_succ_idempot:
assumes "(a \<in> PROB)" "(s \<in> valid_states PROB)"
"(agree (fmrestrict_set vs (snd a)) (fmrestrict_set vs s))"
shows "(fmrestrict_set vs (state_succ s a) = fmrestrict_set vs s)"
proof (cases "fst a \<subseteq>\<^sub>f s")
case True
then have "state_succ s a = s ++\<^sub>f snd a"
unfolding state_succ_def fmap_add_ltr_def
by simp
{
fix v
have "fmlookup (fmrestrict_set vs (s ++\<^sub>f snd a)) v = fmlookup (fmrestrict_set vs s) v"
proof (cases "v \<in> fmdom' (snd a)")
case True
then have 1: "fmdom' (fmrestrict_set vs (s ++\<^sub>f snd a)) = vs \<inter> (fmdom' s \<union> fmdom' (snd a))"
unfolding fmap_add_ltr_def
using fmdom'_fmrestrict_set_fmadd
by metis
then have 2: "fmdom' (fmrestrict_set vs (snd a)) = vs \<inter> fmdom' (snd a)"
using fmdom'_fmrestrict_set
by metis
then show ?thesis
using 1 2
proof (cases "v \<in> vs")
case true: True
then show ?thesis
proof (cases "v \<in> (fmdom' s \<inter> fmdom' (snd a))")
case True
then have "v \<in> vs \<inter> fmdom' s \<inter> fmdom' (snd a)"
using true
by blast
then have "fmlookup (fmrestrict_set vs (snd a)) v = fmlookup (fmrestrict_set vs s) v"
using assms(3) fmrestrict_agree
by fast
then show ?thesis
by fastforce
next
case False
then have "fmdom' (snd a) \<subseteq> fmdom' s"
using assms(1, 2) FDOM_eff_subset_FDOM_valid_states_pair
by metis
then have "v \<notin> fmdom' (snd a)"
using true False
by blast
then show ?thesis
by fastforce
qed
qed auto
qed fastforce
}
then show ?thesis
unfolding state_succ_def fmap_add_ltr_def
using fmap_ext
by metis
next
case False
then show ?thesis
unfolding state_succ_def
by simp
qed
lemma agree_exec_idempot:
assumes "(as \<in> valid_plans PROB)" "(s \<in> valid_states PROB)"
"(\<forall>a. ListMem a as \<longrightarrow> agree (snd a) s)"
shows "(exec_plan s as = s)"
using assms
proof (induction as arbitrary: PROB s)
case (Cons a as)
then have 1: "a \<in> PROB"
using Cons.prems(1) valid_plan_valid_head
by fast
then have 2: "as \<in> valid_plans PROB"
using Cons.prems(1) valid_plan_valid_tail
by fast
then have 3: "\<forall>a. ListMem a as \<longrightarrow> agree (snd a) s"
using Cons.prems(3) ListMem.simps
by metis
then have "ListMem a (a # as)"
using elem
by fast
then have "agree (snd a) s"
using Cons.prems(3)
by blast
then have 4: "state_succ s a = s"
using Cons.prems(1, 2) 1 agree_state_succ_idempot
by blast
then have "exec_plan s as = s"
using Cons.IH Cons.prems(2) 2 3
by blast
then show ?case
using 4
by simp
qed simp
lemma agree_restrict_exec_idempot:
fixes s s'
assumes "(as \<in> valid_plans PROB)" "(s' \<in> valid_states PROB)" "(s \<in> valid_states PROB)"
"(\<forall>a. ListMem a as \<longrightarrow> agree (fmrestrict_set vs (snd a)) (fmrestrict_set vs s))"
"(fmrestrict_set vs s' = fmrestrict_set vs s)"
shows "(fmrestrict_set vs (exec_plan s' as) = fmrestrict_set vs s)"
using assms
proof (induction as arbitrary: PROB s s' vs)
case (Cons a as)
have 1: "as \<in> valid_plans PROB"
using Cons.prems(1) valid_plan_valid_tail
by fast
then have 2: "\<forall>a. ListMem a as \<longrightarrow> agree (fmrestrict_set vs (snd a)) (fmrestrict_set vs s)"
using Cons.prems(4) ListMem.simps
by metis
then have 3: "a \<in> PROB"
using Cons.prems(1) valid_plan_valid_head
by metis
moreover
{
have "ListMem a (a # as)"
using elem
by fast
then have "agree (fmrestrict_set vs (snd a)) (fmrestrict_set vs s)"
using Cons.prems(4) calculation(1)
by blast
then have "agree (fmrestrict_set vs (snd a)) (fmrestrict_set vs s')"
using Cons.prems(5)
by simp
}
ultimately show ?case
using assms
proof (cases "fst a \<subseteq>\<^sub>f s'")
case True
{
have a: "s' \<in> valid_states PROB"
using Cons.prems(2)
by simp
moreover have "state_succ s' a \<in> valid_states PROB"
using 3 a lemma_1_i
by blast
moreover have
" \<forall>a. ListMem a as \<longrightarrow> agree (fmrestrict_set vs (snd a)) (fmrestrict_set vs s)"
using 2
by blast
moreover {
have "ListMem a (a # as)"
using elem
by fast
then have "agree (fmrestrict_set vs (snd a)) (fmrestrict_set vs s)"
using Cons.prems(4) calculation(1)
by blast
then have "fmrestrict_set vs (state_succ s' a) = fmrestrict_set vs s"
using Cons.prems(5) 3 a agree_restrict_state_succ_idempot
by metis
}
ultimately have "fmrestrict_set vs (exec_plan (state_succ s' a) as) = fmrestrict_set vs s"
using assms(3) 1 Cons.IH[where s'="state_succ s' a"]
by auto
}
then show ?thesis
by simp
next case False
moreover have "exec_plan s' (a # as) = exec_plan s' as"
using False
by (simp add: state_succ_def)
ultimately show ?thesis
using Cons.IH Cons.prems(2, 3, 5) 1 2
by presburger
qed
qed simp
lemma agree_restrict_exec_idempot_pair:
fixes s s'
assumes "(as \<in> valid_plans PROB)" "(s' \<in> valid_states PROB)" "(s \<in> valid_states PROB)"
"(\<forall>p e. ListMem (p, e) as \<longrightarrow> agree (fmrestrict_set vs e) (fmrestrict_set vs s))"
"(fmrestrict_set vs s' = fmrestrict_set vs s)"
shows "(fmrestrict_set vs (exec_plan s' as) = fmrestrict_set vs s)"
using assms agree_restrict_exec_idempot
by fastforce
lemma agree_comm: "agree x x' = agree x' x"
unfolding agree_def
by fastforce
lemma restricted_agree_imp_agree:
assumes "(fmdom' s2 \<subseteq> vs)" "(agree (fmrestrict_set vs s1) s2)"
shows "(agree s1 s2)"
using assms contra_subsetD fmlookup_restrict_set Int_iff fmdom'_fmrestrict_set
unfolding agree_def
by metis
lemma agree_imp_submap:
assumes "f1 \<subseteq>\<^sub>f f2"
shows "agree f1 f2"
using assms
unfolding agree_def
by (simp add: as_needed_asses_submap_exec_ii)
lemma agree_FUNION:
assumes "(agree fm fm1)" "(agree fm fm2)"
shows "(agree fm (fm1 ++ fm2))"
unfolding agree_def fmap_add_ltr_def
using assms
by (metis agree_def fmlookup_add fmlookup_dom'_iff)
lemma agree_fm_list_union:
fixes fm
assumes "(\<forall>fm'. ListMem fm' fmList \<longrightarrow> agree fm fm')"
shows "(agree fm (foldr fmap_add_ltr fmList fmempty))"
using assms proof (induction "fmList" arbitrary: fm)
case Nil
then have "foldr fmap_add_ltr [] fmempty = fmempty"
using Nil
by simp
then show ?case
unfolding agree_def
by auto
next
case (Cons a fmList)
then have "\<forall>fm'. ListMem fm' fmList \<longrightarrow> agree fm fm'"
using Cons.prems insert
by fast
then have 1: "agree fm (foldr fmap_add_ltr fmList fmempty)"
using Cons.IH
by blast
then have "agree fm a"
using Cons.prems elem
by fast
then have "agree fm (a ++ foldr fmap_add_ltr fmList fmempty)"
using 1 agree_FUNION
by blast
then show ?case
by simp
qed
lemma DRESTRICT_EQ_AGREE:
assumes "(fmdom' s2 \<subseteq> vs2)" "(fmdom' s1 \<subseteq> vs1)"
shows "((fmrestrict_set vs2 s1 = fmrestrict_set vs1 s2) \<longrightarrow> agree s1 s2)"
using assms fmdom'_restrict_set restricted_agree_imp_agree
by (metis agree_def)
lemma SUBMAPS_AGREE: "(s1 \<subseteq>\<^sub>f s) \<and> (s2 \<subseteq>\<^sub>f s) \<Longrightarrow> (agree s1 s2)"
unfolding agree_def
by (metis as_needed_asses_submap_exec_ii)
\<comment> \<open>NOTE name shortened.\<close>
definition snapshot where
"snapshot PROB s = {a | a. a \<in> PROB \<and> agree (fst a) s \<and> agree (snd a) s}"
lemma snapshot_pair: "snapshot PROB s = {(p, e). (p, e) \<in> PROB \<and> agree p s \<and> agree e s}"
unfolding snapshot_def
by fastforce
lemma action_agree_valid_in_snapshot:
assumes "(a \<in> PROB)" "(agree (fst a) s)" "(agree (snd a) s)"
shows "(a \<in> snapshot PROB s)"
unfolding snapshot_def
using assms
by blast
lemma as_mem_agree_valid_in_snapshot:
assumes "(\<forall>a. ListMem a as \<longrightarrow> agree (fst a) s \<and> agree (snd a) s)" "(as \<in> valid_plans PROB)"
shows "(as \<in> valid_plans (snapshot PROB s))"
using assms
proof (induction as)
case Nil
then show ?case
using empty_plan_is_valid
by blast
next
case (Cons a as)
{
have "\<forall>a. ListMem a as \<longrightarrow> agree (fst a) s \<and> agree (snd a) s"
using Cons.prems(1) insert
by fast
moreover have "(as \<in> valid_plans PROB)"
using Cons.prems(2) valid_plan_valid_tail
by fast
ultimately have "set as \<subseteq> snapshot PROB s"
using Cons.IH valid_plans_def
by fast
}
note 1 = this
{
have a: "a \<in> PROB"
using Cons.prems(2) valid_plan_valid_head
by metis
then have "ListMem a (a # as)"
using elem
by fast
then have "agree (fst a) s \<and> agree (snd a) s"
using Cons.prems(1)
by blast
then have "a \<in> snapshot PROB s"
using a snapshot_def
by auto
}
then have "set (a # as) \<subseteq> snapshot PROB s"
using 1 set_simps(2)
by simp
then show ?case using valid_plans_def
by blast
qed
lemma fmrestrict_agree_monotonous:
fixes f g X
assumes "agree f g"
shows "agree (fmrestrict_set X f) (fmrestrict_set X g)"
proof -
let ?F="fmdom' (fmrestrict_set X f)"
let ?G="fmdom' (fmrestrict_set X g)"
have 1: "?F = X \<inter> fmdom' f" "?G = X \<inter> fmdom' g"
using fmdom'_fmrestrict_set
by metis+
{
fix v
assume "v \<in> ?F" "v \<in> ?G"
then have "v \<in> fmdom' f" "v \<in> fmdom' g"
using 1
by blast+
then have "fmlookup f v = fmlookup g v"
using assms
unfolding agree_def
by blast
then have "fmlookup (fmrestrict_set X f) v = fmlookup (fmrestrict_set X g) v"
unfolding fmlookup_restrict_set
by argo
}
then show ?thesis
using assms
unfolding agree_def
by blast
qed
\<comment> \<open>TODO remove if not used.\<close>
lemma SUBMAP_FUNION_DRESTRICT_i:
fixes v vsa vsb f g
assumes "v \<in> vsa"
shows "
fmlookup (fmrestrict_set ((vsa \<union> vsb) \<inter> vs) f) v
= fmlookup (fmrestrict_set (vsa \<inter> vs) f) v
"
unfolding fmlookup_restrict_set
using assms
by auto
lemma SUBMAP_FUNION_DRESTRICT':
assumes "(agree fma fmb)" "(vsa \<subseteq> fmdom' fma)" "(vsb \<subseteq> fmdom' fmb)"
"(fmrestrict_set vsa fm = fmrestrict_set (vsa \<inter> vs) fma)"
"(fmrestrict_set vsb fm = fmrestrict_set (vsb \<inter> vs) fmb)"
shows "(fmrestrict_set (vsa \<union> vsb) fm = fmrestrict_set ((vsa \<union> vsb) \<inter> vs) (fma ++ fmb))"
proof -
let ?f="fmrestrict_set (vsa \<union> vsb) fm"
let ?g="fmrestrict_set ((vsa \<union> vsb) \<inter> vs) (fma ++ fmb)"
have 1: "?g = fmrestrict_set ((vsa \<union> vsb) \<inter> vs) fmb ++\<^sub>f fmrestrict_set ((vsa \<union> vsb) \<inter> vs) fma"
unfolding fmap_add_ltr_def fmrestrict_set_add_distrib
by simp
have 2: "agree (fmrestrict_set ((vsa \<union> vsb) \<inter> vs) fma) (fmrestrict_set ((vsa \<union> vsb) \<inter> vs) fmb)"
using assms(1) fmrestrict_agree_monotonous
by blast
have 3:
"fmdom' (fmrestrict_set ((vsa \<union> vsb) \<inter> vs) fma) = ((vsa \<union> vsb) \<inter> vs) \<inter> fmdom' fma"
"fmdom' (fmrestrict_set ((vsa \<union> vsb) \<inter> vs) fmb) = ((vsa \<union> vsb) \<inter> vs) \<inter> fmdom' fmb"
using fmdom'_fmrestrict_set
by metis+
{
fix v
have "fmlookup ?f v = fmlookup ?g v"
proof (cases "v \<in> ((vsa \<union> vsb) \<inter> vs)")
case True
\<comment> \<open>TODO unwrap smt proof.\<close>
then show ?thesis
using assms(1, 2, 3, 4, 5) 1
- by (smt IntD1 SUBMAP_FUNION_DRESTRICT_i UnE agree_def domIff fmdom'.rep_eq fmdom'_alt_def
- fmdom'_fmrestrict_set fmlookup_add fmlookup_restrict_set inf_sup_distrib2 notin_fset
+ by (smt (verit) IntD1 SUBMAP_FUNION_DRESTRICT_i UnE agree_def domIff fmdom'.rep_eq fmdom'_alt_def
+ fmdom'_fmrestrict_set fmlookup_add fmlookup_restrict_set inf_sup_distrib2
subset_iff sup_commute)
next
case False
then show ?thesis
proof -
have "v \<notin> vsa \<union> vsb \<or> v \<notin> vs"
using False
by blast
then have "fmlookup (fmrestrict_set (vsa \<union> vsb) fm) v = None"
using assms(4, 5)
by (metis Int_iff Un_iff fmlookup_restrict_set)
then show ?thesis
using False
by auto
qed
qed
}
then show ?thesis
using 1 fmap_ext
by blast
qed
lemma UNION_FUNION_DRESTRICT_SUBMAP:
assumes "(vs1 \<subseteq> fmdom' fma)" "(vs2 \<subseteq> fmdom' fmb)" "(agree fma fmb)"
"(fmrestrict_set vs1 fma \<subseteq>\<^sub>f s)" "(fmrestrict_set vs2 fmb \<subseteq>\<^sub>f s)"
shows "(fmrestrict_set (vs1 \<union> vs2) (fma ++ fmb) \<subseteq>\<^sub>f s)"
proof -
{
let ?f="fmrestrict_set (vs1 \<union> vs2) (fma ++ fmb)"
fix v
assume P: "v \<in> fmdom' ?f"
{
have "v \<in> (vs1 \<union> vs2) \<inter> (fmdom' fma \<union> fmdom' fmb)"
using P
unfolding fmap_add_ltr_def fmdom'_fmrestrict_set fmdom'_add
by force
then have "v \<in> vs1 \<union> vs2" "v \<in> fmdom' fma \<union> fmdom' fmb"
by fast+
}
note 1 = this
then have 2: "fmlookup ?f v = fmlookup (fmb ++\<^sub>f fma) v"
unfolding fmlookup_restrict_set fmap_add_ltr_def
by argo
then consider
(i) "v \<in> vs1"
| (ii) "v \<in> vs2"
| (iii) "\<not>v\<in> vs1 \<and> \<not>v\<in>vs2"
by blast
then have "fmlookup ?f v = fmlookup s v"
proof (cases)
case i
then have "v \<in> fmdom' fma"
using assms(1)
by blast
then have "fmlookup ?f v = fmlookup fma v"
unfolding 2 fmlookup_add
- by (simp add: fmdom'_alt_def notin_fset)
+ by (simp add: fmdom'_alt_def)
also have "\<dots> = fmlookup (fmrestrict_set vs1 fma) v"
unfolding fmlookup_restrict_set
using i
by simp
finally show ?thesis
using assms(4)
by (metis (mono_tags, lifting) P domIff fmdom'_notI fmsubset.rep_eq map_le_def)
next
\<comment> \<open>TODO unwrap smt proof.\<close>
case ii
then show ?thesis
using assms(2, 3, 5) 2 P
by (smt SUBMAP_FUNION_DRESTRICT_i agree_def
fmdom'.rep_eq fmdom'_fmrestrict_set fmdom'_notD fmdom'_notI fmlookup_add
fmrestrict_set_dom fmsubset.rep_eq inf.orderE map_le_def subset_Un_eq)
next
case iii
then show ?thesis
using 1
by blast
qed
}
then show ?thesis
by (simp add: as_needed_asses_submap_exec_vii)
qed
\<comment> \<open>TODO unwrap sledgehammered metis proof.\<close>
(* TODO duplicate lemma? *)
lemma agree_DRESTRICT:
assumes "agree s1 s2"
shows "agree (fmrestrict_set vs s1) (fmrestrict_set vs s2)"
using assms by (fact fmrestrict_agree_monotonous)
lemma agree_DRESTRICT_2:
assumes "(fmdom' s1 \<subseteq> vs1)" "(fmdom' s2 \<subseteq> vs2)" "(agree s1 s2)"
shows "(agree (fmrestrict_set vs2 s1) (fmrestrict_set vs1 s2))"
using assms
unfolding agree_def fmdom'_restrict_set_precise
by auto
\<comment> \<open>NOTE added lemma.\<close>
lemma snapshot_eq_filter:
shows "snapshot PROB s = Set.filter (\<lambda>a. agree (fst a) s \<and> agree (snd a) s) PROB"
unfolding snapshot_def Set.filter_def
by presburger
\<comment> \<open>NOTE moved up.\<close>
corollary snapshot_subset:
shows "snapshot PROB s \<subseteq> PROB"
unfolding snapshot_def
using snapshot_eq_filter
by blast
lemma FINITE_snapshot:
assumes "finite PROB"
shows "finite (snapshot PROB s)"
proof -
have "snapshot PROB s \<subseteq> PROB"
using snapshot_subset
by blast
then show ?thesis
using assms finite_subset[of "snapshot PROB s" PROB]
by blast
qed
\<comment> \<open>NOTE moved up (declared above the previous lemma).
lemma snapshot\_subset\<close>
\<comment> \<open>TODO unwrap metis proof.\<close>
lemma dom_proj_snapshot:
"prob_dom (prob_proj PROB (prob_dom (snapshot PROB s))) = prob_dom (snapshot PROB s)"
by (metis snapshot_subset two_children_parent_mems_le_finite prob_subset_dom_subset)
lemma valid_states_snapshot:
"valid_states (prob_proj PROB (prob_dom (snapshot PROB s))) = valid_states (snapshot PROB s)"
by (metis dom_proj_snapshot valid_states_def)
lemma valid_proj_neq_succ_restricted_neq_succ:
assumes "(x' \<in> prob_proj PROB vs)" "(state_succ s x' \<noteq> s)"
shows "(fmrestrict_set vs (state_succ s x') \<noteq> fmrestrict_set vs s)"
unfolding state_succ_def
using FDOM_eff_subset_prob_dom_pair dom_prob_proj limited_dom_neq_restricted_neq
using assms(1, 2)
by (smt dual_order.trans state_succ_def)
lemma proj_successors: "
((\<lambda>s. fmrestrict_set vs s) ` (state_successors (prob_proj PROB vs) s))
\<subseteq> (state_successors (prob_proj PROB vs) (fmrestrict_set vs s))
"
proof -
let ?A="((\<lambda>s. fmrestrict_set vs s) ` (state_successors (prob_proj PROB vs) s))"
let ?B="(state_successors (prob_proj PROB vs) (fmrestrict_set vs s))"
{
fix x
assume P: "x \<in> ?A"
then obtain x' x'' where a:
"x'' \<in> prob_proj PROB vs" "x' = state_succ s x''" "x' \<noteq> s" "x = fmrestrict_set vs x'"
unfolding state_successors_def subset_iff
by blast
moreover {
have "(\<exists>x''.
x'' \<in> prob_proj PROB vs \<and> x = state_succ (fmrestrict_set vs s) x''
\<and> x \<noteq> fmrestrict_set vs s)"
proof (cases "fst x'' \<subseteq>\<^sub>f s")
case true: True
then show ?thesis
proof (cases "fst x'' \<subseteq>\<^sub>f fmrestrict_set vs s")
case True
{
have "fmdom' (snd x'') \<subseteq> vs"
using a(1) FDOM_eff_subset_prob_dom_pair dom_prob_proj dual_order.trans
by metis
then have "fmrestrict_set vs (snd x'') = snd x''"
using exec_drest_5
by fast
}
note i = this
{
have "x = fmrestrict_set vs (snd x'' ++ s)"
using a(2, 4) true
unfolding state_succ_def
by simp
then have "x = fmrestrict_set vs (snd x'') ++ fmrestrict_set vs s"
unfolding fmap_add_ltr_def
using fmrestrict_set_add_distrib
by simp
then have "x = snd x'' ++ fmrestrict_set vs s"
using i
by simp
then have "x = state_succ (fmrestrict_set vs s) x''"
unfolding state_succ_def
using True
by argo
}
moreover have "x \<noteq> fmrestrict_set vs s"
using a valid_proj_neq_succ_restricted_neq_succ
by fast
ultimately show ?thesis
using a(1)
by blast
next
case False
then show ?thesis
proof -
have "x'' \<in> (\<lambda>p. action_proj p vs) ` PROB"
using calculation(1) prob_proj_def
by auto
then have "action_proj x'' vs = x''"
using action_proj_idempot
by blast
then show ?thesis
by (metis (no_types) False action_proj_pair fmsubset_restrict_set_mono fstI
surjective_pairing true)
qed
qed
next
case False
then show ?thesis
proof (cases "fst x'' \<subseteq>\<^sub>f fmrestrict_set vs s")
case True
then have "fmdom' (snd x'') \<subseteq> vs"
using FDOM_eff_subset_prob_dom_pair dom_prob_proj
using a(1) dual_order.trans
by metis
then have "fmrestrict_set vs (snd x'') = snd x''"
using exec_drest_5
by fast
then show ?thesis
unfolding state_succ_def fmap_add_ltr_def
using False True sublist_as_proj_eq_as_1
by fast
next
case False
then have "fmdom' (fst x'') \<subseteq> vs"
using FDOM_pre_subset_prob_dom_pair dom_prob_proj
using a(1) dual_order.trans
by metis
then have "fmrestrict_set vs (fst x'') = fst x''"
by (simp add: exec_drest_5)
then show ?thesis
unfolding state_succ_def fmap_add_ltr_def
using a False fmsubset_restrict_set_mono
by (metis state_succ_def)
qed
qed
}
then obtain x'' where "x'' \<in> prob_proj PROB vs" "x = state_succ (fmrestrict_set vs s) x''"
"x \<noteq> fmrestrict_set vs s"
by blast
then have "x \<in> ?B" unfolding state_successors_def
by blast
}
then show ?thesis
by blast
qed
lemma state_in_successor_proj_in_state_in_successor: "
(s' \<in> state_successors (prob_proj PROB vs) s)
\<Longrightarrow> (fmrestrict_set vs s' \<in> state_successors (prob_proj PROB vs) (fmrestrict_set vs s))"
using proj_successors
by force
lemma proj_FDOM_eff_subset_FDOM_valid_states:
fixes p e s
assumes "((p, e) \<in> prob_proj PROB vs)" "(s \<in> valid_states PROB)"
shows "(fmdom' e \<subseteq> fmdom' s)"
using assms
proof -
{
obtain p' e' where "(p', e') \<in> PROB" "(p, e) = action_proj (p', e') vs"
using assms(1)
unfolding prob_proj_def
by fast
then have "fmdom' e \<subseteq> prob_dom (prob_proj PROB vs)"
using assms FDOM_eff_subset_prob_dom
by blast
also have "\<dots> = prob_dom PROB \<inter> vs"
using graph_plan_neq_mems_state_set_neq_len
by fast
finally have "fmdom' e \<subseteq> prob_dom PROB"
by simp
}
moreover have "fmdom' s = prob_dom PROB"
using assms(2)
unfolding valid_states_def
by simp
ultimately show ?thesis
by simp
qed
lemma valid_proj_action_valid_succ:
assumes "(h \<in> prob_proj PROB vs)" "(s \<in> valid_states PROB)"
shows "(state_succ s h \<in> valid_states PROB)"
proof -
have "fmdom' (snd h) \<subseteq> fmdom' s"
using assms proj_FDOM_eff_subset_FDOM_valid_states surjective_pairing
by metis
moreover have "fmdom' (state_succ s h) = fmdom' s"
using calculation(1) FDOM_state_succ
by metis
ultimately show ?thesis
using assms(2) valid_states_def
by blast
qed
lemma proj_successors_of_valid_are_valid:
assumes "(s \<in> valid_states PROB)"
shows "(state_successors (prob_proj PROB vs) s \<subseteq> (valid_states PROB))"
unfolding state_successors_def
using assms valid_proj_action_valid_succ
by blast
subsection "State Space Projection"
\<comment> \<open>NOTE name shortened.\<close>
definition ss_proj where
"ss_proj ss vs \<equiv> (\<lambda>s. fmrestrict_set vs s) ` ss"
\<comment> \<open>NOTE added lemma.\<close>
\<comment> \<open>TODO refactor into 'Fmap\_Utils'.\<close>
lemma fmrestrict_set_inter_img:
fixes A X Y
shows "fmrestrict_set (X \<inter> Y) ` A = (fmrestrict_set X \<circ> fmrestrict_set Y) ` A"
proof -
\<comment> \<open>NOTE Proof by mutual inclusion.\<close>
let ?lhs = "fmrestrict_set (X \<inter> Y) ` A"
let ?rhs = "(fmrestrict_set X \<circ> fmrestrict_set Y) ` A"
{
fix a
assume "a \<in> A"
have "(fmrestrict_set X \<circ> fmrestrict_set Y) a = fmrestrict_set X (fmrestrict_set Y a)"
by auto
also have "\<dots> = fmrestrict_set (X \<inter> Y) a"
using action_proj_inter_i
by fast
finally have "(fmrestrict_set X \<circ> fmrestrict_set Y) a = fmrestrict_set (X \<inter> Y) a"
by auto
}
note 1 = this
{
fix a
assume P: "a \<in> A"
then have "fmrestrict_set (X \<inter> Y) a \<in> ?lhs"
by simp
moreover have "(fmrestrict_set X \<circ> fmrestrict_set Y) a \<in> ?rhs"
using P
by blast
ultimately have
"fmrestrict_set (X \<inter> Y) a \<in> ?rhs" "(fmrestrict_set X \<circ> fmrestrict_set Y) a \<in> ?lhs"
using P 1
by metis+
}
then show ?thesis
by blast
qed
lemma invariantStateSpace_thm_9:
fixes ss vs1 vs2
shows "ss_proj ss (vs1 \<inter> vs2) = ss_proj (ss_proj ss vs2) vs1"
proof -
{
have "
ss_proj ss (vs1 \<inter> vs2)
= fmrestrict_set (vs1 \<inter> vs2) ` ss
"
unfolding ss_proj_def
by simp
also have "\<dots> = (fmrestrict_set vs1 \<circ> fmrestrict_set vs2) ` ss"
using fmrestrict_set_inter_img
by metis
finally have "ss_proj ss (vs1 \<inter> vs2) = ss_proj (ss_proj ss vs2) vs1"
unfolding ss_proj_def
by force
}
then show ?thesis
by simp
qed
lemma FINITE_ss_proj:
fixes ss vs
assumes "finite ss"
shows "finite (ss_proj ss vs)"
unfolding ss_proj_def
using assms
by simp
lemma nempty_stateSpace_nempty_ss_proj:
assumes "(ss \<noteq> {})"
shows "(ss_proj ss vs \<noteq> {})"
unfolding ss_proj_def
using assms
by simp
lemma invariantStateSpace_thm_5:
fixes ss vs domain
assumes "(stateSpace ss domain)"
shows "(stateSpace (ss_proj ss vs) (domain \<inter> vs))"
using assms
unfolding stateSpace_def ss_proj_def
by (metis (no_types, lifting) fmdom'_fmrestrict_set imageE inf_commute)
lemma dom_subset_ssproj_eq_ss:
fixes ss domain vs
assumes "(stateSpace ss domain)" "(domain \<subseteq> vs)"
shows "(ss_proj ss vs = ss)"
unfolding ss_proj_def stateSpace_def
using assms exec_drest_5
by (metis (mono_tags, lifting) image_cong image_ident stateSpace_def)
\<comment> \<open>TODO refactor duplicate proof steps in case split.\<close>
lemma neq_vs_neq_ss_proj:
fixes vs
assumes "(ss \<noteq> {})" "(stateSpace ss vs)" "(vs1 \<subseteq> vs)" "(vs2 \<subseteq> vs)" "(vs1 \<noteq> vs2)"
shows "(ss_proj ss vs1 \<noteq> ss_proj ss vs2)"
proof -
{
have 1: "\<exists>f. f \<in> ss"
using assms(1)
by blast
then obtain x where " (x \<in> vs1 \<and> x \<notin> vs2) \<or> (x \<in> vs2 \<and> x \<notin> vs1)"
using assms(5)
by blast
then consider (i) "x \<in> vs1 \<and> x \<notin> vs2" | (ii) "x \<in> vs2 \<and> x \<notin> vs1"
by blast
then have "fmrestrict_set vs1 ` ss \<noteq> fmrestrict_set vs2 ` ss" proof (cases)
case i
{
fix s' t'
assume "s' \<in> fmrestrict_set vs1 ` ss" "t' \<in> fmrestrict_set vs2 ` ss"
then obtain s t where a:
"s \<in> ss" "s' = fmrestrict_set vs1 s" "t \<in> ss" "t' = fmrestrict_set vs2 t"
by blast
then have "fmdom' s = vs"
using assms(2)
by (simp add: stateSpace_def)
then have b: "fmdom' s' = vs1"
using assms(3) a fmdom'_fmrestrict_set inf.order_iff
by metis
then have "fmdom' t = vs"
using assms(2) a(3)
by (simp add: stateSpace_def)
then have "fmdom' t' = vs2"
using assms(4) a(4) fmdom'_fmrestrict_set inf.order_iff
by metis
then have "fmlookup s' x \<noteq> None" "fmlookup t' x = None"
using i b domIff fmdom'_alt_def fmdom.rep_eq
by metis+
then have "s' \<noteq> t'"
by blast
}
then show ?thesis
using 1 neq_funs_neq_images
by blast
next
case ii
{
fix s' t'
assume "s' \<in> fmrestrict_set vs1 ` ss" "t' \<in> fmrestrict_set vs2 ` ss"
then obtain s t where c:
"s \<in> ss" "s' = fmrestrict_set vs1 s" "t \<in> ss" "t' = fmrestrict_set vs2 t"
by blast
then have "fmdom' s = vs"
using assms(2)
by (simp add: stateSpace_def)
then have d: "fmdom' s' = vs1"
using assms(3) c(2) fmdom'_fmrestrict_set inf.order_iff
by metis
then have "fmdom' t = vs"
using assms(2) c(3)
by (simp add: stateSpace_def)
then have "fmdom' t' = vs2"
using assms(4) c(4) fmdom'_fmrestrict_set inf.order_iff
by metis
then have "fmlookup s' x = None" "fmlookup t' x \<noteq> None"
using ii d domIff fmdom'_alt_def fmdom.rep_eq
by metis+
then have "s' \<noteq> t'"
by blast
}
then show ?thesis
using 1 neq_funs_neq_images
by blast
qed
}
then show ?thesis
unfolding ss_proj_def
by blast
qed
lemma subset_dom_stateSpace_ss_proj:
fixes vs1 vs2
assumes "(vs1 \<subseteq> vs2)" "(stateSpace ss vs2)"
shows "(stateSpace (ss_proj ss vs1) vs1)"
using assms
by (metis inf.absorb_iff2 invariantStateSpace_thm_5)
lemma card_proj_leq:
assumes "finite PROB"
shows "card (prob_proj PROB vs) \<le> card PROB"
unfolding prob_proj_def
using assms card_image_le
by blast
end
\ No newline at end of file
diff --git a/thys/Formula_Derivatives/WS1S_Alt_Formula.thy b/thys/Formula_Derivatives/WS1S_Alt_Formula.thy
--- a/thys/Formula_Derivatives/WS1S_Alt_Formula.thy
+++ b/thys/Formula_Derivatives/WS1S_Alt_Formula.thy
@@ -1,363 +1,362 @@
section \<open>Concrete Atomic WS1S Formulas (Singleton Semantics for FO Variables)\<close>
(*<*)
theory WS1S_Alt_Formula
imports
Abstract_Formula
WS1S_Prelim
begin
(*>*)
datatype (FOV0: 'fo, SOV0: 'so) atomic =
Fo 'fo |
Z 'fo |
Less 'fo 'fo |
In 'fo 'so
derive linorder atomic
type_synonym fo = nat
type_synonym so = nat
type_synonym ws1s = "(fo, so) atomic"
type_synonym formula = "(ws1s, order) aformula"
primrec wf0 where
"wf0 idx (Fo m) = LESS FO m idx"
| "wf0 idx (Z m) = LESS FO m idx"
| "wf0 idx (Less m1 m2) = (LESS FO m1 idx \<and> LESS FO m2 idx)"
| "wf0 idx (In m M) = (LESS FO m idx \<and> LESS SO M idx)"
inductive lformula0 where
"lformula0 (Fo m)"
| "lformula0 (Z m)"
| "lformula0 (Less m1 m2)"
| "lformula0 (In m M)"
code_pred lformula0 .
declare lformula0.intros[simp]
inductive_cases lformula0E[elim]: "lformula0 a"
abbreviation "FV0 \<equiv> case_order FOV0 SOV0"
fun find0 where
"find0 FO i (Fo m) = (i = m)"
| "find0 FO i (Z m) = (i = m)"
| "find0 FO i (Less m1 m2) = (i = m1 \<or> i = m2)"
| "find0 FO i (In m _) = (i = m)"
| "find0 SO i (In _ M) = (i = M)"
| "find0 _ _ _ = False"
abbreviation "decr0 ord k \<equiv> map_atomic (case_order (dec k) id ord) (case_order id (dec k) ord)"
primrec satisfies0 where
"satisfies0 \<AA> (Fo m) = (\<exists>x. m\<^bsup>\<AA>\<^esup>FO = {|x|})"
| "satisfies0 \<AA> (Z m) = (m\<^bsup>\<AA>\<^esup>FO = {||})"
| "satisfies0 \<AA> (Less m1 m2) =
(let P1 = m1\<^bsup>\<AA>\<^esup>FO; P2 = m2\<^bsup>\<AA>\<^esup>FO in if \<not>(\<exists>x. P1 = {|x|}) \<or> \<not>(\<exists>x. P2 = {|x|})
then False
else fthe_elem P1 < fthe_elem P2)"
| "satisfies0 \<AA> (In m M) =
(let P = m\<^bsup>\<AA>\<^esup>FO in if \<not>(\<exists>x. P = {|x|}) then False else fMin P |\<in>| M\<^bsup>\<AA>\<^esup>SO)"
fun lderiv0 where
"lderiv0 (bs1, bs2) (Fo m) = (if bs1 ! m then FBase (Z m) else FBase (Fo m))"
| "lderiv0 (bs1, bs2) (Z m) = (if bs1 ! m then FBool False else FBase (Z m))"
| "lderiv0 (bs1, bs2) (Less m1 m2) = (case (bs1 ! m1, bs1 ! m2) of
(False, False) \<Rightarrow> FBase (Less m1 m2)
| (True, False) \<Rightarrow> FAnd (FBase (Z m1)) (FBase (Fo m2))
| _ \<Rightarrow> FBool False)"
| "lderiv0 (bs1, bs2) (In m M) = (case (bs1 ! m, bs2 ! M) of
(False, _) \<Rightarrow> FBase (In m M)
| (True, True) \<Rightarrow> FBase (Z m)
| _ \<Rightarrow> FBool False)"
primrec rev where
"rev (Fo m) = Fo m"
| "rev (Z m) = Z m"
| "rev (Less m1 m2) = Less m2 m1"
| "rev (In m M) = In m M"
abbreviation "rderiv0 v \<equiv> map_aformula rev id o lderiv0 v o rev"
primrec nullable0 where
"nullable0 (Fo m) = False"
| "nullable0 (Z m) = True"
| "nullable0 (Less m1 m2) = False"
| "nullable0 (In m M) = False"
lemma fimage_Suc_fsubset0[simp]: "Suc |`| A |\<subseteq>| {|0|} \<longleftrightarrow> A = {||}"
by blast
lemma fsubset_singleton_iff: "A |\<subseteq>| {|x|} \<longleftrightarrow> A = {||} \<or> A = {|x|}"
by blast
definition "restrict ord P = (case ord of FO \<Rightarrow> \<exists>x. P = {|x|} | SO \<Rightarrow> True)"
definition "Restrict ord i = (case ord of FO \<Rightarrow> FBase (Fo i) | SO \<Rightarrow> FBool True)"
declare [[goals_limit = 50]]
global_interpretation WS1S_Alt: Formula SUC LESS assigns nvars Extend CONS SNOC Length
extend size_atom zero \<sigma> eval downshift upshift finsert cut len restrict Restrict
lformula0 FV0 find0 wf0 decr0 satisfies0 nullable0 lderiv0 rderiv0 undefined
defines norm = "Formula_Operations.norm find0 decr0"
and nFOr = "Formula_Operations.nFOr :: formula \<Rightarrow> _"
and nFAnd = "Formula_Operations.nFAnd :: formula \<Rightarrow> _"
and nFNot = "Formula_Operations.nFNot find0 decr0 :: formula \<Rightarrow> _"
and nFEx = "Formula_Operations.nFEx find0 decr0"
and nFAll = "Formula_Operations.nFAll find0 decr0"
and decr = "Formula_Operations.decr decr0 :: _ \<Rightarrow> _ \<Rightarrow> formula \<Rightarrow> _"
and find = "Formula_Operations.find find0 :: _ \<Rightarrow> _ \<Rightarrow> formula \<Rightarrow> _"
and FV = "Formula_Operations.FV FV0"
and RESTR = "Formula_Operations.RESTR Restrict :: _ \<Rightarrow> formula"
and RESTRICT = "Formula_Operations.RESTRICT Restrict FV0"
and deriv = "\<lambda>d0 (a :: atom) (\<phi> :: formula). Formula_Operations.deriv extend d0 a \<phi>"
and nullable = "\<lambda>\<phi> :: formula. Formula_Operations.nullable nullable0 \<phi>"
and fut_default = "Formula.fut_default extend zero rderiv0"
and fut = "Formula.fut extend zero find0 decr0 rderiv0"
and finalize = "Formula.finalize SUC extend zero find0 decr0 rderiv0"
and final = "Formula.final SUC extend zero find0 decr0
nullable0 rderiv0 :: idx \<Rightarrow> formula \<Rightarrow> _"
and ws1s_wf = "Formula_Operations.wf SUC (wf0 :: idx \<Rightarrow> ws1s \<Rightarrow> _)"
and ws1s_lformula = "Formula_Operations.lformula lformula0 :: formula \<Rightarrow> _"
and check_eqv = "\<lambda>idx. DAs.check_eqv
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
(final idx) (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>)
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
(final idx) (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>) (=)"
and bounded_check_eqv = "\<lambda>idx. DAs.check_eqv
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
nullable (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>)
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
nullable (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>) (=)"
and automaton = "DA.automaton
(\<lambda>a \<phi>. norm (deriv lderiv0 (a :: atom) \<phi> :: formula))"
proof
fix k idx and a :: ws1s and l assume "wf0 (SUC k idx) a" "LESS k l (SUC k idx)" "\<not> find0 k l a"
then show "wf0 idx (decr0 k l a)"
by (induct a) (unfold wf0.simps atomic.map find0.simps,
(transfer, force simp: dec_def split: if_splits order.splits)+) \<comment> \<open>slow\<close>
next
fix k and a :: ws1s and l assume "lformula0 a"
then show "lformula0 (decr0 k l a)" by (induct a) auto
next
fix i k and a :: ws1s and \<AA> :: interp and P assume *: "\<not> find0 k i a" "LESS k i (SUC k (#\<^sub>V \<AA>))"
and disj: "lformula0 a \<or> len P \<le> Length \<AA>"
from disj show "satisfies0 (Extend k i \<AA> P) a = satisfies0 \<AA> (decr0 k i a)"
proof
assume "lformula0 a"
then show ?thesis using *
by (induct a)
(auto simp: dec_def split: if_splits order.split option.splits bool.splits) \<comment> \<open>slow\<close>
next
assume "len P \<le> Length \<AA>"
with * show ?thesis
proof (induct a)
case Fo then show ?case by (cases k) (auto simp: dec_def)
next
case Z then show ?case by (cases k) (auto simp: dec_def)
next
case Less then show ?case by (cases k) (auto simp: dec_def)
next
case In then show ?case by (cases k) (auto simp: dec_def)
qed
qed
next
fix idx and a :: ws1s and x assume "lformula0 a" "wf0 idx a"
then show "Formula_Operations.wf SUC wf0 idx (lderiv0 x a)"
by (induct a rule: lderiv0.induct)
(auto simp: Formula_Operations.wf.simps Let_def split: bool.splits order.splits)
next
fix a :: ws1s and x assume "lformula0 a"
then show "Formula_Operations.lformula lformula0 (lderiv0 x a)"
by (induct a rule: lderiv0.induct)
(auto simp: Formula_Operations.lformula.simps split: bool.splits)
next
fix idx and a :: ws1s and x assume "wf0 idx a"
then show "Formula_Operations.wf SUC wf0 idx (rderiv0 x a)"
by (induct a rule: lderiv0.induct)
(auto simp: Formula_Operations.wf.simps Let_def sorted_append
split: bool.splits order.splits nat.splits)
next
fix \<AA> :: interp and a :: ws1s
- note fmember_iff_member_fset[symmetric, simp]
assume "Length \<AA> = 0"
then show "nullable0 a = satisfies0 \<AA> a"
by (induct a, unfold wf0.simps nullable0.simps satisfies0.simps Let_def)
(transfer, (auto 0 3 dest: MSB_greater split: prod.splits if_splits option.splits bool.splits nat.splits) [])+ \<comment> \<open>slow\<close>
next
note Formula_Operations.satisfies_gen.simps[simp] Let_def[simp] upshift_def[simp]
fix x :: atom and a :: ws1s and \<AA> :: interp
assume "lformula0 a" "wf0 (#\<^sub>V \<AA>) a" "#\<^sub>V \<AA> = size_atom x"
then show "Formula_Operations.satisfies Extend Length satisfies0 \<AA> (lderiv0 x a) =
satisfies0 (CONS x \<AA>) a"
proof (induct a)
qed (auto split: prod.splits bool.splits)
next
note Formula_Operations.satisfies_gen.simps[simp] Let_def[simp] upshift_def[simp]
fix x :: atom and a :: ws1s and \<AA> :: interp
assume "lformula0 a" "wf0 (#\<^sub>V \<AA>) a" "#\<^sub>V \<AA> = size_atom x"
then show "Formula_Operations.satisfies_bounded Extend Length len satisfies0 \<AA> (lderiv0 x a) =
satisfies0 (CONS x \<AA>) a"
by (induct a) (auto split: prod.splits bool.splits)
next
note Formula_Operations.satisfies_gen.simps[simp] Let_def[simp]
fix x :: atom and a :: ws1s and \<AA> :: interp
assume "wf0 (#\<^sub>V \<AA>) a" "#\<^sub>V \<AA> = size_atom x"
then show "Formula_Operations.satisfies_bounded Extend Length len satisfies0 \<AA> (rderiv0 x a) =
satisfies0 (SNOC x \<AA>) a"
proof (induct a)
case Less then show ?case
apply (auto 2 0 split: prod.splits option.splits bool.splits)
apply (auto simp add: fsubset_singleton_iff)
apply (metis assigns_less_Length finsertCI less_not_sym)
apply force
apply (metis assigns_less_Length finsertCI less_not_sym)
apply force
done
next
case In then show ?case by (force split: prod.splits)
qed (auto split: prod.splits)
next
fix a :: ws1s and \<AA> \<BB> :: interp
assume "wf0 (#\<^sub>V \<BB>) a" "#\<^sub>V \<AA> = #\<^sub>V \<BB>" "(\<And>m k. LESS k m (#\<^sub>V \<BB>) \<Longrightarrow> m\<^bsup>\<AA>\<^esup>k = m\<^bsup>\<BB>\<^esup>k)" "lformula0 a"
then show "satisfies0 \<AA> a \<longleftrightarrow> satisfies0 \<BB> a" by (induct a) auto
next
fix a :: ws1s
assume "lformula0 a"
moreover
define d where "d = Formula_Operations.deriv extend lderiv0"
define \<Phi> :: "_ \<Rightarrow> (ws1s, order) aformula set"
where "\<Phi> a =
(case a of
Fo m \<Rightarrow> {FBase (Fo m), FBase (Z m), FBool False}
| Z m \<Rightarrow> {FBase (Z m), FBool False}
| Less m1 m2 \<Rightarrow> {FBase (Less m1 m2),
FAnd (FBase (Z m1)) (FBase (Fo m2)),
FAnd (FBase (Z m1)) (FBase (Z m2)),
FAnd (FBase (Z m1)) (FBool False),
FAnd (FBool False) (FBase (Fo m2)),
FAnd (FBool False) (FBase (Z m2)),
FAnd (FBool False) (FBool False),
FBool False}
| In i I \<Rightarrow> {FBase (In i I), FBase (Z i), FBool False})" for a
{ fix xs
note Formula_Operations.fold_deriv_FBool[simp] Formula_Operations.deriv.simps[simp] \<Phi>_def[simp]
from \<open>lformula0 a\<close> have "FBase a \<in> \<Phi> a" by (cases a) auto
moreover have "\<And>x \<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> d x \<phi> \<in> \<Phi> a"
by (auto simp: d_def split: atomic.splits list.splits bool.splits if_splits option.splits)
then have "\<And>\<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> fold d xs \<phi> \<in> \<Phi> a" by (induct xs) auto
ultimately have "fold d xs (FBase a) \<in> \<Phi> a" by blast
}
moreover have "finite (\<Phi> a)" using \<open>lformula0 a\<close> unfolding \<Phi>_def by (auto split: atomic.splits)
ultimately show "finite {fold d xs (FBase a) | xs. True}" by (blast intro: finite_subset)
next
fix a :: ws1s
define d where "d = Formula_Operations.deriv extend rderiv0"
define \<Phi> :: "_ \<Rightarrow> (ws1s, order) aformula set"
where "\<Phi> a =
(case a of
Fo m \<Rightarrow> {FBase (Fo m), FBase (Z m), FBool False}
| Z m \<Rightarrow> {FBase (Z m), FBool False}
| Less m1 m2 \<Rightarrow> {FBase (Less m1 m2),
FAnd (FBase (Z m2)) (FBase (Fo m1)) ,
FAnd (FBase (Z m2)) (FBase (Z m1)),
FAnd (FBase (Z m2)) (FBool False),
FAnd (FBool False) (FBase (Fo m1)),
FAnd (FBool False) (FBase (Z m1)),
FAnd (FBool False) (FBool False),
FBool False}
| In i I \<Rightarrow> {FBase (In i I), FBase (Z i), FBool False})" for a
{ fix xs
note Formula_Operations.fold_deriv_FBool[simp] Formula_Operations.deriv.simps[simp] \<Phi>_def[simp]
then have "FBase a \<in> \<Phi> a" by (auto split: atomic.splits option.splits)
moreover have "\<And>x \<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> d x \<phi> \<in> \<Phi> a"
by (auto simp add: d_def Let_def not_le gr0_conv_Suc
split: atomic.splits list.splits bool.splits if_splits option.splits nat.splits)
then have "\<And>\<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> fold d xs \<phi> \<in> \<Phi> a"
by (induct xs) auto
ultimately have "fold d xs (FBase a) \<in> \<Phi> a" by blast
}
moreover have "finite (\<Phi> a)" unfolding \<Phi>_def using [[simproc add: finite_Collect]]
by (auto split: atomic.splits)
ultimately show "finite {fold d xs (FBase a) | xs. True}" by (blast intro: finite_subset)
next
fix k l and a :: ws1s
show "find0 k l a \<longleftrightarrow> l \<in> FV0 k a" by (induct a rule: find0.induct) auto
next
fix a :: ws1s and k :: order
show "finite (FV0 k a)" by (cases k) (induct a, auto)+
next
fix idx a k v
assume "wf0 idx a" "v \<in> FV0 k a"
then show "LESS k v idx" by (cases k) (induct a, auto)+
next
fix idx k i
assume "LESS k i idx"
then show "Formula_Operations.wf SUC wf0 idx (Restrict k i)"
unfolding Restrict_def by (cases k) (auto simp: Formula_Operations.wf.simps)
next
fix k and i :: nat
show "Formula_Operations.lformula lformula0 (Restrict k i)"
unfolding Restrict_def by (cases k) (auto simp: Formula_Operations.lformula.simps)
next
fix i \<AA> k P r
assume "i\<^bsup>\<AA>\<^esup>k = P"
then show "restrict k P \<longleftrightarrow>
Formula_Operations.satisfies_gen Extend Length satisfies0 r \<AA> (Restrict k i)"
unfolding restrict_def Restrict_def
by (cases k) (auto simp: Formula_Operations.satisfies_gen.simps)
qed (auto simp: Extend_commute_unsafe downshift_def upshift_def fimage_iff Suc_le_eq len_def
dec_def eval_def cut_def len_downshift_helper CONS_inj dest!: CONS_surj
dest: fMax_ge fMax_ffilter_less fMax_boundedD fsubset_fsingletonD
split: order.splits if_splits)
(*Workaround for code generation*)
lemma check_eqv_code[code]: "check_eqv idx r s =
((ws1s_wf idx r \<and> ws1s_lformula r) \<and> (ws1s_wf idx s \<and> ws1s_lformula s) \<and>
(case rtrancl_while (\<lambda>(p, q). final idx p = final idx q)
(\<lambda>(p, q). map (\<lambda>a. (norm (deriv lderiv0 a p), norm (deriv lderiv0 a q))) (\<sigma> idx))
(norm (RESTRICT r), norm (RESTRICT s)) of
None \<Rightarrow> False
| Some ([], x) \<Rightarrow> True
| Some (a # list, x) \<Rightarrow> False))"
unfolding check_eqv_def WS1S_Alt.check_eqv_def WS1S_Alt.step_alt ..
definition while where [code del, code_abbrev]: "while idx \<phi> = while_default (fut_default idx \<phi>)"
declare while_default_code[of "fut_default idx \<phi>" for idx \<phi>, folded while_def, code]
lemma check_eqv_sound:
"\<lbrakk>#\<^sub>V \<AA> = idx; check_eqv idx \<phi> \<psi>\<rbrakk> \<Longrightarrow> (WS1S_Alt.sat \<AA> \<phi> \<longleftrightarrow> WS1S_Alt.sat \<AA> \<psi>)"
unfolding check_eqv_def by (rule WS1S_Alt.check_eqv_soundness)
lemma bounded_check_eqv_sound:
"\<lbrakk>#\<^sub>V \<AA> = idx; bounded_check_eqv idx \<phi> \<psi>\<rbrakk> \<Longrightarrow> (WS1S_Alt.sat\<^sub>b \<AA> \<phi> \<longleftrightarrow> WS1S_Alt.sat\<^sub>b \<AA> \<psi>)"
unfolding bounded_check_eqv_def by (rule WS1S_Alt.bounded_check_eqv_soundness)
method_setup check_equiv = \<open>
let
fun tac ctxt =
let
val conv = @{computation_check terms: Trueprop
"0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat" Suc
"plus :: nat \<Rightarrow> _" "minus :: nat \<Rightarrow> _"
"times :: nat \<Rightarrow> _" "divide :: nat \<Rightarrow> _" "modulo :: nat \<Rightarrow> _"
"0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int"
check_eqv datatypes: formula "int list" integer idx
"nat \<times> nat" "nat option" "bool option"} ctxt
in
CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 conv)) ctxt) THEN'
resolve_tac ctxt [TrueI]
end
in
Scan.succeed (SIMPLE_METHOD' o tac)
end
\<close>
end
diff --git a/thys/Formula_Derivatives/WS1S_Formula.thy b/thys/Formula_Derivatives/WS1S_Formula.thy
--- a/thys/Formula_Derivatives/WS1S_Formula.thy
+++ b/thys/Formula_Derivatives/WS1S_Formula.thy
@@ -1,798 +1,792 @@
section \<open>Concrete Atomic WS1S Formulas (Minimum Semantics for FO Variables)\<close>
(*<*)
theory WS1S_Formula
imports
Abstract_Formula
WS1S_Prelim
begin
(*>*)
datatype (FOV0: 'fo, SOV0: 'so) atomic =
Fo 'fo |
Eq_Const "nat option" 'fo nat |
Less "bool option" 'fo 'fo |
Plus_FO "nat option" 'fo 'fo nat |
Eq_FO bool 'fo 'fo |
Eq_SO 'so 'so |
Suc_SO bool bool 'so 'so |
Empty 'so |
Singleton 'so |
Subset 'so 'so |
In bool 'fo 'so |
Eq_Max bool 'fo 'so |
Eq_Min bool 'fo 'so |
Eq_Union 'so 'so 'so |
Eq_Inter 'so 'so 'so |
Eq_Diff 'so 'so 'so |
Disjoint 'so 'so |
Eq_Presb "nat option" 'so nat
derive linorder option
derive linorder atomic \<comment> \<open>very slow\<close>
type_synonym fo = nat
type_synonym so = nat
type_synonym ws1s = "(fo, so) atomic"
type_synonym formula = "(ws1s, order) aformula"
primrec wf0 where
"wf0 idx (Fo m) = LESS FO m idx"
| "wf0 idx (Eq_Const i m n) = (LESS FO m idx \<and> (case i of Some i \<Rightarrow> i \<le> n | _ \<Rightarrow> True))"
| "wf0 idx (Less _ m1 m2) = (LESS FO m1 idx \<and> LESS FO m2 idx)"
| "wf0 idx (Plus_FO i m1 m2 n) =
(LESS FO m1 idx \<and> LESS FO m2 idx \<and> (case i of Some i \<Rightarrow> i \<le> n | _ \<Rightarrow> True))"
| "wf0 idx (Eq_FO _ m1 m2) = (LESS FO m1 idx \<and> LESS FO m2 idx)"
| "wf0 idx (Eq_SO M1 M2) = (LESS SO M1 idx \<and> LESS SO M2 idx)"
| "wf0 idx (Suc_SO br bl M1 M2) = (LESS SO M1 idx \<and> LESS SO M2 idx)"
| "wf0 idx (Empty M) = LESS SO M idx"
| "wf0 idx (Singleton M) = LESS SO M idx"
| "wf0 idx (Subset M1 M2) = (LESS SO M1 idx \<and> LESS SO M2 idx)"
| "wf0 idx (In _ m M) = (LESS FO m idx \<and> LESS SO M idx)"
| "wf0 idx (Eq_Max _ m M) = (LESS FO m idx \<and> LESS SO M idx)"
| "wf0 idx (Eq_Min _ m M) = (LESS FO m idx \<and> LESS SO M idx)"
| "wf0 idx (Eq_Union M1 M2 M3) = (LESS SO M1 idx \<and> LESS SO M2 idx \<and> LESS SO M3 idx)"
| "wf0 idx (Eq_Inter M1 M2 M3) = (LESS SO M1 idx \<and> LESS SO M2 idx \<and> LESS SO M3 idx)"
| "wf0 idx (Eq_Diff M1 M2 M3) = (LESS SO M1 idx \<and> LESS SO M2 idx \<and> LESS SO M3 idx)"
| "wf0 idx (Disjoint M1 M2) = (LESS SO M1 idx \<and> LESS SO M2 idx)"
| "wf0 idx (Eq_Presb _ M n) = LESS SO M idx"
inductive lformula0 where
"lformula0 (Fo m)"
| "lformula0 (Eq_Const None m n)"
| "lformula0 (Less None m1 m2)"
| "lformula0 (Plus_FO None m1 m2 n)"
| "lformula0 (Eq_FO False m1 m2)"
| "lformula0 (Eq_SO M1 M2)"
| "lformula0 (Suc_SO False bl M1 M2)"
| "lformula0 (Empty M)"
| "lformula0 (Singleton M)"
| "lformula0 (Subset M1 M2)"
| "lformula0 (In False m M)"
| "lformula0 (Eq_Max False m M)"
| "lformula0 (Eq_Min False m M)"
| "lformula0 (Eq_Union M1 M2 M3)"
| "lformula0 (Eq_Inter M1 M2 M3)"
| "lformula0 (Eq_Diff M1 M2 M3)"
| "lformula0 (Disjoint M1 M2)"
| "lformula0 (Eq_Presb None M n)"
code_pred lformula0 .
declare lformula0.intros[simp]
inductive_cases lformula0E[elim]: "lformula0 a"
abbreviation "FV0 \<equiv> case_order FOV0 SOV0"
fun find0 where
"find0 FO i (Fo m) = (i = m)"
| "find0 FO i (Eq_Const _ m _) = (i = m)"
| "find0 FO i (Less _ m1 m2) = (i = m1 \<or> i = m2)"
| "find0 FO i (Plus_FO _ m1 m2 _) = (i = m1 \<or> i = m2)"
| "find0 FO i (Eq_FO _ m1 m2) = (i = m1 \<or> i = m2)"
| "find0 SO i (Eq_SO M1 M2) = (i = M1 \<or> i = M2)"
| "find0 SO i (Suc_SO _ _ M1 M2) = (i = M1 \<or> i = M2)"
| "find0 SO i (Empty M) = (i = M)"
| "find0 SO i (Singleton M) = (i = M)"
| "find0 SO i (Subset M1 M2) = (i = M1 \<or> i = M2)"
| "find0 FO i (In _ m _) = (i = m)"
| "find0 SO i (In _ _ M) = (i = M)"
| "find0 FO i (Eq_Max _ m _) = (i = m)"
| "find0 SO i (Eq_Max _ _ M) = (i = M)"
| "find0 FO i (Eq_Min _ m _) = (i = m)"
| "find0 SO i (Eq_Min _ _ M) = (i = M)"
| "find0 SO i (Eq_Union M1 M2 M3) = (i = M1 \<or> i = M2 \<or> i = M3)"
| "find0 SO i (Eq_Inter M1 M2 M3) = (i = M1 \<or> i = M2 \<or> i = M3)"
| "find0 SO i (Eq_Diff M1 M2 M3) = (i = M1 \<or> i = M2 \<or> i = M3)"
| "find0 SO i (Disjoint M1 M2) = (i = M1 \<or> i = M2)"
| "find0 SO i (Eq_Presb _ M _) = (i = M)"
| "find0 _ _ _ = False"
abbreviation "decr0 ord k \<equiv> map_atomic (case_order (dec k) id ord) (case_order id (dec k) ord)"
lemma sum_pow2_image_Suc:
"finite X \<Longrightarrow> sum ((^) (2 :: nat)) (Suc ` X) = 2 * sum ((^) 2) X"
by (induct X rule: finite_induct) (auto intro: trans[OF sum.insert])
lemma sum_pow2_insert0:
"\<lbrakk>finite X; 0 \<notin> X\<rbrakk> \<Longrightarrow> sum ((^) (2 :: nat)) (insert 0 X) = Suc (sum ((^) 2) X)"
by (induct X rule: finite_induct) (auto intro: trans[OF sum.insert])
lemma sum_pow2_upto: "sum ((^) (2 :: nat)) {0 ..< x} = 2 ^ x - 1"
by (induct x) (auto simp: algebra_simps)
lemma sum_pow2_inj:
"\<lbrakk>finite X; finite Y; (\<Sum>x\<in>X. 2 ^ x :: nat) = (\<Sum>x\<in>Y. 2 ^ x)\<rbrakk> \<Longrightarrow> X = Y"
(is "_ \<Longrightarrow> _ \<Longrightarrow> ?f X = ?f Y \<Longrightarrow> _")
proof (induct X arbitrary: Y rule: finite_linorder_max_induct)
case (insert x X)
from insert(2) have "?f X \<le> ?f {0 ..< x}" by (intro sum_mono2) auto
also have "\<dots> < 2 ^ x" by (induct x) simp_all
finally have "?f X < 2 ^ x" .
moreover from insert(1,2) have *: "?f X + 2 ^ x = ?f Y"
using trans[OF sym[OF insert(5)] sum.insert] by auto
ultimately have "?f Y < 2 ^ Suc x" by simp
have "\<forall>y \<in> Y. y \<le> x"
proof (rule ccontr)
assume "\<not> (\<forall>y\<in>Y. y \<le> x)"
then obtain y where "y \<in> Y" "Suc x \<le> y" by auto
from this(2) have "2 ^ Suc x \<le> (2 ^ y :: nat)" by (intro power_increasing) auto
also from \<open>y \<in> Y\<close> insert(4) have "\<dots> \<le> ?f Y" by (metis order.refl sum.remove trans_le_add1)
finally show False using \<open>?f Y < 2 ^ Suc x\<close> by simp
qed
{ assume "x \<notin> Y"
with \<open>\<forall>y \<in> Y. y \<le> x\<close> have "?f Y \<le> ?f {0 ..< x}" by (intro sum_mono2) (auto simp: le_less)
also have "\<dots> < 2 ^ x" by (induct x) simp_all
finally have "?f Y < 2 ^ x" .
with * have False by auto
}
then have "x \<in> Y" by blast
from insert(4) have "?f (Y - {x}) + 2 ^ x = ?f (insert x (Y - {x}))"by (subst sum.insert) auto
also have "\<dots> = ?f X + 2 ^ x" unfolding * using \<open>x \<in> Y\<close> by (simp add: insert_absorb)
finally have "?f X = ?f (Y - {x})" by simp
with insert(3,4) have "X = Y - {x}" by simp
with \<open>x \<in> Y\<close> show ?case by auto
qed simp
lemma finite_pow2_eq:
fixes n :: nat
shows "finite {i. 2 ^ i = n}"
proof -
have "((^) 2) ` {i. 2 ^ i = n} \<subseteq> {n}" by auto
then have "finite (((^) (2 :: nat)) ` {i. 2 ^ i = n})" by (rule finite_subset) blast
then show "finite {i. 2 ^ i = n}" by (rule finite_imageD) (auto simp: inj_on_def)
qed
lemma finite_pow2_le[simp]:
fixes n :: nat
shows "finite {i. 2 ^ i \<le> n}"
by (induct n) (auto simp: le_Suc_eq finite_pow2_eq)
lemma le_pow2[simp]: "x \<le> y \<Longrightarrow> x \<le> 2 ^ y"
by (induct x arbitrary: y) (force simp add: Suc_le_eq order.strict_iff_order)+
lemma ld_bounded: "Max {i. 2 ^ i \<le> Suc n} \<le> Suc n" (is "?m \<le> Suc n")
proof -
have "?m \<le> 2 ^ ?m" by (rule le_pow2) simp
moreover
have "?m \<in> {i. 2 ^ i \<le> Suc n}" by (rule Max_in) (auto intro: exI[of _ 0])
then have "2 ^ ?m \<le> Suc n" by simp
ultimately show ?thesis by linarith
qed
primrec satisfies0 where
"satisfies0 \<AA> (Fo m) = (m\<^bsup>\<AA>\<^esup>FO \<noteq> {||})"
| "satisfies0 \<AA> (Eq_Const i m n) =
(let P = m\<^bsup>\<AA>\<^esup>FO in if P = {||}
then (case i of Some i \<Rightarrow> Length \<AA> = i | _ \<Rightarrow> False)
else fMin P = n)"
| "satisfies0 \<AA> (Less b m1 m2) =
(let P1 = m1\<^bsup>\<AA>\<^esup>FO; P2 = m2\<^bsup>\<AA>\<^esup>FO in if P1 = {||} \<or> P2 = {||}
then (case b of None \<Rightarrow> False | Some True \<Rightarrow> P2 = {||} | Some False \<Rightarrow> P1 \<noteq> {||})
else fMin P1 < fMin P2)"
| "satisfies0 \<AA> (Plus_FO i m1 m2 n) =
(let P1 = m1\<^bsup>\<AA>\<^esup>FO; P2 = m2\<^bsup>\<AA>\<^esup>FO in if P1 = {||} \<or> P2 = {||}
then (case i of Some 0 \<Rightarrow> P1 = P2 | Some i \<Rightarrow> P2 \<noteq> {||} \<and> fMin P2 + i = Length \<AA> | _ \<Rightarrow> False)
else fMin P1 = fMin P2 + n)"
| "satisfies0 \<AA> (Eq_FO b m1 m2) =
(let P1 = m1\<^bsup>\<AA>\<^esup>FO; P2 = m2\<^bsup>\<AA>\<^esup>FO in if P1 = {||} \<or> P2 = {||}
then b \<and> P1 = P2
else fMin P1 = fMin P2)"
| "satisfies0 \<AA> (Eq_SO M1 M2) = (M1\<^bsup>\<AA>\<^esup>SO = M2\<^bsup>\<AA>\<^esup>SO)"
| "satisfies0 \<AA> (Suc_SO br bl M1 M2) =
((if br then finsert (Length \<AA>) else id) (M1\<^bsup>\<AA>\<^esup>SO) =
(if bl then finsert 0 else id) (Suc |`| M2\<^bsup>\<AA>\<^esup>SO))"
| "satisfies0 \<AA> (Empty M) = (M\<^bsup>\<AA>\<^esup>SO = {||})"
| "satisfies0 \<AA> (Singleton M) = (\<exists>x. M\<^bsup>\<AA>\<^esup>SO = {|x|})"
| "satisfies0 \<AA> (Subset M1 M2) = (M1\<^bsup>\<AA>\<^esup>SO |\<subseteq>| M2\<^bsup>\<AA>\<^esup>SO)"
| "satisfies0 \<AA> (In b m M) =
(let P = m\<^bsup>\<AA>\<^esup>FO in if P = {||} then b else fMin P |\<in>| M\<^bsup>\<AA>\<^esup>SO)"
| "satisfies0 \<AA> (Eq_Max b m M) =
(let P1 = m\<^bsup>\<AA>\<^esup>FO; P2 = M\<^bsup>\<AA>\<^esup>SO in if b then P1 = {||}
else if P1 = {||} \<or> P2 = {||} then False else fMin P1 = fMax P2)"
| "satisfies0 \<AA> (Eq_Min b m M) =
(let P1 = m\<^bsup>\<AA>\<^esup>FO; P2 = M\<^bsup>\<AA>\<^esup>SO in if P1 = {||} \<or> P2 = {||} then b \<and> P1 = P2 else fMin P1 = fMin P2)"
| "satisfies0 \<AA> (Eq_Union M1 M2 M3) = (M1\<^bsup>\<AA>\<^esup>SO = M2\<^bsup>\<AA>\<^esup>SO |\<union>| M3\<^bsup>\<AA>\<^esup>SO)"
| "satisfies0 \<AA> (Eq_Inter M1 M2 M3) = (M1\<^bsup>\<AA>\<^esup>SO = M2\<^bsup>\<AA>\<^esup>SO |\<inter>| M3\<^bsup>\<AA>\<^esup>SO)"
| "satisfies0 \<AA> (Eq_Diff M1 M2 M3) = (M1\<^bsup>\<AA>\<^esup>SO = M2\<^bsup>\<AA>\<^esup>SO |-| M3\<^bsup>\<AA>\<^esup>SO)"
| "satisfies0 \<AA> (Disjoint M1 M2) = (M1\<^bsup>\<AA>\<^esup>SO |\<inter>| M2\<^bsup>\<AA>\<^esup>SO = {||})"
| "satisfies0 \<AA> (Eq_Presb i M n) = (((\<Sum>x\<in>fset (M\<^bsup>\<AA>\<^esup>SO). 2 ^ x) = n) \<and>
(case i of None \<Rightarrow> True | Some l \<Rightarrow> Length \<AA> = l))"
fun lderiv0 where
"lderiv0 (bs1, bs2) (Fo m) = (if bs1 ! m then FBool True else FBase (Fo m))"
| "lderiv0 (bs1, bs2) (Eq_Const None m n) = (if n = 0 \<and> bs1 ! m then FBool True
else if n = 0 \<or> bs1 ! m then FBool False else FBase (Eq_Const None m (n - 1)))"
| "lderiv0 (bs1, bs2) (Less None m1 m2) = (case (bs1 ! m1, bs1 ! m2) of
(False, False) \<Rightarrow> FBase (Less None m1 m2)
| (True, False) \<Rightarrow> FBase (Fo m2)
| _ \<Rightarrow> FBool False)"|
"lderiv0 (bs1, bs2) (Eq_FO False m1 m2) = (case (bs1 ! m1, bs1 ! m2) of
(False, False) \<Rightarrow> FBase (Eq_FO False m1 m2)
| (True, True) \<Rightarrow> FBool True
| _ \<Rightarrow> FBool False)"
| "lderiv0 (bs1, bs2) (Plus_FO None m1 m2 n) = (if n = 0
then
(case (bs1 ! m1, bs1 ! m2) of
(False, False) \<Rightarrow> FBase (Plus_FO None m1 m2 n)
| (True, True) \<Rightarrow> FBool True
| _ \<Rightarrow> FBool False)
else
(case (bs1 ! m1, bs1 ! m2) of
(False, False) \<Rightarrow> FBase (Plus_FO None m1 m2 n)
| (False, True) \<Rightarrow> FBase (Eq_Const None m1 (n - 1))
| _ \<Rightarrow> FBool False))"
| "lderiv0 (bs1, bs2) (Eq_SO M1 M2) =
(if bs2 ! M1 = bs2 ! M2 then FBase (Eq_SO M1 M2) else FBool False)"
| "lderiv0 (bs1, bs2) (Suc_SO False bl M1 M2) = (if bl = bs2 ! M1
then FBase (Suc_SO False (bs2 ! M2) M1 M2) else FBool False)"
| "lderiv0 (bs1, bs2) (Empty M) = (case bs2 ! M of
True \<Rightarrow> FBool False
| False \<Rightarrow> FBase (Empty M))"
| "lderiv0 (bs1, bs2) (Singleton M) = (case bs2 ! M of
True \<Rightarrow> FBase (Empty M)
| False \<Rightarrow> FBase (Singleton M))"
| "lderiv0 (bs1, bs2) (Subset M1 M2) = (case (bs2 ! M1, bs2 ! M2) of
(True, False) \<Rightarrow> FBool False
| _ \<Rightarrow> FBase (Subset M1 M2))"
| "lderiv0 (bs1, bs2) (In False m M) = (case (bs1 ! m, bs2 ! M) of
(False, _) \<Rightarrow> FBase (In False m M)
| (True, True) \<Rightarrow> FBool True
| _ \<Rightarrow> FBool False)"
| "lderiv0 (bs1, bs2) (Eq_Max False m M) = (case (bs1 ! m, bs2 ! M) of
(False, _) \<Rightarrow> FBase (Eq_Max False m M)
| (True, True) \<Rightarrow> FBase (Empty M)
| _ \<Rightarrow> FBool False)"
| "lderiv0 (bs1, bs2) (Eq_Min False m M) = (case (bs1 ! m, bs2 ! M) of
(False, False) \<Rightarrow> FBase (Eq_Min False m M)
| (True, True) \<Rightarrow> FBool True
| _ \<Rightarrow> FBool False)"
| "lderiv0 (bs1, bs2) (Eq_Union M1 M2 M3) = (if bs2 ! M1 = (bs2 ! M2 \<or> bs2 ! M3)
then FBase (Eq_Union M1 M2 M3) else FBool False)"
| "lderiv0 (bs1, bs2) (Eq_Inter M1 M2 M3) = (if bs2 ! M1 = (bs2 ! M2 \<and> bs2 ! M3)
then FBase (Eq_Inter M1 M2 M3) else FBool False)"
| "lderiv0 (bs1, bs2) (Eq_Diff M1 M2 M3) = (if bs2 ! M1 = (bs2 ! M2 \<and> \<not> bs2 ! M3)
then FBase (Eq_Diff M1 M2 M3) else FBool False)"
| "lderiv0 (bs1, bs2) (Disjoint M1 M2) =
(if bs2 ! M1 \<and> bs2 ! M2 then FBool False else FBase (Disjoint M1 M2))"
| "lderiv0 (bs1, bs2) (Eq_Presb None M n) = (if bs2 ! M = (n mod 2 = 0)
then FBool False else FBase (Eq_Presb None M (n div 2)))"
| "lderiv0 _ _ = undefined"
fun ld where
"ld 0 = 0"
| "ld (Suc 0) = 0"
| "ld n = Suc (ld (n div 2))"
lemma ld_alt[simp]: "n > 0 \<Longrightarrow> ld n = Max {i. 2 ^ i \<le> n}"
proof (safe intro!: Max_eqI[symmetric])
assume "n > 0" then show "2 ^ ld n \<le> n" by (induct n rule: ld.induct) auto
next
fix y
assume "2 ^ y \<le> n"
then show "y \<le> ld n"
proof (induct n arbitrary: y rule: ld.induct)
case (3 z)
then have "y - 1 \<le> ld (Suc (Suc z) div 2)"
by (cases y) simp_all
then show ?case by simp
qed (auto simp: le_eq_less_or_eq)
qed simp
fun rderiv0 where
"rderiv0 (bs1, bs2) (Fo m) = (if bs1 ! m then FBool True else FBase (Fo m))"
| "rderiv0 (bs1, bs2) (Eq_Const i m n) = (case bs1 ! m of
False \<Rightarrow> FBase (Eq_Const (case i of Some (Suc i) \<Rightarrow> Some i | _ \<Rightarrow> None) m n)
| True \<Rightarrow> FBase (Eq_Const (Some n) m n))"
| "rderiv0 (bs1, bs2) (Less b m1 m2) = (case bs1 ! m2 of
False \<Rightarrow> (case b of
Some False \<Rightarrow> (case bs1 ! m1 of
True \<Rightarrow> FBase (Less (Some True) m1 m2)
| False \<Rightarrow> FBase (Less (Some False) m1 m2))
| _ \<Rightarrow> FBase (Less b m1 m2))
| True \<Rightarrow> FBase (Less (Some False) m1 m2))"
| "rderiv0 (bs1, bs2) (Plus_FO i m1 m2 n) = (if n = 0
then
(case (bs1 ! m1, bs1 ! m2) of
(False, False) \<Rightarrow> FBase (Plus_FO i m1 m2 n)
| (True, True) \<Rightarrow> FBase (Plus_FO (Some 0) m1 m2 n)
| _ \<Rightarrow> FBase (Plus_FO None m1 m2 n))
else
(case bs1 ! m1 of
True \<Rightarrow> FBase (Plus_FO (Some n) m1 m2 n)
| False \<Rightarrow> (case bs1 ! m2 of
False \<Rightarrow> (case i of
Some (Suc (Suc i)) \<Rightarrow> FBase (Plus_FO (Some (Suc i)) m1 m2 n)
| Some (Suc 0) \<Rightarrow> FBase (Plus_FO None m1 m2 n)
| _ \<Rightarrow> FBase (Plus_FO i m1 m2 n))
| True \<Rightarrow> (case i of
Some (Suc i) \<Rightarrow> FBase (Plus_FO (Some i) m1 m2 n)
| _ \<Rightarrow> FBase (Plus_FO None m1 m2 n)))))"
| "rderiv0 (bs1, bs2) (Eq_FO b m1 m2) = (case (bs1 ! m1, bs1 ! m2) of
(False, False) \<Rightarrow> FBase (Eq_FO b m1 m2)
| (True, True) \<Rightarrow> FBase (Eq_FO True m1 m2)
| _ \<Rightarrow> FBase (Eq_FO False m1 m2))"
| "rderiv0 (bs1, bs2) (Eq_SO M1 M2) =
(if bs2 ! M1 = bs2 ! M2 then FBase (Eq_SO M1 M2) else FBool False)"
| "rderiv0 (bs1, bs2) (Suc_SO br bl M1 M2) = (if br = bs2 ! M2
then FBase (Suc_SO (bs2 ! M1) bl M1 M2) else FBool False)"
| "rderiv0 (bs1, bs2) (Empty M) = (case bs2 ! M of
True \<Rightarrow> FBool False
| False \<Rightarrow> FBase (Empty M))"
| "rderiv0 (bs1, bs2) (Singleton M) = (case bs2 ! M of
True \<Rightarrow> FBase (Empty M)
| False \<Rightarrow> FBase (Singleton M))"
| "rderiv0 (bs1, bs2) (Subset M1 M2) = (case (bs2 ! M1, bs2 ! M2) of
(True, False) \<Rightarrow> FBool False
| _ \<Rightarrow> FBase (Subset M1 M2))"
| "rderiv0 (bs1, bs2) (In b m M) = (case (bs1 ! m, bs2 ! M) of
(True, True) \<Rightarrow> FBase (In True m M)
| (True, False) \<Rightarrow> FBase (In False m M)
| _ \<Rightarrow> FBase (In b m M))"
| "rderiv0 (bs1, bs2) (Eq_Max b m M) = (case (bs1 ! m, bs2 ! M) of
(True, True) \<Rightarrow> if b then FBool False else FBase (Eq_Max True m M)
| (True, False) \<Rightarrow> if b then FBool False else FBase (Eq_Max False m M)
| (False, True) \<Rightarrow> if b then FBase (Eq_Max True m M) else FBool False
| (False, False) \<Rightarrow> FBase (Eq_Max b m M))"
| "rderiv0 (bs1, bs2) (Eq_Min b m M) = (case (bs1 ! m, bs2 ! M) of
(True, True) \<Rightarrow> FBase (Eq_Min True m M)
| (False, False) \<Rightarrow> FBase (Eq_Min b m M)
| _ \<Rightarrow> FBase (Eq_Min False m M))"
| "rderiv0 (bs1, bs2) (Eq_Union M1 M2 M3) = (if bs2 ! M1 = (bs2 ! M2 \<or> bs2 ! M3)
then FBase (Eq_Union M1 M2 M3) else FBool False)"
| "rderiv0 (bs1, bs2) (Eq_Inter M1 M2 M3) = (if bs2 ! M1 = (bs2 ! M2 \<and> bs2 ! M3)
then FBase (Eq_Inter M1 M2 M3) else FBool False)"
| "rderiv0 (bs1, bs2) (Eq_Diff M1 M2 M3) = (if bs2 ! M1 = (bs2 ! M2 \<and> \<not> bs2 ! M3)
then FBase (Eq_Diff M1 M2 M3) else FBool False)"
| "rderiv0 (bs1, bs2) (Disjoint M1 M2) =
(if bs2 ! M1 \<and> bs2 ! M2 then FBool False else FBase (Disjoint M1 M2))"
| "rderiv0 (bs1, bs2) (Eq_Presb l M n) = (case l of
None \<Rightarrow> if bs2 ! M then
if n = 0 then FBool False
else let l = ld n in FBase (Eq_Presb (Some l) M (n - 2 ^ l))
else FBase (Eq_Presb l M n)
| Some 0 \<Rightarrow> FBool False
| Some (Suc l) \<Rightarrow> if bs2 ! M \<and> n \<ge> 2 ^ l then FBase (Eq_Presb (Some l) M (n - 2 ^ l))
else if \<not> bs2 ! M \<and> n < 2 ^ l then FBase (Eq_Presb (Some l) M n)
else FBool False)"
primrec nullable0 where
"nullable0 (Fo m) = False"
| "nullable0 (Eq_Const i m n) = (i = Some 0)"
| "nullable0 (Less b m1 m2) = (case b of None \<Rightarrow> False | Some b \<Rightarrow> b)"
| "nullable0 (Plus_FO i m1 m2 n) = (i = Some 0)"
| "nullable0 (Eq_FO b m1 m2) = b"
| "nullable0 (Eq_SO M1 M2) = True"
| "nullable0 (Suc_SO br bl M1 M2) = (bl = br)"
| "nullable0 (Empty M) = True"
| "nullable0 (Singleton M) = False"
| "nullable0 (Subset M1 M2) = True"
| "nullable0 (In b m M) = b"
| "nullable0 (Eq_Max b m M) = b"
| "nullable0 (Eq_Min b m M) = b"
| "nullable0 (Eq_Union M1 M2 M3) = True"
| "nullable0 (Eq_Inter M1 M2 M3) = True"
| "nullable0 (Eq_Diff M1 M2 M3) = True"
| "nullable0 (Disjoint M1 M2) = True"
| "nullable0 (Eq_Presb l M n) = (n = 0 \<and> (l = Some 0 \<or> l = None))"
definition "restrict ord P = (case ord of FO \<Rightarrow> P \<noteq> {||} | SO \<Rightarrow> True)"
definition "Restrict ord i = (case ord of FO \<Rightarrow> FBase (Fo i) | SO \<Rightarrow> FBool True)"
declare [[goals_limit = 50]]
global_interpretation WS1S: Formula SUC LESS assigns nvars Extend CONS SNOC Length
extend size_atom zero \<sigma> eval downshift upshift finsert cut len restrict Restrict
lformula0 FV0 find0 wf0 decr0 satisfies0 nullable0 lderiv0 rderiv0 undefined
defines norm = "Formula_Operations.norm find0 decr0"
and nFOr = "Formula_Operations.nFOr :: formula \<Rightarrow> _"
and nFAnd = "Formula_Operations.nFAnd :: formula \<Rightarrow> _"
and nFNot = "Formula_Operations.nFNot find0 decr0 :: formula \<Rightarrow> _"
and nFEx = "Formula_Operations.nFEx find0 decr0"
and nFAll = "Formula_Operations.nFAll find0 decr0"
and decr = "Formula_Operations.decr decr0 :: _ \<Rightarrow> _ \<Rightarrow> formula \<Rightarrow> _"
and find = "Formula_Operations.find find0 :: _ \<Rightarrow> _ \<Rightarrow> formula \<Rightarrow> _"
and FV = "Formula_Operations.FV FV0"
and RESTR = "Formula_Operations.RESTR Restrict :: _ \<Rightarrow> formula"
and RESTRICT = "Formula_Operations.RESTRICT Restrict FV0"
and deriv = "\<lambda>d0 (a :: atom) (\<phi> :: formula). Formula_Operations.deriv extend d0 a \<phi>"
and nullable = "\<lambda>\<phi> :: formula. Formula_Operations.nullable nullable0 \<phi>"
and fut_default = "Formula.fut_default extend zero rderiv0"
and fut = "Formula.fut extend zero find0 decr0 rderiv0"
and finalize = "Formula.finalize SUC extend zero find0 decr0 rderiv0"
and final = "Formula.final SUC extend zero find0 decr0
nullable0 rderiv0 :: idx \<Rightarrow> formula \<Rightarrow> _"
and ws1s_wf = "Formula_Operations.wf SUC (wf0 :: idx \<Rightarrow> ws1s \<Rightarrow> _)"
and ws1s_lformula = "Formula_Operations.lformula lformula0 :: formula \<Rightarrow> _"
and check_eqv = "\<lambda>idx. DAs.check_eqv
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
(final idx) (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>)
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
(final idx) (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>) (=)"
and bounded_check_eqv = "\<lambda>idx. DAs.check_eqv
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
nullable (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>)
(\<sigma> idx) (\<lambda>\<phi>. norm (RESTRICT \<phi>) :: (ws1s, order) aformula)
(\<lambda>a \<phi>. norm (deriv (lderiv0 :: _ \<Rightarrow> _ \<Rightarrow> formula) (a :: atom) \<phi>))
nullable (\<lambda>\<phi> :: formula. ws1s_wf idx \<phi> \<and> ws1s_lformula \<phi>) (=)"
and automaton = "DA.automaton
(\<lambda>a \<phi>. norm (deriv lderiv0 (a :: atom) \<phi> :: formula))"
proof
fix k idx and a :: ws1s and l assume "wf0 (SUC k idx) a" "LESS k l (SUC k idx)" "\<not> find0 k l a"
then show "wf0 idx (decr0 k l a)"
by (induct a) (unfold wf0.simps atomic.map find0.simps,
(transfer, force simp: dec_def split!: if_splits order.splits)+)
next
fix k and a :: ws1s and l assume "lformula0 a"
then show "lformula0 (decr0 k l a)" by (induct a) auto
next
fix i k and a :: ws1s and \<AA> :: interp and P assume *: "\<not> find0 k i a" "LESS k i (SUC k (#\<^sub>V \<AA>))"
and disj: "lformula0 a \<or> len P \<le> Length \<AA>"
from disj show "satisfies0 (Extend k i \<AA> P) a = satisfies0 \<AA> (decr0 k i a)"
proof
assume "lformula0 a"
then show ?thesis using *
by (induct a rule: lformula0.induct)
(auto simp: dec_def split: if_splits order.split option.splits bool.splits) \<comment> \<open>slow\<close>
next
note dec_def[simp]
assume "len P \<le> Length \<AA>"
with * show ?thesis
proof (induct a)
case Fo then show ?case by (cases k) auto
next
case Eq_Const then show ?case
by (cases k) (auto simp: Let_def len_def split: if_splits option.splits)
next
case Less then show ?case by (cases k) auto
next
case Plus_FO then show ?case
by (cases k) (auto simp: max_def len_def Let_def split: option.splits nat.splits)
next
case Eq_FO then show ?case by (cases k) auto
next
case Eq_SO then show ?case by (cases k) auto
next
case (Suc_SO br bl M1 M2) then show ?case
by (cases k) (auto simp: max_def len_def)
next
case Empty then show ?case by (cases k) auto
next
case Singleton then show ?case by (cases k) auto
next
case Subset then show ?case by (cases k) auto
next
case In then show ?case by (cases k) auto
qed (auto simp: len_def max_def split!: option.splits order.splits)
qed
next
fix idx and a :: ws1s and x assume "lformula0 a" "wf0 idx a"
then show "Formula_Operations.wf SUC wf0 idx (lderiv0 x a)"
by (induct a rule: lderiv0.induct)
(auto simp: Formula_Operations.wf.simps Let_def split: bool.splits order.splits)
next
fix a :: ws1s and x assume "lformula0 a"
then show "Formula_Operations.lformula lformula0 (lderiv0 x a)"
by (induct a rule: lderiv0.induct)
(auto simp: Formula_Operations.lformula.simps split: bool.splits)
next
fix idx and a :: ws1s and x assume "wf0 idx a"
then show "Formula_Operations.wf SUC wf0 idx (rderiv0 x a)"
by (induct a rule: lderiv0.induct)
(auto simp: Formula_Operations.wf.simps Let_def sorted_append
split: bool.splits order.splits nat.splits)
next
fix \<AA> :: interp and a :: ws1s
- note fmember_iff_member_fset[symmetric, simp]
assume "Length \<AA> = 0"
then show "nullable0 a = satisfies0 \<AA> a"
by (induct a, unfold wf0.simps nullable0.simps satisfies0.simps Let_def)
(transfer, (auto 0 3 dest: MSB_greater split: prod.splits if_splits option.splits bool.splits nat.splits) [])+ \<comment> \<open>slow\<close>
next
note Formula_Operations.satisfies_gen.simps[simp] Let_def[simp] upshift_def[simp]
fix x :: atom and a :: ws1s and \<AA> :: interp
assume "lformula0 a" "wf0 (#\<^sub>V \<AA>) a" "#\<^sub>V \<AA> = size_atom x"
then show "Formula_Operations.satisfies Extend Length satisfies0 \<AA> (lderiv0 x a) =
satisfies0 (CONS x \<AA>) a"
proof (induct a)
case 18
then show ?case
apply (auto simp: sum_pow2_image_Suc sum_pow2_insert0 image_iff split: prod.splits)
apply presburger+
done
qed (auto split: prod.splits bool.splits)
next
note Formula_Operations.satisfies_gen.simps[simp] Let_def[simp] upshift_def[simp]
fix x :: atom and a :: ws1s and \<AA> :: interp
assume "lformula0 a" "wf0 (#\<^sub>V \<AA>) a" "#\<^sub>V \<AA> = size_atom x"
then show "Formula_Operations.satisfies_bounded Extend Length len satisfies0 \<AA> (lderiv0 x a) =
satisfies0 (CONS x \<AA>) a"
proof (induct a)
case 18
then show ?case
apply (auto simp: sum_pow2_image_Suc sum_pow2_insert0 image_iff split: prod.splits)
apply presburger+
done
qed (auto split: prod.splits bool.splits)
next
note Formula_Operations.satisfies_gen.simps[simp] Let_def[simp]
fix x :: atom and a :: ws1s and \<AA> :: interp
assume "wf0 (#\<^sub>V \<AA>) a" "#\<^sub>V \<AA> = size_atom x"
then show "Formula_Operations.satisfies_bounded Extend Length len satisfies0 \<AA> (rderiv0 x a) =
satisfies0 (SNOC x \<AA>) a"
proof (induct a)
case Eq_Const then show ?case by (auto split: prod.splits option.splits nat.splits)
next
case Less then show ?case
by (auto split: prod.splits option.splits bool.splits) (metis fMin_less_Length less_not_sym)+
next
case (Plus_FO i m1 m2 n) then show ?case
by (auto simp: min.commute dest: fMin_less_Length
split: prod.splits option.splits nat.splits bool.splits)
next
case Eq_FO then show ?case
by (auto split: prod.splits option.splits bool.splits) (metis fMin_less_Length less_not_sym)+
next
case Eq_SO then show ?case
by (auto split: prod.splits option.splits bool.splits)
(metis assigns_less_Length finsertI1 less_not_refl)+
next
case Suc_SO then show ?case
- apply (auto 2 1 split: prod.splits)
- apply (metis finsert_iff gr0_implies_Suc in_fimage_Suc nat.distinct(2))
- apply (metis finsert_iff in_fimage_Suc less_not_refl)
- apply (metis (no_types, opaque_lifting) fimage_finsert finsertE finsertI1 finsert_commute in_fimage_Suc n_not_Suc_n)
- apply (metis (no_types, opaque_lifting) assigns_less_Length order.strict_iff_order finsert_iff in_fimage_Suc not_less_eq_eq order_refl)
- apply (metis assigns_less_Length fimageI finsert_iff less_irrefl_nat nat.inject)
- apply (metis finsertE finsertI1 finsert_commute finsert_fminus_single in_fimage_Suc n_not_Suc_n)
- apply (metis (no_types, opaque_lifting) assigns_less_Length finsertE fminus_finsert2 fminus_iff in_fimage_Suc lessI not_less_iff_gr_or_eq)
- apply (metis assigns_less_Length finsert_iff lessI not_less_iff_gr_or_eq)
- apply (metis assigns_less_Length fimage_finsert finsert_iff not_less_eq not_less_iff_gr_or_eq)
- apply metis
- apply (metis assigns_less_Length order.strict_iff_order finsert_iff in_fimage_Suc not_less_eq_eq order_refl)
- apply (metis Suc_leI assigns_less_Length fimageI finsert_iff le_eq_less_or_eq lessI less_imp_not_less)
- apply (metis assigns_less_Length fimageE finsertI1 finsert_fminus_if fminus_finsert_absorb lessI less_not_sym)
- apply (metis assigns_less_Length order.strict_iff_order finsert_iff not_less_eq_eq order_refl)
- apply (metis assigns_less_Length order.strict_iff_order finsert_iff not_less_eq_eq order_refl)
- apply (metis assigns_less_Length fimage_Suc_inj fimage_finsert finsert_absorb finsert_iff less_not_refl nat.distinct(2))
- apply (metis assigns_less_Length fimage_Suc_inj fimage_finsert finsertI1 finsert_absorb less_not_refl)
- apply (metis assigns_less_Length fimage_Suc_inj fimage_finsert finsert_absorb finsert_iff less_not_refl nat.distinct(2))
- apply (metis assigns_less_Length fimage_Suc_inj fimage_finsert finsertI1 finsert_absorb2 less_not_refl)
- done
+ apply (auto 2 1 dest: assigns_less_Length split: prod.splits)
+ apply (metis fimage.rep_eq finsert_iff less_not_refl)
+ apply (metis fimage.rep_eq finsert_iff less_not_refl)
+ apply (metis fimage.rep_eq finsert_iff n_not_Suc_n)
+ apply (metis Suc_lessD assigns_less_Length fimage.rep_eq finsert_iff not_less_eq)
+ apply (metis Suc_less_eq assigns_less_Length fimageI finsert_iff less_not_refl)
+ apply (metis fimage.rep_eq finsert_fminus_single finsert_iff n_not_Suc_n)
+ apply (metis assigns_less_Length dual_order.strict_trans fimage.rep_eq finsert_fminus_single finsert_iff lessI less_not_refl)
+ apply (metis Suc_n_not_le_n assigns_less_Length finsert_iff less_or_eq_imp_le)
+ apply (metis Suc_n_not_le_n assigns_less_Length finsertE finsertI1 less_or_eq_imp_le)
+ apply (metis assigns_less_Length fimage.rep_eq finsert_iff lessI order_less_imp_not_less)
+ apply (metis Length_notin_assigns finsert_fimage finsert_iff nat.inject)
+ apply (metis assigns_less_Length fimage.rep_eq finsert_fminus_single finsert_iff not_add_less2 plus_1_eq_Suc)
+ apply (metis assigns_less_Length finsertI1 finsert_commute not_add_less2 plus_1_eq_Suc)
+ apply (metis assigns_less_Length finsertI1 not_add_less2 plus_1_eq_Suc)
+ apply (metis Length_notin_assigns Suc_in_fimage_Suc finsert_iff)
+ by (metis Length_notin_assigns Suc_in_fimage_Suc finsertI1)
next
case In then show ?case by (auto split: prod.splits) (metis fMin_less_Length less_not_sym)+
next
case (Eq_Max b m M) then show ?case
by (auto split: prod.splits bool.splits)
(metis fMax_less_Length less_not_sym, (metis fMin_less_Length less_not_sym)+)
next
case Eq_Min then show ?case
by (auto split: prod.splits bool.splits) (metis fMin_less_Length less_not_sym)+
next
case Eq_Union then show ?case
by (auto 0 0 simp add: fset_eq_iff split: prod.splits) (metis assigns_less_Length less_not_refl)+
next
case Eq_Inter then show ?case
by (auto 0 0 simp add: fset_eq_iff split: prod.splits) (metis assigns_less_Length less_not_refl)+
next
case Eq_Diff then show ?case
by (auto 0 1 simp add: fset_eq_iff split: prod.splits) (metis assigns_less_Length less_not_refl)+
next
let ?f = "sum ((^) (2 :: nat))"
- note fmember_iff_member_fset[symmetric, simp]
case (Eq_Presb l M n)
moreover
let ?M = "fset (M\<^bsup>\<AA>\<^esup>SO)" and ?L = "Length \<AA>"
have "?f (insert ?L ?M) = 2 ^ ?L + ?f ?M"
by (subst sum.insert) auto
moreover have "n > 0 \<Longrightarrow> 2 ^ Max {i. 2 ^ i \<le> n} \<le> n"
using Max_in[of "{i. 2 ^ i \<le> n}", simplified, OF exI[of _ 0]] by auto
moreover
{ have "?f ?M \<le> ?f {0 ..< ?L}" by (rule sum_mono2) auto
also have "\<dots> = 2 ^ ?L - 1" by (rule sum_pow2_upto)
also have "\<dots> < 2 ^ ?L" by simp
finally have "?f ?M < 2 ^ ?L" .
}
moreover have "Max {i. 2 ^ i \<le> 2 ^ ?L + ?f ?M} = ?L"
proof (intro Max_eqI, safe)
fix y assume "2 ^ y \<le> 2 ^ ?L + ?f ?M"
{ assume "?L < y"
then have "(2 :: nat) ^ ?L + 2 ^ ?L \<le> 2 ^ y"
by (cases y) (auto simp: less_Suc_eq_le numeral_eq_Suc add_le_mono)
also note \<open>2 ^ y \<le> 2 ^ ?L + ?f ?M\<close>
finally have " 2 ^ ?L \<le> ?f ?M" by simp
with \<open>?f ?M < 2 ^ ?L\<close> have False by auto
} then show "y \<le> ?L" by (intro leI) blast
qed auto
ultimately show ?case by (auto split: prod.splits option.splits nat.splits)
qed (auto split: prod.splits)
next
fix a :: ws1s and \<AA> \<BB> :: interp
assume "wf0 (#\<^sub>V \<BB>) a" "#\<^sub>V \<AA> = #\<^sub>V \<BB>" "(\<And>m k. LESS k m (#\<^sub>V \<BB>) \<Longrightarrow> m\<^bsup>\<AA>\<^esup>k = m\<^bsup>\<BB>\<^esup>k)" "lformula0 a"
then show "satisfies0 \<AA> a \<longleftrightarrow> satisfies0 \<BB> a" by (induct a) auto
next
fix a :: ws1s
assume "lformula0 a"
moreover
define d where "d = Formula_Operations.deriv extend lderiv0"
define \<Phi> :: "_ \<Rightarrow> (ws1s, order) aformula set"
where "\<Phi> a =
(case a of
Fo m \<Rightarrow> {FBase (Fo m), FBool True}
| Eq_Const None m n \<Rightarrow> {FBase (Eq_Const None m i) | i . i \<le> n} \<union> {FBool True, FBool False}
| Less None m1 m2 \<Rightarrow> {FBase (Less None m1 m2), FBase (Fo m2), FBool True, FBool False}
| Plus_FO None m1 m2 n \<Rightarrow> {FBase (Eq_Const None m1 i) | i . i \<le> n} \<union>
{FBase (Plus_FO None m1 m2 n), FBool True, FBool False}
| Eq_FO False m1 m2 \<Rightarrow> {FBase (Eq_FO False m1 m2), FBool True, FBool False}
| Eq_SO M1 M2 \<Rightarrow> {FBase (Eq_SO M1 M2), FBool False}
| Suc_SO False bl M1 M2 \<Rightarrow> {FBase (Suc_SO False True M1 M2), FBase (Suc_SO False False M1 M2),
FBool False}
| Empty M \<Rightarrow> {FBase (Empty M), FBool False}
| Singleton M \<Rightarrow> {FBase (Singleton M), FBase (Empty M), FBool False}
| Subset M1 M2 \<Rightarrow> {FBase (Subset M1 M2), FBool False}
| In False i I \<Rightarrow> {FBase (In False i I), FBool True, FBool False}
| Eq_Max False m M \<Rightarrow> {FBase (Eq_Max False m M), FBase (Empty M), FBool False}
| Eq_Min False m M \<Rightarrow> {FBase (Eq_Min False m M), FBool True, FBool False}
| Eq_Union M1 M2 M3 \<Rightarrow> {FBase (Eq_Union M1 M2 M3), FBool False}
| Eq_Inter M1 M2 M3 \<Rightarrow> {FBase (Eq_Inter M1 M2 M3), FBool False}
| Eq_Diff M1 M2 M3 \<Rightarrow> {FBase (Eq_Diff M1 M2 M3), FBool False}
| Disjoint M1 M2 \<Rightarrow> {FBase (Disjoint M1 M2), FBool False}
| Eq_Presb None M n \<Rightarrow> {FBase (Eq_Presb None M i) | i . i \<le> n} \<union> {FBool False}
| _ \<Rightarrow> {})" for a
{ fix xs
note Formula_Operations.fold_deriv_FBool[simp] Formula_Operations.deriv.simps[simp] \<Phi>_def[simp]
from \<open>lformula0 a\<close> have "FBase a \<in> \<Phi> a" by auto
moreover have "\<And>x \<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> d x \<phi> \<in> \<Phi> a"
by (auto simp: d_def split: atomic.splits list.splits bool.splits if_splits option.splits)
then have "\<And>\<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> fold d xs \<phi> \<in> \<Phi> a" by (induct xs) auto
ultimately have "fold d xs (FBase a) \<in> \<Phi> a" by blast
}
moreover have "finite (\<Phi> a)" using \<open>lformula0 a\<close> unfolding \<Phi>_def by (auto split: atomic.splits)
ultimately show "finite {fold d xs (FBase a) | xs. True}" by (blast intro: finite_subset)
next
fix a :: ws1s
define d where "d = Formula_Operations.deriv extend rderiv0"
define \<Phi> :: "_ \<Rightarrow> (ws1s, order) aformula set"
where "\<Phi> a =
(case a of
Fo m \<Rightarrow> {FBase (Fo m), FBool True}
| Eq_Const i m n \<Rightarrow>
{FBase (Eq_Const (Some j) m n) | j . j \<le> (case i of Some i \<Rightarrow> max i n | _ \<Rightarrow> n)} \<union>
{FBase (Eq_Const None m n)}
| Less b m1 m2 \<Rightarrow> {FBase (Less None m1 m2), FBase (Less (Some True) m1 m2),
FBase (Less (Some False) m1 m2)}
| Plus_FO i m1 m2 n \<Rightarrow>
{FBase (Plus_FO (Some j) m1 m2 n) | j . j \<le> (case i of Some i \<Rightarrow> max i n | _ \<Rightarrow> n)} \<union>
{FBase (Plus_FO None m1 m2 n)}
| Eq_FO b m1 m2 \<Rightarrow> {FBase (Eq_FO True m1 m2), FBase (Eq_FO False m1 m2)}
| Eq_SO M1 M2 \<Rightarrow> {FBase (Eq_SO M1 M2), FBool False}
| Suc_SO br bl M1 M2 \<Rightarrow> {FBase (Suc_SO False True M1 M2), FBase (Suc_SO False False M1 M2),
FBase (Suc_SO True True M1 M2), FBase (Suc_SO True False M1 M2), FBool False}
| Empty M \<Rightarrow> {FBase (Empty M), FBool False}
| Singleton M \<Rightarrow> {FBase (Singleton M), FBase (Empty M), FBool False}
| Subset M1 M2 \<Rightarrow> {FBase (Subset M1 M2), FBool False}
| In b i I \<Rightarrow> {FBase (In True i I), FBase (In False i I)}
| Eq_Max b m M \<Rightarrow> {FBase (Eq_Max False m M), FBase (Eq_Max True m M), FBool False}
| Eq_Min b m M \<Rightarrow> {FBase (Eq_Min False m M), FBase (Eq_Min True m M)}
| Eq_Union M1 M2 M3 \<Rightarrow> {FBase (Eq_Union M1 M2 M3), FBool False}
| Eq_Inter M1 M2 M3 \<Rightarrow> {FBase (Eq_Inter M1 M2 M3), FBool False}
| Eq_Diff M1 M2 M3 \<Rightarrow> {FBase (Eq_Diff M1 M2 M3), FBool False}
| Disjoint M1 M2 \<Rightarrow> {FBase (Disjoint M1 M2), FBool False}
| Eq_Presb i M n \<Rightarrow> {FBase (Eq_Presb (Some l) M j) | j l .
j \<le> (case i of Some i \<Rightarrow> max i n | _ \<Rightarrow> n) \<and> l \<le> (case i of Some i \<Rightarrow> max i n | _ \<Rightarrow> n)} \<union>
{FBase (Eq_Presb None M n), FBool False})" for a
{ fix xs
note Formula_Operations.fold_deriv_FBool[simp] Formula_Operations.deriv.simps[simp] \<Phi>_def[simp]
then have "FBase a \<in> \<Phi> a" by (auto split: atomic.splits option.splits)
moreover have "\<And>x \<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> d x \<phi> \<in> \<Phi> a"
by (auto simp add: d_def Let_def not_le gr0_conv_Suc leD[OF ld_bounded]
split: atomic.splits list.splits bool.splits if_splits option.splits nat.splits)
then have "\<And>\<phi>. \<phi> \<in> \<Phi> a \<Longrightarrow> fold d xs \<phi> \<in> \<Phi> a"
by (induct xs) auto
ultimately have "fold d xs (FBase a) \<in> \<Phi> a" by blast
}
moreover have "finite (\<Phi> a)" unfolding \<Phi>_def using [[simproc add: finite_Collect]]
by (auto split: atomic.splits)
ultimately show "finite {fold d xs (FBase a) | xs. True}" by (blast intro: finite_subset)
next
fix k l and a :: ws1s
show "find0 k l a \<longleftrightarrow> l \<in> FV0 k a" by (induct a rule: find0.induct) auto
next
fix a :: ws1s and k :: order
show "finite (FV0 k a)" by (cases k) (induct a, auto)+
next
fix idx a k v
assume "wf0 idx a" "v \<in> FV0 k a"
then show "LESS k v idx" by (cases k) (induct a, auto)+
next
fix idx k i
assume "LESS k i idx"
then show "Formula_Operations.wf SUC wf0 idx (Restrict k i)"
unfolding Restrict_def by (cases k) (auto simp: Formula_Operations.wf.simps)
next
fix k and i :: nat
show "Formula_Operations.lformula lformula0 (Restrict k i)"
unfolding Restrict_def by (cases k) (auto simp: Formula_Operations.lformula.simps)
next
fix i \<AA> k P r
assume "i\<^bsup>\<AA>\<^esup>k = P"
then show "restrict k P \<longleftrightarrow>
Formula_Operations.satisfies_gen Extend Length satisfies0 r \<AA> (Restrict k i)"
unfolding restrict_def Restrict_def
by (cases k) (auto simp: Formula_Operations.satisfies_gen.simps)
qed (auto simp: Extend_commute_unsafe downshift_def upshift_def fimage_iff Suc_le_eq len_def
dec_def eval_def cut_def len_downshift_helper dest!: CONS_surj
dest: fMax_ge fMax_ffilter_less fMax_boundedD fsubset_fsingletonD
split!: order.splits if_splits)
(*Workaround for code generation*)
lemma check_eqv_code[code]: "check_eqv idx r s =
((ws1s_wf idx r \<and> ws1s_lformula r) \<and> (ws1s_wf idx s \<and> ws1s_lformula s) \<and>
(case rtrancl_while (\<lambda>(p, q). final idx p = final idx q)
(\<lambda>(p, q). map (\<lambda>a. (norm (deriv lderiv0 a p), norm (deriv lderiv0 a q))) (\<sigma> idx))
(norm (RESTRICT r), norm (RESTRICT s)) of
None \<Rightarrow> False
| Some ([], x) \<Rightarrow> True
| Some (a # list, x) \<Rightarrow> False))"
unfolding check_eqv_def WS1S.check_eqv_def WS1S.step_alt ..
definition while where [code del, code_abbrev]: "while idx \<phi> = while_default (fut_default idx \<phi>)"
declare while_default_code[of "fut_default idx \<phi>" for idx \<phi>, folded while_def, code]
lemma check_eqv_sound:
"\<lbrakk>#\<^sub>V \<AA> = idx; check_eqv idx \<phi> \<psi>\<rbrakk> \<Longrightarrow> (WS1S.sat \<AA> \<phi> \<longleftrightarrow> WS1S.sat \<AA> \<psi>)"
unfolding check_eqv_def by (rule WS1S.check_eqv_soundness)
lemma bounded_check_eqv_sound:
"\<lbrakk>#\<^sub>V \<AA> = idx; bounded_check_eqv idx \<phi> \<psi>\<rbrakk> \<Longrightarrow> (WS1S.sat\<^sub>b \<AA> \<phi> \<longleftrightarrow> WS1S.sat\<^sub>b \<AA> \<psi>)"
unfolding bounded_check_eqv_def by (rule WS1S.bounded_check_eqv_soundness)
method_setup check_equiv = \<open>
let
fun tac ctxt =
let
val conv = @{computation_check terms: Trueprop
"0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat" Suc
"plus :: nat \<Rightarrow> _" "minus :: nat \<Rightarrow> _"
"times :: nat \<Rightarrow> _" "divide :: nat \<Rightarrow> _" "modulo :: nat \<Rightarrow> _"
"0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int"
check_eqv datatypes: formula "int list" integer idx
"nat \<times> nat" "nat option" "bool option"} ctxt
in
CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 conv)) ctxt) THEN'
resolve_tac ctxt [TrueI]
end
in
Scan.succeed (SIMPLE_METHOD' o tac)
end
\<close>
end
diff --git a/thys/Formula_Derivatives/WS1S_Prelim.thy b/thys/Formula_Derivatives/WS1S_Prelim.thy
--- a/thys/Formula_Derivatives/WS1S_Prelim.thy
+++ b/thys/Formula_Derivatives/WS1S_Prelim.thy
@@ -1,388 +1,404 @@
section \<open>WS1S Interpretations\<close>
(*<*)
theory WS1S_Prelim
imports
FSet_More
Deriving.Compare_Instances
"List-Index.List_Index"
begin
hide_const (open) cut
(*>*)
definition "eval P x = (x |\<in>| P)"
definition "downshift P = (\<lambda>x. x - Suc 0) |`| (P |-| {|0|})"
definition "upshift P = Suc |`| P"
definition "lift bs i P = (if bs ! i then finsert 0 (upshift P) else upshift P)"
definition "snoc n bs i P = (if bs ! i then finsert n P else P)"
definition "cut n P = ffilter (\<lambda>i. i < n) P"
definition "len P = (if P = {||} then 0 else Suc (fMax P))"
datatype order = FO | SO
derive linorder order
instantiation order :: enum begin
definition "enum_order = [FO, SO]"
definition "enum_all_order P = (P FO \<and> P SO)"
definition "enum_ex_order P = (P FO \<or> P SO)"
lemmas enum_defs = enum_order_def enum_all_order_def enum_ex_order_def
instance proof qed (auto simp: enum_defs, (metis (full_types) order.exhaust)+)
end
typedef idx = "UNIV :: (nat \<times> nat) set" by (rule UNIV_witness)
setup_lifting type_definition_idx
lift_definition SUC :: "order \<Rightarrow> idx \<Rightarrow> idx" is
"\<lambda>ord (m, n). case ord of FO \<Rightarrow> (Suc m, n) | SO \<Rightarrow> (m, Suc n)" .
lift_definition LESS :: "order \<Rightarrow> nat \<Rightarrow> idx \<Rightarrow> bool" is
"\<lambda>ord l (m, n). case ord of FO \<Rightarrow> l < m | SO \<Rightarrow> l < n" .
abbreviation "LEQ ord l idx \<equiv> LESS ord l (SUC ord idx)"
definition "MSB Is \<equiv>
if \<forall>P \<in> set Is. P = {||} then 0 else Suc (Max (\<Union>P \<in> set Is. fset P))"
lemma MSB_Nil[simp]: "MSB [] = 0"
unfolding MSB_def by simp
lemma MSB_Cons[simp]: "MSB (I # Is) = max (if I = {||} then 0 else Suc (fMax I)) (MSB Is)"
unfolding MSB_def including fset.lifting
by transfer (auto simp: Max_Un list_all_iff Sup_bot_conv(2)[symmetric] simp del: Sup_bot_conv(2))
lemma MSB_append[simp]: "MSB (I1 @ I2) = max (MSB I1) (MSB I2)"
by (induct I1) auto
lemma MSB_insert_nth[simp]:
"MSB (insert_nth n P Is) = max (if P = {||} then 0 else Suc (fMax P)) (MSB Is)"
by (subst (2) append_take_drop_id[of n Is, symmetric])
(simp only: insert_nth_take_drop MSB_append MSB_Cons MSB_Nil)
lemma MSB_greater:
"\<lbrakk>i < length Is; p |\<in>| Is ! i\<rbrakk> \<Longrightarrow> p < MSB Is"
unfolding MSB_def by (fastforce simp: Bex_def in_set_conv_nth less_Suc_eq_le intro: Max_ge)
lemma MSB_mono: "set I1 \<subseteq> set I2 \<Longrightarrow> MSB I1 \<le> MSB I2"
unfolding MSB_def including fset.lifting
by transfer (auto simp: list_all_iff intro!: Max_ge)
lemma MSB_map_index'_CONS[simp]:
"MSB (map_index' i (lift bs) Is) =
(if MSB Is = 0 \<and> (\<forall>i \<in> {i ..< i + length Is}. \<not> bs ! i) then 0 else Suc (MSB Is))"
by (induct Is arbitrary: i)
(auto split: if_splits simp: mono_fMax_commute[where f = Suc, symmetric] mono_def
lift_def upshift_def,
metis atLeastLessThan_iff le_antisym not_less_eq_eq)
lemma MSB_map_index'_SNOC[simp]:
"MSB Is \<le> n \<Longrightarrow> MSB (map_index' i (snoc n bs) Is) =
(if (\<forall>i \<in> {i ..< i + length Is}. \<not> bs ! i) then MSB Is else Suc n)"
by (induct Is arbitrary: i)
(auto split: if_splits simp: mono_fMax_commute[where f = Suc, symmetric] mono_def
snoc_def, (metis atLeastLessThan_iff le_antisym not_less_eq_eq)+)
lemma MSB_replicate[simp]: "MSB (replicate n P) = (if P = {||} \<or> n = 0 then 0 else Suc (fMax P))"
by (induct n) auto
typedef interp =
"{(n :: nat, I1 :: nat fset list, I2 :: nat fset list) | n I1 I2. MSB (I1 @ I2) \<le> n}"
by auto
setup_lifting type_definition_interp
lift_definition assigns :: "nat \<Rightarrow> interp \<Rightarrow> order \<Rightarrow> nat fset" ("_\<^bsup>_\<^esup>_" [900, 999, 999] 999)
is "\<lambda>n (_, I1, I2) ord. case ord of FO \<Rightarrow> if n < length I1 then I1 ! n else {||}
| SO \<Rightarrow> if n < length I2 then I2 ! n else {||}" .
lift_definition nvars :: "interp \<Rightarrow> idx" ("#\<^sub>V _" [1000] 900)
is "\<lambda>(_, I1, I2). (length I1, length I2)" .
lift_definition Length :: "interp \<Rightarrow> nat"
is "\<lambda>(n, _, _). n" .
lift_definition Extend :: "order \<Rightarrow> nat \<Rightarrow> interp \<Rightarrow> nat fset \<Rightarrow> interp"
is "\<lambda>ord i (n, I1, I2) P. case ord of
FO \<Rightarrow> (max n (if P = {||} then 0 else Suc (fMax P)), insert_nth i P I1, I2)
| SO \<Rightarrow> (max n (if P = {||} then 0 else Suc (fMax P)), I1, insert_nth i P I2)"
using MSB_mono by (auto simp del: insert_nth_take_drop split: order.splits)
lift_definition CONS :: "(bool list \<times> bool list) \<Rightarrow> interp \<Rightarrow> interp"
is "\<lambda>(bs1, bs2) (n, I1, I2).
(Suc n, map_index (lift bs1) I1, map_index (lift bs2) I2)"
by auto
lift_definition SNOC :: "(bool list \<times> bool list) \<Rightarrow> interp \<Rightarrow> interp"
is "\<lambda>(bs1, bs2) (n, I1, I2).
(Suc n, map_index (snoc n bs1) I1, map_index (snoc n bs2) I2)"
by (auto simp: Let_def)
type_synonym atom = "bool list \<times> bool list"
lift_definition zero :: "idx \<Rightarrow> atom"
is "\<lambda>(m, n). (replicate m False, replicate n False)" .
definition "extend ord b \<equiv>
\<lambda>(bs1, bs2). case ord of FO \<Rightarrow> (b # bs1, bs2) | SO \<Rightarrow> (bs1, b # bs2)"
lift_definition size_atom :: "bool list \<times> bool list \<Rightarrow> idx"
is "\<lambda>(bs1, bs2). (length bs1, length bs2)" .
lift_definition \<sigma> :: "idx \<Rightarrow> atom list"
is "(\<lambda>(n, N). map (\<lambda>bs. (take n bs, drop n bs)) (List.n_lists (n + N) [True, False]))" .
lemma fMin_fimage_Suc[simp]: "x |\<in>| A \<Longrightarrow> fMin (Suc |`| A) = Suc (fMin A)"
by (rule fMin_eqI) (auto intro: fMin_in)
lemma fMin_eq_0[simp]: "0 |\<in>| A \<Longrightarrow> fMin A = (0 :: nat)"
by (rule fMin_eqI) auto
lemma insert_nth_Cons[simp]:
"insert_nth i x (y # xs) = (case i of 0 \<Rightarrow> x # y # xs | Suc i \<Rightarrow> y # insert_nth i x xs)"
by (cases i) simp_all
lemma insert_nth_commute[simp]:
assumes "j \<le> i" "i \<le> length xs"
shows "insert_nth j y (insert_nth i x xs) = insert_nth (Suc i) x (insert_nth j y xs)"
using assms by (induct xs arbitrary: i j) (auto simp del: insert_nth_take_drop split: nat.splits)
lemma SUC_SUC[simp]: "SUC ord (SUC ord' idx) = SUC ord' (SUC ord idx)"
by transfer (auto split: order.splits)
lemma LESS_SUC[simp]:
"LESS ord 0 (SUC ord idx)"
"LESS ord (Suc l) (SUC ord idx) = LESS ord l idx"
"ord \<noteq> ord' \<Longrightarrow> LESS ord l (SUC ord' idx) = LESS ord l idx"
"LESS ord l idx \<Longrightarrow> LESS ord l (SUC ord' idx)"
by (transfer, force split: order.splits)+
lemma nvars_Extend[simp]:
"#\<^sub>V (Extend ord i \<AA> P) = SUC ord (#\<^sub>V \<AA>)"
by (transfer, force split: order.splits)
lemma Length_Extend[simp]:
"Length (Extend k i \<AA> P) = max (Length \<AA>) (if P = {||} then 0 else Suc (fMax P))"
unfolding max_def by (split if_splits, transfer) (force split: order.splits)
lemma assigns_Extend[simp]:
"LEQ ord i (#\<^sub>V \<AA>) \<Longrightarrow>m\<^bsup>Extend ord i \<AA> P\<^esup>ord = (if m = i then P else (if m > i then m - Suc 0 else m)\<^bsup>\<AA>\<^esup>ord)"
"ord \<noteq> ord' \<Longrightarrow> m\<^bsup>Extend ord i \<AA> P\<^esup>ord' = m\<^bsup>\<AA>\<^esup>ord'"
by (transfer, force simp: min_def nth_append split: order.splits)+
lemma Extend_commute_safe[simp]:
"\<lbrakk>j \<le> i; LEQ ord i (#\<^sub>V \<AA>)\<rbrakk> \<Longrightarrow>
Extend ord j (Extend ord i \<AA> P1) P2 = Extend ord (Suc i) (Extend ord j \<AA> P2) P1"
by (transfer,
force simp del: insert_nth_take_drop simp: replicate_add[symmetric] split: order.splits)
lemma Extend_commute_unsafe:
"ord \<noteq> ord' \<Longrightarrow> Extend ord j (Extend ord' i \<AA> P1) P2 = Extend ord' i (Extend ord j \<AA> P2) P1"
by (transfer, force simp: replicate_add[symmetric] split: order.splits)
lemma Length_CONS[simp]:
"Length (CONS x \<AA>) = Suc (Length \<AA>)"
by (transfer, force split: order.splits)
lemma Length_SNOC[simp]:
"Length (SNOC x \<AA>) = Suc (Length \<AA>)"
by (transfer, force simp: Let_def split: order.splits)
lemma nvars_CONS[simp]:
"#\<^sub>V (CONS x \<AA>) = #\<^sub>V \<AA>"
by (transfer, force)
lemma nvars_SNOC[simp]:
"#\<^sub>V (SNOC x \<AA>) = #\<^sub>V \<AA>"
by (transfer, force simp: Let_def)
lemma assigns_CONS[simp]:
assumes "#\<^sub>V \<AA> = size_atom bs1_bs2"
shows "LESS ord x (#\<^sub>V \<AA>) \<Longrightarrow> x\<^bsup>CONS bs1_bs2 \<AA>\<^esup>ord =
(if case_prod case_order bs1_bs2 ord ! x then finsert 0 (upshift (x\<^bsup>\<AA>\<^esup>ord)) else upshift (x\<^bsup>\<AA>\<^esup>ord))"
by (insert assms, transfer) (auto simp: lift_def split: order.splits)
lemma assigns_SNOC[simp]:
assumes "#\<^sub>V \<AA> = size_atom bs1_bs2"
shows "LESS ord x (#\<^sub>V \<AA>) \<Longrightarrow> x\<^bsup>SNOC bs1_bs2 \<AA>\<^esup>ord =
(if case_prod case_order bs1_bs2 ord ! x then finsert (Length \<AA>) (x\<^bsup>\<AA>\<^esup>ord) else x\<^bsup>\<AA>\<^esup>ord)"
by (insert assms, transfer) (force simp: snoc_def Let_def split: order.splits)
lemma map_index'_eq_conv[simp]:
"map_index' i f xs = map_index' j g xs = (\<forall>k < length xs. f (i + k) (xs ! k) = g (j + k) (xs ! k))"
proof (induct xs arbitrary: i j)
case Cons from Cons(1)[of "Suc i" "Suc j"] show ?case by (auto simp: nth_Cons split: nat.splits)
qed simp
lemma fMax_Diff_0[simp]: "Suc m |\<in>| P \<Longrightarrow> fMax (P |-| {|0|}) = fMax P"
by (rule fMax_eqI) (auto intro: fMax_in dest: fMax_ge)
lemma Suc_fMax_pred_fimage[simp]:
assumes "Suc p |\<in>| P" "0 |\<notin>| P"
shows "Suc (fMax ((\<lambda>x. x - Suc 0) |`| P)) = fMax P"
using assms by (subst mono_fMax_commute[of Suc, unfolded mono_def, simplified]) (auto simp: o_def)
lemma Extend_CONS[simp]: "#\<^sub>V \<AA> = size_atom x \<Longrightarrow> Extend ord 0 (CONS x \<AA>) P =
CONS (extend ord (eval P 0) x) (Extend ord 0 \<AA> (downshift P))"
by transfer (auto simp: extend_def o_def gr0_conv_Suc
mono_fMax_commute[of Suc, symmetric, unfolded mono_def, simplified]
lift_def upshift_def downshift_def eval_def
dest!: fsubset_fsingletonD split: order.splits)
lemma size_atom_extend[simp]:
"size_atom (extend ord b x) = SUC ord (size_atom x)"
unfolding extend_def by transfer (simp split: prod.splits order.splits)
lemma size_atom_zero[simp]:
"size_atom (zero idx) = idx"
unfolding extend_def by transfer (simp split: prod.splits order.splits)
lemma interp_eqI:
"\<lbrakk>Length \<AA> = Length \<BB>; #\<^sub>V \<AA> = #\<^sub>V \<BB>; \<And>m k. LESS k m (#\<^sub>V \<AA>) \<Longrightarrow> m\<^bsup>\<AA>\<^esup>k = m\<^bsup>\<BB>\<^esup>k\<rbrakk> \<Longrightarrow> \<AA> = \<BB>"
by transfer (force split: order.splits intro!: nth_equalityI)
lemma Extend_SNOC_cut[unfolded eval_def cut_def Length_SNOC, simp]:
"\<lbrakk>len P \<le> Length (SNOC x \<AA>); #\<^sub>V \<AA> = size_atom x\<rbrakk> \<Longrightarrow>
Extend ord 0 (SNOC x \<AA>) P =
SNOC (extend ord (if eval P (Length \<AA>) then True else False) x) (Extend ord 0 \<AA> (cut (Length \<AA>) P))"
by transfer (fastforce simp: extend_def len_def cut_def ffilter_eq_fempty_iff snoc_def eval_def
split: if_splits order.splits dest: fMax_ge fMax_boundedD intro: fMax_in)
lemma nth_replicate_simp: "replicate m x ! i = (if i < m then x else [] ! (i - m))"
by (induct m arbitrary: i) (auto simp: nth_Cons')
lemma MSB_eq_SucD: "MSB Is = Suc x \<Longrightarrow> \<exists>P\<in>set Is. x |\<in>| P"
using Max_in[of "\<Union>x\<in>set Is. fset x"]
- unfolding MSB_def by (force simp: fmember_def split: if_splits)
+ unfolding MSB_def by (force split: if_splits)
lemma append_replicate_inj:
assumes "xs \<noteq> [] \<Longrightarrow> last xs \<noteq> x" and "ys \<noteq> [] \<Longrightarrow> last ys \<noteq> x"
shows "xs @ replicate m x = ys @ replicate n x \<longleftrightarrow> (xs = ys \<and> m = n)"
proof safe
from assms have assms': "xs \<noteq> [] \<Longrightarrow> rev xs ! 0 \<noteq> x" "ys \<noteq> [] \<Longrightarrow> rev ys ! 0 \<noteq> x"
by (auto simp: hd_rev hd_conv_nth[symmetric])
assume *: "xs @ replicate m x = ys @ replicate n x"
then have "rev (xs @ replicate m x) = rev (ys @ replicate n x)" ..
then have "replicate m x @ rev xs = replicate n x @ rev ys" by simp
then have "take (max m n) (replicate m x @ rev xs) = take (max m n) (replicate n x @ rev ys)" by simp
then have "replicate m x @ take (max m n - m) (rev xs) =
replicate n x @ take (max m n - n) (rev ys)" by (auto simp: min_def max_def split: if_splits)
then have "(replicate m x @ take (max m n - m) (rev xs)) ! min m n =
(replicate n x @ take (max m n - n) (rev ys)) ! min m n" by simp
with arg_cong[OF *, of length, simplified] assms' show "m = n"
by (cases "xs = []" "ys = []" rule: bool.exhaust[case_product bool.exhaust])
(auto simp: min_def nth_append split: if_splits)
with * show "xs = ys" by auto
qed
lemma fin_lift[simp]: "m |\<in>| lift bs i (I ! i) \<longleftrightarrow> (case m of 0 \<Rightarrow> bs ! i | Suc m \<Rightarrow> m |\<in>| I ! i)"
unfolding lift_def upshift_def by (auto split: nat.splits)
lemma ex_Length_0[simp]:
"\<exists>\<AA>. Length \<AA> = 0 \<and> #\<^sub>V \<AA> = idx"
by transfer (auto intro!: exI[of _ "replicate m {||}" for m])
lemma is_empty_inj[simp]: "\<lbrakk>Length \<AA> = 0; Length \<BB> = 0; #\<^sub>V \<AA> = #\<^sub>V \<BB>\<rbrakk> \<Longrightarrow> \<AA> = \<BB>"
by transfer (simp add: list_eq_iff_nth_eq split: prod.splits,
metis MSB_greater fMax_in less_nat_zero_code)
lemma set_\<sigma>_length_atom[simp]: "(x \<in> set (\<sigma> idx)) \<longleftrightarrow> idx = size_atom x"
by transfer (auto simp: set_n_lists enum_UNIV image_iff intro!: exI[of _ "I1 @ I2" for I1 I2])
lemma distinct_\<sigma>[simp]: "distinct (\<sigma> idx)"
by transfer (auto 0 4 simp: distinct_map distinct_n_lists set_n_lists inj_on_def
intro: iffD2[OF append_eq_append_conv]
box_equals[OF _ append_take_drop_id append_take_drop_id, of n _ n for n])
lemma fMin_less_Length[simp]: "x |\<in>| m1\<^bsup>\<AA>\<^esup>k \<Longrightarrow> fMin (m1\<^bsup>\<AA>\<^esup>k) < Length \<AA>"
by transfer
(force elim: order.strict_trans2[OF MSB_greater, rotated -1] intro: fMin_in split: order.splits)
lemma fMax_less_Length[simp]: "x |\<in>| m1\<^bsup>\<AA>\<^esup>k \<Longrightarrow> fMax (m1\<^bsup>\<AA>\<^esup>k) < Length \<AA>"
by transfer
(force elim: order.strict_trans2[OF MSB_greater, rotated -1] intro: fMax_in split: order.splits)
lemma min_Length_fMin[simp]: "x |\<in>| m1\<^bsup>\<AA>\<^esup>k \<Longrightarrow> min (Length \<AA>) (fMin (m1\<^bsup>\<AA>\<^esup>k)) = fMin (m1\<^bsup>\<AA>\<^esup>k)"
using fMin_less_Length[of x m1 \<AA> k] unfolding fMin_def by auto
lemma max_Length_fMin[simp]: "x |\<in>| m1\<^bsup>\<AA>\<^esup>k \<Longrightarrow> max (Length \<AA>) (fMin (m1\<^bsup>\<AA>\<^esup>k)) = Length \<AA>"
using fMin_less_Length[of x m1 \<AA> k] unfolding fMin_def by auto
lemma min_Length_fMax[simp]: "x |\<in>| m1\<^bsup>\<AA>\<^esup>k \<Longrightarrow> min (Length \<AA>) (fMax (m1\<^bsup>\<AA>\<^esup>k)) = fMax (m1\<^bsup>\<AA>\<^esup>k)"
using fMax_less_Length[of x m1 \<AA> k] unfolding fMax_def by auto
lemma max_Length_fMax[simp]: "x |\<in>| m1\<^bsup>\<AA>\<^esup>k \<Longrightarrow> max (Length \<AA>) (fMax (m1\<^bsup>\<AA>\<^esup>k)) = Length \<AA>"
using fMax_less_Length[of x m1 \<AA> k] unfolding fMax_def by auto
lemma assigns_less_Length[simp]: "x |\<in>| m1\<^bsup>\<AA>\<^esup>k \<Longrightarrow> x < Length \<AA>"
by transfer (force dest: MSB_greater split: order.splits if_splits)
lemma Length_notin_assigns[simp]: "Length \<AA> |\<notin>| m\<^bsup>\<AA>\<^esup>k"
by (metis assigns_less_Length less_not_refl)
lemma nth_zero[simp]: "LESS ord m (#\<^sub>V \<AA>) \<Longrightarrow> \<not> case_prod case_order (zero (#\<^sub>V \<AA>)) ord ! m"
by transfer (auto split: order.splits)
lemma in_fimage_Suc[simp]: "x |\<in>| Suc |`| A \<longleftrightarrow> (\<exists>y. y |\<in>| A \<and> x = Suc y)"
by blast
lemma fimage_Suc_inj[simp]: "Suc |`| A = Suc |`| B \<longleftrightarrow> A = B"
by blast
lemma MSB_eq0_D: "MSB I = 0 \<Longrightarrow> x < length I \<Longrightarrow> I ! x = {||}"
unfolding MSB_def by (auto split: if_splits)
lemma Suc_in_fimage_Suc: "Suc x |\<in>| Suc |`| X \<longleftrightarrow> x |\<in>| X"
by auto
lemma Suc_in_fimage_Suc_o_Suc[simp]: "Suc x |\<in>| (Suc \<circ> Suc) |`| X \<longleftrightarrow> x |\<in>| Suc |`| X"
by auto
lemma finsert_same_eq_iff[simp]: "finsert k X = finsert k Y \<longleftrightarrow> X |-| {|k|} = Y |-| {|k|}"
by auto
lemma fimage_Suc_o_Suc_eq_fimage_Suc_iff[simp]:
"((Suc \<circ> Suc) |`| X = Suc |`| Y) \<longleftrightarrow> (Suc |`| X = Y)"
by (metis fimage_Suc_inj fset.map_comp)
lemma fMax_image_Suc[simp]: "x |\<in>| P \<Longrightarrow> fMax (Suc |`| P) = Suc (fMax P)"
by (rule fMax_eqI) (metis Suc_le_mono fMax_ge fimageE, metis fimageI fempty_iff fMax_in)
lemma fimage_Suc_eq_singleton[simp]: "(fimage Suc Z = {|y|}) \<longleftrightarrow> (\<exists>x. Z = {|x|} \<and> Suc x = y)"
by (cases y) auto
lemma len_downshift_helper:
"x |\<in>| P \<Longrightarrow> Suc (fMax ((\<lambda>x. x - Suc 0) |`| (P |-| {|0|}))) \<noteq> fMax P \<Longrightarrow> xa |\<in>| P \<Longrightarrow> xa = 0"
proof -
assume a1: "xa |\<in>| P"
assume a2: "Suc (fMax ((\<lambda>x. x - Suc 0) |`| (P |-| {|0|}))) \<noteq> fMax P"
have "xa |\<in>| {|0|} \<longrightarrow> xa = 0" by fastforce
moreover
{ assume "xa |\<notin>| {|0|}"
hence "0 |\<notin>| P |-| {|0|} \<and> xa |\<notin>| {|0|}" by blast
then obtain esk1\<^sub>1 :: "nat \<Rightarrow> nat" where "xa = 0" using a1 a2 by (metis Suc_fMax_pred_fimage fMax_Diff_0 fminus_iff not0_implies_Suc) }
ultimately show "xa = 0" by blast
qed
lemma CONS_inj[simp]: "size_atom x = #\<^sub>V \<AA> \<Longrightarrow> size_atom y = #\<^sub>V \<AA> \<Longrightarrow> #\<^sub>V \<AA> = #\<^sub>V \<BB> \<Longrightarrow>
CONS x \<AA> = CONS y \<BB> \<longleftrightarrow> (x = y \<and> \<AA> = \<BB>)"
by transfer (auto simp: list_eq_iff_nth_eq lift_def upshift_def split: if_splits; blast)
lemma Suc_minus1: "Suc (x - Suc 0) = (if x = 0 then Suc 0 else x)"
by auto
lemma fset_eq_empty_iff: "(fset X = {}) = (X = {||})"
by (metis bot_fset.rep_eq fset_inverse)
lemma fset_le_singleton_iff: "(fset X \<subseteq> {x}) = (X = {||} \<or> X = {|x|})"
by (metis finsert.rep_eq fset_eq_empty_iff fset_inject order_refl singleton_insert_inj_eq subset_singletonD)
lemma MSB_decreases:
"MSB I \<le> Suc m \<Longrightarrow> MSB (map (\<lambda>X. (\<lambda>I1. I1 - Suc 0) |`| (X |-| {|0|})) I) \<le> m"
unfolding MSB_def
by (auto simp add: not_le less_Suc_eq_le fset_eq_empty_iff fset_le_singleton_iff
split: if_splits dest!: iffD1[OF Max_le_iff, rotated -1] iffD1[OF Max_ge_iff, rotated -1]; force)
-lemma CONS_surj[dest]: "Length \<AA> > 0 \<Longrightarrow>
- \<exists>x \<BB>. \<AA> = CONS x \<BB> \<and> #\<^sub>V \<BB> = #\<^sub>V \<AA> \<and> size_atom x = #\<^sub>V \<AA>"
- by transfer (auto simp: gr0_conv_Suc list_eq_iff_nth_eq lift_def upshift_def split: if_splits
- intro!: exI[of _ "map (\<lambda>X. 0 |\<in>| X) _"] exI[of _ "map (\<lambda>X. (\<lambda>x. x - Suc 0) |`| (X |-| {|0|})) _"],
- auto simp: MSB_decreases upshift_def Suc_minus1 fimage_iff intro: rev_fBexI split: if_splits)
+lemma CONS_surj[dest]:
+ assumes "Length \<AA> > 0"
+ shows "\<exists>x \<BB>. \<AA> = CONS x \<BB> \<and> #\<^sub>V \<BB> = #\<^sub>V \<AA> \<and> size_atom x = #\<^sub>V \<AA>"
+proof -
+ have "MSB I1 \<le> Suc m \<Longrightarrow> MSB I2 \<le> Suc m \<Longrightarrow> \<exists>a b ab ba.
+ (\<exists>I1. length ab = length I1 \<and> (\<forall>i<length ab. ab ! i = I1 ! i) \<and>
+ (\<exists>I2. length ba = length I2 \<and> (\<forall>i<length ba. ba ! i = I2 ! i) \<and>
+ MSB I1 \<le> m \<and> MSB I2 \<le> m)) \<and>
+ I1 = map_index (lift a) ab \<and>
+ I2 = map_index (lift b) ba \<and>
+ length ab = length I1 \<and>
+ length ba = length I2 \<and>
+ length a = length I1 \<and>
+ length b = length I2"
+ for m I1 I2
+ by (auto simp: list_eq_iff_nth_eq lift_def upshift_def split: if_splits
+ intro!: exI[of _ "map (\<lambda>X. 0 |\<in>| X) _"] exI[of _ "map (\<lambda>X. (\<lambda>x. x - Suc 0) |`| (X |-| {|0|})) _"],
+ auto simp: MSB_decreases upshift_def Suc_minus1 fimage_iff intro: rev_fBexI split: if_splits)
+ with assms show ?thesis
+ by transfer (auto simp: gr0_conv_Suc list_eq_iff_nth_eq)
+qed
(*<*)
end
(*>*)
diff --git a/thys/Incredible_Proof_Machine/Abstract_Rules_To_Incredible.thy b/thys/Incredible_Proof_Machine/Abstract_Rules_To_Incredible.thy
--- a/thys/Incredible_Proof_Machine/Abstract_Rules_To_Incredible.thy
+++ b/thys/Incredible_Proof_Machine/Abstract_Rules_To_Incredible.thy
@@ -1,126 +1,126 @@
theory Abstract_Rules_To_Incredible
imports
Main
"HOL-Library.FSet"
"HOL-Library.Stream"
Incredible_Deduction
Abstract_Rules
begin
text \<open>In this theory, the abstract rules given in @{theory Incredible_Proof_Machine.Abstract_Rules} are used to
create a proper signature.\<close>
text \<open>Besides the rules given there, we have nodes for assumptions, conclusions, and the helper
block.\<close>
datatype ('form, 'rule) graph_node = Assumption 'form | Conclusion 'form | Rule 'rule | Helper
type_synonym ('form, 'var) in_port = "('form, 'var) antecedent"
type_synonym 'form reg_out_port = "'form"
type_synonym 'form hyp = "'form"
datatype ('form, 'var) out_port = Reg "'form reg_out_port" | Hyp "'form hyp" "('form, 'var) in_port"
type_synonym ('v, 'form, 'var) edge' = "(('v \<times> ('form, 'var) out_port) \<times> ('v \<times> ('form, 'var) in_port))"
context Abstract_Task
begin
definition nodes :: "('form, 'rule) graph_node stream" where
"nodes = Helper ## shift (map Assumption assumptions) (shift (map Conclusion conclusions) (smap Rule rules))"
lemma Helper_in_nodes[simp]:
"Helper \<in> sset nodes" by (simp add: nodes_def)
lemma Assumption_in_nodes[simp]:
"Assumption a \<in> sset nodes \<longleftrightarrow> a \<in> set assumptions" by (auto simp add: nodes_def stream.set_map)
lemma Conclusion_in_nodes[simp]:
"Conclusion c \<in> sset nodes \<longleftrightarrow> c \<in> set conclusions" by (auto simp add: nodes_def stream.set_map)
lemma Rule_in_nodes[simp]:
"Rule r \<in> sset nodes \<longleftrightarrow> r \<in> sset rules" by (auto simp add: nodes_def stream.set_map)
fun inPorts' :: "('form, 'rule) graph_node \<Rightarrow> ('form, 'var) in_port list" where
"inPorts' (Rule r) = antecedent r"
|"inPorts' (Assumption r) = []"
|"inPorts' (Conclusion r) = [ plain_ant r ]"
|"inPorts' Helper = [ plain_ant anyP ]"
fun inPorts :: "('form, 'rule) graph_node \<Rightarrow> ('form, 'var) in_port fset" where
"inPorts (Rule r) = f_antecedent r"
|"inPorts (Assumption r) = {||}"
|"inPorts (Conclusion r) = {| plain_ant r |}"
|"inPorts Helper = {| plain_ant anyP |}"
lemma inPorts_fset_of:
"inPorts n = fset_from_list (inPorts' n)"
- by (cases n rule: inPorts.cases) (auto simp: fmember_iff_member_fset f_antecedent_def)
+ by (cases n rule: inPorts.cases) (auto simp: f_antecedent_def)
definition outPortsRule where
"outPortsRule r = ffUnion ((\<lambda> a. (\<lambda> h. Hyp h a) |`| a_hyps a) |`| f_antecedent r) |\<union>| Reg |`| f_consequent r"
lemma Reg_in_outPortsRule[simp]: "Reg c |\<in>| outPortsRule r \<longleftrightarrow> c |\<in>| f_consequent r"
- by (auto simp add: outPortsRule_def fmember_iff_member_fset ffUnion.rep_eq)
+ by (auto simp add: outPortsRule_def ffUnion.rep_eq)
lemma Hyp_in_outPortsRule[simp]: "Hyp h c |\<in>| outPortsRule r \<longleftrightarrow> c |\<in>| f_antecedent r \<and> h |\<in>| a_hyps c"
- by (auto simp add: outPortsRule_def fmember_iff_member_fset ffUnion.rep_eq)
+ by (auto simp add: outPortsRule_def ffUnion.rep_eq)
fun outPorts where
"outPorts (Rule r) = outPortsRule r"
|"outPorts (Assumption r) = {|Reg r|}"
|"outPorts (Conclusion r) = {||}"
|"outPorts Helper = {| Reg anyP |}"
fun labelsIn where
"labelsIn _ p = a_conc p"
fun labelsOut where
"labelsOut _ (Reg p) = p"
| "labelsOut _ (Hyp h c) = h"
fun hyps where
"hyps (Rule r) (Hyp h a) = (if a |\<in>| f_antecedent r \<and> h |\<in>| a_hyps a then Some a else None)"
| "hyps _ _ = None"
fun local_vars :: "('form, 'rule) graph_node \<Rightarrow> ('form, 'var) in_port \<Rightarrow> 'var set" where
"local_vars _ a = a_fresh a"
sublocale Labeled_Signature nodes inPorts outPorts hyps labelsIn labelsOut
proof(standard,goal_cases)
case (1 n p1 p2)
thus ?case by(induction n p1 rule: hyps.induct) (auto split: if_splits)
qed
lemma hyps_for_conclusion[simp]: "hyps_for (Conclusion n) p = {||}"
using hyps_for_subset by auto
lemma hyps_for_Helper[simp]: "hyps_for Helper p = {||}"
using hyps_for_subset by auto
lemma hyps_for_Rule[simp]: "ip |\<in>| f_antecedent r \<Longrightarrow> hyps_for (Rule r) ip = (\<lambda> h. Hyp h ip) |`| a_hyps ip"
by (auto elim!: hyps.elims split: if_splits)
end
text \<open>Finally, a given proof graph solves the task at hand if all the given conclusions are present
as conclusion blocks in the graph.\<close>
locale Tasked_Proof_Graph =
Abstract_Task freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP antecedent consequent rules assumptions conclusions +
Scoped_Proof_Graph freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP inPorts outPorts nodeOf hyps nodes vertices labelsIn labelsOut vidx inst edges local_vars
for freshenLC :: "nat \<Rightarrow> 'var \<Rightarrow> 'var"
and renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> 'form \<Rightarrow> 'form"
and lconsts :: "'form \<Rightarrow> 'var set"
and closed :: "'form \<Rightarrow> bool"
and subst :: "'subst \<Rightarrow> 'form \<Rightarrow> 'form"
and subst_lconsts :: "'subst \<Rightarrow> 'var set"
and subst_renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> ('subst \<Rightarrow> 'subst)"
and anyP :: "'form"
and antecedent :: "'rule \<Rightarrow> ('form, 'var) antecedent list"
and consequent :: "'rule \<Rightarrow> 'form list"
and rules :: "'rule stream"
and assumptions :: "'form list"
and conclusions :: "'form list"
and vertices :: "'vertex fset"
and nodeOf :: "'vertex \<Rightarrow> ('form, 'rule) graph_node"
and edges :: "('vertex, 'form, 'var) edge' set"
and vidx :: "'vertex \<Rightarrow> nat"
and inst :: "'vertex \<Rightarrow> 'subst" +
assumes conclusions_present: "set (map Conclusion conclusions) \<subseteq> nodeOf ` fset vertices"
end
diff --git a/thys/Incredible_Proof_Machine/Build_Incredible_Tree.thy b/thys/Incredible_Proof_Machine/Build_Incredible_Tree.thy
--- a/thys/Incredible_Proof_Machine/Build_Incredible_Tree.thy
+++ b/thys/Incredible_Proof_Machine/Build_Incredible_Tree.thy
@@ -1,153 +1,153 @@
theory Build_Incredible_Tree
imports Incredible_Trees Natural_Deduction
begin
text \<open>
This theory constructs an incredible tree (with freshness checked only locally) from a natural
deduction tree.
\<close>
lemma image_eq_to_f:
assumes "f1 ` S1 = f2 ` S2"
obtains f where "\<And> x. x \<in> S2 \<Longrightarrow> f x \<in> S1 \<and> f1 (f x) = f2 x"
proof (atomize_elim)
from assms
have "\<forall>x. x \<in> S2 \<longrightarrow> (\<exists> y. y \<in> S1 \<and> f1 y = f2 x)" by (metis image_iff)
thus "\<exists>f. \<forall>x. x \<in> S2 \<longrightarrow> f x \<in> S1 \<and> f1 (f x) = f2 x" by metis
qed
context includes fset.lifting
begin
lemma fimage_eq_to_f:
assumes "f1 |`| S1 = f2 |`| S2"
obtains f where "\<And> x. x |\<in>| S2 \<Longrightarrow> f x |\<in>| S1 \<and> f1 (f x) = f2 x"
using assms apply transfer using image_eq_to_f by metis
end
context Abstract_Task
begin
lemma build_local_iwf:
fixes t :: "('form entailment \<times> ('rule \<times> 'form) NatRule) tree"
assumes "tfinite t"
assumes "wf t"
shows "\<exists> it. local_iwf it (fst (root t))"
using assms
proof(induction)
case (tfinite t)
from \<open>wf t\<close>
have "snd (root t) \<in> R" using wf.simps by blast
from \<open>wf t\<close>
have "eff (snd (root t)) (fst (root t)) ((fst \<circ> root) |`| cont t)" using wf.simps by blast
from \<open>wf t\<close>
have "\<And> t'. t' |\<in>| cont t \<Longrightarrow> wf t'" using wf.simps by blast
hence IH: "\<And> \<Gamma>' t'. t' |\<in>| cont t \<Longrightarrow> (\<exists>it'. local_iwf it' (fst (root t')))" using tfinite(2) by blast
then obtain its where its: "\<And> t'. t' |\<in>| cont t \<Longrightarrow> local_iwf (its t') (fst (root t'))" by metis
from \<open>eff _ _ _\<close>
show ?case
proof(cases rule: eff.cases[case_names Axiom NatRule Cut])
case (Axiom c \<Gamma>)
show ?thesis
proof (cases "c |\<in>| ass_forms")
case True (* Global assumption *)
then have "c \<in> set assumptions" by (auto simp add: ass_forms_def)
let "?it" = "INode (Assumption c) c undefined undefined [] :: ('form, 'rule, 'subst, 'var) itree"
from \<open>c \<in> set assumptions\<close>
have "local_iwf ?it (\<Gamma> \<turnstile> c)"
by (auto intro!: iwf local_fresh_check.intros)
thus ?thesis unfolding Axiom..
next
case False
obtain s where "subst s anyP = c" by atomize_elim (rule anyP_is_any)
hence [simp]: "subst s (freshen undefined anyP) = c" by (simp add: lconsts_anyP freshen_closed)
let "?it" = "HNode undefined s [] :: ('form, 'rule, 'subst, 'var) itree"
from \<open>c |\<in>| \<Gamma>\<close> False
have "local_iwf ?it (\<Gamma> \<turnstile> c)" by (auto intro: iwfH)
thus ?thesis unfolding Axiom..
qed
next
case (NatRule rule c ants \<Gamma> i s)
from \<open>natEff_Inst rule c ants\<close>
have "snd rule = c" and [simp]: "ants = f_antecedent (fst rule)" and "c \<in> set (consequent (fst rule))"
by (auto simp add: natEff_Inst.simps)
from \<open>(fst \<circ> root) |`| cont t = (\<lambda>ant. (\<lambda>p. subst s (freshen i p)) |`| a_hyps ant |\<union>| \<Gamma> \<turnstile> subst s (freshen i (a_conc ant))) |`| ants\<close>
obtain to_t where "\<And> ant. ant |\<in>| ants \<Longrightarrow> to_t ant |\<in>| cont t \<and> (fst \<circ> root) (to_t ant) = ((\<lambda>p. subst s (freshen i p)) |`| a_hyps ant |\<union>| \<Gamma> \<turnstile> subst s (freshen i (a_conc ant)))"
by (rule fimage_eq_to_f) (rule that)
hence to_t_in_cont: "\<And> ant. ant |\<in>| ants \<Longrightarrow> to_t ant |\<in>| cont t"
and to_t_root: "\<And> ant. ant |\<in>| ants \<Longrightarrow> fst (root (to_t ant)) = ((\<lambda>p. subst s (freshen i p)) |`| a_hyps ant |\<union>| \<Gamma> \<turnstile> subst s (freshen i (a_conc ant)))"
by auto
let ?ants' = "map (\<lambda> ant. its (to_t ant)) (antecedent (fst rule))"
let "?it" = "INode (Rule (fst rule)) c i s ?ants' :: ('form, 'rule, 'subst, 'var) itree"
from \<open>snd (root t) \<in> R\<close>
have "fst rule \<in> sset rules"
unfolding NatRule
by (auto simp add: stream.set_map n_rules_def no_empty_conclusions )
moreover
from \<open>c \<in> set (consequent (fst rule))\<close>
have "c |\<in>| f_consequent (fst rule)" by (simp add: f_consequent_def)
moreover
{ fix ant
assume "ant \<in> set (antecedent (fst rule))"
hence "ant |\<in>| ants" by (simp add: f_antecedent_def)
from its[OF to_t_in_cont[OF this]]
have "local_iwf (its (to_t ant)) (fst (root (to_t ant)))".
also have "fst (root (to_t ant)) =
((\<lambda>p. subst s (freshen i p)) |`| a_hyps ant |\<union>| \<Gamma> \<turnstile> subst s (freshen i (a_conc ant)))"
by (rule to_t_root[OF \<open>ant |\<in>| ants\<close>])
also have "\<dots> =
((\<lambda>h. subst s (freshen i (labelsOut (Rule (fst rule)) h))) |`| hyps_for (Rule (fst rule)) ant |\<union>| \<Gamma>
\<turnstile> subst s (freshen i (a_conc ant)))"
using \<open>ant |\<in>| ants\<close>
by auto
finally
have "local_iwf (its (to_t ant))
((\<lambda>h. subst s (freshen i (labelsOut (Rule (fst rule)) h))) |`| hyps_for (Rule (fst rule)) ant |\<union>|
\<Gamma> \<turnstile> subst s (freshen i (a_conc ant)))".
}
moreover
from NatRule(5,6)
have "local_fresh_check (Rule (fst rule)) i s (\<Gamma> \<turnstile> subst s (freshen i c))"
- by (fastforce intro!: local_fresh_check.intros simp add: all_local_vars_def fmember_iff_member_fset)
+ by (fastforce intro!: local_fresh_check.intros simp add: all_local_vars_def)
ultimately
have "local_iwf ?it ((\<Gamma> \<turnstile> subst s (freshen i c)))"
by (intro iwf ) (auto simp add: list_all2_map2 list_all2_same)
thus ?thesis unfolding NatRule..
next
case (Cut \<Gamma> con)
obtain s where "subst s anyP = con" by atomize_elim (rule anyP_is_any)
hence [simp]: "subst s (freshen undefined anyP) = con" by (simp add: lconsts_anyP freshen_closed)
from \<open>(fst \<circ> root) |`| cont t = {|\<Gamma> \<turnstile> con|}\<close>
obtain t' where "t' |\<in>| cont t" and [simp]: "fst (root t') = (\<Gamma> \<turnstile> con)"
by (cases "cont t") auto
from \<open>t' |\<in>| cont t\<close> obtain "it'" where "local_iwf it' (\<Gamma> \<turnstile> con)" using IH by force
let "?it" = "INode Helper anyP undefined s [it'] :: ('form, 'rule, 'subst, 'var) itree"
from \<open>local_iwf it' (\<Gamma> \<turnstile> con)\<close>
have "local_iwf ?it (\<Gamma> \<turnstile> con)" by (auto intro!: iwf local_fresh_check.intros)
thus ?thesis unfolding Cut..
qed
qed
definition to_it :: "('form entailment \<times> ('rule \<times> 'form) NatRule) tree \<Rightarrow> ('form,'rule,'subst,'var) itree" where
"to_it t = (SOME it. local_iwf it (fst (root t)))"
lemma iwf_to_it:
assumes "tfinite t" and "wf t"
shows "local_iwf (to_it t) (fst (root t))"
unfolding to_it_def using build_local_iwf[OF assms] by (rule someI2_ex)
end
end
diff --git a/thys/Incredible_Proof_Machine/Incredible_Completeness.thy b/thys/Incredible_Proof_Machine/Incredible_Completeness.thy
--- a/thys/Incredible_Proof_Machine/Incredible_Completeness.thy
+++ b/thys/Incredible_Proof_Machine/Incredible_Completeness.thy
@@ -1,689 +1,689 @@
theory Incredible_Completeness
imports Natural_Deduction Incredible_Deduction Build_Incredible_Tree
begin
text \<open>
This theory takes the tree produced in @{theory Incredible_Proof_Machine.Build_Incredible_Tree}, globalizes it using
@{term globalize}, and then builds the incredible proof graph out of it.
\<close>
type_synonym 'form vertex = "('form \<times> nat list)"
type_synonym ('form, 'var) edge'' = "('form vertex, 'form, 'var) edge'"
locale Solved_Task =
Abstract_Task freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP antecedent consequent rules assumptions conclusions
for freshenLC :: "nat \<Rightarrow> 'var \<Rightarrow> 'var"
and renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> 'form \<Rightarrow> 'form"
and lconsts :: "'form \<Rightarrow> 'var set"
and closed :: "'form \<Rightarrow> bool"
and subst :: "'subst \<Rightarrow> 'form \<Rightarrow> 'form"
and subst_lconsts :: "'subst \<Rightarrow> 'var set"
and subst_renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> ('subst \<Rightarrow> 'subst)"
and anyP :: "'form"
and antecedent :: "'rule \<Rightarrow> ('form, 'var) antecedent list"
and consequent :: "'rule \<Rightarrow> 'form list"
and rules :: "'rule stream"
and assumptions :: "'form list"
and conclusions :: "'form list" +
assumes solved: solved
begin
text \<open>Let us get our hand on concrete trees.\<close>
definition ts :: "'form \<Rightarrow> (('form entailment) \<times> ('rule \<times> 'form) NatRule) tree" where
"ts c = (SOME t. snd (fst (root t)) = c \<and> fst (fst (root t)) |\<subseteq>| ass_forms \<and> wf t \<and> tfinite t)"
lemma
assumes "c |\<in>| conc_forms"
shows ts_conc: "snd (fst (root (ts c))) = c"
and ts_context: "fst (fst (root (ts c))) |\<subseteq>| ass_forms"
and ts_wf: "wf (ts c)"
and ts_finite[simp]: "tfinite (ts c)"
unfolding atomize_conj conj_assoc ts_def
apply (rule someI_ex)
using solved assms
by (force simp add: solved_def)
abbreviation it' where
"it' c \<equiv> globalize [fidx conc_forms c, 0] (freshenLC v_away) (to_it (ts c))"
lemma iwf_it:
assumes "c \<in> set conclusions"
shows "plain_iwf (it' c) (fst (root (ts c)))"
using assms
apply (auto simp add: ts_conc conclusions_closed intro!: iwf_globalize' iwf_to_it ts_finite ts_wf)
by (meson assumptions_closed fset_mp mem_ass_forms mem_conc_forms ts_context)
definition vertices :: "'form vertex fset" where
"vertices = Abs_fset (Union ( set (map (\<lambda> c. insert (c, []) ((\<lambda> p. (c, 0 # p)) ` (it_paths (it' c)))) conclusions)))"
lemma mem_vertices: "v |\<in>| vertices \<longleftrightarrow> (fst v \<in> set conclusions \<and> (snd v = [] \<or> snd v \<in> ((#) 0) ` it_paths (it' (fst v))))"
- unfolding vertices_def fmember_iff_member_fset ffUnion.rep_eq
+ unfolding vertices_def ffUnion.rep_eq
by (cases v)(auto simp add: Abs_fset_inverse Bex_def )
lemma prefixeq_vertices: "(c,is) |\<in>| vertices \<Longrightarrow> prefix is' is \<Longrightarrow> (c, is') |\<in>| vertices"
by (cases is') (auto simp add: mem_vertices intro!: imageI elim: it_paths_prefix)
lemma none_vertices[simp]: "(c, []) |\<in>| vertices \<longleftrightarrow> c \<in> set conclusions"
by (simp add: mem_vertices)
lemma some_vertices[simp]: "(c, i#is) |\<in>| vertices \<longleftrightarrow> c \<in> set conclusions \<and> i = 0 \<and> is \<in> it_paths (it' c)"
by (auto simp add: mem_vertices)
lemma vertices_cases[consumes 1, case_names None Some]:
assumes "v |\<in>| vertices"
obtains c where "c \<in> set conclusions" and "v = (c, [])"
| c "is" where "c \<in> set conclusions" and "is \<in> it_paths (it' c)" and "v = (c, 0#is)"
using assms by (cases v; rename_tac "is"; case_tac "is"; auto)
lemma vertices_induct[consumes 1, case_names None Some]:
assumes "v |\<in>| vertices"
assumes "\<And> c. c \<in> set conclusions \<Longrightarrow> P (c, [])"
assumes "\<And> c is . c \<in> set conclusions \<Longrightarrow> is \<in> it_paths (it' c) \<Longrightarrow> P (c, 0#is)"
shows "P v"
using assms by (cases v; rename_tac "is"; case_tac "is"; auto)
fun nodeOf :: "'form vertex \<Rightarrow> ('form, 'rule) graph_node" where
"nodeOf (pf, []) = Conclusion pf"
| "nodeOf (pf, i#is) = iNodeOf (tree_at (it' pf) is)"
fun inst where
"inst (c,[]) = empty_subst"
|"inst (c, i#is) = iSubst (tree_at (it' c) is)"
lemma terminal_is_nil[simp]: "v |\<in>| vertices \<Longrightarrow> outPorts (nodeOf v) = {||} \<longleftrightarrow> snd v = []"
by (induction v rule: nodeOf.induct)
(auto elim: iNodeOf_outPorts[rotated] iwf_it)
sublocale Vertex_Graph nodes inPorts outPorts vertices nodeOf.
definition edge_from :: "'form \<Rightarrow> nat list => ('form vertex \<times> ('form,'var) out_port)" where
"edge_from c is = ((c, 0 # is), Reg (iOutPort (tree_at (it' c) is)))"
lemma fst_edge_from[simp]: "fst (edge_from c is) = (c, 0 # is)"
by (simp add: edge_from_def)
fun in_port_at :: "('form \<times> nat list) \<Rightarrow> nat \<Rightarrow> ('form,'var) in_port" where
"in_port_at (c, []) _ = plain_ant c"
| "in_port_at (c, _#is) i = inPorts' (iNodeOf (tree_at (it' c) is)) ! i"
definition edge_to :: "'form \<Rightarrow> nat list => ('form vertex \<times> ('form,'var) in_port)" where
"edge_to c is =
(case rev is of [] \<Rightarrow> ((c, []), in_port_at (c, []) 0)
| i#is \<Rightarrow> ((c, 0 # (rev is)), in_port_at (c, (0#rev is)) i))"
lemma edge_to_Nil[simp]: "edge_to c [] = ((c, []), plain_ant c)"
by (simp add: edge_to_def)
lemma edge_to_Snoc[simp]: "edge_to c (is@[i]) = ((c, 0 # is), in_port_at ((c, 0 # is)) i)"
by (simp add: edge_to_def)
definition edge_at :: "'form \<Rightarrow> nat list => ('form, 'var) edge''" where
"edge_at c is = (edge_from c is, edge_to c is)"
lemma fst_edge_at[simp]: "fst (edge_at c is) = edge_from c is" by (simp add: edge_at_def)
lemma snd_edge_at[simp]: "snd (edge_at c is) = edge_to c is" by (simp add: edge_at_def)
lemma hyps_exist':
assumes "c \<in> set conclusions"
assumes "is \<in> it_paths (it' c)"
assumes "tree_at (it' c) is = (HNode i s ants)"
shows "subst s (freshen i anyP) \<in> hyps_along (it' c) is"
proof-
from assms(1)
have "plain_iwf (it' c) (fst (root (ts c)))" by (rule iwf_it)
moreover
note assms(2,3)
moreover
have "fst (fst (root (ts c))) |\<subseteq>| ass_forms"
by (simp add: assms(1) ts_context)
ultimately
show ?thesis by (rule iwf_hyps_exist)
qed
definition hyp_edge_to :: "'form \<Rightarrow> nat list => ('form vertex \<times> ('form,'var) in_port)" where
"hyp_edge_to c is = ((c, 0 # is), plain_ant anyP)"
(* TODO: Replace n and s by "subst s (freshen n anyP)" *)
definition hyp_edge_from :: "'form \<Rightarrow> nat list => nat \<Rightarrow> 'subst \<Rightarrow> ('form vertex \<times> ('form,'var) out_port)" where
"hyp_edge_from c is n s =
((c, 0 # hyp_port_path_for (it' c) is (subst s (freshen n anyP))),
hyp_port_h_for (it' c) is (subst s (freshen n anyP)))"
definition hyp_edge_at :: "'form \<Rightarrow> nat list => nat \<Rightarrow> 'subst \<Rightarrow> ('form, 'var) edge''" where
"hyp_edge_at c is n s = (hyp_edge_from c is n s, hyp_edge_to c is)"
lemma fst_hyp_edge_at[simp]:
"fst (hyp_edge_at c is n s) = hyp_edge_from c is n s" by (simp add:hyp_edge_at_def)
lemma snd_hyp_edge_at[simp]:
"snd (hyp_edge_at c is n s) = hyp_edge_to c is" by (simp add:hyp_edge_at_def)
inductive_set edges where
regular_edge: "c \<in> set conclusions \<Longrightarrow> is \<in> it_paths (it' c) \<Longrightarrow> edge_at c is \<in> edges"
| hyp_edge: "c \<in> set conclusions \<Longrightarrow> is \<in> it_paths (it' c) \<Longrightarrow> tree_at (it' c) is = HNode n s ants \<Longrightarrow> hyp_edge_at c is n s \<in> edges"
sublocale Pre_Port_Graph nodes inPorts outPorts vertices nodeOf edges.
lemma edge_from_valid_out_port:
assumes "p \<in> it_paths (it' c)"
assumes "c \<in> set conclusions"
shows "valid_out_port (edge_from c p)"
using assms
by (auto simp add: edge_from_def intro: iwf_outPort iwf_it)
lemma edge_to_valid_in_port:
assumes "p \<in> it_paths (it' c)"
assumes "c \<in> set conclusions"
shows "valid_in_port (edge_to c p)"
using assms
apply (auto simp add: edge_to_def inPorts_fset_of split: list.split elim!: it_paths_SnocE)
apply (rule nth_mem)
apply (drule (1) iwf_length_inPorts[OF iwf_it])
apply auto
done
lemma hyp_edge_from_valid_out_port:
assumes "is \<in> it_paths (it' c)"
assumes "c \<in> set conclusions"
assumes "tree_at (it' c) is = HNode n s ants"
shows "valid_out_port (hyp_edge_from c is n s)"
using assms
by(auto simp add: hyp_edge_from_def intro: hyp_port_outPort it_paths_strict_prefix hyp_port_strict_prefix hyps_exist')
lemma hyp_edge_to_valid_in_port:
assumes "is \<in> it_paths (it' c)"
assumes "c \<in> set conclusions"
assumes "tree_at (it' c) is = HNode n s ants"
shows "valid_in_port (hyp_edge_to c is)"
using assms by (auto simp add: hyp_edge_to_def)
inductive scope' :: "'form vertex \<Rightarrow> ('form,'var) in_port \<Rightarrow> 'form \<times> nat list \<Rightarrow> bool" where
"c \<in> set conclusions \<Longrightarrow>
is' \<in> ((#) 0) ` it_paths (it' c) \<Longrightarrow>
prefix (is@[i]) is' \<Longrightarrow>
ip = in_port_at (c,is) i \<Longrightarrow>
scope' (c, is) ip (c, is')"
inductive_simps scope_simp: "scope' v i v'"
inductive_cases scope_cases: "scope' v i v'"
lemma scope_valid:
"scope' v i v' \<Longrightarrow> v' |\<in>| vertices"
by (auto elim: scope_cases)
lemma scope_valid_inport:
"v' |\<in>| vertices \<Longrightarrow> scope' v ip v' \<longleftrightarrow> (\<exists> i. fst v = fst v' \<and> prefix (snd v@[i]) (snd v') \<and> ip = in_port_at v i)"
by (cases v; cases v') (auto simp add: scope'.simps mem_vertices)
definition terminal_path_from :: "'form \<Rightarrow> nat list => ('form, 'var) edge'' list" where
"terminal_path_from c is = map (edge_at c) (rev (prefixes is))"
lemma terminal_path_from_Nil[simp]:
"terminal_path_from c [] = [edge_at c []]"
by (simp add: terminal_path_from_def)
lemma terminal_path_from_Snoc[simp]:
"terminal_path_from c (is @ [i]) = edge_at c (is@[i]) # terminal_path_from c is"
by (simp add: terminal_path_from_def)
lemma path_terminal_path_from:
"c \<in> set conclusions \<Longrightarrow>
is \<in> it_paths (it' c) \<Longrightarrow>
path (c, 0 # is) (c, []) (terminal_path_from c is)"
by (induction "is" rule: rev_induct)
(auto simp add: path_cons_simp intro!: regular_edge elim: it_paths_SnocE)
lemma edge_step:
assumes "(((a, b), ba), ((aa, bb), bc)) \<in> edges"
obtains
i where "a = aa" and "b = bb@[i]" and "bc = in_port_at (aa,bb) i" and "hyps (nodeOf (a, b)) ba = None"
| i where "a = aa" and "prefix (b@[i]) bb" and "hyps (nodeOf (a, b)) ba = Some (in_port_at (a,b) i)"
using assms
proof(cases rule: edges.cases[consumes 1, case_names Reg Hyp])
case (Reg c "is")
then obtain i where "a = aa" and "b = bb@[i]" and "bc = in_port_at (aa,bb) i" and "hyps (nodeOf (a, b)) ba = None"
by (auto elim!: edges.cases simp add: edge_at_def edge_from_def edge_to_def split: list.split list.split_asm)
thus thesis by (rule that)
next
case (Hyp c "is" n s)
let ?i = "hyp_port_i_for (it' c) is (subst s (freshen n anyP))"
from Hyp have "a = aa" and "prefix (b@[?i]) bb" and
"hyps (nodeOf (a, b)) ba = Some (in_port_at (a,b) ?i)"
by (auto simp add: edge_at_def edge_from_def edge_to_def hyp_edge_at_def hyp_edge_to_def hyp_edge_from_def
intro: hyp_port_prefix hyps_exist' hyp_port_hyps)
thus thesis by (rule that)
qed
lemma path_has_prefixes:
assumes "path v v' pth"
assumes "snd v' = []"
assumes "prefix (is' @ [i]) (snd v)"
shows "((fst v, is'), (in_port_at (fst v, is') i)) \<in> snd ` set pth"
using assms
by (induction rule: path.induct)(auto elim!: edge_step dest: prefix_snocD)
lemma in_scope: "valid_in_port (v', p') \<Longrightarrow> v \<in> scope (v', p') \<longleftrightarrow> scope' v' p' v"
proof
assume "v \<in> scope (v', p')"
hence "v |\<in>| vertices" and "\<And> pth t. path v t pth \<Longrightarrow> terminal_vertex t \<Longrightarrow> (v', p') \<in> snd ` set pth"
by (auto simp add: scope.simps)
from this
show "scope' v' p' v"
proof (induction rule: vertices_induct)
case (None c)
from None(2)[of "(c, [])" "[]", simplified, OF None(1)]
have False.
thus "scope' v' p' (c, [])"..
next
case (Some c "is")
from \<open>c \<in> set conclusions\<close> \<open>is \<in> it_paths (it' c)\<close>
have "path (c, 0#is) (c, []) (terminal_path_from c is)"
by (rule path_terminal_path_from)
moreover
from \<open>c \<in> set conclusions\<close>
have "terminal_vertex (c, [])" by simp
ultimately
have "(v', p') \<in> snd ` set (terminal_path_from c is)"
by (rule Some(3))
hence "(v',p') \<in> set (map (edge_to c) (prefixes is))"
unfolding terminal_path_from_def by auto
then obtain is' where "prefix is' is" and "(v',p') = edge_to c is'"
by auto
show "scope' v' p' (c, 0#is)"
proof(cases "is'" rule: rev_cases)
case Nil
with \<open>(v',p') = edge_to c is'\<close>
have "v' = (c, [])" and "p' = plain_ant c"
by (auto simp add: edge_to_def)
with \<open>c \<in> set conclusions\<close> \<open>is \<in> it_paths (it' c)\<close>
show ?thesis by (auto intro!: scope'.intros)
next
case (snoc is'' i)
with \<open>(v',p') = edge_to c is'\<close>
have "v' = (c, 0 # is'')" and "p' = in_port_at v' i"
by (auto simp add: edge_to_def)
with \<open>c \<in> set conclusions\<close> \<open>is \<in> it_paths (it' c)\<close> \<open>prefix is' is\<close>[unfolded snoc]
show ?thesis
by (auto intro!: scope'.intros)
qed
qed
next
assume "valid_in_port (v', p')"
assume "scope' v' p' v"
then obtain c is' i "is" where
"v' = (c, is')" and "v = (c, is)" and "c \<in> set conclusions" and
"p' = in_port_at v' i" and
"is \<in> (#) 0 ` it_paths (it' c)" and "prefix (is' @ [i]) is"
by (auto simp add: scope'.simps)
from \<open>scope' v' p' v\<close>
have "(c, is) |\<in>| vertices" unfolding \<open>v = _\<close> by (rule scope_valid)
hence "(c, is) \<in> scope ((c, is'), p')"
proof(rule scope.intros)
fix pth t
assume "path (c,is) t pth"
assume "terminal_vertex t"
hence "snd t = []" by auto
from path_has_prefixes[OF \<open>path (c,is) t pth\<close> \<open>snd t = []\<close>, simplified, OF \<open>prefix (is' @ [i]) is\<close>]
show "((c, is'), p') \<in> snd ` set pth" unfolding \<open>p' = _ \<close> \<open>v' = _ \<close>.
qed
thus "v \<in> scope (v', p')" using \<open>v =_\<close> \<open>v' = _\<close> by simp
qed
sublocale Port_Graph nodes inPorts outPorts vertices nodeOf edges
proof
show "nodeOf ` fset vertices \<subseteq> sset nodes"
- apply (auto simp add: fmember_iff_member_fset[symmetric] mem_vertices)
+ apply (auto simp add: mem_vertices)
apply (auto simp add: stream.set_map dest: iNodeOf_tree_at[OF iwf_it])
done
next
have "\<forall> e \<in> edges. valid_out_port (fst e) \<and> valid_in_port (snd e)"
by (auto elim!: edges.cases simp add: edge_at_def
dest: edge_from_valid_out_port edge_to_valid_in_port
dest: hyp_edge_from_valid_out_port hyp_edge_to_valid_in_port)
thus "\<forall>(ps1, ps2)\<in>edges. valid_out_port ps1 \<and> valid_in_port ps2" by auto
qed
sublocale Scoped_Graph nodes inPorts outPorts vertices nodeOf edges hyps..
lemma hyps_free_path_length:
assumes "path v v' pth"
assumes "hyps_free pth"
shows "length pth + length (snd v') = length (snd v)"
using assms by induction (auto elim!: edge_step )
fun vidx :: "'form vertex \<Rightarrow> nat" where
"vidx (c, []) = isidx [fidx conc_forms c]"
|"vidx (c, _#is) = iAnnot (tree_at (it' c) is)"
lemma my_vidx_inj: "inj_on vidx (fset vertices)"
by (rule inj_onI)
- (auto simp add: mem_vertices[unfolded fmember_iff_member_fset] iAnnot_globalize simp del: iAnnot.simps)
+ (auto simp add: mem_vertices iAnnot_globalize simp del: iAnnot.simps)
lemma vidx_not_v_away[simp]: "v |\<in>| vertices \<Longrightarrow> vidx v \<noteq> v_away"
by (cases v rule:vidx.cases) (auto simp add: iAnnot_globalize simp del: iAnnot.simps)
sublocale Instantiation inPorts outPorts nodeOf hyps nodes edges vertices labelsIn labelsOut freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP vidx inst
proof
show "inj_on vidx (fset vertices)" by (rule my_vidx_inj)
qed
sublocale Well_Scoped_Graph nodes inPorts outPorts vertices nodeOf edges hyps
proof
fix v\<^sub>1 p\<^sub>1 v\<^sub>2 p\<^sub>2 p'
assume assms: "((v\<^sub>1, p\<^sub>1), (v\<^sub>2, p\<^sub>2)) \<in> edges" "hyps (nodeOf v\<^sub>1) p\<^sub>1 = Some p'"
from assms(1) hyps_correct[OF assms(2)]
have "valid_out_port (v\<^sub>1, p\<^sub>1)" and "valid_in_port (v\<^sub>2, p\<^sub>2)" and "valid_in_port (v\<^sub>1, p')" and "v\<^sub>2 |\<in>| vertices"
using valid_edges by auto
from assms
have "\<exists> i. fst v\<^sub>1 = fst v\<^sub>2 \<and> prefix (snd v\<^sub>1@[i]) (snd v\<^sub>2) \<and> p' = in_port_at v\<^sub>1 i"
by (cases v\<^sub>1; cases v\<^sub>2; auto elim!: edge_step)
hence "scope' v\<^sub>1 p' v\<^sub>2"
unfolding scope_valid_inport[OF \<open>v\<^sub>2 |\<in>| vertices\<close>].
hence "v\<^sub>2 \<in> scope (v\<^sub>1, p')"
unfolding in_scope[OF \<open>valid_in_port (v\<^sub>1, p')\<close>].
thus "(v\<^sub>2, p\<^sub>2) = (v\<^sub>1, p') \<or> v\<^sub>2 \<in> scope (v\<^sub>1, p')" ..
qed
sublocale Acyclic_Graph nodes inPorts outPorts vertices nodeOf edges hyps
proof
fix v pth
assume "path v v pth" and "hyps_free pth"
from hyps_free_path_length[OF this]
show "pth = []" by simp
qed
sublocale Saturated_Graph nodes inPorts outPorts vertices nodeOf edges
proof
fix v p
assume "valid_in_port (v, p)"
thus "\<exists>e\<in>edges. snd e = (v, p)"
proof(induction v)
fix c cis
assume "valid_in_port ((c, cis), p)"
hence "c \<in> set conclusions" by (auto simp add: mem_vertices)
show "\<exists>e\<in>edges. snd e = ((c, cis), p)"
proof(cases cis)
case Nil
with \<open>valid_in_port ((c, cis), p)\<close>
have [simp]: "p = plain_ant c" by simp
have "[] \<in> it_paths (it' c)" by simp
with \<open>c \<in> set conclusions\<close>
have "edge_at c [] \<in> edges" by (rule regular_edge)
moreover
have "snd (edge_at c []) = ((c, []), plain_ant c)"
by (simp add: edge_to_def)
ultimately
show ?thesis by (auto simp add: Nil simp del: snd_edge_at)
next
case (Cons c' "is")
with \<open>valid_in_port ((c, cis), p)\<close>
have [simp]: "c' = 0" and "is \<in> it_paths (it' c)"
and "p |\<in>| inPorts (iNodeOf (tree_at (it' c) is))" by auto
from this(3) obtain i where
"i < length (inPorts' (iNodeOf (tree_at (it' c) is)))" and
"p = inPorts' (iNodeOf (tree_at (it' c) is)) ! i"
by (auto simp add: inPorts_fset_of in_set_conv_nth)
show ?thesis
proof (cases "tree_at (it' c) is")
case [simp]: (RNode r ants)
show ?thesis
proof(cases r)
case I
hence "\<not> isHNode (tree_at (it' c) is)" by simp
from iwf_length_inPorts_not_HNode[OF iwf_it[OF \<open>c \<in> set conclusions\<close>] \<open>is \<in> it_paths (it' c)\<close> this]
\<open>i < length (inPorts' (iNodeOf (tree_at (it' c) is)))\<close>
have "i < length (children (tree_at (it' c) is))" by simp
with \<open>is \<in> it_paths (it' c)\<close>
have "is@[i] \<in> it_paths (it' c)" by (rule it_path_SnocI)
from \<open>c \<in> set conclusions\<close> this
have "edge_at c (is@[i]) \<in> edges" by (rule regular_edge)
moreover
have "snd (edge_at c (is@[i])) = ((c, 0 # is), inPorts' (iNodeOf (tree_at (it' c) is)) ! i)"
by (simp add: edge_to_def)
ultimately
show ?thesis by (auto simp add: Cons \<open>p = _\<close> simp del: snd_edge_at)
next
case (H n s)
hence "tree_at (it' c) is = HNode n s ants" by simp
from \<open>c \<in> set conclusions\<close> \<open>is \<in> it_paths (it' c)\<close> this
have "hyp_edge_at c is n s \<in> edges"..
moreover
from H \<open>p |\<in>| inPorts (iNodeOf (tree_at (it' c) is))\<close>
have [simp]: "p = plain_ant anyP" by simp
have "snd (hyp_edge_at c is n s) = ((c, 0 # is), p)"
by (simp add: hyp_edge_to_def)
ultimately
show ?thesis by (auto simp add: Cons simp del: snd_hyp_edge_at)
qed
qed
qed
qed
qed
sublocale Pruned_Port_Graph nodes inPorts outPorts vertices nodeOf edges
proof
fix v
assume "v |\<in>| vertices"
thus "\<exists>pth v'. path v v' pth \<and> terminal_vertex v'"
proof(induct rule: vertices_induct)
case (None c)
hence "terminal_vertex (c,[])" by simp
with path.intros(1)
show ?case by blast
next
case (Some c "is")
hence "path (c, 0 # is) (c, []) (terminal_path_from c is)"
by (rule path_terminal_path_from)
moreover
have "terminal_vertex (c,[])" using Some(1) by simp
ultimately
show ?case by blast
qed
qed
sublocale Well_Shaped_Graph nodes inPorts outPorts vertices nodeOf edges hyps..
sublocale sol:Solution inPorts outPorts nodeOf hyps nodes vertices labelsIn labelsOut freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP vidx inst edges
proof
fix v\<^sub>1 p\<^sub>1 v\<^sub>2 p\<^sub>2
assume "((v\<^sub>1, p\<^sub>1), (v\<^sub>2, p\<^sub>2)) \<in> edges"
thus "labelAtOut v\<^sub>1 p\<^sub>1 = labelAtIn v\<^sub>2 p\<^sub>2"
proof(cases rule:edges.cases)
case (regular_edge c "is")
from \<open>((v\<^sub>1, p\<^sub>1), v\<^sub>2, p\<^sub>2) = edge_at c is\<close>
have "(v\<^sub>1,p\<^sub>1) = edge_from c is" using fst_edge_at by (metis fst_conv)
hence [simp]: "v\<^sub>1 = (c, 0 # is)" by (simp add: edge_from_def)
show ?thesis
proof(cases "is" rule:rev_cases)
case Nil
let "?t'" = "it' c"
have "labelAtOut v\<^sub>1 p\<^sub>1 = subst (iSubst ?t') (freshen (vidx v\<^sub>1) (iOutPort ?t'))"
using regular_edge Nil by (simp add: labelAtOut_def edge_at_def edge_from_def)
also have "vidx v\<^sub>1 = iAnnot ?t'" by (simp add: Nil)
also have "subst (iSubst ?t') (freshen (iAnnot ?t') (iOutPort ?t')) = snd (fst (root (ts c)))"
unfolding iwf_subst_freshen_outPort[OF iwf_it[OF \<open>c \<in> set conclusions\<close>]]..
also have "\<dots> = c" using \<open>c \<in> set conclusions\<close> by (simp add: ts_conc)
also have "\<dots> = labelAtIn v\<^sub>2 p\<^sub>2"
using \<open>c \<in> set conclusions\<close> regular_edge Nil
by (simp add: labelAtIn_def edge_at_def freshen_closed conclusions_closed closed_no_lconsts)
finally show ?thesis.
next
case (snoc is' i)
let "?t1" = "tree_at (it' c) (is'@[i])"
let "?t2" = "tree_at (it' c) is'"
have "labelAtOut v\<^sub>1 p\<^sub>1 = subst (iSubst ?t1) (freshen (vidx v\<^sub>1) (iOutPort ?t1))"
using regular_edge snoc by (simp add: labelAtOut_def edge_at_def edge_from_def)
also have "vidx v\<^sub>1 = iAnnot ?t1" using snoc regular_edge(3) by simp
also have "subst (iSubst ?t1) (freshen (iAnnot ?t1) (iOutPort ?t1))
= subst (iSubst ?t2) (freshen (iAnnot ?t2) (a_conc (inPorts' (iNodeOf ?t2) ! i)))"
by (rule iwf_edge_match[OF iwf_it[OF \<open>c \<in> set conclusions\<close>] \<open>is \<in> it_paths (it' c)\<close>[unfolded snoc]])
also have "iAnnot ?t2 = vidx (c, 0 # is')" by simp
also have "subst (iSubst ?t2) (freshen (vidx (c, 0 # is')) (a_conc (inPorts' (iNodeOf ?t2) ! i))) = labelAtIn v\<^sub>2 p\<^sub>2"
using regular_edge snoc by (simp add: labelAtIn_def edge_at_def)
finally show ?thesis.
qed
next
case (hyp_edge c "is" n s ants)
let ?f = "subst s (freshen n anyP)"
let ?h = "hyp_port_h_for (it' c) is ?f"
let ?his = "hyp_port_path_for (it' c) is ?f"
let "?t1" = "tree_at (it' c) ?his"
let "?t2" = "tree_at (it' c) is"
from \<open>c \<in> set conclusions\<close> \<open>is \<in> it_paths (it' c)\<close> \<open>tree_at (it' c) is = HNode n s ants\<close>
have "?f \<in> hyps_along (it' c) is"
by (rule hyps_exist')
from \<open>((v\<^sub>1, p\<^sub>1), v\<^sub>2, p\<^sub>2) = hyp_edge_at c is n s\<close>
have "(v\<^sub>1,p\<^sub>1) = hyp_edge_from c is n s" using fst_hyp_edge_at by (metis fst_conv)
hence [simp]: "v\<^sub>1 = (c, 0 # ?his)" by (simp add: hyp_edge_from_def)
have "labelAtOut v\<^sub>1 p\<^sub>1 = subst (iSubst ?t1) (freshen (vidx v\<^sub>1) (labelsOut (iNodeOf ?t1) ?h))"
using hyp_edge by (simp add: hyp_edge_at_def hyp_edge_from_def labelAtOut_def)
also have "vidx v\<^sub>1 = iAnnot ?t1" by simp
also have "subst (iSubst ?t1) (freshen (iAnnot ?t1) (labelsOut (iNodeOf ?t1) ?h)) = ?f" using \<open>?f \<in> hyps_along (it' c) is\<close> by (rule local.hyp_port_eq[symmetric])
also have "\<dots> = subst (iSubst ?t2) (freshen (iAnnot ?t2) anyP)" using hyp_edge by simp
also have "subst (iSubst ?t2) (freshen (iAnnot ?t2) anyP) = labelAtIn v\<^sub>2 p\<^sub>2"
using hyp_edge by (simp add: labelAtIn_def hyp_edge_at_def hyp_edge_to_def)
finally show ?thesis.
qed
qed
lemma node_disjoint_fresh_vars:
assumes "n \<in> sset nodes"
assumes "i < length (inPorts' n)"
assumes "i' < length (inPorts' n)"
shows "a_fresh (inPorts' n ! i) \<inter> a_fresh (inPorts' n ! i') = {} \<or> i = i'"
using assms no_multiple_local_consts
by (fastforce simp add: nodes_def stream.set_map)
sublocale Well_Scoped_Instantiation freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP inPorts outPorts nodeOf hyps nodes vertices labelsIn labelsOut vidx inst edges local_vars
proof
fix v p var v'
assume "valid_in_port (v, p)"
hence "v |\<in>| vertices" by simp
obtain c "is" where "v = (c,is)" by (cases v, auto)
from \<open>valid_in_port (v, p)\<close> \<open>v= _\<close>
have "(c,is) |\<in>| vertices" and "p |\<in>| inPorts (nodeOf (c, is))" by simp_all
hence "c \<in> set conclusions" by (simp add: mem_vertices)
from \<open>p |\<in>| _\<close> obtain i where
"i < length (inPorts' (nodeOf (c, is)))" and
"p = inPorts' (nodeOf (c, is)) ! i" by (auto simp add: inPorts_fset_of in_set_conv_nth)
hence "p = in_port_at (c, is) i" by (cases "is") auto
assume "v' |\<in>| vertices"
then obtain c' is' where "v' = (c',is')" by (cases v', auto)
assume "var \<in> local_vars (nodeOf v) p"
hence "var \<in> a_fresh p" by simp
assume "freshenLC (vidx v) var \<in> subst_lconsts (inst v')"
then obtain is'' where "is' = 0#is''" and "is'' \<in> it_paths (it' c')"
using \<open>v' |\<in>| vertices\<close>
by (cases is') (auto simp add: \<open>v'=_\<close>)
note \<open>freshenLC (vidx v) var \<in> subst_lconsts (inst v')\<close>
also
have "subst_lconsts (inst v') = subst_lconsts (iSubst (tree_at (it' c') is''))"
by (simp add: \<open>v'=_\<close> \<open>is'=_\<close>)
also
from \<open>is'' \<in> it_paths (it' c')\<close>
have "\<dots> \<subseteq> fresh_at_path (it' c') is'' \<union> range (freshenLC v_away)"
by (rule globalize_local_consts)
finally
have "freshenLC (vidx v) var \<in> fresh_at_path (it' c') is''"
using \<open>v |\<in>| vertices\<close> by auto
then obtain is''' where "prefix is''' is''" and "freshenLC (vidx v) var \<in> fresh_at (it' c') is'''"
unfolding fresh_at_path_def by auto
then obtain i' is'''' where "prefix (is''''@[i']) is''"
and "freshenLC (vidx v) var \<in> fresh_at (it' c') (is''''@[i'])"
using append_butlast_last_id[where xs = is''', symmetric]
apply (cases "is''' = []")
apply (auto simp del: fresh_at_snoc append_butlast_last_id)
apply metis
done
from \<open>is'' \<in> it_paths (it' c')\<close> \<open>prefix (is''''@[i']) is''\<close>
have "(is''''@[i']) \<in> it_paths (it' c')" by (rule it_paths_prefix)
hence "is'''' \<in> it_paths (it' c')" using append_prefixD it_paths_prefix by blast
from this \<open>freshenLC (vidx v) var \<in> fresh_at (it' c') (is''''@[i'])\<close>
have "c = c' \<and> is = 0 # is'''' \<and> var \<in> a_fresh (inPorts' (iNodeOf (tree_at (it' c') is'''')) ! i')"
unfolding fresh_at_def' using \<open>v |\<in>| vertices\<close> \<open>v' |\<in>| vertices\<close>
apply (cases "is")
apply (auto split: if_splits simp add: iAnnot_globalize it_paths_butlast \<open>v=_\<close> \<open>v'=_\<close> \<open>is'=_\<close> simp del: iAnnot.simps)
done
hence "c' = c" and "is = 0 # is''''" and "var \<in> a_fresh (inPorts' (iNodeOf (tree_at (it' c') is'''')) ! i')" by simp_all
from \<open>(is''''@[i']) \<in> it_paths (it' c')\<close>
have "i' < length (inPorts' (nodeOf (c, is)))"
using iwf_length_inPorts[OF iwf_it[OF \<open>c \<in> set conclusions\<close>]]
by (auto elim!: it_paths_SnocE simp add: \<open>is=_\<close> \<open>c' = _\<close> order.strict_trans2)
have "nodeOf (c, is) \<in> sset nodes"
unfolding \<open>is = _\<close> \<open>c' = _\<close> nodeOf.simps
by (rule iNodeOf_tree_at[OF iwf_it[OF \<open>c \<in> set conclusions\<close>] \<open>is'''' \<in> it_paths (it' c')\<close>[unfolded \<open>c' = _\<close>]])
from \<open>var \<in> a_fresh (inPorts' (iNodeOf (tree_at (it' c') is'''')) ! i')\<close>
\<open>var \<in> a_fresh p\<close> \<open>p = inPorts' (nodeOf (c, is)) ! i\<close>
node_disjoint_fresh_vars[OF
\<open>nodeOf (c, is) \<in> sset nodes\<close>
\<open>i < length (inPorts' (nodeOf (c, is)))\<close> \<open>i' < length (inPorts' (nodeOf (c, is)))\<close>]
have "i' = i" by (auto simp add: \<open>is=_\<close> \<open>c'=c\<close>)
from \<open>prefix (is''''@[i']) is''\<close>
have "prefix (is @ [i']) is'" by (simp add: \<open>is'=_\<close> \<open>is=_\<close>)
from \<open>c \<in> set conclusions\<close> \<open>is'' \<in> it_paths (it' c')\<close> \<open>prefix (is @ [i']) is'\<close>
\<open>p = in_port_at (c, is) i\<close>
have "scope' v p v'"
unfolding \<open>v=_\<close> \<open>v'=_\<close> \<open>c' = _\<close> \<open>is' = _\<close> \<open>i'=_\<close> by (auto intro: scope'.intros)
thus "v' \<in> scope (v, p)" using \<open>valid_in_port (v, p)\<close> by (simp add: in_scope)
qed
sublocale Scoped_Proof_Graph freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP inPorts outPorts nodeOf hyps nodes vertices labelsIn labelsOut vidx inst edges local_vars..
(* interpretation of @{term Tasked_Proof_Graph} has to be named to avoid name clashes in @{term Abstract_Task}. *)
sublocale tpg:Tasked_Proof_Graph freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP antecedent consequent rules assumptions conclusions
vertices nodeOf edges vidx inst
proof
show "set (map Conclusion conclusions) \<subseteq> nodeOf ` fset vertices"
proof-
{
fix c
assume "c \<in> set conclusions"
hence "(c, []) |\<in>| vertices" by simp
hence "nodeOf (c, []) \<in> nodeOf ` fset vertices"
- unfolding fmember_iff_member_fset by (rule imageI)
+ by (rule imageI)
hence "Conclusion c \<in> nodeOf ` fset vertices" by simp
} thus ?thesis by auto
qed
qed
end
end
diff --git a/thys/Incredible_Proof_Machine/Incredible_Correctness.thy b/thys/Incredible_Proof_Machine/Incredible_Correctness.thy
--- a/thys/Incredible_Proof_Machine/Incredible_Correctness.thy
+++ b/thys/Incredible_Proof_Machine/Incredible_Correctness.thy
@@ -1,554 +1,554 @@
theory Incredible_Correctness
imports
Abstract_Rules_To_Incredible
Natural_Deduction
begin
text \<open>
In this theory, we prove that if we have a graph that proves a given abstract task (which is
represented as the context @{term Tasked_Proof_Graph}), then we can prove @{term solved}.
\<close>
context Tasked_Proof_Graph
begin
definition adjacentTo :: "'vertex \<Rightarrow> ('form, 'var) in_port \<Rightarrow> ('vertex \<times> ('form, 'var) out_port)" where
"adjacentTo v p = (SOME ps. (ps, (v,p)) \<in> edges)"
fun isReg where
"isReg v p = (case p of Hyp h c \<Rightarrow> False | Reg c \<Rightarrow>
(case nodeOf v of
Conclusion a \<Rightarrow> False
| Assumption a \<Rightarrow> False
| Rule r \<Rightarrow> True
| Helper \<Rightarrow> True
))"
fun toNatRule where
"toNatRule v p = (case p of Hyp h c \<Rightarrow> Axiom | Reg c \<Rightarrow>
(case nodeOf v of
Conclusion a \<Rightarrow> Axiom \<comment> \<open>a lie\<close>
| Assumption a \<Rightarrow> Axiom
| Rule r \<Rightarrow> NatRule (r,c)
| Helper \<Rightarrow> Cut
))"
inductive_set global_assms' :: "'var itself \<Rightarrow> 'form set" for i where
"v |\<in>| vertices \<Longrightarrow> nodeOf v = Assumption p \<Longrightarrow> labelAtOut v (Reg p) \<in> global_assms' i"
lemma finite_global_assms': "finite (global_assms' i)"
proof-
have "finite (fset vertices)" by (rule finite_fset)
moreover
have "global_assms' i \<subseteq> (\<lambda> v. case nodeOf v of Assumption p \<Rightarrow> labelAtOut v (Reg p)) ` fset vertices"
- by (force simp add: global_assms'.simps fmember_iff_member_fset image_iff )
+ by (force simp add: global_assms'.simps image_iff )
ultimately
show ?thesis by (rule finite_surj)
qed
context includes fset.lifting
begin
lift_definition global_assms :: "'var itself \<Rightarrow> 'form fset" is global_assms' by (rule finite_global_assms')
lemmas global_assmsI = global_assms'.intros[Transfer.transferred]
lemmas global_assms_simps = global_assms'.simps[Transfer.transferred]
end
fun extra_assms :: "('vertex \<times> ('form, 'var) in_port) \<Rightarrow> 'form fset" where
"extra_assms (v, p) = (\<lambda> p. labelAtOut v p) |`| hyps_for (nodeOf v) p"
fun hyps_along :: "('vertex, 'form, 'var) edge' list \<Rightarrow> 'form fset" where
"hyps_along pth = ffUnion (extra_assms |`| snd |`| fset_from_list pth) |\<union>| global_assms TYPE('var)"
lemma hyps_alongE[consumes 1, case_names Hyp Assumption]:
assumes "f |\<in>| hyps_along pth"
obtains v p h where "(v,p) \<in> snd ` set pth" and "f = labelAtOut v h " and "h |\<in>| hyps_for (nodeOf v) p"
| v pf where "v |\<in>| vertices" and "nodeOf v = Assumption pf" "f = labelAtOut v (Reg pf)"
using assms
- apply (auto simp add: fmember_iff_member_fset ffUnion.rep_eq global_assms_simps[unfolded fmember_iff_member_fset])
+ apply (auto simp add: ffUnion.rep_eq global_assms_simps)
apply (metis image_iff snd_conv)
done
text \<open>Here we build the natural deduction tree, by walking the graph.\<close>
primcorec tree :: "'vertex \<Rightarrow> ('form, 'var) in_port \<Rightarrow> ('vertex, 'form, 'var) edge' list \<Rightarrow> (('form entailment), ('rule \<times> 'form) NatRule) dtree" where
"root (tree v p pth) =
((hyps_along ((adjacentTo v p,(v,p))#pth) \<turnstile> labelAtIn v p),
(case adjacentTo v p of (v', p') \<Rightarrow> toNatRule v' p'
))"
| "cont (tree v p pth) =
(case adjacentTo v p of (v', p') \<Rightarrow>
(if isReg v' p' then ((\<lambda> p''. tree v' p'' ((adjacentTo v p,(v,p))#pth)) |`| inPorts (nodeOf v')) else {||}
))"
lemma fst_root_tree[simp]: "fst (root (tree v p pth)) = (hyps_along ((adjacentTo v p,(v,p))#pth) \<turnstile> labelAtIn v p)" by simp
lemma out_port_cases[consumes 1, case_names Assumption Hyp Rule Helper]:
assumes "p |\<in>| outPorts n"
obtains
a where "n = Assumption a" and "p = Reg a"
| r h c where "n = Rule r" and "p = Hyp h c"
| r f where "n = Rule r" and "p = Reg f"
| "n = Helper" and "p = Reg anyP"
using assms by (atomize_elim, cases p; cases n) auto
lemma hyps_for_fimage: "hyps_for (Rule r) x = (if x |\<in>| f_antecedent r then (\<lambda> f. Hyp f x) |`| (a_hyps x) else {||})"
apply (rule fset_eqI)
apply (rename_tac p')
apply (case_tac p')
apply (auto simp add: split: if_splits out_port.splits)
done
text \<open>Now we prove that the thus produced tree is well-formed.\<close>
theorem wf_tree:
assumes "valid_in_port (v,p)"
assumes "terminal_path v t pth"
shows "wf (tree v p pth)"
using assms
proof (coinduction arbitrary: v p pth)
case (wf v p pth)
let ?t = "tree v p pth"
from saturated[OF wf(1)]
obtain v' p'
where e:"((v',p'),(v,p)) \<in> edges" and [simp]: "adjacentTo v p = (v',p')"
by (auto simp add: adjacentTo_def, metis (no_types, lifting) eq_fst_iff tfl_some)
let ?e = "((v',p'),(v,p))"
let ?pth' = "?e#pth"
let ?\<Gamma> = "hyps_along ?pth'"
let ?l = "labelAtIn v p"
from e valid_edges have "v' |\<in>| vertices" and "p' |\<in>| outPorts (nodeOf v')" by auto
- hence "nodeOf v' \<in> sset nodes" using valid_nodes by (meson image_eqI notin_fset subsetD)
+ hence "nodeOf v' \<in> sset nodes" using valid_nodes by (meson image_eqI subsetD)
from \<open>?e \<in> edges\<close>
have s: "labelAtOut v' p' = labelAtIn v p" by (rule solved)
from \<open>p' |\<in>| outPorts (nodeOf v')\<close>
show ?case
proof (cases rule: out_port_cases)
case (Hyp r h c)
from Hyp \<open>p' |\<in>| outPorts (nodeOf v')\<close>
have "h |\<in>| a_hyps c" and "c |\<in>| f_antecedent r" by auto
hence "hyps (nodeOf v') (Hyp h c) = Some c" using Hyp by simp
from well_scoped[OF \<open> _ \<in> edges\<close>[unfolded Hyp] this]
have "(v, p) = (v', c) \<or> v \<in> scope (v', c)".
hence "(v', c) \<in> insert (v, p) (snd ` set pth)"
proof
assume "(v, p) = (v', c)"
thus ?thesis by simp
next
assume "v \<in> scope (v', c)"
from this terminal_path_end_is_terminal[OF wf(2)] terminal_path_is_path[OF wf(2)]
have "(v', c) \<in> snd ` set pth" by (rule scope_find)
thus ?thesis by simp
qed
moreover
from \<open>hyps (nodeOf v') (Hyp h c) = Some c\<close>
have "Hyp h c |\<in>| hyps_for (nodeOf v') c" by simp
hence "labelAtOut v' (Hyp h c) |\<in>| extra_assms (v',c)" by auto
ultimately
have "labelAtOut v' (Hyp h c) |\<in>| ?\<Gamma>"
- by (fastforce simp add: fmember_iff_member_fset ffUnion.rep_eq)
+ by (fastforce simp add: ffUnion.rep_eq)
- hence "labelAtIn v p |\<in>| ?\<Gamma>" by (simp add: s[symmetric] Hyp fmember_iff_member_fset)
+ hence "labelAtIn v p |\<in>| ?\<Gamma>" by (simp add: s[symmetric] Hyp)
thus ?thesis
using Hyp
apply (auto intro: exI[where x = ?t] simp add: eff.simps simp del: hyps_along.simps)
done
next
case (Assumption f)
from \<open>v' |\<in>| vertices\<close> \<open>nodeOf v' = Assumption f\<close>
have "labelAtOut v' (Reg f) |\<in>| global_assms TYPE('var)"
by (rule global_assmsI)
hence "labelAtOut v' (Reg f) |\<in>| ?\<Gamma>" by auto
- hence "labelAtIn v p |\<in>| ?\<Gamma>" by (simp add: s[symmetric] Assumption fmember_iff_member_fset)
+ hence "labelAtIn v p |\<in>| ?\<Gamma>" by (simp add: s[symmetric] Assumption)
thus ?thesis using Assumption
by (auto intro: exI[where x = ?t] simp add: eff.simps)
next
case (Rule r f)
with \<open>nodeOf v' \<in> sset nodes\<close>
have "r \<in> sset rules"
by (auto simp add: nodes_def stream.set_map)
from Rule
have "hyps (nodeOf v') p' = None" by simp
with e \<open>terminal_path v t pth\<close>
have "terminal_path v' t ?pth'"..
from Rule \<open>p' |\<in>| outPorts (nodeOf v')\<close>
have "f |\<in>| f_consequent r" by simp
hence "f \<in> set (consequent r)" by (simp add: f_consequent_def)
with \<open>r \<in> sset rules\<close>
have "NatRule (r, f) \<in> sset (smap NatRule n_rules)"
by (auto simp add: stream.set_map n_rules_def no_empty_conclusions)
moreover
{
from \<open>f |\<in>| f_consequent r\<close>
have "f \<in> set (consequent r)" by (simp add: f_consequent_def)
hence "natEff_Inst (r, f) f (f_antecedent r)"
by (rule natEff_Inst.intros)
hence "eff (NatRule (r, f)) (?\<Gamma> \<turnstile> subst (inst v') (freshen (vidx v') f))
((\<lambda>ant. ((\<lambda>p. subst (inst v') (freshen (vidx v') p)) |`| a_hyps ant |\<union>| ?\<Gamma> \<turnstile> subst (inst v') (freshen (vidx v') (a_conc ant)))) |`| f_antecedent r)"
(is "eff _ _ ?ants")
proof (rule eff.intros)
fix ant f
assume "ant |\<in>| f_antecedent r"
from \<open>v' |\<in>| vertices\<close> \<open>ant |\<in>| f_antecedent r\<close>
have "valid_in_port (v',ant)" by (simp add: Rule)
assume "f |\<in>| ?\<Gamma>"
thus "freshenLC (vidx v') ` a_fresh ant \<inter> lconsts f = {}"
proof(induct rule: hyps_alongE)
case (Hyp v'' p'' h'')
from Hyp(1) snd_set_path_verties[OF terminal_path_is_path[OF \<open>terminal_path v' t ?pth'\<close>]]
- have "v'' |\<in>| vertices" by (force simp add: fmember_iff_member_fset)
+ have "v'' |\<in>| vertices" by (force simp add:)
from \<open>terminal_path v' t ?pth'\<close> Hyp(1)
have "v'' \<notin> scope (v', ant)" by (rule hyps_free_path_not_in_scope)
with \<open>valid_in_port (v',ant)\<close> \<open>v'' |\<in>| vertices\<close>
have "freshenLC (vidx v') ` local_vars (nodeOf v') ant \<inter> subst_lconsts (inst v'') = {}"
by (rule out_of_scope)
moreover
from hyps_free_vertices_distinct'[OF \<open>terminal_path v' t ?pth'\<close>] Hyp.hyps(1)
have "v'' \<noteq> v'" by (metis distinct.simps(2) fst_conv image_eqI list.set_map)
- hence "vidx v'' \<noteq> vidx v'" using \<open>v' |\<in>| vertices\<close> \<open>v'' |\<in>| vertices\<close> by (meson vidx_inj inj_onD notin_fset)
+ hence "vidx v'' \<noteq> vidx v'" using \<open>v' |\<in>| vertices\<close> \<open>v'' |\<in>| vertices\<close> by (meson vidx_inj inj_onD)
hence "freshenLC (vidx v') ` a_fresh ant \<inter> freshenLC (vidx v'') ` lconsts (labelsOut (nodeOf v'') h'') = {}"by auto
moreover
have "lconsts f \<subseteq> lconsts (freshen (vidx v'') (labelsOut (nodeOf v'') h'')) \<union> subst_lconsts (inst v'') " using \<open>f = _\<close>
by (simp add: labelAtOut_def fv_subst)
ultimately
show ?thesis
by (fastforce simp add: lconsts_freshen)
next
case (Assumption v pf)
hence "f = subst (inst v) (freshen (vidx v) pf)" by (simp add: labelAtOut_def)
moreover
- from Assumption have "Assumption pf \<in> sset nodes" using valid_nodes by (auto simp add: fmember_iff_member_fset)
+ from Assumption have "Assumption pf \<in> sset nodes" using valid_nodes by auto
hence "pf \<in> set assumptions" unfolding nodes_def by (auto simp add: stream.set_map)
hence "closed pf" by (rule assumptions_closed)
ultimately
have "lconsts f = {}" by (simp add: closed_no_lconsts lconsts_freshen subst_closed freshen_closed)
thus ?thesis by simp
qed
next
fix ant
assume "ant |\<in>| f_antecedent r"
from \<open>v' |\<in>| vertices\<close> \<open>ant |\<in>| f_antecedent r\<close>
have "valid_in_port (v',ant)" by (simp add: Rule)
moreover
note \<open>v' |\<in>| vertices\<close>
moreover
hence "v' \<notin> scope (v', ant)" by (rule scopes_not_refl)
ultimately
have "freshenLC (vidx v') ` local_vars (nodeOf v') ant \<inter> subst_lconsts (inst v') = {}"
by (rule out_of_scope)
thus "freshenLC (vidx v') ` a_fresh ant \<inter> subst_lconsts (inst v') = {}" by simp
qed
also
have "subst (inst v') (freshen (vidx v') f) = labelAtOut v' p'" using Rule by (simp add: labelAtOut_def)
also
note \<open>labelAtOut v' p' = labelAtIn v p\<close>
also
have "?ants = ((\<lambda>x. (extra_assms (v',x) |\<union>| hyps_along ?pth' \<turnstile> labelAtIn v' x)) |`| f_antecedent r)"
by (rule fimage_cong[OF refl])
- (auto simp add: labelAtIn_def labelAtOut_def Rule hyps_for_fimage fmember_iff_member_fset ffUnion.rep_eq)
+ (auto simp add: labelAtIn_def labelAtOut_def Rule hyps_for_fimage ffUnion.rep_eq)
finally
have "eff (NatRule (r, f))
(?\<Gamma>, labelAtIn v p)
((\<lambda>x. extra_assms (v',x) |\<union>| ?\<Gamma> \<turnstile> labelAtIn v' x) |`| f_antecedent r)".
}
moreover
{ fix x
assume "x |\<in>| cont ?t"
then obtain a where "x = tree v' a ?pth'" and "a |\<in>| f_antecedent r"
by (auto simp add: Rule)
note this(1)
moreover
from \<open>v' |\<in>| vertices\<close> \<open>a |\<in>| f_antecedent r\<close>
have "valid_in_port (v',a)" by (simp add: Rule)
moreover
note \<open>terminal_path v' t ?pth'\<close>
ultimately
have "\<exists>v p pth. x = tree v p pth \<and> valid_in_port (v,p) \<and> terminal_path v t pth"
by blast
}
ultimately
show ?thesis using Rule
by (auto intro!: exI[where x = ?t] simp add: comp_def funion_assoc)
next
case Helper
from Helper
have "hyps (nodeOf v') p' = None" by simp
with e \<open>terminal_path v t pth\<close>
have "terminal_path v' t ?pth'"..
have "labelAtIn v' (plain_ant anyP) = labelAtIn v p"
unfolding s[symmetric]
using Helper by (simp add: labelAtIn_def labelAtOut_def)
moreover
{ fix x
assume "x |\<in>| cont ?t"
hence "x = tree v' (plain_ant anyP) ?pth'"
by (auto simp add: Helper)
note this(1)
moreover
from \<open>v' |\<in>| vertices\<close>
have "valid_in_port (v',plain_ant anyP)" by (simp add: Helper)
moreover
note \<open>terminal_path v' t ?pth'\<close>
ultimately
have "\<exists>v p pth. x = tree v p pth \<and> valid_in_port (v,p) \<and> terminal_path v t pth"
by blast
}
ultimately
show ?thesis using Helper
by (auto intro!: exI[where x = ?t] simp add: comp_def funion_assoc )
qed
qed
lemma global_in_ass: "global_assms TYPE('var) |\<subseteq>| ass_forms"
proof
fix x
assume "x |\<in>| global_assms TYPE('var)"
then obtain v pf where "v |\<in>| vertices" and "nodeOf v = Assumption pf" and "x = labelAtOut v (Reg pf)"
by (auto simp add: global_assms_simps)
from this (1,2) valid_nodes
- have "Assumption pf \<in> sset nodes" by (auto simp add: fmember_iff_member_fset)
+ have "Assumption pf \<in> sset nodes" by (auto simp add:)
hence "pf \<in> set assumptions" by (auto simp add: nodes_def stream.set_map)
hence "closed pf" by (rule assumptions_closed)
with \<open>x = labelAtOut v (Reg pf)\<close>
have "x = pf" by (auto simp add: labelAtOut_def lconsts_freshen closed_no_lconsts freshen_closed subst_closed)
thus "x |\<in>| ass_forms" using \<open>pf \<in> set assumptions\<close> by (auto simp add: ass_forms_def)
qed
primcorec edge_tree :: "'vertex \<Rightarrow> ('form, 'var) in_port \<Rightarrow> ('vertex, 'form, 'var) edge' tree" where
"root (edge_tree v p) = (adjacentTo v p, (v,p))"
| "cont (edge_tree v p) =
(case adjacentTo v p of (v', p') \<Rightarrow>
(if isReg v' p' then ((\<lambda> p. edge_tree v' p) |`| inPorts (nodeOf v')) else {||}
))"
lemma tfinite_map_tree: "tfinite (map_tree f t) \<longleftrightarrow> tfinite t"
proof
assume "tfinite (map_tree f t)"
thus "tfinite t"
by (induction "map_tree f t" arbitrary: t rule: tfinite.induct)
(fastforce intro: tfinite.intros simp add: tree.map_sel)
next
assume "tfinite t"
thus "tfinite (map_tree f t)"
by (induction t rule: tfinite.induct)
(fastforce intro: tfinite.intros simp add: tree.map_sel)
qed
lemma finite_tree_edge_tree:
"tfinite (tree v p pth) \<longleftrightarrow> tfinite (edge_tree v p)"
proof-
have "map_tree (\<lambda> _. ()) (tree v p pth) = map_tree (\<lambda> _. ()) (edge_tree v p)"
by(coinduction arbitrary: v p pth)
(fastforce simp add: tree.map_sel rel_fset_def rel_set_def split: prod.split out_port.split graph_node.split option.split)
thus ?thesis by (metis tfinite_map_tree)
qed
coinductive forbidden_path :: "'vertex \<Rightarrow> ('vertex, 'form, 'var) edge' stream \<Rightarrow> bool" where
forbidden_path: "((v\<^sub>1,p\<^sub>1),(v\<^sub>2,p\<^sub>2)) \<in> edges \<Longrightarrow> hyps (nodeOf v\<^sub>1) p\<^sub>1 = None \<Longrightarrow> forbidden_path v\<^sub>1 pth \<Longrightarrow> forbidden_path v\<^sub>2 (((v\<^sub>1,p\<^sub>1),(v\<^sub>2,p\<^sub>2))##pth)"
lemma path_is_forbidden:
assumes "valid_in_port (v,p)"
assumes "ipath (edge_tree v p) es"
shows "forbidden_path v es"
using assms
proof(coinduction arbitrary: v p es)
case forbidden_path
let ?es' = "stl es"
from forbidden_path(2)
obtain t' where "root (edge_tree v p) = shd es" and "t' |\<in>| cont (edge_tree v p)" and "ipath t' ?es'"
by rule blast
from \<open>root (edge_tree v p) = shd es\<close>
have [simp]: "shd es = (adjacentTo v p, (v,p))" by simp
from saturated[OF \<open>valid_in_port (v,p)\<close>]
obtain v' p'
where e:"((v',p'),(v,p)) \<in> edges" and [simp]: "adjacentTo v p = (v',p')"
by (auto simp add: adjacentTo_def, metis (no_types, lifting) eq_fst_iff tfl_some)
let ?e = "((v',p'),(v,p))"
from e have "p' |\<in>| outPorts (nodeOf v')" using valid_edges by auto
thus ?case
proof(cases rule: out_port_cases)
case Hyp
with \<open>t' |\<in>| cont (edge_tree v p)\<close>
have False by auto
thus ?thesis..
next
case Assumption
with \<open>t' |\<in>| cont (edge_tree v p)\<close>
have False by auto
thus ?thesis..
next
case (Rule r f)
from \<open>t' |\<in>| cont (edge_tree v p)\<close> Rule
obtain a where [simp]: "t' = edge_tree v' a" and "a |\<in>| f_antecedent r" by auto
have "es = ?e ## ?es'" by (cases es rule: stream.exhaust_sel) simp
moreover
have "?e \<in> edges" using e by simp
moreover
from \<open>p' = Reg f\<close> \<open>nodeOf v' = Rule r\<close>
have "hyps (nodeOf v') p' = None" by simp
moreover
from e valid_edges have "v' |\<in>| vertices" by auto
with \<open>nodeOf v' = Rule r\<close> \<open>a |\<in>| f_antecedent r\<close>
have "valid_in_port (v', a)" by simp
moreover
have "ipath (edge_tree v' a) ?es'" using \<open>ipath t' _\<close> by simp
ultimately
show ?thesis by metis
next
case Helper
from \<open>t' |\<in>| cont (edge_tree v p)\<close> Helper
have [simp]: "t' = edge_tree v' (plain_ant anyP)" by simp
have "es = ?e ## ?es'" by (cases es rule: stream.exhaust_sel) simp
moreover
have "?e \<in> edges" using e by simp
moreover
from \<open>p' = Reg anyP\<close> \<open>nodeOf v' = Helper\<close>
have "hyps (nodeOf v') p' = None" by simp
moreover
from e valid_edges have "v' |\<in>| vertices" by auto
with \<open>nodeOf v' = Helper\<close>
have "valid_in_port (v', plain_ant anyP)" by simp
moreover
have "ipath (edge_tree v' (plain_ant anyP)) ?es'" using \<open>ipath t' _\<close> by simp
ultimately
show ?thesis by metis
qed
qed
lemma forbidden_path_prefix_is_path:
assumes "forbidden_path v es"
obtains v' where "path v' v (rev (stake n es))"
using assms
apply (atomize_elim)
apply (induction n arbitrary: v es)
apply simp
apply (simp add: path_snoc)
apply (subst (asm) (2) forbidden_path.simps)
apply auto
done
lemma forbidden_path_prefix_is_hyp_free:
assumes "forbidden_path v es"
shows "hyps_free (rev (stake n es))"
using assms
apply (induction n arbitrary: v es)
apply (simp add: hyps_free_def)
apply (subst (asm) (2) forbidden_path.simps)
apply (force simp add: hyps_free_def)
done
text \<open>And now we prove that the tree is finite, which requires the above notion of a
@{term forbidden_path}, i.e.\@ an infinite path.\<close>
theorem finite_tree:
assumes "valid_in_port (v,p)"
assumes "terminal_vertex v"
shows "tfinite (tree v p pth)"
proof(rule ccontr)
let ?n = "Suc (fcard vertices)"
assume "\<not> tfinite (tree v p pth)"
hence "\<not> tfinite (edge_tree v p)" unfolding finite_tree_edge_tree.
then obtain es :: "('vertex, 'form, 'var) edge' stream"
where "ipath (edge_tree v p) es" using Konig by blast
with \<open>valid_in_port (v,p)\<close>
have "forbidden_path v es" by (rule path_is_forbidden)
from forbidden_path_prefix_is_path[OF this] forbidden_path_prefix_is_hyp_free[OF this]
obtain v' where "path v' v (rev (stake ?n es))" and "hyps_free (rev (stake ?n es))"
by blast
from this \<open>terminal_vertex v\<close>
have "terminal_path v' v (rev (stake ?n es))" by (rule terminal_pathI)
hence "length (rev (stake ?n es)) \<le> fcard vertices"
by (rule hyps_free_limited)
thus False by simp
qed
text \<open>The main result of this theory.\<close>
theorem solved
unfolding solved_def
proof(intro ballI allI conjI impI)
fix c
assume "c |\<in>| conc_forms"
hence "c \<in> set conclusions" by (auto simp add: conc_forms_def)
from this(1) conclusions_present
obtain v where "v |\<in>| vertices" and "nodeOf v = Conclusion c"
- by (auto, metis (no_types, lifting) image_iff image_subset_iff notin_fset)
+ by auto
have "valid_in_port (v, (plain_ant c))"
using \<open>v |\<in>| vertices\<close> \<open>nodeOf _ = _ \<close> by simp
have "terminal_vertex v" using \<open>v |\<in>| vertices\<close> \<open>nodeOf v = Conclusion c\<close> by auto
let ?t = "tree v (plain_ant c) []"
have "fst (root ?t) = (global_assms TYPE('var), c)"
using \<open>c \<in> set conclusions\<close> \<open>nodeOf _ = _\<close>
by (auto simp add: labelAtIn_def conclusions_closed closed_no_lconsts freshen_def rename_closed subst_closed)
moreover
have "global_assms TYPE('var) |\<subseteq>| ass_forms" by (rule global_in_ass)
moreover
from \<open>terminal_vertex v\<close>
have "terminal_path v v []" by (rule terminal_path_empty)
with \<open>valid_in_port (v, (plain_ant c))\<close>
have "wf ?t" by (rule wf_tree)
moreover
from \<open>valid_in_port (v, plain_ant c)\<close> \<open>terminal_vertex v\<close>
have "tfinite ?t" by (rule finite_tree)
ultimately
show "\<exists>\<Gamma> t. fst (root t) = (\<Gamma> \<turnstile> c) \<and> \<Gamma> |\<subseteq>| ass_forms \<and> wf t \<and> tfinite t" by blast
qed
end
end
\ No newline at end of file
diff --git a/thys/Incredible_Proof_Machine/Incredible_Deduction.thy b/thys/Incredible_Proof_Machine/Incredible_Deduction.thy
--- a/thys/Incredible_Proof_Machine/Incredible_Deduction.thy
+++ b/thys/Incredible_Proof_Machine/Incredible_Deduction.thy
@@ -1,528 +1,528 @@
theory Incredible_Deduction
imports
Main
"HOL-Library.FSet"
"HOL-Library.Stream"
Incredible_Signatures
"HOL-Eisbach.Eisbach"
begin
text \<open>This theory contains the definition for actual proof graphs, and their various possible
properties.\<close>
text \<open>The following locale first defines graphs, without edges.\<close>
locale Vertex_Graph =
Port_Graph_Signature nodes inPorts outPorts
for nodes :: "'node stream"
and inPorts :: "'node \<Rightarrow> 'inPort fset"
and outPorts :: "'node \<Rightarrow> 'outPort fset" +
fixes vertices :: "'v fset"
fixes nodeOf :: "'v \<Rightarrow> 'node"
begin
fun valid_out_port where "valid_out_port (v,p) \<longleftrightarrow> v |\<in>| vertices \<and> p |\<in>| outPorts (nodeOf v)"
fun valid_in_port where "valid_in_port (v,p) \<longleftrightarrow> v |\<in>| vertices \<and> p |\<in>| inPorts (nodeOf v)"
fun terminal_node where
"terminal_node n \<longleftrightarrow> outPorts n = {||}"
fun terminal_vertex where
"terminal_vertex v \<longleftrightarrow> v |\<in>| vertices \<and> terminal_node (nodeOf v)"
end
text \<open>And now we add the edges. This allows us to define paths and scopes.\<close>
type_synonym ('v, 'outPort, 'inPort) edge = "(('v \<times> 'outPort) \<times> ('v \<times> 'inPort))"
locale Pre_Port_Graph =
Vertex_Graph nodes inPorts outPorts vertices nodeOf
for nodes :: "'node stream"
and inPorts :: "'node \<Rightarrow> 'inPort fset"
and outPorts :: "'node \<Rightarrow> 'outPort fset"
and vertices :: "'v fset"
and nodeOf :: "'v \<Rightarrow> 'node" +
fixes edges :: "('v, 'outPort, 'inPort) edge set"
begin
fun edge_begin :: "(('v \<times> 'outPort) \<times> ('v \<times> 'inPort)) \<Rightarrow> 'v" where
"edge_begin ((v1,p1),(v2,p2)) = v1"
fun edge_end :: "(('v \<times> 'outPort) \<times> ('v \<times> 'inPort)) \<Rightarrow> 'v" where
"edge_end ((v1,p1),(v2,p2)) = v2"
lemma edge_begin_tup: "edge_begin x = fst (fst x)" by (metis edge_begin.simps prod.collapse)
lemma edge_end_tup: "edge_end x = fst (snd x)" by (metis edge_end.simps prod.collapse)
inductive path :: "'v \<Rightarrow> 'v \<Rightarrow> ('v, 'outPort, 'inPort) edge list \<Rightarrow> bool" where
path_empty: "path v v []" |
path_cons: "e \<in> edges \<Longrightarrow> path (edge_end e) v' pth \<Longrightarrow> path (edge_begin e) v' (e#pth)"
inductive_simps path_cons_simp': "path v v' (e#pth)"
inductive_simps path_empty_simp[simp]: "path v v' []"
lemma path_cons_simp: "path v v' (e # pth) \<longleftrightarrow> fst (fst e) = v \<and> e \<in> edges \<and> path (fst (snd e)) v' pth"
by(auto simp add: path_cons_simp', metis prod.collapse)
lemma path_appendI: "path v v' pth1 \<Longrightarrow> path v' v'' pth2 \<Longrightarrow> path v v'' (pth1@pth2)"
by (induction pth1 arbitrary: v) (auto simp add: path_cons_simp )
lemma path_split: "path v v' (pth1@[e]@pth2) \<longleftrightarrow> path v (edge_end e) (pth1@[e]) \<and> path (edge_end e) v' pth2"
by (induction pth1 arbitrary: v) (auto simp add: path_cons_simp edge_end_tup intro: path_empty)
lemma path_split2: "path v v' (pth1@(e#pth2)) \<longleftrightarrow> path v (edge_begin e) pth1 \<and> path (edge_begin e) v' (e#pth2)"
by (induction pth1 arbitrary: v) (auto simp add: path_cons_simp edge_begin_tup intro: path_empty)
lemma path_snoc: "path v v' (pth1@[e]) \<longleftrightarrow> e \<in> edges \<and> path v (edge_begin e) pth1 \<and> edge_end e = v'"
by (auto simp add: path_split2 path_cons_simp edge_end_tup edge_begin_tup)
inductive_set scope :: "'v \<times> 'inPort \<Rightarrow> 'v set" for ps where
"v |\<in>| vertices \<Longrightarrow> (\<And> pth v'. path v v' pth \<Longrightarrow> terminal_vertex v' \<Longrightarrow> ps \<in> snd ` set pth)
\<Longrightarrow> v \<in> scope ps"
lemma scope_find:
assumes "v \<in> scope ps"
assumes "terminal_vertex v'"
assumes "path v v' pth"
shows "ps \<in> snd ` set pth"
using assms by (auto simp add: scope.simps)
lemma snd_set_split:
assumes "ps \<in> snd ` set pth"
obtains pth1 pth2 e where "pth = pth1@[e]@pth2" and "snd e = ps" and "ps \<notin> snd ` set pth1"
using assms
proof (atomize_elim, induction pth)
case Nil thus ?case by simp
next
case (Cons e pth)
show ?case
proof(cases "snd e = ps")
case True
hence "e # pth = [] @ [e] @ pth \<and> snd e = ps \<and> ps \<notin> snd ` set []" by auto
thus ?thesis by (intro exI)
next
case False
with Cons(2)
have "ps \<in> snd ` set pth" by auto
from Cons.IH[OF this this]
obtain pth1 e' pth2 where "pth = pth1 @ [e'] @ pth2 \<and> snd e' = ps \<and> ps \<notin> snd ` set pth1" by auto
with False
have "e#pth = (e# pth1) @ [e'] @ pth2 \<and> snd e' = ps \<and> ps \<notin> snd ` set (e#pth1)" by auto
thus ?thesis by blast
qed
qed
lemma scope_split:
assumes "v \<in> scope ps"
assumes "path v v' pth"
assumes "terminal_vertex v'"
obtains pth1 e pth2
where "pth = (pth1@[e])@pth2" and "path v (fst ps) (pth1@[e])" and "path (fst ps) v' pth2" and "snd e = ps" and "ps \<notin> snd ` set pth1"
proof-
from assms
have "ps \<in> snd ` set pth" by (auto simp add: scope.simps)
then obtain pth1 pth2 e where "pth = pth1@[e]@pth2" and "snd e = ps" and "ps \<notin> snd ` set pth1" by (rule snd_set_split)
from \<open>path _ _ _\<close> and \<open>pth = pth1@[e]@pth2\<close>
have "path v (edge_end e) (pth1@[e])" and "path (edge_end e) v' pth2" by (metis path_split)+
show thesis
proof(rule that)
show "pth = (pth1@[e])@pth2" using \<open>pth= _\<close> by simp
show "path v (fst ps) (pth1@[e])" using \<open>path v (edge_end e) (pth1@[e])\<close> \<open>snd e = ps\<close> by (simp add: edge_end_tup)
show "path (fst ps) v' pth2" using \<open>path (edge_end e) v' pth2\<close> \<open>snd e = ps\<close> by (simp add: edge_end_tup)
show "ps \<notin> snd ` set pth1" by fact
show "snd e = ps" by fact
qed
qed
end
text \<open>This adds well-formedness conditions to the edges and vertices.\<close>
locale Port_Graph = Pre_Port_Graph +
assumes valid_nodes: "nodeOf ` fset vertices \<subseteq> sset nodes"
assumes valid_edges: "\<forall> (ps1,ps2) \<in> edges. valid_out_port ps1 \<and> valid_in_port ps2"
begin
lemma snd_set_path_verties: "path v v' pth \<Longrightarrow> fst ` snd ` set pth \<subseteq> fset vertices"
apply (induction rule: path.induct)
apply auto
- apply (metis valid_in_port.elims(2) edge_end.simps notin_fset case_prodD valid_edges)
+ apply (metis valid_in_port.elims(2) edge_end.simps case_prodD valid_edges)
done
lemma fst_set_path_verties: "path v v' pth \<Longrightarrow> fst ` fst ` set pth \<subseteq> fset vertices"
apply (induction rule: path.induct)
apply auto
- apply (metis valid_out_port.elims(2) edge_begin.simps notin_fset case_prodD valid_edges)
+ apply (metis valid_out_port.elims(2) edge_begin.simps case_prodD valid_edges)
done
end
text \<open>A pruned graph is one where every node has a path to a terminal node (which will be the conclusions).\<close>
locale Pruned_Port_Graph = Port_Graph +
assumes pruned: "\<And>v. v |\<in>| vertices \<Longrightarrow> (\<exists>pth v'. path v v' pth \<and> terminal_vertex v')"
begin
lemma scopes_not_refl:
assumes "v |\<in>| vertices"
shows "v \<notin> scope (v,p)"
proof(rule notI)
assume "v \<in> scope (v,p)"
from pruned[OF assms]
obtain pth t where "terminal_vertex t" and "path v t pth" and least: "\<forall> pth'. path v t pth' \<longrightarrow> length pth \<le> length pth'"
by atomize_elim (auto simp del: terminal_vertex.simps elim: ex_has_least_nat)
from scope_split[OF \<open>v \<in> scope (v,p)\<close> \<open>path v t pth\<close> \<open>terminal_vertex t\<close>]
obtain pth1 e pth2 where "pth = (pth1 @ [e]) @ pth2" "path v t pth2" by (metis fst_conv)
from this(2) least
have "length pth \<le> length pth2" by auto
with \<open>pth = _\<close>
show False by auto
qed
text \<open>This lemma can be found in \<^cite>\<open>"incredible"\<close>, but it is otherwise inconsequential.\<close>
lemma scopes_nest:
fixes ps1 ps2
shows "scope ps1 \<subseteq> scope ps2 \<or> scope ps2 \<subseteq> scope ps1 \<or> scope ps1 \<inter> scope ps2 = {}"
proof(cases "ps1 = ps2")
assume "ps1 \<noteq> ps2"
{
fix v
assume "v \<in> scope ps1 \<inter> scope ps2"
hence "v |\<in>| vertices" using scope.simps by auto
then obtain pth t where "path v t pth" and "terminal_vertex t" using pruned by blast
from \<open>path v t pth\<close> and \<open>terminal_vertex t\<close> and \<open>v \<in> scope ps1 \<inter> scope ps2\<close>
obtain pth1a e1 pth1b where "pth = (pth1a@[e1])@pth1b" and "path v (fst ps1) (pth1a@[e1])" and "snd e1 = ps1" and "ps1 \<notin> snd ` set pth1a"
by (auto elim: scope_split)
from \<open>path v t pth\<close> and \<open>terminal_vertex t\<close> and \<open>v \<in> scope ps1 \<inter> scope ps2\<close>
obtain pth2a e2 pth2b where "pth = (pth2a@[e2])@pth2b" and "path v (fst ps2) (pth2a@[e2])" and "snd e2 = ps2" and "ps2 \<notin> snd ` set pth2a"
by (auto elim: scope_split)
from \<open>pth = (pth1a@[e1])@pth1b\<close> \<open>pth = (pth2a@[e2])@pth2b\<close>
have "set pth1a \<subseteq> set pth2a \<or> set pth2a \<subseteq> set pth1a" by (auto simp add: append_eq_append_conv2)
hence "scope ps1 \<subseteq> scope ps2 \<or> scope ps2 \<subseteq> scope ps1"
proof
assume "set pth1a \<subseteq> set pth2a" with \<open>ps2 \<notin> _\<close>
have "ps2 \<notin> snd ` set (pth1a@[e1])" using \<open>ps1 \<noteq> ps2\<close> \<open>snd e1 = ps1\<close> by auto
have "scope ps1 \<subseteq> scope ps2"
proof
fix v'
assume "v' \<in> scope ps1"
hence "v' |\<in>| vertices" using scope.simps by auto
thus "v' \<in> scope ps2"
proof(rule scope.intros)
fix pth' t'
assume "path v' t' pth'" and "terminal_vertex t'"
with \<open>v' \<in> scope ps1\<close>
obtain pth3a e3 pth3b where "pth' = (pth3a@[e3])@pth3b" and "path (fst ps1) t' pth3b"
by (auto elim: scope_split)
have "path v t' ((pth1a@[e1]) @ pth3b)" using \<open>path v (fst ps1) (pth1a@[e1])\<close> and \<open>path (fst ps1) t' pth3b\<close>
by (rule path_appendI)
- with \<open>terminal_vertex t'\<close> \<open>v \<in> _\<close>
+ with \<open>terminal_vertex t'\<close> \<open>v \<in> scope ps1 \<inter> scope ps2\<close>
have "ps2 \<in> snd ` set ((pth1a@[e1]) @ pth3b)" by (meson IntD2 scope.cases)
hence "ps2 \<in> snd ` set pth3b" using \<open>ps2 \<notin> snd ` set (pth1a@[e1])\<close> by auto
thus "ps2 \<in> snd ` set pth'" using \<open>pth'=_\<close> by auto
qed
qed
thus ?thesis by simp
next
assume "set pth2a \<subseteq> set pth1a" with \<open>ps1 \<notin> _\<close>
have "ps1 \<notin> snd ` set (pth2a@[e2])" using \<open>ps1 \<noteq> ps2\<close> \<open>snd e2 = ps2\<close> by auto
have "scope ps2 \<subseteq> scope ps1"
proof
fix v'
assume "v' \<in> scope ps2"
hence "v' |\<in>| vertices" using scope.simps by auto
thus "v' \<in> scope ps1"
proof(rule scope.intros)
fix pth' t'
assume "path v' t' pth'" and "terminal_vertex t'"
with \<open>v' \<in> scope ps2\<close>
obtain pth3a e3 pth3b where "pth' = (pth3a@[e3])@pth3b" and "path (fst ps2) t' pth3b"
by (auto elim: scope_split)
have "path v t' ((pth2a@[e2]) @ pth3b)" using \<open>path v (fst ps2) (pth2a@[e2])\<close> and \<open>path (fst ps2) t' pth3b\<close>
by (rule path_appendI)
- with \<open>terminal_vertex t'\<close> \<open>v \<in> _\<close>
+ with \<open>terminal_vertex t'\<close> \<open>v \<in> scope ps1 \<inter> scope ps2\<close>
have "ps1 \<in> snd ` set ((pth2a@[e2]) @ pth3b)" by (meson IntD1 scope.cases)
hence "ps1 \<in> snd ` set pth3b" using \<open>ps1 \<notin> snd ` set (pth2a@[e2])\<close> by auto
thus "ps1 \<in> snd ` set pth'" using \<open>pth'=_\<close> by auto
qed
qed
thus ?thesis by simp
qed
}
thus ?thesis by blast
qed simp
end
text \<open>A well-scoped graph is one where a port marked to be a local hypothesis is only connected to
the corresponding input port, either directly or via a path. It must not be, however, that there is
a a path from such a hypothesis to a terminal node that does not pass by the dedicated input port;
this is expressed via scopes.
\<close>
locale Scoped_Graph = Port_Graph + Port_Graph_Signature_Scoped
locale Well_Scoped_Graph = Scoped_Graph +
assumes well_scoped: "((v\<^sub>1,p\<^sub>1),(v\<^sub>2,p\<^sub>2)) \<in> edges \<Longrightarrow> hyps (nodeOf v\<^sub>1) p\<^sub>1 = Some p' \<Longrightarrow> (v\<^sub>2,p\<^sub>2) = (v\<^sub>1,p') \<or> v\<^sub>2 \<in> scope (v\<^sub>1,p')"
context Scoped_Graph
begin
definition hyps_free where
"hyps_free pth = (\<forall> v\<^sub>1 p\<^sub>1 v\<^sub>2 p\<^sub>2. ((v\<^sub>1,p\<^sub>1),(v\<^sub>2,p\<^sub>2)) \<in> set pth \<longrightarrow> hyps (nodeOf v\<^sub>1) p\<^sub>1 = None)"
lemma hyps_free_Nil[simp]: "hyps_free []" by (simp add: hyps_free_def)
lemma hyps_free_Cons[simp]: "hyps_free (e#pth) \<longleftrightarrow> hyps_free pth \<and> hyps (nodeOf (fst (fst e))) (snd (fst e)) = None"
by (auto simp add: hyps_free_def) (metis prod.collapse)
lemma path_vertices_shift:
assumes "path v v' pth"
shows "map fst (map fst pth)@[v'] = v#map fst (map snd pth)"
using assms by induction auto
inductive terminal_path where
terminal_path_empty: "terminal_vertex v \<Longrightarrow> terminal_path v v []" |
terminal_path_cons: "((v\<^sub>1,p\<^sub>1),(v\<^sub>2,p\<^sub>2)) \<in> edges \<Longrightarrow> terminal_path v\<^sub>2 v' pth \<Longrightarrow> hyps (nodeOf v\<^sub>1) p\<^sub>1 = None \<Longrightarrow> terminal_path v\<^sub>1 v' (((v\<^sub>1,p\<^sub>1),(v\<^sub>2,p\<^sub>2))#pth)"
lemma terminal_path_is_path:
assumes "terminal_path v v' pth"
shows "path v v' pth"
using assms by induction (auto simp add: path_cons_simp)
lemma terminal_path_is_hyps_free:
assumes "terminal_path v v' pth"
shows "hyps_free pth"
using assms by induction (auto simp add: hyps_free_def)
lemma terminal_path_end_is_terminal:
assumes "terminal_path v v' pth"
shows "terminal_vertex v'"
using assms by induction
lemma terminal_pathI:
assumes "path v v' pth"
assumes "hyps_free pth"
assumes "terminal_vertex v'"
shows "terminal_path v v' pth"
using assms
by induction (auto intro: terminal_path.intros)
end
text \<open>An acyclic graph is one where there are no non-trivial cyclic paths (disregarding
edges that are local hypotheses -- these are naturally and benignly cyclic).\<close>
locale Acyclic_Graph = Scoped_Graph +
assumes hyps_free_acyclic: "path v v pth \<Longrightarrow> hyps_free pth \<Longrightarrow> pth = []"
begin
lemma hyps_free_vertices_distinct:
assumes "terminal_path v v' pth"
shows "distinct (map fst (map fst pth)@[v'])"
using assms
proof(induction v v' pth)
case terminal_path_empty
show ?case by simp
next
case (terminal_path_cons v\<^sub>1 p\<^sub>1 v\<^sub>2 p\<^sub>2 v' pth)
note terminal_path_cons.IH
moreover
have "v\<^sub>1 \<notin> fst ` fst ` set pth"
proof
assume "v\<^sub>1 \<in> fst ` fst ` set pth"
then obtain pth1 e' pth2 where "pth = pth1@[e']@pth2" and "v\<^sub>1 = fst (fst e')"
apply (atomize_elim)
apply (induction pth)
apply (solves simp)
apply (auto)
apply (solves \<open>rule exI[where x = "[]"]; simp\<close>)
apply (metis Cons_eq_appendI image_eqI prod.sel(1))
done
with terminal_path_is_path[OF \<open>terminal_path v\<^sub>2 v' pth\<close>]
have "path v\<^sub>2 v\<^sub>1 pth1" by (simp add: path_split2 edge_begin_tup)
with \<open>((v\<^sub>1, p\<^sub>1), (v\<^sub>2, p\<^sub>2)) \<in> _\<close>
have "path v\<^sub>1 v\<^sub>1 (((v\<^sub>1, p\<^sub>1), (v\<^sub>2, p\<^sub>2)) # pth1)" by (simp add: path_cons_simp)
moreover
from terminal_path_is_hyps_free[OF \<open>terminal_path v\<^sub>2 v' pth\<close>]
\<open>hyps (nodeOf v\<^sub>1) p\<^sub>1 = None\<close>
\<open>pth = pth1@[e']@pth2\<close>
have "hyps_free(((v\<^sub>1, p\<^sub>1), (v\<^sub>2, p\<^sub>2)) # pth1)"
by (auto simp add: hyps_free_def)
ultimately
show False using hyps_free_acyclic by blast
qed
moreover
have "v\<^sub>1 \<noteq> v'"
using hyps_free_acyclic path_cons terminal_path_cons.hyps(1) terminal_path_cons.hyps(2) terminal_path_cons.hyps(3) terminal_path_is_hyps_free terminal_path_is_path by fastforce
ultimately
show ?case by (auto simp add: comp_def)
qed
lemma hyps_free_vertices_distinct':
assumes "terminal_path v v' pth"
shows "distinct (v # map fst (map snd pth))"
using hyps_free_vertices_distinct[OF assms]
unfolding path_vertices_shift[OF terminal_path_is_path[OF assms]]
.
lemma hyps_free_limited:
assumes "terminal_path v v' pth"
shows "length pth \<le> fcard vertices"
proof-
have "length pth = length (map fst (map fst pth))" by simp
also
from hyps_free_vertices_distinct[OF assms]
have "distinct (map fst (map fst pth))" by simp
hence "length (map fst (map fst pth)) = card (set (map fst (map fst pth)))"
by (rule distinct_card[symmetric])
also have "\<dots> \<le> card (fset vertices)"
proof (rule card_mono[OF finite_fset])
from assms(1)
show "set (map fst (map fst pth)) \<subseteq> fset vertices"
- by (induction v v' pth) (auto, metis valid_edges notin_fset case_prodD valid_out_port.simps)
+ by (induction v v' pth) (auto, metis valid_edges case_prodD valid_out_port.simps)
qed
also have "\<dots> = fcard vertices" by (simp add: fcard.rep_eq)
finally show ?thesis.
qed
lemma hyps_free_path_not_in_scope:
assumes "terminal_path v t pth"
assumes "(v',p') \<in> snd ` set pth"
shows "v' \<notin> scope (v, p)"
proof
assume "v' \<in> scope (v,p)"
from \<open>(v',p') \<in> snd ` set pth\<close>
obtain pth1 pth2 e where "pth = pth1@[e]@pth2" and "snd e = (v',p')" by (rule snd_set_split)
from terminal_path_is_path[OF assms(1), unfolded \<open>pth = _ \<close>] \<open>snd e = _\<close>
have "path v v' (pth1@[e])" and "path v' t pth2" unfolding path_split by (auto simp add: edge_end_tup)
from \<open>v' \<in> scope (v,p)\<close> terminal_path_end_is_terminal[OF assms(1)] \<open>path v' t pth2\<close>
have "(v,p) \<in> snd ` set pth2" by (rule scope_find)
then obtain pth2a e' pth2b where "pth2 = pth2a@[e']@pth2b" and "snd e' = (v,p)" by (rule snd_set_split)
from \<open>path v' t pth2\<close>[unfolded \<open>pth2 = _ \<close>] \<open>snd e' = _\<close>
have "path v' v (pth2a@[e'])" and "path v t pth2b" unfolding path_split by (auto simp add: edge_end_tup)
from \<open>path v v' (pth1@[e])\<close> \<open>path v' v (pth2a@[e'])\<close>
have "path v v ((pth1@[e])@(pth2a@[e']))" by (rule path_appendI)
moreover
from terminal_path_is_hyps_free[OF assms(1)] \<open>pth = _\<close> \<open>pth2 = _\<close>
have "hyps_free ((pth1@[e])@(pth2a@[e']))" by (auto simp add: hyps_free_def)
ultimately
have "((pth1@[e])@(pth2a@[e'])) = []" by (rule hyps_free_acyclic)
thus False by simp
qed
end
text \<open>A saturated graph is one where every input port is incident to an edge.\<close>
locale Saturated_Graph = Port_Graph +
assumes saturated: "valid_in_port (v,p) \<Longrightarrow> \<exists> e \<in> edges . snd e = (v,p)"
text \<open>These four conditions make up a well-shaped graph.\<close>
locale Well_Shaped_Graph = Well_Scoped_Graph + Acyclic_Graph + Saturated_Graph + Pruned_Port_Graph
text \<open>Next we demand an instantiation. This consists of a unique natural number per vertex,
in order to rename the local constants apart, and furthermore a substitution per block
which instantiates the schematic formulas given in @{term Labeled_Signature}.\<close>
locale Instantiation =
Vertex_Graph nodes _ _ vertices _ +
Labeled_Signature nodes _ _ _ labelsIn labelsOut +
Abstract_Formulas freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP
for nodes :: "'node stream" and edges :: "('vertex, 'outPort, 'inPort) edge set" and vertices :: "'vertex fset" and labelsIn :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'form" and labelsOut :: "'node \<Rightarrow> 'outPort \<Rightarrow> 'form"
and freshenLC :: "nat \<Rightarrow> 'var \<Rightarrow> 'var"
and renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> 'form \<Rightarrow> 'form"
and lconsts :: "'form \<Rightarrow> 'var set"
and closed :: "'form \<Rightarrow> bool"
and subst :: "'subst \<Rightarrow> 'form \<Rightarrow> 'form"
and subst_lconsts :: "'subst \<Rightarrow> 'var set"
and subst_renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> ('subst \<Rightarrow> 'subst)"
and anyP :: "'form" +
fixes vidx :: "'vertex \<Rightarrow> nat"
and inst :: "'vertex \<Rightarrow> 'subst"
assumes vidx_inj: "inj_on vidx (fset vertices)"
begin
definition labelAtIn :: "'vertex \<Rightarrow> 'inPort \<Rightarrow> 'form" where
"labelAtIn v p = subst (inst v) (freshen (vidx v) (labelsIn (nodeOf v) p))"
definition labelAtOut :: "'vertex \<Rightarrow> 'outPort \<Rightarrow> 'form" where
"labelAtOut v p = subst (inst v) (freshen (vidx v) (labelsOut (nodeOf v) p))"
end
text \<open>A solution is an instantiation where on every edge, both incident ports are labeld with
the same formula.\<close>
locale Solution =
Instantiation _ _ _ _ _ edges for edges :: "(('vertex \<times> 'outPort) \<times> 'vertex \<times> 'inPort) set" +
assumes solved: "((v\<^sub>1,p\<^sub>1),(v\<^sub>2,p\<^sub>2)) \<in> edges \<Longrightarrow> labelAtOut v\<^sub>1 p\<^sub>1 = labelAtIn v\<^sub>2 p\<^sub>2"
locale Proof_Graph = Well_Shaped_Graph + Solution
text \<open>If we have locally scoped constants, we demand that only blocks in the scope of the
corresponding input port may mention such a locally scoped variable in its substitution.\<close>
locale Well_Scoped_Instantiation =
Pre_Port_Graph nodes inPorts outPorts vertices nodeOf edges +
Instantiation inPorts outPorts nodeOf hyps nodes edges vertices labelsIn labelsOut freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP vidx inst +
Port_Graph_Signature_Scoped_Vars nodes inPorts outPorts freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP local_vars
for freshenLC :: "nat \<Rightarrow> 'var \<Rightarrow> 'var"
and renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> 'form \<Rightarrow> 'form"
and lconsts :: "'form \<Rightarrow> 'var set"
and closed :: "'form \<Rightarrow> bool"
and subst :: "'subst \<Rightarrow> 'form \<Rightarrow> 'form"
and subst_lconsts :: "'subst \<Rightarrow> 'var set"
and subst_renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> ('subst \<Rightarrow> 'subst)"
and anyP :: "'form"
and inPorts :: "'node \<Rightarrow> 'inPort fset"
and outPorts :: "'node \<Rightarrow> 'outPort fset"
and nodeOf :: "'vertex \<Rightarrow> 'node"
and hyps :: "'node \<Rightarrow> 'outPort \<Rightarrow> 'inPort option"
and nodes :: "'node stream"
and vertices :: "'vertex fset"
and labelsIn :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'form"
and labelsOut :: "'node \<Rightarrow> 'outPort \<Rightarrow> 'form"
and vidx :: "'vertex \<Rightarrow> nat"
and inst :: "'vertex \<Rightarrow> 'subst"
and edges :: "('vertex, 'outPort, 'inPort) edge set"
and local_vars :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'var set" +
assumes well_scoped_inst:
"valid_in_port (v,p) \<Longrightarrow>
var \<in> local_vars (nodeOf v) p \<Longrightarrow>
v' |\<in>| vertices \<Longrightarrow>
freshenLC (vidx v) var \<in> subst_lconsts (inst v') \<Longrightarrow>
v' \<in> scope (v,p)"
begin
lemma out_of_scope: "valid_in_port (v,p) \<Longrightarrow> v' |\<in>| vertices \<Longrightarrow> v' \<notin> scope (v,p) \<Longrightarrow> freshenLC (vidx v) ` local_vars (nodeOf v) p \<inter> subst_lconsts (inst v') = {}"
using well_scoped_inst by auto
end
text \<open>The following locale assembles all these conditions.\<close>
locale Scoped_Proof_Graph =
Instantiation inPorts outPorts nodeOf hyps nodes edges vertices labelsIn labelsOut freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP vidx inst +
Well_Shaped_Graph nodes inPorts outPorts vertices nodeOf edges hyps +
Solution inPorts outPorts nodeOf hyps nodes vertices labelsIn labelsOut freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP vidx inst edges +
Well_Scoped_Instantiation freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP inPorts outPorts nodeOf hyps nodes vertices labelsIn labelsOut vidx inst edges local_vars
for freshenLC :: "nat \<Rightarrow> 'var \<Rightarrow> 'var"
and renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> 'form \<Rightarrow> 'form"
and lconsts :: "'form \<Rightarrow> 'var set"
and closed :: "'form \<Rightarrow> bool"
and subst :: "'subst \<Rightarrow> 'form \<Rightarrow> 'form"
and subst_lconsts :: "'subst \<Rightarrow> 'var set"
and subst_renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> ('subst \<Rightarrow> 'subst)"
and anyP :: "'form"
and inPorts :: "'node \<Rightarrow> 'inPort fset"
and outPorts :: "'node \<Rightarrow> 'outPort fset"
and nodeOf :: "'vertex \<Rightarrow> 'node"
and hyps :: "'node \<Rightarrow> 'outPort \<Rightarrow> 'inPort option"
and nodes :: "'node stream"
and vertices :: "'vertex fset"
and labelsIn :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'form"
and labelsOut :: "'node \<Rightarrow> 'outPort \<Rightarrow> 'form"
and vidx :: "'vertex \<Rightarrow> nat"
and inst :: "'vertex \<Rightarrow> 'subst"
and edges :: "('vertex, 'outPort, 'inPort) edge set"
and local_vars :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'var set"
end
diff --git a/thys/Incredible_Proof_Machine/Incredible_Signatures.thy b/thys/Incredible_Proof_Machine/Incredible_Signatures.thy
--- a/thys/Incredible_Proof_Machine/Incredible_Signatures.thy
+++ b/thys/Incredible_Proof_Machine/Incredible_Signatures.thy
@@ -1,69 +1,69 @@
theory Incredible_Signatures
imports
Main
"HOL-Library.FSet"
"HOL-Library.Stream"
Abstract_Formula
begin
text \<open>This theory contains the definition for proof graph signatures, in the variants
\<^item> Plain port graph
\<^item> Port graph with local hypotheses
\<^item> Labeled port graph
\<^item> Port graph with local constants
\<close>
locale Port_Graph_Signature =
fixes nodes :: "'node stream"
fixes inPorts :: "'node \<Rightarrow> 'inPort fset"
fixes outPorts :: "'node \<Rightarrow> 'outPort fset"
locale Port_Graph_Signature_Scoped =
Port_Graph_Signature +
fixes hyps :: "'node \<Rightarrow> 'outPort \<rightharpoonup> 'inPort"
assumes hyps_correct: "hyps n p1 = Some p2 \<Longrightarrow> p1 |\<in>| outPorts n \<and> p2 |\<in>| inPorts n"
begin
inductive_set hyps_for' :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'outPort set" for n p
where "hyps n h = Some p \<Longrightarrow> h \<in> hyps_for' n p"
lemma hyps_for'_subset: "hyps_for' n p \<subseteq> fset (outPorts n)"
- using hyps_correct by (meson hyps_for'.cases notin_fset subsetI)
+ using hyps_correct by (meson hyps_for'.cases subsetI)
context includes fset.lifting
begin
lift_definition hyps_for :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'outPort fset" is hyps_for'
by (meson finite_fset hyps_for'_subset rev_finite_subset)
lemma hyps_for_simp[simp]: "h |\<in>| hyps_for n p \<longleftrightarrow> hyps n h = Some p"
by transfer (simp add: hyps_for'.simps)
lemma hyps_for_simp'[simp]: "h \<in> fset (hyps_for n p) \<longleftrightarrow> hyps n h = Some p"
by transfer (simp add: hyps_for'.simps)
lemma hyps_for_collect: "fset (hyps_for n p) = {h . hyps n h = Some p}"
by auto
end
lemma hyps_for_subset: "hyps_for n p |\<subseteq>| outPorts n"
using hyps_for'_subset
- by (fastforce simp add: fmember_iff_member_fset hyps_for.rep_eq simp del: hyps_for_simp hyps_for_simp')
+ by (fastforce simp add: hyps_for.rep_eq simp del: hyps_for_simp hyps_for_simp')
end
locale Labeled_Signature =
Port_Graph_Signature_Scoped +
fixes labelsIn :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'form"
fixes labelsOut :: "'node \<Rightarrow> 'outPort \<Rightarrow> 'form"
locale Port_Graph_Signature_Scoped_Vars =
Port_Graph_Signature nodes inPorts outPorts +
Abstract_Formulas freshenLC renameLCs lconsts closed subst subst_lconsts subst_renameLCs anyP
for nodes :: "'node stream" and inPorts :: "'node \<Rightarrow> 'inPort fset" and outPorts :: "'node \<Rightarrow> 'outPort fset"
and freshenLC :: "nat \<Rightarrow> 'var \<Rightarrow> 'var"
and renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> 'form \<Rightarrow> 'form"
and lconsts :: "'form \<Rightarrow> 'var set"
and closed :: "'form \<Rightarrow> bool"
and subst :: "'subst \<Rightarrow> 'form \<Rightarrow> 'form"
and subst_lconsts :: "'subst \<Rightarrow> 'var set"
and subst_renameLCs :: "('var \<Rightarrow> 'var) \<Rightarrow> ('subst \<Rightarrow> 'subst)"
and anyP :: "'form" +
fixes local_vars :: "'node \<Rightarrow> 'inPort \<Rightarrow> 'var set"
end
diff --git a/thys/Incredible_Proof_Machine/Incredible_Trees.thy b/thys/Incredible_Proof_Machine/Incredible_Trees.thy
--- a/thys/Incredible_Proof_Machine/Incredible_Trees.thy
+++ b/thys/Incredible_Proof_Machine/Incredible_Trees.thy
@@ -1,642 +1,642 @@
theory Incredible_Trees
imports
"HOL-Library.Sublist"
"HOL-Library.Countable"
Entailment
Rose_Tree
Abstract_Rules_To_Incredible
begin
text \<open>This theory defines incredible trees, which carry roughly the same information
as a (tree-shaped) incredible graph, but where the structure is still given by the data type,
and not by a set of edges etc.\<close>
text \<open>
Tree-shape, but incredible-graph-like content (port names, explicit annotation and substitution)
\<close>
datatype ('form,'rule,'subst,'var) itnode =
I (iNodeOf': "('form, 'rule) graph_node")
(iOutPort': "'form reg_out_port")
(iAnnot': "nat")
(iSubst': "'subst")
| H (iAnnot': "nat")
(iSubst': "'subst")
abbreviation "INode n p i s ants \<equiv> RNode (I n p i s) ants"
abbreviation "HNode i s ants \<equiv> RNode (H i s) ants"
type_synonym ('form,'rule,'subst,'var) itree = "('form,'rule,'subst,'var) itnode rose_tree"
fun iNodeOf where
"iNodeOf (INode n p i s ants) = n"
| "iNodeOf (HNode i s ants) = Helper"
context Abstract_Formulas begin
fun iOutPort where
"iOutPort (INode n p i s ants) = p"
| "iOutPort (HNode i s ants) = anyP"
end
fun iAnnot where "iAnnot it = iAnnot' (root it)"
fun iSubst where "iSubst it = iSubst' (root it)"
fun iAnts where "iAnts it = children it"
type_synonym ('form, 'rule, 'subst) fresh_check = "('form, 'rule) graph_node \<Rightarrow> nat \<Rightarrow> 'subst \<Rightarrow> 'form entailment \<Rightarrow> bool"
context Abstract_Task
begin
text \<open>The well-formedness of the tree. The first argument can be varied, depending on whether we
are interested in the local freshness side-conditions or not.\<close>
inductive iwf :: "('form, 'rule, 'subst) fresh_check \<Rightarrow> ('form,'rule,'subst,'var) itree \<Rightarrow> 'form entailment \<Rightarrow> bool"
for fc
where
iwf: "\<lbrakk>
n \<in> sset nodes;
Reg p |\<in>| outPorts n;
list_all2 (\<lambda> ip t. iwf fc t ((\<lambda> h . subst s (freshen i (labelsOut n h))) |`| hyps_for n ip |\<union>| \<Gamma> \<turnstile> subst s (freshen i (labelsIn n ip))))
(inPorts' n) ants;
fc n i s (\<Gamma> \<turnstile> c);
c = subst s (freshen i p)
\<rbrakk> \<Longrightarrow> iwf fc (INode n p i s ants) (\<Gamma> \<turnstile> c)"
| iwfH: "\<lbrakk>
c |\<notin>| ass_forms;
c |\<in>| \<Gamma>;
c = subst s (freshen i anyP)
\<rbrakk> \<Longrightarrow> iwf fc (HNode i s []) (\<Gamma> \<turnstile> c)"
lemma iwf_subst_freshen_outPort:
"iwf lc ts ent \<Longrightarrow>
snd ent = subst (iSubst ts) (freshen (iAnnot ts) (iOutPort ts))"
by (auto elim: iwf.cases)
definition all_local_vars :: "('form, 'rule) graph_node \<Rightarrow> 'var set" where
"all_local_vars n = \<Union>(local_vars n ` fset (inPorts n))"
lemma all_local_vars_Helper[simp]:
"all_local_vars Helper = {}"
unfolding all_local_vars_def by simp
lemma all_local_vars_Assumption[simp]:
"all_local_vars (Assumption c) = {}"
unfolding all_local_vars_def by simp
text \<open>Local freshness side-conditions, corresponding what we have in the
theory \<open>Natural_Deduction\<close>.\<close>
inductive local_fresh_check :: "('form, 'rule, 'subst) fresh_check" where
"\<lbrakk>\<And> f. f |\<in>| \<Gamma> \<Longrightarrow> freshenLC i ` (all_local_vars n) \<inter> lconsts f = {};
freshenLC i ` (all_local_vars n) \<inter> subst_lconsts s = {}
\<rbrakk> \<Longrightarrow> local_fresh_check n i s (\<Gamma> \<turnstile> c)"
abbreviation "local_iwf \<equiv> iwf local_fresh_check"
text \<open>No freshness side-conditions. Used with the tree that comes out of
\<open>globalize\<close>, where we establish the (global) freshness conditions
separately.\<close>
inductive no_fresh_check :: "('form, 'rule, 'subst) fresh_check" where
"no_fresh_check n i s (\<Gamma> \<turnstile> c)"
abbreviation "plain_iwf \<equiv> iwf no_fresh_check"
fun isHNode where
"isHNode (HNode _ _ _ ) = True"
|"isHNode _ = False"
lemma iwf_edge_match:
assumes "iwf fc t ent"
assumes "is@[i] \<in> it_paths t"
shows "subst (iSubst (tree_at t (is@[i]))) (freshen (iAnnot (tree_at t (is@[i]))) (iOutPort (tree_at t (is@[i]))))
= subst (iSubst (tree_at t is)) (freshen (iAnnot (tree_at t is)) (a_conc (inPorts' (iNodeOf (tree_at t is)) ! i)))"
using assms
apply (induction arbitrary: "is" i)
apply (auto elim!: it_paths_SnocE)[1]
apply (rename_tac "is" i)
apply (case_tac "is")
apply (auto dest!: list_all2_nthD2)[1]
using iwf_subst_freshen_outPort
apply (solves \<open>(auto)[1]\<close>)
apply (auto elim!: it_paths_ConsE dest!: list_all2_nthD2)[1]
using it_path_SnocI
apply (solves blast)
apply (solves auto)
done
lemma iwf_length_inPorts:
assumes "iwf fc t ent"
assumes "is \<in> it_paths t"
shows "length (iAnts (tree_at t is)) \<le> length (inPorts' (iNodeOf (tree_at t is)))"
using assms
by (induction arbitrary: "is" rule: iwf.induct)
(auto elim!: it_paths_RNodeE dest: list_all2_lengthD list_all2_nthD2)
lemma iwf_local_not_in_subst:
assumes "local_iwf t ent"
assumes "is \<in> it_paths t"
assumes "var \<in> all_local_vars (iNodeOf (tree_at t is))"
shows "freshenLC (iAnnot (tree_at t is)) var \<notin> subst_lconsts (iSubst (tree_at t is))"
using assms
by (induction arbitrary: "is" rule: iwf.induct)
(auto 4 4 elim!: it_paths_RNodeE local_fresh_check.cases dest: list_all2_lengthD list_all2_nthD2)
lemma iwf_length_inPorts_not_HNode:
assumes "iwf fc t ent"
assumes "is \<in> it_paths t"
assumes "\<not> (isHNode (tree_at t is))"
shows "length (iAnts (tree_at t is)) = length (inPorts' (iNodeOf (tree_at t is)))"
using assms
by (induction arbitrary: "is" rule: iwf.induct)
(auto 4 4 elim!: it_paths_RNodeE dest: list_all2_lengthD list_all2_nthD2)
lemma iNodeOf_outPorts:
"iwf fc t ent \<Longrightarrow> is \<in> it_paths t \<Longrightarrow> outPorts (iNodeOf (tree_at t is)) = {||} \<Longrightarrow> False"
by (induction arbitrary: "is" rule: iwf.induct)
(auto 4 4 elim!: it_paths_RNodeE dest: list_all2_lengthD list_all2_nthD2)
lemma iNodeOf_tree_at:
"iwf fc t ent \<Longrightarrow> is \<in> it_paths t \<Longrightarrow> iNodeOf (tree_at t is) \<in> sset nodes"
by (induction arbitrary: "is" rule: iwf.induct)
(auto 4 4 elim!: it_paths_RNodeE dest: list_all2_lengthD list_all2_nthD2)
lemma iwf_outPort:
assumes "iwf fc t ent"
assumes "is \<in> it_paths t"
shows "Reg (iOutPort (tree_at t is)) |\<in>| outPorts (iNodeOf (tree_at t is))"
using assms
by (induction arbitrary: "is" rule: iwf.induct)
(auto 4 4 elim!: it_paths_RNodeE dest: list_all2_lengthD list_all2_nthD2)
inductive_set hyps_along for t "is" where
"prefix (is'@[i]) is \<Longrightarrow>
i < length (inPorts' (iNodeOf (tree_at t is'))) \<Longrightarrow>
hyps (iNodeOf (tree_at t is')) h = Some (inPorts' (iNodeOf (tree_at t is')) ! i) \<Longrightarrow>
subst (iSubst (tree_at t is')) (freshen (iAnnot (tree_at t is')) (labelsOut (iNodeOf (tree_at t is')) h)) \<in> hyps_along t is"
lemma hyps_along_Nil[simp]: "hyps_along t [] = {}"
by (auto simp add: hyps_along.simps)
lemma prefix_app_Cons_elim:
assumes "prefix (xs@[y]) (z#zs)"
obtains "xs = []" and "y = z"
| xs' where "xs = z#xs'" and "prefix (xs'@[y]) zs"
using assms by (cases xs) auto
lemma hyps_along_Cons:
assumes "iwf fc t ent"
assumes "i#is \<in> it_paths t"
shows "hyps_along t (i#is) =
(\<lambda>h. subst (iSubst t) (freshen (iAnnot t) (labelsOut (iNodeOf t) h))) ` fset (hyps_for (iNodeOf t) (inPorts' (iNodeOf t) ! i))
\<union> hyps_along (iAnts t ! i) is" (is "?S1 = ?S2 \<union> ?S3")
proof-
from assms
have "i < length (iAnts t)" and "is \<in> it_paths (iAnts t ! i)"
by (auto elim: it_paths_ConsE)
let "?t'" = "iAnts t ! i"
show ?thesis
proof (rule; rule)
fix x
assume "x \<in> hyps_along t (i # is)"
then obtain is' i' h where
"prefix (is'@[i']) (i#is)"
and "i' < length (inPorts' (iNodeOf (tree_at t is')))"
and "hyps (iNodeOf (tree_at t is')) h = Some (inPorts' (iNodeOf (tree_at t is')) ! i')"
and [simp]: "x = subst (iSubst (tree_at t is')) (freshen (iAnnot (tree_at t is')) (labelsOut (iNodeOf (tree_at t is')) h))"
by (auto elim!: hyps_along.cases)
from this(1)
show "x \<in> ?S2 \<union> ?S3"
proof(cases rule: prefix_app_Cons_elim)
assume "is' = []" and "i' = i"
with \<open>hyps (iNodeOf (tree_at t is')) h = Some _\<close>
have "x \<in> ?S2" by auto
thus ?thesis..
next
fix is''
assume [simp]: "is' = i # is''" and "prefix (is'' @ [i']) is"
have "tree_at t is' = tree_at ?t' is''" by simp
note \<open>prefix (is'' @ [i']) is\<close>
\<open>i' < length (inPorts' (iNodeOf (tree_at t is')))\<close>
\<open>hyps (iNodeOf (tree_at t is')) h = Some (inPorts' (iNodeOf (tree_at t is')) ! i')\<close>
from this[unfolded \<open>tree_at t is' = tree_at ?t' is''\<close>]
have "subst (iSubst (tree_at (iAnts t ! i) is'')) (freshen (iAnnot (tree_at (iAnts t ! i) is'')) (labelsOut (iNodeOf (tree_at (iAnts t ! i) is'')) h))
\<in> hyps_along (iAnts t ! i) is" by (rule hyps_along.intros)
hence "x \<in> ?S3" by simp
thus ?thesis..
qed
next
fix x
assume "x \<in> ?S2 \<union> ?S3"
thus "x \<in> ?S1"
proof
have "prefix ([]@[i]) (i#is)" by simp
moreover
from \<open>iwf _ t _\<close>
have "length (iAnts t) \<le> length (inPorts' (iNodeOf (tree_at t []))) "
by cases (auto dest: list_all2_lengthD)
with \<open>i < _\<close>
have "i < length (inPorts' (iNodeOf (tree_at t [])))" by simp
moreover
assume "x \<in> ?S2"
then obtain h where "h |\<in>| hyps_for (iNodeOf t) (inPorts' (iNodeOf t) ! i)"
and [simp]: "x = subst (iSubst t) (freshen (iAnnot t) (labelsOut (iNodeOf t) h))" by auto
from this(1)
have "hyps (iNodeOf (tree_at t [])) h = Some (inPorts' (iNodeOf (tree_at t [])) ! i)" by simp
ultimately
have "subst (iSubst (tree_at t [])) (freshen (iAnnot (tree_at t [])) (labelsOut (iNodeOf (tree_at t [])) h)) \<in> hyps_along t (i # is)"
by (rule hyps_along.intros)
thus "x \<in> hyps_along t (i # is)" by simp
next
assume "x \<in> ?S3"
thus "x \<in> ?S1"
apply (auto simp add: hyps_along.simps)
apply (rule_tac x = "i#is'" in exI)
apply auto
done
qed
qed
qed
lemma iwf_hyps_exist:
assumes "iwf lc it ent"
assumes "is \<in> it_paths it"
assumes "tree_at it is = (HNode i s ants')"
assumes "fst ent |\<subseteq>| ass_forms"
shows "subst s (freshen i anyP) \<in> hyps_along it is"
proof-
from assms(1,2,3)
have "subst s (freshen i anyP) \<in> hyps_along it is
\<or> subst s (freshen i anyP) |\<in>| fst ent
\<and> subst s (freshen i anyP) |\<notin>| ass_forms"
proof(induction arbitrary: "is" rule: iwf.induct)
case (iwf n p s' a' \<Gamma> ants c "is")
have "iwf lc (INode n p a' s' ants) (\<Gamma> \<turnstile> c)"
using iwf(1,2,3,4,5)
by (auto intro!: iwf.intros elim!: list_all2_mono)
show ?case
proof(cases "is")
case Nil
with \<open>tree_at (INode n p a' s' ants) is = HNode i s ants'\<close>
show ?thesis by auto
next
case (Cons i' "is'")
with \<open>is \<in> it_paths (INode n p a' s' ants)\<close>
have "i' < length ants" and "is' \<in> it_paths (ants ! i')"
by (auto elim: it_paths_ConsE)
let ?\<Gamma>' = "(\<lambda>h. subst s' (freshen a' (labelsOut n h))) |`| hyps_for n (inPorts' n ! i')"
from \<open>tree_at (INode n p a' s' ants) is = HNode i s ants'\<close>
have "tree_at (ants ! i') is' = HNode i s ants'" using Cons by simp
from iwf.IH \<open>i' < length ants\<close> \<open>is' \<in> it_paths (ants ! i')\<close> this
have "subst s (freshen i anyP) \<in> hyps_along (ants ! i') is'
\<or> subst s (freshen i anyP) |\<in>| ?\<Gamma>' |\<union>| \<Gamma> \<and> subst s (freshen i anyP) |\<notin>| ass_forms"
by (auto dest: list_all2_nthD2)
moreover
from \<open>is \<in> it_paths (INode n p a' s' ants)\<close>
have "hyps_along (INode n p a' s' ants) is = fset ?\<Gamma>' \<union> hyps_along (ants ! i') is'"
using \<open>is = _\<close>
by (simp add: hyps_along_Cons[OF \<open>iwf lc (INode n p a' s' ants) (\<Gamma> \<turnstile> c)\<close>])
ultimately
show ?thesis by auto
qed
next
case (iwfH c \<Gamma> s' i' "is")
hence [simp]: "is = []" "i' = i" "s' = s" by simp_all
from \<open>c = subst s' (freshen i' anyP)\<close> \<open>c |\<in>| \<Gamma>\<close> \<open>c |\<notin>| ass_forms\<close>
show ?case by simp
qed
with assms(4)
show ?thesis by blast
qed
definition hyp_port_for' :: "('form, 'rule, 'subst, 'var) itree \<Rightarrow> nat list \<Rightarrow> 'form \<Rightarrow> nat list \<times> nat \<times> ('form, 'var) out_port" where
"hyp_port_for' t is f = (SOME x.
(case x of (is', i, h) \<Rightarrow>
prefix (is' @ [i]) is \<and>
i < length (inPorts' (iNodeOf (tree_at t is'))) \<and>
hyps (iNodeOf (tree_at t is')) h = Some (inPorts' (iNodeOf (tree_at t is')) ! i) \<and>
f = subst (iSubst (tree_at t is')) (freshen (iAnnot (tree_at t is')) (labelsOut (iNodeOf (tree_at t is')) h))
))"
lemma hyp_port_for_spec':
assumes "f \<in> hyps_along t is"
shows "(case hyp_port_for' t is f of (is', i, h) \<Rightarrow>
prefix (is' @ [i]) is \<and>
i < length (inPorts' (iNodeOf (tree_at t is'))) \<and>
hyps (iNodeOf (tree_at t is')) h = Some (inPorts' (iNodeOf (tree_at t is')) ! i) \<and>
f = subst (iSubst (tree_at t is')) (freshen (iAnnot (tree_at t is')) (labelsOut (iNodeOf (tree_at t is')) h)))"
using assms unfolding hyps_along.simps hyp_port_for'_def by -(rule someI_ex, blast)
definition hyp_port_path_for :: "('form, 'rule, 'subst, 'var) itree \<Rightarrow> nat list \<Rightarrow> 'form \<Rightarrow> nat list"
where "hyp_port_path_for t is f = fst (hyp_port_for' t is f)"
definition hyp_port_i_for :: "('form, 'rule, 'subst, 'var) itree \<Rightarrow> nat list \<Rightarrow> 'form \<Rightarrow> nat"
where "hyp_port_i_for t is f = fst (snd (hyp_port_for' t is f))"
definition hyp_port_h_for :: "('form, 'rule, 'subst, 'var) itree \<Rightarrow> nat list \<Rightarrow> 'form \<Rightarrow> ('form, 'var) out_port"
where "hyp_port_h_for t is f = snd (snd (hyp_port_for' t is f))"
lemma hyp_port_prefix:
assumes "f \<in> hyps_along t is"
shows "prefix (hyp_port_path_for t is f@[hyp_port_i_for t is f]) is"
using hyp_port_for_spec'[OF assms] unfolding hyp_port_path_for_def hyp_port_i_for_def by auto
lemma hyp_port_strict_prefix:
assumes "f \<in> hyps_along t is"
shows "strict_prefix (hyp_port_path_for t is f) is"
using hyp_port_prefix[OF assms] by (simp add: strict_prefixI' prefix_order.dual_order.strict_trans1)
lemma hyp_port_it_paths:
assumes "is \<in> it_paths t"
assumes "f \<in> hyps_along t is"
shows "hyp_port_path_for t is f \<in> it_paths t"
using assms by (rule it_paths_strict_prefix[OF _ hyp_port_strict_prefix] )
lemma hyp_port_hyps:
assumes "f \<in> hyps_along t is"
shows "hyps (iNodeOf (tree_at t (hyp_port_path_for t is f))) (hyp_port_h_for t is f) = Some (inPorts' (iNodeOf (tree_at t (hyp_port_path_for t is f))) ! hyp_port_i_for t is f)"
using hyp_port_for_spec'[OF assms] unfolding hyp_port_path_for_def hyp_port_i_for_def hyp_port_h_for_def by auto
lemma hyp_port_outPort:
assumes "f \<in> hyps_along t is"
shows "(hyp_port_h_for t is f) |\<in>| outPorts (iNodeOf (tree_at t (hyp_port_path_for t is f)))"
using hyps_correct[OF hyp_port_hyps[OF assms]]..
lemma hyp_port_eq:
assumes "f \<in> hyps_along t is"
shows "f = subst (iSubst (tree_at t (hyp_port_path_for t is f))) (freshen (iAnnot (tree_at t (hyp_port_path_for t is f))) (labelsOut (iNodeOf (tree_at t (hyp_port_path_for t is f))) (hyp_port_h_for t is f)))"
using hyp_port_for_spec'[OF assms] unfolding hyp_port_path_for_def hyp_port_i_for_def hyp_port_h_for_def by auto
definition isidx :: "nat list \<Rightarrow> nat" where "isidx xs = to_nat (Some xs)"
definition v_away :: "nat" where "v_away = to_nat (None :: nat list option)"
lemma isidx_inj[simp]: "isidx xs = isidx ys \<longleftrightarrow> xs = ys"
unfolding isidx_def by simp
lemma isidx_v_away[simp]: "isidx xs \<noteq> v_away"
unfolding isidx_def v_away_def by simp
definition mapWithIndex where "mapWithIndex f xs = map (\<lambda> (i,t) . f i t) (List.enumerate 0 xs)"
lemma mapWithIndex_cong [fundef_cong]:
"xs = ys \<Longrightarrow> (\<And>x i. x \<in> set ys \<Longrightarrow> f i x = g i x) \<Longrightarrow> mapWithIndex f xs = mapWithIndex g ys"
unfolding mapWithIndex_def by (auto simp add: in_set_enumerate_eq)
lemma mapWithIndex_Nil[simp]: "mapWithIndex f [] = []"
unfolding mapWithIndex_def by simp
lemma length_mapWithIndex[simp]: "length (mapWithIndex f xs) = length xs"
unfolding mapWithIndex_def by simp
lemma nth_mapWithIndex[simp]: "i < length xs \<Longrightarrow> mapWithIndex f xs ! i = f i (xs ! i)"
unfolding mapWithIndex_def by (auto simp add: nth_enumerate_eq)
lemma list_all2_mapWithIndex2E:
assumes "list_all2 P as bs"
assumes "\<And> i a b . i < length bs \<Longrightarrow> P a b \<Longrightarrow> Q a (f i b)"
shows "list_all2 Q as (mapWithIndex f bs)"
using assms(1)
by (auto simp add: list_all2_conv_all_nth mapWithIndex_def nth_enumerate_eq intro: assms(2) split: prod.split)
text \<open>The globalize function, which renames all local constants so that they cannot clash with
local constants occurring anywhere else in the tree.\<close>
fun globalize_node :: "nat list \<Rightarrow> ('var \<Rightarrow> 'var) \<Rightarrow> ('form,'rule,'subst,'var) itnode \<Rightarrow> ('form,'rule,'subst,'var) itnode" where
"globalize_node is f (I n p i s) = I n p (isidx is) (subst_renameLCs f s)"
| "globalize_node is f (H i s) = H (isidx is) (subst_renameLCs f s)"
fun globalize :: "nat list \<Rightarrow> ('var \<Rightarrow> 'var) \<Rightarrow> ('form,'rule,'subst,'var) itree \<Rightarrow> ('form,'rule,'subst,'var) itree" where
"globalize is f (RNode r ants) = RNode
(globalize_node is f r)
(mapWithIndex (\<lambda> i' t.
globalize (is@[i'])
(rerename (a_fresh (inPorts' (iNodeOf (RNode r ants)) ! i'))
(iAnnot (RNode r ants)) (isidx is) f)
t
) ants)"
lemma iAnnot'_globalize_node[simp]: "iAnnot' (globalize_node is f n) = isidx is"
by (cases n) auto
lemma iAnnot_globalize:
assumes "is' \<in> it_paths (globalize is f t)"
shows "iAnnot (tree_at (globalize is f t) is') = isidx (is@is')"
using assms
by (induction t arbitrary: f "is" is') (auto elim!: it_paths_RNodeE)
lemma all_local_consts_listed':
assumes "n \<in> sset nodes"
assumes "p |\<in>| inPorts n"
shows "lconsts (a_conc p) \<union> (\<Union>(lconsts ` fset (a_hyps p))) \<subseteq> a_fresh p "
using assms
- by (auto simp add: nodes_def stream.set_map lconsts_anyP closed_no_lconsts conclusions_closed fmember_iff_member_fset f_antecedent_def dest!: all_local_consts_listed)
+ by (auto simp add: nodes_def stream.set_map lconsts_anyP closed_no_lconsts conclusions_closed f_antecedent_def dest!: all_local_consts_listed)
lemma no_local_consts_in_consequences':
"n \<in> sset nodes \<Longrightarrow> Reg p |\<in>| outPorts n \<Longrightarrow> lconsts p = {}"
using no_local_consts_in_consequences
by (auto simp add: nodes_def lconsts_anyP closed_no_lconsts assumptions_closed stream.set_map f_consequent_def)
lemma iwf_globalize:
assumes "local_iwf t (\<Gamma> \<turnstile> c)"
shows "plain_iwf (globalize is f t) (renameLCs f |`| \<Gamma> \<turnstile> renameLCs f c)"
using assms
proof (induction t "\<Gamma> \<turnstile> c" arbitrary: "is" f \<Gamma> c rule: iwf.induct)
case (iwf n p s i \<Gamma> ants c "is" f)
note \<open>n \<in> sset nodes\<close>
moreover
note \<open>Reg p |\<in>| outPorts n\<close>
moreover
{ fix i'
let ?V = "a_fresh (inPorts' n ! i')"
let ?f' = "rerename ?V i (isidx is) f"
let ?t = "globalize (is @ [i']) ?f' (ants ! i')"
let ?ip = "inPorts' n ! i'"
let ?\<Gamma>' = "(\<lambda>h. subst (subst_renameLCs f s) (freshen (isidx is) (labelsOut n h))) |`| hyps_for n ?ip"
let ?c' = "subst (subst_renameLCs f s) (freshen (isidx is) (labelsIn n ?ip))"
assume "i' < length (inPorts' n)"
hence "(inPorts' n ! i') |\<in>| inPorts n" by (simp add: inPorts_fset_of)
from \<open>i' < length (inPorts' n)\<close>
have subset_V: "?V \<subseteq> all_local_vars n"
unfolding all_local_vars_def
by (auto simp add: inPorts_fset_of set_conv_nth)
from \<open>local_fresh_check n i s (\<Gamma> \<turnstile> c)\<close>
have "freshenLC i ` all_local_vars n \<inter> subst_lconsts s = {}"
by (rule local_fresh_check.cases) simp
hence "freshenLC i ` ?V \<inter> subst_lconsts s = {}"
using subset_V by auto
hence rerename_subst: "subst_renameLCs ?f' s = subst_renameLCs f s"
by (rule rerename_subst_noop)
from all_local_consts_listed'[OF \<open> n \<in> sset nodes\<close> \<open>(inPorts' n ! i') |\<in>| inPorts n\<close>]
have subset_conc: "lconsts (a_conc (inPorts' n ! i')) \<subseteq> ?V"
and subset_hyp': "\<And> hyp . hyp |\<in>| a_hyps (inPorts' n ! i') \<Longrightarrow> lconsts hyp \<subseteq> ?V"
- by (auto simp add: fmember_iff_member_fset)
+ by auto
from List.list_all2_nthD[OF \<open>list_all2 _ _ _\<close> \<open>i' < length (inPorts' n)\<close>,simplified]
have "plain_iwf ?t
(renameLCs ?f' |`| ((\<lambda>h. subst s (freshen i (labelsOut n h))) |`| hyps_for n ?ip |\<union>| \<Gamma>) \<turnstile>
renameLCs ?f' (subst s (freshen i (a_conc ?ip))))"
by simp
also have "renameLCs ?f' |`| ((\<lambda>h. subst s (freshen i (labelsOut n h))) |`| hyps_for n ?ip |\<union>| \<Gamma>)
= (\<lambda>x. subst (subst_renameLCs ?f' s) (renameLCs ?f' (freshen i (labelsOut n x)))) |`| hyps_for n ?ip |\<union>| renameLCs ?f' |`| \<Gamma>"
by (simp add: fimage_fimage fimage_funion comp_def rename_subst)
also have "renameLCs ?f' |`| \<Gamma> = renameLCs f |`| \<Gamma>"
proof(rule fimage_cong[OF refl])
fix x
assume "x |\<in>| \<Gamma>"
with \<open>local_fresh_check n i s (\<Gamma> \<turnstile> c)\<close>
have "freshenLC i ` all_local_vars n \<inter> lconsts x = {}"
by (elim local_fresh_check.cases) simp
hence "freshenLC i ` ?V \<inter> lconsts x = {}"
using subset_V by auto
thus "renameLCs ?f' x = renameLCs f x"
by (rule rerename_rename_noop)
qed
also have "(\<lambda>x. subst (subst_renameLCs ?f' s) (renameLCs ?f' (freshen i (labelsOut n x)))) |`| hyps_for n ?ip = ?\<Gamma>'"
proof(rule fimage_cong[OF refl])
fix hyp
assume "hyp |\<in>| hyps_for n (inPorts' n ! i')"
hence "labelsOut n hyp |\<in>| a_hyps (inPorts' n ! i')"
apply (cases hyp)
apply (solves simp)
apply (cases n)
apply (auto split: if_splits)
done
from subset_hyp'[OF this]
have subset_hyp: "lconsts (labelsOut n hyp) \<subseteq> ?V".
show "subst (subst_renameLCs ?f' s) (renameLCs ?f' (freshen i (labelsOut n hyp))) =
subst (subst_renameLCs f s) (freshen (isidx is) (labelsOut n hyp))"
apply (simp add: freshen_def rename_rename rerename_subst)
apply (rule arg_cong[OF renameLCs_cong])
apply (auto dest: subsetD[OF subset_hyp])
done
qed
also have "renameLCs ?f' (subst s (freshen i (a_conc ?ip))) = subst (subst_renameLCs ?f' s) (renameLCs ?f' (freshen i (a_conc ?ip)))" by (simp add: rename_subst)
also have "... = ?c'"
apply (simp add: freshen_def rename_rename rerename_subst)
apply (rule arg_cong[OF renameLCs_cong])
apply (auto dest: subsetD[OF subset_conc])
done
finally
have "plain_iwf ?t (?\<Gamma>' |\<union>| renameLCs f |`| \<Gamma> \<turnstile> ?c')".
}
with list_all2_lengthD[OF \<open>list_all2 _ _ _\<close>]
have "list_all2
(\<lambda>ip t. plain_iwf t ((\<lambda>h. subst (subst_renameLCs f s)
(freshen (isidx is) (labelsOut n h))) |`| hyps_for n ip |\<union>| renameLCs f |`| \<Gamma> \<turnstile> subst (subst_renameLCs f s) (freshen (isidx is) (labelsIn n ip))))
(inPorts' n)
(mapWithIndex (\<lambda> i' t. globalize (is@[i']) (rerename (a_fresh (inPorts' n ! i')) i (isidx is) f) t) ants)"
by (auto simp add: list_all2_conv_all_nth)
moreover
have "no_fresh_check n (isidx is) (subst_renameLCs f s) (renameLCs f |`| \<Gamma> \<turnstile> renameLCs f c)"..
moreover
from \<open>n \<in> sset nodes\<close> \<open>Reg p |\<in>| outPorts n\<close>
have "lconsts p = {}" by (rule no_local_consts_in_consequences')
with \<open>c = subst s (freshen i p)\<close>
have "renameLCs f c = subst (subst_renameLCs f s) (freshen (isidx is) p)"
by (simp add: rename_subst rename_closed freshen_closed)
ultimately
show ?case
unfolding globalize.simps globalize_node.simps iNodeOf.simps iAnnot.simps itnode.sel rose_tree.sel Let_def
by (rule iwf.intros(1))
next
case (iwfH c \<Gamma> s i "is" f)
from \<open>c |\<notin>| ass_forms\<close>
have "renameLCs f c |\<notin>| ass_forms"
using assumptions_closed closed_no_lconsts lconsts_renameLCs rename_closed by fastforce
moreover
from \<open>c |\<in>| \<Gamma>\<close>
have "renameLCs f c |\<in>| renameLCs f |`| \<Gamma>" by auto
moreover
from \<open>c = subst s (freshen i anyP)\<close>
have "renameLCs f c = subst (subst_renameLCs f s) (freshen (isidx is) anyP)"
by (metis freshen_closed lconsts_anyP rename_closed rename_subst)
ultimately
show "plain_iwf (globalize is f (HNode i s [])) (renameLCs f |`| \<Gamma> \<turnstile> renameLCs f c)"
unfolding globalize.simps globalize_node.simps mapWithIndex_Nil Let_def
by (rule iwf.intros(2))
qed
definition fresh_at where
"fresh_at t xs =
(case rev xs of [] \<Rightarrow> {}
| (i#is') \<Rightarrow> freshenLC (iAnnot (tree_at t (rev is'))) ` (a_fresh (inPorts' (iNodeOf (tree_at t (rev is'))) ! i)))"
lemma fresh_at_Nil[simp]:
"fresh_at t [] = {}"
unfolding fresh_at_def by simp
lemma fresh_at_snoc[simp]:
"fresh_at t (is@[i]) = freshenLC (iAnnot (tree_at t is)) ` (a_fresh (inPorts' (iNodeOf (tree_at t is)) ! i))"
unfolding fresh_at_def by simp
lemma fresh_at_def':
"fresh_at t is =
(if is = [] then {}
else freshenLC (iAnnot (tree_at t (butlast is))) ` (a_fresh (inPorts' (iNodeOf (tree_at t (butlast is))) ! last is)))"
unfolding fresh_at_def by (auto split: list.split)
lemma fresh_at_Cons[simp]:
"fresh_at t (i#is) = (if is = [] then freshenLC (iAnnot t) ` (a_fresh (inPorts' (iNodeOf t) ! i)) else (let t' = iAnts t ! i in fresh_at t' is))"
unfolding fresh_at_def'
by (auto simp add: Let_def)
definition fresh_at_path where
"fresh_at_path t is = \<Union>(fresh_at t ` set (prefixes is))"
lemma fresh_at_path_Nil[simp]:
"fresh_at_path t [] = {}"
unfolding fresh_at_path_def by simp
lemma fresh_at_path_Cons[simp]:
"fresh_at_path t (i#is) = fresh_at t [i] \<union> fresh_at_path (iAnts t ! i) is"
unfolding fresh_at_path_def
by (fastforce split: if_splits)
lemma globalize_local_consts:
assumes "is' \<in> it_paths (globalize is f t)"
shows "subst_lconsts (iSubst (tree_at (globalize is f t) is')) \<subseteq>
fresh_at_path (globalize is f t) is' \<union> range f"
using assms
apply (induction "is" f t arbitrary: is' rule:globalize.induct)
apply (rename_tac "is" f r ants is')
apply (case_tac r)
apply (auto simp add: subst_lconsts_subst_renameLCs elim!: it_paths_RNodeE)
apply (solves \<open>force dest!: subsetD[OF range_rerename]\<close>)
apply (solves \<open>force dest!: subsetD[OF range_rerename]\<close>)
done
lemma iwf_globalize':
assumes "local_iwf t ent"
assumes "\<And> x. x |\<in>| fst ent \<Longrightarrow> closed x"
assumes "closed (snd ent)"
shows "plain_iwf (globalize is (freshenLC v_away) t) ent"
using assms
proof(induction ent rule: prod.induct)
case (Pair \<Gamma> c)
have "plain_iwf (globalize is (freshenLC v_away) t) (renameLCs (freshenLC v_away) |`| \<Gamma> \<turnstile> renameLCs (freshenLC v_away) c)"
by (rule iwf_globalize[OF Pair(1)])
also
from Pair(3) have "closed c" by simp
hence "renameLCs (freshenLC v_away) c = c" by (simp add: closed_no_lconsts rename_closed)
also
from Pair(2)
have "renameLCs (freshenLC v_away) |`| \<Gamma> = \<Gamma>"
- by (auto simp add: closed_no_lconsts rename_closed fmember_iff_member_fset image_iff)
+ by (auto simp add: closed_no_lconsts rename_closed image_iff)
finally show ?case.
qed
end
end
diff --git a/thys/Incredible_Proof_Machine/Indexed_FSet.thy b/thys/Incredible_Proof_Machine/Indexed_FSet.thy
--- a/thys/Incredible_Proof_Machine/Indexed_FSet.thy
+++ b/thys/Incredible_Proof_Machine/Indexed_FSet.thy
@@ -1,92 +1,92 @@
theory Indexed_FSet
imports
"HOL-Library.FSet"
begin
text \<open>It is convenient to address the members of a finite set by a natural number, and
also to convert a finite set to a list.\<close>
context includes fset.lifting
begin
lift_definition fset_from_list :: "'a list => 'a fset" is set by (rule finite_set)
lemma mem_fset_from_list[simp]: "x |\<in>| fset_from_list l \<longleftrightarrow> x \<in> set l" by transfer rule
lemma fimage_fset_from_list[simp]: "f |`| fset_from_list l = fset_from_list (map f l)" by transfer auto
lemma fset_fset_from_list[simp]: "fset (fset_from_list l) = set l" by transfer auto
lemmas fset_simps[simp] = set_simps[Transfer.transferred]
lemma size_fset_from_list[simp]: "distinct l \<Longrightarrow> size (fset_from_list l) = length l"
by (induction l) auto
definition list_of_fset :: "'a fset \<Rightarrow> 'a list" where
"list_of_fset s = (SOME l. fset_from_list l = s \<and> distinct l)"
lemma fset_from_list_of_fset[simp]: "fset_from_list (list_of_fset s) = s"
and distinct_list_of_fset[simp]: "distinct (list_of_fset s)"
unfolding atomize_conj list_of_fset_def
by (transfer, rule someI_ex, rule finite_distinct_list)
lemma length_list_of_fset[simp]: "length (list_of_fset s) = size s"
by (metis distinct_list_of_fset fset_from_list_of_fset size_fset_from_list)
lemma nth_list_of_fset_mem[simp]: "i < size s \<Longrightarrow> list_of_fset s ! i |\<in>| s"
by (metis fset_from_list_of_fset length_list_of_fset mem_fset_from_list nth_mem)
inductive indexed_fmember :: "'a \<Rightarrow> nat \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ |\<in>|\<^bsub>_\<^esub> _" [50,50,50] 50 ) where
"i < size s \<Longrightarrow> list_of_fset s ! i |\<in>|\<^bsub>i\<^esub> s"
lemma indexed_fmember_is_fmember: "x |\<in>|\<^bsub>i\<^esub> s \<Longrightarrow> x |\<in>| s"
proof (induction rule: indexed_fmember.induct)
case (1 i s)
hence "i < length (list_of_fset s)" by (metis length_list_of_fset)
hence "list_of_fset s ! i \<in> set (list_of_fset s)" by (rule nth_mem)
thus "list_of_fset s ! i |\<in>| s" by (metis mem_fset_from_list fset_from_list_of_fset)
qed
lemma fmember_is_indexed_fmember:
assumes "x |\<in>| s"
shows "\<exists>i. x |\<in>|\<^bsub>i\<^esub> s"
proof-
from assms
have "x \<in> set (list_of_fset s)" using mem_fset_from_list by fastforce
then obtain i where "i < length (list_of_fset s)" and "x = list_of_fset s ! i" by (metis in_set_conv_nth)
hence "x |\<in>|\<^bsub>i\<^esub> s" by (simp add: indexed_fmember.simps)
thus ?thesis..
qed
lemma indexed_fmember_unique: "x |\<in>|\<^bsub>i\<^esub> s \<Longrightarrow> y |\<in>|\<^bsub>j\<^esub> s \<Longrightarrow> x = y \<longleftrightarrow> i = j"
by (metis distinct_list_of_fset indexed_fmember.cases length_list_of_fset nth_eq_iff_index_eq)
definition indexed_members :: "'a fset \<Rightarrow> (nat \<times> 'a) list" where
"indexed_members s = zip [0..<size s] (list_of_fset s)"
lemma mem_set_indexed_members:
"(i,x) \<in> set (indexed_members s) \<longleftrightarrow> x |\<in>|\<^bsub>i\<^esub> s"
unfolding indexed_members_def indexed_fmember.simps
by (force simp add: set_zip)
lemma mem_set_indexed_members'[simp]:
"t \<in> set (indexed_members s) \<longleftrightarrow> snd t |\<in>|\<^bsub>fst t\<^esub> s"
by (cases t, simp add: mem_set_indexed_members)
definition fnth (infixl "|!|" 100) where
"s |!| n = list_of_fset s ! n"
lemma fnth_indexed_fmember: "i < size s \<Longrightarrow> s |!| i |\<in>|\<^bsub>i\<^esub> s"
unfolding fnth_def by (rule indexed_fmember.intros)
lemma indexed_fmember_fnth: "x |\<in>|\<^bsub>i\<^esub> s \<longleftrightarrow> (s |!| i = x \<and> i < size s)"
unfolding fnth_def by (metis indexed_fmember.simps)
end
definition fidx :: "'a fset \<Rightarrow> 'a \<Rightarrow> nat" where
"fidx s x = (SOME i. x |\<in>|\<^bsub>i\<^esub> s)"
lemma fidx_eq[simp]: "x |\<in>|\<^bsub>i\<^esub> s \<Longrightarrow> fidx s x = i"
unfolding fidx_def
by (rule someI2)(auto simp add: indexed_fmember_fnth fnth_def nth_eq_iff_index_eq)
lemma fidx_inj[simp]: "x |\<in>| s \<Longrightarrow> y |\<in>| s \<Longrightarrow> fidx s x = fidx s y \<longleftrightarrow> x = y"
by (auto dest!: fmember_is_indexed_fmember simp add: indexed_fmember_unique)
lemma inj_on_fidx: "inj_on (fidx vertices) (fset vertices)"
- by (rule inj_onI) (auto simp: fmember_iff_member_fset [symmetric])
+ by (rule inj_onI) simp
end
diff --git a/thys/IsaNet/infrastructure/Abstract_XOR.thy b/thys/IsaNet/infrastructure/Abstract_XOR.thy
--- a/thys/IsaNet/infrastructure/Abstract_XOR.thy
+++ b/thys/IsaNet/infrastructure/Abstract_XOR.thy
@@ -1,150 +1,150 @@
(*******************************************************************************
Project: IsaNet
Author: Tobias Klenze, ETH Zurich <tobias.klenze@inf.ethz.ch>
Version: JCSPaper.1.0
Isabelle Version: Isabelle2021-1
Copyright (c) 2022 Tobias Klenze
Licence: Mozilla Public License 2.0 (MPL) / BSD-3-Clause (dual license)
*******************************************************************************)
section \<open>Abstract XOR\<close>
theory "Abstract_XOR"
imports
"HOL.Finite_Set" "HOL-Library.FSet" "Message"
(*the latter half of this theory uses msgterm. If split it off, we have to add
declare fmember_iff_member_fset[simp]
in order to show the first half*)
begin
(******************************************************************************)
subsection\<open>Abstract XOR definition and lemmas\<close>
(******************************************************************************)
text\<open>We model xor as an operation on finite sets (fset). @{term "{||}"} is defined as the identity element.\<close>
text\<open>xor of two fsets is the symmetric difference\<close>
definition xor :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
"xor xs ys = (xs |\<union>| ys) |-| (xs |\<inter>| ys)"
lemma xor_singleton:
"xor xs {| z |} = (if z |\<in>| xs then xs |-| {| z |} else finsert z xs)"
"xor {| z |} xs = (if z |\<in>| xs then xs |-| {| z |} else finsert z xs)"
by (auto simp add: xor_def)
(*auto loops with this rule declared intro!. We could alternatively try using safe instead of auto*)
declare finsertCI[rule del]
declare finsertCI[intro]
lemma xor_assoc: "xor (xor xs ys) zs = xor xs (xor ys zs)"
by (auto simp add: xor_def)
lemma xor_commut: "xor xs ys = xor ys xs"
by (auto simp add: xor_def)
lemma xor_self_inv: "\<lbrakk>xor xs ys = zs; xs = ys\<rbrakk> \<Longrightarrow> zs = {||}"
by (auto simp add: xor_def)
lemma xor_self_inv': "xor xs xs = {||}"
by (auto simp add: xor_def)
lemma xor_self_inv''[dest!]: "xor xs ys = {||} \<Longrightarrow> xs = ys"
by (auto simp add: xor_def)
lemma xor_identity1[simp]: "xor xs {||} = xs"
by (auto simp add: xor_def)
lemma xor_identity2[simp]: "xor {||} xs = xs"
by (auto simp add: xor_def)
lemma xor_in: "z |\<in>| xs \<Longrightarrow> z |\<notin>| (xor xs {| z |})"
by (auto simp add: xor_singleton)
lemma xor_out: "z |\<notin>| xs \<Longrightarrow> z |\<in>| (xor xs {| z |})"
by (auto simp add: xor_singleton)
lemma xor_elem1[dest]: "\<lbrakk>x \<in> fset (xor X Y); x |\<notin>| X\<rbrakk> \<Longrightarrow> x |\<in>| Y"
by(auto simp add: xor_def)
lemma xor_elem2[dest]: "\<lbrakk>x \<in> fset (xor X Y); x |\<notin>| Y\<rbrakk> \<Longrightarrow> x |\<in>| X"
by(auto simp add: xor_def)
lemma xor_finsert_self: "xor (finsert x xs) {|x|} = xs - {| x |}"
by(auto simp add: xor_def)
(******************************************************************************)
subsection\<open>Lemmas refering to XOR and msgterm\<close>
(******************************************************************************)
lemma FS_contains_elem:
assumes "elem = f (FS zs_s)" "zs_s = xor zs_b {| elem |}" "\<And> x. size (f x) > size x"
shows "elem \<in> fset zs_b"
using assms(1)
apply(auto simp add: xor_def)
- using FS_mono assms notin_fset xor_singleton(1)
+ using FS_mono assms xor_singleton(1)
by (metis)
lemma FS_is_finsert_elem:
assumes "elem = f (FS zs_s)" "zs_s = xor zs_b {| elem |}" "\<And> x. size (f x) > size x"
shows "zs_b = finsert elem zs_s"
using assms FS_contains_elem finsert_fminus xor_singleton(1) FS_mono
by (metis FS_mono)
lemma FS_update_eq:
assumes "xs = f (FS (xor zs {|xs|}))"
and "ys = g (FS (xor zs {|ys|}))"
and "\<And> x. size (f x) > size x"
and "\<And> x. size (g x) > size x"
shows "xs = ys"
proof(rule ccontr)
assume elem_neq: "xs \<noteq> ys"
obtain zs_s1 zs_s2 where zs_defs:
"zs_s1 = xor zs {|xs|}" "zs_s2 = xor zs {|ys|}" by simp
have elems_contained_zs: "xs \<in> fset zs" "ys \<in> fset zs"
using assms FS_contains_elem by blast+
then have elems_elem: "ys |\<in>| zs_s1" "xs |\<in>| zs_s2"
using elem_neq by(auto simp add: xor_def zs_defs)
have zs_finsert: "finsert xs zs_s2 = zs_s2" "finsert ys zs_s1 = zs_s1"
using elems_elem by fastforce+
have f1: "\<forall>m f fa. \<not> sum fa (fset (finsert (m::msgterm) f)) < (fa m::nat)"
by (simp add: sum.insert_remove)
from assms(1-2) have "size xs > size (f (FS {| ys |}))" "size ys > size (g (FS {| xs |}))"
apply(simp_all add: zs_defs[symmetric])
using zs_finsert f1 by (metis (no_types) add_Suc_right assms(3-4) dual_order.strict_trans
less_add_Suc1 msgterm.size(17) not_less_eq size_fset_simps)+
then show False using assms(3,4) elems_elem
by (metis add.right_neutral add_Suc_right f1 less_add_Suc1 msgterm.size(17) not_less_eq
not_less_iff_gr_or_eq order.strict_trans size_fset_simps)
qed
declare fminusE[rule del]
declare finsertCI[rule del]
(*add back the removed rules without !*)
declare fminusE[elim]
declare finsertCI[intro]
(*currently not needed*)
lemma fset_size_le:
assumes "x \<in> fset xs"
shows "size x < Suc (\<Sum>x\<in>fset xs. Suc (size x))"
proof-
have "size x \<le> (\<Sum>x\<in>fset xs. size x)" using assms
by (auto intro: member_le_sum)
moreover have "(\<Sum>x\<in>fset xs. size x) < (\<Sum>x\<in>fset xs. Suc (size x))"
by (metis assms empty_iff finite_fset lessI sum_strict_mono)
ultimately show ?thesis by auto
qed
text\<open>We can show that xor is a commutative function.\<close>
(*not needed*)
locale abstract_xor
begin
sublocale comp_fun_commute xor
by(auto simp add: comp_fun_commute_def xor_def)
end
end
\ No newline at end of file
diff --git a/thys/IsaNet/infrastructure/Message.thy b/thys/IsaNet/infrastructure/Message.thy
--- a/thys/IsaNet/infrastructure/Message.thy
+++ b/thys/IsaNet/infrastructure/Message.thy
@@ -1,1191 +1,1184 @@
(*******************************************************************************
Title: HOL/Auth/Message
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1996 University of Cambridge
Datatypes of agents and messages;
Inductive relations "parts", "analz" and "synth"
********************************************************************************
Edited: Tobias Klenze, ETH Zurich <tobias.klenze@inf.ethz.ch>
Christoph Sprenger, ETH Zurich <sprenger@inf.ethz.ch>
Integrated and adapted for security protocol refinement and to add a constructor for finite sets.
*******************************************************************************)
section \<open>Theory of ASes and Messages for Security Protocols\<close>
theory Message imports Keys "HOL-Library.Sublist" "HOL.Finite_Set" "HOL-Library.FSet"
begin
datatype msgterm =
\<epsilon> \<comment> \<open>Empty message. Used for instance to denote non-existent interface\<close>
| AS as \<comment> \<open>Autonomous System identifier, i.e. agents. Note that AS is an alias of nat\<close>
| Num nat \<comment> \<open>Ordinary integers, timestamps, ...\<close>
| Key key \<comment> \<open>Crypto keys\<close>
| Nonce nonce \<comment> \<open>Unguessable nonces\<close>
| L "msgterm list" \<comment> \<open>Lists\<close>
| FS "msgterm fset" \<comment> \<open>Finite Sets. Used to represent XOR values.\<close>
| MPair msgterm msgterm \<comment> \<open>Compound messages\<close>
| Hash msgterm \<comment> \<open>Hashing\<close>
| Crypt key msgterm \<comment> \<open>Encryption, public- or shared-key\<close>
text \<open>Syntax sugar\<close>
syntax
"_MTuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(2\<langle>_,/ _\<rangle>)")
syntax (xsymbols)
"_MTuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(2\<langle>_,/ _\<rangle>)")
translations
"\<langle>x, y, z\<rangle>" \<rightleftharpoons> "\<langle>x, \<langle>y, z\<rangle>\<rangle>"
"\<langle>x, y\<rangle>" \<rightleftharpoons> "CONST MPair x y"
syntax
"_MHF" :: "['a, 'b , 'c, 'd, 'e] \<Rightarrow> 'a * 'b * 'c * 'd * 'e" ("(5HF\<lhd>_,/ _,/ _,/ _,/ _\<rhd>)")
abbreviation
Mac :: "[msgterm,msgterm] \<Rightarrow> msgterm" ("(4Mac[_] /_)" [0, 1000])
where
\<comment> \<open>Message Y paired with a MAC computed with the help of X\<close>
"Mac[X] Y \<equiv> Hash \<langle>X,Y\<rangle>"
abbreviation macKey where "macKey a \<equiv> Key (macK a)"
definition
keysFor :: "msgterm set \<Rightarrow> key set"
where
\<comment> \<open>Keys useful to decrypt elements of a message set\<close>
"keysFor H \<equiv> invKey ` {K. \<exists>X. Crypt K X \<in> H}"
subsubsection \<open>Inductive Definition of "All Parts" of a Message\<close>
inductive_set
parts :: "msgterm set \<Rightarrow> msgterm set"
for H :: "msgterm set"
where
Inj [intro]: "X \<in> H \<Longrightarrow> X \<in> parts H"
| Fst: "\<langle>X,_\<rangle> \<in> parts H \<Longrightarrow> X \<in> parts H"
| Snd: "\<langle>_,Y\<rangle> \<in> parts H \<Longrightarrow> Y \<in> parts H"
| Lst: "\<lbrakk> L xs \<in> parts H; X \<in> set xs \<rbrakk> \<Longrightarrow> X \<in> parts H"
| FSt: "\<lbrakk> FS xs \<in> parts H; X |\<in>| xs \<rbrakk> \<Longrightarrow> X \<in> parts H"
(*| Hd: "L (X # xs) \<in> parts H \<Longrightarrow> X \<in> parts H"
| Tl: "L (X # xs) \<in> parts H \<Longrightarrow> L xs \<in> parts H" *)
| Body: "Crypt K X \<in> parts H \<Longrightarrow> X \<in> parts H"
text \<open>Monotonicity\<close>
lemma parts_mono: "G \<subseteq> H \<Longrightarrow> parts G \<subseteq> parts H"
apply auto
apply (erule parts.induct)
apply (blast dest: parts.Fst parts.Snd parts.Lst parts.FSt parts.Body)+
done
text \<open>Equations hold because constructors are injective.\<close>
lemma Other_image_eq [simp]: "(AS x \<in> AS`A) = (x:A)"
by auto
lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
by auto
lemma AS_Key_image_eq [simp]: "(AS x \<notin> Key`A)"
by auto
lemma Num_Key_image_eq [simp]: "(Num x \<notin> Key`A)"
by auto
subsection \<open>keysFor operator\<close>
lemma keysFor_empty [simp]: "keysFor {} = {}"
by (unfold keysFor_def, blast)
lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
by (unfold keysFor_def, blast)
lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
by (unfold keysFor_def, blast)
text \<open>Monotonicity\<close>
lemma keysFor_mono: "G \<subseteq> H \<Longrightarrow> keysFor G \<subseteq> keysFor H"
by (unfold keysFor_def, blast)
lemma keysFor_insert_AS [simp]: "keysFor (insert (AS A) H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_Num [simp]: "keysFor (insert (Num N) H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce n) H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_L [simp]: "keysFor (insert (L X) H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_FS [simp]: "keysFor (insert (FS X) H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_MPair [simp]: "keysFor (insert \<langle>X,Y\<rangle> H) = keysFor H"
by (unfold keysFor_def, auto)
lemma keysFor_insert_Crypt [simp]:
"keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
by (unfold keysFor_def, auto)
lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
by (unfold keysFor_def, auto)
lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H \<Longrightarrow> invKey K \<in> keysFor H"
by (unfold keysFor_def, blast)
subsection \<open>Inductive relation "parts"\<close>
lemma MPair_parts:
"\<lbrakk>
\<langle>X,Y\<rangle> \<in> parts H;
\<lbrakk> X \<in> parts H; Y \<in> parts H \<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
by (blast dest: parts.Fst parts.Snd)
lemma L_parts:
"\<lbrakk>
L l \<in> parts H;
\<lbrakk> set l \<subseteq> parts H \<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
by (blast dest: parts.Lst)
lemma FS_parts:
"\<lbrakk>
FS l \<in> parts H;
\<lbrakk> fset l \<subseteq> parts H \<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
- by (simp add: fmember_iff_member_fset parts.FSt subsetI)
-thm fmember_iff_member_fset parts.FSt subsetI
+ by (simp add: parts.FSt subsetI)
+thm parts.FSt subsetI
-declare fmember_iff_member_fset[simp]
declare MPair_parts [elim!] L_parts [elim!] FS_parts [elim] parts.Body [dest!]
text \<open>NB These two rules are UNSAFE in the formal sense, as they discard the
compound message. They work well on THIS FILE.
@{text MPair_parts} is left as SAFE because it speeds up proofs.
The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\<close>
lemma parts_increasing: "H \<subseteq> parts H"
by blast
lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD]
lemma parts_empty [simp]: "parts{} = {}"
apply safe
apply (erule parts.induct, blast+)
done
lemma parts_emptyE [elim!]: "X\<in> parts{} \<Longrightarrow> P"
by simp
text \<open>WARNING: loops if H = {Y}, therefore must not be repeated!\<close>
lemma parts_singleton: "X \<in> parts H \<Longrightarrow> \<exists>Y \<in> H. X \<in> parts {Y}"
apply (erule parts.induct, fast)
using parts.FSt by blast+
lemma parts_singleton_set: "x \<in> parts {s . P s} \<Longrightarrow> \<exists>Y. P Y \<and> x \<in> parts {Y}"
by(auto dest: parts_singleton)
lemma parts_singleton_set_rev: "\<lbrakk>x \<in> parts {Y}; P Y\<rbrakk> \<Longrightarrow> x \<in> parts {s . P s}"
by (induction rule: parts.induct)
(blast dest: parts.Fst parts.Snd parts.Lst parts.FSt parts.Body)+
lemma parts_Hash: "\<lbrakk>\<And>t . t \<in> H \<Longrightarrow> \<exists>t' . t = Hash t'\<rbrakk> \<Longrightarrow> parts H = H"
by (auto, erule parts.induct, fast+)
subsubsection \<open>Unions\<close>
lemma parts_Un_subset1: "parts G \<union> parts H \<subseteq> parts(G \<union> H)"
by (intro Un_least parts_mono Un_upper1 Un_upper2)
lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts G \<union> parts H"
-apply (rule subsetI)
- apply (erule parts.induct, blast+)
- using parts.FSt by auto
+ by (rule subsetI) (erule parts.induct, blast+)
lemma parts_Un [simp]: "parts(G \<union> H) = parts G \<union> parts H"
by (intro equalityI parts_Un_subset1 parts_Un_subset2)
lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
apply (subst insert_is_Un [of _ H])
apply (simp only: parts_Un)
done
text \<open>TWO inserts to avoid looping. This rewrite is better than nothing.
Not suitable for Addsimps: its behaviour can be strange.\<close>
lemma parts_insert2:
"parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
apply (simp add: Un_assoc)
apply (simp add: parts_insert [symmetric])
done
(*needed?!*)
lemma parts_two: "\<lbrakk>x \<in> parts {e1, e2}; x \<notin> parts {e1}\<rbrakk>\<Longrightarrow> x \<in> parts {e2}"
by (simp add: parts_insert2)
text \<open>Added to simplify arguments to parts, analz and synth.\<close>
text \<open>This allows @{text blast} to simplify occurrences of
@{term "parts(G\<union>H)"} in the assumption.\<close>
lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE]
declare in_parts_UnE [elim!]
lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
by (blast intro: parts_mono [THEN [2] rev_subsetD])
subsubsection \<open>Idempotence\<close>
lemma parts_partsD [dest!]: "X\<in> parts (parts H) \<Longrightarrow> X\<in> parts H"
- apply (erule parts.induct, blast+)
- using parts.FSt by auto
+ by (erule parts.induct, blast+)
lemma parts_idem [simp]: "parts (parts H) = parts H"
by blast
lemma parts_subset_iff [simp]: "(parts G \<subseteq> parts H) = (G \<subseteq> parts H)"
apply (rule iffI)
apply (iprover intro: subset_trans parts_increasing)
apply (frule parts_mono, simp)
done
subsubsection \<open>Transitivity\<close>
lemma parts_trans: "\<lbrakk> X\<in> parts G; G \<subseteq> parts H \<rbrakk> \<Longrightarrow> X\<in> parts H"
by (drule parts_mono, blast)
subsubsection \<open>Unions, revisited\<close>
text \<open>You can take the union of parts h for all h in H\<close>
lemma parts_split: "parts H = \<Union> { parts {h} | h . h \<in> H}"
apply auto
apply (erule parts.induct)
apply (blast dest: parts.Fst parts.Snd parts.Lst parts.FSt parts.Body)+
using parts_trans apply blast
done
text \<open>Cut\<close>
lemma parts_cut:
"\<lbrakk> Y\<in> parts (insert X G); X \<in> parts H \<rbrakk> \<Longrightarrow> Y \<in> parts (G \<union> H)"
by (blast intro: parts_trans)
lemma parts_cut_eq [simp]: "X \<in> parts H \<Longrightarrow> parts (insert X H) = parts H"
by (force dest!: parts_cut intro: parts_insertI)
subsubsection \<open>Rewrite rules for pulling out atomic messages\<close>
lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
lemma parts_insert_AS [simp]:
"parts (insert (AS agt) H) = insert (AS agt) (parts H)"
apply (rule parts_insert_eq_I)
by (erule parts.induct, auto elim!: FS_parts)
lemma parts_insert_Epsilon [simp]:
"parts (insert \<epsilon> H) = insert \<epsilon> (parts H)"
apply (rule parts_insert_eq_I)
by (erule parts.induct, auto)
lemma parts_insert_Num [simp]:
"parts (insert (Num N) H) = insert (Num N) (parts H)"
apply (rule parts_insert_eq_I)
by (erule parts.induct, auto)
lemma parts_insert_Key [simp]:
"parts (insert (Key K) H) = insert (Key K) (parts H)"
apply (rule parts_insert_eq_I)
by (erule parts.induct, auto)
lemma parts_insert_Nonce [simp]:
"parts (insert (Nonce n) H) = insert (Nonce n) (parts H)"
apply (rule parts_insert_eq_I)
by (erule parts.induct, auto)
lemma parts_insert_Hash [simp]:
"parts (insert (Hash X) H) = insert (Hash X) (parts H)"
apply (rule parts_insert_eq_I)
by (erule parts.induct, auto)
lemma parts_insert_Crypt [simp]:
"parts (insert (Crypt K X) H) = insert (Crypt K X) (parts (insert X H))"
apply (rule equalityI)
apply (rule subsetI)
apply (erule parts.induct, auto)
by (blast intro: parts.Body)
lemma parts_insert_MPair [simp]:
"parts (insert \<langle>X,Y\<rangle> H) =
insert \<langle>X,Y\<rangle> (parts (insert X (insert Y H)))"
apply (rule equalityI)
apply (rule subsetI)
apply (erule parts.induct, auto)
by (blast intro: parts.Fst parts.Snd)+
lemma parts_insert_L [simp]:
"parts (insert (L xs) H) =
insert (L xs) (parts ((set xs) \<union> H))"
apply (rule equalityI)
apply (rule subsetI)
apply (erule parts.induct, auto)
by (blast intro: parts.Lst)+
lemma parts_insert_FS [simp]:
"parts (insert (FS xs) H) =
insert (FS xs) (parts ((fset xs) \<union> H))"
apply (rule equalityI)
apply (rule subsetI)
apply (erule parts.induct, auto)
by (auto intro: parts.FSt)+
lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
apply auto
apply (erule parts.induct, auto)
done
text \<open>Parts of lists and finite sets.\<close>
lemma parts_list_set (*[simp]*):
"parts (L`ls) = (L`ls) \<union> (\<Union>l \<in> ls. parts (set l)) "
apply (rule equalityI, rule subsetI)
apply (erule parts.induct, auto)
by (meson L_parts image_subset_iff parts_increasing parts_trans)
lemma parts_insert_list_set (*[simp]*):
"parts ((L`ls) \<union> H) = (L`ls) \<union> (\<Union>l \<in> ls. parts ((set l))) \<union> parts H"
apply (rule equalityI, rule subsetI)
by (erule parts.induct, auto simp add: parts_list_set)
(*needed?!*)
lemma parts_fset_set (*[simp]*):
"parts (FS`ls) = (FS`ls) \<union> (\<Union>l \<in> ls. parts (fset l)) "
apply (rule equalityI, rule subsetI)
apply (erule parts.induct, auto)
by (meson FS_parts image_subset_iff parts_increasing parts_trans)
subsubsection \<open>suffix of parts\<close>
lemma suffix_in_parts:
"suffix (x#xs) ys \<Longrightarrow> x \<in> parts {L ys}"
by (auto simp add: suffix_def)
lemma parts_L_set:
"\<lbrakk>x \<in> parts {L ys}; ys \<in> St\<rbrakk> \<Longrightarrow> x \<in> parts (L`St)"
by (metis (no_types, lifting) image_insert insert_iff mk_disjoint_insert parts.Inj
parts_cut_eq parts_insert parts_insert2)
lemma suffix_in_parts_set:
"\<lbrakk>suffix (x#xs) ys; ys \<in> St\<rbrakk> \<Longrightarrow> x \<in> parts (L`St)"
using parts_L_set suffix_in_parts
by blast
subsection \<open>Inductive relation "analz"\<close>
text \<open>Inductive definition of "analz" -- what can be broken down from a set of
messages, including keys. A form of downward closure. Pairs can
be taken apart; messages decrypted with known keys.\<close>
inductive_set
analz :: "msgterm set \<Rightarrow> msgterm set"
for H :: "msgterm set"
where
Inj [intro,simp] : "X \<in> H \<Longrightarrow> X \<in> analz H"
| Fst: "\<langle>X,Y\<rangle> \<in> analz H \<Longrightarrow> X \<in> analz H"
| Snd: "\<langle>X,Y\<rangle> \<in> analz H \<Longrightarrow> Y \<in> analz H"
| Lst: "(L y) \<in> analz H \<Longrightarrow> x \<in> set (y) \<Longrightarrow> x \<in> analz H "
| FSt: "\<lbrakk> FS xs \<in> analz H; X |\<in>| xs \<rbrakk> \<Longrightarrow> X \<in> analz H"
| Decrypt [dest]: "\<lbrakk> Crypt K X \<in> analz H; Key (invKey K) \<in> analz H \<rbrakk> \<Longrightarrow> X \<in> analz H"
text \<open>Monotonicity; Lemma 1 of Lowe's paper\<close>
lemma analz_mono: "G \<subseteq> H \<Longrightarrow> analz(G) \<subseteq> analz(H)"
apply auto
apply (erule analz.induct)
apply (auto dest: analz.Fst analz.Snd analz.Lst analz.FSt )
done
lemmas analz_monotonic = analz_mono [THEN [2] rev_subsetD]
text \<open>Making it safe speeds up proofs\<close>
lemma MPair_analz [elim!]:
"\<lbrakk>
\<langle>X,Y\<rangle> \<in> analz H;
\<lbrakk> X \<in> analz H; Y \<in> analz H \<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
by (blast dest: analz.Fst analz.Snd)
lemma L_analz [elim!]:
"\<lbrakk>
L l \<in> analz H;
\<lbrakk> set l \<subseteq> analz H \<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
by (blast dest: analz.Lst analz.FSt)
lemma FS_analz [elim!]:
"\<lbrakk>
FS l \<in> analz H;
\<lbrakk> fset l \<subseteq> analz H \<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
- by (simp add: fmember_iff_member_fset analz.FSt subsetI)
+ by (simp add: analz.FSt subsetI)
-thm fmember_iff_member_fset parts.FSt subsetI
+thm parts.FSt subsetI
lemma analz_increasing: "H \<subseteq> analz(H)"
by blast
lemma analz_subset_parts: "analz H \<subseteq> parts H"
-apply (rule subsetI)
- apply (erule analz.induct, blast+)
- apply auto
-done
+ by (rule subsetI) (erule analz.induct, blast+)
text \<open>If there is no cryptography, then analz and parts is equivalent.\<close>
lemma no_crypt_analz_is_parts:
"\<not> (\<exists> K X . Crypt K X \<in> parts A) \<Longrightarrow> analz A = parts A"
apply (rule equalityI, simp add: analz_subset_parts)
apply (rule subsetI)
by (erule parts.induct, blast+, auto)
lemmas analz_into_parts = analz_subset_parts [THEN subsetD]
lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD]
lemma parts_analz [simp]: "parts (analz H) = parts H"
apply (rule equalityI)
apply (rule analz_subset_parts [THEN parts_mono, THEN subset_trans], simp)
apply (blast intro: analz_increasing [THEN parts_mono, THEN subsetD])
done
lemma analz_parts [simp]: "analz (parts H) = parts H"
apply auto
apply (erule analz.induct, auto)
done
lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD]
subsubsection \<open>General equational properties\<close>
lemma analz_empty [simp]: "analz {} = {}"
apply safe
apply (erule analz.induct, blast+)
done
text \<open>Converse fails: we can analz more from the union than from the
separate parts, as a key in one might decrypt a message in the other\<close>
lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
by (intro Un_least analz_mono Un_upper1 Un_upper2)
lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
subsubsection \<open>Rewrite rules for pulling out atomic messages\<close>
lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
lemma analz_insert_AS [simp]:
"analz (insert (AS agt) H) = insert (AS agt) (analz H)"
apply (rule analz_insert_eq_I)
by (erule analz.induct, auto)
lemma analz_insert_Num [simp]:
"analz (insert (Num N) H) = insert (Num N) (analz H)"
apply (rule analz_insert_eq_I)
by (erule analz.induct, auto)
text \<open>Can only pull out Keys if they are not needed to decrypt the rest\<close>
lemma analz_insert_Key [simp]:
"K \<notin> keysFor (analz H) \<Longrightarrow>
analz (insert (Key K) H) = insert (Key K) (analz H)"
apply (unfold keysFor_def)
apply (rule analz_insert_eq_I)
by (erule analz.induct, auto)
lemma analz_insert_LEmpty [simp]:
"analz (insert (L []) H) = insert (L []) (analz H)"
apply (rule analz_insert_eq_I)
by (erule analz.induct, auto)
lemma analz_insert_L [simp]:
"analz (insert (L l) H) = insert (L l) (analz (set l \<union> H))"
apply (rule equalityI)
apply (rule subsetI)
apply (erule analz.induct, auto)
apply (erule analz.induct, auto)
using analz.Inj by blast
lemma analz_insert_FS [simp]:
"analz (insert (FS l) H) = insert (FS l) (analz (fset l \<union> H))"
apply (rule equalityI)
apply (rule subsetI)
apply (erule analz.induct, auto)
apply (erule analz.induct, auto)
using analz.Inj by blast
lemma "L[] \<in> analz {L[L[]]}"
using analz.Inj by simp
lemma analz_insert_Hash [simp]:
"analz (insert (Hash X) H) = insert (Hash X) (analz H)"
apply (rule analz_insert_eq_I)
by (erule analz.induct, auto)
lemma analz_insert_MPair [simp]:
"analz (insert \<langle>X,Y\<rangle> H) =
insert \<langle>X,Y\<rangle> (analz (insert X (insert Y H)))"
apply (rule equalityI)
apply (rule subsetI)
apply (erule analz.induct, auto)
apply (erule analz.induct, auto)
using Fst Snd analz.Inj insertI1
by (metis)+
text \<open>Can pull out enCrypted message if the Key is not known\<close>
lemma analz_insert_Crypt:
"Key (invKey K) \<notin> analz H
\<Longrightarrow> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
apply (rule analz_insert_eq_I)
by (erule analz.induct, auto)
lemma analz_insert_Decrypt1:
"Key (invKey K) \<in> analz H \<Longrightarrow>
analz (insert (Crypt K X) H) \<subseteq>
insert (Crypt K X) (analz (insert X H))"
apply (rule subsetI)
by (erule_tac x = x in analz.induct, auto)
lemma analz_insert_Decrypt2:
"Key (invKey K) \<in> analz H \<Longrightarrow>
insert (Crypt K X) (analz (insert X H)) \<subseteq>
analz (insert (Crypt K X) H)"
apply auto
apply (erule_tac x = x in analz.induct, auto)
by (blast intro: analz_insertI analz.Decrypt)
lemma analz_insert_Decrypt:
"Key (invKey K) \<in> analz H \<Longrightarrow>
analz (insert (Crypt K X) H) =
insert (Crypt K X) (analz (insert X H))"
by (intro equalityI analz_insert_Decrypt1 analz_insert_Decrypt2)
text \<open>Case analysis: either the message is secure, or it is not! Effective,
but can cause subgoals to blow up! Use with @{text "split_if"}; apparently
@{text "split_tac"} does not cope with patterns such as @{term"analz (insert
(Crypt K X) H)"}\<close>
lemma analz_Crypt_if [simp]:
"analz (insert (Crypt K X) H) =
(if (Key (invKey K) \<in> analz H)
then insert (Crypt K X) (analz (insert X H))
else insert (Crypt K X) (analz H))"
by (simp add: analz_insert_Crypt analz_insert_Decrypt)
text \<open>This rule supposes "for the sake of argument" that we have the key.\<close>
lemma analz_insert_Crypt_subset:
"analz (insert (Crypt K X) H) \<subseteq>
insert (Crypt K X) (analz (insert X H))"
apply (rule subsetI)
by (erule analz.induct, auto)
lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
apply auto
apply (erule analz.induct, auto)
done
subsubsection \<open>Idempotence and transitivity\<close>
lemma analz_analzD [dest!]: "X\<in> analz (analz H) \<Longrightarrow> X\<in> analz H"
by (erule analz.induct, auto)
lemma analz_idem [simp]: "analz (analz H) = analz H"
by blast
lemma analz_subset_iff [simp]: "(analz G \<subseteq> analz H) = (G \<subseteq> analz H)"
apply (rule iffI)
apply (iprover intro: subset_trans analz_increasing)
apply (frule analz_mono, simp)
done
lemma analz_trans: "\<lbrakk> X\<in> analz G; G \<subseteq> analz H \<rbrakk> \<Longrightarrow> X\<in> analz H"
by (drule analz_mono, blast)
text \<open>Cut; Lemma 2 of Lowe\<close>
lemma analz_cut: "\<lbrakk> Y\<in> analz (insert X H); X\<in> analz H \<rbrakk> \<Longrightarrow> Y\<in> analz H"
by (erule analz_trans, blast)
(*Cut can be proved easily by induction on
"Y: analz (insert X H) \<Longrightarrow> X: analz H --> Y: analz H"
*)
text \<open>This rewrite rule helps in the simplification of messages that involve
the forwarding of unknown components (X). Without it, removing occurrences
of X can be very complicated.\<close>
lemma analz_insert_eq: "X\<in> analz H \<Longrightarrow> analz (insert X H) = analz H"
by (blast intro: analz_cut analz_insertI)
text \<open>A congruence rule for "analz"\<close>
lemma analz_subset_cong:
"\<lbrakk> analz G \<subseteq> analz G'; analz H \<subseteq> analz H' \<rbrakk>
\<Longrightarrow> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
apply simp
apply (iprover intro: conjI subset_trans analz_mono Un_upper1 Un_upper2)
done
lemma analz_cong:
"\<lbrakk> analz G = analz G'; analz H = analz H' \<rbrakk>
\<Longrightarrow> analz (G \<union> H) = analz (G' \<union> H')"
by (intro equalityI analz_subset_cong, simp_all)
lemma analz_insert_cong:
"analz H = analz H' \<Longrightarrow> analz(insert X H) = analz(insert X H')"
by (force simp only: insert_def intro!: analz_cong)
text \<open>If there are no pairs, lists or encryptions then analz does nothing\<close>
(*needed?*)
lemma analz_trivial:
"\<lbrakk>
\<forall>X Y. \<langle>X,Y\<rangle> \<notin> H; \<forall>xs. L xs \<notin> H; \<forall>xs. FS xs \<notin> H;
\<forall>X K. Crypt K X \<notin> H
\<rbrakk> \<Longrightarrow> analz H = H"
apply safe
by (erule analz.induct, auto)
text \<open>These two are obsolete (with a single Spy) but cost little to prove...\<close>
lemma analz_UN_analz_lemma:
"X\<in> analz (\<Union>i\<in>A. analz (H i)) \<Longrightarrow> X\<in> analz (\<Union>i\<in>A. H i)"
apply (erule analz.induct)
by (auto intro: analz_mono [THEN [2] rev_subsetD])
lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
subsubsection \<open>Lemmas assuming absense of keys\<close>
text \<open>If there are no keys in analz H, you can take the union of analz h for all h in H\<close>
lemma analz_split:
"\<not>(\<exists> K . Key K \<in> analz H)
\<Longrightarrow> analz H = \<Union> { analz {h} | h . h \<in> H}"
apply auto
subgoal
apply (erule analz.induct)
apply (auto dest: analz.Fst analz.Snd analz.Lst analz.FSt)
done
apply (erule analz.induct)
apply (auto dest: analz.Fst analz.Snd analz.Lst analz.FSt)
done
lemma analz_Un_eq:
assumes "\<not>(\<exists> K . Key K \<in> analz H)" and "\<not>(\<exists> K . Key K \<in> analz G)"
shows "analz (H \<union> G) = analz H \<union> analz G"
apply (intro equalityI, rule subsetI)
apply (erule analz.induct)
using assms by auto
lemma analz_Un_eq_Crypt:
assumes "\<not>(\<exists> K . Key K \<in> analz G)" and "\<not>(\<exists> K X . Crypt K X \<in> analz G)"
shows "analz (H \<union> G) = analz H \<union> analz G"
apply (intro equalityI, rule subsetI)
apply (erule analz.induct)
using assms by auto
lemma analz_list_set (*[simp]*):
"\<not>(\<exists> K . Key K \<in> analz (L`ls))
\<Longrightarrow> analz (L`ls) = (L`ls) \<union> (\<Union>l \<in> ls. analz (set l)) "
apply (rule equalityI, rule subsetI)
apply (erule analz.induct, auto)
using L_analz image_subset_iff analz_increasing analz_trans by metis
lemma analz_fset_set (*[simp]*):
"\<not>(\<exists> K . Key K \<in> analz (FS`ls))
\<Longrightarrow> analz (FS`ls) = (FS`ls) \<union> (\<Union>l \<in> ls. analz (fset l)) "
apply (rule equalityI, rule subsetI)
apply (erule analz.induct, auto)
using FS_analz image_subset_iff analz_increasing analz_trans by metis
subsection \<open>Inductive relation "synth"\<close>
text \<open>Inductive definition of "synth" -- what can be built up from a set of
messages. A form of upward closure. Pairs can be built, messages
encrypted with known keys. AS names are public domain.
Nums can be guessed, but Nonces cannot be.\<close>
inductive_set
synth :: "msgterm set \<Rightarrow> msgterm set"
for H :: "msgterm set"
where
Inj [intro]: "X \<in> H \<Longrightarrow> X \<in> synth H"
| \<epsilon> [simp,intro!]: "\<epsilon> \<in> synth H"
| AS [simp,intro!]: "AS agt \<in> synth H"
| Num [simp,intro!]: "Num n \<in> synth H"
| Lst [intro]: "\<lbrakk> \<And>x . x \<in> set xs \<Longrightarrow> x \<in> synth H \<rbrakk> \<Longrightarrow> L xs \<in> synth H"
| FSt [intro]: "\<lbrakk> \<And>x . x \<in> fset xs \<Longrightarrow> x \<in> synth H;
\<And>x ys . x \<in> fset xs \<Longrightarrow> x \<noteq> FS ys \<rbrakk>
\<Longrightarrow> FS xs \<in> synth H"
| Hash [intro]: "X \<in> synth H \<Longrightarrow> Hash X \<in> synth H"
| MPair [intro]: "\<lbrakk> X \<in> synth H; Y \<in> synth H \<rbrakk> \<Longrightarrow> \<langle>X,Y\<rangle> \<in> synth H"
| Crypt [intro]: "\<lbrakk> X \<in> synth H; Key K \<in> H \<rbrakk> \<Longrightarrow> Crypt K X \<in> synth H"
(*removed fcard xs \<noteq> Suc 0 from FSt premise*)
text \<open>Monotonicity\<close>
lemma synth_mono: "G \<subseteq> H \<Longrightarrow> synth(G) \<subseteq> synth(H)"
apply (auto, erule synth.induct, auto)
by blast
text \<open>NO @{text AS_synth}, as any AS name can be synthesized.
The same holds for @{term Num}\<close>
inductive_cases Key_synth [elim!]: "Key K \<in> synth H"
inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
inductive_cases Hash_synth [elim!]: "Hash X \<in> synth H"
inductive_cases MPair_synth [elim!]: "\<langle>X,Y\<rangle> \<in> synth H"
inductive_cases L_synth [elim!]: "L X \<in> synth H"
inductive_cases FS_synth [elim!]: "FS X \<in> synth H"
inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
lemma synth_increasing: "H \<subseteq> synth(H)"
by blast
lemma synth_analz_self: "x \<in> H \<Longrightarrow> x \<in> synth (analz H)"
by blast
subsubsection \<open>Unions\<close>
text \<open>Converse fails: we can synth more from the union than from the
separate parts, building a compound message using elements of each.\<close>
lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
by (intro Un_least synth_mono Un_upper1 Un_upper2)
lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
by (blast intro: synth_mono [THEN [2] rev_subsetD])
subsubsection \<open>Idempotence and transitivity\<close>
lemma synth_synthD [dest!]: "X\<in> synth (synth H) \<Longrightarrow> X \<in> synth H"
apply (erule synth.induct, blast)
apply auto by blast
lemma synth_idem: "synth (synth H) = synth H"
by blast
lemma synth_subset_iff [simp]: "(synth G \<subseteq> synth H) = (G \<subseteq> synth H)"
apply (rule iffI)
apply (iprover intro: subset_trans synth_increasing)
apply (frule synth_mono, simp add: synth_idem)
done
lemma synth_trans: "\<lbrakk> X\<in> synth G; G \<subseteq> synth H \<rbrakk> \<Longrightarrow> X\<in> synth H"
by (drule synth_mono, blast)
text \<open>Cut; Lemma 2 of Lowe\<close>
lemma synth_cut: "\<lbrakk> Y\<in> synth (insert X H); X\<in> synth H \<rbrakk> \<Longrightarrow> Y\<in> synth H"
by (erule synth_trans, blast)
lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
by blast
lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
by blast
lemma Crypt_synth_eq [simp]:
"Key K \<notin> H \<Longrightarrow> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
by blast
lemma keysFor_synth [simp]:
"keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
by (unfold keysFor_def, blast)
lemma L_cons_synth [simp]:
"(set xs \<subseteq> H) \<Longrightarrow> (L xs \<in> synth H)"
by auto
lemma FS_cons_synth [simp]:
"\<lbrakk>fset xs \<subseteq> H; \<And>x ys. x \<in> fset xs \<Longrightarrow> x \<noteq> FS ys; fcard xs \<noteq> Suc 0 \<rbrakk> \<Longrightarrow> (FS xs \<in> synth H)"
by auto
subsubsection \<open>Combinations of parts, analz and synth\<close>
lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
proof (safe del: UnCI)
fix X
assume "X \<in> parts (synth H)"
thus "X \<in> parts H \<union> synth H"
by (induct rule: parts.induct)
(auto intro: parts.Fst parts.Snd parts.Lst parts.FSt parts.Body)
next
fix X
assume "X \<in> parts H"
thus "X \<in> parts (synth H)"
by (induction rule: parts.induct)
(auto intro: parts.Fst parts.Snd parts.Lst parts.FSt parts.Body)
next
fix X
assume "X \<in> synth H"
thus "X \<in> parts (synth H)"
apply (induction rule: synth.induct)
apply(auto intro: parts.Fst parts.Snd parts.Lst parts.FSt parts.Body)
by blast
qed
lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
apply (intro equalityI analz_subset_cong)+
apply simp_all
done
lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
proof (safe del: UnCI)
fix X
assume "X \<in> analz (synth G \<union> H)"
thus "X \<in> analz (G \<union> H) \<union> synth G"
by (induction rule: analz.induct)
(auto intro: analz.Fst analz.Snd analz.Lst analz.FSt analz.Decrypt)
qed (auto elim: analz_mono [THEN [2] rev_subsetD])
lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
apply (cut_tac H = "{}" in analz_synth_Un)
apply (simp (no_asm_use))
done
lemma analz_Un_analz [simp]: "analz (G \<union> analz H) = analz (G \<union> H)"
by (subst Un_commute, auto)+
lemma analz_synth_Un2 [simp]: "analz (G \<union> synth H) = analz (G \<union> H) \<union> synth H"
by (subst Un_commute, auto)+
subsubsection \<open>For reasoning about the Fake rule in traces\<close>
lemma parts_insert_subset_Un: "X\<in> G \<Longrightarrow> parts(insert X H) \<subseteq> parts G \<union> parts H"
by (rule subset_trans [OF parts_mono parts_Un_subset2], blast)
text \<open>More specifically for Fake. Very occasionally we could do with a version
of the form @{term"parts{X} \<subseteq> synth (analz H) \<union> parts H"}\<close>
lemma Fake_parts_insert:
"X \<in> synth (analz H) \<Longrightarrow>
parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
apply (drule parts_insert_subset_Un)
apply (simp (no_asm_use))
apply blast
done
lemma Fake_parts_insert_in_Un:
"\<lbrakk>Z \<in> parts (insert X H); X \<in> synth (analz H)\<rbrakk>
\<Longrightarrow> Z \<in> synth (analz H) \<union> parts H"
by (blast dest: Fake_parts_insert [THEN subsetD, dest])
text \<open>@{term H} is sometimes @{term"Key ` KK \<union> spies evs"}, so can't put @{term "G=H"}.\<close>
lemma Fake_analz_insert:
"X\<in> synth (analz G) \<Longrightarrow>
analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
apply (rule subsetI)
apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
prefer 2
apply (blast intro: analz_mono [THEN [2] rev_subsetD]
analz_mono [THEN synth_mono, THEN [2] rev_subsetD])
apply (simp (no_asm_use))
apply blast
done
lemma analz_conj_parts [simp]:
"(X \<in> analz H & X \<in> parts H) = (X \<in> analz H)"
by (blast intro: analz_subset_parts [THEN subsetD])
lemma analz_disj_parts [simp]:
"(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
by (blast intro: analz_subset_parts [THEN subsetD])
text \<open>Without this equation, other rules for synth and analz would yield
redundant cases\<close>
lemma MPair_synth_analz [iff]:
"(\<langle>X,Y\<rangle> \<in> synth (analz H)) =
(X \<in> synth (analz H) & Y \<in> synth (analz H))"
by blast
lemma L_cons_synth_analz [iff]:
"(L xs \<in> synth (analz H)) =
(set xs \<subseteq> synth (analz H))"
by blast
lemma L_cons_synth_parts [iff]:
"(L xs \<in> synth (parts H)) =
(set xs \<subseteq> synth (parts H))"
by blast
lemma FS_cons_synth_analz [iff]:
"\<lbrakk>\<And>x ys . x \<in> fset xs \<Longrightarrow> x \<noteq> FS ys; fcard xs \<noteq> Suc 0 \<rbrakk> \<Longrightarrow>
(FS xs \<in> synth (analz H)) =
(fset xs \<subseteq> synth (analz H))"
by blast
lemma FS_cons_synth_parts [iff]:
"\<lbrakk>\<And>x ys . x \<in> fset xs \<Longrightarrow> x \<noteq> FS ys; fcard xs \<noteq> Suc 0 \<rbrakk> \<Longrightarrow>
(FS xs \<in> synth (parts H)) =
(fset xs \<subseteq> synth (parts H))"
by blast
lemma Crypt_synth_analz:
"\<lbrakk> Key K \<in> analz H; Key (invKey K) \<in> analz H \<rbrakk>
\<Longrightarrow> (Crypt K X \<in> synth (analz H)) = (X \<in> synth (analz H))"
by blast
lemma Hash_synth_analz [simp]:
"X \<notin> synth (analz H)
\<Longrightarrow> (Hash\<langle>X,Y\<rangle> \<in> synth (analz H)) = (Hash\<langle>X,Y\<rangle> \<in> analz H)"
by blast
subsection \<open>HPair: a combination of Hash and MPair\<close>
text \<open>We do NOT want Crypt... messages broken up in protocols!!\<close>
declare parts.Body [rule del]
text \<open>Rewrites to push in Key and Crypt messages, so that other messages can
be pulled out using the @{text analz_insert} rules\<close>
(*needed?*)
lemmas pushKeys =
insert_commute [of "Key K" "AS C" for K C]
insert_commute [of "Key K" "Nonce N" for K N]
insert_commute [of "Key K" "Num N" for K N]
insert_commute [of "Key K" "Hash X" for K X]
insert_commute [of "Key K" "MPair X Y" for K X Y]
insert_commute [of "Key K" "Crypt X K'" for K K' X]
(*needed?*)
lemmas pushCrypts =
insert_commute [of "Crypt X K" "AS C" for X K C]
insert_commute [of "Crypt X K" "AS C" for X K C]
insert_commute [of "Crypt X K" "Nonce N" for X K N]
insert_commute [of "Crypt X K" "Num N" for X K N]
insert_commute [of "Crypt X K" "Hash X'" for X K X']
insert_commute [of "Crypt X K" "MPair X' Y" for X K X' Y]
text \<open>Cannot be added with @{text "[simp]"} -- messages should not always be
re-ordered.\<close>
lemmas pushes = pushKeys pushCrypts
text \<open>By default only @{text o_apply} is built-in. But in the presence of
eta-expansion this means that some terms displayed as @{term "f o g"} will be
rewritten, and others will not!\<close>
declare o_def [simp]
lemma Crypt_notin_image_Key [simp]: "Crypt K X \<notin> Key ` A"
by auto
lemma Hash_notin_image_Key [simp] :"Hash X \<notin> Key ` A"
by auto
lemma synth_analz_mono: "G\<subseteq>H \<Longrightarrow> synth (analz(G)) \<subseteq> synth (analz(H))"
by (iprover intro: synth_mono analz_mono)
lemma synth_parts_mono: "G\<subseteq>H \<Longrightarrow> synth (parts G) \<subseteq> synth (parts H)"
by (iprover intro: synth_mono parts_mono)
lemma Fake_analz_eq [simp]:
"X \<in> synth(analz H) \<Longrightarrow> synth (analz (insert X H)) = synth (analz H)"
apply (drule Fake_analz_insert[of _ _ "H"])
apply (simp add: synth_increasing[THEN Un_absorb2])
apply (drule synth_mono)
apply (simp add: synth_idem)
apply (rule equalityI)
apply (simp add: )
apply (rule synth_analz_mono, blast)
done
text \<open>Two generalizations of @{text analz_insert_eq}\<close>
lemma gen_analz_insert_eq [rule_format]:
"X \<in> analz H \<Longrightarrow> ALL G. H \<subseteq> G --> analz (insert X G) = analz G"
by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
lemma Fake_parts_sing:
"X \<in> synth (analz H) \<Longrightarrow> parts{X} \<subseteq> synth (analz H) \<union> parts H"
apply (rule subset_trans)
apply (erule_tac [2] Fake_parts_insert)
apply (rule parts_mono, blast)
done
lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
text \<open>For some reason, moving this up can make some proofs loop!\<close>
declare invKey_K [simp]
lemma synth_analz_insert:
assumes "analz H \<subseteq> synth (analz H')"
shows "analz (insert X H) \<subseteq> synth (analz (insert X H'))"
proof
fix x
assume "x \<in> analz (insert X H)"
then have "x \<in> analz (insert X (synth (analz H')))"
using assms by (meson analz_increasing analz_monotonic insert_mono)
then show "x \<in> synth (analz (insert X H'))"
by (metis (no_types) Un_iff analz_idem analz_insert analz_monotonic analz_synth synth.Inj
synth_insert synth_mono)
qed
lemma synth_parts_insert:
assumes "parts H \<subseteq> synth (parts H')"
shows "parts (insert X H) \<subseteq> synth (parts (insert X H'))"
proof
fix x
assume "x \<in> parts (insert X H)"
then have "x \<in> parts (insert X (synth (parts H')))"
using assms parts_increasing
by (metis UnE UnI1 analz_monotonic analz_parts parts_insert parts_insertI)
then show "x \<in> synth (parts (insert X H'))"
using Un_iff parts_idem parts_insert parts_synth synth.Inj
by (metis Un_subset_iff synth_increasing synth_trans)
qed
lemma parts_insert_subset_impl:
"\<lbrakk>x \<in> parts (insert a G); x \<in> parts G \<Longrightarrow> x \<in> synth (parts H); a \<in> synth (parts H)\<rbrakk>
\<Longrightarrow> x \<in> synth (parts H)"
using Fake_parts_sing in_parts_UnE insert_is_Un
parts_idem parts_synth subsetCE sup.absorb2 synth_idem synth_increasing
by (metis (no_types, lifting) analz_parts)
lemma synth_parts_subset_elem:
"\<lbrakk>A \<subseteq> synth (parts B); x \<in> parts A\<rbrakk> \<Longrightarrow> x \<in> synth (parts B)"
by (meson parts_emptyE parts_insert_subset_impl parts_singleton subset_iff)
lemma synth_parts_subset:
"A \<subseteq> synth (parts B) \<Longrightarrow> parts A \<subseteq> synth (parts B)"
by (auto simp add: synth_parts_subset_elem)
lemma parts_synth_parts[simp]: "parts (synth (parts H)) = synth (parts H)"
by auto
lemma synth_parts_trans:
assumes "A \<subseteq> synth (parts B)" and "B \<subseteq> synth (parts C)"
shows "A \<subseteq> synth (parts C)"
using assms by (metis order_trans parts_synth_parts synth_idem synth_parts_mono)
lemma synth_parts_trans_elem:
assumes "x \<in> A" and "A \<subseteq> synth (parts B)" and "B \<subseteq> synth (parts C)"
shows "x \<in> synth (parts C)"
using synth_parts_trans assms by auto
lemma synth_un_parts_split:
assumes "x \<in> synth (parts A \<union> parts B)"
and "\<And>x . x \<in> A \<Longrightarrow> x \<in> synth (parts C)"
and "\<And>x . x \<in> B \<Longrightarrow> x \<in> synth (parts C)"
shows "x \<in> synth (parts C)"
proof -
have "parts A \<subseteq> synth (parts C)" "parts B \<subseteq> synth (parts C)"
using assms(2) assms(3) synth_parts_subset by blast+
then have "x \<in> synth (synth (parts C) \<union> synth (parts C))" using assms(1)
using synth_trans by auto
then show ?thesis by auto
qed
subsubsection \<open>Normalization of Messages\<close>
text\<open>Prevent FS from being contained directly in other FS.
For instance, a term @{term "FS {| FS {| Num 0 |}, Num 0 |}"} is
not normalized, whereas @{term "FS {| Hash (FS {| Num 0 |}), Num 0 |}"} is normalized.\<close>
inductive normalized :: "msgterm \<Rightarrow> bool" where
\<epsilon> [simp,intro!]: "normalized \<epsilon>"
| AS [simp,intro!]: "normalized (AS agt)"
| Num [simp,intro!]: "normalized (Num n)"
| Key [simp,intro!]: "normalized (Key n)"
| Nonce [simp,intro!]: "normalized (Nonce n)"
| Lst [intro]: "\<lbrakk> \<And>x . x \<in> set xs \<Longrightarrow> normalized x \<rbrakk> \<Longrightarrow> normalized (L xs)"
| FSt [intro]: "\<lbrakk> \<And>x . x \<in> fset xs \<Longrightarrow> normalized x;
\<And>x ys . x \<in> fset xs \<Longrightarrow> x \<noteq> FS ys \<rbrakk>
\<Longrightarrow> normalized (FS xs)"
| Hash [intro]: "normalized X \<Longrightarrow> normalized (Hash X)"
| MPair [intro]: "\<lbrakk> normalized X; normalized Y \<rbrakk> \<Longrightarrow> normalized \<langle>X,Y\<rangle>"
| Crypt [intro]: "\<lbrakk> normalized X \<rbrakk> \<Longrightarrow> normalized (Crypt K X)"
thm normalized.simps
find_theorems normalized
text\<open>Examples\<close>
lemma "normalized (FS {| Hash (FS {| Num 0 |}), Num 0 |})" by fastforce
lemma "\<not> normalized (FS {| FS {| Num 0 |}, Num 0 |})" by (auto elim: normalized.cases)
subsubsection\<open>Closure of @{text "normalized"} under @{text "parts"}, @{text "analz"} and @{text "synth"}\<close>
text\<open>All synthesized terms are normalized (since @{text "synth"} prevents directly nested FSets).\<close>
lemma normalized_synth[elim!]: "\<lbrakk>t \<in> synth H; \<And>t. t \<in> H \<Longrightarrow> normalized t\<rbrakk> \<Longrightarrow> normalized t"
by(induction t, auto 3 4)
lemma normalized_parts[elim!]: "\<lbrakk>t \<in> parts H; \<And>t. t \<in> H \<Longrightarrow> normalized t\<rbrakk> \<Longrightarrow> normalized t"
by(induction t rule: parts.induct)
(auto elim: normalized.cases)
lemma normalized_analz[elim!]: "\<lbrakk>t \<in> analz H; \<And>t. t \<in> H \<Longrightarrow> normalized t\<rbrakk> \<Longrightarrow> normalized t"
by(induction t rule: analz.induct)
(auto elim: normalized.cases)
subsubsection\<open>Properties of @{text "normalized"}\<close>
lemma normalized_FS[elim]: "\<lbrakk>normalized (FS xs); x |\<in>| xs\<rbrakk> \<Longrightarrow> normalized x"
by(auto simp add: normalized.simps[of "FS xs"])
lemma normalized_FS_FS[elim]: "\<lbrakk>normalized (FS xs); x |\<in>| xs; x = FS ys\<rbrakk> \<Longrightarrow> False"
by(auto simp add: normalized.simps[of "FS xs"])
lemma normalized_subset: "\<lbrakk>normalized (FS xs); ys |\<subseteq>| xs\<rbrakk> \<Longrightarrow> normalized (FS ys)"
by (auto intro!: normalized.FSt)
lemma normalized_insert[elim!]: "normalized (FS (finsert x xs)) \<Longrightarrow> normalized (FS xs)"
by(auto elim!: normalized_subset)
lemma normalized_union:
assumes "normalized (FS xs)" "normalized (FS ys)" "zs |\<subseteq>| xs |\<union>| ys"
shows "normalized (FS zs)"
using assms by(auto intro!: normalized.FSt)
lemma normalized_minus[elim]:
assumes "normalized (FS (ys |-| xs))" "normalized (FS xs)"
shows "normalized (FS ys)"
using normalized_union assms by blast
subsubsection\<open>Lemmas that do not use @{text "normalized"}, but are helpful in proving its properties\<close>
lemma FS_mono: "\<lbrakk>zs_s = finsert (f (FS zs_s)) zs_b; \<And> x. size (f x) > size x\<rbrakk> \<Longrightarrow> False"
by (metis (no_types) add.right_neutral add_Suc_right finite_fset finsert.rep_eq less_add_Suc1
msgterm.size(17) not_less_eq size_fset_simps sum.insert_remove)
lemma FS_contr: "\<lbrakk>zs = f (FS {|zs|}); \<And> x. size (f x) > size x\<rbrakk> \<Longrightarrow> False"
using FS_mono by blast
end
diff --git a/thys/IsaNet/instances/Anapaya_SCION.thy b/thys/IsaNet/instances/Anapaya_SCION.thy
--- a/thys/IsaNet/instances/Anapaya_SCION.thy
+++ b/thys/IsaNet/instances/Anapaya_SCION.thy
@@ -1,567 +1,567 @@
(*******************************************************************************
Project: IsaNet
Author: Tobias Klenze, ETH Zurich <tobias.klenze@inf.ethz.ch>
Version: JCSPaper.1.0
Isabelle Version: Isabelle2021-1
Copyright (c) 2022 Tobias Klenze
Licence: Mozilla Public License 2.0 (MPL) / BSD-3-Clause (dual license)
*******************************************************************************)
section \<open>Anapaya-SCION\<close>
text\<open>This is the "new" SCION protocol, as specified on the website of Anapaya:
\url{https://scion.docs.anapaya.net/en/latest/protocols/scion-header.html} (Accessed 2021-03-02).
It does not use the next hop field in its MAC computation, but instead refers uses a mutable uinfo
field which acts as an XOR-based accumulator for all upstream MACs.
This protocol instance requires the use of the extensions of our formalization that provide mutable
uinfo field and an XOR abstraction.\<close>
theory "Anapaya_SCION"
imports
"../Parametrized_Dataplane_3_directed"
"../infrastructure/Abstract_XOR"
begin
locale scion_defs = network_assums_direct _ _ _ auth_seg0
for auth_seg0 :: "(msgterm \<times> ahi list) set"
begin
sublocale comp_fun_commute xor
by(auto simp add: comp_fun_commute_def xor_def)
(******************************************************************************)
subsection \<open>Hop validation check and extract functions\<close>
(******************************************************************************)
type_synonym SCION_HF = "(unit, unit) HF"
text\<open>The predicate @{term "hf_valid"} is given to the concrete parametrized model as a parameter.
It ensures the authenticity of the hop authenticator in the hop field. The predicate takes an authenticated
info field (in this model always a numeric value, hence the matching on Num ts), the unauthenticated
info field and the hop field to be validated. The next hop field is not used in this instance.\<close>
fun hf_valid :: "msgterm \<Rightarrow> msgterm fset
\<Rightarrow> SCION_HF
\<Rightarrow> SCION_HF option \<Rightarrow> bool" where
"hf_valid (Num ts) uinfo \<lparr>AHI = ahi, UHI = _, HVF = x\<rparr> nxt \<longleftrightarrow>
(\<exists>upif downif. x = Mac[macKey (ASID ahi)] (L [Num ts, upif, downif, FS uinfo]) \<and>
ASIF (DownIF ahi) downif \<and> ASIF (UpIF ahi) upif)"
| "hf_valid _ _ _ _ = False"
text\<open>Updating the uinfo field involves XORin the current hop validation field onto it. Note that in all
authorized segments, the hvf will already have been contained in segid, hence this operation only
removes terms from the fset in the forwarding of honestly created packets.\<close>
definition upd_uinfo :: "msgterm fset \<Rightarrow> SCION_HF \<Rightarrow> msgterm fset" where
"upd_uinfo segid hf = xor segid {| HVF hf |}"
declare upd_uinfo_def[simp]
(*What can uinfo be?
hfs = [], uinfo = FS {||}
hfs = [x], uinfo = FS {||}
hfs = [y, x], uinfo = FS {| MAC-of-x |}
hfs = [z, y, x], uinfo = FS {| MAC-of-x, MAC-of-y |}
*)
text\<open>The following lemma is needed to show the termination of extr, defined below.\<close>
lemma extr_helper:
"\<lbrakk>x = Mac[macKey asid'a] (L [ts, upif'a, downif'a, FS segid']);
fcard segid' = fcard (xor segid {|x|}); x |\<in>| segid\<rbrakk>
\<Longrightarrow> (case x of Hash \<langle>Key (macK asid), L []\<rangle> \<Rightarrow> 0 | Hash \<langle>Key (macK asid), L [ts]\<rangle> \<Rightarrow> 0
| Hash \<langle>Key (macK asid), L [ts, upif]\<rangle> \<Rightarrow> 0 | Hash \<langle>Key (macK asid), L [ts, upif, downif]\<rangle> \<Rightarrow> 0
| Hash \<langle>Key (macK asid), L [ts, upif, downif, FS segid]\<rangle> \<Rightarrow> Suc (fcard segid)
| Hash \<langle>Key (macK asid), L (ts # upif # downif # FS segid # ac # lista)\<rangle> \<Rightarrow> 0
| Hash \<langle>Key (macK asid), L (ts # upif # downif # _ # list)\<rangle> \<Rightarrow> 0
| Hash \<langle>Key (macK asid), _\<rangle> \<Rightarrow> 0 | Hash \<langle>Key _, msgterm2\<rangle> \<Rightarrow> 0 | Hash \<langle>_, msgterm2\<rangle> \<Rightarrow> 0
| Hash _ \<Rightarrow> 0 | _ \<Rightarrow> 0)
< Suc (fcard segid)"
apply auto
- using fcard_fminus1_less notin_fset xor_singleton(1) by (metis)
+ using fcard_fminus1_less xor_singleton(1) by (metis)
text\<open>We can extract the entire path from the hvf field, which includes the local forwarding
information as well as, recursively, all upstream hvf fields and their hop information.\<close>
function (sequential) extr :: "msgterm \<Rightarrow> ahi list" where
"extr (Mac[macKey asid] (L [ts, upif, downif, FS segid]))
= \<lparr>UpIF = term2if upif, DownIF = term2if downif, ASID = asid\<rparr> # (if (\<exists>nextmac asid' upif' downif' segid'.
segid' = xor segid {| nextmac |} \<and>
nextmac = Mac[macKey asid'] (L [ts, upif', downif', FS segid']))
then extr (THE nextmac. (\<exists>asid' upif' downif' segid'.
segid' = xor segid {| nextmac |} \<and>
nextmac = Mac[macKey asid'] (L [ts, upif', downif', FS segid'])))
else [])"
| "extr _ = []"
by pat_completeness auto
termination
apply (relation "measure (\<lambda>x. (case x of Mac[macKey asid] (L [ts, upif, downif, FS segid])
\<Rightarrow> Suc (fcard segid)
| _ \<Rightarrow> 0))")
apply auto
apply(rule theI2)
by(auto elim: FS_update_eq elim!: FS_contains_elem intro!: extr_helper)
text\<open>Extract the authenticated info field from a hop validation field.\<close>
fun extr_ainfo :: "msgterm \<Rightarrow> msgterm" where
"extr_ainfo (Mac[macKey asid] (L (Num ts # xs))) = Num ts"
| "extr_ainfo _ = \<epsilon>"
abbreviation term_ainfo :: "msgterm \<Rightarrow> msgterm" where
"term_ainfo \<equiv> id"
text\<open>The ainfo field must be a Num, since it represents the timestamp (this is only needed for
authorized segments (ainfo, []), since for all other segments, @{text "hf_valid"} enforces this.
Furthermore, we require that the last hop field on l has a MAC that is computed with the empty uinfo
field. This restriction cannot be introduced via @{text "hf_valid"}, since it is not a check
performed by the on-path routers, but rather results from the way that authorized paths are set up
on the control plane. We need this restriction to ensure that the uinfo field of the top node does
not contain extra terms (e.g. secret keys).\<close>
definition auth_restrict where
"auth_restrict ainfo uinfo l \<equiv>
(\<exists>ts. ainfo = Num ts)
\<and> (case l of [] \<Rightarrow> (uinfo = {||}) |
_ \<Rightarrow> hf_valid ainfo {||} (last l) None)"
text\<open>When observing a hop field, an attacker learns the HVF. UHI is empty and the AHI only contains
public information that are not terms.\<close>
fun terms_hf :: "SCION_HF \<Rightarrow> msgterm set" where
"terms_hf hf = {HVF hf}"
text\<open>When analyzing a uinfo field (which is an fset of message terms), the attacker learns all
elements of the fset.\<close>
abbreviation terms_uinfo :: "msgterm fset \<Rightarrow> msgterm set" where
"terms_uinfo \<equiv> fset"
abbreviation no_oracle :: "'ainfo \<Rightarrow> msgterm fset \<Rightarrow> bool" where "no_oracle \<equiv> (\<lambda> _ _. True)"
subsubsection\<open>Properties following from definitions\<close>
text\<open>We now define useful properties of the above definition.\<close>
lemma hf_valid_invert:
"hf_valid tsn uinfo hf nxt \<longleftrightarrow>
((\<exists>ahi ts upif downif asid x.
hf = \<lparr>AHI = ahi, UHI = (), HVF = x\<rparr> \<and>
ASID ahi = asid \<and> ASIF (DownIF ahi) downif \<and> ASIF (UpIF ahi) upif \<and>
x = Mac[macKey asid] (L [tsn, upif, downif, FS uinfo]) \<and>
tsn = Num ts)
)"
by(auto elim!: hf_valid.elims)
lemma info_hvf:
assumes "hf_valid ainfo uinfo m z" "hf_valid ainfo' uinfo' m' z'" "HVF m = HVF m'"
shows "ainfo' = ainfo" "m' = m"
using assms by(auto simp add: hf_valid_invert intro: ahi_eq)
(******************************************************************************)
subsection\<open>Definitions and properties of the added intruder knowledge\<close>
(******************************************************************************)
text\<open>Here we define @{text "ik_add"} and @{text "ik_oracle"} as being empty, as these features are
not used in this instance model.\<close>
print_locale dataplane_3_directed_defs
sublocale dataplane_3_directed_defs _ _ _ auth_seg0 hf_valid auth_restrict extr extr_ainfo term_ainfo
terms_hf terms_uinfo upd_uinfo no_oracle
by unfold_locales
declare TWu.holds_set_list[dest]
declare TWu.holds_takeW_is_identity[simp]
declare parts_singleton[dest]
abbreviation ik_add :: "msgterm set" where "ik_add \<equiv> {}"
abbreviation ik_oracle :: "msgterm set" where "ik_oracle \<equiv> {}"
(******************************************************************************)
subsection\<open>Properties of the intruder knowledge, including @{term "terms_uinfo"}.\<close>
(******************************************************************************)
text\<open>We now instantiate the parametrized model's definition of the intruder knowledge, using the
definitions of @{term "terms_uinfo"} and @{text "ik_oracle"} from above. We then prove the properties
that we need to instantiate the @{text "dataplane_3_directed"} locale.\<close>
print_locale dataplane_3_directed_ik_defs
sublocale
dataplane_3_directed_ik_defs _ _ _ auth_seg0 terms_uinfo no_oracle hf_valid auth_restrict extr extr_ainfo term_ainfo
terms_hf upd_uinfo ik_add ik_oracle
by unfold_locales
text\<open>For this instance model, the neighboring hop field is irrelevant. Hence, if we are interested
in establishing the first hop field's validity given @{text "hfs_valid"}, we do not need to make
a case distinction on the rest of the hop fields (which would normally be required by @{text "TWu"}.\<close>
lemma hfs_valid_first[elim]: "hfs_valid ainfo uinfo (hf # post) nxt \<Longrightarrow> hf_valid ainfo uinfo hf nxt'"
by(cases post, auto simp add: hf_valid_invert TWu.holds.simps)
text\<open>Properties of HVF of valid hop fields that fulfill the restriction.\<close>
lemma auth_properties:
assumes "hf \<in> set hfs" "hfs_valid ainfo uinfo hfs nxt" "auth_restrict ainfo uinfo hfs"
"t = HVF hf"
shows "(\<exists>t' . t = Hash t')
\<and> (\<exists>uinfo'. auth_restrict ainfo uinfo' hfs
\<and> (\<exists>nxt. hf_valid ainfo uinfo' hf nxt))"
using assms
proof(induction uinfo hfs nxt arbitrary: hf rule: TWu.holds.induct[where ?upd=upd_uinfo])
case (1 info x y ys nxt)
then show ?case
proof(cases "hf = x")
case True
then show ?thesis
using 1(2-5) by (auto simp add: TWu.holds.simps(1) hf_valid_invert)
next
case False
then have "hf \<in> set (y # ys)" using 1 by auto
then show ?thesis
apply- apply(drule 1)
subgoal using assms 1(2-5) by (simp add: TWu.holds.simps(1))
using assms(3) 1(2-5) False
by(auto simp add: auth_restrict_def hf_valid_invert)
qed
qed(auto simp add: auth_restrict_def hf_valid_invert intro!: exI)
lemma ik_hfs_form: "t \<in> parts ik_hfs \<Longrightarrow> \<exists> t' . t = Hash t'"
by(auto 3 4 simp add: auth_seg2_def dest: auth_properties)
declare ik_hfs_def[simp del]
lemma parts_ik_hfs[simp]: "parts ik_hfs = ik_hfs"
by (auto intro!: parts_Hash ik_hfs_form)
text\<open>This lemma allows us not only to expand the definition of @{term "ik_hfs"}, but also
to obtain useful properties, such as a term being a Hash, and it being part of a valid hop field.\<close>
lemma ik_hfs_simp:
"t \<in> ik_hfs \<longleftrightarrow> (\<exists>t' . t = Hash t') \<and> (\<exists>hf . t = HVF hf
\<and> (\<exists>hfs uinfo. hf \<in> set hfs \<and> (\<exists>ainfo . (ainfo, hfs) \<in> (auth_seg2 uinfo)
\<and> (\<exists> nxt uinfo'. hf_valid ainfo uinfo' hf nxt))))" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume asm: "?lhs"
then obtain ainfo uinfo hf hfs where
dfs: "hf \<in> set hfs" "(ainfo, hfs) \<in> (auth_seg2 uinfo)" "t = HVF hf"
by(auto simp add: ik_hfs_def)
then have "hfs_valid_None ainfo uinfo hfs" "(ainfo, AHIS hfs) \<in> auth_seg0"
by(auto simp add: auth_seg2_def)
then show "?rhs" using asm dfs by(fast intro: ik_hfs_form)
qed(auto simp add: ik_hfs_def)
text\<open>The following lemma is one of the conditions. We already prove it here, since it is helpful
elsewhere. \<close>
lemma auth_restrict_upd:
"auth_restrict ainfo uinfo (x#y#hfs)
\<Longrightarrow> auth_restrict ainfo (upd_uinfo uinfo y) (y#hfs)"
by (auto simp add: auth_restrict_def)
text\<open>We now show that @{text "ik_uinfo"} is redundant, since all of its terms are already contained in
@{text "ik_hfs"}. To this end, we first show that a term contained in the uinfo field of an authorized
paths is also contained in the HVF of the same path.\<close>
lemma uinfo_contained_in_HVF:
assumes "t \<in> fset uinfo" "(ainfo, hfs) \<in> (auth_seg2 uinfo)"
shows "\<exists>hf. t = HVF hf \<and> hf \<in> set hfs"
proof-
from assms(2) have hfs_defs: "hfs_valid_None ainfo uinfo hfs" "auth_restrict ainfo uinfo hfs"
by(auto simp add: auth_seg2_def)
obtain nxt::"SCION_HF option" where nxt_None[intro]: "nxt = None" by simp
then show ?thesis using hfs_defs assms(1)
proof(induction uinfo hfs nxt rule: TWu.holds.induct[where ?upd=upd_uinfo])
case (1 info x y ys nxt)
then have hf_valid_x: "hf_valid ainfo info x (Some y)" by(auto simp only: TWu.holds.simps)
from 1(2-3,5) show ?case
proof(cases "t = HVF y")
case False
then have "t \<in> fset (upd_uinfo info y)" using 1(2,5) by (simp add: xor_singleton(1))
moreover have "hfs_valid_None ainfo (upd_uinfo info y) (y # ys)"
"auth_restrict ainfo (upd_uinfo info y) (y # ys)"
using 1(3-4) by(auto simp only: TWu.holds.simps elim: auth_restrict_upd)
ultimately have "\<exists>hf. t = HVF hf \<and> hf \<in> set (y # ys)" using 1(1) "1.prems"(1) by blast
then show ?thesis by auto
qed(auto)
next
case (2 info x nxt)
then show ?case
apply(auto simp only: TWu.holds.simps auth_restrict_def)
by (auto simp add: hf_valid_invert)
next
case (3 info nxt)
then show ?case by(auto simp add: auth_restrict_def)
qed
qed
text\<open>The following lemma allows us to ignore @{text "ik_uinfo"} when we unfold @{text "ik"}.\<close>
lemma ik_uinfo_in_ik_hfs: "t \<in> ik_uinfo \<Longrightarrow> t \<in> ik_hfs"
by(auto simp add: ik_hfs_def dest!: uinfo_contained_in_HVF)
(******************************************************************************)
subsubsection \<open>Properties of Intruder Knowledge\<close>
(******************************************************************************)
lemma auth_ainfo[dest]: "\<lbrakk>(ainfo, hfs) \<in> (auth_seg2 uinfo)\<rbrakk> \<Longrightarrow> \<exists> ts . ainfo = Num ts"
by(auto simp add: auth_seg2_def auth_restrict_def)
text\<open>This lemma unfolds the definition of the intruder knowledge but also already applies some
simplifications, such as ignoring @{text "ik_uinfo"}.\<close>
lemma ik_simpler:
"ik = ik_hfs
\<union> {term_ainfo ainfo | ainfo hfs uinfo. (ainfo, hfs) \<in> (auth_seg2 uinfo)}
\<union> Key`(macK`bad)"
by(auto simp add: ik_def simp del: ik_uinfo_def dest: ik_uinfo_in_ik_hfs)
text \<open>There are no ciphertexts (or signatures) in @{term "parts ik"}. Thus, @{term "analz ik"}
and @{term "parts ik"} are identical.\<close>
lemma analz_parts_ik[simp]: "analz ik = parts ik"
by(rule no_crypt_analz_is_parts)
(auto simp add: ik_simpler auth_seg2_def ik_hfs_simp auth_restrict_def)
lemma parts_ik[simp]: "parts ik = ik"
by(auto 3 4 simp add: ik_simpler auth_seg2_def auth_restrict_def)
lemma key_ik_bad: "Key (macK asid) \<in> ik \<Longrightarrow> asid \<in> bad"
by(auto simp add: ik_simpler)
(auto 3 4 simp add: auth_seg2_def ik_hfs_simp hf_valid_invert)
lemma MAC_synth_helper:
assumes "hf_valid ainfo uinfo m z" "HVF m = Mac[Key (macK asid)] j" "HVF m \<in> ik"
shows "\<exists>hfs. m \<in> set hfs \<and> (\<exists>uinfo'. (ainfo, hfs) \<in> auth_seg2 uinfo')"
proof-
from assms(2-3) obtain ainfo' uinfo' uinfo'' m' hfs' nxt' where dfs:
"m' \<in> set hfs'" "(ainfo', hfs') \<in> auth_seg2 uinfo''" "hf_valid ainfo' uinfo' m' nxt'"
"HVF m = HVF m'"
by(auto simp add: ik_simpler ik_hfs_simp)
then have eqs[simp]: "ainfo' = ainfo" "m' = m" using assms(1) by(auto elim!: info_hvf)
have "auth_restrict ainfo' uinfo'' hfs'" using dfs by(auto simp add: auth_seg2_def)
then show ?thesis using dfs assms by auto
qed
text\<open>This definition helps with the limiting the number of cases generated. We don't require it,
but it is convenient. Given a hop validation field and an asid, return if the hvf has the expected
format.\<close>
definition mac_format :: "msgterm \<Rightarrow> as \<Rightarrow> bool" where
"mac_format m asid \<equiv> \<exists> j . m = Mac[macKey asid] j"
text\<open>If a valid hop field is derivable by the attacker, but does not belong to the attacker, then
the hop field is already contained in the set of authorized segments.\<close>
lemma MAC_synth:
assumes "hf_valid ainfo uinfo m z" "HVF m \<in> synth ik" "mac_format (HVF m) asid" "asid \<notin> bad"
shows "\<exists>hfs. m \<in> set hfs \<and>
(\<exists>uinfo'. (ainfo, hfs) \<in> auth_seg2 uinfo')"
using assms
by(auto simp add: mac_format_def ik_simpler ik_hfs_simp elim!: MAC_synth_helper dest!: key_ik_bad)
(******************************************************************************)
subsection\<open>Lemmas helping with conditions relating to extract\<close>
(******************************************************************************)
text\<open>Resolve the definite descriptor operator THE.\<close>
lemma THE_nextmac:
assumes "hvf = Mac[macKey askey] (L [Num ts, upif, downif, FS (xor info {|hvf|})])"
shows "(THE nextmac. \<exists>asid' upif' downif'.
nextmac = Mac[macKey asid'] (L [Num ts, upif', downif', FS (xor info {|nextmac|})]))
= hvf"
apply(rule theI2[of _ hvf])
using assms
by(auto elim!: FS_update_eq[of _ _ info "hvf"])
lemma hf_valid_uinfo:
assumes "hf_valid ainfo (upd_uinfo uinfo y) y nxt" "hvfy = HVF y"
shows "hvfy \<in> fset uinfo"
apply (cases y)
using assms by(auto simp add: hf_valid_invert elim!: FS_contains_elem)
text\<open>A single step of extract. Extract on a single valid hop field is equivalent to that hop field's
hop info field concat extract on the next hop field, where the next hop field has to be valid
with uinfo updated.\<close>
lemma extr_hf_valid:
assumes "hf_valid ainfo uinfo x nxt" "hf_valid ainfo (upd_uinfo uinfo y) y nxt'"
shows "extr (HVF x) = AHI x # extr (HVF y)"
proof-
obtain uinfo' where info'_def: "uinfo' = xor uinfo {|HVF y|}" by simp
obtain ts ahi upif downif hvfx ahi' upif' downif' hvfy where unfolded_defs:
"x = \<lparr>AHI = ahi, UHI = (), HVF = hvfx\<rparr>"
"ASIF (UpIF ahi) upif"
"ASIF (DownIF ahi) downif"
"hvfx = Mac[macKey (ASID ahi)] (L [Num ts, upif, downif, FS uinfo])"
"y = \<lparr>AHI = ahi', UHI = (), HVF = hvfy\<rparr>"
"ASIF (UpIF ahi') upif'"
"ASIF (DownIF ahi') downif'"
"hvfy = Mac[macKey (ASID ahi')] (L [Num ts, upif', downif', FS (uinfo')])"
using assms apply(auto simp only: hf_valid_invert) by (auto simp add: info'_def)
have hvfy_in_uinfo: "hvfy \<in> fset uinfo"
using assms(2) apply(auto intro!: hf_valid_uinfo) using unfolded_defs by simp
then obtain fcard_uinfo_minus1 where "fcard uinfo = Suc fcard_uinfo_minus1"
- by (metis fcard_Suc_fminus1 notin_fset)
+ by (metis fcard_Suc_fminus1)
then show ?thesis
apply(cases y)
using unfolded_defs(1-7) apply (auto intro!: ahi_eq)
subgoal for nextmac (*difficult case*)
apply(auto simp add: THE_nextmac dest!: FS_update_eq[of "nextmac" _
_ "hvfy" "(\<lambda>x. Mac[macKey (ASID ahi')] (L [Num ts, upif', downif', x]))"])
using unfolded_defs(8) info'_def by fastforce+
using hvfy_in_uinfo unfolded_defs(8) info'_def
by (fastforce dest!: fcard_Suc_fminus1[simplified] elim!: allE[of _ "HVF y"])
qed
(******************************************************************************)
subsection\<open>Direct proof goals for interpretation of @{text "dataplane_3_directed"}\<close>
(******************************************************************************)
lemma COND_honest_hf_analz:
assumes "ASID (AHI hf) \<notin> bad" "hf_valid ainfo uinfo hf nxt" "terms_hf hf \<subseteq> synth (analz ik)"
"no_oracle ainfo uinfo"
shows "terms_hf hf \<subseteq> analz ik"
proof-
let ?asid = "ASID (AHI hf)"
from assms(3) have hf_synth_ik: "HVF hf \<in> synth ik" by auto
from assms(2) have "mac_format (HVF hf) ?asid"
by(auto simp add: mac_format_def hf_valid_invert)
then obtain hfs uinfo' where
"hf \<in> set hfs" "(ainfo, hfs) \<in> auth_seg2 uinfo'"
using assms(1,2) hf_synth_ik by(auto dest!: MAC_synth)
then have "HVF hf \<in> ik"
using assms(2)
by(auto simp add: ik_hfs_def intro!: ik_ik_hfs intro!: exI)
then show ?thesis by auto
qed
lemma COND_terms_hf:
assumes "hf_valid ainfo uinfo hf nxt" "terms_hf hf \<subseteq> analz ik"
"no_oracle ainfo uinfo"
shows "\<exists>hfs. hf \<in> set hfs \<and> (\<exists>uinfo' . (ainfo, hfs) \<in> (auth_seg2 uinfo'))"
proof-
obtain hfs ainfo uinfo where hfs_def: "hf \<in> set hfs" "(ainfo, hfs) \<in> auth_seg2 uinfo"
using assms
by(simp only: analz_parts_ik parts_ik)
(auto 3 4 simp add: hf_valid_invert ik_hfs_simp ik_simpler dest: ahi_eq)
show ?thesis
using hfs_def apply (auto simp add: auth_seg2_def dest!: TWu.holds_set_list)
using hfs_def assms(1) by (auto simp add: auth_seg2_def dest: info_hvf)
qed
lemmas COND_auth_restrict_upd = auth_restrict_upd
lemma COND_extr_prefix_path:
"\<lbrakk>hfs_valid ainfo uinfo l nxt; auth_restrict ainfo uinfo l\<rbrakk> \<Longrightarrow> prefix (extr_from_hd l) (AHIS l)"
proof(induction l nxt rule: TWu.holds.induct[where ?upd=upd_uinfo])
case (1 info x y ys nxt)
from 1(2-3) have hfs_valid:
"hfs_valid ainfo (upd_uinfo info y) (y # ys) nxt"
"auth_restrict ainfo (upd_uinfo info y) (y # ys)"
by(auto simp only: TWu.holds.simps intro!: auth_restrict_upd)
then have prefix_y: "prefix (extr_from_hd (y # ys)) (AHIS (y # ys))" by(rule 1(1))
have "extr_from_hd (x # y # ys) = AHI x # extr_from_hd (y # ys)"
apply(cases ys)
using 1(2) by(auto simp only: extr_from_hd.simps TWu.holds.simps intro!: extr_hf_valid)
then show ?case
using prefix_y by (auto)
qed(auto simp only: TWu.holds.simps hf_valid_invert,
auto simp add: fcard_fempty auth_restrict_def intro!: ahi_eq dest!: FS_contr)
lemma COND_path_prefix_extr:
"prefix (AHIS (hfs_valid_prefix ainfo uinfo l nxt))
(extr_from_hd l)"
proof(induction l nxt rule: TWu.takeW.induct[where ?Pa="hf_valid ainfo", where ?upd=upd_uinfo])
case (2 info x xo)
then show ?case
apply(cases "fcard info")
by(auto 3 4 intro!: ahi_eq simp add: fcard_fempty TWu.takeW_split_tail TWu.takeW.simps(1) hf_valid_invert)
next
case (4 info x y xs xo)
from 4(1) show ?case
proof (cases "hf_valid ainfo (upd_uinfo info y) y (head xs)")
case hf_valid_y: True
obtain info' where info'_def: "info' = xor info {|HVF y|}" by simp
show ?thesis
proof(rule prefix_cons[where ?ys="extr_from_hd (y # xs)", where ?x = "AHI x"])
show "extr_from_hd (x # y # xs) = AHI x # extr_from_hd (y # xs)"
using hf_valid_y 4(1)
by(auto simp del: upd_uinfo_def elim!: extr_hf_valid[rotated])
next
have "prefix (map AHI (hfs_valid_prefix ainfo (upd_uinfo info y) (y # xs) xo)) (extr (HVF y))"
using 4(2) by (auto simp del: upd_uinfo_def)
then show "prefix (AHIS (hfs_valid_prefix ainfo info (x # y # xs) xo)) (AHI x # extr (HVF y))"
by (auto simp add: TWu.takeW_split_tail[where ?x = x] TWu.takeW.simps(1) simp del: upd_uinfo_def)
qed(auto)
next
case False
then show ?thesis
by(auto simp add: TWu.takeW_split_tail hf_valid_invert)
(auto simp add: fcard_fempty intro!: ahi_eq)
qed
qed(auto simp add: TWu.takeW_split_tail TWu.takeW.simps(1))
lemma COND_hf_valid_uinfo:
"\<lbrakk>hf_valid ainfo uinfo hf nxt; hf_valid ainfo' uinfo' hf nxt'\<rbrakk> \<Longrightarrow> uinfo' = uinfo"
by(auto simp add: hf_valid_invert)
lemma COND_upd_uinfo_ik:
"\<lbrakk>terms_uinfo uinfo \<subseteq> synth (analz ik); terms_hf hf \<subseteq> synth (analz ik)\<rbrakk>
\<Longrightarrow> terms_uinfo (upd_uinfo uinfo hf) \<subseteq> synth (analz ik)"
by fastforce
lemma COND_upd_uinfo_no_oracle:
"no_oracle ainfo uinfo \<Longrightarrow> no_oracle ainfo (upd_uinfo uinfo fld)"
by (auto)
(******************************************************************************)
subsection\<open>Instantiation of @{text "dataplane_3_directed"} locale\<close>
(******************************************************************************)
print_locale dataplane_3_directed
sublocale
dataplane_3_directed _ _ _ auth_seg0 terms_uinfo terms_hf hf_valid auth_restrict extr extr_ainfo term_ainfo
upd_uinfo ik_add
ik_oracle no_oracle
apply unfold_locales
using COND_terms_hf COND_honest_hf_analz COND_extr_prefix_path
COND_path_prefix_extr COND_hf_valid_uinfo COND_upd_uinfo_ik COND_upd_uinfo_no_oracle
COND_auth_restrict_upd by auto
(******************************************************************************)
subsection\<open>Normalization of terms\<close>
(******************************************************************************)
text\<open>We now show that all terms that occur in reachable states are normalized, meaning that they
do not have directly nested FSets. For instance, a term @{term "FS {| FS {| Num 0 |}, Num 0 |}"} is
not normalized, whereas @{term "FS {| Hash (FS {| Num 0 |}), Num 0 |}"} is normalized.\<close>
lemma normalized_upd:
"\<lbrakk>normalized (FS (upd_uinfo info y)); normalized (FS {| HVF y |})\<rbrakk>
\<Longrightarrow> normalized (FS info)"
by(auto simp add: xor_singleton)
declare normalized.Lst[intro!] normalized.FSt[intro!] normalized.Hash[intro!] normalized.MPair[intro!]
lemma auth_uinfo_normalized:
"\<lbrakk>hfs_valid ainfo uinfo hfs nxt; auth_restrict ainfo uinfo hfs\<rbrakk> \<Longrightarrow> normalized (FS uinfo)"
proof(induction hfs nxt arbitrary: rule: TWu.holds.induct[where ?upd=upd_uinfo])
case (1 info x y ys nxt)
from 1 have hfs_valid: "hf_valid ainfo info x (Some y)"
"hfs_valid ainfo (upd_uinfo info y) (y # ys) nxt"
"auth_restrict ainfo (upd_uinfo info y) (y # ys)"
by(auto simp only: TWu.holds.simps intro!: auth_restrict_upd)
then have normalized_upd_y: "normalized (FS (upd_uinfo info y))" by (elim 1(1))
obtain z where hfy_valid: "hf_valid ainfo (upd_uinfo info y) y z"
using hfs_valid(2) by(auto dest: hfs_valid_first)
obtain info_s where info_s_def[simp]: "xor info {|HVF y|} = info_s" by simp
from normalized_upd_y show ?case
apply(auto elim!: normalized_upd simp only:)
using hfy_valid info_s_def normalized_upd_y
by (auto 3 4 simp add: hf_valid_invert elim: ASIF.elims(2))
qed(auto simp only: TWu.holds.simps auth_restrict_def,
auto simp add: hf_valid_invert)
lemma auth_normalized_hf:
assumes "auth_restrict ainfo uinfo (pre @ hf # post)"
"hfs_valid ainfo (upds_uinfo_shifted uinfo pre hf) (hf # post) nxt"
"upds_uinfo_shifted uinfo pre hf = hf_uinfo"
shows "normalized (HVF hf)"
using assms(1-2)
apply(auto dest!: hfs_valid_first simp add: hf_valid_invert assms(3))
using assms(2-3)
by(auto dest!: auth_uinfo_normalized dest: auth_restrict_app elim: ASIF.elims(2))
lemma auth_normalized:
"\<lbrakk>hf \<in> set hfs; hfs_valid ainfo uinfo hfs nxt; auth_restrict ainfo uinfo hfs\<rbrakk>
\<Longrightarrow> normalized (HVF hf)"
by(auto dest: TWu.holds_intermediate_ex intro: auth_normalized_hf)
text\<open>All terms derivable by the intruder are normalized. Note that (i) the dynamic intruder knowledge
@{text "ik_dyn"} contains all terms of messages contained in the state and (ii) the dynamic intruder
knowledge remains constant. Hence this lemma suffices to show that all terms contained in @{text "int"}
and @{text "ext"} channels of reachable states are normalized as well.\<close>
lemma ik_synth_normalized: "t \<in> synth (analz ik) \<Longrightarrow> normalized t"
by (auto, auto simp add: ik_simpler)
(auto simp add: ik_hfs_def auth_seg2_def hfs_valid_prefix_generic_def
elim!: auth_normalized)
end
end
diff --git a/thys/LTL/Disjunctive_Normal_Form.thy b/thys/LTL/Disjunctive_Normal_Form.thy
--- a/thys/LTL/Disjunctive_Normal_Form.thy
+++ b/thys/LTL/Disjunctive_Normal_Form.thy
@@ -1,1092 +1,1092 @@
(*
Author: Benedikt Seidl
Author: Salomon Sickert
License: BSD
*)
section \<open>Disjunctive Normal Form of LTL formulas\<close>
theory Disjunctive_Normal_Form
imports
LTL Equivalence_Relations "HOL-Library.FSet"
begin
text \<open>
We use the propositional representation of LTL formulas to define
the minimal disjunctive normal form of our formulas. For this purpose
we define the minimal product \<open>\<otimes>\<^sub>m\<close> and union \<open>\<union>\<^sub>m\<close>.
In the end we show that for a set \<open>\<A>\<close> of literals,
@{term "\<A> \<Turnstile>\<^sub>P \<phi>"} if, and only if, there exists a subset
of \<open>\<A>\<close> in the minimal DNF of \<open>\<phi>\<close>.
\<close>
subsection \<open>Definition of Minimum Sets\<close>
definition (in ord) min_set :: "'a set \<Rightarrow> 'a set" where
"min_set X = {y \<in> X. \<forall>x \<in> X. x \<le> y \<longrightarrow> x = y}"
lemma min_set_iff:
"x \<in> min_set X \<longleftrightarrow> x \<in> X \<and> (\<forall>y \<in> X. y \<le> x \<longrightarrow> y = x)"
unfolding min_set_def by blast
lemma min_set_subset:
"min_set X \<subseteq> X"
by (auto simp: min_set_def)
lemma min_set_idem[simp]:
"min_set (min_set X) = min_set X"
by (auto simp: min_set_def)
lemma min_set_empty[simp]:
"min_set {} = {}"
using min_set_subset by blast
lemma min_set_singleton[simp]:
"min_set {x} = {x}"
by (auto simp: min_set_def)
lemma min_set_finite:
"finite X \<Longrightarrow> finite (min_set X)"
by (simp add: min_set_def)
lemma min_set_obtains_helper:
"A \<in> B \<Longrightarrow> \<exists>C. C |\<subseteq>| A \<and> C \<in> min_set B"
proof (induction "fcard A" arbitrary: A rule: less_induct)
case less
then have "(\<forall>A'. A' \<notin> B \<or> \<not> A' |\<subseteq>| A \<or> A' = A) \<or> (\<exists>A'. A' |\<subseteq>| A \<and> A' \<in> min_set B)"
by (metis (no_types) dual_order.trans order.not_eq_order_implies_strict pfsubset_fcard_mono)
then show ?case
using less.prems min_set_def by auto
qed
lemma min_set_obtains:
assumes "A \<in> B"
obtains C where "C |\<subseteq>| A" and "C \<in> min_set B"
using min_set_obtains_helper assms by metis
subsection \<open>Minimal operators on sets\<close>
definition product :: "'a fset set \<Rightarrow> 'a fset set \<Rightarrow> 'a fset set" (infixr "\<otimes>" 65)
where "A \<otimes> B = {a |\<union>| b | a b. a \<in> A \<and> b \<in> B}"
definition min_product :: "'a fset set \<Rightarrow> 'a fset set \<Rightarrow> 'a fset set" (infixr "\<otimes>\<^sub>m" 65)
where "A \<otimes>\<^sub>m B = min_set (A \<otimes> B)"
definition min_union :: "'a fset set \<Rightarrow> 'a fset set \<Rightarrow> 'a fset set" (infixr "\<union>\<^sub>m" 65)
where "A \<union>\<^sub>m B = min_set (A \<union> B)"
definition product_set :: "'a fset set set \<Rightarrow> 'a fset set" ("\<Otimes>")
where "\<Otimes> X = Finite_Set.fold product {{||}} X"
definition min_product_set :: "'a fset set set \<Rightarrow> 'a fset set" ("\<Otimes>\<^sub>m")
where "\<Otimes>\<^sub>m X = Finite_Set.fold min_product {{||}} X"
lemma min_product_idem[simp]:
"A \<otimes>\<^sub>m A = min_set A"
by (auto simp: min_product_def product_def min_set_def) fastforce
lemma min_union_idem[simp]:
"A \<union>\<^sub>m A = min_set A"
by (simp add: min_union_def)
lemma product_empty[simp]:
"A \<otimes> {} = {}"
"{} \<otimes> A = {}"
by (simp_all add: product_def)
lemma min_product_empty[simp]:
"A \<otimes>\<^sub>m {} = {}"
"{} \<otimes>\<^sub>m A = {}"
by (simp_all add: min_product_def)
lemma min_union_empty[simp]:
"A \<union>\<^sub>m {} = min_set A"
"{} \<union>\<^sub>m A = min_set A"
by (simp_all add: min_union_def)
lemma product_empty_singleton[simp]:
"A \<otimes> {{||}} = A"
"{{||}} \<otimes> A = A"
by (simp_all add: product_def)
lemma min_product_empty_singleton[simp]:
"A \<otimes>\<^sub>m {{||}} = min_set A"
"{{||}} \<otimes>\<^sub>m A = min_set A"
by (simp_all add: min_product_def)
lemma product_singleton_singleton:
"A \<otimes> {{|x|}} = finsert x ` A"
"{{|x|}} \<otimes> A = finsert x ` A"
unfolding product_def by blast+
lemma product_mono:
"A \<subseteq> B \<Longrightarrow> A \<otimes> C \<subseteq> B \<otimes> C"
"B \<subseteq> C \<Longrightarrow> A \<otimes> B \<subseteq> A \<otimes> C"
unfolding product_def by auto
lemma product_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<otimes> B)"
by (simp add: product_def finite_image_set2)
lemma min_product_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<otimes>\<^sub>m B)"
by (metis min_product_def product_finite min_set_finite)
lemma min_union_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<union>\<^sub>m B)"
by (simp add: min_union_def min_set_finite)
lemma product_set_infinite[simp]:
"infinite X \<Longrightarrow> \<Otimes> X = {{||}}"
by (simp add: product_set_def)
lemma min_product_set_infinite[simp]:
"infinite X \<Longrightarrow> \<Otimes>\<^sub>m X = {{||}}"
by (simp add: min_product_set_def)
lemma product_comm:
"A \<otimes> B = B \<otimes> A"
unfolding product_def by blast
lemma min_product_comm:
"A \<otimes>\<^sub>m B = B \<otimes>\<^sub>m A"
unfolding min_product_def
by (simp add: product_comm)
lemma min_union_comm:
"A \<union>\<^sub>m B = B \<union>\<^sub>m A"
unfolding min_union_def
by (simp add: sup.commute)
lemma product_iff:
"x \<in> A \<otimes> B \<longleftrightarrow> (\<exists>a \<in> A. \<exists>b \<in> B. x = a |\<union>| b)"
unfolding product_def by blast
lemma min_product_iff:
"x \<in> A \<otimes>\<^sub>m B \<longleftrightarrow> (\<exists>a \<in> A. \<exists>b \<in> B. x = a |\<union>| b) \<and> (\<forall>a \<in> A. \<forall>b \<in> B. a |\<union>| b |\<subseteq>| x \<longrightarrow> a |\<union>| b = x)"
unfolding min_product_def min_set_iff product_iff product_def by blast
lemma min_union_iff:
"x \<in> A \<union>\<^sub>m B \<longleftrightarrow> x \<in> A \<union> B \<and> (\<forall>a \<in> A. a |\<subseteq>| x \<longrightarrow> a = x) \<and> (\<forall>b \<in> B. b |\<subseteq>| x \<longrightarrow> b = x)"
unfolding min_union_def min_set_iff by blast
lemma min_set_min_product_helper:
"x \<in> (min_set A) \<otimes>\<^sub>m B \<longleftrightarrow> x \<in> A \<otimes>\<^sub>m B"
proof
fix x
assume "x \<in> (min_set A) \<otimes>\<^sub>m B"
then obtain a b where "a \<in> min_set A" and "b \<in> B" and "x = a |\<union>| b" and 1: "\<forall>a \<in> min_set A. \<forall>b \<in> B. a |\<union>| b |\<subseteq>| x \<longrightarrow> a |\<union>| b = x"
unfolding min_product_iff by blast
moreover
{
fix a' b'
assume "a' \<in> A" and "b' \<in> B" and "a' |\<union>| b' |\<subseteq>| x"
then obtain a'' where "a'' |\<subseteq>| a'" and "a'' \<in> min_set A"
using min_set_obtains by metis
then have "a'' |\<union>| b' = x"
by (metis (full_types) 1 \<open>b' \<in> B\<close> \<open>a' |\<union>| b' |\<subseteq>| x\<close> dual_order.trans le_sup_iff)
then have "a' |\<union>| b' = x"
using \<open>a' |\<union>| b' |\<subseteq>| x\<close> \<open>a'' |\<subseteq>| a'\<close> by blast
}
ultimately show "x \<in> A \<otimes>\<^sub>m B"
by (metis min_product_iff min_set_iff)
next
fix x
assume "x \<in> A \<otimes>\<^sub>m B"
then have 1: "x \<in> A \<otimes> B" and "\<forall>y \<in> A \<otimes> B. y |\<subseteq>| x \<longrightarrow> y = x"
unfolding min_product_def min_set_iff by simp+
then have 2: "\<forall>y\<in>min_set A \<otimes> B. y |\<subseteq>| x \<longrightarrow> y = x"
by (metis product_iff min_set_iff)
then have "x \<in> min_set A \<otimes> B"
by (metis 1 funion_mono min_set_obtains order_refl product_iff)
then show "x \<in> min_set A \<otimes>\<^sub>m B"
by (simp add: 2 min_product_def min_set_iff)
qed
lemma min_set_min_product[simp]:
"(min_set A) \<otimes>\<^sub>m B = A \<otimes>\<^sub>m B"
"A \<otimes>\<^sub>m (min_set B) = A \<otimes>\<^sub>m B"
using min_product_comm min_set_min_product_helper by blast+
lemma min_set_min_union[simp]:
"(min_set A) \<union>\<^sub>m B = A \<union>\<^sub>m B"
"A \<union>\<^sub>m (min_set B) = A \<union>\<^sub>m B"
proof (unfold min_union_def min_set_def, safe)
show "\<And>x xa xb. \<lbrakk>\<forall>xa\<in>{y \<in> A. \<forall>x\<in>A. x |\<subseteq>| y \<longrightarrow> x = y} \<union> B. xa |\<subseteq>| x \<longrightarrow> xa = x; x \<in> B; xa |\<subseteq>| x; xb |\<in>| x; xa \<in> A\<rbrakk> \<Longrightarrow> xb |\<in>| xa"
by (metis (mono_tags) UnCI dual_order.trans fequalityI min_set_def min_set_obtains)
next
show "\<And>x xa xb. \<lbrakk>\<forall>xa\<in>A \<union> {y \<in> B. \<forall>x\<in>B. x |\<subseteq>| y \<longrightarrow> x = y}. xa |\<subseteq>| x \<longrightarrow> xa = x; x \<in> A; xa |\<subseteq>| x; xb |\<in>| x; xa \<in> B\<rbrakk> \<Longrightarrow> xb |\<in>| xa"
by (metis (mono_tags) UnCI dual_order.trans fequalityI min_set_def min_set_obtains)
qed blast+
lemma product_assoc[simp]:
"(A \<otimes> B) \<otimes> C = A \<otimes> (B \<otimes> C)"
proof (unfold product_def, safe)
fix a b c
assume "a \<in> A" and "c \<in> C" and "b \<in> B"
then have "b |\<union>| c \<in> {b |\<union>| c |b c. b \<in> B \<and> c \<in> C}"
by blast
then show "\<exists>a' bc. a |\<union>| b |\<union>| c = a' |\<union>| bc \<and> a' \<in> A \<and> bc \<in> {b |\<union>| c |b c. b \<in> B \<and> c \<in> C}"
using \<open>a \<in> A\<close> by (metis (no_types) inf_sup_aci(5) sup_left_commute)
qed (metis (mono_tags, lifting) mem_Collect_eq sup_assoc)
lemma min_product_assoc[simp]:
"(A \<otimes>\<^sub>m B) \<otimes>\<^sub>m C = A \<otimes>\<^sub>m (B \<otimes>\<^sub>m C)"
unfolding min_product_def[of A B] min_product_def[of B C]
by simp (simp add: min_product_def)
lemma min_union_assoc[simp]:
"(A \<union>\<^sub>m B) \<union>\<^sub>m C = A \<union>\<^sub>m (B \<union>\<^sub>m C)"
unfolding min_union_def[of A B] min_union_def[of B C]
by simp (simp add: min_union_def sup_assoc)
lemma min_product_comp:
"a \<in> A \<Longrightarrow> b \<in> B \<Longrightarrow> \<exists>c. c |\<subseteq>| (a |\<union>| b) \<and> c \<in> A \<otimes>\<^sub>m B"
by (metis (mono_tags, lifting) mem_Collect_eq min_product_def product_def min_set_obtains)
lemma min_union_comp:
"a \<in> A \<Longrightarrow> \<exists>c. c |\<subseteq>| a \<and> c \<in> A \<union>\<^sub>m B"
by (metis Un_iff min_set_obtains min_union_def)
interpretation product_set_thms: Finite_Set.comp_fun_commute product
proof unfold_locales
have "\<And>x y z. x \<otimes> (y \<otimes> z) = y \<otimes> (x \<otimes> z)"
by (simp only: product_assoc[symmetric]) (simp only: product_comm)
then show "\<And>x y. (\<otimes>) y \<circ> (\<otimes>) x = (\<otimes>) x \<circ> (\<otimes>) y"
by fastforce
qed
interpretation min_product_set_thms: Finite_Set.comp_fun_idem min_product
proof unfold_locales
have "\<And>x y z. x \<otimes>\<^sub>m (y \<otimes>\<^sub>m z) = y \<otimes>\<^sub>m (x \<otimes>\<^sub>m z)"
by (simp only: min_product_assoc[symmetric]) (simp only: min_product_comm)
then show "\<And>x y. (\<otimes>\<^sub>m) y \<circ> (\<otimes>\<^sub>m) x = (\<otimes>\<^sub>m) x \<circ> (\<otimes>\<^sub>m) y"
by fastforce
next
have "\<And>x y. x \<otimes>\<^sub>m (x \<otimes>\<^sub>m y) = x \<otimes>\<^sub>m y"
by (simp add: min_product_assoc[symmetric])
then show "\<And>x. (\<otimes>\<^sub>m) x \<circ> (\<otimes>\<^sub>m) x = (\<otimes>\<^sub>m) x"
by fastforce
qed
interpretation min_union_set_thms: Finite_Set.comp_fun_idem min_union
proof unfold_locales
have "\<And>x y z. x \<union>\<^sub>m (y \<union>\<^sub>m z) = y \<union>\<^sub>m (x \<union>\<^sub>m z)"
by (simp only: min_union_assoc[symmetric]) (simp only: min_union_comm)
then show "\<And>x y. (\<union>\<^sub>m) y \<circ> (\<union>\<^sub>m) x = (\<union>\<^sub>m) x \<circ> (\<union>\<^sub>m) y"
by fastforce
next
have "\<And>x y. x \<union>\<^sub>m (x \<union>\<^sub>m y) = x \<union>\<^sub>m y"
by (simp add: min_union_assoc[symmetric])
then show "\<And>x. (\<union>\<^sub>m) x \<circ> (\<union>\<^sub>m) x = (\<union>\<^sub>m) x"
by fastforce
qed
lemma product_set_empty[simp]:
"\<Otimes> {} = {{||}}"
"\<Otimes> {{}} = {}"
"\<Otimes> {{{||}}} = {{||}}"
by (simp_all add: product_set_def)
lemma min_product_set_empty[simp]:
"\<Otimes>\<^sub>m {} = {{||}}"
"\<Otimes>\<^sub>m {{}} = {}"
"\<Otimes>\<^sub>m {{{||}}} = {{||}}"
by (simp_all add: min_product_set_def)
lemma product_set_code[code]:
"\<Otimes> (set xs) = fold product (remdups xs) {{||}}"
by (simp add: product_set_def product_set_thms.fold_set_fold_remdups)
lemma min_product_set_code[code]:
"\<Otimes>\<^sub>m (set xs) = fold min_product (remdups xs) {{||}}"
by (simp add: min_product_set_def min_product_set_thms.fold_set_fold_remdups)
lemma product_set_insert[simp]:
"finite X \<Longrightarrow> \<Otimes> (insert x X) = x \<otimes> (\<Otimes> (X - {x}))"
unfolding product_set_def product_set_thms.fold_insert_remove ..
lemma min_product_set_insert[simp]:
"finite X \<Longrightarrow> \<Otimes>\<^sub>m (insert x X) = x \<otimes>\<^sub>m (\<Otimes>\<^sub>m X)"
unfolding min_product_set_def min_product_set_thms.fold_insert_idem ..
lemma min_product_subseteq:
"x \<in> A \<otimes>\<^sub>m B \<Longrightarrow> \<exists>a. a |\<subseteq>| x \<and> a \<in> A"
by (metis funion_upper1 min_product_iff)
lemma min_product_set_subseteq:
"finite X \<Longrightarrow> x \<in> \<Otimes>\<^sub>m X \<Longrightarrow> A \<in> X \<Longrightarrow> \<exists>a \<in> A. a |\<subseteq>| x"
by (induction X rule: finite_induct) (blast, metis finite_insert insert_absorb min_product_set_insert min_product_subseteq)
lemma min_set_product_set:
"\<Otimes>\<^sub>m A = min_set (\<Otimes> A)"
by (cases "finite A", induction A rule: finite_induct) (simp_all add: min_product_set_def product_set_def, metis min_product_def)
lemma min_product_min_set[simp]:
"min_set (A \<otimes>\<^sub>m B) = A \<otimes>\<^sub>m B"
by (simp add: min_product_def)
lemma min_union_min_set[simp]:
"min_set (A \<union>\<^sub>m B) = A \<union>\<^sub>m B"
by (simp add: min_union_def)
lemma min_product_set_min_set[simp]:
"finite X \<Longrightarrow> min_set (\<Otimes>\<^sub>m X) = \<Otimes>\<^sub>m X"
by (induction X rule: finite_induct, auto simp add: min_product_set_def min_set_iff)
lemma min_set_min_product_set[simp]:
"finite X \<Longrightarrow> \<Otimes>\<^sub>m (min_set ` X) = \<Otimes>\<^sub>m X"
by (induction X rule: finite_induct) simp_all
lemma min_product_set_union[simp]:
"finite X \<Longrightarrow> finite Y \<Longrightarrow> \<Otimes>\<^sub>m (X \<union> Y) = (\<Otimes>\<^sub>m X) \<otimes>\<^sub>m (\<Otimes>\<^sub>m Y)"
by (induction X rule: finite_induct) simp_all
lemma product_set_finite:
"(\<And>x. x \<in> X \<Longrightarrow> finite x) \<Longrightarrow> finite (\<Otimes> X)"
by (cases "finite X", rotate_tac, induction X rule: finite_induct) (simp_all add: product_set_def, insert product_finite, blast)
lemma min_product_set_finite:
"(\<And>x. x \<in> X \<Longrightarrow> finite x) \<Longrightarrow> finite (\<Otimes>\<^sub>m X)"
by (cases "finite X", rotate_tac, induction X rule: finite_induct) (simp_all add: min_product_set_def, insert min_product_finite, blast)
subsection \<open>Disjunctive Normal Form\<close>
fun dnf :: "'a ltln \<Rightarrow> 'a ltln fset set"
where
"dnf true\<^sub>n = {{||}}"
| "dnf false\<^sub>n = {}"
| "dnf (\<phi> and\<^sub>n \<psi>) = (dnf \<phi>) \<otimes> (dnf \<psi>)"
| "dnf (\<phi> or\<^sub>n \<psi>) = (dnf \<phi>) \<union> (dnf \<psi>)"
| "dnf \<phi> = {{|\<phi>|}}"
fun min_dnf :: "'a ltln \<Rightarrow> 'a ltln fset set"
where
"min_dnf true\<^sub>n = {{||}}"
| "min_dnf false\<^sub>n = {}"
| "min_dnf (\<phi> and\<^sub>n \<psi>) = (min_dnf \<phi>) \<otimes>\<^sub>m (min_dnf \<psi>)"
| "min_dnf (\<phi> or\<^sub>n \<psi>) = (min_dnf \<phi>) \<union>\<^sub>m (min_dnf \<psi>)"
| "min_dnf \<phi> = {{|\<phi>|}}"
lemma dnf_min_set:
"min_dnf \<phi> = min_set (dnf \<phi>)"
by (induction \<phi>) (simp_all, simp_all only: min_product_def min_union_def)
lemma dnf_finite:
"finite (dnf \<phi>)"
by (induction \<phi>) (auto simp: product_finite)
lemma min_dnf_finite:
"finite (min_dnf \<phi>)"
by (induction \<phi>) (auto simp: min_product_finite min_union_finite)
lemma dnf_Abs_fset[simp]:
"fset (Abs_fset (dnf \<phi>)) = dnf \<phi>"
by (simp add: dnf_finite Abs_fset_inverse)
lemma min_dnf_Abs_fset[simp]:
"fset (Abs_fset (min_dnf \<phi>)) = min_dnf \<phi>"
by (simp add: min_dnf_finite Abs_fset_inverse)
lemma dnf_prop_atoms:
"\<Phi> \<in> dnf \<phi> \<Longrightarrow> fset \<Phi> \<subseteq> prop_atoms \<phi>"
- by (induction \<phi> arbitrary: \<Phi>) (auto simp: product_def, blast+)
+ by (induction \<phi> arbitrary: \<Phi>) (auto simp: product_def)
lemma min_dnf_prop_atoms:
"\<Phi> \<in> min_dnf \<phi> \<Longrightarrow> fset \<Phi> \<subseteq> prop_atoms \<phi>"
using dnf_min_set dnf_prop_atoms min_set_subset by blast
lemma min_dnf_atoms_dnf:
"\<Phi> \<in> min_dnf \<psi> \<Longrightarrow> \<phi> \<in> fset \<Phi> \<Longrightarrow> dnf \<phi> = {{|\<phi>|}}"
proof (induction \<phi>)
case True_ltln
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(1) by blast
next
case False_ltln
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(2) by blast
next
case (And_ltln \<phi>1 \<phi>2)
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(3) by force
next
case (Or_ltln \<phi>1 \<phi>2)
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(4) by force
qed auto
lemma min_dnf_min_set[simp]:
"min_set (min_dnf \<phi>) = min_dnf \<phi>"
by (induction \<phi>) (simp_all add: min_set_def min_product_def min_union_def, blast+)
lemma min_dnf_iff_prop_assignment_subset:
"\<A> \<Turnstile>\<^sub>P \<phi> \<longleftrightarrow> (\<exists>B. fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>)"
proof
assume "\<A> \<Turnstile>\<^sub>P \<phi>"
then show "\<exists>B. fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>"
proof (induction \<phi> arbitrary: \<A>)
case (And_ltln \<phi>\<^sub>1 \<phi>\<^sub>2)
then obtain B\<^sub>1 B\<^sub>2 where 1: "fset B\<^sub>1 \<subseteq> \<A> \<and> B\<^sub>1 \<in> min_dnf \<phi>\<^sub>1" and 2: "fset B\<^sub>2 \<subseteq> \<A> \<and> B\<^sub>2 \<in> min_dnf \<phi>\<^sub>2"
by fastforce
then obtain C where "C |\<subseteq>| B\<^sub>1 |\<union>| B\<^sub>2" and "C \<in> min_dnf \<phi>\<^sub>1 \<otimes>\<^sub>m min_dnf \<phi>\<^sub>2"
using min_product_comp by metis
then show ?case
by (metis 1 2 le_sup_iff min_dnf.simps(3) sup.absorb_iff1 sup_fset.rep_eq)
next
case (Or_ltln \<phi>\<^sub>1 \<phi>\<^sub>2)
{
assume "\<A> \<Turnstile>\<^sub>P \<phi>\<^sub>1"
then obtain B where 1: "fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>\<^sub>1"
using Or_ltln by fastforce
then obtain C where "C |\<subseteq>| B" and "C \<in> min_dnf \<phi>\<^sub>1 \<union>\<^sub>m min_dnf \<phi>\<^sub>2"
using min_union_comp by metis
then have ?case
by (metis 1 dual_order.trans less_eq_fset.rep_eq min_dnf.simps(4))
}
moreover
{
assume "\<A> \<Turnstile>\<^sub>P \<phi>\<^sub>2"
then obtain B where 2: "fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>\<^sub>2"
using Or_ltln by fastforce
then obtain C where "C |\<subseteq>| B" and "C \<in> min_dnf \<phi>\<^sub>1 \<union>\<^sub>m min_dnf \<phi>\<^sub>2"
using min_union_comp min_union_comm by metis
then have ?case
by (metis 2 dual_order.trans less_eq_fset.rep_eq min_dnf.simps(4))
}
ultimately show ?case
using Or_ltln.prems by auto
qed simp_all
next
assume "\<exists>B. fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>"
then obtain B where "fset B \<subseteq> \<A>" and "B \<in> min_dnf \<phi>"
by auto
then have "fset B \<Turnstile>\<^sub>P \<phi>"
- by (induction \<phi> arbitrary: B) (auto simp: min_set_def min_product_def product_def min_union_def, blast+)
+ by (induction \<phi> arbitrary: B) (auto simp: min_set_def min_product_def product_def min_union_def)
then show "\<A> \<Turnstile>\<^sub>P \<phi>"
using \<open>fset B \<subseteq> \<A>\<close> by blast
qed
lemma ltl_prop_implies_min_dnf:
"\<phi> \<longrightarrow>\<^sub>P \<psi> = (\<forall>A \<in> min_dnf \<phi>. \<exists>B \<in> min_dnf \<psi>. B |\<subseteq>| A)"
by (meson less_eq_fset.rep_eq ltl_prop_implies_def min_dnf_iff_prop_assignment_subset order_refl dual_order.trans)
lemma ltl_prop_equiv_min_dnf:
"\<phi> \<sim>\<^sub>P \<psi> = (min_dnf \<phi> = min_dnf \<psi>)"
proof
assume "\<phi> \<sim>\<^sub>P \<psi>"
then have "\<And>x. x \<in> min_set (min_dnf \<phi>) \<longleftrightarrow> x \<in> min_set (min_dnf \<psi>)"
unfolding ltl_prop_implies_equiv ltl_prop_implies_min_dnf min_set_iff
by fastforce
then show "min_dnf \<phi> = min_dnf \<psi>"
by auto
qed (simp add: ltl_prop_equiv_def min_dnf_iff_prop_assignment_subset)
lemma min_dnf_rep_abs[simp]:
"min_dnf (rep_ltln\<^sub>P (abs_ltln\<^sub>P \<phi>)) = min_dnf \<phi>"
by (simp add: ltl_prop_equiv_min_dnf[symmetric] Quotient3_ltln\<^sub>P rep_abs_rsp_left)
subsection \<open>Folding of \<open>and\<^sub>n\<close> and \<open>or\<^sub>n\<close> over Finite Sets\<close>
definition And\<^sub>n :: "'a ltln set \<Rightarrow> 'a ltln"
where
"And\<^sub>n \<Phi> \<equiv> SOME \<phi>. fold_graph And_ltln True_ltln \<Phi> \<phi>"
definition Or\<^sub>n :: "'a ltln set \<Rightarrow> 'a ltln"
where
"Or\<^sub>n \<Phi> \<equiv> SOME \<phi>. fold_graph Or_ltln False_ltln \<Phi> \<phi>"
lemma fold_graph_And\<^sub>n:
"finite \<Phi> \<Longrightarrow> fold_graph And_ltln True_ltln \<Phi> (And\<^sub>n \<Phi>)"
unfolding And\<^sub>n_def by (rule someI2_ex[OF finite_imp_fold_graph])
lemma fold_graph_Or\<^sub>n:
"finite \<Phi> \<Longrightarrow> fold_graph Or_ltln False_ltln \<Phi> (Or\<^sub>n \<Phi>)"
unfolding Or\<^sub>n_def by (rule someI2_ex[OF finite_imp_fold_graph])
lemma Or\<^sub>n_empty[simp]:
"Or\<^sub>n {} = False_ltln"
by (metis empty_fold_graphE finite.emptyI fold_graph_Or\<^sub>n)
lemma And\<^sub>n_empty[simp]:
"And\<^sub>n {} = True_ltln"
by (metis empty_fold_graphE finite.emptyI fold_graph_And\<^sub>n)
interpretation dnf_union_thms: Finite_Set.comp_fun_commute "\<lambda>\<phi>. (\<union>) (f \<phi>)"
by unfold_locales fastforce
interpretation dnf_product_thms: Finite_Set.comp_fun_commute "\<lambda>\<phi>. (\<otimes>) (f \<phi>)"
by unfold_locales (simp add: product_set_thms.comp_fun_commute)
\<comment> \<open>Copied from locale @{locale comp_fun_commute}\<close>
lemma fold_graph_finite:
assumes "fold_graph f z A y"
shows "finite A"
using assms by induct simp_all
text \<open>Taking the DNF of @{const And\<^sub>n} and @{const Or\<^sub>n} is the same as folding over the individual DNFs.\<close>
lemma And\<^sub>n_dnf:
"finite \<Phi> \<Longrightarrow> dnf (And\<^sub>n \<Phi>) = Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) (dnf \<phi>)) {{||}} \<Phi>"
proof (drule fold_graph_And\<^sub>n, induction rule: fold_graph.induct)
case (insertI x A y)
then have "finite A"
using fold_graph_finite by fast
then show ?case
using insertI by auto
qed simp
lemma Or\<^sub>n_dnf:
"finite \<Phi> \<Longrightarrow> dnf (Or\<^sub>n \<Phi>) = Finite_Set.fold (\<lambda>\<phi>. (\<union>) (dnf \<phi>)) {} \<Phi>"
proof (drule fold_graph_Or\<^sub>n, induction rule: fold_graph.induct)
case (insertI x A y)
then have "finite A"
using fold_graph_finite by fast
then show ?case
using insertI by auto
qed simp
text \<open>@{const And\<^sub>n} and @{const Or\<^sub>n} are injective on finite sets.\<close>
lemma And\<^sub>n_inj:
"inj_on And\<^sub>n {s. finite s}"
proof (standard, simp)
fix x y :: "'a ltln set"
assume "finite x" and "finite y"
then have 1: "fold_graph And_ltln True_ltln x (And\<^sub>n x)" and 2: "fold_graph And_ltln True_ltln y (And\<^sub>n y)"
using fold_graph_And\<^sub>n by blast+
assume "And\<^sub>n x = And\<^sub>n y"
with 1 show "x = y"
proof (induction rule: fold_graph.induct)
case emptyI
then show ?case
using 2 fold_graph.cases by force
next
case (insertI x A y)
with 2 show ?case
proof (induction arbitrary: x A y rule: fold_graph.induct)
case (insertI x A y)
then show ?case
by (metis fold_graph.cases insertI1 ltln.distinct(7) ltln.inject(3))
qed blast
qed
qed
lemma Or\<^sub>n_inj:
"inj_on Or\<^sub>n {s. finite s}"
proof (standard, simp)
fix x y :: "'a ltln set"
assume "finite x" and "finite y"
then have 1: "fold_graph Or_ltln False_ltln x (Or\<^sub>n x)" and 2: "fold_graph Or_ltln False_ltln y (Or\<^sub>n y)"
using fold_graph_Or\<^sub>n by blast+
assume "Or\<^sub>n x = Or\<^sub>n y"
with 1 show "x = y"
proof (induction rule: fold_graph.induct)
case emptyI
then show ?case
using 2 fold_graph.cases by force
next
case (insertI x A y)
with 2 show ?case
proof (induction arbitrary: x A y rule: fold_graph.induct)
case (insertI x A y)
then show ?case
by (metis fold_graph.cases insertI1 ltln.distinct(27) ltln.inject(4))
qed blast
qed
qed
text \<open>The semantics of @{const And\<^sub>n} and @{const Or\<^sub>n} can be expressed using quantifiers.\<close>
lemma And\<^sub>n_semantics:
"finite \<Phi> \<Longrightarrow> w \<Turnstile>\<^sub>n And\<^sub>n \<Phi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph And_ltln True_ltln \<Phi> \<psi> \<Longrightarrow> w \<Turnstile>\<^sub>n \<psi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_And\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma Or\<^sub>n_semantics:
"finite \<Phi> \<Longrightarrow> w \<Turnstile>\<^sub>n Or\<^sub>n \<Phi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph Or_ltln False_ltln \<Phi> \<psi> \<Longrightarrow> w \<Turnstile>\<^sub>n \<psi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_Or\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma And\<^sub>n_prop_semantics:
"finite \<Phi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P And\<^sub>n \<Phi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph And_ltln True_ltln \<Phi> \<psi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<psi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_And\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma Or\<^sub>n_prop_semantics:
"finite \<Phi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P Or\<^sub>n \<Phi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph Or_ltln False_ltln \<Phi> \<psi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<psi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_Or\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma Or\<^sub>n_And\<^sub>n_image_semantics:
assumes "finite \<A>" and "\<And>\<Phi>. \<Phi> \<in> \<A> \<Longrightarrow> finite \<Phi>"
shows "w \<Turnstile>\<^sub>n Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
proof -
have "w \<Turnstile>\<^sub>n Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. w \<Turnstile>\<^sub>n And\<^sub>n \<Phi>)"
using Or\<^sub>n_semantics assms by auto
then show ?thesis
using And\<^sub>n_semantics assms by fast
qed
lemma Or\<^sub>n_And\<^sub>n_image_prop_semantics:
assumes "finite \<A>" and "\<And>\<Phi>. \<Phi> \<in> \<A> \<Longrightarrow> finite \<Phi>"
shows "\<I> \<Turnstile>\<^sub>P Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi> \<in> \<Phi>. \<I> \<Turnstile>\<^sub>P \<phi>)"
proof -
have "\<I> \<Turnstile>\<^sub>P Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<I> \<Turnstile>\<^sub>P And\<^sub>n \<Phi>)"
using Or\<^sub>n_prop_semantics assms by blast
then show ?thesis
using And\<^sub>n_prop_semantics assms by metis
qed
subsection \<open>DNF to LTL conversion\<close>
definition ltln_of_dnf :: "'a ltln fset set \<Rightarrow> 'a ltln"
where
"ltln_of_dnf \<A> = Or\<^sub>n (And\<^sub>n ` fset ` \<A>)"
lemma ltln_of_dnf_semantics:
assumes "finite \<A>"
shows "w \<Turnstile>\<^sub>n ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi>. \<phi> |\<in>| \<Phi> \<longrightarrow> w \<Turnstile>\<^sub>n \<phi>)"
proof -
have "finite (fset ` \<A>)"
using assms by blast
then have "w \<Turnstile>\<^sub>n ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> fset ` \<A>. \<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
unfolding ltln_of_dnf_def using Or\<^sub>n_And\<^sub>n_image_semantics by fastforce
then show ?thesis
- by (metis image_iff notin_fset)
+ by (metis image_iff)
qed
lemma ltln_of_dnf_prop_semantics:
assumes "finite \<A>"
shows "\<I> \<Turnstile>\<^sub>P ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi>. \<phi> |\<in>| \<Phi> \<longrightarrow> \<I> \<Turnstile>\<^sub>P \<phi>)"
proof -
have "finite (fset ` \<A>)"
using assms by blast
then have "\<I> \<Turnstile>\<^sub>P ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> fset ` \<A>. \<forall>\<phi> \<in> \<Phi>. \<I> \<Turnstile>\<^sub>P \<phi>)"
unfolding ltln_of_dnf_def using Or\<^sub>n_And\<^sub>n_image_prop_semantics by fastforce
then show ?thesis
- by (metis image_iff notin_fset)
+ by (metis image_iff)
qed
lemma ltln_of_dnf_prop_equiv:
"ltln_of_dnf (min_dnf \<phi>) \<sim>\<^sub>P \<phi>"
unfolding ltl_prop_equiv_def
proof
fix \<A>
have "\<A> \<Turnstile>\<^sub>P ltln_of_dnf (min_dnf \<phi>) \<longleftrightarrow> (\<exists>\<Phi> \<in> min_dnf \<phi>. \<forall>\<phi>. \<phi> |\<in>| \<Phi> \<longrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>)"
using ltln_of_dnf_prop_semantics min_dnf_finite by metis
also have "\<dots> \<longleftrightarrow> (\<exists>\<Phi> \<in> min_dnf \<phi>. fset \<Phi> \<subseteq> \<A>)"
- by (metis min_dnf_prop_atoms prop_atoms_entailment_iff notin_fset subset_eq)
+ by (metis min_dnf_prop_atoms prop_atoms_entailment_iff subset_eq)
also have "\<dots> \<longleftrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>"
using min_dnf_iff_prop_assignment_subset by blast
finally show "\<A> \<Turnstile>\<^sub>P ltln_of_dnf (min_dnf \<phi>) = \<A> \<Turnstile>\<^sub>P \<phi>" .
qed
lemma min_dnf_ltln_of_dnf[simp]:
"min_dnf (ltln_of_dnf (min_dnf \<phi>)) = min_dnf \<phi>"
using ltl_prop_equiv_min_dnf ltln_of_dnf_prop_equiv by blast
subsection \<open>Substitution in DNF formulas\<close>
definition subst_clause :: "'a ltln fset \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln fset set"
where
"subst_clause \<Phi> m = \<Otimes>\<^sub>m {min_dnf (subst \<phi> m) | \<phi>. \<phi> \<in> fset \<Phi>}"
definition subst_dnf :: "'a ltln fset set \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln fset set"
where
"subst_dnf \<A> m = (\<Union>\<Phi> \<in> \<A>. subst_clause \<Phi> m)"
lemma subst_clause_empty[simp]:
"subst_clause {||} m = {{||}}"
by (simp add: subst_clause_def)
lemma subst_dnf_empty[simp]:
"subst_dnf {} m = {}"
by (simp add: subst_dnf_def)
lemma subst_clause_inner_finite:
"finite {min_dnf (subst \<phi> m) | \<phi>. \<phi> \<in> \<Phi>}" if "finite \<Phi>"
using that by simp
lemma subst_clause_finite:
"finite (subst_clause \<Phi> m)"
unfolding subst_clause_def
by (auto intro: min_dnf_finite min_product_set_finite)
lemma subst_dnf_finite:
"finite \<A> \<Longrightarrow> finite (subst_dnf \<A> m)"
unfolding subst_dnf_def using subst_clause_finite by blast
lemma subst_dnf_mono:
"\<A> \<subseteq> \<B> \<Longrightarrow> subst_dnf \<A> m \<subseteq> subst_dnf \<B> m"
unfolding subst_dnf_def by blast
lemma subst_clause_min_set[simp]:
"min_set (subst_clause \<Phi> m) = subst_clause \<Phi> m"
unfolding subst_clause_def by simp
lemma subst_clause_finsert[simp]:
"subst_clause (finsert \<phi> \<Phi>) m = (min_dnf (subst \<phi> m)) \<otimes>\<^sub>m (subst_clause \<Phi> m)"
proof -
have "{min_dnf (subst \<psi> m) | \<psi>. \<psi> \<in> fset (finsert \<phi> \<Phi>)}
= insert (min_dnf (subst \<phi> m)) {min_dnf (subst \<psi> m) | \<psi>. \<psi> \<in> fset \<Phi>}"
by auto
then show ?thesis
by (simp add: subst_clause_def)
qed
lemma subst_clause_funion[simp]:
"subst_clause (\<Phi> |\<union>| \<Psi>) m = (subst_clause \<Phi> m) \<otimes>\<^sub>m (subst_clause \<Psi> m)"
proof (induction \<Psi>)
case (insert x F)
then show ?case
using min_product_set_thms.fun_left_comm by fastforce
qed simp
text \<open>For the proof of correctness, we redefine the @{const product} operator on lists.\<close>
definition list_product :: "'a list set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" (infixl "\<otimes>\<^sub>l" 65)
where
"A \<otimes>\<^sub>l B = {a @ b | a b. a \<in> A \<and> b \<in> B}"
lemma list_product_fset_of_list[simp]:
"fset_of_list ` (A \<otimes>\<^sub>l B) = (fset_of_list ` A) \<otimes> (fset_of_list ` B)"
unfolding list_product_def product_def image_def by fastforce
lemma list_product_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<otimes>\<^sub>l B)"
unfolding list_product_def by (simp add: finite_image_set2)
lemma list_product_iff:
"x \<in> A \<otimes>\<^sub>l B \<longleftrightarrow> (\<exists>a b. a \<in> A \<and> b \<in> B \<and> x = a @ b)"
unfolding list_product_def by blast
lemma list_product_assoc[simp]:
"A \<otimes>\<^sub>l (B \<otimes>\<^sub>l C) = A \<otimes>\<^sub>l B \<otimes>\<^sub>l C"
unfolding set_eq_iff list_product_iff by fastforce
text \<open>Furthermore, we introduct DNFs where the clauses are represented as lists.\<close>
fun list_dnf :: "'a ltln \<Rightarrow> 'a ltln list set"
where
"list_dnf true\<^sub>n = {[]}"
| "list_dnf false\<^sub>n = {}"
| "list_dnf (\<phi> and\<^sub>n \<psi>) = (list_dnf \<phi>) \<otimes>\<^sub>l (list_dnf \<psi>)"
| "list_dnf (\<phi> or\<^sub>n \<psi>) = (list_dnf \<phi>) \<union> (list_dnf \<psi>)"
| "list_dnf \<phi> = {[\<phi>]}"
definition list_dnf_to_dnf :: "'a list set \<Rightarrow> 'a fset set"
where
"list_dnf_to_dnf X = fset_of_list ` X"
lemma list_dnf_to_dnf_list_dnf[simp]:
"list_dnf_to_dnf (list_dnf \<phi>) = dnf \<phi>"
by (induction \<phi>) (simp_all add: list_dnf_to_dnf_def image_Un)
lemma list_dnf_finite:
"finite (list_dnf \<phi>)"
by (induction \<phi>) (simp_all add: list_product_finite)
text \<open>We use this to redefine @{const subst_clause} and @{const subst_dnf} on list DNFs.\<close>
definition subst_clause' :: "'a ltln list \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln list set"
where
"subst_clause' \<Phi> m = fold (\<lambda>\<phi> acc. acc \<otimes>\<^sub>l list_dnf (subst \<phi> m)) \<Phi> {[]}"
definition subst_dnf' :: "'a ltln list set \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln list set"
where
"subst_dnf' \<A> m = (\<Union>\<Phi> \<in> \<A>. subst_clause' \<Phi> m)"
lemma subst_clause'_finite:
"finite (subst_clause' \<Phi> m)"
by (induction \<Phi> rule: rev_induct) (simp_all add: subst_clause'_def list_dnf_finite list_product_finite)
lemma subst_clause'_nil[simp]:
"subst_clause' [] m = {[]}"
by (simp add: subst_clause'_def)
lemma subst_clause'_cons[simp]:
"subst_clause' (xs @ [x]) m = subst_clause' xs m \<otimes>\<^sub>l list_dnf (subst x m)"
by (simp add: subst_clause'_def)
lemma subst_clause'_append[simp]:
"subst_clause' (A @ B) m = subst_clause' A m \<otimes>\<^sub>l subst_clause' B m"
proof (induction B rule: rev_induct)
case (snoc x xs)
then show ?case
by simp (metis append_assoc subst_clause'_cons)
qed(simp add: list_product_def)
lemma subst_dnf'_iff:
"x \<in> subst_dnf' A m \<longleftrightarrow> (\<exists>\<Phi> \<in> A. x \<in> subst_clause' \<Phi> m)"
by (simp add: subst_dnf'_def)
lemma subst_dnf'_product:
"subst_dnf' (A \<otimes>\<^sub>l B) m = (subst_dnf' A m) \<otimes>\<^sub>l (subst_dnf' B m)" (is "?lhs = ?rhs")
proof (unfold set_eq_iff, safe)
fix x
assume "x \<in> ?lhs"
then obtain \<Phi> where "\<Phi> \<in> A \<otimes>\<^sub>l B" and "x \<in> subst_clause' \<Phi> m"
unfolding subst_dnf'_iff by blast
then obtain a b where "a \<in> A" and "b \<in> B" and "\<Phi> = a @ b"
unfolding list_product_def by blast
then have "x \<in> (subst_clause' a m) \<otimes>\<^sub>l (subst_clause' b m)"
using \<open>x \<in> subst_clause' \<Phi> m\<close> by simp
then obtain a' b' where "a' \<in> subst_clause' a m" and "b' \<in> subst_clause' b m" and "x = a' @ b'"
unfolding list_product_iff by blast
then have "a' \<in> subst_dnf' A m" and "b' \<in> subst_dnf' B m"
unfolding subst_dnf'_iff using \<open>a \<in> A\<close> \<open>b \<in> B\<close> by auto
then have "\<exists>a\<in>subst_dnf' A m. \<exists>b\<in>subst_dnf' B m. x = a @ b"
using \<open>x = a' @ b'\<close> by blast
then show "x \<in> ?rhs"
unfolding list_product_iff by blast
next
fix x
assume "x \<in> ?rhs"
then obtain a b where "a \<in> subst_dnf' A m" and "b \<in> subst_dnf' B m" and "x = a @ b"
unfolding list_product_iff by blast
then obtain a' b' where "a' \<in> A" and "b' \<in> B" and a: "a \<in> subst_clause' a' m" and b: "b \<in> subst_clause' b' m"
unfolding subst_dnf'_iff by blast
then have "x \<in> (subst_clause' a' m) \<otimes>\<^sub>l (subst_clause' b' m)"
unfolding list_product_iff using \<open>x = a @ b\<close> by blast
moreover
have "a' @ b' \<in> A \<otimes>\<^sub>l B"
unfolding list_product_iff using \<open>a' \<in> A\<close> \<open>b' \<in> B\<close> by blast
ultimately show "x \<in> ?lhs"
unfolding subst_dnf'_iff by force
qed
lemma subst_dnf'_list_dnf:
"subst_dnf' (list_dnf \<phi>) m = list_dnf (subst \<phi> m)"
proof (induction \<phi>)
case (And_ltln \<phi>1 \<phi>2)
then show ?case
by (simp add: subst_dnf'_product)
qed (simp_all add: subst_dnf'_def subst_clause'_def list_product_def)
lemma min_set_Union:
"finite X \<Longrightarrow> min_set (\<Union> (min_set ` X)) = min_set (\<Union> X)" for X :: "'a fset set set"
by (induction X rule: finite_induct) (force, metis Sup_insert image_insert min_set_min_union min_union_def)
lemma min_set_Union_image:
"finite X \<Longrightarrow> min_set (\<Union>x \<in> X. min_set (f x)) = min_set (\<Union>x \<in> X. f x)" for f :: "'b \<Rightarrow> 'a fset set"
proof -
assume "finite X"
then have *: "finite (f ` X)" by auto
with min_set_Union show ?thesis
unfolding image_image by fastforce
qed
lemma subst_clause_fset_of_list:
"subst_clause (fset_of_list \<Phi>) m = min_set (list_dnf_to_dnf (subst_clause' \<Phi> m))"
unfolding list_dnf_to_dnf_def subst_clause'_def
proof (induction \<Phi> rule: rev_induct)
case (snoc x xs)
then show ?case
by simp (metis (no_types, lifting) dnf_min_set list_dnf_to_dnf_def list_dnf_to_dnf_list_dnf min_product_comm min_product_def min_set_min_product(1))
qed simp
lemma min_set_list_dnf_to_dnf_subst_dnf':
"finite X \<Longrightarrow> min_set (list_dnf_to_dnf (subst_dnf' X m)) = min_set (subst_dnf (list_dnf_to_dnf X) m)"
by (simp add: subst_dnf'_def subst_dnf_def subst_clause_fset_of_list list_dnf_to_dnf_def min_set_Union_image image_Union)
lemma subst_dnf_dnf:
"min_set (subst_dnf (dnf \<phi>) m) = min_dnf (subst \<phi> m)"
unfolding dnf_min_set
unfolding list_dnf_to_dnf_list_dnf[symmetric]
unfolding subst_dnf'_list_dnf[symmetric]
unfolding min_set_list_dnf_to_dnf_subst_dnf'[OF list_dnf_finite]
by simp
text \<open>This is almost the lemma we need. However, we need to show that the same holds for @{term "min_dnf \<phi>"}, too.\<close>
lemma fold_product:
"Finite_Set.fold (\<lambda>x. (\<otimes>) {{|x|}}) {{||}} (fset x) = {x}"
- by (induction x) (simp_all add: notin_fset, simp add: product_singleton_singleton)
+ by (induction x) (simp_all, simp add: product_singleton_singleton)
lemma fold_union:
"Finite_Set.fold (\<lambda>x. (\<union>) {x}) {} (fset x) = fset x"
- by (induction x) (simp_all add: notin_fset comp_fun_idem_on.fold_insert_idem[OF comp_fun_idem_insert[unfolded comp_fun_idem_def']])
+ by (induction x) (simp_all add: comp_fun_idem_on.fold_insert_idem[OF comp_fun_idem_insert[unfolded comp_fun_idem_def']])
lemma fold_union_fold_product:
assumes "finite X" and "\<And>\<Psi> \<psi>. \<Psi> \<in> X \<Longrightarrow> \<psi> \<in> fset \<Psi> \<Longrightarrow> dnf \<psi> = {{|\<psi>|}}"
shows "Finite_Set.fold (\<lambda>x. (\<union>) (Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) (dnf \<phi>)) {{||}} (fset x))) {} X = X" (is "?lhs = X")
proof -
from assms have "?lhs = Finite_Set.fold (\<lambda>x. (\<union>) (Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) {{|\<phi>|}}) {{||}} (fset x))) {} X"
proof (induction X rule: finite_induct)
case (insert \<Phi> X)
from insert.prems have 1: "\<And>\<Psi> \<psi>. \<lbrakk>\<Psi> \<in> X; \<psi> \<in> fset \<Psi>\<rbrakk> \<Longrightarrow> dnf \<psi> = {{|\<psi>|}}"
by force
from insert.prems have "Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) (dnf \<phi>)) {{||}} (fset \<Phi>) = Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) {{|\<phi>|}}) {{||}} (fset \<Phi>)"
- by (induction \<Phi>) (force simp: notin_fset)+
+ by (induction \<Phi>) force+
with insert 1 show ?case
by simp
qed simp
with \<open>finite X\<close> show ?thesis
unfolding fold_product by (metis fset_to_fset fold_union)
qed
lemma dnf_ltln_of_dnf_min_dnf:
"dnf (ltln_of_dnf (min_dnf \<phi>)) = min_dnf \<phi>"
proof -
have 1: "finite (And\<^sub>n ` fset ` min_dnf \<phi>)"
using min_dnf_finite by blast
have 2: "inj_on And\<^sub>n (fset ` min_dnf \<phi>)"
by (metis (mono_tags, lifting) And\<^sub>n_inj f_inv_into_f fset inj_onI inj_on_contraD)
have 3: "inj_on fset (min_dnf \<phi>)"
by (meson fset_inject inj_onI)
show ?thesis
unfolding ltln_of_dnf_def
unfolding Or\<^sub>n_dnf[OF 1]
unfolding fold_image[OF 2]
unfolding fold_image[OF 3]
unfolding comp_def
unfolding And\<^sub>n_dnf[OF finite_fset]
by (metis fold_union_fold_product min_dnf_finite min_dnf_atoms_dnf)
qed
lemma min_dnf_subst:
"min_set (subst_dnf (min_dnf \<phi>) m) = min_dnf (subst \<phi> m)" (is "?lhs = ?rhs")
proof -
let ?\<phi>' = "ltln_of_dnf (min_dnf \<phi>)"
have "?lhs = min_set (subst_dnf (dnf ?\<phi>') m)"
unfolding dnf_ltln_of_dnf_min_dnf ..
also have "\<dots> = min_dnf (subst ?\<phi>' m)"
unfolding subst_dnf_dnf ..
also have "\<dots> = min_dnf (subst \<phi> m)"
using ltl_prop_equiv_min_dnf ltln_of_dnf_prop_equiv subst_respects_ltl_prop_entailment(2) by blast
finally show ?thesis .
qed
end
diff --git a/thys/LTL_Normal_Form/Normal_Form_Code_Export.thy b/thys/LTL_Normal_Form/Normal_Form_Code_Export.thy
--- a/thys/LTL_Normal_Form/Normal_Form_Code_Export.thy
+++ b/thys/LTL_Normal_Form/Normal_Form_Code_Export.thy
@@ -1,134 +1,134 @@
(*
Author: Salomon Sickert
License: BSD
*)
section \<open>Code Export\<close>
theory Normal_Form_Code_Export imports
LTL.Code_Equations
LTL.Rewriting
LTL.Disjunctive_Normal_Form
HOL.String
Normal_Form
begin
fun flatten_pi_1_list :: "String.literal ltln \<Rightarrow> String.literal ltln list \<Rightarrow> String.literal ltln"
where
"flatten_pi_1_list (\<psi>\<^sub>1 U\<^sub>n \<psi>\<^sub>2) M = (if (\<psi>\<^sub>1 U\<^sub>n \<psi>\<^sub>2) \<in> set M then (flatten_pi_1_list \<psi>\<^sub>1 M) W\<^sub>n (flatten_pi_1_list \<psi>\<^sub>2 M) else false\<^sub>n)"
| "flatten_pi_1_list (\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2) M = (flatten_pi_1_list \<psi>\<^sub>1 M) W\<^sub>n (flatten_pi_1_list \<psi>\<^sub>2 M)"
| "flatten_pi_1_list (\<psi>\<^sub>1 M\<^sub>n \<psi>\<^sub>2) M = (if (\<psi>\<^sub>1 M\<^sub>n \<psi>\<^sub>2) \<in> set M then (flatten_pi_1_list \<psi>\<^sub>1 M) R\<^sub>n (flatten_pi_1_list \<psi>\<^sub>2 M) else false\<^sub>n)"
| "flatten_pi_1_list (\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2) M = (flatten_pi_1_list \<psi>\<^sub>1 M) R\<^sub>n (flatten_pi_1_list \<psi>\<^sub>2 M)"
| "flatten_pi_1_list (\<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2) M = (flatten_pi_1_list \<psi>\<^sub>1 M) and\<^sub>n (flatten_pi_1_list \<psi>\<^sub>2 M)"
| "flatten_pi_1_list (\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2) M = (flatten_pi_1_list \<psi>\<^sub>1 M) or\<^sub>n (flatten_pi_1_list \<psi>\<^sub>2 M)"
| "flatten_pi_1_list (X\<^sub>n \<psi>) M = X\<^sub>n (flatten_pi_1_list \<psi> M)"
| "flatten_pi_1_list \<phi> _ = \<phi>"
fun flatten_sigma_1_list :: "String.literal ltln \<Rightarrow> String.literal ltln list \<Rightarrow> String.literal ltln"
where
"flatten_sigma_1_list (\<psi>\<^sub>1 U\<^sub>n \<psi>\<^sub>2) N = (flatten_sigma_1_list \<psi>\<^sub>1 N) U\<^sub>n (flatten_sigma_1_list \<psi>\<^sub>2 N)"
| "flatten_sigma_1_list (\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2) N = (if (\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2) \<in> set N then true\<^sub>n else (flatten_sigma_1_list \<psi>\<^sub>1 N) U\<^sub>n (flatten_sigma_1_list \<psi>\<^sub>2 N))"
| "flatten_sigma_1_list (\<psi>\<^sub>1 M\<^sub>n \<psi>\<^sub>2) N = (flatten_sigma_1_list \<psi>\<^sub>1 N) M\<^sub>n (flatten_sigma_1_list \<psi>\<^sub>2 N)"
| "flatten_sigma_1_list (\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2) N = (if (\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2) \<in> set N then true\<^sub>n else (flatten_sigma_1_list \<psi>\<^sub>1 N) M\<^sub>n (flatten_sigma_1_list \<psi>\<^sub>2 N))"
| "flatten_sigma_1_list (\<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2) N = (flatten_sigma_1_list \<psi>\<^sub>1 N) and\<^sub>n (flatten_sigma_1_list \<psi>\<^sub>2 N)"
| "flatten_sigma_1_list (\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2) N = (flatten_sigma_1_list \<psi>\<^sub>1 N) or\<^sub>n (flatten_sigma_1_list \<psi>\<^sub>2 N)"
| "flatten_sigma_1_list (X\<^sub>n \<psi>) N = X\<^sub>n (flatten_sigma_1_list \<psi> N)"
| "flatten_sigma_1_list \<phi> _ = \<phi>"
fun flatten_sigma_2_list :: "String.literal ltln \<Rightarrow> String.literal ltln list \<Rightarrow> String.literal ltln"
where
"flatten_sigma_2_list (\<phi> U\<^sub>n \<psi>) M = (flatten_sigma_2_list \<phi> M) U\<^sub>n (flatten_sigma_2_list \<psi> M)"
| "flatten_sigma_2_list (\<phi> W\<^sub>n \<psi>) M = (flatten_sigma_2_list \<phi> M) U\<^sub>n ((flatten_sigma_2_list \<psi> M) or\<^sub>n (G\<^sub>n (flatten_pi_1_list \<phi> M)))"
| "flatten_sigma_2_list (\<phi> M\<^sub>n \<psi>) M = (flatten_sigma_2_list \<phi> M) M\<^sub>n (flatten_sigma_2_list \<psi> M)"
| "flatten_sigma_2_list (\<phi> R\<^sub>n \<psi>) M = ((flatten_sigma_2_list \<phi> M) or\<^sub>n (G\<^sub>n (flatten_pi_1_list \<psi> M))) M\<^sub>n (flatten_sigma_2_list \<psi> M)"
| "flatten_sigma_2_list (\<phi> and\<^sub>n \<psi>) M = (flatten_sigma_2_list \<phi> M) and\<^sub>n (flatten_sigma_2_list \<psi> M)"
| "flatten_sigma_2_list (\<phi> or\<^sub>n \<psi>) M = (flatten_sigma_2_list \<phi> M) or\<^sub>n (flatten_sigma_2_list \<psi> M)"
| "flatten_sigma_2_list (X\<^sub>n \<phi>) M = X\<^sub>n (flatten_sigma_2_list \<phi> M)"
| "flatten_sigma_2_list \<phi> _ = \<phi>"
lemma flatten_code_equations[simp]:
"\<phi>[set M]\<^sub>\<Pi>\<^sub>1 = flatten_pi_1_list \<phi> M"
"\<phi>[set M]\<^sub>\<Sigma>\<^sub>1 = flatten_sigma_1_list \<phi> M"
"\<phi>[set M]\<^sub>\<Sigma>\<^sub>2 = flatten_sigma_2_list \<phi> M"
by (induction \<phi>) auto
abbreviation "and_list \<equiv> foldl And_ltln true\<^sub>n"
abbreviation "or_list \<equiv> foldl Or_ltln false\<^sub>n"
definition "normal_form_disjunct (\<phi> :: String.literal ltln) M N
\<equiv> (flatten_sigma_2_list \<phi> M)
and\<^sub>n (and_list (map (\<lambda>\<psi>. G\<^sub>n (F\<^sub>n (flatten_sigma_1_list \<psi> N))) M)
and\<^sub>n (and_list (map (\<lambda>\<psi>. F\<^sub>n (G\<^sub>n (flatten_pi_1_list \<psi> M))) N)))"
definition "normal_form (\<phi> :: String.literal ltln)
\<equiv> or_list (map (\<lambda>(M, N). normal_form_disjunct \<phi> M N) (advice_sets \<phi>))"
lemma and_list_semantic: "w \<Turnstile>\<^sub>n and_list xs \<longleftrightarrow> (\<forall>x \<in> set xs. w \<Turnstile>\<^sub>n x)"
by (induction xs rule: rev_induct) auto
lemma or_list_semantic: "w \<Turnstile>\<^sub>n or_list xs \<longleftrightarrow> (\<exists>x \<in> set xs. w \<Turnstile>\<^sub>n x)"
by (induction xs rule: rev_induct) auto
theorem normal_form_correct:
"w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n normal_form \<phi>"
proof
assume "w \<Turnstile>\<^sub>n \<phi>"
then obtain M N where "M \<subseteq> subformulas\<^sub>\<mu> \<phi>" and "N \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and c1: "w \<Turnstile>\<^sub>n \<phi>[M]\<^sub>\<Sigma>\<^sub>2" and c2: "\<forall>\<psi> \<in> M. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[N]\<^sub>\<Sigma>\<^sub>1)" and c3: "\<forall>\<psi> \<in> N. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[M]\<^sub>\<Pi>\<^sub>1)"
using normal_form_with_flatten_sigma_2 by metis
then obtain ms ns where "M = set ms" and "N = set ns" and ms_ns_in: "(ms, ns) \<in> set (advice_sets \<phi>)"
by (meson advice_sets_subformulas)
then have "w \<Turnstile>\<^sub>n normal_form_disjunct \<phi> ms ns"
using c1 c2 c3 by (simp add: and_list_semantic normal_form_disjunct_def)
then show "w \<Turnstile>\<^sub>n normal_form \<phi>"
using normal_form_def or_list_semantic ms_ns_in by fastforce
next
assume "w \<Turnstile>\<^sub>n normal_form \<phi>"
then obtain ms ns where "(ms, ns) \<in> set (advice_sets \<phi>)"
and "w \<Turnstile>\<^sub>n normal_form_disjunct \<phi> ms ns"
unfolding normal_form_def or_list_semantic by force
then have "set ms \<subseteq> subformulas\<^sub>\<mu> \<phi>" and "set ns \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and c1: "w \<Turnstile>\<^sub>n \<phi>[set ms]\<^sub>\<Sigma>\<^sub>2" and c2: "\<forall>\<psi> \<in> set ms. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[set ns]\<^sub>\<Sigma>\<^sub>1)" and c3: "\<forall>\<psi> \<in> set ns. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[set ms]\<^sub>\<Pi>\<^sub>1)"
using advice_sets_element_subfrmlsn
by (auto simp: and_list_semantic normal_form_disjunct_def) blast
then show "w \<Turnstile>\<^sub>n \<phi>"
using normal_form_with_flatten_sigma_2 by metis
qed
definition "normal_form_with_simplifier (\<phi> :: String.literal ltln)
\<equiv> min_dnf (simplify Slow (normal_form (simplify Slow \<phi>)))"
lemma ltl_semantics_min_dnf:
"w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> (\<exists>C \<in> min_dnf \<phi>. \<forall>\<psi>. \<psi> |\<in>| C \<longrightarrow> w \<Turnstile>\<^sub>n \<psi>)" (is "?lhs \<longleftrightarrow> ?rhs")
proof
let ?M = "{\<psi>. w \<Turnstile>\<^sub>n \<psi>}"
assume ?lhs
hence "?M \<Turnstile>\<^sub>P \<phi>"
using ltl_models_equiv_prop_entailment by blast
then obtain M' where "fset M' \<subseteq> ?M" and "M' \<in> min_dnf \<phi>"
using min_dnf_iff_prop_assignment_subset by blast
thus ?rhs
- by (meson in_mono mem_Collect_eq notin_fset)
+ by (meson in_mono mem_Collect_eq)
next
let ?M = "{\<psi>. w \<Turnstile>\<^sub>n \<psi>}"
assume ?rhs
then obtain M' where "fset M' \<subseteq> ?M" and "M' \<in> min_dnf \<phi>"
- using notin_fset by fastforce
+ by auto
hence "?M \<Turnstile>\<^sub>P \<phi>"
using min_dnf_iff_prop_assignment_subset by blast
thus ?lhs
using ltl_models_equiv_prop_entailment by blast
qed
theorem
"w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> (\<exists>C \<in> (normal_form_with_simplifier \<phi>). \<forall>\<psi>. \<psi> |\<in>| C \<longrightarrow> w \<Turnstile>\<^sub>n \<psi>)" (is "?lhs \<longleftrightarrow> ?rhs")
unfolding normal_form_with_simplifier_def ltl_semantics_min_dnf[symmetric]
using normal_form_correct by simp
text \<open>In order to export the code run \texttt{isabelle build -D [PATH] -e}.\<close>
export_code normal_form in SML
export_code normal_form_with_simplifier in SML
end
\ No newline at end of file
diff --git a/thys/Linear_Programming/Matrix_LinPoly.thy b/thys/Linear_Programming/Matrix_LinPoly.thy
--- a/thys/Linear_Programming/Matrix_LinPoly.thy
+++ b/thys/Linear_Programming/Matrix_LinPoly.thy
@@ -1,506 +1,506 @@
theory Matrix_LinPoly
imports
Jordan_Normal_Form.Matrix_Impl
Farkas.Simplex_for_Reals
Farkas.Matrix_Farkas
begin
text \<open> Add this to linear polynomials in Simplex \<close>
lemma eval_poly_with_sum: "(v \<lbrace> X \<rbrace>) = (\<Sum>x\<in> vars v. coeff v x * X x)"
- using linear_poly_sum sum.cong by fastforce
+ by (auto simp: linear_poly_sum intro: sum.cong)
lemma eval_poly_with_sum_superset:
assumes "finite S"
assumes "S \<supseteq> vars v"
shows "(v \<lbrace>X\<rbrace>) = (\<Sum>x\<in>S. coeff v x * X x)"
proof -
define D where D: "D = S - vars v"
have zeros: "\<forall>x \<in> D. coeff v x = 0"
using D coeff_zero by auto
have fnt: "finite (vars v)"
using finite_vars by auto
have "(v \<lbrace>X\<rbrace>) = (\<Sum>x\<in> vars v. coeff v x * X x)"
- using linear_poly_sum sum.cong by fastforce
+ by (auto simp add: linear_poly_sum intro: sum.cong)
also have "... = (\<Sum>x\<in> vars v. coeff v x * X x) + (\<Sum>x\<in>D. coeff v x * X x)"
using zeros by auto
also have "... = (\<Sum>x\<in> vars v \<union> D. coeff v x * X x)"
using assms(1) fnt Diff_partition[of "vars v" S, OF assms(2)]
sum.subset_diff[of "vars v" S, OF assms(2) assms(1)]
by (simp add: \<open>\<And>g. sum g S = sum g (S - vars v) + sum g (vars v)\<close> D)
also have "... = (\<Sum>x\<in>S. coeff v x * X x)"
using D Diff_partition assms(2) by fastforce
finally show ?thesis .
qed
text \<open> Get rid of these synonyms \<close>
section \<open> Translations of Jordan Normal Forms Matrix Library to Simplex polynomials \<close>
subsection \<open> Vectors \<close>
(* Translate rat list to linear polynomial with same coefficients *)
definition list_to_lpoly where
"list_to_lpoly cs = sum_list (map2 (\<lambda> i c. lp_monom c i) [0..<length cs] cs)"
lemma empty_list_0poly:
shows "list_to_lpoly [] = 0"
unfolding list_to_lpoly_def by simp
lemma sum_list_map_upto_coeff_limit:
assumes "i \<ge> length L"
shows "coeff (list_to_lpoly L) i = 0"
using assms by (induction L rule: rev_induct) (auto simp: list_to_lpoly_def)
lemma rl_lpoly_coeff_nth_non_empty:
assumes "i < length cs"
assumes "cs \<noteq> []"
shows "coeff (list_to_lpoly cs) i = cs!i"
using assms(2) assms(1)
proof (induction cs rule: rev_nonempty_induct)
fix x ::rat
assume "i < length [x]"
have "(list_to_lpoly [x]) = lp_monom x 0"
by (simp add: list_to_lpoly_def)
then show "coeff (list_to_lpoly [x]) i = [x] ! i"
using \<open>i < length [x]\<close> list_to_lpoly_def by auto
next
fix x :: rat
fix xs :: "rat list"
assume "xs \<noteq> []"
assume IH: "i < length xs \<Longrightarrow> coeff (list_to_lpoly xs) i = xs ! i"
assume "i < length (xs @ [x])"
consider (le) "i < length xs" | (eq) "i = length xs"
using \<open>i < length (xs @ [x])\<close> less_Suc_eq by auto
then show "coeff (list_to_lpoly (xs @ [x])) i = (xs @ [x]) ! i"
proof (cases)
case le
have "coeff (lp_monom x (length xs)) i = 0"
using le by auto
have "coeff (sum_list (map2 (\<lambda>x y. lp_monom y x)
[0..<length (xs @ [x])] (xs @ [x]))) i = (xs @ [x]) ! i"
apply(simp add: IH le nth_append)
using IH le list_to_lpoly_def by auto
then show ?thesis
unfolding list_to_lpoly_def by simp
next
case eq
then have *: "coeff (sum_list (map2 (\<lambda>x y. lp_monom y x) [0..<length xs] xs)) i = 0"
using sum_list_map_upto_coeff_limit[of xs i]
by (simp add: list_to_lpoly_def)
have **: "(sum_list (map2 (\<lambda> x y. lp_monom y x) [0..<length (xs @ [x])] (xs @ [x]))) =
sum_list (map (\<lambda>(x,y). lp_monom y x) (zip [0..<length xs] xs)) + lp_monom x (length xs)"
by simp
have "coeff ((list_to_lpoly xs) + lp_monom x (length xs)) i = x"
unfolding list_to_lpoly_def using * ** by (simp add: eq)
then show ?thesis
by (simp add: eq list_to_lpoly_def)
qed
qed
lemma list_to_lpoly_coeff_nth:
assumes "i < length cs "
shows "coeff (list_to_lpoly cs) i = cs ! i"
using gr_implies_not0 rl_lpoly_coeff_nth_non_empty assms by fastforce
lemma rat_list_outside_zero:
assumes "length cs \<le> i"
shows "coeff (list_to_lpoly cs) i = 0"
using sum_list_map_upto_coeff_limit[of cs i, OF assms] by simp
text \<open> Transform linear polynomials to rational vectors \<close>
fun dim_poly where
"dim_poly p = (if (vars p) = {} then 0 else Max (vars p)+1)"
(* 0, 0, 0, 3, 0, 0, \<dots> has dimension 4 , consistent with dim vec *)
definition max_dim_poly_list where
"max_dim_poly_list lst = Max {Max (vars p) |p. p \<in> set lst}"
fun lpoly_to_vec where
"lpoly_to_vec p = vec (dim_poly p) (coeff p)"
lemma all_greater_dim_poly_zero[simp]:
assumes "x \<ge> dim_poly p"
shows "coeff p x = 0"
using Max_ge[of "vars p" x, OF finite_vars[of p]] coeff_zero[of p x]
by (metis add_cancel_left_right assms dim_poly.elims empty_iff leD le_eq_less_or_eq
trans_less_add1 zero_neq_one_class.zero_neq_one)
lemma lpoly_to_vec_0_iff_zero_poly [iff]:
shows "(lpoly_to_vec p) = 0\<^sub>v 0 \<longleftrightarrow> p = 0"
proof(standard)
show "lpoly_to_vec p = 0\<^sub>v 0 \<Longrightarrow> p = 0"
proof (rule contrapos_pp)
assume "p \<noteq> 0"
then have "vars p \<noteq> {}"
by (simp add: vars_empty_zero)
then have "dim_poly p > 0"
by (simp)
then show "lpoly_to_vec p \<noteq> 0\<^sub>v 0"
using vec_of_dim_0[of "lpoly_to_vec p"] by simp
qed
next
qed (auto simp: vars_empty_zero)
lemma dim_poly_dim_vec_equiv:
"dim_vec (lpoly_to_vec p) = dim_poly p"
using lpoly_to_vec.simps by auto
lemma dim_poly_greater_ex_coeff: "dim_poly x > d \<Longrightarrow> \<exists>i\<ge>d. coeff x i \<noteq> 0"
by (simp split: if_splits) (meson Max_in coeff_zero finite_vars less_Suc_eq_le)
lemma dimpoly_all_zero_limit:
assumes "\<And>i. i \<ge> d \<Longrightarrow> coeff x i = 0"
shows "dim_poly x \<le> d"
proof -
have "(\<forall>i\<ge> d. coeff x i = 0) \<Longrightarrow> dim_poly x \<le> d "
proof (rule contrapos_pp)
assume "\<not> dim_poly x \<le> d"
then have "dim_poly x > d" by linarith
then have "\<exists>i \<ge> d. coeff x i \<noteq> 0"
using dim_poly_greater_ex_coeff[of d x] by blast
then show "\<not> (\<forall>i\<ge>d. coeff x i = 0)"
by blast
qed
then show ?thesis
using assms by blast
qed
lemma construct_poly_from_lower_dim_poly:
assumes "dim_poly x = d+1"
obtains p c where "dim_poly p \<le> d" "x = p + lp_monom c d"
proof -
define c' where c': "c' = coeff x d"
have f: "\<forall>i>d. coeff x i = 0"
using assms by auto
have *: "x = x - (lp_monom c' d) + (lp_monom c' d)"
by simp
have "coeff (x - (lp_monom c' d)) d = 0"
using c' by simp
then have "\<forall>i\<ge>d. coeff (x - (lp_monom c' d)) i = 0"
using f by auto
then have **: "dim_poly (x - (lp_monom c' d)) \<le> d"
using dimpoly_all_zero_limit[of d "(x - (lp_monom c' d))"] by auto
define p' where p': "p' = x - (lp_monom c' d)"
have "\<exists>p c. dim_poly p \<le> d \<and> x = p + lp_monom c d"
using "*" "**" by blast
then show ?thesis
using * p' c' that by blast
qed
lemma vars_subset_0_dim_poly:
"vars z \<subseteq> {0..<dim_poly z}"
by (simp add: finite_vars less_Suc_eq_le subsetI)
lemma in_dim_and_not_var_zero: "x \<in> {0..<dim_poly z} - vars z \<Longrightarrow> coeff z x = 0"
using coeff_zero by auto
lemma valuate_with_dim_poly: "z \<lbrace> X \<rbrace> = (\<Sum>i\<in>{0..<dim_poly z}. coeff z i * X i)"
using eval_poly_with_sum_superset[of "{0..<dim_poly z}" z X] using vars_subset_0_dim_poly by blast
lemma lin_poly_to_vec_coeff_access:
assumes "x < dim_poly y"
shows "(lpoly_to_vec y) $ x = coeff y x"
proof -
have "x < dim_vec (lpoly_to_vec y)"
using dim_poly_dim_vec_equiv[of y] assms by auto
then show ?thesis
by (simp add: coeff_def)
qed
lemma addition_over_lin_poly_to_vec:
fixes x y
assumes "a < dim_poly x"
assumes "dim_poly x = dim_poly y"
shows "(lpoly_to_vec x + lpoly_to_vec y) $ a = coeff (x + y) a"
using assms(1) assms(2) lin_poly_to_vec_coeff_access by (simp add: dim_poly_dim_vec_equiv)
lemma list_to_lpoly_dim_less: "length cs \<ge> dim_poly (list_to_lpoly cs)"
using dimpoly_all_zero_limit sum_list_map_upto_coeff_limit by blast
text \<open> Transform rational vectors to linear polynomials \<close>
fun vec_to_lpoly where
"vec_to_lpoly rv = list_to_lpoly (list_of_vec rv)"
lemma vec_to_lin_poly_coeff_access:
assumes "x < dim_vec y"
shows "y $ x = coeff (vec_to_lpoly y) x"
by (simp add: assms list_to_lpoly_coeff_nth)
lemma addition_over_vec_to_lin_poly:
fixes x y
assumes "a < dim_vec x"
assumes "dim_vec x = dim_vec y"
shows "(x + y) $ a = coeff (vec_to_lpoly x + vec_to_lpoly y) a"
using assms(1) assms(2) coeff_plus index_add_vec(1)
by (metis vec_to_lin_poly_coeff_access)
lemma outside_list_coeff0:
assumes "i \<ge> dim_vec xs"
shows "coeff (vec_to_lpoly xs) i = 0"
by (simp add: assms sum_list_map_upto_coeff_limit)
lemma vec_to_poly_dim_less:
"dim_poly (vec_to_lpoly x) \<le> dim_vec x"
using list_to_lpoly_dim_less[of "list_of_vec x"] by simp
lemma vec_to_lpoly_from_lpoly_coeff_dual1:
"coeff (vec_to_lpoly (lpoly_to_vec p)) i = coeff p i"
by (metis all_greater_dim_poly_zero dim_poly_dim_vec_equiv lin_poly_to_vec_coeff_access
not_less outside_list_coeff0 vec_to_lin_poly_coeff_access)
lemma vec_to_lpoly_from_lpoly_coeff_dual2:
assumes "i < dim_vec (lpoly_to_vec (vec_to_lpoly v))"
shows "(lpoly_to_vec (vec_to_lpoly v)) $ i = v $ i"
by (metis assms dim_poly_dim_vec_equiv less_le_trans lin_poly_to_vec_coeff_access
vec_to_lin_poly_coeff_access vec_to_poly_dim_less)
lemma vars_subset_dim_vec_to_lpoly_dim: "vars (vec_to_lpoly v) \<subseteq> {0..<dim_vec v}"
by (meson ivl_subset le_numeral_extra(3) order.trans vec_to_poly_dim_less
vars_subset_0_dim_poly)
lemma sum_dim_vec_equals_sum_dim_poly:
shows "(\<Sum>a = 0..<dim_vec A. coeff (vec_to_lpoly A) a * X a) =
(\<Sum>a = 0..<dim_poly (vec_to_lpoly A). coeff (vec_to_lpoly A) a * X a)"
proof -
consider (eq) "dim_vec A = dim_poly (vec_to_lpoly A)" |
(le) "dim_vec A > dim_poly (vec_to_lpoly A)"
using vec_to_poly_dim_less[of "A"] by fastforce
then show ?thesis
proof (cases)
case le
define dp where dp: "dp = dim_poly (vec_to_lpoly A)"
have "(\<Sum>a = 0..<dim_vec A. coeff (vec_to_lpoly A) a * X a) =
(\<Sum>a = 0..<dp. coeff (vec_to_lpoly A) a * X a) +
(\<Sum>a = dp..<dim_vec A. coeff (vec_to_lpoly A) a * X a)"
by (metis (no_types, lifting) dp vec_to_poly_dim_less sum.atLeastLessThan_concat zero_le)
also have "... = (\<Sum>a = 0..<dp. coeff (vec_to_lpoly A) a * X a)"
using all_greater_dim_poly_zero by (simp add: dp)
also have "... = (\<Sum>a = 0..<dim_poly (vec_to_lpoly A).coeff (vec_to_lpoly A) a * X a)"
using dp by auto
finally show ?thesis
by blast
qed (auto)
qed
lemma vec_to_lpoly_vNil [simp]: "vec_to_lpoly vNil = 0"
by (simp add: empty_list_0poly)
lemma zero_vector_is_zero_poly: "coeff (vec_to_lpoly (0\<^sub>v n)) i = 0"
by (metis index_zero_vec(1) index_zero_vec(2) not_less
outside_list_coeff0 vec_to_lin_poly_coeff_access)
lemma coeff_nonzero_dim_vec_non_zero:
assumes "coeff (vec_to_lpoly v) i \<noteq> 0"
shows "v $ i \<noteq> 0" "i < dim_vec v"
apply (metis assms leI outside_list_coeff0 vec_to_lin_poly_coeff_access)
using assms leI outside_list_coeff0 by blast
lemma lpoly_of_v_equals_v_append0:
"vec_to_lpoly v = vec_to_lpoly (v @\<^sub>v 0\<^sub>v a)" (is "?lhs = ?rhs")
proof -
have "\<forall>i. coeff ?lhs i = coeff ?rhs i"
proof
fix i
consider (le) "i < dim_vec v" | (ge) "i \<ge> dim_vec v"
using leI by blast
then show "coeff (vec_to_lpoly v) i = coeff (vec_to_lpoly (v @\<^sub>v 0\<^sub>v a)) i"
proof (cases)
case le
then show ?thesis using vec_to_lin_poly_coeff_access[of i v] index_append_vec(1)
by (metis index_append_vec(2) vec_to_lin_poly_coeff_access trans_less_add1)
next
case ge
then have "coeff (vec_to_lpoly v) i = 0"
using outside_list_coeff0 by blast
moreover have "coeff (vec_to_lpoly (v @\<^sub>v 0\<^sub>v a)) i = 0"
proof (rule ccontr)
assume na: "\<not> coeff (vec_to_lpoly (v @\<^sub>v 0\<^sub>v a)) i = 0"
define va where v: "va = coeff (vec_to_lpoly (v @\<^sub>v 0\<^sub>v a)) i"
have "i < dim_vec (v @\<^sub>v 0\<^sub>v a)"
using coeff_nonzero_dim_vec_non_zero[of "(v @\<^sub>v 0\<^sub>v a)" i] na by blast
moreover have "(0\<^sub>v a) $ (i - dim_vec v) = va"
by (metis ge diff_is_0_eq' index_append_vec(1) index_append_vec(2)
not_less_zero vec_to_lin_poly_coeff_access v zero_less_diff calculation)
moreover have "va \<noteq> 0" using v na by linarith
ultimately show False
using ge by auto
qed
then show "coeff (vec_to_lpoly v) i = coeff (vec_to_lpoly (v @\<^sub>v 0\<^sub>v a)) i"
using not_less using calculation by linarith
qed
qed
then show ?thesis
using Abstract_Linear_Poly.poly_eqI by blast
qed
lemma vec_to_lpoly_eval_dot_prod:
"(vec_to_lpoly v) \<lbrace> x \<rbrace> = v \<bullet> (vec (dim_vec v) x)"
proof -
have "(vec_to_lpoly v) \<lbrace> x \<rbrace> = (\<Sum>i\<in>{0..<dim_vec v}. coeff (vec_to_lpoly v) i * x i)"
using eval_poly_with_sum_superset[of "{0..<dim_vec v}" "vec_to_lpoly v" x]
vars_subset_dim_vec_to_lpoly_dim by blast
also have "... = (\<Sum>i\<in>{0..<dim_vec v}. v$i * x i)"
using list_to_lpoly_coeff_nth by auto
also have "... = v \<bullet> (vec (dim_vec v) x)"
unfolding scalar_prod_def by auto
finally show ?thesis .
qed
lemma dim_poly_of_append_vec:
"dim_poly (vec_to_lpoly (a@\<^sub>vb)) \<le> dim_vec a + dim_vec b"
using vec_to_poly_dim_less[of "a@\<^sub>vb"] index_append_vec(2)[of a b] by auto
lemma vec_coeff_append1: "i \<in> {0..<dim_vec a} \<Longrightarrow> coeff (vec_to_lpoly (a@\<^sub>vb)) i = a$i"
by (metis atLeastLessThan_iff index_append_vec(1) index_append_vec(2) vec_to_lin_poly_coeff_access trans_less_add1)
lemma vec_coeff_append2:
"i \<in> {dim_vec a..<dim_vec (a@\<^sub>vb)} \<Longrightarrow> coeff (vec_to_lpoly (a@\<^sub>vb)) i = b$(i-dim_vec a)"
by (metis atLeastLessThan_iff index_append_vec(1) index_append_vec(2) leD vec_to_lin_poly_coeff_access)
text \<open> Maybe Code Equation \<close>
lemma vec_to_lpoly_poly_of_vec_eq: "vec_to_lpoly v = poly_of_vec v"
proof -
have "\<And>i. i < dim_vec v \<Longrightarrow> coeff (poly_of_vec v) i = v $ i"
by (simp add: coeff.rep_eq poly_of_vec.rep_eq)
moreover have "\<And>i. i < dim_vec v \<Longrightarrow> coeff (vec_to_lpoly v) i = v $ i"
by (simp add: vec_to_lin_poly_coeff_access)
moreover have "\<And>i. i \<ge> dim_vec v \<Longrightarrow> coeff (poly_of_vec v) i = 0"
by (simp add: coeff.rep_eq poly_of_vec.rep_eq)
moreover have "\<And>i. i \<ge> dim_vec v \<Longrightarrow> coeff (vec_to_lpoly v) i = 0"
using outside_list_coeff0 by blast
ultimately show ?thesis
by (metis Abstract_Linear_Poly.poly_eq_iff le_less_linear)
qed
lemma vars_vec_append_subset: "vars (vec_to_lpoly (0\<^sub>v n @\<^sub>v v)) \<subseteq> {n..<n+dim_vec v}"
proof -
let ?p = "(vec_to_lpoly (0\<^sub>v n @\<^sub>v v))"
have "dim_poly ?p \<le> n+dim_vec v"
using dim_poly_of_append_vec[of "0\<^sub>v n" "v"] by auto
have "vars (vec_to_lpoly (0\<^sub>v n @\<^sub>v v)) \<subseteq> {0..<n+dim_vec v}"
using vars_subset_dim_vec_to_lpoly_dim[of "(0\<^sub>v n @\<^sub>v v)"] by auto
moreover have "\<forall>i < n. coeff ?p i = 0"
using vec_coeff_append1[of _ "0\<^sub>v n" v] by auto
ultimately show "vars (vec_to_lpoly (0\<^sub>v n @\<^sub>v v)) \<subseteq> {n..<n+dim_vec v}"
by (meson atLeastLessThan_iff coeff_zero not_le subsetCE subsetI)
qed
section \<open> Matrices \<close>
(* \<open> From \<open> mat \<close> to \<open> linear_poly list \<close> *)
fun matrix_to_lpolies where
"matrix_to_lpolies A = map vec_to_lpoly (rows A)"
lemma matrix_to_lpolies_vec_of_row:
"i <dim_row A \<Longrightarrow> matrix_to_lpolies A ! i = vec_to_lpoly (row A i)"
using matrix_to_lpolies.simps[of A] by simp
lemma outside_of_col_range_is_0:
assumes "i < dim_row A" and "j \<ge> dim_col A"
shows "coeff ((matrix_to_lpolies A)!i) j = 0"
using outside_list_coeff0[of "col A i" j]
by (metis assms(1) assms(2) index_row(2) length_rows matrix_to_lpolies.simps nth_map nth_rows outside_list_coeff0)
lemma polys_greater_col_zero:
assumes "x \<in> set (matrix_to_lpolies A)"
assumes "j \<ge> dim_col A"
shows "coeff x j = 0"
using assms(1) assms(2) outside_of_col_range_is_0[of _ A j]
assms(2) matrix_to_lpolies.simps by (metis in_set_conv_nth length_map length_rows)
lemma matrix_to_lp_vec_to_lpoly_row [simp]:
assumes "i < dim_row A"
shows "(matrix_to_lpolies A)!i = vec_to_lpoly (row A i)"
by (simp add: assms)
lemma matrix_to_lpolies_coeff_access:
assumes "i < dim_row A" and "j < dim_col A"
shows "coeff (matrix_to_lpolies A ! i) j = A $$ (i,j)"
using matrix_to_lp_vec_to_lpoly_row[of i A, OF assms(1)]
by (metis assms(1) assms(2) index_row(1) index_row(2) vec_to_lin_poly_coeff_access)
text \<open> From linear polynomial list to matrix \<close>
definition lin_polies_to_mat where
"lin_polies_to_mat lst = mat (length lst) (max_dim_poly_list lst) (\<lambda>(x,y).coeff (lst!x) y)"
lemma lin_polies_to_rat_mat_coeff_index:
assumes "i < length L" and "j < (max_dim_poly_list L)"
shows "coeff (L ! i) j = (lin_polies_to_mat L) $$ (i,j)"
unfolding lin_polies_to_mat_def by (simp add: assms(1) assms(2))
lemma vec_to_lpoly_valuate_equiv_dot_prod:
assumes "dim_vec y = dim_vec x" (* Can be \<ge> *)
shows "(vec_to_lpoly y) \<lbrace> ($)x \<rbrace> = y \<bullet> x"
proof -
let ?p = "vec_to_lpoly y"
have 2: "?p\<lbrace> ($)x \<rbrace> = (\<Sum>j\<in>vars?p. coeff ?p j * x$j)"
using eval_poly_with_sum[of ?p "($)x"] by blast
have "vars ?p \<subseteq> {0..<dim_vec y}"
using vars_subset_dim_vec_to_lpoly_dim by blast
have 2: "?p\<lbrace> ($)x \<rbrace> = (\<Sum>j\<in>vars?p. coeff ?p j * x$j)"
using eval_poly_with_sum[of ?p "($)x"] by blast
also have *: "... = (\<Sum>i\<in>{0..<dim_poly ?p}. coeff ?p i * x$i)"
using valuate_with_dim_poly by (metis (no_types, lifting) calculation sum.cong)
also have "... = y \<bullet> x"
proof -
have "\<And>j. j < dim_vec x \<Longrightarrow> coeff (vec_to_lpoly y) j = y $ j"
using assms vec_to_lin_poly_coeff_access by auto
then show ?thesis
using vec_to_lpoly_eval_dot_prod[of "y" "($)x"]
by (metis assms calculation dim_vec index_vec vec_eq_iff)
qed
finally show ?thesis unfolding scalar_prod_def .
qed
lemma matrix_to_lpolies_valuate_scalarP:
assumes "i < dim_row A"
assumes "dim_col A = dim_vec x"
shows "(matrix_to_lpolies A!i) \<lbrace> ($)x \<rbrace> = (row A i) \<bullet> x"
using vec_to_lpoly_valuate_equiv_dot_prod[of "row A i" x]
by (simp add: assms(1) assms(2))
lemma matrix_to_lpolies_lambda_valuate_scalarP:
assumes "i < dim_row A"
assumes "dim_col A = dim_vec x"
shows "(matrix_to_lpolies A!i) \<lbrace> (\<lambda>i. (if i < dim_vec x then x$i else 0)) \<rbrace> = (row A i) \<bullet> x"
proof -
have "\<And>j. j < dim_vec x \<Longrightarrow> x$j = (\<lambda>i. (if i < dim_vec x then x$i else 0)) j"
by simp
let ?p = "(matrix_to_lpolies A!i)"
have "\<And>j. coeff (matrix_to_lpolies A!i) j \<noteq> 0 \<Longrightarrow> j < dim_vec x"
using outside_of_col_range_is_0[of i A] assms(1) assms(2) leI by auto
then have subs: "vars ?p \<subseteq> {0..<dim_vec x}"
using \<open>\<And>j. Abstract_Linear_Poly.coeff (matrix_to_lpolies A ! i) j \<noteq> 0 \<Longrightarrow> j < dim_vec x\<close> atLeastLessThan_iff coeff_zero by blast
then have *: "\<And>j. j \<in> vars ?p \<Longrightarrow> x$j = (\<lambda>i. (if i < dim_vec x then x$i else 0)) j"
by (simp add: \<open>\<And>j. Abstract_Linear_Poly.coeff (matrix_to_lpolies A ! i) j \<noteq> 0 \<Longrightarrow> j < dim_vec x\<close> coeff_zero)
have "row A i \<bullet> x = (?p \<lbrace> ($) x \<rbrace>)"
using assms(1) assms(2) matrix_to_lpolies_valuate_scalarP[of i A x] by linarith
also have "... = (\<Sum>j\<in> vars ?p. coeff ?p j * x$j)"
using eval_poly_with_sum by blast
also have "... = (\<Sum>j\<in> vars ?p. coeff ?p j * (\<lambda>i. (if i < dim_vec x then x$i else 0)) j)"
by (metis (full_types, opaque_lifting) \<open>\<And>j. Abstract_Linear_Poly.coeff (matrix_to_lpolies A ! i) j \<noteq> 0 \<Longrightarrow> j < dim_vec x\<close> mult.commute mult_zero_right)
also have "... = (?p \<lbrace> (\<lambda>i. (if i < dim_vec x then x$i else 0)) \<rbrace>)"
using eval_poly_with_sum by presburger
finally show ?thesis
by linarith
qed
end
\ No newline at end of file
diff --git a/thys/MiniSail/SyntaxL.thy b/thys/MiniSail/SyntaxL.thy
--- a/thys/MiniSail/SyntaxL.thy
+++ b/thys/MiniSail/SyntaxL.thy
@@ -1,383 +1,383 @@
(*<*)
theory SyntaxL
imports Syntax IVSubst
begin
(*>*)
chapter \<open>Syntax Lemmas\<close>
section \<open>Support, lookup and contexts\<close>
lemma supp_v_tau [simp]:
assumes "atom z \<sharp> v"
shows "supp (\<lbrace> z : b | CE_val (V_var z) == CE_val v \<rbrace>) = supp v \<union> supp b"
using assms \<tau>.supp c.supp ce.supp
by (simp add: fresh_def supp_at_base)
lemma supp_v_var_tau [simp]:
assumes "z \<noteq> x"
shows "supp (\<lbrace> z : b | CE_val (V_var z) == CE_val (V_var x) \<rbrace>) = { atom x } \<union> supp b"
using supp_v_tau assms
using supp_at_base by fastforce
text \<open> Sometimes we need to work with a version of a binder where the variable is fresh
in something else, such as a bigger context. I think these could be generated automatically \<close>
lemma obtain_fresh_fun_def:
fixes t::"'b::fs"
shows "\<exists>y::x. atom y \<sharp> (s,c,\<tau>,t) \<and> (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c ) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s))))"
proof -
obtain y::x where y: "atom y \<sharp> (s,c,\<tau>,t)" using obtain_fresh by blast
moreover have "AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c ) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)))"
proof(cases "x=y")
case True
then show ?thesis using fun_def.eq_iff Abs1_eq_iff(3) flip_commute flip_fresh_fresh fresh_PairD by auto
next
case False
have "(AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s)) =(AF_fun_typ x b c \<tau> s)" proof(subst fun_typ.eq_iff, subst Abs1_eq_iff(3))
show \<open>(y = x \<and> (((y \<leftrightarrow> x) \<bullet> c, (y \<leftrightarrow> x) \<bullet> \<tau>), (y \<leftrightarrow> x) \<bullet> s) = ((c, \<tau>), s) \<or>
y \<noteq> x \<and> (((y \<leftrightarrow> x) \<bullet> c, (y \<leftrightarrow> x) \<bullet> \<tau>), (y \<leftrightarrow> x) \<bullet> s) = (y \<leftrightarrow> x) \<bullet> ((c, \<tau>), s) \<and> atom y \<sharp> ((c, \<tau>), s)) \<and>
b = b\<close> using False flip_commute flip_fresh_fresh fresh_PairD y by auto
qed
thus ?thesis by metis
qed
ultimately show ?thesis using y fresh_Pair by metis
qed
lemma lookup_fun_member:
assumes "Some (AF_fundef f ft) = lookup_fun \<Phi> f"
shows "AF_fundef f ft \<in> set \<Phi>"
using assms proof (induct \<Phi>)
case Nil
then show ?case by auto
next
case (Cons a \<Phi>)
then show ?case using lookup_fun.simps
by (metis fun_def.exhaust insert_iff list.simps(15) option.inject)
qed
lemma rig_dom_eq:
"dom (G[x \<longmapsto> c]) = dom G"
proof(induct G rule: \<Gamma>.induct)
case GNil
then show ?case using replace_in_g.simps by presburger
next
case (GCons xbc \<Gamma>')
obtain x' and b' and c' where xbc: "xbc=(x',b',c')" using prod_cases3 by blast
then show ?case using replace_in_g.simps GCons by simp
qed
lemma lookup_in_rig_eq:
assumes "Some (b,c) = lookup \<Gamma> x"
shows "Some (b,c') = lookup (\<Gamma>[x\<longmapsto>c']) x"
using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
case GNil
then show ?case by auto
next
case (GCons x b c \<Gamma>')
then show ?case using replace_in_g.simps lookup.simps by auto
qed
lemma lookup_in_rig_neq:
assumes "Some (b,c) = lookup \<Gamma> y" and "x\<noteq>y"
shows "Some (b,c) = lookup (\<Gamma>[x\<longmapsto>c']) y"
using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
case GNil
then show ?case by auto
next
case (GCons x' b' c' \<Gamma>')
then show ?case using replace_in_g.simps lookup.simps by auto
qed
lemma lookup_in_rig:
assumes "Some (b,c) = lookup \<Gamma> y"
shows "\<exists>c''. Some (b,c'') = lookup (\<Gamma>[x\<longmapsto>c']) y"
proof(cases "x=y")
case True
then show ?thesis using lookup_in_rig_eq using assms by blast
next
case False
then show ?thesis using lookup_in_rig_neq using assms by blast
qed
lemma lookup_inside[simp]:
assumes "x \<notin> fst ` toSet \<Gamma>'"
shows "Some (b1,c1) = lookup (\<Gamma>'@(x,b1,c1)#\<^sub>\<Gamma>\<Gamma>) x"
using assms by(induct \<Gamma>',auto)
lemma lookup_inside2:
assumes "Some (b1,c1) = lookup (\<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)) y" and "x\<noteq>y"
shows "Some (b1,c1) = lookup (\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)) y"
using assms by(induct \<Gamma>' rule: \<Gamma>.induct,auto+)
fun tail:: "'a list \<Rightarrow> 'a list" where
"tail [] = []"
| "tail (x#xs) = xs"
lemma lookup_options:
assumes "Some (b,c) = lookup (xt#\<^sub>\<Gamma>G) x"
shows "((x,b,c) = xt) \<or> (Some (b,c) = lookup G x)"
by (metis assms lookup.simps(2) option.inject surj_pair)
lemma lookup_x:
assumes "Some (b,c) = lookup G x"
shows "x \<in> fst ` toSet G"
using assms
by(induct "G" rule: \<Gamma>.induct ,auto+)
lemma GCons_eq_appendI:
fixes xs1::\<Gamma>
shows "[| x #\<^sub>\<Gamma> xs1 = ys; xs = xs1 @ zs |] ==> x #\<^sub>\<Gamma> xs = ys @ zs"
by (drule sym) simp
lemma split_G: "x : toSet xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x #\<^sub>\<Gamma> zs"
proof (induct xs)
case GNil thus ?case by simp
next
case GCons thus ?case using GCons_eq_appendI
by (metis Un_iff append_g.simps(1) singletonD toSet.simps(2))
qed
lemma lookup_not_empty:
assumes "Some \<tau> = lookup G x"
shows "G \<noteq> GNil"
using assms by auto
lemma lookup_in_g:
assumes "Some (b,c) = lookup \<Gamma> x"
shows "(x,b,c) \<in> toSet \<Gamma>"
using assms apply(induct \<Gamma>, simp)
using lookup_options by fastforce
lemma lookup_split:
fixes \<Gamma>::\<Gamma>
assumes "Some (b,c) = lookup \<Gamma> x"
shows "\<exists>G G' . \<Gamma> = G'@(x,b,c)#\<^sub>\<Gamma>G"
by (meson assms(1) lookup_in_g split_G)
lemma toSet_splitU[simp]:
"(x',b',c') \<in> toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<longleftrightarrow> (x',b',c') \<in> (toSet \<Gamma>' \<union> {(x, b, c)} \<union> toSet \<Gamma>)"
using append_g_toSetU toSet.simps by auto
lemma toSet_splitP[simp]:
"(\<forall>(x', b', c')\<in>toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>). P x' b' c') \<longleftrightarrow> (\<forall> (x', b', c')\<in>toSet \<Gamma>'. P x' b' c') \<and> P x b c \<and> (\<forall> (x', b', c')\<in>toSet \<Gamma>. P x' b' c')" (is "?A \<longleftrightarrow> ?B")
using toSet_splitU by force
lemma lookup_restrict:
assumes "Some (b',c') = lookup (\<Gamma>'@(x,b,c)#\<^sub>\<Gamma>\<Gamma>) y" and "x \<noteq> y"
shows "Some (b',c') = lookup (\<Gamma>'@\<Gamma>) y"
using assms proof(induct \<Gamma>' rule:\<Gamma>_induct)
case GNil
then show ?case by auto
next
case (GCons x1 b1 c1 \<Gamma>')
then show ?case by auto
qed
lemma supp_list_member:
fixes x::"'a::fs" and l::"'a list"
assumes "x \<in> set l"
shows "supp x \<subseteq> supp l"
using assms apply(induct l, auto)
using supp_Cons by auto
lemma GNil_append:
assumes "GNil = G1@G2"
shows "G1 = GNil \<and> G2 = GNil"
proof(rule ccontr)
assume "\<not> (G1 = GNil \<and> G2 = GNil)"
hence "G1@G2 \<noteq> GNil" using append_g.simps by (metis \<Gamma>.distinct(1) \<Gamma>.exhaust)
thus False using assms by auto
qed
lemma GCons_eq_append_conv:
fixes xs::\<Gamma>
shows "x#\<^sub>\<Gamma>xs = ys@zs = (ys = GNil \<and> x#\<^sub>\<Gamma>xs = zs \<or> (\<exists>ys'. x#\<^sub>\<Gamma>ys' = ys \<and> xs = ys'@zs))"
by(cases ys) auto
lemma dclist_distinct_unique:
assumes "(dc, const) \<in> set dclist2" and "(cons, const1) \<in> set dclist2" and "dc=cons" and "distinct (List.map fst dclist2)"
shows "(const) = const1"
proof -
have "(cons, const) = (dc, const1)"
using assms by (metis (no_types, lifting) assms(3) assms(4) distinct.simps(1) distinct.simps(2) empty_iff insert_iff list.set(1) list.simps(15) list.simps(8) list.simps(9) map_of_eq_Some_iff)
thus ?thesis by auto
qed
lemma fresh_d_fst_d:
assumes "atom u \<sharp> \<delta>"
shows "u \<notin> fst ` set \<delta>"
using assms proof(induct \<delta>)
case Nil
then show ?case by auto
next
case (Cons ut \<delta>')
obtain u' and t' where *:"ut = (u',t') " by fastforce
hence "atom u \<sharp> ut \<and> atom u \<sharp> \<delta>'" using fresh_Cons Cons by auto
moreover hence "atom u \<sharp> fst ut" using * fresh_Pair[of "atom u" u' t'] Cons by auto
ultimately show ?case using Cons by auto
qed
lemma bv_not_in_bset_supp:
fixes bv::bv
assumes "bv |\<notin>| B"
shows "atom bv \<notin> supp B"
proof -
have *:"supp B = fset (fimage atom B)"
by (metis fimage.rep_eq finite_fset supp_finite_set_at_base supp_fset)
thus ?thesis using assms
- using notin_fset by fastforce
+ by fastforce
qed
lemma u_fresh_d:
assumes "atom u \<sharp> D"
shows "u \<notin> fst ` setD D"
using assms proof(induct D rule: \<Delta>_induct)
case DNil
then show ?case by auto
next
case (DCons u' t' \<Delta>')
then show ?case unfolding setD.simps
using fresh_DCons fresh_Pair by (simp add: fresh_Pair fresh_at_base(2))
qed
section \<open>Type Definitions\<close>
lemma exist_fresh_bv:
fixes tm::"'a::fs"
shows "\<exists>bva2 dclist2. AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
atom bva2 \<sharp> tm"
proof -
obtain bva2::bv where *:"atom bva2 \<sharp> (bva, dclist,tyid,tm)" using obtain_fresh by metis
moreover hence "bva2 \<noteq> bva" using fresh_at_base by auto
moreover have " dclist = (bva \<leftrightarrow> bva2) \<bullet> (bva2 \<leftrightarrow> bva) \<bullet> dclist" by simp
moreover have "atom bva \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist" proof -
have "atom bva2 \<sharp> dclist" using * fresh_prodN by auto
hence "atom ((bva2 \<leftrightarrow> bva) \<bullet> bva2) \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist" using fresh_eqvt True_eqvt
proof -
have "(bva2 \<leftrightarrow> bva) \<bullet> atom bva2 \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist"
by (metis True_eqvt \<open>atom bva2 \<sharp> dclist\<close> fresh_eqvt) (* 62 ms *)
then show ?thesis
by simp (* 125 ms *)
qed
thus ?thesis by auto
qed
ultimately have "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 ((bva2 \<leftrightarrow> bva ) \<bullet> dclist)"
unfolding type_def.eq_iff Abs1_eq_iff by metis
thus ?thesis using * fresh_prodN by metis
qed
lemma obtain_fresh_bv:
fixes tm::"'a::fs"
obtains bva2::bv and dclist2 where "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
atom bva2 \<sharp> tm"
using exist_fresh_bv by metis
section \<open>Function Definitions\<close>
lemma fun_typ_flip:
fixes bv1::bv and c::bv
shows "(bv1 \<leftrightarrow> c) \<bullet> AF_fun_typ x1 b1 c1 \<tau>1 s1 = AF_fun_typ x1 ((bv1 \<leftrightarrow> c) \<bullet> b1) ((bv1 \<leftrightarrow> c) \<bullet> c1) ((bv1 \<leftrightarrow> c) \<bullet> \<tau>1) ((bv1 \<leftrightarrow> c) \<bullet> s1)"
using fun_typ.perm_simps flip_fresh_fresh supp_at_base fresh_def
flip_fresh_fresh fresh_def supp_at_base
by (simp add: flip_fresh_fresh)
lemma fun_def_eq:
assumes "AF_fundef fa (AF_fun_typ_none (AF_fun_typ xa ba ca \<tau>a sa)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))"
shows "f = fa" and "b = ba" and "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \<tau>a = [[atom x]]lst. \<tau>" and
" [[atom xa]]lst. ca = [[atom x]]lst. c"
using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis
using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis
proof -
have "([[atom xa]]lst. ((ca, \<tau>a), sa) = [[atom x]]lst. ((c, \<tau>), s))" using assms fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff by auto
thus "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \<tau>a = [[atom x]]lst. \<tau>" and
" [[atom xa]]lst. ca = [[atom x]]lst. c" using lst_snd lst_fst by metis+
qed
lemma fun_arg_unique_aux:
assumes "AF_fun_typ x1 b1 c1 \<tau>1' s1' = AF_fun_typ x2 b2 c2 \<tau>2' s2'"
shows "\<lbrace> x1 : b1 | c1 \<rbrace> = \<lbrace> x2 : b2 | c2\<rbrace>"
proof -
have " ([[atom x1]]lst. c1 = [[atom x2]]lst. c2)" using fun_def_eq assms by metis
moreover have "b1 = b2" using fun_typ.eq_iff assms by metis
ultimately show ?thesis using \<tau>.eq_iff by fast
qed
lemma fresh_x_neq:
fixes x::x and y::x
shows "atom x \<sharp> y = (x \<noteq> y)"
using fresh_at_base fresh_def by auto
lemma obtain_fresh_z3:
fixes tm::"'b::fs"
obtains z::x where "\<lbrace> x : b | c \<rbrace> = \<lbrace> z : b | c[x::=V_var z]\<^sub>c\<^sub>v \<rbrace> \<and> atom z \<sharp> tm \<and> atom z \<sharp> (x,c)"
proof -
obtain z::x and c'::c where z:"\<lbrace> x : b | c \<rbrace> = \<lbrace> z : b | c' \<rbrace> \<and> atom z \<sharp> (tm,x,c)" using obtain_fresh_z2 b_of.simps by metis
hence "c' = c[x::=V_var z]\<^sub>c\<^sub>v" proof -
have "([[atom z]]lst. c' = [[atom x]]lst. c)" using z \<tau>.eq_iff by metis
hence "c' = (z \<leftrightarrow> x) \<bullet> c" using Abs1_eq_iff[of z c' x c] fresh_x_neq fresh_prodN by fastforce
also have "... = c[x::=V_var z]\<^sub>c\<^sub>v"
using subst_v_c_def flip_subst_v[of z c x] z fresh_prod3 by metis
finally show ?thesis by auto
qed
thus ?thesis using z fresh_prodN that by metis
qed
lemma u_fresh_v:
fixes u::u and t::v
shows "atom u \<sharp> t"
by(nominal_induct t rule:v.strong_induct,auto)
lemma u_fresh_ce:
fixes u::u and t::ce
shows "atom u \<sharp> t"
apply(nominal_induct t rule:ce.strong_induct)
using u_fresh_v pure_fresh
apply (auto simp add: opp.fresh ce.fresh opp.fresh opp.exhaust)
unfolding ce.fresh opp.fresh opp.exhaust by (simp add: fresh_opp_all)
lemma u_fresh_c:
fixes u::u and t::c
shows "atom u \<sharp> t"
by(nominal_induct t rule:c.strong_induct,auto simp add: c.fresh u_fresh_ce)
lemma u_fresh_g:
fixes u::u and t::\<Gamma>
shows "atom u \<sharp> t"
by(induct t rule:\<Gamma>_induct, auto simp add: u_fresh_b u_fresh_c fresh_GCons fresh_GNil)
lemma u_fresh_t:
fixes u::u and t::\<tau>
shows "atom u \<sharp> t"
by(nominal_induct t rule:\<tau>.strong_induct,auto simp add: \<tau>.fresh u_fresh_c u_fresh_b)
lemma b_of_c_of_eq:
assumes "atom z \<sharp> \<tau>"
shows "\<lbrace> z : b_of \<tau> | c_of \<tau> z \<rbrace> = \<tau>"
using assms proof(nominal_induct \<tau> avoiding: z rule: \<tau>.strong_induct)
case (T_refined_type x1a x2a x3a)
hence " \<lbrace> z : b_of \<lbrace> x1a : x2a | x3a \<rbrace> | c_of \<lbrace> x1a : x2a | x3a \<rbrace> z \<rbrace> = \<lbrace> z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \<rbrace>"
using b_of.simps c_of.simps c_of_eq by auto
moreover have "\<lbrace> z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> x1a : x2a | x3a \<rbrace>" using T_refined_type \<tau>.fresh by auto
ultimately show ?case by auto
qed
lemma fresh_d_not_in:
assumes "atom u2 \<sharp> \<Delta>'"
shows "u2 \<notin> fst ` setD \<Delta>'"
using assms proof(induct \<Delta>' rule: \<Delta>_induct)
case DNil
then show ?case by simp
next
case (DCons u t \<Delta>')
hence *: "atom u2 \<sharp> \<Delta>' \<and> atom u2 \<sharp> (u,t)"
by (simp add: fresh_def supp_DCons)
hence "u2 \<notin> fst ` setD \<Delta>'" using DCons by auto
moreover have "u2 \<noteq> u" using * fresh_Pair
by (metis eq_fst_iff not_self_fresh)
ultimately show ?case by simp
qed
end
\ No newline at end of file
diff --git a/thys/Nested_Multisets_Ordinals/Unary_PCF.thy b/thys/Nested_Multisets_Ordinals/Unary_PCF.thy
--- a/thys/Nested_Multisets_Ordinals/Unary_PCF.thy
+++ b/thys/Nested_Multisets_Ordinals/Unary_PCF.thy
@@ -1,659 +1,659 @@
(* Title: Towards Decidability of Behavioral Equivalence for Unary PCF
Author: Dmitriy Traytel <traytel at inf.ethz.ch>, 2017
Maintainer: Dmitriy Traytel <traytel at inf.ethz.ch>
*)
section \<open>Towards Decidability of Behavioral Equivalence for Unary PCF\<close>
theory Unary_PCF
imports
"HOL-Library.FSet"
"HOL-Library.Countable_Set_Type"
"HOL-Library.Nat_Bijection"
Hereditary_Multiset
"List-Index.List_Index"
begin
subsection \<open>Preliminaries\<close>
lemma prod_UNIV: "UNIV = UNIV \<times> UNIV"
by auto
lemma infinite_cartesian_productI1: "infinite A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> infinite (A \<times> B)"
by (auto dest!: finite_cartesian_productD1)
subsection \<open>Types\<close>
datatype type = \<B> ("\<B>") | Fun type type (infixr "\<rightarrow>" 65)
definition mk_fun (infixr "\<rightarrow>\<rightarrow>" 65) where
"Ts \<rightarrow>\<rightarrow> T = fold (\<rightarrow>) (rev Ts) T"
primrec dest_fun where
"dest_fun \<B> = []"
| "dest_fun (T \<rightarrow> U) = T # dest_fun U"
definition arity where
"arity T = length (dest_fun T)"
lemma mk_fun_dest_fun[simp]: "dest_fun T \<rightarrow>\<rightarrow> \<B> = T"
by (induct T) (auto simp: mk_fun_def)
lemma dest_fun_mk_fun[simp]: "dest_fun (Ts \<rightarrow>\<rightarrow> T) = Ts @ dest_fun T"
by (induct Ts) (auto simp: mk_fun_def)
primrec \<delta> where
"\<delta> \<B> = HMSet {#}"
| "\<delta> (T \<rightarrow> U) = HMSet (add_mset (\<delta> T) (hmsetmset (\<delta> U)))"
lemma \<delta>_mk_fun: "\<delta> (Ts \<rightarrow>\<rightarrow> T) = HMSet (hmsetmset (\<delta> T) + mset (map \<delta> Ts))"
by (induct Ts) (auto simp: mk_fun_def)
lemma type_induct [case_names Fun]:
assumes
"(\<And>T. (\<And>T1 T2. T = T1 \<rightarrow> T2 \<Longrightarrow> P T1) \<Longrightarrow>
(\<And>T1 T2. T = T1 \<rightarrow> T2 \<Longrightarrow> P T2) \<Longrightarrow> P T)"
shows "P T"
proof (induct T)
case \<B>
show ?case by (rule assms) simp_all
next
case Fun
show ?case by (rule assms) (insert Fun, simp_all)
qed
subsection \<open>Terms\<close>
type_synonym name = string
type_synonym idx = nat
datatype expr =
Var "name * type" ("\<langle>_\<rangle>") | Bound idx | B bool
| Seq expr expr (infixr "?" 75) | App expr expr (infixl "\<cdot>" 75)
| Abs type expr ("\<Lambda>\<langle>_\<rangle> _" [100, 100] 800)
declare [[coercion_enabled]]
declare [[coercion B]]
declare [[coercion Bound]]
notation (output) B ("_")
notation (output) Bound ("_")
primrec "open" :: "idx \<Rightarrow> expr \<Rightarrow> expr \<Rightarrow> expr" where
"open i t (j :: idx) = (if i = j then t else j)"
| "open i t \<langle>yU\<rangle> = \<langle>yU\<rangle>"
| "open i t (b :: bool) = b"
| "open i t (e1 ? e2) = open i t e1 ? open i t e2"
| "open i t (e1 \<cdot> e2) = open i t e1 \<cdot> open i t e2"
| "open i t (\<Lambda>\<langle>U\<rangle> e) = \<Lambda>\<langle>U\<rangle> (open (i + 1) t e)"
abbreviation "open0 \<equiv> open 0"
abbreviation "open_Var i xT \<equiv> open i \<langle>xT\<rangle>"
abbreviation "open0_Var xT \<equiv> open 0 \<langle>xT\<rangle>"
primrec "close_Var" :: "idx \<Rightarrow> name \<times> type \<Rightarrow> expr \<Rightarrow> expr" where
"close_Var i xT (j :: idx) = j"
| "close_Var i xT \<langle>yU\<rangle> = (if xT = yU then i else \<langle>yU\<rangle>)"
| "close_Var i xT (b :: bool) = b"
| "close_Var i xT (e1 ? e2) = close_Var i xT e1 ? close_Var i xT e2"
| "close_Var i xT (e1 \<cdot> e2) = close_Var i xT e1 \<cdot> close_Var i xT e2"
| "close_Var i xT (\<Lambda>\<langle>U\<rangle> e) = \<Lambda>\<langle>U\<rangle> (close_Var (i + 1) xT e)"
abbreviation "close0_Var \<equiv> close_Var 0"
primrec "fv" :: "expr \<Rightarrow> (name \<times> type) fset" where
"fv (j :: idx) = {||}"
| "fv \<langle>yU\<rangle> = {|yU|}"
| "fv (b :: bool) = {||}"
| "fv (e1 ? e2) = fv e1 |\<union>| fv e2"
| "fv (e1 \<cdot> e2) = fv e1 |\<union>| fv e2"
| "fv (\<Lambda>\<langle>U\<rangle> e) = fv e"
abbreviation "fresh x e \<equiv> x |\<notin>| fv e"
lemma ex_fresh: "\<exists>x. (x :: char list, T) |\<notin>| A"
proof (rule ccontr, unfold not_ex not_not)
assume "\<forall>x. (x, T) |\<in>| A"
then have "infinite {x. (x, T) |\<in>| A}" (is "infinite ?P")
by (auto simp add: infinite_UNIV_listI)
moreover
have "?P \<subseteq> fst ` fset A"
- by (force simp: fmember_iff_member_fset)
+ by force
from finite_surj[OF _ this] have "finite ?P"
by simp
ultimately show False by blast
qed
inductive lc where
lc_Var[simp]: "lc \<langle>xT\<rangle>"
| lc_B[simp]: "lc (b :: bool)"
| lc_Seq: "lc e1 \<Longrightarrow> lc e2 \<Longrightarrow> lc (e1 ? e2)"
| lc_App: "lc e1 \<Longrightarrow> lc e2 \<Longrightarrow> lc (e1 \<cdot> e2)"
| lc_Abs: "(\<forall>x. (x, T) |\<notin>| X \<longrightarrow> lc (open0_Var (x, T) e)) \<Longrightarrow> lc (\<Lambda>\<langle>T\<rangle> e)"
declare lc.intros[intro]
definition "body T t \<equiv> (\<exists>X. \<forall>x. (x, T) |\<notin>| X \<longrightarrow> lc (open0_Var (x, T) t))"
lemma lc_Abs_iff_body: "lc (\<Lambda>\<langle>T\<rangle> t) \<longleftrightarrow> body T t"
unfolding body_def by (subst lc.simps) simp
lemma fv_open_Var: "fresh xT t \<Longrightarrow> fv (open_Var i xT t) |\<subseteq>| finsert xT (fv t)"
by (induct t arbitrary: i) auto
lemma fv_close_Var[simp]: "fv (close_Var i xT t) = fv t |-| {|xT|}"
by (induct t arbitrary: i) auto
lemma close_Var_open_Var[simp]: "fresh xT t \<Longrightarrow> close_Var i xT (open_Var i xT t) = t"
by (induct t arbitrary: i) auto
lemma open_Var_inj: "fresh xT t \<Longrightarrow> fresh xT u \<Longrightarrow> open_Var i xT t = open_Var i xT u \<Longrightarrow> t = u"
by (metis close_Var_open_Var)
context begin
private lemma open_Var_open_Var_close_Var: "i \<noteq> j \<Longrightarrow> xT \<noteq> yU \<Longrightarrow> fresh yU t \<Longrightarrow>
open_Var i yU (open_Var j zV (close_Var j xT t)) = open_Var j zV (close_Var j xT (open_Var i yU t))"
by (induct t arbitrary: i j) auto
lemma open_Var_close_Var[simp]: "lc t \<Longrightarrow> open_Var i xT (close_Var i xT t) = t"
proof (induction t arbitrary: i rule: lc.induct)
case (lc_Abs T X e i)
obtain x where x: "fresh (x, T) e" "(x, T) \<noteq> xT" "(x, T) |\<notin>| X"
using ex_fresh[of _ "fv e |\<union>| finsert xT X"] by blast
with lc_Abs.IH have "lc (open0_Var (x, T) e)"
"open_Var (i + 1) xT (close_Var (i + 1) xT (open0_Var (x, T) e)) = open0_Var (x, T) e"
by auto
with x show ?case
by (auto simp: open_Var_open_Var_close_Var
dest: fset_mp[OF fv_open_Var, rotated]
intro!: open_Var_inj[of "(x, T)" _ e 0])
qed auto
end
lemma close_Var_inj: "lc t \<Longrightarrow> lc u \<Longrightarrow> close_Var i xT t = close_Var i xT u \<Longrightarrow> t = u"
by (metis open_Var_close_Var)
primrec Apps (infixl "\<bullet>" 75) where
"f \<bullet> [] = f"
| "f \<bullet> (x # xs) = f \<cdot> x \<bullet> xs"
lemma Apps_snoc: "f \<bullet> (xs @ [x]) = f \<bullet> xs \<cdot> x"
by (induct xs arbitrary: f) auto
lemma Apps_append: "f \<bullet> (xs @ ys) = f \<bullet> xs \<bullet> ys"
by (induct xs arbitrary: f) auto
lemma Apps_inj[simp]: "f \<bullet> ts = g \<bullet> ts \<longleftrightarrow> f = g"
by (induct ts arbitrary: f g) auto
lemma eq_Apps_conv[simp]:
fixes i :: idx and b :: bool and f :: expr and ts :: "expr list"
shows
"(\<langle>m\<rangle> = f \<bullet> ts) = (\<langle>m\<rangle> = f \<and> ts = [])"
"(f \<bullet> ts = \<langle>m\<rangle>) = (\<langle>m\<rangle> = f \<and> ts = [])"
"(i = f \<bullet> ts) = (i = f \<and> ts = [])"
"(f \<bullet> ts = i) = (i = f \<and> ts = [])"
"(b = f \<bullet> ts) = (b = f \<and> ts = [])"
"(f \<bullet> ts = b) = (b = f \<and> ts = [])"
"(e1 ? e2 = f \<bullet> ts) = (e1 ? e2 = f \<and> ts = [])"
"(f \<bullet> ts = e1 ? e2) = (e1 ? e2 = f \<and> ts = [])"
"(\<Lambda>\<langle>T\<rangle> t = f \<bullet> ts) = (\<Lambda>\<langle>T\<rangle> t = f \<and> ts = [])"
"(f \<bullet> ts = \<Lambda>\<langle>T\<rangle> t) = (\<Lambda>\<langle>T\<rangle> t = f \<and> ts = [])"
by (induct ts arbitrary: f) auto
lemma Apps_Var_eq[simp]: "\<langle>xT\<rangle> \<bullet> ss = \<langle>yU\<rangle> \<bullet> ts \<longleftrightarrow> xT = yU \<and> ss = ts"
proof (induct ss arbitrary: ts rule: rev_induct)
case snoc
then show ?case by (induct ts rule: rev_induct) (auto simp: Apps_snoc)
qed auto
lemma Apps_Abs_neq_Apps[simp, symmetric, simp]:
"\<Lambda>\<langle>T\<rangle> r \<cdot> t \<noteq> \<langle>xT\<rangle> \<bullet> ss"
"\<Lambda>\<langle>T\<rangle> r \<cdot> t \<noteq> (i :: idx) \<bullet> ss"
"\<Lambda>\<langle>T\<rangle> r \<cdot> t \<noteq> (b :: bool) \<bullet> ss"
"\<Lambda>\<langle>T\<rangle> r \<cdot> t \<noteq> (e1 ? e2) \<bullet> ss"
by (induct ss rule: rev_induct) (auto simp: Apps_snoc)
lemma App_Abs_eq_Apps_Abs[simp]: "\<Lambda>\<langle>T\<rangle> r \<cdot> t = \<Lambda>\<langle>T'\<rangle> r' \<bullet> ss \<longleftrightarrow> T = T' \<and> r = r' \<and> ss = [t]"
by (induct ss rule: rev_induct) (auto simp: Apps_snoc)
lemma Apps_Var_neq_Apps_Abs[simp, symmetric, simp]: "\<langle>xT\<rangle> \<bullet> ss \<noteq> \<Lambda>\<langle>T\<rangle> r \<bullet> ts"
proof (induct ss arbitrary: ts rule: rev_induct)
case (snoc a ss)
then show ?case by (induct ts rule: rev_induct) (auto simp: Apps_snoc)
qed simp
lemma Apps_Var_neq_Apps_beta[simp, THEN not_sym, simp]:
"\<langle>xT\<rangle> \<bullet> ss \<noteq> \<Lambda>\<langle>T\<rangle> r \<cdot> s \<bullet> ts"
by (metis Apps_Var_neq_Apps_Abs Apps_append Apps_snoc eq_Apps_conv(9))
lemma [simp]:
"(\<Lambda>\<langle>T\<rangle> r \<bullet> ts = \<Lambda>\<langle>T'\<rangle> r' \<cdot> s' \<bullet> ts') = (T = T' \<and> r = r' \<and> ts = s' # ts')"
proof (induct ts arbitrary: ts' rule: rev_induct)
case Nil
then show ?case by (induct ts' rule: rev_induct) (auto simp: Apps_snoc)
next
case snoc
then show ?case by (induct ts' rule: rev_induct) (auto simp: Apps_snoc)
qed
lemma fold_eq_Bool_iff[simp]:
"fold (\<rightarrow>) (rev Ts) T = \<B> \<longleftrightarrow> Ts = [] \<and> T = \<B>"
"\<B> = fold (\<rightarrow>) (rev Ts) T \<longleftrightarrow> Ts = [] \<and> T = \<B>"
by (induct Ts) auto
lemma fold_eq_Fun_iff[simp]:
"fold (\<rightarrow>) (rev Ts) T = U \<rightarrow> V \<longleftrightarrow>
(Ts = [] \<and> T = U \<rightarrow> V \<or> (\<exists>Us. Ts = U # Us \<and> fold (\<rightarrow>) (rev Us) T = V))"
by (induct Ts) auto
subsection \<open>Substitution\<close>
primrec subst where
"subst xT t \<langle>yU\<rangle> = (if xT = yU then t else \<langle>yU\<rangle>)"
| "subst xT t (i :: idx) = i"
| "subst xT t (b :: bool) = b"
| "subst xT t (e1 ? e2) = subst xT t e1 ? subst xT t e2"
| "subst xT t (e1 \<cdot> e2) = subst xT t e1 \<cdot> subst xT t e2"
| "subst xT t (\<Lambda>\<langle>T\<rangle> e) = \<Lambda>\<langle>T\<rangle> (subst xT t e)"
lemma fv_subst:
"fv (subst xT t u) = fv u |-| {|xT|} |\<union>| (if xT |\<in>| fv u then fv t else {||})"
by (induct u) auto
lemma subst_fresh: "fresh xT u \<Longrightarrow> subst xT t u = u"
by (induct u) auto
context begin
private lemma open_open_id: "i \<noteq> j \<Longrightarrow> open i t (open j t' u) = open j t' u \<Longrightarrow> open i t u = u"
by (induct u arbitrary: i j) (auto 6 0)
lemma lc_open_id: "lc u \<Longrightarrow> open k t u = u"
proof (induct u arbitrary: k rule: lc.induct)
case (lc_Abs T X e)
obtain x where x: "fresh (x, T) e" "(x, T) |\<notin>| X"
using ex_fresh[of _ "fv e |\<union>| X"] by blast
with lc_Abs show ?case
by (auto intro: open_open_id dest: spec[of _ x] spec[of _ "Suc k"])
qed auto
lemma subst_open: "lc u \<Longrightarrow> subst xT u (open i t v) = open i (subst xT u t) (subst xT u v)"
by (induction v arbitrary: i) (auto intro: lc_open_id[symmetric])
lemma subst_open_Var:
"xT \<noteq> yU \<Longrightarrow> lc u \<Longrightarrow> subst xT u (open_Var i yU v) = open_Var i yU (subst xT u v)"
by (auto simp: subst_open)
lemma subst_Apps[simp]:
"subst xT u (f \<bullet> xs) = subst xT u f \<bullet> map (subst xT u) xs"
by (induct xs arbitrary: f) auto
end
context begin
private lemma fresh_close_Var_id: "fresh xT t \<Longrightarrow> close_Var k xT t = t"
by (induct t arbitrary: k) auto
lemma subst_close_Var:
"xT \<noteq> yU \<Longrightarrow> fresh yU u \<Longrightarrow> subst xT u (close_Var i yU t) = close_Var i yU (subst xT u t)"
by (induct t arbitrary: i) (auto simp: fresh_close_Var_id)
end
lemma subst_intro: "fresh xT t \<Longrightarrow> lc u \<Longrightarrow> open0 u t = subst xT u (open0_Var xT t)"
by (auto simp: subst_fresh subst_open)
lemma lc_subst[simp]: "lc u \<Longrightarrow> lc t \<Longrightarrow> lc (subst xT t u)"
proof (induct u rule: lc.induct)
case (lc_Abs T X e)
then show ?case
by (auto simp: subst_open_Var intro!: lc.lc_Abs[of _ "fv e |\<union>| X |\<union>| {|xT|}"])
qed auto
lemma body_subst[simp]: "body U u \<Longrightarrow> lc t \<Longrightarrow> body U (subst xT t u)"
proof (subst (asm) body_def, elim conjE exE)
fix X
assume [simp]: "lc t" "\<forall>x. (x, U) |\<notin>| X \<longrightarrow> lc (open0_Var (x, U) u)"
show "body U (subst xT t u)"
proof (unfold body_def, intro exI[of _ "finsert xT X"] conjI allI impI)
fix x
assume "(x, U) |\<notin>| finsert xT X"
then show "lc (open0_Var (x, U) (subst xT t u))"
by (auto simp: subst_open_Var[symmetric])
qed
qed
lemma lc_open_Var: "lc u \<Longrightarrow> lc (open_Var i xT u)"
by (simp add: lc_open_id)
lemma lc_open[simp]: "body U u \<Longrightarrow> lc t \<Longrightarrow> lc (open0 t u)"
proof (unfold body_def, elim conjE exE)
fix X
assume [simp]: "lc t" "\<forall>x. (x, U) |\<notin>| X \<longrightarrow> lc (open0_Var (x, U) u)"
with ex_fresh[of _ "fv u |\<union>| X"] obtain x where [simp]: "fresh (x, U) u" "(x, U) |\<notin>| X" by blast
show ?thesis by (subst subst_intro[of "(x, U)"]) auto
qed
subsection \<open>Typing\<close>
inductive welltyped :: "expr \<Rightarrow> type \<Rightarrow> bool" (infix ":::" 60) where
welltyped_Var[intro!]: "\<langle>(x, T)\<rangle> ::: T"
| welltyped_B[intro!]: "(b :: bool) ::: \<B>"
| welltyped_Seq[intro!]: "e1 ::: \<B> \<Longrightarrow> e2 ::: \<B> \<Longrightarrow> e1 ? e2 ::: \<B>"
| welltyped_App[intro]: "e1 ::: T \<rightarrow> U \<Longrightarrow> e2 ::: T \<Longrightarrow> e1 \<cdot> e2 ::: U"
| welltyped_Abs[intro]: "(\<forall>x. (x, T) |\<notin>| X \<longrightarrow> open0_Var (x, T) e ::: U) \<Longrightarrow> \<Lambda>\<langle>T\<rangle> e ::: T \<rightarrow> U"
inductive_cases welltypedE[elim!]:
"\<langle>x\<rangle> ::: T"
"(i :: idx) ::: T"
"(b :: bool) ::: T"
"e1 ? e2 ::: T"
"e1 \<cdot> e2 ::: T"
"\<Lambda>\<langle>T\<rangle> e ::: U"
lemma welltyped_unique: "t ::: T \<Longrightarrow> t ::: U \<Longrightarrow> T = U"
proof (induction t T arbitrary: U rule: welltyped.induct)
case (welltyped_Abs T X t U T')
from welltyped_Abs.prems show ?case
proof (elim welltypedE)
fix Y U'
obtain x where [simp]: "(x, T) |\<notin>| X" "(x, T) |\<notin>| Y"
using ex_fresh[of _ "X |\<union>| Y"] by blast
assume [simp]: "T' = T \<rightarrow> U'" "\<forall>x. (x, T) |\<notin>| Y \<longrightarrow> open0_Var (x, T) t ::: U'"
show "T \<rightarrow> U = T'"
by (auto intro: conjunct2[OF welltyped_Abs.IH[rule_format], rule_format, of x])
qed
qed blast+
lemma welltyped_lc[simp]: "t ::: T \<Longrightarrow> lc t"
by (induction t T rule: welltyped.induct) auto
lemma welltyped_subst[intro]:
"u ::: U \<Longrightarrow> t ::: snd xT \<Longrightarrow> subst xT t u ::: U"
proof (induction u U rule: welltyped.induct)
case (welltyped_Abs T' X u U)
then show ?case unfolding subst.simps
by (intro welltyped.welltyped_Abs[of _ "finsert xT X"]) (auto simp: subst_open_Var[symmetric])
qed auto
lemma rename_welltyped: "u ::: U \<Longrightarrow> subst (x, T) \<langle>(y, T)\<rangle> u ::: U"
by (rule welltyped_subst) auto
lemma welltyped_Abs_fresh:
assumes "fresh (x, T) u" "open0_Var (x, T) u ::: U"
shows "\<Lambda>\<langle>T\<rangle> u ::: T \<rightarrow> U"
proof (intro welltyped_Abs[of _ "fv u"] allI impI)
fix y
assume "fresh (y, T) u"
with assms(2) have "subst (x, T) \<langle>(y, T)\<rangle> (open0_Var (x, T) u) ::: U" (is "?t ::: _")
by (auto intro: rename_welltyped)
also have "?t = open0_Var (y, T) u"
by (subst subst_intro[symmetric]) (auto simp: assms(1))
finally show "open0_Var (y, T) u ::: U" .
qed
lemma Apps_alt: "f \<bullet> ts ::: T \<longleftrightarrow>
(\<exists>Ts. f ::: fold (\<rightarrow>) (rev Ts) T \<and> list_all2 (:::) ts Ts)"
proof (induct ts arbitrary: f)
case (Cons t ts)
from Cons(1)[of "f \<cdot> t"] show ?case
by (force simp: list_all2_Cons1)
qed simp
subsection \<open>Definition 10 and Lemma 11 from Schmidt-Schau{\ss}'s paper\<close>
abbreviation "closed t \<equiv> fv t = {||}"
primrec constant0 where
"constant0 \<B> = Var (''bool'', \<B>)"
| "constant0 (T \<rightarrow> U) = \<Lambda>\<langle>T\<rangle> (constant0 U)"
definition "constant T = \<Lambda>\<langle>\<B>\<rangle> (close0_Var (''bool'', \<B>) (constant0 T))"
lemma fv_constant0[simp]: "fv (constant0 T) = {|(''bool'', \<B>)|}"
by (induct T) auto
lemma closed_constant[simp]: "closed (constant T)"
unfolding constant_def by auto
lemma welltyped_constant0[simp]: "constant0 T ::: T"
by (induct T) (auto simp: lc_open_id)
lemma lc_constant0[simp]: "lc (constant0 T)"
using welltyped_constant0 welltyped_lc by blast
lemma welltyped_constant[simp]: "constant T ::: \<B> \<rightarrow> T"
unfolding constant_def by (auto intro: welltyped_Abs_fresh[of "''bool''"])
definition nth_drop where
"nth_drop i xs \<equiv> take i xs @ drop (Suc i) xs"
definition nth_arg (infixl "!-" 100) where
"nth_arg T i \<equiv> nth (dest_fun T) i"
abbreviation ar where
"ar T \<equiv> length (dest_fun T)"
lemma size_nth_arg[simp]: "i < ar T \<Longrightarrow> size (T !- i) < size T"
by (induct T arbitrary: i) (force simp: nth_Cons' nth_arg_def gr0_conv_Suc)+
fun \<pi> :: "type \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> type" where
"\<pi> T i 0 = (if i < ar T then nth_drop i (dest_fun T) \<rightarrow>\<rightarrow> \<B> else \<B>)"
| "\<pi> T i (Suc j) = (if i < ar T \<and> j < ar (T!-i)
then \<pi> (T!-i) j 0 \<rightarrow>
map (\<pi> (T!-i) j o Suc) [0 ..< ar (T!-i!-j)] \<rightarrow>\<rightarrow> \<pi> T i 0 else \<B>)"
theorem \<pi>_induct[rotated -2, consumes 2, case_names 0 Suc]:
assumes "\<And>T i. i < ar T \<Longrightarrow> P T i 0"
and "\<And>T i j. i < ar T \<Longrightarrow> j < ar (T !- i) \<Longrightarrow> P (T !- i) j 0 \<Longrightarrow>
(\<forall>x < ar (T !- i !- j). P (T !- i) j (x + 1)) \<Longrightarrow> P T i (j + 1)"
shows "i < ar T \<Longrightarrow> j \<le> ar (T !- i) \<Longrightarrow> P T i j"
by (induct T i j rule: \<pi>.induct) (auto intro: assms[simplified])
definition \<epsilon> :: "type \<Rightarrow> nat \<Rightarrow> type" where
"\<epsilon> T i = \<pi> T i 0 \<rightarrow> map (\<pi> T i o Suc) [0 ..< ar (T!-i)] \<rightarrow>\<rightarrow> T"
definition Abss ("\<Lambda>[_] _" [100, 100] 800) where
"\<Lambda>[xTs] b = fold (\<lambda>xT t. \<Lambda>\<langle>snd xT\<rangle> close0_Var xT t) (rev xTs) b"
definition Seqs (infixr "??" 75) where
"ts ?? t = fold (\<lambda>u t. u ? t) (rev ts) t"
definition "variant k base = base @ replicate k CHR ''*''"
lemma variant_inj: "variant i base = variant j base \<Longrightarrow> i = j"
unfolding variant_def by auto
lemma variant_inj2:
"CHR ''*'' \<notin> set b1 \<Longrightarrow> CHR ''*'' \<notin> set b2 \<Longrightarrow> variant i b1 = variant j b2 \<Longrightarrow> b1 = b2"
unfolding variant_def
by (auto simp: append_eq_append_conv2)
(metis Nil_is_append_conv hd_append2 hd_in_set hd_rev last_ConsR
last_snoc replicate_append_same rev_replicate)+
fun E :: "type \<Rightarrow> nat \<Rightarrow> expr" and P :: "type \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> expr" where
"E T i = (if i < ar T then (let
Ti = T!-i;
x = \<lambda>k. (variant k ''x'', T!-k);
xs = map x [0 ..< ar T];
xx_var = \<langle>nth xs i\<rangle>;
x_vars = map (\<lambda>x. \<langle>x\<rangle>) (nth_drop i xs);
yy = (''z'', \<pi> T i 0);
yy_var = \<langle>yy\<rangle>;
y = \<lambda>j. (variant j ''y'', \<pi> T i (j + 1));
ys = map y [0 ..< ar Ti];
e = \<lambda>j. \<langle>y j\<rangle> \<bullet> (P Ti j 0 \<cdot> xx_var # map (\<lambda>k. P Ti j (k + 1) \<cdot> xx_var) [0 ..< ar (Ti!-j)]);
guards = map (\<lambda>i. xx_var \<bullet>
map (\<lambda>j. constant (Ti!-j) \<cdot> (if i = j then e i \<bullet> x_vars else True)) [0 ..< ar Ti])
[0 ..< ar Ti]
in \<Lambda>[(yy # ys @ xs)] (guards ?? (yy_var \<bullet> x_vars))) else constant (\<epsilon> T i) \<cdot> False)"
| "P T i 0 =
(if i < ar T then (let
f = (''f'', T);
f_var = \<langle>f\<rangle>;
x = \<lambda>k. (variant k ''x'', T!-k);
xs = nth_drop i (map x [0 ..< ar T]);
x_vars = insert_nth i (constant (T!-i) \<cdot> True) (map (\<lambda>x. \<langle>x\<rangle>) xs)
in \<Lambda>[(f # xs)] (f_var \<bullet> x_vars)) else constant (T \<rightarrow> \<pi> T i 0) \<cdot> False)"
| "P T i (Suc j) = (if i < ar T \<and> j < ar (T!-i) then (let
Ti = T!-i;
Tij = Ti!-j;
f = (''f'', T);
f_var = \<langle>f\<rangle>;
x = \<lambda>k. (variant k ''x'', T!-k);
xs = nth_drop i (map x [0 ..< ar T]);
yy = (''z'', \<pi> Ti j 0);
yy_var = \<langle>yy\<rangle>;
y = \<lambda>k. (variant k ''y'', \<pi> Ti j (k + 1));
ys = map y [0 ..< ar Tij];
y_vars = yy_var # map (\<lambda>x. \<langle>x\<rangle>) ys;
x_vars = insert_nth i (E Ti j \<bullet> y_vars) (map (\<lambda>x. \<langle>x\<rangle>) xs)
in \<Lambda>[(f # yy # ys @ xs)] (f_var \<bullet> x_vars)) else constant (T \<rightarrow> \<pi> T i (j + 1)) \<cdot> False)"
lemma Abss_Nil[simp]: "\<Lambda>[[]] b = b"
unfolding Abss_def by simp
lemma Abss_Cons[simp]: "\<Lambda>[(x#xs)] b = \<Lambda>\<langle>snd x\<rangle> (close0_Var x (\<Lambda>[xs] b))"
unfolding Abss_def by simp
lemma welltyped_Abss: "b ::: U \<Longrightarrow> T = map snd xTs \<rightarrow>\<rightarrow> U \<Longrightarrow> \<Lambda>[xTs] b ::: T"
by (hypsubst_thin, induct xTs) (auto simp: mk_fun_def intro!: welltyped_Abs_fresh)
lemma welltyped_Apps: "list_all2 (:::) ts Ts \<Longrightarrow> f ::: Ts \<rightarrow>\<rightarrow> U \<Longrightarrow> f \<bullet> ts ::: U"
by (induct ts Ts arbitrary: f rule: list.rel_induct) (auto simp: mk_fun_def)
lemma welltyped_open_Var_close_Var[intro!]:
"t ::: T \<Longrightarrow> open0_Var xT (close0_Var xT t) ::: T"
by auto
lemma welltyped_Var_iff[simp]:
"\<langle>(x, T)\<rangle> ::: U \<longleftrightarrow> T = U"
by auto
lemma welltyped_bool_iff[simp]: "(b :: bool) ::: T \<longleftrightarrow> T = \<B>"
by auto
lemma welltyped_constant0_iff[simp]: "constant0 T ::: U \<longleftrightarrow> (U = T)"
by (induct T arbitrary: U) (auto simp: ex_fresh lc_open_id)
lemma welltyped_constant_iff[simp]: "constant T ::: U \<longleftrightarrow> (U = \<B> \<rightarrow> T)"
unfolding constant_def
proof (intro iffI, elim welltypedE, hypsubst_thin, unfold type.inject simp_thms)
fix X U
assume "\<forall>x. (x, \<B>) |\<notin>| X \<longrightarrow> open0_Var (x, \<B>) (close0_Var (''bool'', \<B>) (constant0 T)) ::: U"
moreover obtain x where "(x, \<B>) |\<notin>| X" using ex_fresh[of \<B> X] by blast
ultimately have "open0_Var (x, \<B>) (close0_Var (''bool'', \<B>) (constant0 T)) ::: U" by simp
then have "open0_Var (''bool'', \<B>) (close0_Var (''bool'', \<B>) (constant0 T)) ::: U"
using rename_welltyped[of \<open>open0_Var (x, \<B>) (close0_Var (''bool'', \<B>) (constant0 T))\<close>
U x \<B> "''bool''"]
by (auto simp: subst_open subst_fresh)
then show "U = T" by auto
qed (auto intro!: welltyped_Abs_fresh)
lemma welltyped_Seq_iff[simp]: "e1 ? e2 ::: T \<longleftrightarrow> (T = \<B> \<and> e1 ::: \<B> \<and> e2 ::: \<B>)"
by auto
lemma welltyped_Seqs_iff[simp]: "es ?? e ::: T \<longleftrightarrow>
((es \<noteq> [] \<longrightarrow> T = \<B>) \<and> (\<forall>e \<in> set es. e ::: \<B>) \<and> e ::: T)"
by (induct es arbitrary: e) (auto simp: Seqs_def)
lemma welltyped_App_iff[simp]: "f \<cdot> t ::: U \<longleftrightarrow> (\<exists>T. f ::: T \<rightarrow> U \<and> t ::: T)"
by auto
lemma welltyped_Apps_iff[simp]: "f \<bullet> ts ::: U \<longleftrightarrow> (\<exists>Ts. f ::: Ts \<rightarrow>\<rightarrow> U \<and> list_all2 (:::) ts Ts)"
by (induct ts arbitrary: f) (auto 0 3 simp: mk_fun_def list_all2_Cons1 intro: exI[of _ "_ # _"])
lemma eq_mk_fun_iff[simp]: "T = Ts \<rightarrow>\<rightarrow> \<B> \<longleftrightarrow> Ts = dest_fun T"
by auto
lemma map_nth_eq_drop_take[simp]: "j \<le> length xs \<Longrightarrow> map (nth xs) [i ..< j] = drop i (take j xs)"
by (induct j) (auto simp: take_Suc_conv_app_nth)
lemma dest_fun_\<pi>_0: "i < ar T \<Longrightarrow> dest_fun (\<pi> T i 0) = nth_drop i (dest_fun T)"
by auto
lemma welltyped_E: "E T i ::: \<epsilon> T i" and welltyped_P: "P T i j ::: T \<rightarrow> \<pi> T i j"
proof (induct T i and T i j rule: E_P.induct)
case (1 T i)
note P.simps[simp del] \<pi>.simps[simp del] \<epsilon>_def[simp] nth_drop_def[simp] nth_arg_def[simp]
from 1(1)[OF _ refl refl refl refl refl refl refl refl refl]
1(2)[OF _ refl refl refl refl refl refl refl refl refl]
show ?case
by (auto 0 4 simp: Let_def o_def take_map[symmetric] drop_map[symmetric]
list_all2_conv_all_nth nth_append min_def dest_fun_\<pi>_0 \<pi>.simps[of T i]
intro!: welltyped_Abs_fresh welltyped_Abss[of _ \<B>])
next
case (2 T i)
show ?case
by (auto simp: Let_def take_map drop_map o_def list_all2_conv_all_nth nth_append nth_Cons'
nth_drop_def nth_arg_def
intro!: welltyped_constant welltyped_Abs_fresh welltyped_Abss[of _ \<B>])
next
case (3 T i j)
note E.simps[simp del] \<pi>.simps[simp del] Abss_Cons[simp del] \<epsilon>_def[simp]
nth_drop_def[simp] nth_arg_def[simp]
from 3(1)[OF _ refl refl refl refl refl refl refl refl refl refl refl]
show ?case
by (auto 0 3 simp: Let_def o_def take_map[symmetric] drop_map[symmetric]
list_all2_conv_all_nth nth_append nth_Cons' min_def \<pi>.simps[of T i]
intro!: welltyped_Abs_fresh welltyped_Abss[of _ \<B>])
qed
lemma \<delta>_gt_0[simp]: "T \<noteq> \<B> \<Longrightarrow> HMSet {#} < \<delta> T"
by (cases T) auto
lemma mset_nth_drop_less: "i < length xs \<Longrightarrow> mset (nth_drop i xs) < mset xs"
by (induct xs arbitrary: i) (auto simp: take_Cons' nth_drop_def gr0_conv_Suc)
lemma map_nth_drop: "i < length xs \<Longrightarrow> map f (nth_drop i xs) = nth_drop i (map f xs)"
by (induct xs arbitrary: i) (auto simp: take_Cons' nth_drop_def gr0_conv_Suc)
lemma empty_less_mset: "{#} < mset xs \<longleftrightarrow> xs \<noteq> []"
by auto
lemma dest_fun_alt: "dest_fun T = map (\<lambda>i. T !- i) [0..<ar T]"
unfolding list_eq_iff_nth_eq nth_arg_def by auto
context notes \<pi>.simps[simp del] notes One_nat_def[simp del] begin
lemma \<delta>_\<pi>:
assumes "i < ar T" "j \<le> ar (T !- i)"
shows "\<delta> (\<pi> T i j) < \<delta> T"
using assms proof (induct T i j rule: \<pi>_induct)
fix T i
assume "i < ar T"
then show "\<delta> (\<pi> T i 0) < \<delta> T"
by (subst (2) mk_fun_dest_fun[symmetric, of T], unfold \<delta>_mk_fun)
(auto simp: \<delta>_mk_fun mset_map[symmetric] take_map[symmetric] drop_map[symmetric] \<pi>.simps
mset_nth_drop_less map_nth_drop simp del: mset_map)
next
fix T i j
let ?Ti = "T !- i"
assume [rule_format, simp]: "i < ar T" "j < ar ?Ti" "\<delta> (\<pi> ?Ti j 0) < \<delta> ?Ti"
"\<forall>k < ar (?Ti !- j). \<delta> (\<pi> ?Ti j (k + 1)) < \<delta> ?Ti"
define X and Y and M where
[simp]: "X = {#\<delta> ?Ti#}" and
[simp]: "Y = {#\<delta> (\<pi> ?Ti j 0)#} + {#\<delta> (\<pi> ?Ti j (k + 1)). k \<in># mset [0 ..< ar (?Ti !- j)]#}" and
[simp]: "M \<equiv> {# \<delta> z. z \<in># mset (nth_drop i (dest_fun T))#}"
have "\<delta> (\<pi> T i (j + 1)) = HMSet (Y + M)"
by (auto simp: One_nat_def \<pi>.simps \<delta>_mk_fun)
also have "Y + M < X + M"
unfolding less_multiset\<^sub>D\<^sub>M by (rule exI[of _ "X"], rule exI[of _ "Y"]) auto
also have "HMSet (X + M) = \<delta> T"
unfolding M_def
by (subst (2) mk_fun_dest_fun[symmetric, of T], subst (2) id_take_nth_drop[of i "dest_fun T"])
(auto simp: \<delta>_mk_fun nth_arg_def nth_drop_def)
finally show "\<delta> (\<pi> T i (j + 1)) < \<delta> T" by simp
qed
end
end
diff --git a/thys/Nominal2/Nominal2_Base.thy b/thys/Nominal2/Nominal2_Base.thy
--- a/thys/Nominal2/Nominal2_Base.thy
+++ b/thys/Nominal2/Nominal2_Base.thy
@@ -1,3435 +1,3435 @@
(* Title: Nominal2_Base
Authors: Christian Urban, Brian Huffman, Cezary Kaliszyk
Basic definitions and lemma infrastructure for
Nominal Isabelle.
*)
theory Nominal2_Base
imports "HOL-Library.Infinite_Set"
"HOL-Library.Multiset"
"HOL-Library.FSet"
FinFun.FinFun
keywords
"atom_decl" "equivariance" :: thy_decl
begin
declare [[typedef_overloaded]]
section \<open>Atoms and Sorts\<close>
text \<open>A simple implementation for \<open>atom_sorts\<close> is strings.\<close>
(* types atom_sort = string *)
text \<open>To deal with Church-like binding we use trees of
strings as sorts.\<close>
datatype atom_sort = Sort "string" "atom_sort list"
datatype atom = Atom atom_sort nat
text \<open>Basic projection function.\<close>
primrec
sort_of :: "atom \<Rightarrow> atom_sort"
where
"sort_of (Atom s n) = s"
primrec
nat_of :: "atom \<Rightarrow> nat"
where
"nat_of (Atom s n) = n"
text \<open>There are infinitely many atoms of each sort.\<close>
lemma INFM_sort_of_eq:
shows "INFM a. sort_of a = s"
proof -
have "INFM i. sort_of (Atom s i) = s" by simp
moreover have "inj (Atom s)" by (simp add: inj_on_def)
ultimately show "INFM a. sort_of a = s" by (rule INFM_inj)
qed
lemma infinite_sort_of_eq:
shows "infinite {a. sort_of a = s}"
using INFM_sort_of_eq unfolding INFM_iff_infinite .
lemma atom_infinite [simp]:
shows "infinite (UNIV :: atom set)"
using subset_UNIV infinite_sort_of_eq
by (rule infinite_super)
lemma obtain_atom:
fixes X :: "atom set"
assumes X: "finite X"
obtains a where "a \<notin> X" "sort_of a = s"
proof -
from X have "MOST a. a \<notin> X"
unfolding MOST_iff_cofinite by simp
with INFM_sort_of_eq
have "INFM a. sort_of a = s \<and> a \<notin> X"
by (rule INFM_conjI)
then obtain a where "a \<notin> X" "sort_of a = s"
by (auto elim: INFM_E)
then show ?thesis ..
qed
lemma atom_components_eq_iff:
fixes a b :: atom
shows "a = b \<longleftrightarrow> sort_of a = sort_of b \<and> nat_of a = nat_of b"
by (induct a, induct b, simp)
section \<open>Sort-Respecting Permutations\<close>
definition
"perm \<equiv> {f. bij f \<and> finite {a. f a \<noteq> a} \<and> (\<forall>a. sort_of (f a) = sort_of a)}"
typedef perm = "perm"
proof
show "id \<in> perm" unfolding perm_def by simp
qed
lemma permI:
assumes "bij f" and "MOST x. f x = x" and "\<And>a. sort_of (f a) = sort_of a"
shows "f \<in> perm"
using assms unfolding perm_def MOST_iff_cofinite by simp
lemma perm_is_bij: "f \<in> perm \<Longrightarrow> bij f"
unfolding perm_def by simp
lemma perm_is_finite: "f \<in> perm \<Longrightarrow> finite {a. f a \<noteq> a}"
unfolding perm_def by simp
lemma perm_is_sort_respecting: "f \<in> perm \<Longrightarrow> sort_of (f a) = sort_of a"
unfolding perm_def by simp
lemma perm_MOST: "f \<in> perm \<Longrightarrow> MOST x. f x = x"
unfolding perm_def MOST_iff_cofinite by simp
lemma perm_id: "id \<in> perm"
unfolding perm_def by simp
lemma perm_comp:
assumes f: "f \<in> perm" and g: "g \<in> perm"
shows "(f \<circ> g) \<in> perm"
apply (rule permI)
apply (rule bij_comp)
apply (rule perm_is_bij [OF g])
apply (rule perm_is_bij [OF f])
apply (rule MOST_rev_mp [OF perm_MOST [OF g]])
apply (rule MOST_rev_mp [OF perm_MOST [OF f]])
apply (simp)
apply (simp add: perm_is_sort_respecting [OF f])
apply (simp add: perm_is_sort_respecting [OF g])
done
lemma perm_inv:
assumes f: "f \<in> perm"
shows "(inv f) \<in> perm"
apply (rule permI)
apply (rule bij_imp_bij_inv)
apply (rule perm_is_bij [OF f])
apply (rule MOST_mono [OF perm_MOST [OF f]])
apply (erule subst, rule inv_f_f)
apply (rule bij_is_inj [OF perm_is_bij [OF f]])
apply (rule perm_is_sort_respecting [OF f, THEN sym, THEN trans])
apply (simp add: surj_f_inv_f [OF bij_is_surj [OF perm_is_bij [OF f]]])
done
lemma bij_Rep_perm: "bij (Rep_perm p)"
using Rep_perm [of p] unfolding perm_def by simp
lemma finite_Rep_perm: "finite {a. Rep_perm p a \<noteq> a}"
using Rep_perm [of p] unfolding perm_def by simp
lemma sort_of_Rep_perm: "sort_of (Rep_perm p a) = sort_of a"
using Rep_perm [of p] unfolding perm_def by simp
lemma Rep_perm_ext:
"Rep_perm p1 = Rep_perm p2 \<Longrightarrow> p1 = p2"
by (simp add: fun_eq_iff Rep_perm_inject [symmetric])
instance perm :: size ..
subsection \<open>Permutations form a (multiplicative) group\<close>
instantiation perm :: group_add
begin
definition
"0 = Abs_perm id"
definition
"- p = Abs_perm (inv (Rep_perm p))"
definition
"p + q = Abs_perm (Rep_perm p \<circ> Rep_perm q)"
definition
"(p1::perm) - p2 = p1 + - p2"
lemma Rep_perm_0: "Rep_perm 0 = id"
unfolding zero_perm_def
by (simp add: Abs_perm_inverse perm_id)
lemma Rep_perm_add:
"Rep_perm (p1 + p2) = Rep_perm p1 \<circ> Rep_perm p2"
unfolding plus_perm_def
by (simp add: Abs_perm_inverse perm_comp Rep_perm)
lemma Rep_perm_uminus:
"Rep_perm (- p) = inv (Rep_perm p)"
unfolding uminus_perm_def
by (simp add: Abs_perm_inverse perm_inv Rep_perm)
instance
apply standard
unfolding Rep_perm_inject [symmetric]
unfolding minus_perm_def
unfolding Rep_perm_add
unfolding Rep_perm_uminus
unfolding Rep_perm_0
by (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]])
end
section \<open>Implementation of swappings\<close>
definition
swap :: "atom \<Rightarrow> atom \<Rightarrow> perm" ("'(_ \<rightleftharpoons> _')")
where
"(a \<rightleftharpoons> b) =
Abs_perm (if sort_of a = sort_of b
then (\<lambda>c. if a = c then b else if b = c then a else c)
else id)"
lemma Rep_perm_swap:
"Rep_perm (a \<rightleftharpoons> b) =
(if sort_of a = sort_of b
then (\<lambda>c. if a = c then b else if b = c then a else c)
else id)"
unfolding swap_def
apply (rule Abs_perm_inverse)
apply (rule permI)
apply (auto simp: bij_def inj_on_def surj_def)[1]
apply (rule MOST_rev_mp [OF MOST_neq(1) [of a]])
apply (rule MOST_rev_mp [OF MOST_neq(1) [of b]])
apply (simp)
apply (simp)
done
lemmas Rep_perm_simps =
Rep_perm_0
Rep_perm_add
Rep_perm_uminus
Rep_perm_swap
lemma swap_different_sorts [simp]:
"sort_of a \<noteq> sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) = 0"
by (rule Rep_perm_ext) (simp add: Rep_perm_simps)
lemma swap_cancel:
shows "(a \<rightleftharpoons> b) + (a \<rightleftharpoons> b) = 0"
and "(a \<rightleftharpoons> b) + (b \<rightleftharpoons> a) = 0"
by (rule_tac [!] Rep_perm_ext)
(simp_all add: Rep_perm_simps fun_eq_iff)
lemma swap_self [simp]:
"(a \<rightleftharpoons> a) = 0"
by (rule Rep_perm_ext, simp add: Rep_perm_simps fun_eq_iff)
lemma minus_swap [simp]:
"- (a \<rightleftharpoons> b) = (a \<rightleftharpoons> b)"
by (rule minus_unique [OF swap_cancel(1)])
lemma swap_commute:
"(a \<rightleftharpoons> b) = (b \<rightleftharpoons> a)"
by (rule Rep_perm_ext)
(simp add: Rep_perm_swap fun_eq_iff)
lemma swap_triple:
assumes "a \<noteq> b" and "c \<noteq> b"
assumes "sort_of a = sort_of b" "sort_of b = sort_of c"
shows "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
using assms
by (rule_tac Rep_perm_ext)
(auto simp: Rep_perm_simps fun_eq_iff)
section \<open>Permutation Types\<close>
text \<open>
Infix syntax for \<open>permute\<close> has higher precedence than
addition, but lower than unary minus.
\<close>
class pt =
fixes permute :: "perm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet> _" [76, 75] 75)
assumes permute_zero [simp]: "0 \<bullet> x = x"
assumes permute_plus [simp]: "(p + q) \<bullet> x = p \<bullet> (q \<bullet> x)"
begin
lemma permute_diff [simp]:
shows "(p - q) \<bullet> x = p \<bullet> - q \<bullet> x"
using permute_plus [of p "- q" x] by simp
lemma permute_minus_cancel [simp]:
shows "p \<bullet> - p \<bullet> x = x"
and "- p \<bullet> p \<bullet> x = x"
unfolding permute_plus [symmetric] by simp_all
lemma permute_swap_cancel [simp]:
shows "(a \<rightleftharpoons> b) \<bullet> (a \<rightleftharpoons> b) \<bullet> x = x"
unfolding permute_plus [symmetric]
by (simp add: swap_cancel)
lemma permute_swap_cancel2 [simp]:
shows "(a \<rightleftharpoons> b) \<bullet> (b \<rightleftharpoons> a) \<bullet> x = x"
unfolding permute_plus [symmetric]
by (simp add: swap_commute)
lemma inj_permute [simp]:
shows "inj (permute p)"
by (rule inj_on_inverseI)
(rule permute_minus_cancel)
lemma surj_permute [simp]:
shows "surj (permute p)"
by (rule surjI, rule permute_minus_cancel)
lemma bij_permute [simp]:
shows "bij (permute p)"
by (rule bijI [OF inj_permute surj_permute])
lemma inv_permute:
shows "inv (permute p) = permute (- p)"
by (rule inv_equality) (simp_all)
lemma permute_minus:
shows "permute (- p) = inv (permute p)"
by (simp add: inv_permute)
lemma permute_eq_iff [simp]:
shows "p \<bullet> x = p \<bullet> y \<longleftrightarrow> x = y"
by (rule inj_permute [THEN inj_eq])
end
subsection \<open>Permutations for atoms\<close>
instantiation atom :: pt
begin
definition
"p \<bullet> a = (Rep_perm p) a"
instance
apply standard
apply(simp_all add: permute_atom_def Rep_perm_simps)
done
end
lemma sort_of_permute [simp]:
shows "sort_of (p \<bullet> a) = sort_of a"
unfolding permute_atom_def by (rule sort_of_Rep_perm)
lemma swap_atom:
shows "(a \<rightleftharpoons> b) \<bullet> c =
(if sort_of a = sort_of b
then (if c = a then b else if c = b then a else c) else c)"
unfolding permute_atom_def
by (simp add: Rep_perm_swap)
lemma swap_atom_simps [simp]:
"sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> a = b"
"sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> b = a"
"c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> c = c"
unfolding swap_atom by simp_all
lemma perm_eq_iff:
fixes p q :: "perm"
shows "p = q \<longleftrightarrow> (\<forall>a::atom. p \<bullet> a = q \<bullet> a)"
unfolding permute_atom_def
by (metis Rep_perm_ext ext)
subsection \<open>Permutations for permutations\<close>
instantiation perm :: pt
begin
definition
"p \<bullet> q = p + q - p"
instance
apply standard
apply (simp add: permute_perm_def)
apply (simp add: permute_perm_def algebra_simps)
done
end
lemma permute_self:
shows "p \<bullet> p = p"
unfolding permute_perm_def
by (simp add: add.assoc)
lemma pemute_minus_self:
shows "- p \<bullet> p = p"
unfolding permute_perm_def
by (simp add: add.assoc)
subsection \<open>Permutations for functions\<close>
instantiation "fun" :: (pt, pt) pt
begin
definition
"p \<bullet> f = (\<lambda>x. p \<bullet> (f (- p \<bullet> x)))"
instance
apply standard
apply (simp add: permute_fun_def)
apply (simp add: permute_fun_def minus_add)
done
end
lemma permute_fun_app_eq:
shows "p \<bullet> (f x) = (p \<bullet> f) (p \<bullet> x)"
unfolding permute_fun_def by simp
lemma permute_fun_comp:
shows "p \<bullet> f = (permute p) o f o (permute (-p))"
by (simp add: comp_def permute_fun_def)
subsection \<open>Permutations for booleans\<close>
instantiation bool :: pt
begin
definition "p \<bullet> (b::bool) = b"
instance
apply standard
apply(simp_all add: permute_bool_def)
done
end
lemma permute_boolE:
fixes P::"bool"
shows "p \<bullet> P \<Longrightarrow> P"
by (simp add: permute_bool_def)
lemma permute_boolI:
fixes P::"bool"
shows "P \<Longrightarrow> p \<bullet> P"
by(simp add: permute_bool_def)
subsection \<open>Permutations for sets\<close>
instantiation "set" :: (pt) pt
begin
definition
"p \<bullet> X = {p \<bullet> x | x. x \<in> X}"
instance
apply standard
apply (auto simp: permute_set_def)
done
end
lemma permute_set_eq:
shows "p \<bullet> X = {x. - p \<bullet> x \<in> X}"
unfolding permute_set_def
by (auto) (metis permute_minus_cancel(1))
lemma permute_set_eq_image:
shows "p \<bullet> X = permute p ` X"
unfolding permute_set_def by auto
lemma permute_set_eq_vimage:
shows "p \<bullet> X = permute (- p) -` X"
unfolding permute_set_eq vimage_def
by simp
lemma permute_finite [simp]:
shows "finite (p \<bullet> X) = finite X"
unfolding permute_set_eq_vimage
using bij_permute by (rule finite_vimage_iff)
lemma swap_set_not_in:
assumes a: "a \<notin> S" "b \<notin> S"
shows "(a \<rightleftharpoons> b) \<bullet> S = S"
unfolding permute_set_def
using a by (auto simp: swap_atom)
lemma swap_set_in:
assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b"
shows "(a \<rightleftharpoons> b) \<bullet> S \<noteq> S"
unfolding permute_set_def
using a by (auto simp: swap_atom)
lemma swap_set_in_eq:
assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b"
shows "(a \<rightleftharpoons> b) \<bullet> S = (S - {a}) \<union> {b}"
unfolding permute_set_def
using a by (auto simp: swap_atom)
lemma swap_set_both_in:
assumes a: "a \<in> S" "b \<in> S"
shows "(a \<rightleftharpoons> b) \<bullet> S = S"
unfolding permute_set_def
using a by (auto simp: swap_atom)
lemma mem_permute_iff:
shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X"
unfolding permute_set_def
by auto
lemma empty_eqvt:
shows "p \<bullet> {} = {}"
unfolding permute_set_def
by (simp)
lemma insert_eqvt:
shows "p \<bullet> (insert x A) = insert (p \<bullet> x) (p \<bullet> A)"
unfolding permute_set_eq_image image_insert ..
subsection \<open>Permutations for @{typ unit}\<close>
instantiation unit :: pt
begin
definition "p \<bullet> (u::unit) = u"
instance
by standard (simp_all add: permute_unit_def)
end
subsection \<open>Permutations for products\<close>
instantiation prod :: (pt, pt) pt
begin
primrec
permute_prod
where
Pair_eqvt: "p \<bullet> (x, y) = (p \<bullet> x, p \<bullet> y)"
instance
by standard auto
end
subsection \<open>Permutations for sums\<close>
instantiation sum :: (pt, pt) pt
begin
primrec
permute_sum
where
Inl_eqvt: "p \<bullet> (Inl x) = Inl (p \<bullet> x)"
| Inr_eqvt: "p \<bullet> (Inr y) = Inr (p \<bullet> y)"
instance
by standard (case_tac [!] x, simp_all)
end
subsection \<open>Permutations for @{typ "'a list"}\<close>
instantiation list :: (pt) pt
begin
primrec
permute_list
where
Nil_eqvt: "p \<bullet> [] = []"
| Cons_eqvt: "p \<bullet> (x # xs) = p \<bullet> x # p \<bullet> xs"
instance
by standard (induct_tac [!] x, simp_all)
end
lemma set_eqvt:
shows "p \<bullet> (set xs) = set (p \<bullet> xs)"
by (induct xs) (simp_all add: empty_eqvt insert_eqvt)
subsection \<open>Permutations for @{typ "'a option"}\<close>
instantiation option :: (pt) pt
begin
primrec
permute_option
where
None_eqvt: "p \<bullet> None = None"
| Some_eqvt: "p \<bullet> (Some x) = Some (p \<bullet> x)"
instance
by standard (induct_tac [!] x, simp_all)
end
subsection \<open>Permutations for @{typ "'a multiset"}\<close>
instantiation multiset :: (pt) pt
begin
definition
"p \<bullet> M = {# p \<bullet> x. x :# M #}"
instance
proof
fix M :: "'a multiset" and p q :: "perm"
show "0 \<bullet> M = M"
unfolding permute_multiset_def
by (induct_tac M) (simp_all)
show "(p + q) \<bullet> M = p \<bullet> q \<bullet> M"
unfolding permute_multiset_def
by (induct_tac M) (simp_all)
qed
end
lemma permute_multiset [simp]:
fixes M N::"('a::pt) multiset"
shows "(p \<bullet> {#}) = ({#} ::('a::pt) multiset)"
and "(p \<bullet> add_mset x M) = add_mset (p \<bullet> x) (p \<bullet> M)"
and "(p \<bullet> (M + N)) = (p \<bullet> M) + (p \<bullet> N)"
unfolding permute_multiset_def
by (simp_all)
subsection \<open>Permutations for @{typ "'a fset"}\<close>
instantiation fset :: (pt) pt
begin
context includes fset.lifting begin
lift_definition
"permute_fset" :: "perm \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
is "permute :: perm \<Rightarrow> 'a set \<Rightarrow> 'a set" by simp
end
context includes fset.lifting begin
instance
proof
fix x :: "'a fset" and p q :: "perm"
show "0 \<bullet> x = x" by transfer simp
show "(p + q) \<bullet> x = p \<bullet> q \<bullet> x" by transfer simp
qed
end
end
context includes fset.lifting
begin
lemma permute_fset [simp]:
fixes S::"('a::pt) fset"
shows "(p \<bullet> {||}) = ({||} ::('a::pt) fset)"
and "(p \<bullet> finsert x S) = finsert (p \<bullet> x) (p \<bullet> S)"
apply (transfer, simp add: empty_eqvt)
apply (transfer, simp add: insert_eqvt)
done
lemma fset_eqvt:
shows "p \<bullet> (fset S) = fset (p \<bullet> S)"
by transfer simp
end
subsection \<open>Permutations for @{typ "('a, 'b) finfun"}\<close>
instantiation finfun :: (pt, pt) pt
begin
lift_definition
permute_finfun :: "perm \<Rightarrow> ('a, 'b) finfun \<Rightarrow> ('a, 'b) finfun"
is
"permute :: perm \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
apply(simp add: permute_fun_comp)
apply(rule finfun_right_compose)
apply(rule finfun_left_compose)
apply(assumption)
apply(simp)
done
instance
apply standard
apply(transfer)
apply(simp)
apply(transfer)
apply(simp)
done
end
subsection \<open>Permutations for @{typ char}, @{typ nat}, and @{typ int}\<close>
instantiation char :: pt
begin
definition "p \<bullet> (c::char) = c"
instance
by standard (simp_all add: permute_char_def)
end
instantiation nat :: pt
begin
definition "p \<bullet> (n::nat) = n"
instance
by standard (simp_all add: permute_nat_def)
end
instantiation int :: pt
begin
definition "p \<bullet> (i::int) = i"
instance
by standard (simp_all add: permute_int_def)
end
section \<open>Pure types\<close>
text \<open>Pure types will have always empty support.\<close>
class pure = pt +
assumes permute_pure: "p \<bullet> x = x"
text \<open>Types @{typ unit} and @{typ bool} are pure.\<close>
instance unit :: pure
proof qed (rule permute_unit_def)
instance bool :: pure
proof qed (rule permute_bool_def)
text \<open>Other type constructors preserve purity.\<close>
instance "fun" :: (pure, pure) pure
by standard (simp add: permute_fun_def permute_pure)
instance set :: (pure) pure
by standard (simp add: permute_set_def permute_pure)
instance prod :: (pure, pure) pure
by standard (induct_tac x, simp add: permute_pure)
instance sum :: (pure, pure) pure
by standard (induct_tac x, simp_all add: permute_pure)
instance list :: (pure) pure
by standard (induct_tac x, simp_all add: permute_pure)
instance option :: (pure) pure
by standard (induct_tac x, simp_all add: permute_pure)
subsection \<open>Types @{typ char}, @{typ nat}, and @{typ int}\<close>
instance char :: pure
proof qed (rule permute_char_def)
instance nat :: pure
proof qed (rule permute_nat_def)
instance int :: pure
proof qed (rule permute_int_def)
section \<open>Infrastructure for Equivariance and \<open>Perm_simp\<close>\<close>
subsection \<open>Basic functions about permutations\<close>
ML_file \<open>nominal_basics.ML\<close>
subsection \<open>Eqvt infrastructure\<close>
text \<open>Setup of the theorem attributes \<open>eqvt\<close> and \<open>eqvt_raw\<close>.\<close>
ML_file \<open>nominal_thmdecls.ML\<close>
lemmas [eqvt] =
(* pt types *)
permute_prod.simps
permute_list.simps
permute_option.simps
permute_sum.simps
(* sets *)
empty_eqvt insert_eqvt set_eqvt
(* fsets *)
permute_fset fset_eqvt
(* multisets *)
permute_multiset
subsection \<open>\<open>perm_simp\<close> infrastructure\<close>
definition
"unpermute p = permute (- p)"
lemma eqvt_apply:
fixes f :: "'a::pt \<Rightarrow> 'b::pt"
and x :: "'a::pt"
shows "p \<bullet> (f x) \<equiv> (p \<bullet> f) (p \<bullet> x)"
unfolding permute_fun_def by simp
lemma eqvt_lambda:
fixes f :: "'a::pt \<Rightarrow> 'b::pt"
shows "p \<bullet> f \<equiv> (\<lambda>x. p \<bullet> (f (unpermute p x)))"
unfolding permute_fun_def unpermute_def by simp
lemma eqvt_bound:
shows "p \<bullet> unpermute p x \<equiv> x"
unfolding unpermute_def by simp
text \<open>provides \<open>perm_simp\<close> methods\<close>
ML_file \<open>nominal_permeq.ML\<close>
method_setup perm_simp =
\<open>Nominal_Permeq.args_parser >> Nominal_Permeq.perm_simp_meth\<close>
\<open>pushes permutations inside.\<close>
method_setup perm_strict_simp =
\<open>Nominal_Permeq.args_parser >> Nominal_Permeq.perm_strict_simp_meth\<close>
\<open>pushes permutations inside, raises an error if it cannot solve all permutations.\<close>
simproc_setup perm_simproc ("p \<bullet> t") = \<open>fn _ => fn ctxt => fn ctrm =>
case Thm.term_of (Thm.dest_arg ctrm) of
Free _ => NONE
| Var _ => NONE
| \<^Const_>\<open>permute _ for _ _\<close> => NONE
| _ =>
let
val thm = Nominal_Permeq.eqvt_conv ctxt Nominal_Permeq.eqvt_strict_config ctrm
handle ERROR _ => Thm.reflexive ctrm
in
if Thm.is_reflexive thm then NONE else SOME(thm)
end
\<close>
subsubsection \<open>Equivariance for permutations and swapping\<close>
lemma permute_eqvt:
shows "p \<bullet> (q \<bullet> x) = (p \<bullet> q) \<bullet> (p \<bullet> x)"
unfolding permute_perm_def by simp
(* the normal version of this lemma would cause loops *)
lemma permute_eqvt_raw [eqvt_raw]:
shows "p \<bullet> permute \<equiv> permute"
apply(simp add: fun_eq_iff permute_fun_def)
apply(subst permute_eqvt)
apply(simp)
done
lemma zero_perm_eqvt [eqvt]:
shows "p \<bullet> (0::perm) = 0"
unfolding permute_perm_def by simp
lemma add_perm_eqvt [eqvt]:
fixes p p1 p2 :: perm
shows "p \<bullet> (p1 + p2) = p \<bullet> p1 + p \<bullet> p2"
unfolding permute_perm_def
by (simp add: perm_eq_iff)
lemma swap_eqvt [eqvt]:
shows "p \<bullet> (a \<rightleftharpoons> b) = (p \<bullet> a \<rightleftharpoons> p \<bullet> b)"
unfolding permute_perm_def
by (auto simp: swap_atom perm_eq_iff)
lemma uminus_eqvt [eqvt]:
fixes p q::"perm"
shows "p \<bullet> (- q) = - (p \<bullet> q)"
unfolding permute_perm_def
by (simp add: diff_add_eq_diff_diff_swap)
subsubsection \<open>Equivariance of Logical Operators\<close>
lemma eq_eqvt [eqvt]:
shows "p \<bullet> (x = y) \<longleftrightarrow> (p \<bullet> x) = (p \<bullet> y)"
unfolding permute_eq_iff permute_bool_def ..
lemma Not_eqvt [eqvt]:
shows "p \<bullet> (\<not> A) \<longleftrightarrow> \<not> (p \<bullet> A)"
by (simp add: permute_bool_def)
lemma conj_eqvt [eqvt]:
shows "p \<bullet> (A \<and> B) \<longleftrightarrow> (p \<bullet> A) \<and> (p \<bullet> B)"
by (simp add: permute_bool_def)
lemma imp_eqvt [eqvt]:
shows "p \<bullet> (A \<longrightarrow> B) \<longleftrightarrow> (p \<bullet> A) \<longrightarrow> (p \<bullet> B)"
by (simp add: permute_bool_def)
declare imp_eqvt[folded HOL.induct_implies_def, eqvt]
lemma all_eqvt [eqvt]:
shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. (p \<bullet> P) x)"
unfolding All_def
by (perm_simp) (rule refl)
declare all_eqvt[folded HOL.induct_forall_def, eqvt]
lemma ex_eqvt [eqvt]:
shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. (p \<bullet> P) x)"
unfolding Ex_def
by (perm_simp) (rule refl)
lemma ex1_eqvt [eqvt]:
shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. (p \<bullet> P) x)"
unfolding Ex1_def
by (perm_simp) (rule refl)
lemma if_eqvt [eqvt]:
shows "p \<bullet> (if b then x else y) = (if p \<bullet> b then p \<bullet> x else p \<bullet> y)"
by (simp add: permute_fun_def permute_bool_def)
lemma True_eqvt [eqvt]:
shows "p \<bullet> True = True"
unfolding permute_bool_def ..
lemma False_eqvt [eqvt]:
shows "p \<bullet> False = False"
unfolding permute_bool_def ..
lemma disj_eqvt [eqvt]:
shows "p \<bullet> (A \<or> B) \<longleftrightarrow> (p \<bullet> A) \<or> (p \<bullet> B)"
by (simp add: permute_bool_def)
lemma all_eqvt2:
shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. p \<bullet> P (- p \<bullet> x))"
by (perm_simp add: permute_minus_cancel) (rule refl)
lemma ex_eqvt2:
shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. p \<bullet> P (- p \<bullet> x))"
by (perm_simp add: permute_minus_cancel) (rule refl)
lemma ex1_eqvt2:
shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. p \<bullet> P (- p \<bullet> x))"
by (perm_simp add: permute_minus_cancel) (rule refl)
lemma the_eqvt:
assumes unique: "\<exists>!x. P x"
shows "(p \<bullet> (THE x. P x)) = (THE x. (p \<bullet> P) x)"
apply(rule the1_equality [symmetric])
apply(rule_tac p="-p" in permute_boolE)
apply(perm_simp add: permute_minus_cancel)
apply(rule unique)
apply(rule_tac p="-p" in permute_boolE)
apply(perm_simp add: permute_minus_cancel)
apply(rule theI'[OF unique])
done
lemma the_eqvt2:
assumes unique: "\<exists>!x. P x"
shows "(p \<bullet> (THE x. P x)) = (THE x. p \<bullet> P (- p \<bullet> x))"
apply(rule the1_equality [symmetric])
apply(simp only: ex1_eqvt2[symmetric])
apply(simp add: permute_bool_def unique)
apply(simp add: permute_bool_def)
apply(rule theI'[OF unique])
done
subsubsection \<open>Equivariance of Set operators\<close>
lemma mem_eqvt [eqvt]:
shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)"
unfolding permute_bool_def permute_set_def
by (auto)
lemma Collect_eqvt [eqvt]:
shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}"
unfolding permute_set_eq permute_fun_def
by (auto simp: permute_bool_def)
lemma Bex_eqvt [eqvt]:
shows "p \<bullet> (\<exists>x \<in> S. P x) = (\<exists>x \<in> (p \<bullet> S). (p \<bullet> P) x)"
unfolding Bex_def by simp
lemma Ball_eqvt [eqvt]:
shows "p \<bullet> (\<forall>x \<in> S. P x) = (\<forall>x \<in> (p \<bullet> S). (p \<bullet> P) x)"
unfolding Ball_def by simp
lemma image_eqvt [eqvt]:
shows "p \<bullet> (f ` A) = (p \<bullet> f) ` (p \<bullet> A)"
unfolding image_def by simp
lemma Image_eqvt [eqvt]:
shows "p \<bullet> (R `` A) = (p \<bullet> R) `` (p \<bullet> A)"
unfolding Image_def by simp
lemma UNIV_eqvt [eqvt]:
shows "p \<bullet> UNIV = UNIV"
unfolding UNIV_def
by (perm_simp) (rule refl)
lemma inter_eqvt [eqvt]:
shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)"
unfolding Int_def by simp
lemma Inter_eqvt [eqvt]:
shows "p \<bullet> \<Inter>S = \<Inter>(p \<bullet> S)"
unfolding Inter_eq by simp
lemma union_eqvt [eqvt]:
shows "p \<bullet> (A \<union> B) = (p \<bullet> A) \<union> (p \<bullet> B)"
unfolding Un_def by simp
lemma Union_eqvt [eqvt]:
shows "p \<bullet> \<Union>A = \<Union>(p \<bullet> A)"
unfolding Union_eq
by perm_simp rule
lemma Diff_eqvt [eqvt]:
fixes A B :: "'a::pt set"
shows "p \<bullet> (A - B) = (p \<bullet> A) - (p \<bullet> B)"
unfolding set_diff_eq by simp
lemma Compl_eqvt [eqvt]:
fixes A :: "'a::pt set"
shows "p \<bullet> (- A) = - (p \<bullet> A)"
unfolding Compl_eq_Diff_UNIV by simp
lemma subset_eqvt [eqvt]:
shows "p \<bullet> (S \<subseteq> T) \<longleftrightarrow> (p \<bullet> S) \<subseteq> (p \<bullet> T)"
unfolding subset_eq by simp
lemma psubset_eqvt [eqvt]:
shows "p \<bullet> (S \<subset> T) \<longleftrightarrow> (p \<bullet> S) \<subset> (p \<bullet> T)"
unfolding psubset_eq by simp
lemma vimage_eqvt [eqvt]:
shows "p \<bullet> (f -` A) = (p \<bullet> f) -` (p \<bullet> A)"
unfolding vimage_def by simp
lemma foldr_eqvt[eqvt]:
"p \<bullet> foldr f xs = foldr (p \<bullet> f) (p \<bullet> xs)"
apply(induct xs)
apply(simp_all)
apply(perm_simp exclude: foldr)
apply(simp)
done
(* FIXME: eqvt attribute *)
lemma Sigma_eqvt:
shows "(p \<bullet> (X \<times> Y)) = (p \<bullet> X) \<times> (p \<bullet> Y)"
unfolding Sigma_def
by (perm_simp) (rule refl)
text \<open>
In order to prove that lfp is equivariant we need two
auxiliary classes which specify that (<=) and
Inf are equivariant. Instances for bool and fun are
given.
\<close>
class le_eqvt = pt +
assumes le_eqvt [eqvt]: "p \<bullet> (x \<le> y) = ((p \<bullet> x) \<le> (p \<bullet> (y :: 'a :: {order, pt})))"
class inf_eqvt = pt +
assumes inf_eqvt [eqvt]: "p \<bullet> (Inf X) = Inf (p \<bullet> (X :: 'a :: {complete_lattice, pt} set))"
instantiation bool :: le_eqvt
begin
instance
apply standard
unfolding le_bool_def
apply(perm_simp)
apply(rule refl)
done
end
instantiation "fun" :: (pt, le_eqvt) le_eqvt
begin
instance
apply standard
unfolding le_fun_def
apply(perm_simp)
apply(rule refl)
done
end
instantiation bool :: inf_eqvt
begin
instance
apply standard
unfolding Inf_bool_def
apply(perm_simp)
apply(rule refl)
done
end
instantiation "fun" :: (pt, inf_eqvt) inf_eqvt
begin
instance
apply standard
unfolding Inf_fun_def
apply(perm_simp)
apply(rule refl)
done
end
lemma lfp_eqvt [eqvt]:
fixes F::"('a \<Rightarrow> 'b) \<Rightarrow> ('a::pt \<Rightarrow> 'b::{inf_eqvt, le_eqvt})"
shows "p \<bullet> (lfp F) = lfp (p \<bullet> F)"
unfolding lfp_def
by simp
lemma finite_eqvt [eqvt]:
shows "p \<bullet> finite A = finite (p \<bullet> A)"
unfolding finite_def
by simp
lemma fun_upd_eqvt[eqvt]:
shows "p \<bullet> (f(x := y)) = (p \<bullet> f)((p \<bullet> x) := (p \<bullet> y))"
unfolding fun_upd_def
by simp
lemma comp_eqvt [eqvt]:
shows "p \<bullet> (f \<circ> g) = (p \<bullet> f) \<circ> (p \<bullet> g)"
unfolding comp_def
by simp
subsubsection \<open>Equivariance for product operations\<close>
lemma fst_eqvt [eqvt]:
shows "p \<bullet> (fst x) = fst (p \<bullet> x)"
by (cases x) simp
lemma snd_eqvt [eqvt]:
shows "p \<bullet> (snd x) = snd (p \<bullet> x)"
by (cases x) simp
lemma split_eqvt [eqvt]:
shows "p \<bullet> (case_prod P x) = case_prod (p \<bullet> P) (p \<bullet> x)"
unfolding split_def
by simp
subsubsection \<open>Equivariance for list operations\<close>
lemma append_eqvt [eqvt]:
shows "p \<bullet> (xs @ ys) = (p \<bullet> xs) @ (p \<bullet> ys)"
by (induct xs) auto
lemma rev_eqvt [eqvt]:
shows "p \<bullet> (rev xs) = rev (p \<bullet> xs)"
by (induct xs) (simp_all add: append_eqvt)
lemma map_eqvt [eqvt]:
shows "p \<bullet> (map f xs) = map (p \<bullet> f) (p \<bullet> xs)"
by (induct xs) (simp_all)
lemma removeAll_eqvt [eqvt]:
shows "p \<bullet> (removeAll x xs) = removeAll (p \<bullet> x) (p \<bullet> xs)"
by (induct xs) (auto)
lemma filter_eqvt [eqvt]:
shows "p \<bullet> (filter f xs) = filter (p \<bullet> f) (p \<bullet> xs)"
apply(induct xs)
apply(simp)
apply(simp only: filter.simps permute_list.simps if_eqvt)
apply(simp only: permute_fun_app_eq)
done
lemma distinct_eqvt [eqvt]:
shows "p \<bullet> (distinct xs) = distinct (p \<bullet> xs)"
apply(induct xs)
apply(simp add: permute_bool_def)
apply(simp add: conj_eqvt Not_eqvt mem_eqvt set_eqvt)
done
lemma length_eqvt [eqvt]:
shows "p \<bullet> (length xs) = length (p \<bullet> xs)"
by (induct xs) (simp_all add: permute_pure)
subsubsection \<open>Equivariance for @{typ "'a option"}\<close>
lemma map_option_eqvt[eqvt]:
shows "p \<bullet> (map_option f x) = map_option (p \<bullet> f) (p \<bullet> x)"
by (cases x) (simp_all)
subsubsection \<open>Equivariance for @{typ "'a fset"}\<close>
context includes fset.lifting begin
-lemma in_fset_eqvt [eqvt]:
+lemma in_fset_eqvt:
shows "(p \<bullet> (x |\<in>| S)) = ((p \<bullet> x) |\<in>| (p \<bullet> S))"
by transfer simp
lemma union_fset_eqvt [eqvt]:
shows "(p \<bullet> (S |\<union>| T)) = ((p \<bullet> S) |\<union>| (p \<bullet> T))"
by (induct S) (simp_all)
lemma inter_fset_eqvt [eqvt]:
shows "(p \<bullet> (S |\<inter>| T)) = ((p \<bullet> S) |\<inter>| (p \<bullet> T))"
by transfer simp
lemma subset_fset_eqvt [eqvt]:
shows "(p \<bullet> (S |\<subseteq>| T)) = ((p \<bullet> S) |\<subseteq>| (p \<bullet> T))"
by transfer simp
lemma map_fset_eqvt [eqvt]:
shows "p \<bullet> (f |`| S) = (p \<bullet> f) |`| (p \<bullet> S)"
by transfer simp
end
subsubsection \<open>Equivariance for @{typ "('a, 'b) finfun"}\<close>
lemma finfun_update_eqvt [eqvt]:
shows "(p \<bullet> (finfun_update f a b)) = finfun_update (p \<bullet> f) (p \<bullet> a) (p \<bullet> b)"
by (transfer) (simp)
lemma finfun_const_eqvt [eqvt]:
shows "(p \<bullet> (finfun_const b)) = finfun_const (p \<bullet> b)"
by (transfer) (simp)
lemma finfun_apply_eqvt [eqvt]:
shows "(p \<bullet> (finfun_apply f b)) = finfun_apply (p \<bullet> f) (p \<bullet> b)"
by (transfer) (simp)
section \<open>Supp, Freshness and Supports\<close>
context pt
begin
definition
supp :: "'a \<Rightarrow> atom set"
where
"supp x = {a. infinite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}}"
definition
fresh :: "atom \<Rightarrow> 'a \<Rightarrow> bool" ("_ \<sharp> _" [55, 55] 55)
where
"a \<sharp> x \<equiv> a \<notin> supp x"
end
lemma supp_conv_fresh:
shows "supp x = {a. \<not> a \<sharp> x}"
unfolding fresh_def by simp
lemma swap_rel_trans:
assumes "sort_of a = sort_of b"
assumes "sort_of b = sort_of c"
assumes "(a \<rightleftharpoons> c) \<bullet> x = x"
assumes "(b \<rightleftharpoons> c) \<bullet> x = x"
shows "(a \<rightleftharpoons> b) \<bullet> x = x"
proof (cases)
assume "a = b \<or> c = b"
with assms show "(a \<rightleftharpoons> b) \<bullet> x = x" by auto
next
assume *: "\<not> (a = b \<or> c = b)"
have "((a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c)) \<bullet> x = x"
using assms by simp
also have "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
using assms * by (simp add: swap_triple)
finally show "(a \<rightleftharpoons> b) \<bullet> x = x" .
qed
lemma swap_fresh_fresh:
assumes a: "a \<sharp> x"
and b: "b \<sharp> x"
shows "(a \<rightleftharpoons> b) \<bullet> x = x"
proof (cases)
assume asm: "sort_of a = sort_of b"
have "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}" "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}"
using a b unfolding fresh_def supp_def by simp_all
then have "finite ({c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x} \<union> {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x})" by simp
then obtain c
where "(a \<rightleftharpoons> c) \<bullet> x = x" "(b \<rightleftharpoons> c) \<bullet> x = x" "sort_of c = sort_of b"
by (rule obtain_atom) (auto)
then show "(a \<rightleftharpoons> b) \<bullet> x = x" using asm by (rule_tac swap_rel_trans) (simp_all)
next
assume "sort_of a \<noteq> sort_of b"
then show "(a \<rightleftharpoons> b) \<bullet> x = x" by simp
qed
subsection \<open>supp and fresh are equivariant\<close>
lemma supp_eqvt [eqvt]:
shows "p \<bullet> (supp x) = supp (p \<bullet> x)"
unfolding supp_def by simp
lemma fresh_eqvt [eqvt]:
shows "p \<bullet> (a \<sharp> x) = (p \<bullet> a) \<sharp> (p \<bullet> x)"
unfolding fresh_def by simp
lemma fresh_permute_iff:
shows "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x"
by (simp only: fresh_eqvt[symmetric] permute_bool_def)
lemma fresh_permute_left:
shows "a \<sharp> p \<bullet> x \<longleftrightarrow> - p \<bullet> a \<sharp> x"
proof
assume "a \<sharp> p \<bullet> x"
then have "- p \<bullet> a \<sharp> - p \<bullet> p \<bullet> x" by (simp only: fresh_permute_iff)
then show "- p \<bullet> a \<sharp> x" by simp
next
assume "- p \<bullet> a \<sharp> x"
then have "p \<bullet> - p \<bullet> a \<sharp> p \<bullet> x" by (simp only: fresh_permute_iff)
then show "a \<sharp> p \<bullet> x" by simp
qed
section \<open>supports\<close>
definition
supports :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" (infixl "supports" 80)
where
"S supports x \<equiv> \<forall>a b. (a \<notin> S \<and> b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x)"
lemma supp_is_subset:
fixes S :: "atom set"
and x :: "'a::pt"
assumes a1: "S supports x"
and a2: "finite S"
shows "(supp x) \<subseteq> S"
proof (rule ccontr)
assume "\<not> (supp x \<subseteq> S)"
then obtain a where b1: "a \<in> supp x" and b2: "a \<notin> S" by auto
from a1 b2 have "\<forall>b. b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" unfolding supports_def by auto
then have "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x} \<subseteq> S" by auto
with a2 have "finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}" by (simp add: finite_subset)
then have "a \<notin> (supp x)" unfolding supp_def by simp
with b1 show False by simp
qed
lemma supports_finite:
fixes S :: "atom set"
and x :: "'a::pt"
assumes a1: "S supports x"
and a2: "finite S"
shows "finite (supp x)"
proof -
have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
then show "finite (supp x)" using a2 by (simp add: finite_subset)
qed
lemma supp_supports:
fixes x :: "'a::pt"
shows "(supp x) supports x"
unfolding supports_def
proof (intro strip)
fix a b
assume "a \<notin> (supp x) \<and> b \<notin> (supp x)"
then have "a \<sharp> x" and "b \<sharp> x" by (simp_all add: fresh_def)
then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh)
qed
lemma supports_fresh:
fixes x :: "'a::pt"
assumes a1: "S supports x"
and a2: "finite S"
and a3: "a \<notin> S"
shows "a \<sharp> x"
unfolding fresh_def
proof -
have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
then show "a \<notin> (supp x)" using a3 by auto
qed
lemma supp_is_least_supports:
fixes S :: "atom set"
and x :: "'a::pt"
assumes a1: "S supports x"
and a2: "finite S"
and a3: "\<And>S'. finite S' \<Longrightarrow> (S' supports x) \<Longrightarrow> S \<subseteq> S'"
shows "(supp x) = S"
proof (rule equalityI)
show "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
with a2 have fin: "finite (supp x)" by (rule rev_finite_subset)
have "(supp x) supports x" by (rule supp_supports)
with fin a3 show "S \<subseteq> supp x" by blast
qed
lemma subsetCI:
shows "(\<And>x. x \<in> A \<Longrightarrow> x \<notin> B \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> B"
by auto
lemma finite_supp_unique:
assumes a1: "S supports x"
assumes a2: "finite S"
assumes a3: "\<And>a b. \<lbrakk>a \<in> S; b \<notin> S; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
shows "(supp x) = S"
using a1 a2
proof (rule supp_is_least_supports)
fix S'
assume "finite S'" and "S' supports x"
show "S \<subseteq> S'"
proof (rule subsetCI)
fix a
assume "a \<in> S" and "a \<notin> S'"
have "finite (S \<union> S')"
using \<open>finite S\<close> \<open>finite S'\<close> by simp
then obtain b where "b \<notin> S \<union> S'" and "sort_of b = sort_of a"
by (rule obtain_atom)
then have "b \<notin> S" and "b \<notin> S'" and "sort_of a = sort_of b"
by simp_all
then have "(a \<rightleftharpoons> b) \<bullet> x = x"
using \<open>a \<notin> S'\<close> \<open>S' supports x\<close> by (simp add: supports_def)
moreover have "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
using \<open>a \<in> S\<close> \<open>b \<notin> S\<close> \<open>sort_of a = sort_of b\<close>
by (rule a3)
ultimately show "False" by simp
qed
qed
section \<open>Support w.r.t. relations\<close>
text \<open>
This definition is used for unquotient types, where
alpha-equivalence does not coincide with equality.
\<close>
definition
"supp_rel R x = {a. infinite {b. \<not>(R ((a \<rightleftharpoons> b) \<bullet> x) x)}}"
section \<open>Finitely-supported types\<close>
class fs = pt +
assumes finite_supp: "finite (supp x)"
lemma pure_supp:
fixes x::"'a::pure"
shows "supp x = {}"
unfolding supp_def by (simp add: permute_pure)
lemma pure_fresh:
fixes x::"'a::pure"
shows "a \<sharp> x"
unfolding fresh_def by (simp add: pure_supp)
instance pure < fs
by standard (simp add: pure_supp)
subsection \<open>Type \<^typ>\<open>atom\<close> is finitely-supported.\<close>
lemma supp_atom:
shows "supp a = {a}"
apply (rule finite_supp_unique)
apply (clarsimp simp add: supports_def)
apply simp
apply simp
done
lemma fresh_atom:
shows "a \<sharp> b \<longleftrightarrow> a \<noteq> b"
unfolding fresh_def supp_atom by simp
instance atom :: fs
by standard (simp add: supp_atom)
section \<open>Type \<^typ>\<open>perm\<close> is finitely-supported.\<close>
lemma perm_swap_eq:
shows "(a \<rightleftharpoons> b) \<bullet> p = p \<longleftrightarrow> (p \<bullet> (a \<rightleftharpoons> b)) = (a \<rightleftharpoons> b)"
unfolding permute_perm_def
by (metis add_diff_cancel minus_perm_def)
lemma supports_perm:
shows "{a. p \<bullet> a \<noteq> a} supports p"
unfolding supports_def
unfolding perm_swap_eq
by (simp add: swap_eqvt)
lemma finite_perm_lemma:
shows "finite {a::atom. p \<bullet> a \<noteq> a}"
using finite_Rep_perm [of p]
unfolding permute_atom_def .
lemma supp_perm:
shows "supp p = {a. p \<bullet> a \<noteq> a}"
apply (rule finite_supp_unique)
apply (rule supports_perm)
apply (rule finite_perm_lemma)
apply (simp add: perm_swap_eq swap_eqvt)
apply (auto simp: perm_eq_iff swap_atom)
done
lemma fresh_perm:
shows "a \<sharp> p \<longleftrightarrow> p \<bullet> a = a"
unfolding fresh_def
by (simp add: supp_perm)
lemma supp_swap:
shows "supp (a \<rightleftharpoons> b) = (if a = b \<or> sort_of a \<noteq> sort_of b then {} else {a, b})"
by (auto simp: supp_perm swap_atom)
lemma fresh_swap:
shows "a \<sharp> (b \<rightleftharpoons> c) \<longleftrightarrow> (sort_of b \<noteq> sort_of c) \<or> b = c \<or> (a \<sharp> b \<and> a \<sharp> c)"
by (simp add: fresh_def supp_swap supp_atom)
lemma fresh_zero_perm:
shows "a \<sharp> (0::perm)"
unfolding fresh_perm by simp
lemma supp_zero_perm:
shows "supp (0::perm) = {}"
unfolding supp_perm by simp
lemma fresh_plus_perm:
fixes p q::perm
assumes "a \<sharp> p" "a \<sharp> q"
shows "a \<sharp> (p + q)"
using assms
unfolding fresh_def
by (auto simp: supp_perm)
lemma supp_plus_perm:
fixes p q::perm
shows "supp (p + q) \<subseteq> supp p \<union> supp q"
by (auto simp: supp_perm)
lemma fresh_minus_perm:
fixes p::perm
shows "a \<sharp> (- p) \<longleftrightarrow> a \<sharp> p"
unfolding fresh_def
unfolding supp_perm
apply(simp)
apply(metis permute_minus_cancel)
done
lemma supp_minus_perm:
fixes p::perm
shows "supp (- p) = supp p"
unfolding supp_conv_fresh
by (simp add: fresh_minus_perm)
lemma plus_perm_eq:
fixes p q::"perm"
assumes asm: "supp p \<inter> supp q = {}"
shows "p + q = q + p"
unfolding perm_eq_iff
proof
fix a::"atom"
show "(p + q) \<bullet> a = (q + p) \<bullet> a"
proof -
{ assume "a \<notin> supp p" "a \<notin> supp q"
then have "(p + q) \<bullet> a = (q + p) \<bullet> a"
by (simp add: supp_perm)
}
moreover
{ assume a: "a \<in> supp p" "a \<notin> supp q"
then have "p \<bullet> a \<in> supp p" by (simp add: supp_perm)
then have "p \<bullet> a \<notin> supp q" using asm by auto
with a have "(p + q) \<bullet> a = (q + p) \<bullet> a"
by (simp add: supp_perm)
}
moreover
{ assume a: "a \<notin> supp p" "a \<in> supp q"
then have "q \<bullet> a \<in> supp q" by (simp add: supp_perm)
then have "q \<bullet> a \<notin> supp p" using asm by auto
with a have "(p + q) \<bullet> a = (q + p) \<bullet> a"
by (simp add: supp_perm)
}
ultimately show "(p + q) \<bullet> a = (q + p) \<bullet> a"
using asm by blast
qed
qed
lemma supp_plus_perm_eq:
fixes p q::perm
assumes asm: "supp p \<inter> supp q = {}"
shows "supp (p + q) = supp p \<union> supp q"
proof -
{ fix a::"atom"
assume "a \<in> supp p"
then have "a \<notin> supp q" using asm by auto
then have "a \<in> supp (p + q)" using \<open>a \<in> supp p\<close>
by (simp add: supp_perm)
}
moreover
{ fix a::"atom"
assume "a \<in> supp q"
then have "a \<notin> supp p" using asm by auto
then have "a \<in> supp (q + p)" using \<open>a \<in> supp q\<close>
by (simp add: supp_perm)
then have "a \<in> supp (p + q)" using asm plus_perm_eq
by metis
}
ultimately have "supp p \<union> supp q \<subseteq> supp (p + q)"
by blast
then show "supp (p + q) = supp p \<union> supp q" using supp_plus_perm
by blast
qed
lemma perm_eq_iff2:
fixes p q :: "perm"
shows "p = q \<longleftrightarrow> (\<forall>a::atom \<in> supp p \<union> supp q. p \<bullet> a = q \<bullet> a)"
unfolding perm_eq_iff
apply(auto)
apply(case_tac "a \<sharp> p \<and> a \<sharp> q")
apply(simp add: fresh_perm)
apply(simp add: fresh_def)
done
instance perm :: fs
by standard (simp add: supp_perm finite_perm_lemma)
section \<open>Finite Support instances for other types\<close>
subsection \<open>Type @{typ "'a \<times> 'b"} is finitely-supported.\<close>
lemma supp_Pair:
shows "supp (x, y) = supp x \<union> supp y"
by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
lemma fresh_Pair:
shows "a \<sharp> (x, y) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> y"
by (simp add: fresh_def supp_Pair)
lemma supp_Unit:
shows "supp () = {}"
by (simp add: supp_def)
lemma fresh_Unit:
shows "a \<sharp> ()"
by (simp add: fresh_def supp_Unit)
instance prod :: (fs, fs) fs
apply standard
apply (case_tac x)
apply (simp add: supp_Pair finite_supp)
done
subsection \<open>Type @{typ "'a + 'b"} is finitely supported\<close>
lemma supp_Inl:
shows "supp (Inl x) = supp x"
by (simp add: supp_def)
lemma supp_Inr:
shows "supp (Inr x) = supp x"
by (simp add: supp_def)
lemma fresh_Inl:
shows "a \<sharp> Inl x \<longleftrightarrow> a \<sharp> x"
by (simp add: fresh_def supp_Inl)
lemma fresh_Inr:
shows "a \<sharp> Inr y \<longleftrightarrow> a \<sharp> y"
by (simp add: fresh_def supp_Inr)
instance sum :: (fs, fs) fs
apply standard
apply (case_tac x)
apply (simp_all add: supp_Inl supp_Inr finite_supp)
done
subsection \<open>Type @{typ "'a option"} is finitely supported\<close>
lemma supp_None:
shows "supp None = {}"
by (simp add: supp_def)
lemma supp_Some:
shows "supp (Some x) = supp x"
by (simp add: supp_def)
lemma fresh_None:
shows "a \<sharp> None"
by (simp add: fresh_def supp_None)
lemma fresh_Some:
shows "a \<sharp> Some x \<longleftrightarrow> a \<sharp> x"
by (simp add: fresh_def supp_Some)
instance option :: (fs) fs
apply standard
apply (induct_tac x)
apply (simp_all add: supp_None supp_Some finite_supp)
done
subsubsection \<open>Type @{typ "'a list"} is finitely supported\<close>
lemma supp_Nil:
shows "supp [] = {}"
by (simp add: supp_def)
lemma fresh_Nil:
shows "a \<sharp> []"
by (simp add: fresh_def supp_Nil)
lemma supp_Cons:
shows "supp (x # xs) = supp x \<union> supp xs"
by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
lemma fresh_Cons:
shows "a \<sharp> (x # xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
by (simp add: fresh_def supp_Cons)
lemma supp_append:
shows "supp (xs @ ys) = supp xs \<union> supp ys"
by (induct xs) (auto simp: supp_Nil supp_Cons)
lemma fresh_append:
shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
by (induct xs) (simp_all add: fresh_Nil fresh_Cons)
lemma supp_rev:
shows "supp (rev xs) = supp xs"
by (induct xs) (auto simp: supp_append supp_Cons supp_Nil)
lemma fresh_rev:
shows "a \<sharp> rev xs \<longleftrightarrow> a \<sharp> xs"
by (induct xs) (auto simp: fresh_append fresh_Cons fresh_Nil)
lemma supp_removeAll:
fixes x::"atom"
shows "supp (removeAll x xs) = supp xs - {x}"
by (induct xs)
(auto simp: supp_Nil supp_Cons supp_atom)
lemma supp_of_atom_list:
fixes as::"atom list"
shows "supp as = set as"
by (induct as)
(simp_all add: supp_Nil supp_Cons supp_atom)
instance list :: (fs) fs
apply standard
apply (induct_tac x)
apply (simp_all add: supp_Nil supp_Cons finite_supp)
done
section \<open>Support and Freshness for Applications\<close>
lemma fresh_conv_MOST:
shows "a \<sharp> x \<longleftrightarrow> (MOST b. (a \<rightleftharpoons> b) \<bullet> x = x)"
unfolding fresh_def supp_def
unfolding MOST_iff_cofinite by simp
lemma fresh_fun_app:
assumes "a \<sharp> f" and "a \<sharp> x"
shows "a \<sharp> f x"
using assms
unfolding fresh_conv_MOST
unfolding permute_fun_app_eq
by (elim MOST_rev_mp) (simp)
lemma supp_fun_app:
shows "supp (f x) \<subseteq> (supp f) \<union> (supp x)"
using fresh_fun_app
unfolding fresh_def
by auto
subsection \<open>Equivariance Predicate \<open>eqvt\<close> and \<open>eqvt_at\<close>\<close>
definition
"eqvt f \<equiv> \<forall>p. p \<bullet> f = f"
lemma eqvt_boolI:
fixes f::"bool"
shows "eqvt f"
unfolding eqvt_def by (simp add: permute_bool_def)
text \<open>equivariance of a function at a given argument\<close>
definition
"eqvt_at f x \<equiv> \<forall>p. p \<bullet> (f x) = f (p \<bullet> x)"
lemma eqvtI:
shows "(\<And>p. p \<bullet> f \<equiv> f) \<Longrightarrow> eqvt f"
unfolding eqvt_def
by simp
lemma eqvt_at_perm:
assumes "eqvt_at f x"
shows "eqvt_at f (q \<bullet> x)"
proof -
{ fix p::"perm"
have "p \<bullet> (f (q \<bullet> x)) = p \<bullet> q \<bullet> (f x)"
using assms by (simp add: eqvt_at_def)
also have "\<dots> = (p + q) \<bullet> (f x)" by simp
also have "\<dots> = f ((p + q) \<bullet> x)"
using assms by (simp only: eqvt_at_def)
finally have "p \<bullet> (f (q \<bullet> x)) = f (p \<bullet> q \<bullet> x)" by simp }
then show "eqvt_at f (q \<bullet> x)" unfolding eqvt_at_def
by simp
qed
lemma supp_fun_eqvt:
assumes a: "eqvt f"
shows "supp f = {}"
using a
unfolding eqvt_def
unfolding supp_def
by simp
lemma fresh_fun_eqvt:
assumes a: "eqvt f"
shows "a \<sharp> f"
using a
unfolding fresh_def
by (simp add: supp_fun_eqvt)
lemma fresh_fun_eqvt_app:
assumes a: "eqvt f"
shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
proof -
from a have "supp f = {}" by (simp add: supp_fun_eqvt)
then show "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
unfolding fresh_def
using supp_fun_app by auto
qed
lemma supp_fun_app_eqvt:
assumes a: "eqvt f"
shows "supp (f x) \<subseteq> supp x"
using fresh_fun_eqvt_app[OF a]
unfolding fresh_def
by auto
lemma supp_eqvt_at:
assumes asm: "eqvt_at f x"
and fin: "finite (supp x)"
shows "supp (f x) \<subseteq> supp x"
apply(rule supp_is_subset)
unfolding supports_def
unfolding fresh_def[symmetric]
using asm
apply(simp add: eqvt_at_def)
apply(simp add: swap_fresh_fresh)
apply(rule fin)
done
lemma finite_supp_eqvt_at:
assumes asm: "eqvt_at f x"
and fin: "finite (supp x)"
shows "finite (supp (f x))"
apply(rule finite_subset)
apply(rule supp_eqvt_at[OF asm fin])
apply(rule fin)
done
lemma fresh_eqvt_at:
assumes asm: "eqvt_at f x"
and fin: "finite (supp x)"
and fresh: "a \<sharp> x"
shows "a \<sharp> f x"
using fresh
unfolding fresh_def
using supp_eqvt_at[OF asm fin]
by auto
text \<open>for handling of freshness of functions\<close>
simproc_setup fresh_fun_simproc ("a \<sharp> (f::'a::pt \<Rightarrow>'b::pt)") = \<open>fn _ => fn ctxt => fn ctrm =>
let
val _ $ _ $ f = Thm.term_of ctrm
in
case (Term.add_frees f [], Term.add_vars f []) of
([], []) => SOME(@{thm fresh_fun_eqvt[simplified eqvt_def, THEN Eq_TrueI]})
| (x::_, []) =>
let
val argx = Free x
val absf = absfree x f
val cty_inst =
[SOME (Thm.ctyp_of ctxt (fastype_of argx)), SOME (Thm.ctyp_of ctxt (fastype_of f))]
val ctrm_inst = [NONE, SOME (Thm.cterm_of ctxt absf), SOME (Thm.cterm_of ctxt argx)]
val thm = Thm.instantiate' cty_inst ctrm_inst @{thm fresh_fun_app}
in
SOME(thm RS @{thm Eq_TrueI})
end
| (_, _) => NONE
end
\<close>
subsection \<open>helper functions for \<open>nominal_functions\<close>\<close>
lemma THE_defaultI2:
assumes "\<exists>!x. P x" "\<And>x. P x \<Longrightarrow> Q x"
shows "Q (THE_default d P)"
by (iprover intro: assms THE_defaultI')
lemma the_default_eqvt:
assumes unique: "\<exists>!x. P x"
shows "(p \<bullet> (THE_default d P)) = (THE_default (p \<bullet> d) (p \<bullet> P))"
apply(rule THE_default1_equality [symmetric])
apply(rule_tac p="-p" in permute_boolE)
apply(simp add: ex1_eqvt)
apply(rule unique)
apply(rule_tac p="-p" in permute_boolE)
apply(rule subst[OF permute_fun_app_eq])
apply(simp)
apply(rule THE_defaultI'[OF unique])
done
lemma fundef_ex1_eqvt:
fixes x::"'a::pt"
assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (G x))"
assumes eqvt: "eqvt G"
assumes ex1: "\<exists>!y. G x y"
shows "(p \<bullet> (f x)) = f (p \<bullet> x)"
apply(simp only: f_def)
apply(subst the_default_eqvt)
apply(rule ex1)
apply(rule THE_default1_equality [symmetric])
apply(rule_tac p="-p" in permute_boolE)
apply(perm_simp add: permute_minus_cancel)
using eqvt[simplified eqvt_def]
apply(simp)
apply(rule ex1)
apply(rule THE_defaultI2)
apply(rule_tac p="-p" in permute_boolE)
apply(perm_simp add: permute_minus_cancel)
apply(rule ex1)
apply(perm_simp)
using eqvt[simplified eqvt_def]
apply(simp)
done
lemma fundef_ex1_eqvt_at:
fixes x::"'a::pt"
assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (G x))"
assumes eqvt: "eqvt G"
assumes ex1: "\<exists>!y. G x y"
shows "eqvt_at f x"
unfolding eqvt_at_def
using assms
by (auto intro: fundef_ex1_eqvt)
lemma fundef_ex1_prop:
fixes x::"'a::pt"
assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (G x))"
assumes P_all: "\<And>x y. G x y \<Longrightarrow> P x y"
assumes ex1: "\<exists>!y. G x y"
shows "P x (f x)"
unfolding f_def
using ex1
apply(erule_tac ex1E)
apply(rule THE_defaultI2)
apply(blast)
apply(rule P_all)
apply(assumption)
done
section \<open>Support of Finite Sets of Finitely Supported Elements\<close>
text \<open>support and freshness for atom sets\<close>
lemma supp_finite_atom_set:
fixes S::"atom set"
assumes "finite S"
shows "supp S = S"
apply(rule finite_supp_unique)
apply(simp add: supports_def)
apply(simp add: swap_set_not_in)
apply(rule assms)
apply(simp add: swap_set_in)
done
lemma supp_cofinite_atom_set:
fixes S::"atom set"
assumes "finite (UNIV - S)"
shows "supp S = (UNIV - S)"
apply(rule finite_supp_unique)
apply(simp add: supports_def)
apply(simp add: swap_set_both_in)
apply(rule assms)
apply(subst swap_commute)
apply(simp add: swap_set_in)
done
lemma fresh_finite_atom_set:
fixes S::"atom set"
assumes "finite S"
shows "a \<sharp> S \<longleftrightarrow> a \<notin> S"
unfolding fresh_def
by (simp add: supp_finite_atom_set[OF assms])
lemma fresh_minus_atom_set:
fixes S::"atom set"
assumes "finite S"
shows "a \<sharp> S - T \<longleftrightarrow> (a \<notin> T \<longrightarrow> a \<sharp> S)"
unfolding fresh_def
by (auto simp: supp_finite_atom_set assms)
lemma Union_supports_set:
shows "(\<Union>x \<in> S. supp x) supports S"
proof -
{ fix a b
have "\<forall>x \<in> S. (a \<rightleftharpoons> b) \<bullet> x = x \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> S = S"
unfolding permute_set_def by force
}
then show "(\<Union>x \<in> S. supp x) supports S"
unfolding supports_def
by (simp add: fresh_def[symmetric] swap_fresh_fresh)
qed
lemma Union_of_finite_supp_sets:
fixes S::"('a::fs set)"
assumes fin: "finite S"
shows "finite (\<Union>x\<in>S. supp x)"
using fin by (induct) (auto simp: finite_supp)
lemma Union_included_in_supp:
fixes S::"('a::fs set)"
assumes fin: "finite S"
shows "(\<Union>x\<in>S. supp x) \<subseteq> supp S"
proof -
have eqvt: "eqvt (\<lambda>S. \<Union>x \<in> S. supp x)"
unfolding eqvt_def by simp
have "(\<Union>x\<in>S. supp x) = supp (\<Union>x\<in>S. supp x)"
by (rule supp_finite_atom_set[symmetric]) (rule Union_of_finite_supp_sets[OF fin])
also have "\<dots> \<subseteq> supp S" using eqvt
by (rule supp_fun_app_eqvt)
finally show "(\<Union>x\<in>S. supp x) \<subseteq> supp S" .
qed
lemma supp_of_finite_sets:
fixes S::"('a::fs set)"
assumes fin: "finite S"
shows "(supp S) = (\<Union>x\<in>S. supp x)"
apply(rule subset_antisym)
apply(rule supp_is_subset)
apply(rule Union_supports_set)
apply(rule Union_of_finite_supp_sets[OF fin])
apply(rule Union_included_in_supp[OF fin])
done
lemma finite_sets_supp:
fixes S::"('a::fs set)"
assumes "finite S"
shows "finite (supp S)"
using assms
by (simp only: supp_of_finite_sets Union_of_finite_supp_sets)
lemma supp_of_finite_union:
fixes S T::"('a::fs) set"
assumes fin1: "finite S"
and fin2: "finite T"
shows "supp (S \<union> T) = supp S \<union> supp T"
using fin1 fin2
by (simp add: supp_of_finite_sets)
lemma fresh_finite_union:
fixes S T::"('a::fs) set"
assumes fin1: "finite S"
and fin2: "finite T"
shows "a \<sharp> (S \<union> T) \<longleftrightarrow> a \<sharp> S \<and> a \<sharp> T"
unfolding fresh_def
by (simp add: supp_of_finite_union[OF fin1 fin2])
lemma supp_of_finite_insert:
fixes S::"('a::fs) set"
assumes fin: "finite S"
shows "supp (insert x S) = supp x \<union> supp S"
using fin
by (simp add: supp_of_finite_sets)
lemma fresh_finite_insert:
fixes S::"('a::fs) set"
assumes fin: "finite S"
shows "a \<sharp> (insert x S) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> S"
using fin unfolding fresh_def
by (simp add: supp_of_finite_insert)
lemma supp_set_empty:
shows "supp {} = {}"
unfolding supp_def
by (simp add: empty_eqvt)
lemma fresh_set_empty:
shows "a \<sharp> {}"
by (simp add: fresh_def supp_set_empty)
lemma supp_set:
fixes xs :: "('a::fs) list"
shows "supp (set xs) = supp xs"
apply(induct xs)
apply(simp add: supp_set_empty supp_Nil)
apply(simp add: supp_Cons supp_of_finite_insert)
done
lemma fresh_set:
fixes xs :: "('a::fs) list"
shows "a \<sharp> (set xs) \<longleftrightarrow> a \<sharp> xs"
unfolding fresh_def
by (simp add: supp_set)
subsection \<open>Type @{typ "'a multiset"} is finitely supported\<close>
lemma set_mset_eqvt [eqvt]:
shows "p \<bullet> (set_mset M) = set_mset (p \<bullet> M)"
by (induct M) (simp_all add: insert_eqvt empty_eqvt)
lemma supp_set_mset:
shows "supp (set_mset M) \<subseteq> supp M"
apply (rule supp_fun_app_eqvt)
unfolding eqvt_def
apply(perm_simp)
apply(simp)
done
lemma Union_finite_multiset:
fixes M::"'a::fs multiset"
shows "finite (\<Union>{supp x | x. x \<in># M})"
proof -
have "finite (\<Union>(supp ` {x. x \<in># M}))"
by (induct M) (simp_all add: Collect_imp_eq Collect_neg_eq finite_supp)
then show "finite (\<Union>{supp x | x. x \<in># M})"
by (simp only: image_Collect)
qed
lemma Union_supports_multiset:
shows "\<Union>{supp x | x. x \<in># M} supports M"
proof -
have sw: "\<And>a b. ((\<And>x. x \<in># M \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x) \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> M = M)"
unfolding permute_multiset_def by (induct M) simp_all
have "(\<Union>x\<in>set_mset M. supp x) supports M"
by (auto intro!: sw swap_fresh_fresh simp add: fresh_def supports_def)
also have "(\<Union>x\<in>set_mset M. supp x) = (\<Union>{supp x | x. x \<in># M})"
by auto
finally show "(\<Union>{supp x | x. x \<in># M}) supports M" .
qed
lemma Union_included_multiset:
fixes M::"('a::fs multiset)"
shows "(\<Union>{supp x | x. x \<in># M}) \<subseteq> supp M"
proof -
have "(\<Union>{supp x | x. x \<in># M}) = (\<Union>x \<in> set_mset M. supp x)" by auto
also have "... = supp (set_mset M)"
by (simp add: supp_of_finite_sets)
also have " ... \<subseteq> supp M" by (rule supp_set_mset)
finally show "(\<Union>{supp x | x. x \<in># M}) \<subseteq> supp M" .
qed
lemma supp_of_multisets:
fixes M::"('a::fs multiset)"
shows "(supp M) = (\<Union>{supp x | x. x \<in># M})"
apply(rule subset_antisym)
apply(rule supp_is_subset)
apply(rule Union_supports_multiset)
apply(rule Union_finite_multiset)
apply(rule Union_included_multiset)
done
lemma multisets_supp_finite:
fixes M::"('a::fs multiset)"
shows "finite (supp M)"
by (simp only: supp_of_multisets Union_finite_multiset)
lemma supp_of_multiset_union:
fixes M N::"('a::fs) multiset"
shows "supp (M + N) = supp M \<union> supp N"
by (auto simp: supp_of_multisets)
lemma supp_empty_mset [simp]:
shows "supp {#} = {}"
unfolding supp_def
by simp
instance multiset :: (fs) fs
by standard (rule multisets_supp_finite)
subsection \<open>Type @{typ "'a fset"} is finitely supported\<close>
lemma supp_fset [simp]:
shows "supp (fset S) = supp S"
unfolding supp_def
by (simp add: fset_eqvt fset_cong)
lemma supp_empty_fset [simp]:
shows "supp {||} = {}"
unfolding supp_def
by simp
lemma fresh_empty_fset:
shows "a \<sharp> {||}"
unfolding fresh_def
by (simp)
lemma supp_finsert [simp]:
fixes x::"'a::fs"
and S::"'a fset"
shows "supp (finsert x S) = supp x \<union> supp S"
apply(subst supp_fset[symmetric])
apply(simp add: supp_of_finite_insert)
done
lemma fresh_finsert:
fixes x::"'a::fs"
and S::"'a fset"
shows "a \<sharp> finsert x S \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> S"
unfolding fresh_def
by simp
lemma fset_finite_supp:
fixes S::"('a::fs) fset"
shows "finite (supp S)"
by (induct S) (simp_all add: finite_supp)
lemma supp_union_fset:
fixes S T::"'a::fs fset"
shows "supp (S |\<union>| T) = supp S \<union> supp T"
by (induct S) (auto)
lemma fresh_union_fset:
fixes S T::"'a::fs fset"
shows "a \<sharp> S |\<union>| T \<longleftrightarrow> a \<sharp> S \<and> a \<sharp> T"
unfolding fresh_def
by (simp add: supp_union_fset)
instance fset :: (fs) fs
by standard (rule fset_finite_supp)
subsection \<open>Type @{typ "('a, 'b) finfun"} is finitely supported\<close>
lemma fresh_finfun_const:
shows "a \<sharp> (finfun_const b) \<longleftrightarrow> a \<sharp> b"
by (simp add: fresh_def supp_def)
lemma fresh_finfun_update:
shows "\<lbrakk>a \<sharp> f; a \<sharp> x; a \<sharp> y\<rbrakk> \<Longrightarrow> a \<sharp> finfun_update f x y"
unfolding fresh_conv_MOST
unfolding finfun_update_eqvt
by (elim MOST_rev_mp) (simp)
lemma supp_finfun_const:
shows "supp (finfun_const b) = supp(b)"
by (simp add: supp_def)
lemma supp_finfun_update:
shows "supp (finfun_update f x y) \<subseteq> supp(f, x, y)"
using fresh_finfun_update
by (auto simp: fresh_def supp_Pair)
instance finfun :: (fs, fs) fs
apply standard
apply(induct_tac x rule: finfun_weak_induct)
apply(simp add: supp_finfun_const finite_supp)
apply(rule finite_subset)
apply(rule supp_finfun_update)
apply(simp add: supp_Pair finite_supp)
done
section \<open>Freshness and Fresh-Star\<close>
lemma fresh_Unit_elim:
shows "(a \<sharp> () \<Longrightarrow> PROP C) \<equiv> PROP C"
by (simp add: fresh_Unit)
lemma fresh_Pair_elim:
shows "(a \<sharp> (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp> x \<Longrightarrow> a \<sharp> y \<Longrightarrow> PROP C)"
by rule (simp_all add: fresh_Pair)
(* this rule needs to be added before the fresh_prodD is *)
(* added to the simplifier with mksimps *)
lemma [simp]:
shows "a \<sharp> x1 \<Longrightarrow> a \<sharp> x2 \<Longrightarrow> a \<sharp> (x1, x2)"
by (simp add: fresh_Pair)
lemma fresh_PairD:
shows "a \<sharp> (x, y) \<Longrightarrow> a \<sharp> x"
and "a \<sharp> (x, y) \<Longrightarrow> a \<sharp> y"
by (simp_all add: fresh_Pair)
declaration \<open>fn _ =>
let
val mksimps_pairs = (@{const_name Nominal2_Base.fresh}, @{thms fresh_PairD}) :: mksimps_pairs
in
Simplifier.map_ss (fn ss => Simplifier.set_mksimps (mksimps mksimps_pairs) ss)
end
\<close>
text \<open>The fresh-star generalisation of fresh is used in strong
induction principles.\<close>
definition
fresh_star :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp>* _" [80,80] 80)
where
"as \<sharp>* x \<equiv> \<forall>a \<in> as. a \<sharp> x"
lemma fresh_star_supp_conv:
shows "supp x \<sharp>* y \<Longrightarrow> supp y \<sharp>* x"
by (auto simp: fresh_star_def fresh_def)
lemma fresh_star_perm_set_conv:
fixes p::"perm"
assumes fresh: "as \<sharp>* p"
and fin: "finite as"
shows "supp p \<sharp>* as"
apply(rule fresh_star_supp_conv)
apply(simp add: supp_finite_atom_set fin fresh)
done
lemma fresh_star_atom_set_conv:
assumes fresh: "as \<sharp>* bs"
and fin: "finite as" "finite bs"
shows "bs \<sharp>* as"
using fresh
unfolding fresh_star_def fresh_def
by (auto simp: supp_finite_atom_set fin)
lemma atom_fresh_star_disjoint:
assumes fin: "finite bs"
shows "as \<sharp>* bs \<longleftrightarrow> (as \<inter> bs = {})"
unfolding fresh_star_def fresh_def
by (auto simp: supp_finite_atom_set fin)
lemma fresh_star_Pair:
shows "as \<sharp>* (x, y) = (as \<sharp>* x \<and> as \<sharp>* y)"
by (auto simp: fresh_star_def fresh_Pair)
lemma fresh_star_list:
shows "as \<sharp>* (xs @ ys) \<longleftrightarrow> as \<sharp>* xs \<and> as \<sharp>* ys"
and "as \<sharp>* (x # xs) \<longleftrightarrow> as \<sharp>* x \<and> as \<sharp>* xs"
and "as \<sharp>* []"
by (auto simp: fresh_star_def fresh_Nil fresh_Cons fresh_append)
lemma fresh_star_set:
fixes xs::"('a::fs) list"
shows "as \<sharp>* set xs \<longleftrightarrow> as \<sharp>* xs"
unfolding fresh_star_def
by (simp add: fresh_set)
lemma fresh_star_singleton:
fixes a::"atom"
shows "as \<sharp>* {a} \<longleftrightarrow> as \<sharp>* a"
by (simp add: fresh_star_def fresh_finite_insert fresh_set_empty)
lemma fresh_star_fset:
fixes xs::"('a::fs) list"
shows "as \<sharp>* fset S \<longleftrightarrow> as \<sharp>* S"
by (simp add: fresh_star_def fresh_def)
lemma fresh_star_Un:
shows "(as \<union> bs) \<sharp>* x = (as \<sharp>* x \<and> bs \<sharp>* x)"
by (auto simp: fresh_star_def)
lemma fresh_star_insert:
shows "(insert a as) \<sharp>* x = (a \<sharp> x \<and> as \<sharp>* x)"
by (auto simp: fresh_star_def)
lemma fresh_star_Un_elim:
"((as \<union> bs) \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (as \<sharp>* x \<Longrightarrow> bs \<sharp>* x \<Longrightarrow> PROP C)"
unfolding fresh_star_def
apply(rule)
apply(erule meta_mp)
apply(auto)
done
lemma fresh_star_insert_elim:
"(insert a as \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (a \<sharp> x \<Longrightarrow> as \<sharp>* x \<Longrightarrow> PROP C)"
unfolding fresh_star_def
by rule (simp_all add: fresh_star_def)
lemma fresh_star_empty_elim:
"({} \<sharp>* x \<Longrightarrow> PROP C) \<equiv> PROP C"
by (simp add: fresh_star_def)
lemma fresh_star_Unit_elim:
shows "(a \<sharp>* () \<Longrightarrow> PROP C) \<equiv> PROP C"
by (simp add: fresh_star_def fresh_Unit)
lemma fresh_star_Pair_elim:
shows "(a \<sharp>* (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp>* x \<Longrightarrow> a \<sharp>* y \<Longrightarrow> PROP C)"
by (rule, simp_all add: fresh_star_Pair)
lemma fresh_star_zero:
shows "as \<sharp>* (0::perm)"
unfolding fresh_star_def
by (simp add: fresh_zero_perm)
lemma fresh_star_plus:
fixes p q::perm
shows "\<lbrakk>a \<sharp>* p; a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)"
unfolding fresh_star_def
by (simp add: fresh_plus_perm)
lemma fresh_star_permute_iff:
shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x"
unfolding fresh_star_def
by (metis mem_permute_iff permute_minus_cancel(1) fresh_permute_iff)
lemma fresh_star_eqvt [eqvt]:
shows "p \<bullet> (as \<sharp>* x) \<longleftrightarrow> (p \<bullet> as) \<sharp>* (p \<bullet> x)"
unfolding fresh_star_def by simp
section \<open>Induction principle for permutations\<close>
lemma smaller_supp:
assumes a: "a \<in> supp p"
shows "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subset> supp p"
proof -
have "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subseteq> supp p"
unfolding supp_perm by (auto simp: swap_atom)
moreover
have "a \<notin> supp ((p \<bullet> a \<rightleftharpoons> a) + p)" by (simp add: supp_perm)
then have "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<noteq> supp p" using a by auto
ultimately
show "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subset> supp p" by auto
qed
lemma perm_struct_induct[consumes 1, case_names zero swap]:
assumes S: "supp p \<subseteq> S"
and zero: "P 0"
and swap: "\<And>p a b. \<lbrakk>P p; supp p \<subseteq> S; a \<in> S; b \<in> S; a \<noteq> b; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> P ((a \<rightleftharpoons> b) + p)"
shows "P p"
proof -
have "finite (supp p)" by (simp add: finite_supp)
then show "P p" using S
proof(induct A\<equiv>"supp p" arbitrary: p rule: finite_psubset_induct)
case (psubset p)
then have ih: "\<And>q. supp q \<subset> supp p \<Longrightarrow> P q" by auto
have as: "supp p \<subseteq> S" by fact
{ assume "supp p = {}"
then have "p = 0" by (simp add: supp_perm perm_eq_iff)
then have "P p" using zero by simp
}
moreover
{ assume "supp p \<noteq> {}"
then obtain a where a0: "a \<in> supp p" by blast
then have a1: "p \<bullet> a \<in> S" "a \<in> S" "sort_of (p \<bullet> a) = sort_of a" "p \<bullet> a \<noteq> a"
using as by (auto simp: supp_atom supp_perm swap_atom)
let ?q = "(p \<bullet> a \<rightleftharpoons> a) + p"
have a2: "supp ?q \<subset> supp p" using a0 smaller_supp by simp
then have "P ?q" using ih by simp
moreover
have "supp ?q \<subseteq> S" using as a2 by simp
ultimately have "P ((p \<bullet> a \<rightleftharpoons> a) + ?q)" using as a1 swap by simp
moreover
have "p = (p \<bullet> a \<rightleftharpoons> a) + ?q" by (simp add: perm_eq_iff)
ultimately have "P p" by simp
}
ultimately show "P p" by blast
qed
qed
lemma perm_simple_struct_induct[case_names zero swap]:
assumes zero: "P 0"
and swap: "\<And>p a b. \<lbrakk>P p; a \<noteq> b; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> P ((a \<rightleftharpoons> b) + p)"
shows "P p"
by (rule_tac S="supp p" in perm_struct_induct)
(auto intro: zero swap)
lemma perm_struct_induct2[consumes 1, case_names zero swap plus]:
assumes S: "supp p \<subseteq> S"
assumes zero: "P 0"
assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b; a \<in> S; b \<in> S\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)"
assumes plus: "\<And>p1 p2. \<lbrakk>P p1; P p2; supp p1 \<subseteq> S; supp p2 \<subseteq> S\<rbrakk> \<Longrightarrow> P (p1 + p2)"
shows "P p"
using S
by (induct p rule: perm_struct_induct)
(auto intro: zero plus swap simp add: supp_swap)
lemma perm_simple_struct_induct2[case_names zero swap plus]:
assumes zero: "P 0"
assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)"
assumes plus: "\<And>p1 p2. \<lbrakk>P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)"
shows "P p"
by (rule_tac S="supp p" in perm_struct_induct2)
(auto intro: zero swap plus)
lemma supp_perm_singleton:
fixes p::"perm"
shows "supp p \<subseteq> {b} \<longleftrightarrow> p = 0"
proof -
{ assume "supp p \<subseteq> {b}"
then have "p = 0"
by (induct p rule: perm_struct_induct) (simp_all)
}
then show "supp p \<subseteq> {b} \<longleftrightarrow> p = 0" by (auto simp: supp_zero_perm)
qed
lemma supp_perm_pair:
fixes p::"perm"
shows "supp p \<subseteq> {a, b} \<longleftrightarrow> p = 0 \<or> p = (b \<rightleftharpoons> a)"
proof -
{ assume "supp p \<subseteq> {a, b}"
then have "p = 0 \<or> p = (b \<rightleftharpoons> a)"
apply (induct p rule: perm_struct_induct)
apply (auto simp: swap_cancel supp_zero_perm supp_swap)
apply (simp add: swap_commute)
done
}
then show "supp p \<subseteq> {a, b} \<longleftrightarrow> p = 0 \<or> p = (b \<rightleftharpoons> a)"
by (auto simp: supp_zero_perm supp_swap split: if_splits)
qed
lemma supp_perm_eq:
assumes "(supp x) \<sharp>* p"
shows "p \<bullet> x = x"
proof -
from assms have "supp p \<subseteq> {a. a \<sharp> x}"
unfolding supp_perm fresh_star_def fresh_def by auto
then show "p \<bullet> x = x"
proof (induct p rule: perm_struct_induct)
case zero
show "0 \<bullet> x = x" by simp
next
case (swap p a b)
then have "a \<sharp> x" "b \<sharp> x" "p \<bullet> x = x" by simp_all
then show "((a \<rightleftharpoons> b) + p) \<bullet> x = x" by (simp add: swap_fresh_fresh)
qed
qed
text \<open>same lemma as above, but proved with a different induction principle\<close>
lemma supp_perm_eq_test:
assumes "(supp x) \<sharp>* p"
shows "p \<bullet> x = x"
proof -
from assms have "supp p \<subseteq> {a. a \<sharp> x}"
unfolding supp_perm fresh_star_def fresh_def by auto
then show "p \<bullet> x = x"
proof (induct p rule: perm_struct_induct2)
case zero
show "0 \<bullet> x = x" by simp
next
case (swap a b)
then have "a \<sharp> x" "b \<sharp> x" by simp_all
then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh)
next
case (plus p1 p2)
have "p1 \<bullet> x = x" "p2 \<bullet> x = x" by fact+
then show "(p1 + p2) \<bullet> x = x" by simp
qed
qed
lemma perm_supp_eq:
assumes a: "(supp p) \<sharp>* x"
shows "p \<bullet> x = x"
proof -
from assms have "supp p \<subseteq> {a. a \<sharp> x}"
unfolding supp_perm fresh_star_def fresh_def by auto
then show "p \<bullet> x = x"
proof (induct p rule: perm_struct_induct2)
case zero
show "0 \<bullet> x = x" by simp
next
case (swap a b)
then have "a \<sharp> x" "b \<sharp> x" by simp_all
then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh)
next
case (plus p1 p2)
have "p1 \<bullet> x = x" "p2 \<bullet> x = x" by fact+
then show "(p1 + p2) \<bullet> x = x" by simp
qed
qed
lemma supp_perm_perm_eq:
assumes a: "\<forall>a \<in> supp x. p \<bullet> a = q \<bullet> a"
shows "p \<bullet> x = q \<bullet> x"
proof -
from a have "\<forall>a \<in> supp x. (-q + p) \<bullet> a = a" by simp
then have "\<forall>a \<in> supp x. a \<notin> supp (-q + p)"
unfolding supp_perm by simp
then have "supp x \<sharp>* (-q + p)"
unfolding fresh_star_def fresh_def by simp
then have "(-q + p) \<bullet> x = x" by (simp only: supp_perm_eq)
then show "p \<bullet> x = q \<bullet> x"
by (metis permute_minus_cancel permute_plus)
qed
text \<open>disagreement set\<close>
definition
dset :: "perm \<Rightarrow> perm \<Rightarrow> atom set"
where
"dset p q = {a::atom. p \<bullet> a \<noteq> q \<bullet> a}"
lemma ds_fresh:
assumes "dset p q \<sharp>* x"
shows "p \<bullet> x = q \<bullet> x"
using assms
unfolding dset_def fresh_star_def fresh_def
by (auto intro: supp_perm_perm_eq)
lemma atom_set_perm_eq:
assumes a: "as \<sharp>* p"
shows "p \<bullet> as = as"
proof -
from a have "supp p \<subseteq> {a. a \<notin> as}"
unfolding supp_perm fresh_star_def fresh_def by auto
then show "p \<bullet> as = as"
proof (induct p rule: perm_struct_induct)
case zero
show "0 \<bullet> as = as" by simp
next
case (swap p a b)
then have "a \<notin> as" "b \<notin> as" "p \<bullet> as = as" by simp_all
then show "((a \<rightleftharpoons> b) + p) \<bullet> as = as" by (simp add: swap_set_not_in)
qed
qed
section \<open>Avoiding of atom sets\<close>
text \<open>
For every set of atoms, there is another set of atoms
avoiding a finitely supported c and there is a permutation
which 'translates' between both sets.
\<close>
lemma at_set_avoiding_aux:
fixes Xs::"atom set"
and As::"atom set"
assumes b: "Xs \<subseteq> As"
and c: "finite As"
shows "\<exists>p. (p \<bullet> Xs) \<inter> As = {} \<and> (supp p) = (Xs \<union> (p \<bullet> Xs))"
proof -
from b c have "finite Xs" by (rule finite_subset)
then show ?thesis using b
proof (induct rule: finite_subset_induct)
case empty
have "0 \<bullet> {} \<inter> As = {}" by simp
moreover
have "supp (0::perm) = {} \<union> 0 \<bullet> {}" by (simp add: supp_zero_perm)
ultimately show ?case by blast
next
case (insert x Xs)
then obtain p where
p1: "(p \<bullet> Xs) \<inter> As = {}" and
p2: "supp p = (Xs \<union> (p \<bullet> Xs))" by blast
from \<open>x \<in> As\<close> p1 have "x \<notin> p \<bullet> Xs" by fast
with \<open>x \<notin> Xs\<close> p2 have "x \<notin> supp p" by fast
hence px: "p \<bullet> x = x" unfolding supp_perm by simp
have "finite (As \<union> p \<bullet> Xs \<union> supp p)"
using \<open>finite As\<close> \<open>finite Xs\<close>
by (simp add: permute_set_eq_image finite_supp)
then obtain y where "y \<notin> (As \<union> p \<bullet> Xs \<union> supp p)" "sort_of y = sort_of x"
by (rule obtain_atom)
hence y: "y \<notin> As" "y \<notin> p \<bullet> Xs" "y \<notin> supp p" "sort_of y = sort_of x"
by simp_all
hence py: "p \<bullet> y = y" "x \<noteq> y" using \<open>x \<in> As\<close>
by (auto simp: supp_perm)
let ?q = "(x \<rightleftharpoons> y) + p"
have q: "?q \<bullet> insert x Xs = insert y (p \<bullet> Xs)"
unfolding insert_eqvt
using \<open>p \<bullet> x = x\<close> \<open>sort_of y = sort_of x\<close>
using \<open>x \<notin> p \<bullet> Xs\<close> \<open>y \<notin> p \<bullet> Xs\<close>
by (simp add: swap_atom swap_set_not_in)
have "?q \<bullet> insert x Xs \<inter> As = {}"
using \<open>y \<notin> As\<close> \<open>p \<bullet> Xs \<inter> As = {}\<close>
unfolding q by simp
moreover
have "supp (x \<rightleftharpoons> y) \<inter> supp p = {}" using px py \<open>sort_of y = sort_of x\<close>
unfolding supp_swap by (simp add: supp_perm)
then have "supp ?q = (supp (x \<rightleftharpoons> y) \<union> supp p)"
by (simp add: supp_plus_perm_eq)
then have "supp ?q = insert x Xs \<union> ?q \<bullet> insert x Xs"
using p2 \<open>sort_of y = sort_of x\<close> \<open>x \<noteq> y\<close> unfolding q supp_swap
by auto
ultimately show ?case by blast
qed
qed
lemma at_set_avoiding:
assumes a: "finite Xs"
and b: "finite (supp c)"
obtains p::"perm" where "(p \<bullet> Xs)\<sharp>*c" and "(supp p) = (Xs \<union> (p \<bullet> Xs))"
using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \<union> supp c"]
unfolding fresh_star_def fresh_def by blast
lemma at_set_avoiding1:
assumes "finite xs"
and "finite (supp c)"
shows "\<exists>p. (p \<bullet> xs) \<sharp>* c"
using assms
apply(erule_tac c="c" in at_set_avoiding)
apply(auto)
done
lemma at_set_avoiding2:
assumes "finite xs"
and "finite (supp c)" "finite (supp x)"
and "xs \<sharp>* x"
shows "\<exists>p. (p \<bullet> xs) \<sharp>* c \<and> supp x \<sharp>* p"
using assms
apply(erule_tac c="(c, x)" in at_set_avoiding)
apply(simp add: supp_Pair)
apply(rule_tac x="p" in exI)
apply(simp add: fresh_star_Pair)
apply(rule fresh_star_supp_conv)
apply(auto simp: fresh_star_def)
done
lemma at_set_avoiding3:
assumes "finite xs"
and "finite (supp c)" "finite (supp x)"
and "xs \<sharp>* x"
shows "\<exists>p. (p \<bullet> xs) \<sharp>* c \<and> supp x \<sharp>* p \<and> supp p = xs \<union> (p \<bullet> xs)"
using assms
apply(erule_tac c="(c, x)" in at_set_avoiding)
apply(simp add: supp_Pair)
apply(rule_tac x="p" in exI)
apply(simp add: fresh_star_Pair)
apply(rule fresh_star_supp_conv)
apply(auto simp: fresh_star_def)
done
lemma at_set_avoiding2_atom:
assumes "finite (supp c)" "finite (supp x)"
and b: "a \<sharp> x"
shows "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p"
proof -
have a: "{a} \<sharp>* x" unfolding fresh_star_def by (simp add: b)
obtain p where p1: "(p \<bullet> {a}) \<sharp>* c" and p2: "supp x \<sharp>* p"
using at_set_avoiding2[of "{a}" "c" "x"] assms a by blast
have c: "(p \<bullet> a) \<sharp> c" using p1
unfolding fresh_star_def Ball_def
by(erule_tac x="p \<bullet> a" in allE) (simp add: permute_set_def)
hence "p \<bullet> a \<sharp> c \<and> supp x \<sharp>* p" using p2 by blast
then show "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p" by blast
qed
section \<open>Renaming permutations\<close>
lemma set_renaming_perm:
assumes b: "finite bs"
shows "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)"
using b
proof (induct)
case empty
have "(\<forall>b \<in> {}. 0 \<bullet> b = p \<bullet> b) \<and> supp (0::perm) \<subseteq> {} \<union> p \<bullet> {}"
by (simp add: permute_set_def supp_perm)
then show "\<exists>q. (\<forall>b \<in> {}. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> {} \<union> p \<bullet> {}" by blast
next
case (insert a bs)
then have " \<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> p \<bullet> bs" by simp
then obtain q where *: "\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> bs \<union> p \<bullet> bs"
by (metis empty_subsetI insert(3) supp_swap)
{ assume 1: "q \<bullet> a = p \<bullet> a"
have "\<forall>b \<in> (insert a bs). q \<bullet> b = p \<bullet> b" using 1 * by simp
moreover
have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
using ** by (auto simp: insert_eqvt)
ultimately
have "\<exists>q. (\<forall>b \<in> insert a bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by blast
}
moreover
{ assume 2: "q \<bullet> a \<noteq> p \<bullet> a"
define q' where "q' = ((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q"
have "\<forall>b \<in> insert a bs. q' \<bullet> b = p \<bullet> b" using 2 * \<open>a \<notin> bs\<close> unfolding q'_def
by (auto simp: swap_atom)
moreover
{ have "{q \<bullet> a, p \<bullet> a} \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
using **
apply (auto simp: supp_perm insert_eqvt)
apply (subgoal_tac "q \<bullet> a \<in> bs \<union> p \<bullet> bs")
apply(auto)[1]
apply(subgoal_tac "q \<bullet> a \<in> {a. q \<bullet> a \<noteq> a}")
apply(blast)
apply(simp)
done
then have "supp (q \<bullet> a \<rightleftharpoons> p \<bullet> a) \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
unfolding supp_swap by auto
moreover
have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
using ** by (auto simp: insert_eqvt)
ultimately
have "supp q' \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
unfolding q'_def using supp_plus_perm by blast
}
ultimately
have "\<exists>q. (\<forall>b \<in> insert a bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by blast
}
ultimately show "\<exists>q. (\<forall>b \<in> insert a bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
by blast
qed
lemma set_renaming_perm2:
shows "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)"
proof -
have "finite (bs \<inter> supp p)" by (simp add: finite_supp)
then obtain q
where *: "\<forall>b \<in> bs \<inter> supp p. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> (bs \<inter> supp p) \<union> (p \<bullet> (bs \<inter> supp p))"
using set_renaming_perm by blast
from ** have "supp q \<subseteq> bs \<union> (p \<bullet> bs)" by (auto simp: inter_eqvt)
moreover
have "\<forall>b \<in> bs - supp p. q \<bullet> b = p \<bullet> b"
apply(auto)
apply(subgoal_tac "b \<notin> supp q")
apply(simp add: fresh_def[symmetric])
apply(simp add: fresh_perm)
apply(clarify)
apply(rotate_tac 2)
apply(drule subsetD[OF **])
apply(simp add: inter_eqvt supp_eqvt permute_self)
done
ultimately have "(\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" using * by auto
then show "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" by blast
qed
lemma list_renaming_perm:
shows "\<exists>q. (\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set bs \<union> (p \<bullet> set bs)"
proof (induct bs)
case (Cons a bs)
then have " \<exists>q. (\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set bs \<union> p \<bullet> (set bs)" by simp
then obtain q where *: "\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> set bs \<union> p \<bullet> (set bs)"
by (blast)
{ assume 1: "a \<in> set bs"
have "q \<bullet> a = p \<bullet> a" using * 1 by (induct bs) (auto)
then have "\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b" using * by simp
moreover
have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" using ** by (auto simp: insert_eqvt)
ultimately
have "\<exists>q. (\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" by blast
}
moreover
{ assume 2: "a \<notin> set bs"
define q' where "q' = ((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q"
have "\<forall>b \<in> set (a # bs). q' \<bullet> b = p \<bullet> b"
unfolding q'_def using 2 * \<open>a \<notin> set bs\<close> by (auto simp: swap_atom)
moreover
{ have "{q \<bullet> a, p \<bullet> a} \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))"
using **
apply (auto simp: supp_perm insert_eqvt)
apply (subgoal_tac "q \<bullet> a \<in> set bs \<union> p \<bullet> set bs")
apply(auto)[1]
apply(subgoal_tac "q \<bullet> a \<in> {a. q \<bullet> a \<noteq> a}")
apply(blast)
apply(simp)
done
then have "supp (q \<bullet> a \<rightleftharpoons> p \<bullet> a) \<subseteq> set (a # bs) \<union> p \<bullet> set (a # bs)"
unfolding supp_swap by auto
moreover
have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))"
using ** by (auto simp: insert_eqvt)
ultimately
have "supp q' \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))"
unfolding q'_def using supp_plus_perm by blast
}
ultimately
have "\<exists>q. (\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" by blast
}
ultimately show "\<exists>q. (\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))"
by blast
next
case Nil
have "(\<forall>b \<in> set []. 0 \<bullet> b = p \<bullet> b) \<and> supp (0::perm) \<subseteq> set [] \<union> p \<bullet> set []"
by (simp add: supp_zero_perm)
then show "\<exists>q. (\<forall>b \<in> set []. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set [] \<union> p \<bullet> (set [])" by blast
qed
section \<open>Concrete Atoms Types\<close>
text \<open>
Class \<open>at_base\<close> allows types containing multiple sorts of atoms.
Class \<open>at\<close> only allows types with a single sort.
\<close>
class at_base = pt +
fixes atom :: "'a \<Rightarrow> atom"
assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b"
assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)"
declare atom_eqvt [eqvt]
class at = at_base +
assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)"
lemma sort_ineq [simp]:
assumes "sort_of (atom a) \<noteq> sort_of (atom b)"
shows "atom a \<noteq> atom b"
using assms by metis
lemma supp_at_base:
fixes a::"'a::at_base"
shows "supp a = {atom a}"
by (simp add: supp_atom [symmetric] supp_def atom_eqvt)
lemma fresh_at_base:
shows "sort_of a \<noteq> sort_of (atom b) \<Longrightarrow> a \<sharp> b"
and "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b"
unfolding fresh_def
apply(simp_all add: supp_at_base)
apply(metis)
done
(* solves the freshness only if the inequality can be shown by the
simproc below *)
lemma fresh_ineq_at_base [simp]:
shows "a \<noteq> atom b \<Longrightarrow> a \<sharp> b"
by (simp add: fresh_at_base)
lemma fresh_atom_at_base [simp]:
fixes b::"'a::at_base"
shows "a \<sharp> atom b \<longleftrightarrow> a \<sharp> b"
by (simp add: fresh_def supp_at_base supp_atom)
lemma fresh_star_atom_at_base:
fixes b::"'a::at_base"
shows "as \<sharp>* atom b \<longleftrightarrow> as \<sharp>* b"
by (simp add: fresh_star_def fresh_atom_at_base)
lemma if_fresh_at_base [simp]:
shows "atom a \<sharp> x \<Longrightarrow> P (if a = x then t else s) = P s"
and "atom a \<sharp> x \<Longrightarrow> P (if x = a then t else s) = P s"
by (simp_all add: fresh_at_base)
simproc_setup fresh_ineq ("x \<noteq> (y::'a::at_base)") = \<open>fn _ => fn ctxt => fn ctrm =>
case Thm.term_of ctrm of \<^Const_>\<open>Not for \<^Const_>\<open>HOL.eq _ for lhs rhs\<close>\<close> =>
let
fun first_is_neg lhs rhs [] = NONE
| first_is_neg lhs rhs (thm::thms) =
(case Thm.prop_of thm of
_ $ \<^Const_>\<open>Not for \<^Const_>\<open>HOL.eq _ for l r\<close>\<close> =>
(if l = lhs andalso r = rhs then SOME(thm)
else if r = lhs andalso l = rhs then SOME(thm RS @{thm not_sym})
else first_is_neg lhs rhs thms)
| _ => first_is_neg lhs rhs thms)
val simp_thms = @{thms fresh_Pair fresh_at_base atom_eq_iff}
val prems = Simplifier.prems_of ctxt
|> filter (fn thm => case Thm.prop_of thm of
_ $ \<^Const_>\<open>fresh _ for \<open>_ $ a\<close> b\<close> =>
(let
val atms = a :: HOLogic.strip_tuple b
in
member (op =) atms lhs andalso member (op =) atms rhs
end)
| _ => false)
|> map (simplify (put_simpset HOL_basic_ss ctxt addsimps simp_thms))
|> map (HOLogic.conj_elims ctxt)
|> flat
in
case first_is_neg lhs rhs prems of
SOME(thm) => SOME(thm RS @{thm Eq_TrueI})
| NONE => NONE
end
| _ => NONE
\<close>
instance at_base < fs
proof qed (simp add: supp_at_base)
lemma at_base_infinite [simp]:
shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U")
proof
obtain a :: 'a where "True" by auto
assume "finite ?U"
hence "finite (atom ` ?U)"
by (rule finite_imageI)
then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)"
by (rule obtain_atom)
from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)"
unfolding atom_eqvt [symmetric]
by (simp add: swap_atom)
hence "b \<in> atom ` ?U" by simp
with b(1) show "False" by simp
qed
lemma swap_at_base_simps [simp]:
fixes x y::"'a::at_base"
shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y"
and "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x"
and "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x"
unfolding atom_eq_iff [symmetric]
unfolding atom_eqvt [symmetric]
by simp_all
lemma obtain_at_base:
assumes X: "finite X"
obtains a::"'a::at_base" where "atom a \<notin> X"
proof -
have "inj (atom :: 'a \<Rightarrow> atom)"
by (simp add: inj_on_def)
with X have "finite (atom -` X :: 'a set)"
by (rule finite_vimageI)
with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)"
by auto
then obtain a :: 'a where "atom a \<notin> X"
by auto
thus ?thesis ..
qed
lemma obtain_fresh':
assumes fin: "finite (supp x)"
obtains a::"'a::at_base" where "atom a \<sharp> x"
using obtain_at_base[where X="supp x"]
by (auto simp: fresh_def fin)
lemma obtain_fresh:
fixes x::"'b::fs"
obtains a::"'a::at_base" where "atom a \<sharp> x"
by (rule obtain_fresh') (auto simp: finite_supp)
lemma supp_finite_set_at_base:
assumes a: "finite S"
shows "supp S = atom ` S"
apply(simp add: supp_of_finite_sets[OF a])
apply(simp add: supp_at_base)
apply(auto)
done
(* FIXME
lemma supp_cofinite_set_at_base:
assumes a: "finite (UNIV - S)"
shows "supp S = atom ` (UNIV - S)"
apply(rule finite_supp_unique)
*)
lemma fresh_finite_set_at_base:
fixes a::"'a::at_base"
assumes a: "finite S"
shows "atom a \<sharp> S \<longleftrightarrow> a \<notin> S"
unfolding fresh_def
apply(simp add: supp_finite_set_at_base[OF a])
apply(subst inj_image_mem_iff)
apply(simp add: inj_on_def)
apply(simp)
done
lemma fresh_at_base_permute_iff [simp]:
fixes a::"'a::at_base"
shows "atom (p \<bullet> a) \<sharp> p \<bullet> x \<longleftrightarrow> atom a \<sharp> x"
unfolding atom_eqvt[symmetric]
by (simp only: fresh_permute_iff)
lemma fresh_at_base_permI:
shows "atom a \<sharp> p \<Longrightarrow> p \<bullet> a = a"
by (simp add: fresh_def supp_perm)
section \<open>Infrastructure for concrete atom types\<close>
definition
flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
where
"(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)"
lemma flip_fresh_fresh:
assumes "atom a \<sharp> x" "atom b \<sharp> x"
shows "(a \<leftrightarrow> b) \<bullet> x = x"
using assms
by (simp add: flip_def swap_fresh_fresh)
lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0"
unfolding flip_def by (rule swap_self)
lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)"
unfolding flip_def by (rule swap_commute)
lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)"
unfolding flip_def by (rule minus_swap)
lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0"
unfolding flip_def by (rule swap_cancel)
lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x"
unfolding permute_plus [symmetric] add_flip_cancel by simp
lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x"
by (simp add: flip_commute)
lemma flip_eqvt [eqvt]:
shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)"
unfolding flip_def
by (simp add: swap_eqvt atom_eqvt)
lemma flip_at_base_simps [simp]:
shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b"
and "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a"
and "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c"
and "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
unfolding flip_def
unfolding atom_eq_iff [symmetric]
unfolding atom_eqvt [symmetric]
by simp_all
text \<open>the following two lemmas do not hold for \<open>at_base\<close>,
only for single sort atoms from at\<close>
lemma flip_triple:
fixes a b c::"'a::at"
assumes "a \<noteq> b" and "c \<noteq> b"
shows "(a \<leftrightarrow> c) + (b \<leftrightarrow> c) + (a \<leftrightarrow> c) = (a \<leftrightarrow> b)"
unfolding flip_def
by (rule swap_triple) (simp_all add: assms)
lemma permute_flip_at:
fixes a b c::"'a::at"
shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)"
unfolding flip_def
apply (rule atom_eq_iff [THEN iffD1])
apply (subst atom_eqvt [symmetric])
apply (simp add: swap_atom)
done
lemma flip_at_simps [simp]:
fixes a b::"'a::at"
shows "(a \<leftrightarrow> b) \<bullet> a = b"
and "(a \<leftrightarrow> b) \<bullet> b = a"
unfolding permute_flip_at by simp_all
subsection \<open>Syntax for coercing at-elements to the atom-type\<close>
syntax
"_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
translations
"_atom_constrain a t" => "CONST atom (_constrain a t)"
subsection \<open>A lemma for proving instances of class \<open>at\<close>.\<close>
setup \<open>Sign.add_const_constraint (@{const_name "permute"}, NONE)\<close>
setup \<open>Sign.add_const_constraint (@{const_name "atom"}, NONE)\<close>
text \<open>
New atom types are defined as subtypes of \<^typ>\<open>atom\<close>.
\<close>
lemma exists_eq_simple_sort:
shows "\<exists>a. a \<in> {a. sort_of a = s}"
by (rule_tac x="Atom s 0" in exI, simp)
lemma exists_eq_sort:
shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
by (rule_tac x="Atom (sort_fun x) y" in exI, simp)
lemma at_base_class:
fixes sort_fun :: "'b \<Rightarrow> atom_sort"
fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
assumes atom_def: "\<And>a. atom a = Rep a"
assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
shows "OFCLASS('a, at_base_class)"
proof
interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp
fix a b :: 'a and p p1 p2 :: perm
show "0 \<bullet> a = a"
unfolding permute_def by (simp add: Rep_inverse)
show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
show "atom a = atom b \<longleftrightarrow> a = b"
unfolding atom_def by (simp add: Rep_inject)
show "p \<bullet> atom a = atom (p \<bullet> a)"
unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
qed
(*
lemma at_class:
fixes s :: atom_sort
fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
assumes atom_def: "\<And>a. atom a = Rep a"
assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
shows "OFCLASS('a, at_class)"
proof
interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
fix a b :: 'a and p p1 p2 :: perm
show "0 \<bullet> a = a"
unfolding permute_def by (simp add: Rep_inverse)
show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
show "sort_of (atom a) = sort_of (atom b)"
unfolding atom_def by (simp add: sort_of_Rep)
show "atom a = atom b \<longleftrightarrow> a = b"
unfolding atom_def by (simp add: Rep_inject)
show "p \<bullet> atom a = atom (p \<bullet> a)"
unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
qed
*)
lemma at_class:
fixes s :: atom_sort
fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
assumes type: "type_definition Rep Abs {a. sort_of a = s}"
assumes atom_def: "\<And>a. atom a = Rep a"
assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
shows "OFCLASS('a, at_class)"
proof
interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
fix a b :: 'a and p p1 p2 :: perm
show "0 \<bullet> a = a"
unfolding permute_def by (simp add: Rep_inverse)
show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
show "sort_of (atom a) = sort_of (atom b)"
unfolding atom_def by (simp add: sort_of_Rep)
show "atom a = atom b \<longleftrightarrow> a = b"
unfolding atom_def by (simp add: Rep_inject)
show "p \<bullet> atom a = atom (p \<bullet> a)"
unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
qed
lemma at_class_sort:
fixes s :: atom_sort
fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
fixes a::"'a"
assumes type: "type_definition Rep Abs {a. sort_of a = s}"
assumes atom_def: "\<And>a. atom a = Rep a"
shows "sort_of (atom a) = s"
using atom_def type
unfolding type_definition_def by simp
setup \<open>Sign.add_const_constraint
(@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"})\<close>
setup \<open>Sign.add_const_constraint
(@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"})\<close>
section \<open>Library functions for the nominal infrastructure\<close>
ML_file \<open>nominal_library.ML\<close>
section \<open>The freshness lemma according to Andy Pitts\<close>
lemma freshness_lemma:
fixes h :: "'a::at \<Rightarrow> 'b::pt"
assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
shows "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
proof -
from a obtain b where a1: "atom b \<sharp> h" and a2: "atom b \<sharp> h b"
by (auto simp: fresh_Pair)
show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
proof (intro exI allI impI)
fix a :: 'a
assume a3: "atom a \<sharp> h"
show "h a = h b"
proof (cases "a = b")
assume "a = b"
thus "h a = h b" by simp
next
assume "a \<noteq> b"
hence "atom a \<sharp> b" by (simp add: fresh_at_base)
with a3 have "atom a \<sharp> h b"
by (rule fresh_fun_app)
with a2 have d1: "(atom b \<rightleftharpoons> atom a) \<bullet> (h b) = (h b)"
by (rule swap_fresh_fresh)
from a1 a3 have d2: "(atom b \<rightleftharpoons> atom a) \<bullet> h = h"
by (rule swap_fresh_fresh)
from d1 have "h b = (atom b \<rightleftharpoons> atom a) \<bullet> (h b)" by simp
also have "\<dots> = ((atom b \<rightleftharpoons> atom a) \<bullet> h) ((atom b \<rightleftharpoons> atom a) \<bullet> b)"
by (rule permute_fun_app_eq)
also have "\<dots> = h a"
using d2 by simp
finally show "h a = h b" by simp
qed
qed
qed
lemma freshness_lemma_unique:
fixes h :: "'a::at \<Rightarrow> 'b::pt"
assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
shows "\<exists>!x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
proof (rule ex_ex1I)
from a show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
by (rule freshness_lemma)
next
fix x y
assume x: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
assume y: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = y"
from a x y show "x = y"
by (auto simp: fresh_Pair)
qed
text \<open>packaging the freshness lemma into a function\<close>
definition
Fresh :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b"
where
"Fresh h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)"
lemma Fresh_apply:
fixes h :: "'a::at \<Rightarrow> 'b::pt"
assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
assumes b: "atom a \<sharp> h"
shows "Fresh h = h a"
unfolding Fresh_def
proof (rule the_equality)
show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a"
proof (intro strip)
fix a':: 'a
assume c: "atom a' \<sharp> h"
from a have "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" by (rule freshness_lemma)
with b c show "h a' = h a" by auto
qed
next
fix fr :: 'b
assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr"
with b show "fr = h a" by auto
qed
lemma Fresh_apply':
fixes h :: "'a::at \<Rightarrow> 'b::pt"
assumes a: "atom a \<sharp> h" "atom a \<sharp> h a"
shows "Fresh h = h a"
apply (rule Fresh_apply)
apply (auto simp: fresh_Pair intro: a)
done
simproc_setup Fresh_simproc ("Fresh (h::'a::at \<Rightarrow> 'b::pt)") = \<open>fn _ => fn ctxt => fn ctrm =>
let
val _ $ h = Thm.term_of ctrm
val atoms = Simplifier.prems_of ctxt
|> map_filter (fn thm => case Thm.prop_of thm of
_ $ \<^Const_>\<open>fresh _ for \<^Const_>\<open>atom _ for atm\<close> _\<close> => SOME atm | _ => NONE)
|> distinct (op =)
fun get_thm atm =
let
val goal1 = HOLogic.mk_Trueprop (mk_fresh (mk_atom atm) h)
val goal2 = HOLogic.mk_Trueprop (mk_fresh (mk_atom atm) (h $ atm))
val thm1 = Goal.prove ctxt [] [] goal1 (fn {context = ctxt', ...} => asm_simp_tac ctxt' 1)
val thm2 = Goal.prove ctxt [] [] goal2 (fn {context = ctxt', ...} => asm_simp_tac ctxt' 1)
in
SOME (@{thm Fresh_apply'} OF [thm1, thm2] RS eq_reflection)
end handle ERROR _ => NONE
in
get_first get_thm atoms
end
\<close>
lemma Fresh_eqvt:
fixes h :: "'a::at \<Rightarrow> 'b::pt"
assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
shows "p \<bullet> (Fresh h) = Fresh (p \<bullet> h)"
proof -
from a obtain a::"'a::at" where fr: "atom a \<sharp> h" "atom a \<sharp> h a"
by (metis fresh_Pair)
then have fr_p: "atom (p \<bullet> a) \<sharp> (p \<bullet> h)" "atom (p \<bullet> a) \<sharp> (p \<bullet> h) (p \<bullet> a)"
by (metis atom_eqvt fresh_permute_iff eqvt_apply)+
have "p \<bullet> (Fresh h) = p \<bullet> (h a)" using fr by simp
also have "... = (p \<bullet> h) (p \<bullet> a)" by simp
also have "... = Fresh (p \<bullet> h)" using fr_p by simp
finally show "p \<bullet> (Fresh h) = Fresh (p \<bullet> h)" .
qed
lemma Fresh_supports:
fixes h :: "'a::at \<Rightarrow> 'b::pt"
assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
shows "(supp h) supports (Fresh h)"
apply (simp add: supports_def fresh_def [symmetric])
apply (simp add: Fresh_eqvt [OF a] swap_fresh_fresh)
done
notation Fresh (binder "FRESH " 10)
lemma FRESH_f_iff:
fixes P :: "'a::at \<Rightarrow> 'b::pure"
fixes f :: "'b \<Rightarrow> 'c::pure"
assumes P: "finite (supp P)"
shows "(FRESH x. f (P x)) = f (FRESH x. P x)"
proof -
obtain a::'a where "atom a \<sharp> P" using P by (rule obtain_fresh')
then show "(FRESH x. f (P x)) = f (FRESH x. P x)"
by (simp add: pure_fresh)
qed
lemma FRESH_binop_iff:
fixes P :: "'a::at \<Rightarrow> 'b::pure"
fixes Q :: "'a::at \<Rightarrow> 'c::pure"
fixes binop :: "'b \<Rightarrow> 'c \<Rightarrow> 'd::pure"
assumes P: "finite (supp P)"
and Q: "finite (supp Q)"
shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)"
proof -
from assms have "finite (supp (P, Q))" by (simp add: supp_Pair)
then obtain a::'a where "atom a \<sharp> (P, Q)" by (rule obtain_fresh')
then show ?thesis
by (simp add: pure_fresh)
qed
lemma FRESH_conj_iff:
fixes P Q :: "'a::at \<Rightarrow> bool"
assumes P: "finite (supp P)" and Q: "finite (supp Q)"
shows "(FRESH x. P x \<and> Q x) \<longleftrightarrow> (FRESH x. P x) \<and> (FRESH x. Q x)"
using P Q by (rule FRESH_binop_iff)
lemma FRESH_disj_iff:
fixes P Q :: "'a::at \<Rightarrow> bool"
assumes P: "finite (supp P)" and Q: "finite (supp Q)"
shows "(FRESH x. P x \<or> Q x) \<longleftrightarrow> (FRESH x. P x) \<or> (FRESH x. Q x)"
using P Q by (rule FRESH_binop_iff)
section \<open>Automation for creating concrete atom types\<close>
text \<open>At the moment only single-sort concrete atoms are supported.\<close>
ML_file \<open>nominal_atoms.ML\<close>
section \<open>Automatic equivariance procedure for inductive definitions\<close>
ML_file \<open>nominal_eqvt.ML\<close>
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,229 +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"
Nested_Multisets_Ordinals.Duplicate_Free_Multiset
begin
notation image_mset (infixr "`#" 90)
section \<open>Finite maps and multisets\<close>
subsection \<open>Finite sets and multisets\<close>
abbreviation mset_fset :: \<open>'a fset \<Rightarrow> 'a multiset\<close> where
\<open>mset_fset N \<equiv> mset_set (fset N)\<close>
definition fset_mset :: \<open>'a multiset \<Rightarrow> 'a fset\<close> where
\<open>fset_mset N \<equiv> Abs_fset (set_mset N)\<close>
lemma fset_mset_mset_fset: \<open>fset_mset (mset_fset N) = N\<close>
by (auto simp: fset.fset_inverse fset_mset_def)
lemma mset_fset_fset_mset[simp]:
\<open>mset_fset (fset_mset N) = remdups_mset N\<close>
by (auto simp: fset.fset_inverse fset_mset_def Abs_fset_inverse remdups_mset_def)
lemma in_mset_fset_fmember[simp]: \<open>x \<in># mset_fset N \<longleftrightarrow> x |\<in>| N\<close>
- by (auto simp: fmember_iff_member_fset)
+ by simp
lemma in_fset_mset_mset[simp]: \<open>x |\<in>| fset_mset N \<longleftrightarrow> x \<in># N\<close>
- by (auto simp: fmember_iff_member_fset fset_mset_def Abs_fset_inverse)
+ by (simp add: fset_mset_def Abs_fset_inverse)
subsection \<open>Finite map and multisets\<close>
text \<open>Roughly the same as \<^term>\<open>ran\<close> and \<^term>\<open>dom\<close>, but with duplication in the content (unlike their
finite sets counterpart) while still working on finite domains (unlike a function mapping).
Remark that \<^term>\<open>dom_m\<close> (the keys) does not contain duplicates, but we keep for symmetry (and for
easier use of multiset operators as in the definition of \<^term>\<open>ran_m\<close>).
\<close>
definition dom_m where
\<open>dom_m N = mset_fset (fmdom N)\<close>
definition ran_m where
\<open>ran_m N = the `# fmlookup N `# dom_m N\<close>
lemma dom_m_fmdrop[simp]: \<open>dom_m (fmdrop C N) = remove1_mset C (dom_m N)\<close>
unfolding dom_m_def
by (cases \<open>C |\<in>| fmdom N\<close>)
- (auto simp: mset_set.remove fmember_iff_member_fset)
+ (auto simp: mset_set.remove)
lemma dom_m_fmdrop_All: \<open>dom_m (fmdrop C N) = removeAll_mset C (dom_m N)\<close>
unfolding dom_m_def
by (cases \<open>C |\<in>| fmdom N\<close>)
- (auto simp: mset_set.remove fmember_iff_member_fset)
+ (auto simp: mset_set.remove)
lemma dom_m_fmupd[simp]: \<open>dom_m (fmupd k C N) = add_mset k (remove1_mset k (dom_m N))\<close>
unfolding dom_m_def
by (cases \<open>k |\<in>| fmdom N\<close>)
- (auto simp: mset_set.remove fmember_iff_member_fset mset_set.insert_remove)
+ (auto simp: mset_set.remove mset_set.insert_remove)
lemma distinct_mset_dom: \<open>distinct_mset (dom_m N)\<close>
by (simp add: distinct_mset_mset_set dom_m_def)
lemma in_dom_m_lookup_iff: \<open>C \<in># dom_m N' \<longleftrightarrow> fmlookup N' C \<noteq> None\<close>
by (auto simp: dom_m_def fmdom.rep_eq fmlookup_dom'_iff)
lemma in_dom_in_ran_m[simp]: \<open>i \<in># dom_m N \<Longrightarrow> the (fmlookup N i) \<in># ran_m N\<close>
by (auto simp: ran_m_def)
lemma fmupd_same[simp]:
\<open>x1 \<in># dom_m x1aa \<Longrightarrow> fmupd x1 (the (fmlookup x1aa x1)) x1aa = x1aa\<close>
by (metis fmap_ext fmupd_lookup in_dom_m_lookup_iff option.collapse)
lemma ran_m_fmempty[simp]: \<open>ran_m fmempty = {#}\<close> and
dom_m_fmempty[simp]: \<open>dom_m fmempty = {#}\<close>
by (auto simp: ran_m_def dom_m_def)
lemma fmrestrict_set_fmupd:
\<open>a \<in> xs \<Longrightarrow> fmrestrict_set xs (fmupd a C N) = fmupd a C (fmrestrict_set xs N)\<close>
\<open>a \<notin> xs \<Longrightarrow> fmrestrict_set xs (fmupd a C N) = fmrestrict_set xs N\<close>
by (auto simp: fmfilter_alt_defs)
lemma fset_fmdom_fmrestrict_set:
\<open>fset (fmdom (fmrestrict_set xs N)) = fset (fmdom N) \<inter> xs\<close>
by (auto simp: fmfilter_alt_defs)
lemma dom_m_fmrestrict_set: \<open>dom_m (fmrestrict_set (set xs) N) = mset xs \<inter># dom_m N\<close>
using fset_fmdom_fmrestrict_set[of \<open>set xs\<close> N] distinct_mset_dom[of N]
distinct_mset_inter_remdups_mset[of \<open>mset_fset (fmdom N)\<close> \<open>mset xs\<close>]
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': \<open>dom_m (fmrestrict_set xs N) = mset_set (xs \<inter> set_mset (dom_m N))\<close>
using fset_fmdom_fmrestrict_set[of \<open>xs\<close> 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: \<open>fmlookup m x = Some y \<Longrightarrow> x \<in># dom_m m\<close>
- by (drule fmdomI) (auto simp: dom_m_def fmember_iff_member_fset)
+ by (drule fmdomI) (auto simp: dom_m_def)
lemma fmupd_fmdrop_id:
assumes \<open>k |\<in>| fmdom N'\<close>
shows \<open>fmupd k (the (fmlookup N' k)) (fmdrop k N') = N'\<close>
proof -
have [simp]: \<open>map_upd k (the (fmlookup N' k))
(\<lambda>x. if x \<noteq> k then fmlookup N' x else None) =
map_upd k (the (fmlookup N' k))
(fmlookup N')\<close>
by (auto intro!: ext simp: map_upd_def)
have [simp]: \<open>map_upd k (the (fmlookup N' k)) (fmlookup N') = fmlookup N'\<close>
using assms
by (auto intro!: ext simp: map_upd_def)
have [simp]: \<open>finite (dom (\<lambda>x. if x = k then None else fmlookup N' x))\<close>
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: \<open>k |\<in>| fmdom N' \<Longrightarrow> \<exists>N'' v. N' = fmupd k v N'' \<and> the (fmlookup N' k) = v \<and>
k |\<notin>| fmdom N''\<close>
by (rule exI[of _ \<open>fmdrop k N'\<close>])
(auto simp: fmupd_fmdrop_id)
lemma \<open>fmdrop k (fmupd k va N'') = fmdrop k N''\<close>
by (simp add: fmap_ext)
lemma fmap_ext_fmdom:
\<open>(fmdom N = fmdom N') \<Longrightarrow> (\<And> x. x |\<in>| fmdom N \<Longrightarrow> fmlookup N x = fmlookup N' x) \<Longrightarrow>
N = N'\<close>
by (rule fmap_ext)
(case_tac \<open>x |\<in>| fmdom N\<close>, auto simp: fmdom_notD)
lemma fmrestrict_set_insert_in:
\<open>xa \<in> fset (fmdom N) \<Longrightarrow>
fmrestrict_set (insert xa l1) N = fmupd xa (the (fmlookup N xa)) (fmrestrict_set l1 N)\<close>
apply (rule fmap_ext_fmdom)
- apply (auto simp: fset_fmdom_fmrestrict_set fmember_iff_member_fset notin_fset; fail)[]
+ apply (auto simp: fset_fmdom_fmrestrict_set; fail)[]
apply (auto simp: fmlookup_dom_iff; fail)
done
lemma fmrestrict_set_insert_notin:
\<open>xa \<notin> fset (fmdom N) \<Longrightarrow>
fmrestrict_set (insert xa l1) N = fmrestrict_set l1 N\<close>
by (rule fmap_ext_fmdom)
- (auto simp: fset_fmdom_fmrestrict_set fmember_iff_member_fset notin_fset)
+ (auto simp: fset_fmdom_fmrestrict_set)
lemma fmrestrict_set_insert_in_dom_m[simp]:
\<open>xa \<in># dom_m N \<Longrightarrow>
fmrestrict_set (insert xa l1) N = fmupd xa (the (fmlookup N xa)) (fmrestrict_set l1 N)\<close>
by (simp add: fmrestrict_set_insert_in dom_m_def)
lemma fmrestrict_set_insert_notin_dom_m[simp]:
\<open>xa \<notin># dom_m N \<Longrightarrow>
fmrestrict_set (insert xa l1) N = fmrestrict_set l1 N\<close>
by (simp add: fmrestrict_set_insert_notin dom_m_def)
lemma fmlookup_restrict_set_id: \<open>fset (fmdom N) \<subseteq> A \<Longrightarrow> fmrestrict_set A N = N\<close>
by (metis fmap_ext fmdom'_alt_def fmdom'_notD fmlookup_restrict_set subset_iff)
lemma fmlookup_restrict_set_id': \<open>set_mset (dom_m N) \<subseteq> A \<Longrightarrow> fmrestrict_set A N = N\<close>
by (rule fmlookup_restrict_set_id)
(auto simp: dom_m_def)
lemma ran_m_mapsto_upd:
assumes
NC: \<open>C \<in># dom_m N\<close>
shows \<open>ran_m (fmupd C C' N) = add_mset C' (remove1_mset (the (fmlookup N C)) (ran_m N))\<close>
proof -
define N' where
\<open>N' = fmdrop C N\<close>
have N_N': \<open>dom_m N = add_mset C (dom_m N')\<close>
using NC unfolding N'_def by auto
have \<open>C \<notin># dom_m N'\<close>
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: \<open>C \<notin># dom_m N\<close>
shows \<open>ran_m (fmupd C C' N) = add_mset C' (ran_m N)\<close>
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:
\<open>C \<notin># A \<Longrightarrow> {#f (if x = C then a x else b x). x \<in># A#} = {# f(b x). x \<in># A #}\<close>
by (induction A) auto
lemma filter_mset_cong2:
"(\<And>x. x \<in># M \<Longrightarrow> f x = g x) \<Longrightarrow> M = N \<Longrightarrow> filter_mset f M = filter_mset g N"
by (hypsubst, rule filter_mset_cong, simp)
lemma ran_m_fmdrop:
\<open>C \<in># dom_m N \<Longrightarrow> ran_m (fmdrop C N) = remove1_mset (the (fmlookup N C)) (ran_m N)\<close>
using distinct_mset_dom[of N]
by (cases \<open>fmlookup N C\<close>)
(auto simp: ran_m_def image_mset_If_eq_notin[of C _ \<open>\<lambda>x. fst (the x)\<close>]
dest!: multi_member_split
intro!: filter_mset_cong2 image_mset_cong2)
lemma ran_m_fmdrop_notin:
\<open>C \<notin># dom_m N \<Longrightarrow> ran_m (fmdrop C N) = ran_m N\<close>
using distinct_mset_dom[of N]
by (auto simp: ran_m_def image_mset_If_eq_notin[of C _ \<open>\<lambda>x. fst (the x)\<close>]
dest!: multi_member_split
intro!: filter_mset_cong2 image_mset_cong2)
lemma ran_m_fmdrop_If:
\<open>ran_m (fmdrop C N) = (if C \<in># dom_m N then remove1_mset (the (fmlookup N C)) (ran_m N) else ran_m N)\<close>
using distinct_mset_dom[of N]
by (auto simp: ran_m_def image_mset_If_eq_notin[of C _ \<open>\<lambda>x. fst (the x)\<close>]
dest!: multi_member_split
intro!: filter_mset_cong2 image_mset_cong2)
lemma dom_m_empty_iff[iff]:
\<open>dom_m NU = {#} \<longleftrightarrow> NU = fmempty\<close>
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_Map_Rel.thy b/thys/PAC_Checker/PAC_Map_Rel.thy
--- a/thys/PAC_Checker/PAC_Map_Rel.thy
+++ b/thys/PAC_Checker/PAC_Map_Rel.thy
@@ -1,321 +1,319 @@
(*
File: PAC_Map_Rel.thy
Author: Mathias Fleury, Daniela Kaufmann, JKU
Maintainer: Mathias Fleury, JKU
*)
theory PAC_Map_Rel
imports
Refine_Imperative_HOL.IICF Finite_Map_Multiset
begin
section \<open>Hash-Map for finite mappings\<close>
text \<open>
This function declares hash-maps for \<^typ>\<open>('a, 'b)fmap\<close>, that are nicer
to use especially here where everything is finite.
\<close>
definition fmap_rel where
[to_relAPP]:
"fmap_rel K V \<equiv> {(m1, m2).
(\<forall>i j. i |\<in>| fmdom m2 \<longrightarrow> (j, i) \<in> K \<longrightarrow> (the (fmlookup m1 j), the (fmlookup m2 i)) \<in> V) \<and>
fset (fmdom m1) \<subseteq> Domain K \<and> fset (fmdom m2) \<subseteq> Range K \<and>
(\<forall>i j. (i, j) \<in> K \<longrightarrow> j |\<in>| fmdom m2 \<longleftrightarrow> i |\<in>| fmdom m1)}"
lemma fmap_rel_alt_def:
\<open>\<langle>K, V\<rangle>fmap_rel \<equiv>
{(m1, m2).
(\<forall>i j. i \<in># dom_m m2 \<longrightarrow>
(j, i) \<in> K \<longrightarrow> (the (fmlookup m1 j), the (fmlookup m2 i)) \<in> V) \<and>
fset (fmdom m1) \<subseteq> Domain K \<and>
fset (fmdom m2) \<subseteq> Range K \<and>
(\<forall>i j. (i, j) \<in> K \<longrightarrow> (j \<in># dom_m m2) = (i \<in># dom_m m1))}
\<close>
- unfolding fmap_rel_def dom_m_def fmember_iff_member_fset
+ unfolding fmap_rel_def dom_m_def
by auto
lemma fmdom_empty_fmempty_iff[simp]: \<open>fmdom m = {||} \<longleftrightarrow> m = fmempty\<close>
by (metis fmdom_empty fmdrop_fset_fmdom fmdrop_fset_null)
lemma fmap_rel_empty1_simp[simp]:
"(fmempty,m)\<in>\<langle>K,V\<rangle>fmap_rel \<longleftrightarrow> m=fmempty"
apply (cases \<open>fmdom m = {||}\<close>)
apply (auto simp: fmap_rel_def)[]
- by (auto simp add: fmember_iff_member_fset fmap_rel_def simp del: fmdom_empty_fmempty_iff)
+ by (auto simp add: fmap_rel_def simp del: fmdom_empty_fmempty_iff)
lemma fmap_rel_empty2_simp[simp]:
"(m,fmempty)\<in>\<langle>K,V\<rangle>fmap_rel \<longleftrightarrow> m=fmempty"
apply (cases \<open>fmdom m = {||}\<close>)
apply (auto simp: fmap_rel_def)[]
- by (fastforce simp add: fmember_iff_member_fset fmap_rel_def simp del: fmdom_empty_fmempty_iff)
+ by (fastforce simp add: fmap_rel_def simp del: fmdom_empty_fmempty_iff)
sepref_decl_intf ('k,'v) f_map is "('k, 'v) fmap"
lemma [synth_rules]: "\<lbrakk>INTF_OF_REL K TYPE('k); INTF_OF_REL V TYPE('v)\<rbrakk>
\<Longrightarrow> INTF_OF_REL (\<langle>K,V\<rangle>fmap_rel) TYPE(('k,'v) f_map)" by simp
subsection \<open>Operations\<close>
sepref_decl_op fmap_empty: "fmempty" :: "\<langle>K,V\<rangle>fmap_rel" .
sepref_decl_op fmap_is_empty: "(=) fmempty" :: "\<langle>K,V\<rangle>fmap_rel \<rightarrow> bool_rel"
apply (rule fref_ncI)
apply parametricity
apply (rule fun_relI; auto)
done
lemma fmap_rel_fmupd_fmap_rel:
\<open>(A, B) \<in> \<langle>K, R\<rangle>fmap_rel \<Longrightarrow> (p, p') \<in> K \<Longrightarrow> (q, q') \<in> R \<Longrightarrow>
(fmupd p q A, fmupd p' q' B) \<in> \<langle>K, R\<rangle>fmap_rel\<close>
if "single_valued K" "single_valued (K\<inverse>)"
using that
unfolding fmap_rel_alt_def
apply (case_tac \<open>p' \<in># dom_m B\<close>)
apply (auto simp add: all_conj_distrib IS_RIGHT_UNIQUED dest!: multi_member_split)
done
sepref_decl_op fmap_update: "fmupd" :: "K \<rightarrow> V \<rightarrow> \<langle>K,V\<rangle>fmap_rel \<rightarrow> \<langle>K,V\<rangle>fmap_rel"
where "single_valued K" "single_valued (K\<inverse>)"
apply (rule fref_ncI)
apply parametricity
apply (intro fun_relI)
by (rule fmap_rel_fmupd_fmap_rel)
lemma remove1_mset_eq_add_mset_iff:
\<open>remove1_mset a A = add_mset a A' \<longleftrightarrow> A = add_mset a (add_mset a A')\<close>
by (metis add_mset_add_single add_mset_diff_bothsides diff_zero remove1_mset_eqE)
lemma fmap_rel_fmdrop_fmap_rel:
\<open>(fmdrop p A, fmdrop p' B) \<in> \<langle>K, R\<rangle>fmap_rel\<close>
if single: "single_valued K" "single_valued (K\<inverse>)" and
H0: \<open>(A, B) \<in> \<langle>K, R\<rangle>fmap_rel\<close> \<open>(p, p') \<in> K\<close>
proof -
have H: \<open>\<And>Aa j.
\<forall>i. i \<in># dom_m B \<longrightarrow> (\<forall>j. (j, i) \<in> K \<longrightarrow> (the (fmlookup A j), the (fmlookup B i)) \<in> R) \<Longrightarrow>
remove1_mset p' (dom_m B) = add_mset p' Aa \<Longrightarrow> (j, p') \<in> K \<Longrightarrow> False\<close>
by (metis dom_m_fmdrop fmlookup_drop in_dom_m_lookup_iff union_single_eq_member)
have H2: \<open>\<And>i Aa j.
(p, p') \<in> K \<Longrightarrow>
\<forall>i. i \<in># dom_m B \<longrightarrow> (\<forall>j. (j, i) \<in> K \<longrightarrow> (the (fmlookup A j), the (fmlookup B i)) \<in> R) \<Longrightarrow>
\<forall>i j. (i, j) \<in> K \<longrightarrow> (j \<in># dom_m B) = (i \<in># dom_m A) \<Longrightarrow>
remove1_mset p' (dom_m B) = add_mset i Aa \<Longrightarrow>
(j, i) \<in> K \<Longrightarrow>
(the (fmlookup A j), the (fmlookup B i)) \<in> R \<and> j \<in># remove1_mset p (dom_m A) \<and>
i \<in># remove1_mset p' (dom_m B)\<close>
\<open>\<And>i j Aa.
(p, p') \<in> K \<Longrightarrow>
single_valued K \<Longrightarrow>
single_valued (K\<inverse>) \<Longrightarrow>
\<forall>i. i \<in># dom_m B \<longrightarrow> (\<forall>j. (j, i) \<in> K \<longrightarrow> (the (fmlookup A j), the (fmlookup B i)) \<in> R) \<Longrightarrow>
fset (fmdom A) \<subseteq> Domain K \<Longrightarrow>
fset (fmdom B) \<subseteq> Range K \<Longrightarrow>
\<forall>i j. (i, j) \<in> K \<longrightarrow> (j \<in># dom_m B) = (i \<in># dom_m A) \<Longrightarrow>
(i, j) \<in> K \<Longrightarrow> remove1_mset p (dom_m A) = add_mset i Aa \<Longrightarrow> j \<in># remove1_mset p' (dom_m B)\<close>
using single
by (metis IS_RIGHT_UNIQUED converse.intros dom_m_fmdrop fmlookup_drop in_dom_m_lookup_iff
union_single_eq_member)+
show \<open>(fmdrop p A, fmdrop p' B) \<in> \<langle>K, R\<rangle>fmap_rel\<close>
using that
unfolding fmap_rel_alt_def
by (auto simp add: all_conj_distrib IS_RIGHT_UNIQUED
dest!: multi_member_split dest: H H2)
qed
sepref_decl_op fmap_delete: "fmdrop" :: "K \<rightarrow> \<langle>K,V\<rangle>fmap_rel \<rightarrow> \<langle>K,V\<rangle>fmap_rel"
where "single_valued K" "single_valued (K\<inverse>)"
apply (rule fref_ncI)
apply parametricity
by (auto simp add: fmap_rel_fmdrop_fmap_rel)
lemma fmap_rel_nat_the_fmlookup[intro]:
\<open>(A, B) \<in> \<langle>S, R\<rangle>fmap_rel \<Longrightarrow> (p, p') \<in> S \<Longrightarrow> p' \<in># dom_m B \<Longrightarrow>
(the (fmlookup A p), the (fmlookup B p')) \<in> R\<close>
by (auto simp: fmap_rel_alt_def distinct_mset_dom)
lemma fmap_rel_in_dom_iff:
\<open>(aa, a'a) \<in> \<langle>K, V\<rangle>fmap_rel \<Longrightarrow>
(a, a') \<in> K \<Longrightarrow>
a' \<in># dom_m a'a \<longleftrightarrow>
a \<in># dom_m aa\<close>
unfolding fmap_rel_alt_def
by auto
lemma fmap_rel_fmlookup_rel:
\<open>(a, a') \<in> K \<Longrightarrow> (aa, a'a) \<in> \<langle>K, V\<rangle>fmap_rel \<Longrightarrow>
(fmlookup aa a, fmlookup a'a a') \<in> \<langle>V\<rangle>option_rel\<close>
using fmap_rel_nat_the_fmlookup[of aa a'a K V a a']
fmap_rel_in_dom_iff[of aa a'a K V a a']
in_dom_m_lookup_iff[of a' a'a]
in_dom_m_lookup_iff[of a aa]
by (cases \<open>a' \<in># dom_m a'a\<close>)
(auto simp del: fmap_rel_nat_the_fmlookup)
sepref_decl_op fmap_lookup: "fmlookup" :: "\<langle>K,V\<rangle>fmap_rel \<rightarrow> K \<rightarrow> \<langle>V\<rangle>option_rel"
apply (rule fref_ncI)
apply parametricity
apply (intro fun_relI)
apply (rule fmap_rel_fmlookup_rel; assumption)
done
lemma in_fdom_alt: "k\<in>#dom_m m \<longleftrightarrow> \<not>is_None (fmlookup m k)"
- apply (auto split: option.split intro: fmdom_notI simp: dom_m_def fmember_iff_member_fset)
- apply (meson fmdom_notI notin_fset)
- using notin_fset by fastforce
+ by (auto split: option.split intro: fmdom_notI fmdomI simp: dom_m_def)
sepref_decl_op fmap_contains_key: "\<lambda>k m. k\<in>#dom_m m" :: "K \<rightarrow> \<langle>K,V\<rangle>fmap_rel \<rightarrow> bool_rel"
unfolding in_fdom_alt
apply (rule fref_ncI)
apply parametricity
apply (rule fmap_rel_fmlookup_rel; assumption)
done
subsection \<open>Patterns\<close>
lemma pat_fmap_empty[pat_rules]: "fmempty \<equiv> op_fmap_empty" by simp
lemma pat_map_is_empty[pat_rules]:
"(=) $m$fmempty \<equiv> op_fmap_is_empty$m"
"(=) $fmempty$m \<equiv> op_fmap_is_empty$m"
"(=) $(dom_m$m)${#} \<equiv> op_fmap_is_empty$m"
"(=) ${#}$(dom_m$m) \<equiv> op_fmap_is_empty$m"
unfolding atomize_eq
by (auto dest: sym)
lemma op_map_contains_key[pat_rules]:
"(\<in>#) $ k $ (dom_m$m) \<equiv> op_fmap_contains_key$'k$'m"
by (auto intro!: eq_reflection)
subsection \<open>Mapping to Normal Hashmaps\<close>
abbreviation map_of_fmap :: \<open>('k \<Rightarrow> 'v option) \<Rightarrow> ('k, 'v) fmap\<close> where
\<open>map_of_fmap h \<equiv> Abs_fmap h\<close>
definition map_fmap_rel where
\<open>map_fmap_rel = br map_of_fmap (\<lambda>a. finite (dom a))\<close>
lemma fmdrop_set_None:
\<open>(op_map_delete, fmdrop) \<in> Id \<rightarrow> map_fmap_rel \<rightarrow> map_fmap_rel\<close>
apply (auto simp: map_fmap_rel_def br_def)
apply (subst fmdrop.abs_eq)
apply (auto simp: eq_onp_def fmap.Abs_fmap_inject
map_drop_def map_filter_finite
intro!: ext)
apply (auto simp: map_filter_def)
done
lemma map_upd_fmupd:
\<open>(op_map_update, fmupd) \<in> Id \<rightarrow> Id \<rightarrow> map_fmap_rel \<rightarrow> map_fmap_rel\<close>
apply (auto simp: map_fmap_rel_def br_def)
apply (subst fmupd.abs_eq)
apply (auto simp: eq_onp_def fmap.Abs_fmap_inject
map_drop_def map_filter_finite map_upd_def
intro!: ext)
done
text \<open>Technically @{term op_map_lookup} has the arguments in the wrong direction.\<close>
definition fmlookup' where
[simp]: \<open>fmlookup' A k = fmlookup k A\<close>
lemma [def_pat_rules]:
\<open>((\<in>#)$k$(dom_m$A)) \<equiv> Not$(is_None$(fmlookup'$k$A))\<close>
by (simp add: fold_is_None in_fdom_alt)
lemma op_map_lookup_fmlookup:
\<open>(op_map_lookup, fmlookup') \<in> Id \<rightarrow> map_fmap_rel \<rightarrow> \<langle>Id\<rangle>option_rel\<close>
by (auto simp: map_fmap_rel_def br_def fmap.Abs_fmap_inverse)
abbreviation hm_fmap_assn where
\<open>hm_fmap_assn K V \<equiv> hr_comp (hm.assn K V) map_fmap_rel\<close>
lemmas fmap_delete_hnr [sepref_fr_rules] =
hm.delete_hnr[FCOMP fmdrop_set_None]
lemmas fmap_update_hnr [sepref_fr_rules] =
hm.update_hnr[FCOMP map_upd_fmupd]
lemmas fmap_lookup_hnr [sepref_fr_rules] =
hm.lookup_hnr[FCOMP op_map_lookup_fmlookup]
lemma fmempty_empty:
\<open>(uncurry0 (RETURN op_map_empty), uncurry0 (RETURN fmempty)) \<in> unit_rel \<rightarrow>\<^sub>f \<langle>map_fmap_rel\<rangle>nres_rel\<close>
by (auto simp: map_fmap_rel_def br_def fmempty_def frefI nres_relI)
lemmas [sepref_fr_rules] =
hm.empty_hnr[FCOMP fmempty_empty, unfolded op_fmap_empty_def[symmetric]]
abbreviation iam_fmap_assn where
\<open>iam_fmap_assn K V \<equiv> hr_comp (iam.assn K V) map_fmap_rel\<close>
lemmas iam_fmap_delete_hnr [sepref_fr_rules] =
iam.delete_hnr[FCOMP fmdrop_set_None]
lemmas iam_ffmap_update_hnr [sepref_fr_rules] =
iam.update_hnr[FCOMP map_upd_fmupd]
lemmas iam_ffmap_lookup_hnr [sepref_fr_rules] =
iam.lookup_hnr[FCOMP op_map_lookup_fmlookup]
definition op_iam_fmap_empty where
\<open>op_iam_fmap_empty = fmempty\<close>
lemma iam_fmempty_empty:
\<open>(uncurry0 (RETURN op_map_empty), uncurry0 (RETURN op_iam_fmap_empty)) \<in> unit_rel \<rightarrow>\<^sub>f \<langle>map_fmap_rel\<rangle>nres_rel\<close>
by (auto simp: map_fmap_rel_def br_def fmempty_def frefI nres_relI op_iam_fmap_empty_def)
lemmas [sepref_fr_rules] =
iam.empty_hnr[FCOMP fmempty_empty, unfolded op_iam_fmap_empty_def[symmetric]]
definition upper_bound_on_dom where
\<open>upper_bound_on_dom A = SPEC(\<lambda>n. \<forall>i \<in>#(dom_m A). i < n)\<close>
lemma [sepref_fr_rules]:
\<open>((Array.len), upper_bound_on_dom) \<in> (iam_fmap_assn nat_assn V)\<^sup>k \<rightarrow>\<^sub>a nat_assn\<close>
proof -
have [simp]: \<open>finite (dom b) \<Longrightarrow> i \<in> fset (fmdom (map_of_fmap b)) \<longleftrightarrow> i \<in> dom b\<close> for i b
by (subst fmdom.abs_eq)
(auto simp: eq_onp_def fset.Abs_fset_inverse)
have 2: \<open>nat_rel = the_pure (nat_assn)\<close> and
3: \<open>nat_assn = pure nat_rel\<close>
by auto
have [simp]: \<open>the_pure (\<lambda>a c :: nat. \<up> (c = a)) = nat_rel\<close>
apply (subst 2)
apply (subst 3)
apply (subst pure_def)
apply auto
done
have [simp]: \<open>(iam_of_list l, b) \<in> the_pure (\<lambda>a c :: nat. \<up> (c = a)) \<rightarrow> \<langle>the_pure V\<rangle>option_rel \<Longrightarrow>
b i = Some y \<Longrightarrow> i < length l\<close> for i b l y
by (auto dest!: fun_relD[of _ _ _ _ i i] simp: option_rel_def
iam_of_list_def split: if_splits)
show ?thesis
by sepref_to_hoare
(sep_auto simp: upper_bound_on_dom_def hr_comp_def iam.assn_def map_rel_def
map_fmap_rel_def is_iam_def br_def dom_m_def)
qed
lemma fmap_rel_nat_rel_dom_m[simp]:
\<open>(A, B) \<in> \<langle>nat_rel, R\<rangle>fmap_rel \<Longrightarrow> dom_m A = dom_m B\<close>
by (subst distinct_set_mset_eq_iff[symmetric])
(auto simp: fmap_rel_alt_def distinct_mset_dom
simp del: fmap_rel_nat_the_fmlookup)
lemma ref_two_step':
\<open>A \<le> B \<Longrightarrow> \<Down> R A \<le> \<Down> R B\<close>
using ref_two_step by auto
end
diff --git a/thys/Query_Optimization/Dtree.thy b/thys/Query_Optimization/Dtree.thy
--- a/thys/Query_Optimization/Dtree.thy
+++ b/thys/Query_Optimization/Dtree.thy
@@ -1,4547 +1,4572 @@
(* Author: Bernhard Stöckl *)
theory Dtree
imports Complex_Main "Directed_Tree_Additions" "HOL-Library.FSet"
begin
section \<open>Algebraic Type for Directed Trees\<close>
datatype (dverts:'a, darcs: 'b) dtree = Node (root: 'a) (sucs: "(('a,'b) dtree \<times> 'b) fset")
subsection \<open>Termination Proofs\<close>
lemma fset_sum_ge_elem: "finite xs \<Longrightarrow> x \<in> xs \<Longrightarrow> (\<Sum>u\<in>xs. (f::'a \<Rightarrow> nat) u) \<ge> f x"
by (simp add: sum_nonneg_leq_bound)
lemma dtree_size_decr_aux:
assumes "(x,y) \<in> fset xs"
shows "size x < size (Node r xs)"
proof -
have 0: "((x,size x),y) \<in> (map_prod (\<lambda>u. (u, size u)) (\<lambda>u. u)) ` fset xs" using assms by fast
have "size x < Suc (size_prod snd (\<lambda>_. 0) ((x,size x),y))" by simp
also have
"\<dots> \<le> (\<Sum>u\<in>(map_prod (\<lambda>x. (x, size x)) (\<lambda>y. y)) ` fset xs. Suc (size_prod snd (\<lambda>_. 0) u)) + 1"
using fset_sum_ge_elem 0 finite_fset finite_imageI
by (metis (mono_tags, lifting) add_increasing2 zero_le_one)
finally show ?thesis by simp
qed
lemma dtree_size_decr_aux': "t1 \<in> fst ` fset xs \<Longrightarrow> size t1 < size (Node r xs)"
using dtree_size_decr_aux by fastforce
lemma dtree_size_decr[termination_simp]:
assumes "(x, y) \<in> fset (xs:: (('a, 'b) dtree \<times> 'b) fset)"
shows "size x < Suc (\<Sum>u\<in>map_prod (\<lambda>x. (x, size x)) (\<lambda>y. y) ` fset xs. Suc (Suc (snd (fst u))))"
proof -
let ?xs = "(map_prod (\<lambda>x. (x, size x)) (\<lambda>y. y)) ` fset xs"
have "size x < (\<Sum>u\<in>?xs. Suc (size_prod snd (\<lambda>_. 0) u)) + 1"
using dtree_size_decr_aux assms by fastforce
also have "\<dots> = Suc (\<Sum>u\<in>?xs. Suc (Suc (snd (fst u))))" by (simp add: size_prod_simp)
finally show ?thesis by blast
qed
subsection "Dtree Basic Functions"
fun darcs_mset :: "('a,'b) dtree \<Rightarrow> 'b multiset" where
"darcs_mset (Node r xs) = (\<Sum>(t,e) \<in> fset xs. {#e#} + darcs_mset t)"
fun dverts_mset :: "('a,'b) dtree \<Rightarrow> 'a multiset" where
"dverts_mset (Node r xs) = {#r#} + (\<Sum>(t,e) \<in> fset xs. dverts_mset t)"
(* disjoint_darcs & wf_darcs' are old definitions equivalent to wf_darcs; still used for proofs *)
abbreviation disjoint_darcs :: "(('a,'b) dtree \<times> 'b) fset \<Rightarrow> bool" where
"disjoint_darcs xs \<equiv> (\<forall>(x,e1) \<in> fset xs. e1 \<notin> darcs x \<and> (\<forall>(y,e2) \<in> fset xs.
(darcs x \<union> {e1}) \<inter> (darcs y \<union> {e2}) = {} \<or> (x,e1)=(y,e2)))"
fun wf_darcs' :: "('a,'b) dtree \<Rightarrow> bool" where
"wf_darcs' (Node r xs) = (disjoint_darcs xs \<and> (\<forall>(x,e) \<in> fset xs. wf_darcs' x))"
definition wf_darcs :: "('a,'b) dtree \<Rightarrow> bool" where
"wf_darcs t = (\<forall>x \<in># darcs_mset t. count (darcs_mset t) x = 1)"
(* same here as with wf_darcs' *)
fun wf_dverts' :: "('a,'b) dtree \<Rightarrow> bool" where
"wf_dverts' (Node r xs) = (\<forall>(x,e1) \<in> fset xs.
r \<notin> dverts x \<and> (\<forall>(y,e2) \<in> fset xs. (dverts x \<inter> dverts y = {} \<or> (x,e1)=(y,e2))) \<and> wf_dverts' x)"
definition wf_dverts :: "('a,'b) dtree \<Rightarrow> bool" where
"wf_dverts t = (\<forall>x \<in># dverts_mset t. count (dverts_mset t) x = 1)"
fun dtail :: "('a,'b) dtree \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'a" where
"dtail (Node r xs) def = (\<lambda>e. if e \<in> snd ` fset xs then r
else (ffold (\<lambda>(x,e2) b.
if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r xs)
then b else dtail x def) def xs) e)"
(* (x,y) \<in> fset case required for termination proof (always fulfilled)
disjointness requirement for commutativity *)
fun dhead :: "('a,'b) dtree \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'a" where
"dhead (Node r xs) def = (\<lambda>e. (ffold (\<lambda>(x,e2) b.
if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs)
then b else if e=e2 then root x else dhead x def e) (def e) xs))"
abbreviation from_dtree :: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> ('a,'b) dtree \<Rightarrow> ('a,'b) pre_digraph" where
"from_dtree deft defh t \<equiv>
\<lparr>verts = dverts t, arcs = darcs t, tail = dtail t deft, head = dhead t defh\<rparr>"
abbreviation from_dtree' :: "('a,'b) dtree \<Rightarrow> ('a,'b) pre_digraph" where
"from_dtree' t \<equiv> from_dtree (\<lambda>_. root t) (\<lambda>_. root t) t"
fun is_subtree :: "('a,'b) dtree \<Rightarrow> ('a,'b) dtree \<Rightarrow> bool" where
"is_subtree x (Node r xs) =
(x = Node r xs \<or> (\<exists>(y,e) \<in> fset xs. is_subtree x y))"
definition strict_subtree :: "('a,'b) dtree \<Rightarrow> ('a,'b) dtree \<Rightarrow> bool" where
"strict_subtree t1 t2 \<longleftrightarrow> is_subtree t1 t2 \<and> t1 \<noteq> t2"
fun num_leaves :: "('a,'b) dtree \<Rightarrow> nat" where
"num_leaves (Node r xs) = (if xs = {||} then 1 else (\<Sum>(t,e)\<in> fset xs. num_leaves t))"
subsection "Dtree Basic Proofs"
lemma finite_dverts: "finite (dverts t)"
by(induction t) auto
lemma finite_darcs: "finite (darcs t)"
by(induction t) auto
lemma dverts_child_subseteq: "x \<in> fst ` fset xs \<Longrightarrow> dverts x \<subseteq> dverts (Node r xs)"
by fastforce
lemma dverts_suc_subseteq: "x \<in> fst ` fset (sucs t) \<Longrightarrow> dverts x \<subseteq> dverts t"
using dverts_child_subseteq[of x "sucs t" "root t"] by simp
lemma dverts_root_or_child: "v \<in> dverts (Node r xs) \<Longrightarrow> v = r \<or> v \<in> (\<Union>(t,e) \<in> fset xs. dverts t)"
by auto
lemma dverts_root_or_suc: "v \<in> dverts t \<Longrightarrow> v = root t \<or> (\<exists>(t,e) \<in> fset (sucs t).v \<in> dverts t)"
using dverts_root_or_child[of v "root t" "sucs t"] by auto
lemma dverts_child_if_not_root:
"\<lbrakk>v \<in> dverts (Node r xs); v \<noteq> r\<rbrakk> \<Longrightarrow> \<exists>t\<in>fst ` fset xs. v \<in> dverts t"
by force
lemma dverts_suc_if_not_root:
"\<lbrakk>v \<in> dverts t; v \<noteq> root t\<rbrakk> \<Longrightarrow> \<exists>t\<in>fst ` fset (sucs t). v \<in> dverts t"
using dverts_root_or_suc by force
lemma darcs_child_subseteq: "x \<in> fst ` fset xs \<Longrightarrow> darcs x \<subseteq> darcs (Node r xs)"
by force
lemma mset_sum_elem: "x \<in># (\<Sum>y \<in> fset Y. f y) \<Longrightarrow> \<exists>y \<in> fset Y. x \<in># f y"
- by (induction Y) (auto simp: notin_fset)
+ by (induction Y) auto
lemma mset_sum_elem_iff: "x \<in># (\<Sum>y \<in> fset Y. f y) \<longleftrightarrow> (\<exists>y \<in> fset Y. x \<in># f y)"
- by (induction Y) (auto simp: notin_fset)
+ by (induction Y) auto
lemma mset_sum_elemI: "\<lbrakk>y \<in> fset Y; x \<in># f y\<rbrakk> \<Longrightarrow> x \<in># (\<Sum>y \<in> fset Y. f y)"
- by (induction Y) (auto simp: notin_fset)
+ by (induction Y) auto
lemma darcs_mset_elem:
"x \<in># darcs_mset (Node r xs) \<Longrightarrow> \<exists>(t,e) \<in> fset xs. x \<in># darcs_mset t \<or> x = e"
using mset_sum_elem by fastforce
lemma darcs_mset_if_nsnd:
"\<lbrakk>x \<in># darcs_mset (Node r xs); x \<notin> snd ` fset xs\<rbrakk> \<Longrightarrow> \<exists>(t1,e1) \<in> fset xs. x \<in># darcs_mset t1"
using darcs_mset_elem[of x r xs] by force
lemma darcs_mset_suc_if_nsnd:
"\<lbrakk>x \<in># darcs_mset t; x \<notin> snd ` fset (sucs t)\<rbrakk> \<Longrightarrow> \<exists>(t1,e1) \<in> fset (sucs t). x \<in># darcs_mset t1"
using darcs_mset_if_nsnd[of x "root t" "sucs t"] by simp
lemma darcs_mset_if_nchild:
"\<lbrakk>x \<in># darcs_mset (Node r xs); \<nexists>t1 e1. (t1,e1) \<in> fset xs \<and> x \<in># darcs_mset t1\<rbrakk>
\<Longrightarrow> x \<in> snd ` fset xs"
using mset_sum_elem by force
lemma darcs_mset_if_nsuc:
"\<lbrakk>x \<in># darcs_mset t; \<nexists>t1 e1. (t1,e1) \<in> fset (sucs t) \<and> x \<in># darcs_mset t1\<rbrakk>
\<Longrightarrow> x \<in> snd ` fset (sucs t)"
using darcs_mset_if_nchild[of x "root t" "sucs t"] by simp
lemma darcs_mset_if_snd[intro]: "x \<in> snd ` fset xs \<Longrightarrow> x \<in># darcs_mset (Node r xs)"
- by (induction xs) (auto simp: notin_fset)
+ by (induction xs) auto
lemma darcs_mset_suc_if_snd[intro]: "x \<in> snd ` fset (sucs t) \<Longrightarrow> x \<in># darcs_mset t"
using darcs_mset_if_snd[of x "sucs t" "root t"] by simp
lemma darcs_mset_if_child[intro]:
"\<lbrakk>(t1,e1) \<in> fset xs; x \<in># darcs_mset t1\<rbrakk> \<Longrightarrow> x \<in># darcs_mset (Node r xs)"
- by (induction xs) (auto simp: notin_fset)
+ by (induction xs) auto
lemma darcs_mset_if_suc[intro]:
"\<lbrakk>(t1,e1) \<in> fset (sucs t); x \<in># darcs_mset t1\<rbrakk> \<Longrightarrow> x \<in># darcs_mset t"
using darcs_mset_if_child[of t1 e1 "sucs t" x "root t"] by simp
lemma darcs_mset_sub_darcs: "set_mset (darcs_mset t) \<subseteq> darcs t"
proof(standard, induction t rule: darcs_mset.induct)
case (1 r xs)
then show ?case
proof(cases "x \<in> snd ` fset xs")
case False
then obtain t1 e1 where "(t1,e1) \<in> fset xs \<and> x \<in># darcs_mset t1"
using "1.prems" darcs_mset_if_nsnd[of x r] by blast
then show ?thesis using "1.IH" by force
qed(force)
qed
lemma darcs_sub_darcs_mset: "darcs t \<subseteq> set_mset (darcs_mset t)"
proof(standard, induction t rule: darcs_mset.induct)
case (1 r xs)
then show ?case
proof(cases "x \<in> snd ` fset xs")
case False
then obtain t1 e1 where "(t1,e1) \<in> fset xs \<and> x \<in> darcs t1"
using "1.prems" by force
then show ?thesis using "1.IH" by blast
qed(blast)
qed
lemma darcs_mset_eq_darcs[simp]: "set_mset (darcs_mset t) = darcs t"
using darcs_mset_sub_darcs darcs_sub_darcs_mset by force
lemma dverts_mset_elem:
"x \<in># dverts_mset (Node r xs) \<Longrightarrow> (\<exists>(t,e) \<in> fset xs. x \<in># dverts_mset t) \<or> x = r"
using mset_sum_elem by fastforce
lemma dverts_mset_if_nroot:
"\<lbrakk>x \<in># dverts_mset (Node r xs); x \<noteq> r\<rbrakk> \<Longrightarrow> \<exists>(t1,e1) \<in> fset xs. x \<in># dverts_mset t1"
using dverts_mset_elem[of x r xs] by blast
lemma dverts_mset_suc_if_nroot:
"\<lbrakk>x \<in># dverts_mset t; x \<noteq> root t\<rbrakk> \<Longrightarrow> \<exists>(t1,e1) \<in> fset (sucs t). x \<in># dverts_mset t1"
using dverts_mset_if_nroot[of x "root t" "sucs t"] by simp
lemma dverts_mset_if_nchild:
"\<lbrakk>x \<in># dverts_mset (Node r xs); \<nexists>t1 e1. (t1,e1) \<in> fset xs \<and> x \<in># dverts_mset t1\<rbrakk> \<Longrightarrow> x = r"
using mset_sum_elem by force
lemma dverts_mset_if_nsuc:
"\<lbrakk>x \<in># dverts_mset t; \<nexists>t1 e1. (t1,e1) \<in> fset (sucs t) \<and> x \<in># dverts_mset t1\<rbrakk> \<Longrightarrow> x = root t"
using dverts_mset_if_nchild[of x "root t" "sucs t"] by simp
lemma dverts_mset_if_root[intro]: "x = r \<Longrightarrow> x \<in># dverts_mset (Node r xs)"
by simp
lemma dverts_mset_suc_if_root[intro]: "x = root t \<Longrightarrow> x \<in># dverts_mset t"
using dverts_mset_if_root[of x "root t" "sucs t"] by simp
lemma dverts_mset_if_child[intro]:
"\<lbrakk>(t1,e1) \<in> fset xs; x \<in># dverts_mset t1\<rbrakk> \<Longrightarrow> x \<in># dverts_mset (Node r xs)"
- by (induction xs) (auto simp: notin_fset)
+ by (induction xs) auto
lemma dverts_mset_if_suc[intro]:
"\<lbrakk>(t1,e1) \<in> fset (sucs t); x \<in># dverts_mset t1\<rbrakk> \<Longrightarrow> x \<in># dverts_mset t"
using dverts_mset_if_child[of t1 e1 "sucs t" x "root t"] by simp
lemma dverts_mset_sub_dverts: "set_mset (dverts_mset t) \<subseteq> dverts t"
proof(standard, induction t)
case (Node r xs)
then show ?case
proof(cases "x = r")
case False
then obtain t1 e1 where "(t1,e1) \<in> fset xs \<and> x \<in># dverts_mset t1"
using Node.prems dverts_mset_if_nroot by fast
then show ?thesis using Node.IH by fastforce
qed(simp)
qed
lemma dverts_sub_dverts_mset: "dverts t \<subseteq> set_mset (dverts_mset t)"
proof(standard, induction t rule: dverts_mset.induct)
case (1 r xs)
then show ?case
proof(cases "x = r")
case False
then obtain t1 e1 where "(t1,e1) \<in> fset xs \<and> x \<in> dverts t1"
using "1.prems" by force
then show ?thesis using "1.IH" by blast
qed(simp)
qed
lemma dverts_mset_eq_dverts[simp]: "set_mset (dverts_mset t) = dverts t"
using dverts_mset_sub_dverts dverts_sub_dverts_mset by force
lemma mset_sum_count_le: "y \<in> fset Y \<Longrightarrow> count (f y) x \<le> count (\<Sum>y \<in> fset Y. f y) x"
- by (induction Y) (auto simp: notin_fset)
+ by (induction Y) auto
lemma darcs_mset_alt:
"darcs_mset (Node r xs) = (\<Sum>(t,e) \<in> fset xs. {#e#}) + (\<Sum>(t,e) \<in> fset xs. darcs_mset t)"
- by (induction xs) (auto simp: notin_fset)
+ by (induction xs) auto
lemma darcs_mset_ge_child:
"t1 \<in> fst ` fset xs \<Longrightarrow> count (darcs_mset t1) x \<le> count (darcs_mset (Node r xs)) x"
- by (induction xs) (force simp: notin_fset)+
+ by (induction xs) force+
lemma darcs_mset_ge_suc:
"t1 \<in> fst ` fset (sucs t) \<Longrightarrow> count (darcs_mset t1) x \<le> count (darcs_mset t) x"
using darcs_mset_ge_child[of t1 "sucs t" x "root t"] by simp
lemma darcs_mset_count_sum_aux:
"(\<Sum>(t1,e1) \<in> fset xs. count (darcs_mset t1) x) = count ((\<Sum>(t,e) \<in> fset xs. darcs_mset t)) x"
by (smt (verit, ccfv_SIG) count_add_mset count_sum multi_self_add_other_not_self
prod.case prod.case_distrib split_cong sum.cong)
lemma darcs_mset_count_sum_aux0:
"x \<notin> snd ` fset xs \<Longrightarrow> count ((\<Sum>(t, e)\<in>fset xs. {#e#})) x = 0"
- by (induction xs) (auto simp: notin_fset)
+ by (induction xs) auto
lemma darcs_mset_count_sum_eq:
"x \<notin> snd ` fset xs
\<Longrightarrow> (\<Sum>(t1,e1) \<in> fset xs. count (darcs_mset t1) x) = count (darcs_mset (Node r xs)) x"
unfolding darcs_mset_alt using darcs_mset_count_sum_aux darcs_mset_count_sum_aux0 by fastforce
lemma darcs_mset_count_sum_ge:
"(\<Sum>(t1,e1) \<in> fset xs. count (darcs_mset t1) x) \<le> count (darcs_mset (Node r xs)) x"
- by (induction xs) (auto simp: notin_fset split: prod.splits)
+ by (induction xs) (auto split: prod.splits)
lemma wf_darcs_alt: "wf_darcs t \<longleftrightarrow> (\<forall>x. count (darcs_mset t) x \<le> 1)"
unfolding wf_darcs_def by (metis count_greater_eq_one_iff dual_order.eq_iff linorder_le_cases)
lemma disjoint_darcs_simp:
"\<lbrakk>(t1,e1) \<in> fset xs; (t2,e2) \<in> fset xs; (t1,e1) \<noteq> (t2,e2); disjoint_darcs xs\<rbrakk>
\<Longrightarrow> (darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
by fast
lemma disjoint_darcs_single: "e \<notin> darcs t \<longleftrightarrow> disjoint_darcs {|(t,e)|}"
by simp
lemma disjoint_darcs_insert: "disjoint_darcs (finsert x xs) \<Longrightarrow> disjoint_darcs xs"
by simp fast
lemma wf_darcs_rec[dest]:
assumes "wf_darcs (Node r xs)" and "t1 \<in> fst ` fset xs"
shows "wf_darcs t1"
unfolding wf_darcs_def proof (rule ccontr)
assume asm: "\<not> (\<forall>x \<in># darcs_mset t1. count (darcs_mset t1) x = 1)"
then obtain x where x_def: "x \<in># darcs_mset t1" "count (darcs_mset t1) x \<noteq> 1"
by blast
then have "count (darcs_mset t1) x > 1" by (simp add: order_le_neq_trans)
then have "count (darcs_mset (Node r xs)) x > 1"
using assms(2) darcs_mset_ge_child[of t1 xs x] by simp
moreover have "x \<in># (darcs_mset (Node r xs))"
using x_def(1) assms(2) by fastforce
ultimately show False using assms(1) unfolding wf_darcs_def by simp
qed
lemma disjoint_darcs_if_wf_aux1: "\<lbrakk>wf_darcs (Node r xs); (t1,e1) \<in> fset xs\<rbrakk> \<Longrightarrow> e1 \<notin> darcs t1"
apply (induction xs)
- apply(auto simp: notin_fset wf_darcs_def split: if_splits prod.splits)[2]
+ apply(auto simp: wf_darcs_def split: if_splits prod.splits)[2]
by (metis UnI2 add_is_1 count_eq_zero_iff)
lemma fset_sum_ge_elem2:
"\<lbrakk>x \<in> fset X; y \<in> fset X; x \<noteq> y\<rbrakk> \<Longrightarrow> (f :: 'a \<Rightarrow> nat) x + f y \<le> (\<Sum>x \<in> fset X. f x)"
- by (induction X) (auto simp: notin_fset fset_sum_ge_elem)
+ by (induction X) (auto simp: fset_sum_ge_elem)
lemma darcs_children_count_ge2_aux:
assumes "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs" and "(t1,e1) \<noteq> (t2,e2)"
and "e \<in> darcs t1" and "e \<in> darcs t2"
shows "(\<Sum>(t1, e1)\<in>fset xs. count (darcs_mset t1) e) \<ge> 2"
proof -
have "2 \<le> 1 + count (darcs_mset t2) e"
using assms(2,5) by simp
also have "\<dots> \<le> count (darcs_mset t1) e + count (darcs_mset t2) e"
using assms(1,4) by simp
finally show ?thesis
using fset_sum_ge_elem2[OF assms(1-3), of "\<lambda>(t1,e1). count (darcs_mset t1) e"] by simp
qed
lemma darcs_children_count_ge2:
assumes "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs" and "(t1,e1) \<noteq> (t2,e2)"
and "e \<in> darcs t1" and "e \<in> darcs t2"
shows "count (darcs_mset (Node r xs)) e \<ge> 2"
using darcs_children_count_ge2_aux[OF assms] darcs_mset_count_sum_ge dual_order.trans by fast
lemma darcs_children_count_not1:
"\<lbrakk>(t1,e1) \<in> fset xs; (t2,e2) \<in> fset xs; (t1,e1) \<noteq> (t2,e2); e \<in> darcs t1; e \<in> darcs t2\<rbrakk>
\<Longrightarrow> count (darcs_mset (Node r xs)) e \<noteq> 1"
using darcs_children_count_ge2 by fastforce
lemma disjoint_darcs_if_wf_aux2:
assumes "wf_darcs (Node r xs)"
and "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs" and "(t1,e1) \<noteq> (t2,e2)"
shows "darcs t1 \<inter> darcs t2 = {}"
proof(rule ccontr)
assume "darcs t1 \<inter> darcs t2 \<noteq> {}"
then obtain e where e_def: "e \<in> darcs t1" "e \<in> darcs t2" by blast
then have "e \<in> darcs (Node r xs)" using assms(2) by force
then have "e \<in># darcs_mset (Node r xs)" using darcs_mset_eq_darcs by fast
then show False
using darcs_children_count_ge2[OF assms(2-4) e_def] assms(1) unfolding wf_darcs_def by simp
qed
lemma darcs_child_count_ge1:
"\<lbrakk>(t1,e1) \<in> fset xs; e2 \<in> darcs t1\<rbrakk> \<Longrightarrow> count (\<Sum>(t, e)\<in>fset xs. darcs_mset t) e2 \<ge> 1"
by (simp add: mset_sum_elemI)
lemma darcs_snd_count_ge1:
"(t2,e2) \<in> fset xs \<Longrightarrow> count (\<Sum>(t, e)\<in>fset xs. {#e#}) e2 \<ge> 1"
by (simp add: mset_sum_elemI)
lemma darcs_child_count_ge2:
"\<lbrakk>(t1,e1) \<in> fset xs; (t2,e2) \<in> fset xs; e2 \<in> darcs t1\<rbrakk> \<Longrightarrow> count (darcs_mset (Node r xs)) e2 \<ge> 2"
unfolding darcs_mset_alt
by (metis darcs_child_count_ge1 darcs_snd_count_ge1 add_mono count_union one_add_one)
lemma disjoint_darcs_if_wf_aux3:
assumes "wf_darcs (Node r xs)" and "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs"
shows "e2 \<notin> darcs t1"
proof
assume asm: "e2 \<in> darcs t1"
then have "e2 \<in> darcs (Node r xs)" using assms(2) by force
then have "e2 \<in># darcs_mset (Node r xs)" using darcs_mset_eq_darcs by fast
then show False using darcs_child_count_ge2 asm assms(1-3) unfolding wf_darcs_def by fastforce
qed
lemma darcs_snds_count_ge2_aux:
assumes "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs" and "(t1,e1) \<noteq> (t2,e2)" and "e1 = e2"
shows "count (\<Sum>(t, e)\<in>fset xs. {#e#}) e2 \<ge> 2"
using assms proof(induction xs)
case (insert x xs)
then consider "x = (t1,e1)" | "x = (t2,e2)" | "(t1,e1) \<in> fset xs" "(t2,e2) \<in> fset xs" by auto
then show ?case
proof(cases)
case 1
then have "count (\<Sum>(t, e)\<in>fset xs. {#e#}) e2 \<ge> 1"
using insert.prems(2,3) darcs_snd_count_ge1 by auto
- then show ?thesis using insert.prems(4) insert.hyps 1 by (auto simp: notin_fset)
+ then show ?thesis using insert.prems(4) insert.hyps 1 by auto
next
case 2
then have "count (\<Sum>(t, e)\<in>fset xs. {#e#}) e2 \<ge> 1"
using insert.prems(1,3,4) darcs_snd_count_ge1 by auto
- then show ?thesis using insert.prems(4) insert.hyps 2 by (auto simp: notin_fset)
+ then show ?thesis using insert.prems(4) insert.hyps 2 by auto
next
case 3
- then show ?thesis using insert.IH insert.prems(3,4) insert.hyps by (auto simp: notin_fset)
+ then show ?thesis using insert.IH insert.prems(3,4) insert.hyps by auto
qed
qed(simp)
lemma darcs_snds_count_ge2:
"\<lbrakk>(t1,e1) \<in> fset xs; (t2,e2) \<in> fset xs; (t1,e1) \<noteq> (t2,e2); e1 = e2\<rbrakk>
\<Longrightarrow> count (darcs_mset (Node r xs)) e2 \<ge> 2"
using darcs_snds_count_ge2_aux unfolding darcs_mset_alt by fastforce
lemma disjoint_darcs_if_wf_aux4:
assumes "wf_darcs (Node r xs)"
and "(t1,e1) \<in> fset xs"
and "(t2,e2) \<in> fset xs"
and "(t1,e1) \<noteq> (t2,e2)"
shows "e1 \<noteq> e2"
proof
assume asm: "e1 = e2"
have "e2 \<in># darcs_mset (Node r xs)" using assms(3) darcs_mset_if_snd by fastforce
then show False
using assms(1) darcs_snds_count_ge2[OF assms(2-4) asm] unfolding wf_darcs_def by simp
qed
lemma disjoint_darcs_if_wf_aux5:
"\<lbrakk>wf_darcs (Node r xs); (t1,e1) \<in> fset xs; (t2,e2) \<in> fset xs; (t1,e1) \<noteq> (t2,e2)\<rbrakk>
\<Longrightarrow>(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
by (auto dest: disjoint_darcs_if_wf_aux4 disjoint_darcs_if_wf_aux3 disjoint_darcs_if_wf_aux2)
lemma disjoint_darcs_if_wf_xs: "wf_darcs (Node r xs) \<Longrightarrow> disjoint_darcs xs"
by (auto dest: disjoint_darcs_if_wf_aux1 disjoint_darcs_if_wf_aux5)
lemma disjoint_darcs_if_wf: "wf_darcs t \<Longrightarrow> disjoint_darcs (sucs t)"
using disjoint_darcs_if_wf_xs[of "root t" "sucs t"] by simp
lemma wf_darcs'_if_darcs: "wf_darcs t \<Longrightarrow> wf_darcs' t"
proof(induction t)
case (Node r xs)
then show ?case using disjoint_darcs_if_wf_xs[OF Node.prems] by fastforce
qed
lemma wf_darcs_if_darcs'_aux:
"\<lbrakk>\<forall>(x,e) \<in> fset xs. wf_darcs x; disjoint_darcs xs\<rbrakk> \<Longrightarrow> wf_darcs (Node r xs)"
apply(simp split: prod.splits)
apply(induction xs)
- apply(auto simp: notin_fset wf_darcs_def count_eq_zero_iff)[2]
+ apply(auto simp: wf_darcs_def count_eq_zero_iff)[2]
by (fastforce dest: mset_sum_elem)+
lemma wf_darcs_if_darcs': "wf_darcs' t \<Longrightarrow> wf_darcs t"
proof(induction t)
case (Node r xs)
then show ?case using wf_darcs_if_darcs'_aux[of xs] by fastforce
qed
corollary wf_darcs_iff_darcs': "wf_darcs t \<longleftrightarrow> wf_darcs' t"
using wf_darcs_if_darcs' wf_darcs'_if_darcs by blast
lemma disjoint_darcs_subset:
assumes "xs |\<subseteq>| ys" and "disjoint_darcs ys"
shows "disjoint_darcs xs"
proof (rule ccontr)
assume "\<not> disjoint_darcs xs"
then obtain x e1 y e2 where x_def: "(x,e1) \<in> fset xs" "(y,e2) \<in> fset xs"
"e1 \<in> darcs x \<or> (darcs x \<union> {e1}) \<inter> (darcs y \<union> {e2}) \<noteq> {} \<and> (x,e1)\<noteq>(y,e2)"
by blast
have "(x,e1) \<in> fset ys" "(y,e2) \<in> fset ys" using x_def(1,2) assms(1) less_eq_fset.rep_eq by fast+
then show False using assms(2) x_def(3) by fast
qed
lemma disjoint_darcs_img:
assumes "disjoint_darcs xs" and "\<forall>(t,e) \<in> fset xs. darcs (f t) \<subseteq> darcs t"
shows "disjoint_darcs ((\<lambda>(t,e). (f t,e)) |`| xs)" (is "disjoint_darcs ?xs")
proof (rule ccontr)
assume "\<not> disjoint_darcs ?xs"
then obtain x1 e1 y1 e2 where asm: "(x1,e1) \<in> fset ?xs" "(y1,e2) \<in> fset ?xs"
"e1 \<in> darcs x1 \<or> (darcs x1 \<union> {e1}) \<inter> (darcs y1 \<union> {e2}) \<noteq> {} \<and> (x1,e1)\<noteq>(y1,e2)"
by blast
then obtain x2 where x2_def: "f x2 = x1" "(x2,e1) \<in> fset xs" by auto
obtain y2 where y2_def: "f y2 = y1" "(y2,e2) \<in> fset xs" using asm(2) by auto
have "darcs x1 \<subseteq> darcs x2" using assms(2) x2_def by fast
moreover have "darcs y1 \<subseteq> darcs y2" using assms(2) y2_def by fast
ultimately have "\<not> disjoint_darcs xs" using asm(3) x2_def y2_def by fast
then show False using assms(1) by blast
qed
lemma dverts_mset_count_sum_ge:
"(\<Sum>(t1,e1) \<in> fset xs. count (dverts_mset t1) x) \<le> count (dverts_mset (Node r xs)) x"
- by (induction xs) (auto simp: notin_fset)
+ by (induction xs) auto
lemma dverts_children_count_ge2_aux:
assumes "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs" and "(t1,e1) \<noteq> (t2,e2)"
and "x \<in> dverts t1" and "x \<in> dverts t2"
shows "(\<Sum>(t1, e1)\<in>fset xs. count (dverts_mset t1) x) \<ge> 2"
proof -
have "2 \<le> count (dverts_mset t1) x + 1" using assms(4) by simp
also have "\<dots> \<le> count (dverts_mset t1) x + count (dverts_mset t2) x" using assms(5) by simp
finally show ?thesis
using fset_sum_ge_elem2[OF assms(1-3), of "\<lambda>(t1,e1). count (dverts_mset t1) x"] by simp
qed
lemma dverts_children_count_ge2:
assumes "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs" and "(t1,e1) \<noteq> (t2,e2)"
and "x \<in> dverts t1" and "x \<in> dverts t2"
shows "count (dverts_mset (Node r xs)) x \<ge> 2"
using dverts_children_count_ge2_aux[OF assms] dverts_mset_count_sum_ge le_trans by fast
lemma disjoint_dverts_if_wf_aux:
assumes "wf_dverts (Node r xs)"
and "(t1,e1) \<in> fset xs" and "(t2,e2) \<in> fset xs" and "(t1,e1) \<noteq> (t2,e2)"
shows "dverts t1 \<inter> dverts t2 = {}"
proof (rule ccontr)
assume "dverts t1 \<inter> dverts t2 \<noteq> {}"
then obtain x where x_def: "x \<in> dverts t1" "x \<in> dverts t2" by blast
then have "2 \<le> count (dverts_mset (Node r xs)) x"
using dverts_children_count_ge2[OF assms(2-4)] by blast
moreover have "x \<in># (dverts_mset (Node r xs))" using assms(2) x_def(1) by fastforce
ultimately show False using assms(1) unfolding wf_dverts_def by fastforce
qed
lemma disjoint_dverts_if_wf:
"wf_dverts (Node r xs)
\<Longrightarrow> \<forall>(x,e1) \<in> fset xs. \<forall>(y,e2) \<in> fset xs. (dverts x \<inter> dverts y = {} \<or> (x,e1)=(y,e2))"
using disjoint_dverts_if_wf_aux by fast
lemma disjoint_dverts_if_wf_sucs:
"wf_dverts t
\<Longrightarrow> \<forall>(x,e1) \<in> fset (sucs t). \<forall>(y,e2) \<in> fset (sucs t).
(dverts x \<inter> dverts y = {} \<or> (x,e1)=(y,e2))"
using disjoint_dverts_if_wf[of "root t" "sucs t"] by simp
lemma dverts_child_count_ge1:
"\<lbrakk>(t1,e1) \<in> fset xs; x \<in> dverts t1\<rbrakk> \<Longrightarrow> count (\<Sum>(t, e)\<in>fset xs. dverts_mset t) x \<ge> 1"
by (simp add: mset_sum_elemI)
lemma root_not_child_if_wf_dverts: "\<lbrakk>wf_dverts (Node r xs); (t1,e1) \<in> fset xs\<rbrakk> \<Longrightarrow> r \<notin> dverts t1"
by (fastforce dest: dverts_child_count_ge1 simp: wf_dverts_def)
lemma root_not_child_if_wf_dverts': "wf_dverts (Node r xs) \<Longrightarrow> \<forall>(t1,e1) \<in> fset xs. r \<notin> dverts t1"
by (fastforce dest: dverts_child_count_ge1 simp: wf_dverts_def)
lemma dverts_mset_ge_child:
"t1 \<in> fst ` fset xs \<Longrightarrow> count (dverts_mset t1) x \<le> count (dverts_mset (Node r xs)) x"
- by (induction xs) (force simp: notin_fset)+
+ by (induction xs) force+
lemma wf_dverts_rec[dest]:
assumes "wf_dverts (Node r xs)" and "t1 \<in> fst ` fset xs"
shows "wf_dverts t1"
unfolding wf_dverts_def proof (rule ccontr)
assume asm: "\<not> (\<forall>x \<in># dverts_mset t1. count (dverts_mset t1) x = 1)"
then obtain x where x_def: "x \<in># dverts_mset t1" "count (dverts_mset t1) x \<noteq> 1"
by blast
then have "count (dverts_mset t1) x > 1" by (simp add: order_le_neq_trans)
then have "count (dverts_mset (Node r xs)) x > 1"
using assms(2) dverts_mset_ge_child[of t1 xs x r] by simp
moreover have "x \<in># (dverts_mset (Node r xs))"
using x_def(1) assms(2) by fastforce
ultimately show False using assms(1) unfolding wf_dverts_def by fastforce
qed
lemma wf_dverts'_if_dverts: "wf_dverts t \<Longrightarrow> wf_dverts' t"
proof(induction t)
case (Node r xs)
then have "\<forall>(x,e1)\<in>fset xs. wf_dverts' x" by auto
then show ?case
using disjoint_dverts_if_wf[OF Node.prems] root_not_child_if_wf_dverts'[OF Node.prems]
by fastforce
qed
lemma wf_dverts_if_dverts'_aux:
"\<lbrakk>\<forall>(x,e) \<in> fset xs. wf_dverts x;
\<forall>(x,e1) \<in> fset xs. r \<notin> dverts x \<and> (\<forall>(y,e2) \<in> fset xs.
(dverts x \<inter> dverts y = {} \<or> (x,e1)=(y,e2)))\<rbrakk>
\<Longrightarrow> wf_dverts (Node r xs)"
apply(simp split: prod.splits)
apply(induction xs)
- apply(auto simp: notin_fset wf_dverts_def count_eq_zero_iff)[2]
+ apply(auto simp: wf_dverts_def count_eq_zero_iff)[2]
by (fastforce dest: mset_sum_elem)+
lemma wf_dverts_if_dverts': "wf_dverts' t \<Longrightarrow> wf_dverts t"
proof(induction t)
case (Node r xs)
then show ?case using wf_dverts_if_dverts'_aux[of xs] by fastforce
qed
corollary wf_dverts_iff_dverts': "wf_dverts t \<longleftrightarrow> wf_dverts' t"
using wf_dverts_if_dverts' wf_dverts'_if_dverts by blast
lemma wf_dverts_sub:
assumes "xs |\<subseteq>| ys" and "wf_dverts (Node r ys)"
shows "wf_dverts (Node r xs)"
proof -
have "ys |\<union>| xs = ys" using assms(1) by blast
then have "wf_dverts (Node r (ys |\<union>| xs))" using assms(2) by simp
then show ?thesis unfolding wf_dverts_iff_dverts' by fastforce
qed
lemma count_subset_le:
"xs |\<subseteq>| ys \<Longrightarrow> count (\<Sum>x \<in> fset xs. f x) a \<le> count (\<Sum>x \<in> fset ys. f x) a"
proof(induction ys arbitrary: xs)
case (insert y ys)
then show ?case
proof(cases "y |\<in>| xs")
case True
then obtain xs' where xs'_def: "finsert y xs' = xs" "y |\<notin>| xs'"
by blast
then have "xs' |\<subseteq>| ys" using insert.prems by blast
have "count (\<Sum>x \<in> fset xs. f x) a = count (\<Sum>x \<in> fset xs'. f x) a + count (f y) a"
- using xs'_def by (auto simp: notin_fset)
+ using xs'_def by auto
also have "\<dots> \<le> count (\<Sum>x \<in> fset ys. f x) a + count (f y) a"
using \<open>xs' |\<subseteq>| ys\<close> insert.IH by simp
also have "\<dots> = count (\<Sum>x \<in> fset (finsert y ys). f x) a"
- using insert.hyps by (auto simp: notin_fset)
+ using insert.hyps by auto
finally show ?thesis .
next
case False
then have "count (\<Sum>x \<in> fset xs. f x) a \<le> count (\<Sum>x \<in> fset ys. f x) a"
using insert.prems insert.IH by blast
- then show ?thesis using insert.hyps by (auto simp: notin_fset)
+ then show ?thesis using insert.hyps by auto
qed
qed(simp)
lemma darcs_mset_count_le_subset:
"xs |\<subseteq>| ys \<Longrightarrow> count (darcs_mset (Node r' xs)) x \<le> count (darcs_mset (Node r ys)) x"
using count_subset_le by fastforce
lemma wf_darcs_sub: "\<lbrakk>xs |\<subseteq>| ys; wf_darcs (Node r' ys)\<rbrakk> \<Longrightarrow> wf_darcs (Node r xs)"
unfolding wf_darcs_def using darcs_mset_count_le_subset
by (smt (verit, best) count_greater_eq_one_iff le_trans verit_la_disequality)
lemma wf_darcs_sucs: "\<lbrakk>wf_darcs t; x \<in> fset (sucs t)\<rbrakk> \<Longrightarrow> wf_darcs (Node r {|x|})"
using wf_darcs_sub[of "{|x|}" "sucs t" "root t"] by (simp add: less_eq_fset.rep_eq)
lemma size_fset_alt:
"size_fset (size_prod snd (\<lambda>_. 0)) (map_prod (\<lambda>t. (t, size t)) (\<lambda>x. x) |`| xs)
= (\<Sum>(x,y)\<in> fset xs. size x + 2)"
proof -
have "size_fset (size_prod snd (\<lambda>_. 0)) (map_prod (\<lambda>t. (t, size t)) (\<lambda>x. x) |`| xs)
= (\<Sum>u\<in>(\<lambda>(x,y). ((x,size x), y)) ` fset xs. snd (fst u) + 2)"
by (simp add: size_prod_simp map_prod_def)
also have "\<dots> = (\<Sum>(x,y) \<in> fset xs. size x + 2)"
using case_prod_beta' comm_monoid_add_class.sum.eq_general
by (smt (verit, del_insts) Pair_inject fstI imageE imageI prod_eqI snd_conv)
finally show ?thesis .
qed
lemma dtree_size_alt: "size (Node r xs) = (\<Sum>(x,y)\<in> fset xs. size x + 2) + 1"
using size_fset_alt by auto
lemma dtree_size_eq_root: "size (Node r xs) = size (Node r' xs)"
by auto
lemma size_combine_decr: "size (Node (r@root t1) (sucs t1)) < size (Node r {|(t1, e1)|})"
using dtree_size_eq_root[of "r@root t1" "sucs t1" "root t1"] by simp
lemma size_le_if_child_subset: "xs |\<subseteq>| ys \<Longrightarrow> size (Node r xs) \<le> size (Node v ys)"
unfolding dtree_size_alt by (simp add: dtree_size_alt less_eq_fset.rep_eq sum.subset_diff)
lemma size_le_if_sucs_subset: "sucs t1 |\<subseteq>| sucs t2 \<Longrightarrow> size t1 \<le> size t2"
using size_le_if_child_subset[of "sucs t1" "sucs t2" "root t1" "root t2"] by simp
lemma combine_uneq: "Node r {|(t1, e1)|} \<noteq> Node (r@root t1) (sucs t1)"
using size_combine_decr[of r t1 e1] by fastforce
lemma child_uneq: "t \<in> fst ` fset xs \<Longrightarrow> Node r xs \<noteq> t"
- using dtree_size_decr_aux' by fast
+ using dtree_size_decr_aux' by fastforce
lemma suc_uneq: "t1 \<in> fst ` fset (sucs t) \<Longrightarrow> t \<noteq> t1"
using child_uneq[of t1 "sucs t" "root t"] by simp
lemma singleton_uneq: "Node r {|(t,e)|} \<noteq> t"
using child_uneq[of t] by simp
lemma child_uneq': "t \<in> fst ` fset xs \<Longrightarrow> Node r xs \<noteq> Node v (sucs t)"
using dtree_size_decr_aux'[of t] dtree_size_eq_root[of "root t" "sucs t"] by auto
lemma suc_uneq': "t1 \<in> fst ` fset (sucs t) \<Longrightarrow> t \<noteq> Node v (sucs t1)"
using child_uneq'[of t1 "sucs t" "root t"] by simp
lemma singleton_uneq': "Node r {|(t,e)|} \<noteq> Node v (sucs t)"
using child_uneq'[of t] by simp
lemma singleton_suc: "t \<in> fst ` fset (sucs (Node r {|(t,e)|}))"
by simp
lemma fcard_image_le: "fcard (f |`| xs) \<le> fcard xs"
by (simp add: FSet.fcard.rep_eq card_image_le)
lemma sum_img_le:
assumes "\<forall>t \<in> fst ` fset xs. (g::'a \<Rightarrow> nat) (f t) \<le> g t"
shows "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x) \<le> (\<Sum>(x,y)\<in> fset xs. g x)"
using assms proof(induction xs)
case (insert x xs)
obtain t e where t_def: "x = (t,e)" by fastforce
then show ?case
proof(cases "(f t,e) \<notin> fset ((\<lambda>(t,e). (f t, e)) |`| xs)")
case True
then have "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| (finsert x xs)). g x)
= g (f t) + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using t_def by auto
also have "\<dots> \<le> g t + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using insert.prems t_def by auto
also have "\<dots> \<le> g t + (\<Sum>(x,y)\<in> fset xs. g x)" using insert by simp
- finally show ?thesis using insert.hyps t_def notin_fset by fastforce
+ finally show ?thesis using insert.hyps t_def by fastforce
next
case False
then have "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| (finsert x xs)). g x)
= (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
- by (metis (no_types, lifting) t_def fimage_finsert finsert_absorb notin_fset prod.case)
+ by (metis (no_types, lifting) t_def fimage_finsert finsert_absorb prod.case)
also have "\<dots> \<le> (\<Sum>(x,y)\<in> fset xs. g x)" using insert by simp
- finally show ?thesis using insert.hyps t_def notin_fset by fastforce
+ finally show ?thesis using insert.hyps t_def by fastforce
qed
qed (simp)
lemma dtree_size_img_le:
assumes "\<forall>t \<in> fst ` fset xs. size (f t) \<le> size t"
shows "size (Node r ((\<lambda>(t,e). (f t, e)) |`| xs)) \<le> size (Node r xs)"
using sum_img_le[of xs "\<lambda>x. size x + 2"] dtree_size_alt assms
by (metis (mono_tags, lifting) add_right_mono)
lemma sum_img_lt:
assumes "\<forall>t \<in> fst ` fset xs. (g::'a \<Rightarrow> nat) (f t) \<le> g t"
and "\<exists>t \<in> fst ` fset xs. g (f t) < g t"
and "\<forall>t \<in> fst ` fset xs. g t > 0"
shows "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x) < (\<Sum>(x,y)\<in> fset xs. g x)"
using assms proof(induction xs)
case (insert x xs)
obtain t e where t_def: "x = (t,e)" by fastforce
then show ?case
proof(cases "(f t,e) \<notin> fset ((\<lambda>(t,e). (f t, e)) |`| xs)")
case f_notin_xs: True
show ?thesis
proof(cases "g (f t) < g t")
case True
have "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| (finsert x xs)). g x)
= g (f t) + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using t_def f_notin_xs by auto
also have "\<dots> < g t + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using True by simp
also have "\<dots> \<le> g t + (\<Sum>(x,y)\<in> fset xs. g x)" using sum_img_le insert.prems(1) by auto
- finally show ?thesis using insert.hyps t_def notin_fset by fastforce
+ finally show ?thesis using insert.hyps t_def by fastforce
next
case False
then have 0: "\<exists>t \<in> fst ` fset xs. g (f t) < g t" using insert.prems(2) t_def by simp
have "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| (finsert x xs)). g x)
= g (f t) + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using t_def f_notin_xs by auto
also have "\<dots> \<le> g t + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using t_def insert.prems(1) by simp
also have "\<dots> < g t + (\<Sum>(x,y)\<in> fset xs. g x)" using insert.IH insert.prems(1,3) 0 by simp
- finally show ?thesis using insert.hyps t_def notin_fset by fastforce
+ finally show ?thesis using insert.hyps t_def by fastforce
qed
next
case False
then have "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| (finsert x xs)). g x)
= (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
- by (metis (no_types, lifting) t_def fimage_finsert finsert_absorb notin_fset prod.case)
+ by (metis (no_types, lifting) t_def fimage_finsert finsert_absorb prod.case)
also have "\<dots> \<le> (\<Sum>(x,y)\<in> fset xs. g x)" using sum_img_le insert.prems(1) by auto
also have "\<dots> < g t + (\<Sum>(x,y)\<in> fset xs. g x)" using insert.prems(3) t_def by simp
- finally show ?thesis using insert.hyps t_def notin_fset by fastforce
+ finally show ?thesis using insert.hyps t_def by fastforce
qed
qed (simp)
lemma dtree_size_img_lt:
assumes "\<forall>t \<in> fst ` fset xs. size (f t) \<le> size t"
and "\<exists>t \<in> fst ` fset xs. size (f t) < size t"
shows "size (Node r ((\<lambda>(t,e). (f t, e)) |`| xs)) < size (Node r xs)"
proof -
have 0: "\<forall>t \<in> fst ` fset xs. size (f t) + 2 \<le> size t + 2" using assms(1) by simp
have "\<forall>t\<in>fst ` fset xs. 0 < size t + 2" by simp
then show ?thesis using sum_img_lt[OF 0] dtree_size_alt assms(2) by (smt (z3) add_less_mono1)
qed
lemma sum_img_eq:
assumes "\<forall>t \<in> fst ` fset xs. (g::'a \<Rightarrow> nat) (f t) = g t"
and "fcard ((\<lambda>(t,e). (f t, e)) |`| xs) = fcard xs"
shows "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x) = (\<Sum>(x,y)\<in> fset xs. g x)"
using assms proof(induction xs)
case (insert x xs)
obtain t e where t_def: "x = (t,e)" by fastforce
then have 0: "(f t,e) \<notin> fset ((\<lambda>(t,e). (f t, e)) |`| xs)"
- using insert.prems(2) insert.hyps notin_fset fcard_finsert_if fcard_image_le
+ using insert.prems(2) insert.hyps fcard_finsert_if fcard_image_le
by (metis (mono_tags, lifting) case_prod_conv fimage_finsert leD lessI)
then have 1: "fcard ((\<lambda>(t,e). (f t, e)) |`| xs) = fcard xs "
- using insert.prems(2) insert.hyps t_def notin_fset Suc_inject
+ using insert.prems(2) insert.hyps t_def Suc_inject
by (metis (mono_tags, lifting) fcard_finsert_if fimage_finsert old.prod.case)
have "(\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| (finsert x xs)). g x)
= g (f t) + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using t_def 0 by auto
also have "\<dots> = g t + (\<Sum>(x,y)\<in> fset ((\<lambda>(t,e). (f t, e)) |`| xs). g x)"
using insert.prems t_def by auto
also have "\<dots> = g t + (\<Sum>(x,y)\<in> fset xs. g x)" using insert.IH 1 insert.prems(1) by simp
- finally show ?case using insert.hyps t_def notin_fset by fastforce
+ finally show ?case using insert.hyps t_def by fastforce
qed (simp)
lemma elem_neq_if_fset_neq:
"((\<lambda>(t,e). (f t, e)) |`| xs) \<noteq> xs \<Longrightarrow> \<exists>t \<in> fst ` fset xs. f t \<noteq> t"
by (smt (verit, ccfv_threshold) case_prod_eta case_prod_eta fimage.rep_eq fset_inject fst_conv
image_cong image_ident image_subset_iff old.prod.case prod.case_distrib split_cong subsetI)
lemma ffold_commute_supset:
"\<lbrakk>xs |\<subseteq>| ys; P ys; \<And>ys xs. \<lbrakk>xs |\<subseteq>| ys; P ys\<rbrakk> \<Longrightarrow> P xs;
\<And>xs. comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not>Q a b \<or> \<not>P xs then b else R a b)\<rbrakk>
\<Longrightarrow> ffold (\<lambda>a b. if a \<notin> fset ys \<or> \<not>Q a b \<or> \<not>P ys then b else R a b) acc xs
= ffold (\<lambda>a b. if a \<notin> fset xs \<or> \<not>Q a b \<or> \<not>P xs then b else R a b) acc xs"
proof(induction xs arbitrary: ys)
+ case empty
+ show ?case
+ unfolding empty.prems(4)[THEN comp_fun_commute.ffold_empty]
+ by simp
+next
case (insert x xs)
let ?f = "\<lambda>a b. if a \<notin> fset ys \<or> \<not>Q a b \<or> \<not>P ys then b else R a b"
let ?f' = "\<lambda>a b. if a \<notin> fset xs \<or> \<not>Q a b \<or> \<not>P xs then b else R a b"
let ?f1 = "\<lambda>a b. if a \<notin> fset (finsert x xs) \<or> \<not>Q a b \<or> \<not>P (finsert x xs) then b else R a b"
have 0: "P (finsert x xs)" using insert.prems by simp
have 1: "xs |\<subseteq>| (finsert x xs)" by blast
have 2: "comp_fun_commute ?f1" using insert.prems(4) by blast
- have 3: "x \<in> fset ys" using insert.prems(1) notin_fset by fastforce
+ have 3: "x \<in> fset ys" using insert.prems(1) by fastforce
have "ffold ?f acc (finsert x xs) = ?f x (ffold ?f acc xs)"
using comp_fun_commute.ffold_finsert[of ?f] insert.prems(4) insert.hyps by blast
also have "\<dots> = ?f x (ffold ?f' acc xs)" using insert.IH[of ys] insert.prems by fastforce
also have "\<dots> = ?f x (ffold ?f1 acc xs)" using insert.IH[OF 1 0] insert.prems(3,4) by presburger
also have "\<dots> = ?f1 x (ffold ?f1 acc xs)" using 0 3 insert.prems(2) by fastforce
also have "\<dots> = ffold ?f1 acc (finsert x xs)"
using comp_fun_commute.ffold_finsert[of ?f1 x xs] 2 insert.hyps by presburger
finally show ?case .
-qed (smt (z3) comp_fun_commute.ffold_empty)
+qed
lemma ffold_eq_fold: "\<lbrakk>finite xs; f = g\<rbrakk> \<Longrightarrow> ffold f acc (Abs_fset xs) = Finite_Set.fold g acc xs"
unfolding ffold_def by (simp add: Abs_fset_inverse)
lemma Abs_fset_sub_if_sub:
assumes "finite ys" and "xs \<subseteq> ys"
shows "Abs_fset xs |\<subseteq>| Abs_fset ys"
proof (rule ccontr)
assume "\<not>(Abs_fset xs |\<subseteq>| Abs_fset ys)"
then obtain x where x_def: "x |\<in>| Abs_fset xs" "x |\<notin>| Abs_fset ys" by blast
- then have "x \<in> fset (Abs_fset xs) \<and> x \<notin> fset (Abs_fset ys)" using notin_fset by fast
+ then have "x \<in> fset (Abs_fset xs) \<and> x \<notin> fset (Abs_fset ys)" by fast
moreover have "finite xs" using assms finite_subset by auto
ultimately show False using assms Abs_fset_inverse by blast
qed
lemma fold_commute_supset:
assumes "finite ys" and "xs \<subseteq> ys" and "P ys" and "\<And>ys xs. \<lbrakk>xs \<subseteq> ys; P ys\<rbrakk> \<Longrightarrow> P xs"
and "\<And>xs. comp_fun_commute (\<lambda>a b. if a \<notin> xs \<or> \<not>Q a b \<or> \<not>P xs then b else R a b)"
shows "Finite_Set.fold (\<lambda>a b. if a \<notin> ys \<or> \<not>Q a b \<or> \<not>P ys then b else R a b) acc xs
= Finite_Set.fold (\<lambda>a b. if a \<notin> xs \<or> \<not>Q a b \<or> \<not>P xs then b else R a b) acc xs"
proof -
let ?f = "\<lambda>a b. if a \<notin> ys \<or> \<not>Q a b \<or> \<not>P ys then b else R a b"
let ?f' = "\<lambda>a b. if a \<notin> xs \<or> \<not>Q a b \<or> \<not>P xs then b else R a b"
let ?P = "\<lambda>xs. P (fset xs)"
let ?g = "\<lambda>a b. if a \<notin> fset (Abs_fset ys) \<or> \<not>Q a b \<or> \<not>(?P (Abs_fset ys)) then b else R a b"
let ?g' = "\<lambda>a b. if a \<notin> fset (Abs_fset xs) \<or> \<not>Q a b \<or> \<not>(?P (Abs_fset xs)) then b else R a b"
have 0: "finite xs" using assms(1,2) finite_subset by auto
then have 1: "Abs_fset xs |\<subseteq>| (Abs_fset ys)" using Abs_fset_sub_if_sub[OF assms(1,2)] by blast
have 2: "?P (Abs_fset ys)" by (simp add: Abs_fset_inverse assms(1,3))
have 3: "\<And>ys xs. \<lbrakk>xs |\<subseteq>| ys; ?P ys\<rbrakk> \<Longrightarrow> ?P xs" by (simp add: assms(4) less_eq_fset.rep_eq)
have 4: "\<And>xs. comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not>Q a b \<or> \<not>(?P xs) then b else R a b)"
using assms(5) by (simp add: less_eq_fset.rep_eq)
have "?f' = ?g'" by (simp add: Abs_fset_inverse 0)
have "?f = ?g" by (simp add: Abs_fset_inverse assms(1))
then have "Finite_Set.fold (\<lambda>a b. if a \<notin> ys \<or> \<not>Q a b \<or> \<not>P ys then b else R a b) acc xs
= ffold ?g acc (Abs_fset xs)" by (simp add: 0 ffold_eq_fold)
also have "\<dots> = ffold ?g' acc (Abs_fset xs)"
using ffold_commute_supset[OF 1, of ?P, OF 2 3 4] by simp
finally show ?thesis using \<open>?f' = ?g'\<close> by (simp add: 0 ffold_eq_fold)
qed
lemma dtail_commute_aux:
fixes r xs e def
defines "f \<equiv> (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r xs)
then b else dtail x def)"
shows "(f y \<circ> f x) z = (f x \<circ> f y) z"
proof -
obtain y1 y2 where y_def: "y = (y1,y2)" by fastforce
obtain x1 x2 where x_def: "x = (x1,x2)" by fastforce
show ?thesis
proof(cases "(x1,x2) \<in> fset xs \<and> (y1,y2) \<in> fset xs")
case 0: True
then show ?thesis
proof(cases "e \<in> darcs x1 \<and> e \<in> darcs y1")
case True
then have 1: "x1 = y1 \<or> \<not>wf_darcs (Node r xs)" using 0 disjoint_darcs_if_wf_aux2 by fast
then show ?thesis using assms by (cases "x1=y1")(auto simp: x_def y_def)
next
case False
then show ?thesis using assms by (simp add: x_def y_def)
qed
next
case False
then show ?thesis using assms by (simp add: x_def y_def)
qed
qed
lemma dtail_commute:
"comp_fun_commute (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r xs)
then b else dtail x def)"
using dtail_commute_aux[of xs] by unfold_locales blast
lemma dtail_f_alt:
assumes "P = (\<lambda>xs. wf_darcs (Node r xs))"
and "Q = (\<lambda>(t1,e1) b. e \<in> darcs t1)"
and "R = (\<lambda>(t1,e1) b. dtail t1 def)"
shows "(\<lambda>(t1,e1) b. if (t1,e1) \<notin> fset xs \<or> e \<notin> darcs t1\<or> \<not>wf_darcs (Node r xs)
then b else dtail t1 def)
= (\<lambda>a b. if a \<notin> fset xs \<or> \<not> Q a b \<or> \<not> P xs then b else R a b)"
using assms by fast
lemma dtail_f_alt_commute:
assumes "P = (\<lambda>xs. wf_darcs (Node r xs))"
and "Q = (\<lambda>(t1,e1) b. e \<in> darcs t1)"
and "R = (\<lambda>(t1,e1) b. dtail t1 def)"
shows "comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not> Q a b \<or> \<not> P xs then b else R a b)"
using dtail_commute[of xs e r def] dtail_f_alt[OF assms] by simp
lemma dtail_ffold_supset:
assumes "xs |\<subseteq>| ys" and "wf_darcs (Node r ys)"
shows "ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset ys \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r ys)
then b else dtail x def) def xs
= ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r xs)
then b else dtail x def) def xs"
proof -
let ?P = "\<lambda>xs. wf_darcs (Node r xs)"
let ?Q = "\<lambda>(t1,e1) b. e \<in> darcs t1"
let ?R = "\<lambda>(t1,e1) b. dtail t1 def"
have 0: "\<And>xs. comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not> ?Q a b \<or> \<not> ?P xs then b else ?R a b)"
using dtail_f_alt_commute by fast
have "ffold (\<lambda>a b. if a \<notin> fset ys \<or> \<not> ?Q a b \<or> \<not> ?P ys then b else ?R a b) def xs
= ffold (\<lambda>a b. if a \<notin> fset xs \<or> \<not> ?Q a b \<or> \<not> ?P xs then b else ?R a b) def xs"
using ffold_commute_supset[OF assms(1),of ?P ?Q ?R,OF assms(2) wf_darcs_sub 0] by simp
then show ?thesis using dtail_f_alt[of ?P r ?Q e ?R] by simp
qed
lemma dtail_in_child_eq_child_ffold:
assumes "(t,e1) \<in> fset xs" and "e \<in> darcs t" and "wf_darcs (Node r xs)"
shows "ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r xs)
then b else dtail x def) def xs
= dtail t def"
using assms proof(induction xs)
case (insert x' xs)
let ?f = "(\<lambda>(x,e2) b.
if (x,e2) \<notin> fset (finsert x' xs) \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r (finsert x' xs))
then b else dtail x def)"
let ?f' = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r xs)
then b else dtail x def)"
obtain x e3 where x_def: "x' = (x,e3)" by fastforce
show ?case
proof(cases "x=t")
case True
have "ffold ?f def (finsert x' xs) = (?f x' (ffold ?f def xs))"
using comp_fun_commute.ffold_finsert[of ?f x' xs def] dtail_commute insert.hyps by fast
also have "\<dots> = (?f (x,e3) (ffold ?f def xs))" using x_def by blast
also have "\<dots> = dtail x def" using x_def insert.prems(2,3) True by fastforce
finally show ?thesis using True by blast
next
case False
then have 0: "(t,e1) \<in> fset xs" using insert.prems(1) x_def by simp
have 1: "wf_darcs (Node r xs)" using wf_darcs_sub[OF fsubset_finsertI insert.prems(3)] .
have 2: "xs |\<subseteq>| (finsert x' xs)" by blast
have "(x,e3) \<in> fset (finsert x' xs)" using x_def by simp
have 3: "e \<notin> darcs x" using insert.prems(1-3) disjoint_darcs_if_wf x_def False by fastforce
have "ffold ?f def (finsert x' xs) = (?f x' (ffold ?f def xs))"
using comp_fun_commute.ffold_finsert[of ?f x' xs def] dtail_commute insert.hyps by fast
also have "\<dots> = (?f (x,e3) (ffold ?f def xs))" using x_def by blast
also have "\<dots> = (ffold ?f def xs)" using 3 by fastforce
also have "\<dots> = (ffold ?f' def xs)"
using dtail_ffold_supset[of xs "finsert x' xs"] insert.prems(3) 2 by simp
also have "\<dots> = dtail t def" using insert.IH 0 1 insert.prems(2) by fast
finally show ?thesis .
qed
qed(simp)
lemma dtail_in_child_eq_child:
assumes "(t,e1) \<in> fset xs" and "e \<in> darcs t" and "wf_darcs (Node r xs)"
shows "dtail (Node r xs) def e = dtail t def e"
using assms dtail_in_child_eq_child_ffold[OF assms] disjoint_darcs_if_wf_aux3 by fastforce
lemma dtail_ffold_notelem_eq_def:
assumes "\<forall>(t,e1) \<in> fset xs. e \<notin> darcs t"
shows "ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset ys \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r ys)
then b else dtail x def) def xs = def"
using assms proof(induction xs)
+ case empty
+ show ?case
+ unfolding dtail_commute[THEN comp_fun_commute.ffold_empty]
+ by simp
+next
case (insert x' xs)
obtain x e3 where x_def: "x' = (x,e3)" by fastforce
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset ys \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r ys)
then b else dtail x def)"
have "ffold ?f def (finsert x' xs) = ?f x' (ffold ?f def xs)"
using comp_fun_commute.ffold_finsert[of ?f x' xs] dtail_commute insert.hyps by fast
also have "\<dots> = (ffold ?f def xs)" using insert.prems by auto
also have "\<dots> = def" using insert.IH insert.prems by simp
finally show ?case .
-qed(auto intro: dtail_commute comp_fun_commute.ffold_empty)
+qed
lemma dtail_notelem_eq_def:
assumes "e \<notin> darcs t"
shows "dtail t def e = def e"
proof -
obtain r xs where xs_def[simp]: "t = Node r xs" using dtree.exhaust by auto
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>wf_darcs (Node r xs)
then b else dtail x def)"
have 0: "\<forall>(t, e1)\<in>fset xs. e \<notin> darcs t" using assms by auto
have "dtail (Node r xs) def e = ffold ?f def xs e" using assms by auto
then show ?thesis using dtail_ffold_notelem_eq_def 0 by fastforce
qed
lemma dhead_commute_aux:
fixes r xs e def
defines "f \<equiv> (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs)
then b else if e=e2 then root x else dhead x def e)"
shows "(f y \<circ> f x) z = (f x \<circ> f y) z"
proof -
obtain x1 x2 where x_def: "x = (x1,x2)" by fastforce
obtain y1 y2 where y_def: "y = (y1,y2)" by fastforce
show ?thesis
proof(cases "(x1,x2) \<in> fset xs \<and> (y1,y2) \<in> fset xs")
case 0: True
then show ?thesis
proof(cases "e \<in> darcs x1 \<and> e \<in> darcs y1")
case True
then have 1: "(x1,x2) = (y1,y2) \<or> \<not>wf_darcs (Node r xs)"
using 0 disjoint_darcs_if_wf_aux2 by fast
then show ?thesis using assms x_def y_def by (smt (z3) case_prod_conv comp_apply)
next
case False
then show ?thesis
proof(cases "x2=e")
case True
then show ?thesis using assms x_def y_def disjoint_darcs_if_wf by force
next
case False
then show ?thesis using assms x_def y_def disjoint_darcs_if_wf by fastforce
qed
qed
next
case False
then show ?thesis using assms by (simp add: x_def y_def)
qed
qed
lemma dhead_commute:
"comp_fun_commute (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs)
then b else if e=e2 then root x else dhead x def e)"
using dhead_commute_aux[of xs] by unfold_locales blast
lemma dhead_ffold_f_alt:
assumes "P = (\<lambda>xs. wf_darcs (Node r xs))" and "Q = (\<lambda>(x,e2) _. e \<in> (darcs x \<union> {e2}))"
and "R = (\<lambda>(x,e2) _. if e=e2 then root x else dhead x def e)"
shows "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs) then b
else if e=e2 then root x else dhead x def e)
= (\<lambda>a b. if a \<notin> fset xs \<or> \<not> Q a b \<or> \<not> P xs then b else R a b)"
using assms by fast
lemma dhead_ffold_f_alt_commute:
assumes "P = (\<lambda>xs. wf_darcs (Node r xs))" and "Q = (\<lambda>(x,e2) _. e \<in> (darcs x \<union> {e2}))"
and "R = (\<lambda>(x,e2) _. if e=e2 then root x else dhead x def e)"
shows "comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not> Q a b \<or> \<not> P xs then b else R a b)"
using dhead_commute[of xs e r def] dhead_ffold_f_alt[OF assms] by simp
lemma dhead_ffold_supset:
assumes "xs |\<subseteq>| ys" and "wf_darcs (Node r ys)"
shows "ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset ys \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r ys) then b
else if e=e2 then root x else dhead x def e) (def e) xs
= ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs) then b
else if e=e2 then root x else dhead x def e) (def e) xs"
(is "ffold ?f _ _ = ffold ?g _ _")
proof -
let ?P = "\<lambda>xs. wf_darcs (Node r xs)"
let ?Q = "\<lambda>(x,e2) _. e \<in> (darcs x \<union> {e2})"
let ?R = "\<lambda>(x,e2) _. if e=e2 then root x else dhead x def e"
have 0: "\<And>xs. comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not> ?Q a b \<or> \<not> ?P xs then b else ?R a b)"
using dhead_ffold_f_alt_commute by fast
have "ffold (\<lambda>a b. if a \<notin> fset ys \<or> \<not> ?Q a b \<or> \<not> ?P ys then b else ?R a b) (def e) xs
= ffold (\<lambda>a b. if a \<notin> fset xs \<or> \<not> ?Q a b \<or> \<not> ?P xs then b else ?R a b) (def e) xs"
using ffold_commute_supset[OF assms(1), of ?P ?Q ?R, OF assms(2) wf_darcs_sub 0] by simp
moreover have "?f = (\<lambda>a b. if a \<notin> fset ys \<or> \<not> ?Q a b \<or> \<not> ?P ys then b else ?R a b)" by fast
moreover have "?g = (\<lambda>a b. if a \<notin> fset xs \<or> \<not> ?Q a b \<or> \<not> ?P xs then b else ?R a b)" by fast
ultimately show ?thesis by argo
qed
lemma dhead_in_child_eq_child_ffold:
assumes "(t,e1) \<in> fset xs" and "e \<in> darcs t" and "wf_darcs (Node r xs)"
shows "ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs)
then b else if e=e2 then root x else dhead x def e) (def e) xs
= dhead t def e"
using assms proof(induction xs)
case (insert x' xs)
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset (finsert x' xs) \<or> e \<notin> (darcs x \<union> {e2})
\<or> \<not>wf_darcs (Node r (finsert x' xs))
then b else if e=e2 then root x else dhead x def e)"
let ?f' = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs) then b
else if e=e2 then root x else dhead x def e)"
obtain x e3 where x_def: "x' = (x,e3)" by fastforce
show ?case
proof(cases "x=t")
case True
have "ffold ?f (def e) (finsert x' xs) = (?f x' (ffold ?f (def e) xs))"
using comp_fun_commute.ffold_finsert[of ?f x' xs "def e"] dhead_commute insert.hyps by fast
also have "\<dots> = (?f (x,e3) (ffold ?f (def e) xs))" using x_def by blast
also have "\<dots> = dhead x def e"
using x_def insert.prems(2,3) True disjoint_darcs_if_wf by fastforce
finally show ?thesis using True by blast
next
case False
then have 0: "(t,e1) \<in> fset xs" using insert.prems(1) x_def by simp
have 1: "wf_darcs (Node r xs)" using wf_darcs_sub[OF fsubset_finsertI insert.prems(3)] .
have 2: "xs |\<subseteq>| (finsert x' xs)" by blast
have 3: "e3 \<noteq> e" "e \<notin> darcs x"
using insert.prems(1-3) disjoint_darcs_if_wf x_def False by fastforce+
have "ffold ?f (def e) (finsert x' xs) = (?f x' (ffold ?f (def e) xs))"
using comp_fun_commute.ffold_finsert[of ?f x' xs "def e"] dhead_commute insert.hyps by fast
also have "\<dots> = (?f (x,e3) (ffold ?f (def e) xs))" using x_def by blast
also have "\<dots> = (ffold ?f (def e) xs)" using 3 by simp
also have "\<dots> = (ffold ?f' (def e) xs)"
using dhead_ffold_supset[of xs "finsert x' xs"] insert.prems(3) 2 by simp
also have "\<dots> = dhead t def e" using insert.IH 0 1 insert.prems(2) by fast
finally show ?thesis .
qed
qed(simp)
lemma dhead_in_child_eq_child:
assumes "(t,e1) \<in> fset xs" and "e \<in> darcs t" and "wf_darcs (Node r xs)"
shows "dhead (Node r xs) def e = dhead t def e"
using assms dhead_in_child_eq_child_ffold[of t] by simp
lemma dhead_ffold_notelem_eq_def:
assumes "\<forall>(t,e1) \<in> fset xs. e \<notin> darcs t \<and> e \<noteq> e1"
shows "ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset ys \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r ys) then b
else if e=e2 then root x else dhead x def e) (def e) xs = def e"
using assms proof(induction xs)
+ case empty
+ show ?case
+ apply (rule comp_fun_commute.ffold_empty)
+ using dhead_commute by force
+next
case (insert x' xs)
obtain x e3 where x_def: "x' = (x,e3)" by fastforce
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset ys \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r ys)
then b else if e=e2 then root x else dhead x def e)"
have "ffold ?f (def e) (finsert x' xs) = ?f x' (ffold ?f (def e) xs)"
using comp_fun_commute.ffold_finsert[of ?f x' xs] dhead_commute insert.hyps by fast
also have "\<dots> = (ffold ?f (def e) xs)" using insert.prems by auto
also have "\<dots> = def e" using insert.IH insert.prems by simp
finally show ?case .
-qed(auto intro: dtail_commute comp_fun_commute.ffold_empty)
+qed
lemma dhead_notelem_eq_def:
assumes "e \<notin> darcs t"
shows "dhead t def e = def e"
proof -
obtain r xs where xs_def[simp]: "t = Node r xs" using dtree.exhaust by auto
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs)
then b else if e=e2 then root x else dhead x def e)"
have 0: "\<forall>(t, e1)\<in>fset xs. e \<notin> darcs t \<and> e1\<noteq>e" using assms by auto
have "dhead (Node r xs) def e = ffold ?f (def e) xs" by simp
then show ?thesis using dhead_ffold_notelem_eq_def 0 by fastforce
qed
lemma dhead_in_set_eq_root_ffold:
assumes "(t,e) \<in> fset xs" and "wf_darcs (Node r xs)"
shows "ffold (\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs)
then b else if e=e2 then root x else dhead x def e) (def e) xs
= root t" (is "ffold ?f' _ _ = _")
using assms proof(induction xs)
case (insert x' xs)
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset (finsert x' xs) \<or> e \<notin> (darcs x \<union> {e2})
\<or> \<not>wf_darcs (Node r (finsert x' xs))
then b else if e=e2 then root x else dhead x def e)"
let ?f' = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>wf_darcs (Node r xs) then b
else if e=e2 then root x else dhead x def e)"
obtain x e3 where x_def: "x' = (x,e3)" by fastforce
show ?case
proof(cases "e3=e")
case True
then have "x=t" using insert.prems(1,2) x_def disjoint_darcs_if_wf by fastforce
have "ffold ?f (def e) (finsert x' xs) = (?f x' (ffold ?f (def e) xs))"
using comp_fun_commute.ffold_finsert[of ?f x' xs "def e"] dhead_commute insert.hyps by fast
also have "\<dots> = (?f (x,e3) (ffold ?f (def e) xs))" using x_def by blast
also have "\<dots> = root x" using x_def insert.prems(1,2) True by simp
finally show ?thesis using True \<open>x=t\<close> by blast
next
case False
then have 0: "(t,e) \<in> fset xs" using insert.prems(1) x_def by simp
have 1: "wf_darcs (Node r xs)" using wf_darcs_sub[OF fsubset_finsertI insert.prems(2)] .
have 2: "xs |\<subseteq>| (finsert x' xs)" by blast
have 3: "e3 \<noteq> e" using insert.prems(2) False by simp
have 4: "e \<notin> (darcs x \<union> {e3})"
using insert.prems(1-2) False x_def disjoint_darcs_if_wf by fastforce
have "ffold ?f (def e) (finsert x' xs) = (?f x' (ffold ?f (def e) xs))"
using comp_fun_commute.ffold_finsert[of ?f x' xs "def e"] dhead_commute insert.hyps by fast
also have "\<dots> = (?f (x,e3) (ffold ?f (def e) xs))" using x_def by blast
also have "\<dots> = (ffold ?f (def e) xs)" using 4 by auto
also have "\<dots> = (ffold ?f' (def e) xs)"
using dhead_ffold_supset[of xs "finsert x' xs"] insert.prems(2) 2 by simp
also have "\<dots> = root t" using insert.IH 0 1 insert.prems(2) by blast
finally show ?thesis .
qed
qed(simp)
lemma dhead_in_set_eq_root:
"\<lbrakk>(t,e) \<in> fset xs; wf_darcs (Node r xs)\<rbrakk> \<Longrightarrow> dhead (Node r xs) def e = root t"
using dhead_in_set_eq_root_ffold[of t] by simp
lemma self_subtree: "is_subtree t t"
using is_subtree.elims(3) by blast
lemma subtree_trans: "is_subtree x y \<Longrightarrow> is_subtree y z \<Longrightarrow> is_subtree x z"
by (induction z) fastforce+
lemma subtree_trans': "transp is_subtree"
using subtree_trans transpI by auto
lemma subtree_if_child: "x \<in> fst ` fset xs \<Longrightarrow> is_subtree x (Node r xs)"
using is_subtree.elims(3) by force
lemma subtree_if_suc: "t1 \<in> fst ` fset (sucs t2) \<Longrightarrow> is_subtree t1 t2"
using subtree_if_child[of t1 "sucs t2" "root t2"] by simp
lemma child_sub_if_strict_subtree:
"\<lbrakk>strict_subtree t1 (Node r xs)\<rbrakk> \<Longrightarrow> \<exists>t3 \<in> fst ` fset xs. is_subtree t1 t3"
unfolding strict_subtree_def by force
lemma suc_sub_if_strict_subtree:
"strict_subtree t1 t2 \<Longrightarrow> \<exists>t3 \<in> fst ` fset (sucs t2). is_subtree t1 t3"
using child_sub_if_strict_subtree[of t1 "root t2"] by simp
lemma subtree_size_decr: "\<lbrakk>is_subtree t1 t2; t1 \<noteq> t2\<rbrakk> \<Longrightarrow> size t1 < size t2"
using dtree_size_decr_aux by(induction t2) fastforce
lemma subtree_size_decr': "strict_subtree t1 t2 \<Longrightarrow> size t1 < size t2"
unfolding strict_subtree_def using dtree_size_decr_aux by(induction t2) fastforce
lemma subtree_size_le: "is_subtree t1 t2 \<Longrightarrow> size t1 \<le> size t2"
using subtree_size_decr by fastforce
lemma subtree_antisym: "\<lbrakk>is_subtree t1 t2; is_subtree t2 t1\<rbrakk> \<Longrightarrow> t1 = t2"
using subtree_size_le subtree_size_decr by fastforce
lemma subtree_antisym': "antisymp is_subtree"
using antisympI subtree_antisym by blast
corollary subtree_eq_if_trans_eq1: "\<lbrakk>is_subtree t1 t2; is_subtree t2 t3; t1 = t3\<rbrakk> \<Longrightarrow> t1 = t2"
using subtree_antisym by blast
corollary subtree_eq_if_trans_eq2: "\<lbrakk>is_subtree t1 t2; is_subtree t2 t3; t1 = t3\<rbrakk> \<Longrightarrow> t2 = t3"
using subtree_antisym by blast
lemma subtree_partial_ord: "class.order is_subtree strict_subtree"
by standard (auto simp: self_subtree subtree_antisym strict_subtree_def intro: subtree_trans)
lemma finite_subtrees: "finite {x. is_subtree x t}"
by (induction t) auto
lemma subtrees_insert_union:
"{x. is_subtree x (Node r xs)} = insert (Node r xs) (\<Union>t1 \<in> fst ` fset xs. {x. is_subtree x t1})"
by fastforce
lemma subtrees_insert_union_suc:
"{x. is_subtree x t} = insert t (\<Union>t1 \<in> fst ` fset (sucs t). {x. is_subtree x t1})"
using subtrees_insert_union[of "root t" "sucs t"] by simp
lemma darcs_subtree_subset: "is_subtree x y \<Longrightarrow> darcs x \<subseteq> darcs y"
by(induction y) force
lemma dverts_subtree_subset: "is_subtree x y \<Longrightarrow> dverts x \<subseteq> dverts y"
by(induction y) force
lemma single_subtree_root_dverts:
"is_subtree (Node v2 {|(t2, e2)|}) t1 \<Longrightarrow> v2 \<in> dverts t1"
by (fastforce dest: dverts_subtree_subset)
lemma single_subtree_child_root_dverts:
"is_subtree (Node v2 {|(t2, e2)|}) t1 \<Longrightarrow> root t2 \<in> dverts t1"
by (fastforce simp: dtree.set_sel(1) dest: dverts_subtree_subset)
lemma subtree_root_if_dverts: "x \<in> dverts t \<Longrightarrow> \<exists>xs. is_subtree (Node x xs) t"
by(induction t) fastforce
lemma subtree_child_if_strict_subtree:
"strict_subtree t1 t2 \<Longrightarrow> \<exists>r xs. is_subtree (Node r xs) t2 \<and> t1 \<in> fst ` fset xs"
proof(induction t2)
case (Node r xs)
then obtain t e where t_def: "(t,e) \<in> fset xs" "is_subtree t1 t"
unfolding strict_subtree_def by auto
show ?case
proof(cases "t1 = t")
case True
then show ?thesis using t_def by force
next
case False
then show ?thesis using Node.IH[OF t_def(1)] t_def unfolding strict_subtree_def by auto
qed
qed
lemma subtree_child_if_dvert_notroot:
assumes "v \<noteq> r" and "v \<in> dverts (Node r xs)"
shows "\<exists>r' ys zs. is_subtree (Node r' ys) (Node r xs) \<and> Node v zs \<in> fst ` fset ys"
proof -
obtain zs where sub: "is_subtree (Node v zs) (Node r xs)"
using assms(2) subtree_root_if_dverts by fast
then show ?thesis using subtree_child_if_strict_subtree strict_subtree_def assms(1) by fast
qed
lemma subtree_child_if_dvert_notelem:
"\<lbrakk>v \<noteq> root t; v \<in> dverts t\<rbrakk> \<Longrightarrow> \<exists>r' ys zs. is_subtree (Node r' ys) t \<and> Node v zs \<in> fst ` fset ys"
using subtree_child_if_dvert_notroot[of v "root t" "sucs t"] by simp
lemma strict_subtree_subset:
assumes "strict_subtree t (Node r xs)" and "xs |\<subseteq>| ys"
shows "strict_subtree t (Node r ys)"
proof -
obtain t1 e1 where t1_def: "(t1,e1) \<in> fset xs" "is_subtree t t1"
using assms(1) unfolding strict_subtree_def by auto
have "size t < size (Node r xs)" using subtree_size_decr'[OF assms(1)] by blast
then have "size t < size (Node r ys)" using size_le_if_child_subset[OF assms(2)] by simp
- moreover have "is_subtree t (Node r ys)" using assms(2) t1_def notin_fset[of "(t1,e1)"] by auto
+ moreover have "is_subtree t (Node r ys)" using assms(2) t1_def by auto
ultimately show ?thesis unfolding strict_subtree_def by blast
qed
lemma strict_subtree_singleton:
"\<lbrakk>strict_subtree t (Node r {|x|}); x |\<in>| xs\<rbrakk>
\<Longrightarrow> strict_subtree t (Node r xs)"
using strict_subtree_subset by fast
subsubsection "Finite Directed Trees to Dtree"
context finite_directed_tree
begin
lemma child_subtree:
assumes "e \<in> out_arcs T r"
shows "{x. (head T e) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x} \<subseteq> {x. r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x}"
proof -
have "r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> (head T e)" using assms in_arcs_imp_in_arcs_ends by auto
then show ?thesis by (metis Collect_mono reachable_trans)
qed
lemma child_strict_subtree:
assumes "e \<in> out_arcs T r"
shows "{x. (head T e) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x} \<subset> {x. r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x}"
proof -
have "r \<rightarrow>\<^bsub>T\<^esub> (head T e)" using assms in_arcs_imp_in_arcs_ends by auto
then have "\<not> ((head T e) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> r)" using reachable1_not_reverse by blast
then show ?thesis using child_subtree assms by auto
qed
lemma child_card_decr:
assumes "e \<in> out_arcs T r"
shows "Finite_Set.card {x. (head T e) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x} < Finite_Set.card {x. r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x}"
using assms child_strict_subtree by (meson psubset_card_mono reachable_verts_finite)
function to_dtree_aux :: "'a \<Rightarrow> ('a,'b) dtree" where
"to_dtree_aux r = Node r (Abs_fset {(x,e).
(if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)})"
by auto
termination
by(relation "measure (\<lambda>r. Finite_Set.card {x. r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x})") (auto simp: child_card_decr)
definition to_dtree :: "('a,'b) dtree" where
"to_dtree = to_dtree_aux root"
abbreviation from_dtree :: "('a,'b) dtree \<Rightarrow> ('a,'b) pre_digraph" where
"from_dtree t \<equiv> Dtree.from_dtree (tail T) (head T) t"
lemma to_dtree_root_eq_root[simp]: "Dtree.root to_dtree = root"
unfolding to_dtree_def by simp
lemma verts_fset_id: "fset (Abs_fset (verts T)) = verts T"
by (simp add: Abs_fset_inverse)
lemma arcs_fset_id: "fset (Abs_fset (arcs T)) = arcs T"
by (simp add: Abs_fset_inverse)
lemma dtree_leaf_child_empty:
"leaf r \<Longrightarrow> {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)} = {}"
unfolding leaf_def by simp
lemma dtree_leaf_no_children: "leaf r \<Longrightarrow> to_dtree_aux r = Node r {||}"
using dtree_leaf_child_empty by (simp add: bot_fset.abs_eq)
lemma dtree_children_alt:
"{(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}
= {(x,e). e \<in> out_arcs T r \<and> x = to_dtree_aux (head T e)}"
by metis
lemma dtree_children_img_alt:
"(\<lambda>e. (to_dtree_aux (head T e),e)) ` (out_arcs T r)
= {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}"
using dtree_children_alt by blast
lemma dtree_children_fin:
"finite {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}"
using finite_imageI[of "out_arcs T r" "(\<lambda>e. (to_dtree_aux (head T e),e))"]
dtree_children_img_alt finite_out_arcs by fastforce
lemma dtree_children_fset_id:
assumes "to_dtree_aux r = Node r xs"
shows "fset xs = {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}"
proof -
let ?xs = "{(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}"
have "finite ?xs" using dtree_children_fin by simp
then have "fset (Abs_fset ?xs) = ?xs" using Abs_fset_inverse by blast
then show ?thesis using assms Abs_fset_inverse by simp
qed
lemma to_dtree_aux_empty_if_notT:
assumes "r \<notin> verts T"
shows "to_dtree_aux r = Node r {||}"
proof(rule ccontr)
assume asm: "to_dtree_aux r \<noteq> Node r {||}"
then obtain xs where xs_def: "Node r xs = to_dtree_aux r" by simp
then have "xs \<noteq> {||}" using asm by simp
- then obtain x e where x_def: "(x,e) \<in> fset xs" using notin_fset by fast
+ then obtain x e where x_def: "(x,e) \<in> fset xs" by fast
then have "e \<in> out_arcs T r" using xs_def dtree_children_fset_id[of r] by (auto split: if_splits)
then show False using assms by auto
qed
lemma to_dtree_aux_root: "Dtree.root (to_dtree_aux r) = r"
by simp
lemma out_arc_if_child:
assumes "x \<in> (fst ` {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)})"
shows "\<exists>e. e \<in> out_arcs T r \<and> x = to_dtree_aux (head T e)"
proof -
let ?xs = "{(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}"
have "\<exists>y. y \<in> ?xs \<and> fst y = x" using assms by blast
then show ?thesis by (smt (verit, best) case_prodE fst_conv mem_Collect_eq)
qed
lemma dominated_if_child_aux:
assumes "x \<in> (fst ` {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)})"
shows "r \<rightarrow>\<^bsub>T\<^esub> (Dtree.root x)"
proof -
obtain e where "e \<in> out_arcs T r \<and> x = to_dtree_aux (head T e)"
using assms out_arc_if_child by blast
then show ?thesis using in_arcs_imp_in_arcs_ends by force
qed
lemma dominated_if_child:
"\<lbrakk>to_dtree_aux r = Node r xs; x \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> r \<rightarrow>\<^bsub>T\<^esub> (Dtree.root x)"
using dominated_if_child_aux dtree_children_fset_id by simp
lemma image_add_snd_snd_id: "snd ` ((\<lambda>e. (to_dtree_aux (head T e),e)) ` x) = x"
by (intro equalityI subsetI) (force simp: image_iff)+
lemma to_dtree_aux_child_in_verts:
assumes "Node r' xs = to_dtree_aux r" and "x \<in> fst ` fset xs"
shows "Dtree.root x \<in> verts T"
proof -
have "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root x" using assms dominated_if_child by auto
then show ?thesis using adj_in_verts(2) by auto
qed
lemma to_dtree_aux_parent_in_verts:
assumes "Node r' xs = to_dtree_aux r" and "x \<in> fst ` fset xs"
shows "r \<in> verts T"
proof -
have "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root x" using assms dominated_if_child by auto
then show ?thesis using adj_in_verts(2) by auto
qed
lemma dtree_out_arcs:
"snd ` {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)} = out_arcs T r"
using dtree_children_img_alt by (metis image_add_snd_snd_id)
lemma dtree_out_arcs_eq_snd:
assumes "to_dtree_aux r = Node r xs"
shows "(snd ` (fset xs)) = out_arcs T r"
using assms dtree_out_arcs dtree_children_fset_id by blast
lemma dtree_aux_fst_head_snd_aux:
assumes "x \<in> {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}"
shows "Dtree.root (fst x) = (head T (snd x))"
using assms by (metis (mono_tags, lifting) Collect_case_prodD to_dtree_aux_root)
lemma dtree_aux_fst_head_snd:
assumes "to_dtree_aux r = Node r xs" and "x \<in> fset xs"
shows "Dtree.root (fst x) = (head T (snd x))"
using assms dtree_children_fset_id dtree_aux_fst_head_snd_aux by simp
lemma child_if_dominated_aux:
assumes "r \<rightarrow>\<^bsub>T\<^esub> x"
shows "\<exists>y \<in> (fst ` {(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}).
Dtree.root y = x"
proof -
let ?xs = "{(x,e). (if e \<in> out_arcs T r then x = to_dtree_aux (head T e) else False)}"
obtain e where e_def: "e \<in> out_arcs T r \<and> head T e = x" using assms by auto
then have "e \<in> snd ` ?xs" using dtree_out_arcs by auto
then obtain y where y_def: "y \<in> ?xs \<and> snd y = e" by blast
then have "Dtree.root (fst y) = head T e" using dtree_aux_fst_head_snd_aux by blast
then show ?thesis using e_def y_def by blast
qed
lemma child_if_dominated:
assumes "to_dtree_aux r = Node r xs" and "r \<rightarrow>\<^bsub>T\<^esub> x"
shows "\<exists>y \<in> (fst ` (fset xs)). Dtree.root y = x"
using assms child_if_dominated_aux dtree_children_fset_id by presburger
lemma to_dtree_aux_reach_in_dverts: "\<lbrakk>t = to_dtree_aux r; r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x\<rbrakk> \<Longrightarrow> x \<in> dverts t"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r = r'" by simp
then show ?case
proof(cases "r=x")
case True
then show ?thesis using \<open>r = r'\<close> by simp
next
case False
then have "r \<rightarrow>\<^sup>+\<^bsub>T\<^esub> x" using "1.prems"(2) by blast
then have "\<exists>r'. r \<rightarrow>\<^bsub>T\<^esub> r' \<and> r' \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x"
by (metis False converse_reachable_cases reachable1_reachable)
then obtain x' where x'_def: "r \<rightarrow>\<^bsub>T\<^esub> x' \<and> x' \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x" by blast
then obtain y where y_def: "y \<in> fst ` fset xs \<and> Dtree.root y = x'"
using "1.prems"(1) child_if_dominated by fastforce
then obtain yp where yp_def: "fst yp = y \<and> yp \<in> fset xs" using y_def by blast
from y_def have "y = to_dtree_aux x'"
using "1.prems"(1) dtree_children_fset_id \<open>r=r'\<close>
by (metis (no_types, lifting) out_arc_if_child to_dtree_aux_root)
then have "x \<in> dverts y" using "1.IH" prod.exhaust_sel yp_def x'_def by metis
then show ?thesis using dtree.set_intros(2) y_def by auto
qed
qed
lemma to_dtree_aux_dverts_reachable:
"\<lbrakk>t = to_dtree_aux r; x \<in> dverts t; r \<in> verts T\<rbrakk> \<Longrightarrow> r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r = r'" by simp
then show ?case
proof(cases "r=x")
case True
then show ?thesis using "1.prems"(3) by auto
next
case False
then obtain y where y_def: "y \<in> fst ` fset xs \<and> x \<in> dverts y"
using "1.prems"(2) \<open>r = r'\<close> by fastforce
then have 0: "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root y" using "1.prems"(1) \<open>r = r'\<close> dominated_if_child by simp
then have 2: "Dtree.root y \<in> verts T" using adj_in_verts(2) by auto
obtain yp where yp_def: "fst yp = y \<and> yp \<in> fset xs" using y_def by blast
have "\<exists>yr. y = to_dtree_aux yr"
using "1.prems"(1) y_def dtree_children_fset_id
by (metis (no_types, lifting) \<open>r = r'\<close> out_arc_if_child)
then have "Dtree.root y \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x"
using "1.IH" 2 y_def yp_def surjective_pairing to_dtree_aux_root by metis
then show ?thesis using 0 adj_reachable_trans by auto
qed
qed
lemma dverts_eq_reachable: "r \<in> verts T \<Longrightarrow> dverts (to_dtree_aux r) = {x. r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x}"
using to_dtree_aux_reach_in_dverts to_dtree_aux_dverts_reachable by blast
lemma dverts_eq_reachable': "\<lbrakk>r \<in> verts T; t = to_dtree_aux r\<rbrakk> \<Longrightarrow> dverts t = {x. r \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x}"
using dverts_eq_reachable by blast
lemma dverts_eq_verts: "dverts to_dtree = verts T"
unfolding to_dtree_def using dverts_eq_reachable reachable_from_root reachable_in_verts(2)
by (metis mem_Collect_eq root_in_T subsetI subset_antisym)
lemma arc_out_arc: "e \<in> arcs T \<Longrightarrow> \<exists>v \<in> verts T. e \<in> out_arcs T v"
by simp
lemma darcs_in_out_arcs: "t = to_dtree_aux r \<Longrightarrow> e \<in> darcs t \<Longrightarrow> \<exists>v\<in>dverts t. e \<in> out_arcs T v"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then show ?case
proof(cases "e \<in> snd ` fset xs")
case True
then show ?thesis
using "1.prems"(1) dtree_out_arcs_eq_snd to_dtree_aux_root
by (metis dtree.set_intros(1) dtree.sel(1))
next
case False
then have "\<exists>y \<in> fst ` fset xs. e \<in> darcs y" using "1.prems"(2) by force
then obtain y where y_def: "y \<in> fst ` fset xs \<and> e \<in> darcs y" by blast
obtain yp where yp_def: "fst yp = y \<and> yp \<in> fset xs" using y_def by blast
have 0: "(y, snd yp) = yp" using yp_def by auto
have "\<exists>yr. y = to_dtree_aux yr"
using "1.prems"(1) y_def dtree_children_fset_id
by (metis (no_types, lifting) dtree.sel(1) out_arc_if_child to_dtree_aux_root)
then have "\<exists>v\<in>dverts y. e \<in> out_arcs T v" using "1.IH" 0 y_def yp_def by blast
then obtain v where "v \<in> dverts y \<and> e \<in> out_arcs T v" by blast
then show ?thesis using y_def by auto
qed
qed
lemma darcs_in_arcs: "e \<in> darcs to_dtree \<Longrightarrow> e \<in> arcs T"
using darcs_in_out_arcs out_arcs_in_arcs to_dtree_def by fast
lemma out_arcs_in_darcs: "t = to_dtree_aux r \<Longrightarrow> \<exists>v\<in>dverts t. e \<in> out_arcs T v \<Longrightarrow> e \<in> darcs t"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r' = r" by simp
then obtain v where v_def: "v\<in>dverts (Node r xs) \<and> e \<in> out_arcs T v" using "1.prems"(2) by blast
then show ?case
proof(cases "e \<in> snd ` fset xs")
case True
then show ?thesis by force
next
case False
then have "e \<notin> out_arcs T r" using "1.prems"(1) \<open>r' = r\<close> dtree_out_arcs_eq_snd by metis
then have "v \<noteq> r" using v_def by blast
then obtain y where y_def: "y \<in> fst ` fset xs \<and> v \<in> dverts y" using v_def by force
then obtain yp where yp_def: "fst yp = y \<and> yp \<in> fset xs" by blast
have 0: "(y, snd yp) = yp" using yp_def by auto
have "\<exists>yr. y = to_dtree_aux yr"
using "1.prems"(1) y_def dtree_children_fset_id
by (metis (no_types, lifting) dtree.sel(1) out_arc_if_child to_dtree_aux_root)
then have "e \<in> darcs y" using "1.IH" 0 v_def y_def yp_def by blast
then show ?thesis using y_def by force
qed
qed
lemma arcs_in_darcs: "e \<in> arcs T \<Longrightarrow> e \<in> darcs to_dtree"
using arc_out_arc out_arcs_in_darcs dverts_eq_verts to_dtree_def by fast
lemma darcs_eq_arcs: "darcs to_dtree = arcs T"
using arcs_in_darcs darcs_in_arcs by blast
lemma to_dtree_aux_self:
assumes "Node r xs = to_dtree_aux r" and "(y,e) \<in> fset xs"
shows "y = to_dtree_aux (Dtree.root y)"
proof -
have "\<exists>y'. y = to_dtree_aux y'"
using assms dtree_children_fset_id by (metis (mono_tags, lifting) case_prodD mem_Collect_eq)
then obtain y' where "y = to_dtree_aux y'" by blast
then show ?thesis by simp
qed
lemma to_dtree_aux_self_subtree:
"\<lbrakk>t1 = to_dtree_aux r; is_subtree t2 t1\<rbrakk> \<Longrightarrow> t2 = to_dtree_aux (Dtree.root t2)"
proof(induction t1 arbitrary: r)
case (Node r' xs)
then show ?case
proof(cases "Node r' xs = t2")
case True
then show ?thesis using Node.prems(1) by force
next
case False
then obtain t e where t_def: "(t,e) \<in> fset xs" "is_subtree t2 t" using Node.prems(2) by auto
then have "t = to_dtree_aux (Dtree.root t)" using Node.prems(1) to_dtree_aux_self by simp
then show ?thesis using Node.IH[of "(t,e)" t "Dtree.root t"] t_def by simp
qed
qed
lemma to_dtree_self_subtree: "is_subtree t to_dtree \<Longrightarrow> t = to_dtree_aux (Dtree.root t)"
unfolding to_dtree_def using to_dtree_aux_self_subtree by blast
lemma to_dtree_self_subtree': "is_subtree (Node r xs) to_dtree \<Longrightarrow> (Node r xs) = to_dtree_aux r"
using to_dtree_self_subtree[of "Node r xs"] by simp
lemma child_if_dominated_to_dtree:
"\<lbrakk>is_subtree (Node r xs) to_dtree; r \<rightarrow>\<^bsub>T\<^esub> v\<rbrakk> \<Longrightarrow> \<exists>t. t \<in> fst ` fset xs \<and> Dtree.root t = v"
using child_if_dominated[of r] to_dtree_self_subtree' by simp
lemma child_if_dominated_to_dtree':
"\<lbrakk>is_subtree (Node r xs) to_dtree; r \<rightarrow>\<^bsub>T\<^esub> v\<rbrakk> \<Longrightarrow> \<exists>ys. Node v ys \<in> fst ` fset xs"
using child_if_dominated_to_dtree dtree.exhaust dtree.sel(1) by metis
lemma child_darc_tail_parent:
assumes "Node r xs = to_dtree_aux r" and "(x,e) \<in> fset xs"
shows "tail T e = r"
proof -
have "e \<in> out_arcs T r"
using assms dtree_children_fset_id by (metis (no_types, lifting) case_prodD mem_Collect_eq)
then show ?thesis by simp
qed
lemma child_darc_head_root:
"\<lbrakk>Node r xs = to_dtree_aux r; (t,e) \<in> fset xs\<rbrakk> \<Longrightarrow> head T e = Dtree.root t"
using dtree_aux_fst_head_snd by force
lemma child_darc_in_arcs:
assumes "Node r xs = to_dtree_aux r" and "(x,e) \<in> fset xs"
shows "e \<in> arcs T"
proof -
have "e \<in> out_arcs T r"
using assms dtree_children_fset_id by (metis (no_types, lifting) case_prodD mem_Collect_eq)
then show ?thesis by simp
qed
lemma darcs_neq_if_dtrees_neq:
"\<lbrakk>Node r xs = to_dtree_aux r; (x,e1) \<in> fset xs; (y,e2) \<in> fset xs; x\<noteq>y\<rbrakk> \<Longrightarrow> e1 \<noteq> e2"
using dtree_children_fset_id by (metis (mono_tags, lifting) case_prodD mem_Collect_eq)
lemma dtrees_neq_if_darcs_neq:
"\<lbrakk>Node r xs = to_dtree_aux r; (x,e1) \<in> fset xs; (y,e2) \<in> fset xs; e1\<noteq>e2\<rbrakk> \<Longrightarrow> x \<noteq> y"
using dtree_children_fset_id case_prodD dtree_aux_fst_head_snd fst_conv
by (metis (no_types, lifting) mem_Collect_eq out_arcs_in_arcs snd_conv two_in_arcs_contr)
lemma dverts_disjoint:
assumes "Node r xs = to_dtree_aux r" and "(x,e1) \<in> fset xs" and "(y,e2) \<in> fset xs"
and "(x,e1)\<noteq>(y,e2)"
shows "dverts x \<inter> dverts y = {}"
proof (rule ccontr)
assume "dverts x \<inter> dverts y \<noteq> {}"
then obtain v where v_def: "v \<in> dverts x \<and> v \<in> dverts y" by blast
have "x \<noteq> y" using dtrees_neq_if_darcs_neq assms by blast
have 0: "x = to_dtree_aux (Dtree.root x)" using to_dtree_aux_self assms(1,2) by blast
have 1: "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root x"
using assms(1,2) dominated_if_child by (metis (no_types, opaque_lifting) fst_conv image_iff)
then have 2: "Dtree.root x \<in> verts T" using adj_in_verts(2) by simp
have 3: "y = to_dtree_aux (Dtree.root y)" using to_dtree_aux_self assms(1,3) by blast
have 4: "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root y"
using assms(1,3) dominated_if_child by (metis (no_types, opaque_lifting) fst_conv image_iff)
then have 5: "Dtree.root y \<in> verts T" using adj_in_verts(2) by simp
have "Dtree.root x \<rightarrow>\<^sup>*\<^bsub>T\<^esub> v" using 0 2 to_dtree_aux_dverts_reachable v_def by blast
moreover have "Dtree.root y \<rightarrow>\<^sup>*\<^bsub>T\<^esub> v" using 3 5 to_dtree_aux_dverts_reachable v_def by blast
moreover have "Dtree.root x \<noteq> Dtree.root y" using 0 3 assms(4) \<open>x\<noteq>y\<close> by auto
ultimately show False using 1 4 reachable_via_child_impl_same by simp
qed
lemma wf_dverts_to_dtree_aux1: "r \<notin> verts T \<Longrightarrow> wf_dverts (to_dtree_aux r)"
using to_dtree_aux_empty_if_notT unfolding wf_dverts_iff_dverts' by simp
lemma wf_dverts_to_dtree_aux2: "r \<in> verts T \<Longrightarrow> t = to_dtree_aux r \<Longrightarrow> wf_dverts t"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r = r'" by simp
have "\<forall>(x,e) \<in> fset xs. wf_dverts x \<and> r \<notin> dverts x"
proof (standard, standard, standard)
fix xp x e
assume asm: "xp \<in> fset xs" "xp = (x,e)"
then have 0: "x = to_dtree_aux (Dtree.root x)" using to_dtree_aux_self "1.prems"(2) by simp
have 2: "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root x" using asm "1.prems" \<open>r = r'\<close>
by (metis (no_types, opaque_lifting) dominated_if_child fst_conv image_iff)
then have 3: "Dtree.root x \<in> verts T" using adj_in_verts(2) by simp
then show "wf_dverts x" using "1.IH" asm 0 by blast
show "r \<notin> dverts x"
proof
assume "r \<in> dverts x"
then have "Dtree.root x \<rightarrow>\<^sup>*\<^bsub>T\<^esub> r" using 0 3 to_dtree_aux_dverts_reachable by blast
then have "r \<rightarrow>\<^sup>+\<^bsub>T\<^esub> r" using 2 by auto
then show False using reachable1_not_reverse by blast
qed
qed
then show ?case using dverts_disjoint \<open>r=r'\<close> "1.prems"(1,2) unfolding wf_dverts_iff_dverts'
by (smt (verit, del_insts) wf_dverts'.simps case_prodI2 case_prod_conv)
qed
lemma wf_dverts_to_dtree_aux: "wf_dverts (to_dtree_aux r)"
using wf_dverts_to_dtree_aux1 wf_dverts_to_dtree_aux2 by blast
lemma wf_dverts_to_dtree_aux': "t = to_dtree_aux r \<Longrightarrow> wf_dverts t"
using wf_dverts_to_dtree_aux by blast
lemma wf_dverts_to_dtree: "wf_dverts to_dtree"
using to_dtree_def wf_dverts_to_dtree_aux by simp
lemma darcs_not_in_subtree:
assumes "Node r xs = to_dtree_aux r" and "(x,e) \<in> fset xs" and "(y,e2) \<in> fset xs"
shows "e \<notin> darcs y"
proof
assume asm: "e \<in> darcs y"
have 0: "y = to_dtree_aux (Dtree.root y)" using to_dtree_aux_self assms(1,3) by blast
then obtain v where v_def: "v \<in> dverts y \<and> e \<in> out_arcs T v" using darcs_in_out_arcs asm by blast
have 1: "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root y"
using assms(1,3) by (metis (no_types, opaque_lifting) dominated_if_child fst_conv image_iff)
then have "Dtree.root y \<in> verts T" using adj_in_verts(2) by auto
then have "Dtree.root y \<rightarrow>\<^sup>*\<^bsub>T\<^esub> v" using to_dtree_aux_dverts_reachable 0 v_def by blast
then have "r \<rightarrow>\<^sup>+\<^bsub>T\<^esub> v" using 1 by auto
then have "r \<noteq> v" using reachable1_not_reverse two_in_arcs_contr by blast
moreover have "tail T e = v" using v_def by simp
moreover have "tail T e = r" using assms(1,2) child_darc_tail_parent by blast
ultimately show False by blast
qed
lemma darcs_disjoint:
assumes "Node r xs = to_dtree_aux r" and "r \<in> verts T"
and "(x,e1) \<in> fset xs" and "(y,e2) \<in> fset xs" and "(x,e1)\<noteq>(y,e2)"
shows "(darcs x \<union> {e1}) \<inter> (darcs y \<union> {e2}) = {}"
proof (rule ccontr)
assume "(darcs x \<union> {e1}) \<inter> (darcs y \<union> {e2}) \<noteq> {}"
moreover have "e1 \<notin> darcs y" using darcs_not_in_subtree assms(1-4) by blast
moreover have "e2 \<notin> darcs x" using darcs_not_in_subtree assms(1-4) by blast
moreover have "e1 \<noteq> e2" using darcs_neq_if_dtrees_neq assms by blast
ultimately have "darcs x \<inter> darcs y \<noteq> {}" by blast
then obtain e where e_def: "e \<in> darcs x \<and> e \<in> darcs y" by blast
have "x = to_dtree_aux (Dtree.root x)" using to_dtree_aux_self assms(1,3) by blast
then obtain v1 where v1_def: "v1 \<in> dverts x \<and> e \<in> out_arcs T v1"
using darcs_in_out_arcs e_def by blast
have "y = to_dtree_aux (Dtree.root y)" using to_dtree_aux_self assms(1,4) by blast
then obtain v2 where v2_def: "v2 \<in> dverts y \<and> e \<in> out_arcs T v2"
using darcs_in_out_arcs e_def by blast
then have "v2 \<noteq> v1" using v1_def v2_def dverts_disjoint assms dtrees_neq_if_darcs_neq by blast
then show False using v1_def v2_def by simp
qed
lemma wf_darcs_to_dtree_aux1: "r \<notin> verts T \<Longrightarrow> wf_darcs (to_dtree_aux r)"
using to_dtree_aux_empty_if_notT unfolding wf_darcs_def by simp
lemma wf_darcs_to_dtree_aux2: "r \<in> verts T \<Longrightarrow> t = to_dtree_aux r \<Longrightarrow> wf_darcs t"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r = r'" by simp
have "\<forall>(x,e) \<in> fset xs. wf_darcs x"
proof (standard, standard)
fix xp x e
assume asm: "xp \<in> fset xs" "xp = (x,e)"
then have 0: "x = to_dtree_aux (Dtree.root x)" using to_dtree_aux_self "1.prems"(2) by simp
have "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root x" using asm "1.prems" \<open>r = r'\<close>
by (metis (no_types, opaque_lifting) dominated_if_child fst_conv image_iff)
then have "Dtree.root x \<in> verts T" using adj_in_verts(2) by simp
then show "wf_darcs x" using "1.IH" asm 0 by blast
qed
moreover have "\<forall>(x,e1) \<in> fset xs. (\<forall>(y,e2) \<in> fset xs.
(darcs x \<union> {e1}) \<inter> (darcs y \<union> {e2}) = {} \<or> (x,e1)=(y,e2))"
using darcs_disjoint "1.prems" \<open>r = r'\<close> by blast
ultimately show ?case using darcs_not_in_subtree "1.prems" \<open>r = r'\<close>
by (smt (verit) case_prodD case_prodI2 wf_darcs_if_darcs'_aux)
qed
lemma wf_darcs_to_dtree_aux: "wf_darcs (to_dtree_aux r)"
using wf_darcs_to_dtree_aux1 wf_darcs_to_dtree_aux2 by blast
lemma wf_darcs_to_dtree_aux': "t = to_dtree_aux r \<Longrightarrow> wf_darcs t"
using wf_darcs_to_dtree_aux by blast
lemma wf_darcs_to_dtree: "wf_darcs to_dtree"
using to_dtree_def wf_darcs_to_dtree_aux by simp
lemma dtail_aux_elem_eq_tail:
"t = to_dtree_aux r \<Longrightarrow> e \<in> darcs t \<Longrightarrow> dtail t def e = tail T e"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r = r'" by simp
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> darcs x \<or> \<not>disjoint_darcs xs
then b else dtail x def)"
show ?case
proof(cases "e \<in> snd ` fset xs")
case True
then have 0: "dtail (Node r' xs) def e = r" using \<open>r=r'\<close> by simp
have "e \<in> out_arcs T r" using dtree_out_arcs_eq_snd "1.prems"(1) True by simp
then have "tail T e = r" by simp
then show ?thesis using 0 by blast
next
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using "1.prems"(2) by force
then have "x = to_dtree_aux (Dtree.root x)" using "1.prems"(1) \<open>r = r'\<close> to_dtree_aux_self by blast
then have 0: "dtail x def e = tail T e" using "1.IH" x_def by blast
have "wf_darcs (Node r xs)" using "1.prems"(1) wf_darcs_to_dtree_aux by simp
then have "dtail (Node r' xs) def e = dtail x def e"
using dtail_in_child_eq_child[of x] x_def "1.prems" by force
then show ?thesis using 0 by simp
qed
qed
lemma dtail_elem_eq_tail: "e \<in> darcs to_dtree \<Longrightarrow> dtail to_dtree def e = tail T e"
using dtail_aux_elem_eq_tail to_dtree_def by blast
lemma to_dtree_dtail_eq_tail_aux: "dtail to_dtree (tail T) e = tail T e"
using dtail_notelem_eq_def dtail_elem_eq_tail by metis
lemma to_dtree_dtail_eq_tail: "dtail to_dtree (tail T) = tail T"
using to_dtree_dtail_eq_tail_aux by blast
lemma dhead_aux_elem_eq_head:
"t = to_dtree_aux r \<Longrightarrow> e \<in> darcs t \<Longrightarrow> dhead t def e = head T e"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r = r'" by simp
let ?f = "(\<lambda>(x,e2) b. if (x,e2) \<notin> fset xs \<or> e \<notin> (darcs x \<union> {e2}) \<or> \<not>disjoint_darcs xs
then b else if e=e2 then Dtree.root x else dhead x def e)"
obtain child where "child \<in> fset xs" using "1.prems"(2) by auto
then have wf: "wf_darcs (Node r xs)" using "1.prems"(1) wf_darcs_to_dtree_aux by simp
show ?case
proof(cases "e \<in> snd ` fset xs")
case True
then obtain x where x_def: "(x,e) \<in> fset xs" by force
then have 0: "dhead (Node r' xs) def e = Dtree.root x"
using dhead_in_set_eq_root wf \<open>r=r'\<close> by fast
have "e \<in> out_arcs T r" using dtree_out_arcs_eq_snd "1.prems"(1) True by simp
then have "head T e = Dtree.root x" using x_def "1.prems"(1) dtree_aux_fst_head_snd by force
then show ?thesis using 0 by simp
next
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using "1.prems"(2) by force
then have "x = to_dtree_aux (Dtree.root x)" using "1.prems"(1) \<open>r = r'\<close> to_dtree_aux_self by blast
then have 0: "dhead x def e = head T e" using "1.IH" x_def by blast
have "dhead (Node r' xs) def e = dhead x def e"
using dhead_in_child_eq_child[of x] x_def wf \<open>r=r'\<close> by blast
then show ?thesis using 0 by simp
qed
qed
lemma dhead_elem_eq_head: "e \<in> darcs to_dtree \<Longrightarrow> dhead to_dtree def e = head T e"
using dhead_aux_elem_eq_head to_dtree_def by blast
lemma to_dtree_dhead_eq_head_aux: "dhead to_dtree (head T) e = head T e"
using dhead_notelem_eq_def dhead_elem_eq_head by metis
lemma to_dtree_dhead_eq_head: "dhead to_dtree (head T) = head T"
using to_dtree_dhead_eq_head_aux by blast
lemma from_to_dtree_eq_orig: "from_dtree (to_dtree) = T"
using to_dtree_dhead_eq_head to_dtree_dtail_eq_tail darcs_eq_arcs dverts_eq_verts by simp
lemma subtree_darc_tail_parent:
"\<lbrakk>is_subtree (Node r xs) to_dtree; (t,e) \<in> fset xs\<rbrakk> \<Longrightarrow> tail T e = r"
using child_darc_tail_parent to_dtree_self_subtree' by blast
lemma subtree_darc_head_root:
"\<lbrakk>is_subtree (Node r xs) to_dtree; (t,e) \<in> fset xs\<rbrakk> \<Longrightarrow> head T e = Dtree.root t"
using child_darc_head_root to_dtree_self_subtree' by blast
lemma subtree_darc_in_arcs:
"\<lbrakk>is_subtree (Node r xs) to_dtree; (t,e) \<in> fset xs\<rbrakk> \<Longrightarrow> e \<in> arcs T"
using to_dtree_self_subtree' child_darc_in_arcs by blast
lemma subtree_child_dom: "\<lbrakk>is_subtree (Node r xs) to_dtree; (t,e) \<in> fset xs\<rbrakk> \<Longrightarrow> r \<rightarrow>\<^bsub>T\<^esub> Dtree.root t"
using subtree_darc_tail_parent subtree_darc_head_root subtree_darc_in_arcs
in_arcs_imp_in_arcs_ends by fastforce
end
subsubsection "Well-Formed Dtrees"
locale wf_dtree =
fixes t :: "('a,'b) dtree"
assumes wf_arcs: "wf_darcs t"
and wf_verts: "wf_dverts t"
begin
lemma wf_dtree_rec: "Node r xs = t \<Longrightarrow> (x,e) \<in> fset xs \<Longrightarrow> wf_dtree x"
using wf_arcs wf_verts by (unfold_locales) auto
lemma wf_dtree_sub: "is_subtree x t \<Longrightarrow> wf_dtree x"
using wf_dtree_axioms proof(induction t rule: darcs_mset.induct)
case (1 r xs)
then interpret wf_dtree "Node r xs" by blast
show ?case
proof(cases "x = Node r xs")
case True
then show ?thesis by (simp add: wf_dtree_axioms)
next
case False
then show ?thesis using "1.IH" wf_dtree_rec "1.prems"(1) by auto
qed
qed
lemma root_not_subtree: "\<lbrakk>(Node r xs) = t; x \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> r \<notin> dverts x"
using wf_verts root_not_child_if_wf_dverts by fastforce
lemma dverts_child_subset: "\<lbrakk>(Node r xs) = t; x \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> dverts x \<subset> dverts t"
using root_not_subtree by fastforce
lemma child_arc_not_subtree: "\<lbrakk>(Node r xs) = t; (x,e1) \<in> fset xs\<rbrakk> \<Longrightarrow> e1 \<notin> darcs x"
using wf_arcs disjoint_darcs_if_wf_aux3 by fast
lemma darcs_child_subset: "\<lbrakk>(Node r xs) = t; x \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> darcs x \<subset> darcs t"
using child_arc_not_subtree by force
lemma dtail_in_dverts: "e \<in> darcs t \<Longrightarrow> dtail t def e \<in> dverts t"
using wf_arcs proof(induction t rule: darcs_mset.induct)
case (1 r xs)
show ?case
proof(cases "e \<in> snd ` fset xs")
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using "1.prems"(1) by force
then have "wf_darcs x" using "1.prems"(2) by auto
then have "dtail x def e \<in> dverts x" using "1.IH" x_def by blast
then have 0: "dtail x def e \<in> dverts (Node r xs)"
using x_def by (auto simp: dverts_child_subseteq)
have "dtail (Node r xs) def e = dtail x def e"
using dtail_in_child_eq_child[of x] x_def "1.prems"(2) by blast
then show ?thesis using 0 by argo
qed (simp)
qed
lemma dtail_in_childverts:
assumes "e \<in> darcs x" and "(x,e') \<in> fset xs" and "Node r xs = t"
shows "dtail t def e \<in> dverts x"
proof -
interpret X: wf_dtree x using assms(2,3) wf_dtree_rec by blast
have "dtail t def e = dtail x def e"
using dtail_in_child_eq_child[of x] assms wf_arcs by force
then show ?thesis using assms(1) X.dtail_in_dverts by simp
qed
lemma dhead_in_dverts: "e \<in> darcs t \<Longrightarrow> dhead t def e \<in> dverts t"
using wf_arcs proof(induction t rule: darcs_mset.induct)
case (1 r xs)
show ?case
proof(cases "e \<in> snd ` fset xs")
case True
then obtain x where x_def: "(x,e) \<in> fset xs" by force
then have "dhead (Node r xs) def e = root x"
using dhead_in_set_eq_root[of x] "1.prems"(2) by blast
then show ?thesis using dtree.set_sel(1) x_def by fastforce
next
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using "1.prems"(1) by force
then have "wf_darcs x" using "1.prems"(2) by auto
then have "dhead x def e \<in> dverts x" using "1.IH" x_def by blast
then have 0: "dhead x def e \<in> dverts (Node r xs)"
using x_def by (auto simp: dverts_child_subseteq)
have "dhead (Node r xs) def e = dhead x def e"
using dhead_in_child_eq_child[of x] x_def "1.prems"(2) by force
then show ?thesis using 0 by argo
qed
qed
lemma dhead_in_childverts:
assumes "e \<in> darcs x" and "(x,e') \<in> fset xs" and "Node r xs = t"
shows "dhead t def e \<in> dverts x"
proof -
interpret X: wf_dtree x using wf_arcs wf_verts assms(2,3) by(unfold_locales) auto
have "dhead t def e = dhead x def e"
using dhead_in_child_eq_child[of x] assms wf_arcs by auto
then show ?thesis using assms(1) X.dhead_in_dverts by simp
qed
lemma dhead_in_dverts_no_root: "e \<in> darcs t \<Longrightarrow> dhead t def e \<in> (dverts t - {root t})"
using wf_arcs wf_verts proof(induction t rule: darcs_mset.induct)
case (1 r xs)
interpret wf_dtree "Node r xs" using "1.prems"(2,3) by (unfold_locales) auto
show ?case
proof(cases "e \<in> snd ` fset xs")
case True
then obtain x where x_def: "(x,e) \<in> fset xs" by force
then have "dhead (Node r xs) def e = root x"
using dhead_in_set_eq_root[of x] "1.prems"(2) by simp
then show ?thesis using dtree.set_sel(1) x_def "1.prems"(3) wf_dverts_iff_dverts' by fastforce
next
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using "1.prems"(1) by force
then have "wf_darcs x" using "1.prems"(2) by auto
then have "dhead x def e \<in> dverts x" using "1.IH" x_def "1.prems"(3) by auto
moreover have "r \<notin> dverts x" using root_not_subtree x_def by fastforce
ultimately have 0: "dhead x def e \<in> dverts (Node r xs) - {root (Node r xs)}"
using x_def dverts_child_subseteq by fastforce
have "dhead (Node r xs) def e = dhead x def e"
using dhead_in_child_eq_child[of x] x_def "1.prems"(2) by force
then show ?thesis using 0 by argo
qed
qed
lemma dhead_in_childverts_no_root:
assumes "e \<in> darcs x" and "(x,e') \<in> fset xs" and "Node r xs = t"
shows "dhead t def e \<in> (dverts x - {root x})"
proof -
interpret X: wf_dtree x using assms(2,3) wf_dtree_rec by blast
have "dhead t def e = dhead x def e"
using dhead_in_child_eq_child[of x] assms wf_arcs by auto
then show ?thesis using assms(1) X.dhead_in_dverts_no_root by simp
qed
lemma dtree_cas_iff_subtree:
assumes "(x,e1) \<in> fset xs" and "Node r xs = t" and "set p \<subseteq> darcs x"
shows "pre_digraph.cas (from_dtree dt dh x) u p v
\<longleftrightarrow> pre_digraph.cas (from_dtree dt dh t) u p v"
(is "pre_digraph.cas ?X _ _ _ \<longleftrightarrow> pre_digraph.cas ?T _ _ _")
using assms proof(induction p arbitrary: u)
case Nil
then show ?case by(simp add: pre_digraph.cas.simps(1))
next
case (Cons p ps)
note pre_digraph.cas.simps[simp]
have "pre_digraph.cas ?T u (p # ps) v = (tail ?T p = u \<and> pre_digraph.cas ?T (head ?T p) ps v)"
by simp
also have "\<dots> = (tail ?T p = u \<and> pre_digraph.cas ?X (head ?T p) ps v)"
using Cons.IH Cons.prems by simp
also have "\<dots> = (tail ?X p = u \<and> pre_digraph.cas ?X (head ?T p) ps v)"
using dtail_in_child_eq_child[of x] Cons.prems(1-3) wf_arcs by force
also have "\<dots> = (tail ?X p = u \<and> pre_digraph.cas ?X (head ?X p) ps v)"
using dhead_in_child_eq_child[of x] Cons.prems(1-3) wf_arcs by force
finally show ?case by simp
qed
lemma dtree_cas_exists:
"v \<in> dverts t \<Longrightarrow> \<exists>p. set p \<subseteq> darcs t \<and> pre_digraph.cas (from_dtree dt dh t) (root t) p v"
using wf_dtree_axioms proof(induction t)
case (Node r xs)
then show ?case
proof(cases "r=v")
case True
then have "pre_digraph.cas (from_dtree dt dh (Node r xs)) (root (Node r xs)) [] v"
by (simp add: pre_digraph.cas.simps(1))
then show ?thesis by force
next
case False
then obtain x e where x_def: "(x,e) \<in> fset xs \<and> v \<in> dverts x" using Node.prems by auto
let ?T = "from_dtree dt dh (Node r xs)"
let ?X = "from_dtree dt dh x"
interpret wf_dtree "Node r xs" by (rule Node.prems(2))
have "wf_dtree x" using x_def wf_dtree_rec by blast
then obtain p where p_def: "set p \<subseteq> darcs x \<and> pre_digraph.cas ?X (root x) p v"
using Node.IH x_def by fastforce
then have "pre_digraph.cas ?T (root x) p v"
using dtree_cas_iff_subtree x_def Node.prems(2) by blast
moreover have "head ?T e = root x"
using x_def dhead_in_set_eq_root[of x] wf_arcs by simp
moreover have "tail ?T e = r" using x_def by force
ultimately have "pre_digraph.cas ?T (root (Node r xs)) (e#p) v"
by (simp add: pre_digraph.cas.simps(2))
moreover have "set (e#p) \<subseteq> darcs (Node r xs)" using p_def x_def by force
ultimately show ?thesis by blast
qed
qed
lemma dtree_awalk_exists:
assumes "v \<in> dverts t"
shows "\<exists>p. pre_digraph.awalk (from_dtree dt dh t) (root t) p v"
unfolding pre_digraph.awalk_def using dtree_cas_exists assms dtree.set_sel(1) by fastforce
lemma subtree_root_not_root: "t = Node r xs \<Longrightarrow> (x,e) \<in> fset xs \<Longrightarrow> root x \<noteq> r"
using dtree.set_sel(1) root_not_subtree by fastforce
lemma dhead_not_root:
assumes "e \<in> darcs t"
shows "dhead t def e \<noteq> root t"
proof -
obtain r xs where xs_def[simp]: "t = Node r xs" using dtree.exhaust by auto
show ?thesis
proof(cases "e \<in> snd ` fset xs")
case True
then obtain x where x_def: "(x,e) \<in> fset xs" by force
then have "dhead (Node r xs) def e = root x"
using dhead_in_set_eq_root[of x] wf_arcs by simp
then show ?thesis using x_def subtree_root_not_root by simp
next
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using assms by force
then interpret X: wf_dtree x using wf_dtree_rec by auto
have "dhead x def e \<in> dverts x" using x_def X.dhead_in_dverts by blast
moreover have "dhead (Node r xs) def e = dhead x def e"
using x_def dhead_in_child_eq_child[of x] wf_arcs by force
ultimately show ?thesis using x_def root_not_subtree by fastforce
qed
qed
lemma nohead_cas_no_arc_in_subset:
"\<lbrakk>\<forall>e\<in>darcs t. dhead t dh e \<noteq> v; p\<noteq>[]; pre_digraph.cas (from_dtree dt dh t) u p v\<rbrakk>
\<Longrightarrow> \<not>set p \<subseteq> darcs t"
by(induction p arbitrary: u) (fastforce simp: pre_digraph.cas.simps)+
lemma dtail_root_in_set:
assumes "e \<in> darcs t" and "t = Node r xs" and "dtail t dt e = r"
shows "e \<in> snd ` fset xs"
proof (rule ccontr)
assume "e \<notin> snd ` fset xs"
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using assms(1,2) by force
interpret X: wf_dtree x using assms(2) x_def wf_dtree_rec by blast
have "dtail t dt e = dtail x dt e"
using dtail_in_child_eq_child[of x] wf_arcs assms(2) x_def by force
then have "dtail t dt e \<in> dverts x" using X.dtail_in_dverts x_def by simp
then show False using assms(2,3) wf_verts x_def unfolding wf_dverts_iff_dverts' by auto
qed
lemma dhead_notin_subtree_wo_root:
assumes "(x,e) \<in> fset xs" and "p \<notin> darcs x" and "p \<in> darcs t" and "t = Node r xs"
shows "dhead t dh p \<notin> (dverts x - {root x})"
proof(cases "p \<in> snd ` fset xs")
case True
then obtain x' where x'_def: "(x',p) \<in> fset xs" by auto
then have 0: "dhead t dh p = root x'"
using dhead_in_set_eq_root[of x'] wf_arcs assms(4) by auto
have "root x' \<notin> (dverts x - {root x})"
proof(cases "x'=x")
case True
then show ?thesis by blast
next
case False
have "root x' \<in> dverts x'" by (simp add: dtree.set_sel(1))
then show ?thesis using wf_verts x'_def assms(1,4) unfolding wf_dverts_iff_dverts' by fastforce
qed
then show ?thesis using 0 by simp
next
case False
then obtain x' e1 where x'_def: "(x',e1) \<in> fset xs \<and> p \<in> darcs x'" using assms(3,4) by force
then have 0: "dhead t dh p = dhead x' dh p"
using dhead_in_child_eq_child[of x'] wf_arcs assms(4) by auto
interpret X: wf_dtree x' using assms(4) x'_def wf_dtree_rec by blast
have 1: "dhead x' dh p \<in> dverts x'" using X.dhead_in_dverts x'_def by blast
moreover have "dverts x' \<inter> dverts x = {}"
using wf_verts x'_def assms(1,2,4) unfolding wf_dverts_iff_dverts' by fastforce
ultimately show ?thesis using 0 by auto
qed
lemma subtree_uneq_if_arc_uneq:
"\<lbrakk>(x1,e1) \<in> fset xs; (x2,e2) \<in> fset xs; e1\<noteq>e2; Node r xs = t\<rbrakk> \<Longrightarrow> x1 \<noteq> x2"
using dtree.set_sel(1) wf_verts disjoint_dverts_if_wf_aux by fast
lemma arc_uneq_if_subtree_uneq:
"\<lbrakk>(x1,e1) \<in> fset xs; (x2,e2) \<in> fset xs; x1\<noteq>x2; Node r xs = t\<rbrakk> \<Longrightarrow> e1 \<noteq> e2"
using disjoint_darcs_if_wf[OF wf_arcs] by fastforce
lemma dhead_unique: "e \<in> darcs t \<Longrightarrow> p \<in> darcs t \<Longrightarrow> e \<noteq> p \<Longrightarrow> dhead t dh e \<noteq> dhead t dh p"
using wf_dtree_axioms proof(induction t rule: darcs_mset.induct)
case ind: (1 r xs)
then interpret wf_dtree "Node r xs" by blast
show ?case
proof(cases "\<exists>x \<in> fst ` fset xs. e \<in> darcs x \<and> p \<in> darcs x")
case True
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x \<and> p \<in> darcs x" by force
then have "wf_dtree x" using ind.prems(4) wf_dtree_rec by blast
then have "dhead x dh e \<noteq> dhead x dh p" using ind x_def by blast
then show ?thesis using True dhead_in_child_eq_child[of x] wf_arcs x_def by force
next
case False
then consider "\<exists>x \<in> fst ` fset xs. e \<in> darcs x" | "\<exists>x \<in> fst ` fset xs. p \<in> darcs x"
| "e \<in> snd ` fset xs \<and> p \<in> snd ` fset xs"
using ind.prems(1,2) by force
then show ?thesis
proof(cases)
case 1
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x \<and> p \<notin> darcs x"
using False by force
then interpret X: wf_dtree x using wf_dtree_rec by blast
have "dhead x dh e \<in> (dverts x - {root x})" using X.dhead_in_dverts_no_root x_def by blast
then have "dhead (Node r xs) dh e \<in> (dverts x - {root x})"
using dhead_in_child_eq_child[of x] wf_arcs x_def by force
moreover have "dhead (Node r xs) dh p \<notin> (dverts x - {root x})"
using x_def dhead_notin_subtree_wo_root ind.prems(2) by blast
ultimately show ?thesis by auto
next
case 2
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> p \<in> darcs x \<and> e \<notin> darcs x"
using False by force
then interpret X: wf_dtree x using wf_dtree_rec by blast
have "dhead x dh p \<in> (dverts x - {root x})" using X.dhead_in_dverts_no_root x_def by blast
then have "dhead (Node r xs) dh p \<in> (dverts x - {root x})"
using dhead_in_child_eq_child[of x] wf_arcs x_def by force
moreover have "dhead (Node r xs) dh e \<notin> (dverts x - {root x})"
using x_def dhead_notin_subtree_wo_root ind.prems(1) by blast
ultimately show ?thesis by auto
next
case 3
then obtain x1 x2 where x_def: "(x1,p) \<in> fset xs \<and> (x2,e) \<in> fset xs" by force
then have 0: "dhead (Node r xs) dh p = root x1 \<and> dhead (Node r xs) dh e = root x2"
using dhead_in_set_eq_root[of x1] dhead_in_set_eq_root[of x2] wf_arcs by simp
have "x1 \<noteq> x2" using subtree_uneq_if_arc_uneq x_def ind.prems(3) by blast
then have "root x1 \<noteq> root x2"
using wf_verts x_def dtree.set_sel(1) unfolding wf_dverts_iff_dverts' by fastforce
then show ?thesis using 0 by argo
qed
qed
qed
lemma arc_in_subtree_if_tail_in_subtree:
assumes "dtail t dt p \<in> dverts x"
and "p \<in> darcs t"
and "t = Node r xs"
and "(x,e) \<in> fset xs"
shows "p \<in> darcs x"
proof (rule ccontr)
assume asm: "p \<notin> darcs x"
show False
proof(cases "p \<in> snd ` fset xs")
case True
then have "dtail t dt p = r" using assms(2,3) by simp
then show ?thesis using assms(1,3,4) root_not_subtree by force
next
case False
then obtain x' e1 where x'_def: "(x',e1) \<in> fset xs \<and> p \<in> darcs x'" using assms(2,3) by force
then have "x \<noteq> x'" using asm by blast
interpret X: wf_dtree x' using x'_def assms(3) wf_dtree_rec by blast
have "dtail t dt p = dtail x' dt p"
using dtail_in_child_eq_child[of x'] x'_def wf_arcs assms(3) by force
then have "dtail t dt p \<in> dverts x'" using X.dtail_in_dverts by (simp add: x'_def)
then have "dtail t dt p \<notin> dverts x"
using \<open>x\<noteq>x'\<close> wf_verts assms(3,4) x'_def unfolding wf_dverts_iff_dverts' by fastforce
then show ?thesis using assms(1) by blast
qed
qed
lemma dhead_in_verts_if_dtail:
assumes "dtail t dt p \<in> dverts x"
and "p \<in> darcs t"
and "t = Node r xs"
and "(x,e) \<in> fset xs"
shows "dhead t dh p \<in> dverts x"
proof -
interpret X: wf_dtree x using assms(3,4) wf_dtree_rec by blast
have 0: "p \<in> darcs x" using assms arc_in_subtree_if_tail_in_subtree by blast
then have "dhead t dh p = dhead x dh p"
using dhead_in_child_eq_child[of x] assms(3,4) wf_arcs by simp
then show ?thesis using X.dhead_in_dverts 0 by simp
qed
lemma cas_darcs_in_subtree:
assumes "pre_digraph.cas (from_dtree dt dh t) u ps v"
and "set ps \<subseteq> darcs t"
and "t = Node r xs"
and "(x,e) \<in> fset xs"
and "u \<in> dverts x"
shows "set ps \<subseteq> darcs x"
using assms proof(induction ps arbitrary: u)
case Nil
then show ?case by simp
next
case (Cons p ps)
note pre_digraph.cas.simps[simp]
then have u_p: "dtail t dt p = u" using Cons.prems(1) by simp
have "p \<in> darcs t" using Cons.prems(2) by simp
then have 0: "p \<in> darcs x" using arc_in_subtree_if_tail_in_subtree Cons.prems(3-5) u_p by blast
have 1: "dhead t dh p \<in> dverts x" using dhead_in_verts_if_dtail Cons.prems(2-5) u_p by force
have "set ps \<subseteq> darcs t" using Cons.prems(2) by simp
have "pre_digraph.cas (from_dtree dt dh t) (dhead t dh p) ps v" using Cons.prems(1) by simp
then have "set ps \<subseteq> darcs x" using Cons.IH Cons.prems(2,3,4) 1 by simp
then show ?case using 0 by simp
qed
lemma dtree_cas_in_subtree:
assumes "pre_digraph.cas (from_dtree dt dh t) u ps v"
and "set ps \<subseteq> darcs t"
and "t = Node r xs"
and "(x,e) \<in> fset xs"
and "u \<in> dverts x"
shows "pre_digraph.cas (from_dtree dt dh x) u ps v"
using assms cas_darcs_in_subtree dtree_cas_iff_subtree by fast
lemma cas_to_end_subtree:
assumes "set (p#ps) \<subseteq> darcs t" and "pre_digraph.cas (from_dtree dt dh t) (root t) (p#ps) v"
and "t = Node r xs" and "(x,e) \<in> fset xs" and "v \<in> dverts x"
shows "p = e"
proof (rule ccontr)
assume asm: "p \<noteq> e"
note pre_digraph.cas.simps[simp]
have "dtail t dt p = r" using assms(2,3) by simp
then have "p \<in> snd ` fset xs" using dtail_root_in_set assms(1,3) list.set_intros(1) by fast
then obtain x' where x'_def: "(x',p) \<in> fset xs" by force
show False
proof(cases "ps=[]")
case True
then have "root x' = v"
using dhead_in_set_eq_root[of x'] x'_def assms(2,3) wf_arcs by simp
then have "x = x'"
using wf_verts x'_def assms(3,4,5) dtree.set_sel(1) by (fastforce simp: wf_dverts_iff_dverts')
then show ?thesis using asm assms(3,4) subtree_uneq_if_arc_uneq x'_def by blast
next
case False
interpret X: wf_dtree x' using wf_dtree_rec x'_def assms(3) by blast
have "x' \<noteq> x" using asm assms(3,4) subtree_uneq_if_arc_uneq x'_def by blast
then have x'_no_v: "\<forall>e\<in>darcs x'. dhead x' dh e \<noteq> v"
using X.dhead_in_dverts assms(3,4,5) x'_def wf_verts
by (fastforce simp: wf_dverts_iff_dverts')
have 0: "pre_digraph.cas (from_dtree dt dh t) (dhead t dh p) ps v" using assms(2) by simp
have 1: "dhead t dh p \<in> dverts x'"
using dhead_in_set_eq_root[of x'] x'_def assms(3) dtree.set_sel(1) wf_arcs by auto
then have "pre_digraph.cas (from_dtree dt dh x') (dhead t dh p) ps v"
using dtree_cas_in_subtree x'_def assms(1,3) 0 by force
then have "\<not> set ps \<subseteq> darcs x'" using X.nohead_cas_no_arc_in_subset x'_no_v False by blast
moreover have "set ps \<subseteq> darcs x'" using cas_darcs_in_subtree assms(1,3) x'_def 0 1 by simp
ultimately show ?thesis by blast
qed
qed
lemma cas_unique_in_darcs: "\<lbrakk>v \<in> dverts t; pre_digraph.cas (from_dtree dt dh t) (root t) ps v;
pre_digraph.cas (from_dtree dt dh t) (root t) es v\<rbrakk>
\<Longrightarrow> ps = es \<or> \<not>set ps \<subseteq> darcs t \<or> \<not>set es \<subseteq> darcs t"
using wf_dtree_axioms proof(induction t arbitrary: ps es rule: darcs_mset.induct)
case ind: (1 r xs)
interpret wf_dtree "Node r xs" by (rule ind.prems(4))
show ?case
proof(cases "r=v")
case True
have 0: "\<forall>e \<in> darcs (Node r xs). dhead (Node r xs) dh e \<noteq> r" using dhead_not_root by force
consider "ps = [] \<and> es = []" | "ps \<noteq> []" | "es \<noteq> []" by blast
then show ?thesis
proof(cases)
case 1
then show ?thesis by blast
next
case 2
then show ?thesis using nohead_cas_no_arc_in_subset 0 ind.prems(2) True by blast
next
case 3
then show ?thesis using nohead_cas_no_arc_in_subset 0 ind.prems(3) True by blast
qed
next
case False
then obtain x e where x_def: "(x,e) \<in> fset xs" "v \<in> dverts x" using ind.prems by auto
then have wf_x: "wf_dtree x" using wf_dtree_rec by blast
note pre_digraph.cas.simps[simp]
have nempty: "ps \<noteq> [] \<and> es \<noteq> []" using ind.prems(2,3) False by force
then obtain p ps' where p_def: "ps = p # ps'" using list.exhaust_sel by auto
obtain e' es' where e'_def: "es = e' # es'" using list.exhaust_sel nempty by auto
show ?thesis
proof (rule ccontr)
assume "\<not>(ps = es \<or> \<not>set ps \<subseteq> darcs (Node r xs) \<or> \<not>set es \<subseteq> darcs (Node r xs))"
then have asm: "ps \<noteq> es \<and> set ps \<subseteq> darcs (Node r xs) \<and> set es \<subseteq> darcs (Node r xs)" by blast
then have "p = e" using cas_to_end_subtree p_def ind.prems(2) x_def by blast
moreover have "e' = e" using cas_to_end_subtree e'_def ind.prems(3) x_def asm by blast
ultimately have "p = e'" by blast
have "dhead (Node r xs) dh p = root x"
using dhead_in_set_eq_root[of x] x_def(1) \<open>p=e\<close> wf_arcs by simp
then have cas_p_r: "pre_digraph.cas (from_dtree dt dh (Node r xs)) (root x) ps' v"
using ind.prems(2) p_def by fastforce
moreover have 0: "root x \<in> dverts x" using dtree.set_sel(1) by blast
ultimately have cas_ps: "pre_digraph.cas (from_dtree dt dh x) (root x) ps' v"
using dtree_cas_in_subtree asm x_def(1) p_def dtree.set_sel(1) by force
have "dhead (Node r xs) dh e' = root x"
using dhead_in_set_eq_root[of x] x_def \<open>e'=e\<close> wf_arcs by simp
then have cas_e_r: "pre_digraph.cas (from_dtree dt dh (Node r xs)) (root x) es' v"
using ind.prems(3) e'_def by fastforce
then have "pre_digraph.cas (from_dtree dt dh x) (root x) es' v"
using dtree_cas_in_subtree asm x_def(1) e'_def 0 by force
then have "ps' = es' \<or> \<not> set ps' \<subseteq> darcs x \<or> \<not> set es' \<subseteq> darcs x"
using ind.IH cas_ps x_def wf_x by blast
moreover have "set ps' \<subseteq> darcs x"
using cas_darcs_in_subtree cas_p_r x_def(1) asm p_def 0 set_subset_Cons by fast
moreover have "set es' \<subseteq> darcs x"
using cas_darcs_in_subtree cas_e_r x_def(1) asm e'_def 0 set_subset_Cons by fast
ultimately have "ps' = es'" by blast
then show False using asm p_def e'_def \<open>p=e'\<close> by blast
qed
qed
qed
lemma dtree_awalk_unique:
"\<lbrakk>v \<in> dverts t; pre_digraph.awalk (from_dtree dt dh t) (root t) ps v;
pre_digraph.awalk (from_dtree dt dh t) (root t) es v\<rbrakk>
\<Longrightarrow> ps = es"
unfolding pre_digraph.awalk_def using cas_unique_in_darcs by fastforce
lemma dtree_unique_awalk_exists:
assumes "v \<in> dverts t"
shows "\<exists>!p. pre_digraph.awalk (from_dtree dt dh t) (root t) p v"
using dtree_awalk_exists dtree_awalk_unique assms by blast
lemma from_dtree_directed: "directed_tree (from_dtree dt dh t) (root t)"
apply(unfold_locales)
by(auto simp: dtail_in_dverts dhead_in_dverts dtree.set_sel(1) dtree_unique_awalk_exists)
theorem from_dtree_fin_directed: "finite_directed_tree (from_dtree dt dh t) (root t)"
apply(unfold_locales)
by(auto simp: dtail_in_dverts dhead_in_dverts dtree.set_sel(1) dtree_unique_awalk_exists
finite_dverts finite_darcs)
subsubsection "Identity of Transformation Operations"
lemma dhead_img_eq_root_img:
"Node r xs = t
\<Longrightarrow> (\<lambda>e. ((dhead (Node r xs) dh e), e)) ` snd ` fset xs = (\<lambda>(x,e). (root x, e)) ` fset xs"
using dhead_in_set_eq_root wf_arcs snd_conv image_image disjoint_darcs_if_wf_xs
by (smt (verit) case_prodE case_prod_conv image_cong)
lemma childarcs_in_out_arcs:
"\<lbrakk>Node r xs = t; e \<in> snd ` fset xs\<rbrakk> \<Longrightarrow> e \<in> out_arcs (from_dtree dt dh t) r"
by force
lemma out_arcs_in_childarcs:
assumes "Node r xs = t" and "e \<in> out_arcs (from_dtree dt dh t) r"
shows "e \<in> snd ` fset xs"
proof (rule ccontr)
assume asm: "e \<notin> snd ` fset xs"
have "e \<in> darcs t" using assms(2) by simp
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> e \<in> darcs x" using assms(1) asm by force
then have "dtail t dt e \<in> dverts x" using assms(1) dtail_in_childverts by blast
moreover have "r \<notin> dverts x" using assms(1) wf_verts x_def by (auto simp: wf_dverts_iff_dverts')
ultimately show False using assms(2) by simp
qed
lemma childarcs_eq_out_arcs:
"Node r xs = t \<Longrightarrow> snd ` fset xs = out_arcs (from_dtree dt dh t) r"
using childarcs_in_out_arcs out_arcs_in_childarcs by fast
lemma dtail_in_subtree_eq_subtree:
"\<lbrakk>is_subtree t1 t; e \<in> darcs t1\<rbrakk> \<Longrightarrow> dtail t def e = dtail t1 def e"
using wf_arcs proof(induction t rule: darcs_mset.induct)
case (1 r xs)
show ?case
proof(cases "Node r xs=t1")
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> is_subtree t1 x" using "1.prems"(1) by auto
then have "e \<in> darcs x" using "1.prems"(2) darcs_subtree_subset by blast
then have "dtail (Node r xs) def e = dtail x def e"
using dtail_in_child_eq_child[of x] x_def "1.prems"(3) by blast
then show ?thesis using "1.IH" x_def "1.prems"(2-3) by fastforce
qed (simp)
qed
lemma dtail_in_subdverts:
assumes "e \<in> darcs x" and "is_subtree x t"
shows "dtail t def e \<in> dverts x"
proof -
interpret X: wf_dtree x by (simp add: assms(2) wf_dtree_sub)
have "dtail t def e = dtail x def e" using dtail_in_subtree_eq_subtree assms(1,2) by blast
then show ?thesis using assms(1) X.dtail_in_dverts by simp
qed
lemma dhead_in_subtree_eq_subtree:
"\<lbrakk>is_subtree t1 t; e \<in> darcs t1\<rbrakk> \<Longrightarrow> dhead t def e = dhead t1 def e"
using wf_arcs proof(induction t)
case (Node r xs)
show ?case
proof(cases "Node r xs=t1")
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs \<and> is_subtree t1 x" using Node.prems(1) by auto
then have "e \<in> darcs x" using Node.prems(2) darcs_subtree_subset by blast
then have "dhead (Node r xs) def e = dhead x def e"
using dhead_in_child_eq_child[of x] x_def Node.prems(3) by force
then show ?thesis using Node.IH x_def Node.prems(2-3) by fastforce
qed (simp)
qed
lemma subarcs_in_out_arcs:
assumes "is_subtree (Node r xs) t" and "e \<in> snd ` fset xs"
shows "e \<in> out_arcs (from_dtree dt dh t) r"
proof -
have "e \<in> darcs (Node r xs)" using assms(2) by force
then have "tail (from_dtree dt dh t) e = r"
using dtail_in_subtree_eq_subtree assms(1,2) by auto
then show ?thesis using darcs_subtree_subset assms(1,2) by fastforce
qed
lemma darc_in_sub_if_dtail_in_sub:
assumes "dtail t dt e = v" and "e \<in> darcs t" and "(x,e1) \<in> fset xs"
and "is_subtree t1 x" and "Node r xs = t" and "v \<in> dverts t1"
shows "e \<in> darcs x"
proof (rule ccontr)
assume asm: "e \<notin> darcs x"
have "e \<notin> snd ` fset xs"
- using assms(1-6) asm arc_in_subtree_if_tail_in_subtree dverts_subtree_subset by blast
+ using assms(1-6) asm arc_in_subtree_if_tail_in_subtree dverts_subtree_subset
+ by (metis subset_eq)
then obtain x2 e2 where x2_def: "(x2,e2) \<in> fset xs \<and> e \<in> darcs x2" using assms(2,5) by force
then have "v \<in> dverts x" using assms(4,6) dverts_subtree_subset by fastforce
then have "v \<notin> dverts x2" using assms(1-3,5) arc_in_subtree_if_tail_in_subtree asm by blast
then have "dtail x2 dt e \<noteq> v" using assms(1,5) dtail_in_childverts x2_def by fast
then have "dtail t dt e = dtail x2 dt e"
using assms(1,5) x2_def \<open>v \<notin> dverts x2\<close> dtail_in_childverts by blast
then show False using assms(1) \<open>dtail x2 dt e \<noteq> v\<close> by simp
qed
lemma out_arcs_in_subarcs_aux:
assumes "is_subtree (Node r xs) t" and "dtail t dt e = r" and "e \<in> darcs t"
shows "e \<in> snd ` fset xs"
using assms wf_dtree_axioms proof(induction t)
case (Node v ys)
then interpret wf_dtree "Node v ys" by blast
show ?case
proof(cases "Node v ys = Node r xs")
case True
then show ?thesis using dtail_root_in_set Node.prems(2,3) by blast
next
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset ys \<and> is_subtree (Node r xs) x"
using Node.prems(1) by auto
then have "e \<in> darcs x"
using darc_in_sub_if_dtail_in_sub Node.prems(2,3) dtree.set_intros(1) by fast
moreover from this have "dtail x dt e = r"
using dtail_in_child_eq_child[of x] x_def Node.prems(2) wf_arcs by force
moreover from this have "wf_dtree x" using wf_verts wf_arcs x_def by(unfold_locales) auto
ultimately show ?thesis using Node.IH x_def by force
qed
qed
lemma out_arcs_in_subarcs:
assumes "is_subtree (Node r xs) t" and "e \<in> out_arcs (from_dtree dt dh t) r"
shows "e \<in> snd ` fset xs"
using assms out_arcs_in_subarcs_aux by auto
lemma subarcs_eq_out_arcs:
"is_subtree (Node r xs) t \<Longrightarrow> snd ` fset xs = out_arcs (from_dtree dt dh t) r"
using subarcs_in_out_arcs out_arcs_in_subarcs by fast
lemma dhead_sub_img_eq_root_img:
"is_subtree (Node v ys) t
\<Longrightarrow> (\<lambda>e. ((dhead t dh e), e)) ` snd ` fset ys = (\<lambda>(x,e). (root x, e)) ` fset ys"
using wf_dtree_axioms proof(induction t)
case (Node r xs)
then interpret wf_dtree "Node r xs" by blast
show ?case
proof(cases "Node v ys = Node r xs")
case True
then show ?thesis using dhead_img_eq_root_img by simp
next
case False
then obtain x e where x_def: "(x,e) \<in> fset xs \<and> is_subtree (Node v ys) x"
using Node.prems(1) by auto
then interpret X: wf_dtree x using wf_verts wf_arcs by(unfold_locales) auto
have "\<forall>a \<in> snd ` fset ys. (\<lambda>e. ((dhead (Node r xs) dh e), e)) a = (\<lambda>e. ((dhead x dh e), e)) a"
proof
fix a
assume asm: "a \<in> snd ` fset ys"
then have "a \<in> darcs x" using x_def darcs_subtree_subset by fastforce
then show "(\<lambda>e. ((dhead (Node r xs) dh e), e)) a = (\<lambda>e. ((dhead x dh e), e)) a"
using dhead_in_child_eq_child[of x] x_def wf_arcs by auto
qed
then have "(\<lambda>e. ((dhead (Node r xs) dh e), e)) ` snd ` fset ys
= (\<lambda>e. ((dhead x dh e), e)) ` snd ` fset ys"
by (meson image_cong)
then show ?thesis using Node.IH x_def X.wf_dtree_axioms by force
qed
qed
lemma subtree_to_dtree_aux_eq:
assumes "is_subtree x t" and "v \<in> dverts x"
shows "finite_directed_tree.to_dtree_aux (from_dtree dt dh t) v
= finite_directed_tree.to_dtree_aux (from_dtree dt dh x) v
\<and> finite_directed_tree.to_dtree_aux (from_dtree dt dh x) (root x) = x"
using assms wf_dtree_axioms proof(induction x arbitrary: t v rule: darcs_mset.induct)
case ind: (1 r xs)
then interpret wf_dtree t by blast
obtain r' xs' where r'_def: "t = Node r' xs'" using dtree.exhaust by auto
interpret R_xs: wf_dtree "Node r xs" using ind.prems(1,3) wf_dtree_sub by simp
let ?todt = "finite_directed_tree.to_dtree_aux"
let ?T = "(from_dtree dt dh t)"
let ?X = "(from_dtree dt dh (Node r xs))"
interpret DT: finite_directed_tree ?T "root t" using from_dtree_fin_directed by blast
interpret XT: finite_directed_tree ?X "root (Node r xs)"
using R_xs.from_dtree_fin_directed by blast
(* goal 2 *)
have ih: "\<forall>y \<in> fset xs. (\<lambda>(x,e). (XT.to_dtree_aux (root x), e)) y = y"
proof
fix y
assume asm: "y \<in> fset xs"
obtain x e where x_def: "y = (x,e)" by fastforce
- then have "is_subtree x (Node r xs)" using subtree_if_child asm by fastforce
+ then have "is_subtree x (Node r xs)" using subtree_if_child asm
+ by (metis image_iff prod.sel(1))
then have "?todt (from_dtree dt dh x) (root x) = x
\<and> XT.to_dtree_aux (root x) = ?todt (from_dtree dt dh x) (root x)"
using ind.IH R_xs.wf_dtree_axioms asm x_def dtree.set_sel(1) by blast
then have "XT.to_dtree_aux (root x) = x" by simp
then show "(\<lambda>(x,e). (XT.to_dtree_aux (root x), e)) y = y" using x_def by fast
qed
let ?f = "\<lambda>(x,e). (XT.to_dtree_aux x, e)"
let ?g = "\<lambda>e. ((dhead (Node r xs) dh e), e)"
obtain ys where ys_def: "XT.to_dtree_aux (root (Node r xs)) = Node r ys"
using dtree.exhaust dtree.sel(1) XT.to_dtree_aux_root by metis
then have "fset ys = (\<lambda>e. (XT.to_dtree_aux (head ?X e), e)) ` out_arcs ?X r"
using XT.dtree_children_img_alt XT.dtree_children_fset_id dtree.sel(1) by (smt (verit))
also have "\<dots> = (\<lambda>e. (XT.to_dtree_aux (dhead (Node r xs) dh e), e)) ` (snd ` fset xs)"
using R_xs.childarcs_eq_out_arcs by simp
also have "\<dots> = ?f ` ?g ` (snd ` fset xs)" by fast
also have "\<dots> = ?f ` (\<lambda>(x,e). (root x, e)) ` fset xs" using R_xs.dhead_img_eq_root_img by simp
also have "\<dots> = (\<lambda>(x,e). (XT.to_dtree_aux (root x), e)) ` fset xs" by fast
also have "\<dots> = fset xs" using ih by simp
finally have g2: "ys = xs" by (simp add: fset_inject)
show ?case
proof(cases "v = r")
case True
(* goal 1 *)
have 0: "\<forall>y \<in> fset xs. (\<lambda>(x,e). (DT.to_dtree_aux (root x), e)) y = y"
proof
fix y
assume asm: "y \<in> fset xs"
obtain x e where x_def: "y = (x,e)" by fastforce
- then have "is_subtree x (Node r xs)" using subtree_if_child asm by fastforce
+ then have "is_subtree x (Node r xs)" using subtree_if_child asm
+ by (metis image_iff prod.sel(1))
then have "is_subtree x t" using asm subtree_trans ind.prems(1) by blast
then have "?todt (from_dtree dt dh x) (root x) = x
\<and> DT.to_dtree_aux (root x) = ?todt (from_dtree dt dh x) (root x)"
using ind.IH wf_dtree_axioms asm x_def dtree.set_sel(1) by blast
then have "DT.to_dtree_aux (root x) = x" by simp
then show "(\<lambda>(x,e). (DT.to_dtree_aux (root x), e)) y = y" using x_def by fast
qed
let ?f = "\<lambda>(x,e). (DT.to_dtree_aux x, e)"
let ?g = "\<lambda>e. ((dhead (Node r' xs') dh e), e)"
obtain zs where zs_def: "DT.to_dtree_aux v = Node v zs"
using dtree.exhaust by simp
then have "fset zs = (\<lambda>e. (DT.to_dtree_aux (head ?T e), e)) ` out_arcs ?T r"
using DT.dtree_children_img_alt DT.dtree_children_fset_id True by presburger
also have "\<dots> = (\<lambda>e. (DT.to_dtree_aux (dhead t dh e), e)) ` (snd ` fset xs)"
using ind.prems(1) subarcs_eq_out_arcs by force
also have "\<dots> = ?f ` ?g ` (snd ` fset xs)" using r'_def by fast
also have "\<dots> = ?f ` (\<lambda>(x,e). (root x, e)) ` fset xs"
using dhead_sub_img_eq_root_img ind.prems(1) r'_def by blast
also have "\<dots> = (\<lambda>(x,e). (DT.to_dtree_aux (root x), e)) ` fset xs" by fast
also have "\<dots> = fset xs" using 0 by simp
finally have g1: "zs = xs" by (simp add: fset_inject)
then show ?thesis using zs_def True g2 ys_def by simp
next
case False
(* goal 1 *)
then obtain x1 e1 where x_def: "(x1,e1) \<in> fset xs" "v \<in> dverts x1" using ind.prems(2) by auto
- then have "is_subtree x1 (Node r xs)" using subtree_if_child by fastforce
+ then have "is_subtree x1 (Node r xs)" using subtree_if_child
+ by (metis image_iff prod.sel(1))
moreover from this have "is_subtree x1 t" using ind.prems(1) subtree_trans by blast
ultimately have g1: "DT.to_dtree_aux v = XT.to_dtree_aux v"
using ind.IH x_def by (metis R_xs.wf_dtree_axioms wf_dtree_axioms)
then show ?thesis using g1 g2 ys_def by blast
qed
qed
interpretation T: finite_directed_tree "from_dtree dt dh t" "root t"
using from_dtree_fin_directed by simp
lemma to_from_dtree_aux_id: "T.to_dtree_aux dt dh (root t) = t"
using subtree_to_dtree_aux_eq dtree.set_sel(1) self_subtree by blast
theorem to_from_dtree_id: "T.to_dtree dt dh = t"
using to_from_dtree_aux_id T.to_dtree_def by simp
end
context finite_directed_tree
begin
lemma wf_to_dtree_aux: "wf_dtree (to_dtree_aux r)"
unfolding wf_dtree_def using wf_dverts_to_dtree_aux wf_darcs_to_dtree_aux by blast
theorem wf_to_dtree: "wf_dtree to_dtree"
unfolding to_dtree_def using wf_to_dtree_aux by blast
end
subsection \<open>Degrees of Nodes\<close>
fun max_deg :: "('a,'b) dtree \<Rightarrow> nat" where
"max_deg (Node r xs) = (if xs = {||} then 0 else max (Max (max_deg ` fst ` fset xs)) (fcard xs))"
lemma mdeg_eq_fcard_if_empty: "xs = {||} \<Longrightarrow> max_deg (Node r xs) = fcard xs"
by simp
lemma mdeg0_if_fcard0: "fcard xs = 0 \<Longrightarrow> max_deg (Node r xs) = 0"
by simp
lemma mdeg0_iff_fcard0: "fcard xs = 0 \<longleftrightarrow> max_deg (Node r xs) = 0"
by simp
lemma nempty_if_mdeg_gt_fcard: "max_deg (Node r xs) > fcard xs \<Longrightarrow> xs \<noteq> {||}"
by auto
lemma mdeg_img_nempty: "max_deg (Node r xs) > fcard xs \<Longrightarrow> max_deg ` fst ` fset xs \<noteq> {}"
- using nempty_if_mdeg_gt_fcard notin_fset[where S=xs] by fast
+ using nempty_if_mdeg_gt_fcard[of xs] by fast
lemma mdeg_img_fin: "finite (max_deg ` fst ` fset xs)"
by simp
lemma mdeg_Max_if_gt_fcard:
"max_deg (Node r xs) > fcard xs \<Longrightarrow> max_deg (Node r xs) = Max (max_deg ` fst ` fset xs)"
by (auto split: if_splits)
lemma mdeg_child_if_gt_fcard:
"max_deg (Node r xs) > fcard xs \<Longrightarrow> \<exists>t \<in> fst ` fset xs. max_deg t = max_deg (Node r xs)"
unfolding mdeg_Max_if_gt_fcard using Max_in[OF mdeg_img_fin mdeg_img_nempty] by force
lemma mdeg_child_if_wedge:
"\<lbrakk>max_deg (Node r xs) > n; fcard xs \<le> n \<or> \<not>(\<forall>t \<in> fst ` fset xs. max_deg t \<le> n)\<rbrakk>
\<Longrightarrow> \<exists>t \<in> fst ` fset xs. max_deg t > n"
- using mdeg_child_if_gt_fcard by force
+ using mdeg_child_if_gt_fcard[of xs] by force
lemma maxif_eq_Max: "finite X \<Longrightarrow> (if X \<noteq> {} then max x (Max X) else x) = Max (insert x X)"
by simp
lemma mdeg_img_empty_iff: "max_deg ` fst ` fset xs = {} \<longleftrightarrow> xs = {||}"
- using notin_fset by fast
+ by fast
lemma mdeg_alt: "max_deg (Node r xs) = Max (insert (fcard xs) (max_deg ` fst ` fset xs))"
using maxif_eq_Max[OF mdeg_img_fin, of xs "fcard xs"] mdeg_img_empty_iff[of xs]
by (auto split: if_splits)
lemma finite_fMax_union: "finite Y \<Longrightarrow> finite (\<Union>y\<in>Y. {Max (f y)})"
by blast
lemma Max_union_Max_out:
assumes "finite Y" and "\<forall>y \<in> Y. finite (f y)" and "\<forall>y \<in> Y. f y \<noteq> {}" and "Y \<noteq> {}"
shows "Max (\<Union>y\<in>Y. {Max (f y)}) = Max (\<Union>y\<in>Y. f y)" (is "?M1=_")
proof -
have "\<forall>y \<in> Y. \<forall>x \<in> f y. Max (f y) \<ge> x" using assms(2) by simp
moreover have "\<forall>x \<in> (\<Union>y\<in>Y. {Max (f y)}). ?M1 \<ge> x" using assms(1) by simp
moreover have M1_in: "?M1 \<in> (\<Union>y\<in>Y. {Max (f y)})"
using assms(1,4) Max_in[OF finite_fMax_union] by auto
ultimately have "\<forall>y \<in> Y. \<forall>x \<in> f y. ?M1 \<ge> x" by force
then have "\<forall>x \<in> (\<Union>y\<in>Y. f y). ?M1 \<ge> x" by blast
moreover have "?M1 \<in> (\<Union>y\<in>Y. (f y))" using M1_in assms(2-4) by force
ultimately show ?thesis using assms(1,2) Max_eqI finite_UN_I by metis
qed
lemma Max_union_Max_out_insert:
"\<lbrakk>finite Y; \<forall>y \<in> Y. finite (f y); \<forall>y \<in> Y. f y \<noteq> {}; Y \<noteq> {}\<rbrakk>
\<Longrightarrow> Max (insert x (\<Union>y\<in>Y. {Max (f y)})) = Max (insert x (\<Union>y\<in>Y. f y))"
using Max_union_Max_out[of Y f] by simp
lemma mdeg_alt2: "max_deg t = Max {fcard (sucs x)|x. is_subtree x t}"
proof(induction t rule: max_deg.induct)
case (1 r xs)
then show ?case
proof(cases "xs = {||}")
case False
let ?f = "(\<lambda>t1. {fcard (sucs x)|x. is_subtree x t1})"
let ?f' = "(\<lambda>t1. (\<lambda>x. fcard (sucs x)) ` {x. is_subtree x t1})"
have fin: "finite (fst ` fset xs)" by simp
have f_eq1: "?f = ?f'" by blast
then have f_eq: "\<forall>y\<in>fst ` fset xs. (?f y = ?f' y)" by blast
moreover have "\<forall>y\<in>fst ` fset xs. finite (?f' y)" using finite_subtrees by blast
ultimately have fin': "\<forall>y\<in>fst ` fset xs. finite (?f y)" by simp
have nempty: "\<forall>y\<in>fst ` fset xs. {fcard (sucs x) |x. is_subtree x y} \<noteq> {}"
using self_subtree by blast
have "max_deg ` fst ` fset xs = (\<Union>t1\<in>fst ` fset xs. {Max (?f t1)})"
using "1.IH"[OF False] by auto
then have "max_deg (Node r xs) = Max (insert (fcard xs) (\<Union>t1\<in>fst ` fset xs. {Max (?f t1)}))"
using mdeg_alt[of r xs] by simp
also have "\<dots> = Max (insert (fcard xs) (\<Union>t1\<in>fst ` fset xs. ?f t1))"
using Max_union_Max_out_insert[OF fin fin' nempty] by fastforce
also have "\<dots> = Max (insert (fcard xs) ((\<Union>t1\<in>fst ` fset xs. ?f' t1)))"
using f_eq by simp
also have "\<dots>
= Max (insert (fcard xs) ((\<Union>t1\<in>fst ` fset xs. fcard ` sucs ` {x. is_subtree x t1})))"
using image_image by metis
also have "\<dots>
= Max (insert (fcard xs) (fcard ` sucs ` (\<Union>t1\<in>fst ` fset xs. {x. is_subtree x t1})))"
by (metis image_UN)
also have "\<dots>
= Max (fcard ` sucs ` (insert (Node r xs) (\<Union>t1\<in>fst ` fset xs. {x. is_subtree x t1})))"
by force
also have "\<dots> = Max (fcard ` sucs ` {x. is_subtree x (Node r xs)})"
unfolding subtrees_insert_union by blast
finally show ?thesis using f_eq1 image_image by metis
qed(simp)
qed
lemma mdeg_singleton: "max_deg (Node r {|(t1,e1)|}) = max (max_deg t1) (fcard {|(t1,e1)|})"
by simp
lemma mdeg_ge_child_aux: "(t1,e1) \<in> fset xs \<Longrightarrow> max_deg t1 \<le> Max (max_deg ` fst ` fset xs)"
using Max_ge[OF mdeg_img_fin] by fastforce
lemma mdeg_ge_child: "(t1,e1) \<in> fset xs \<Longrightarrow> max_deg t1 \<le> max_deg (Node r xs)"
using mdeg_ge_child_aux by fastforce
lemma mdeg_ge_child': "t1 \<in> fst ` fset xs \<Longrightarrow> max_deg t1 \<le> max_deg (Node r xs)"
using mdeg_ge_child[of t1] by force
lemma mdeg_ge_sub: "is_subtree t1 t2 \<Longrightarrow> max_deg t1 \<le> max_deg t2"
proof(induction t2)
case (Node r xs)
show ?case
proof(cases "Node r xs=t1")
case False
then obtain x e1 where x_def: "(x,e1) \<in> fset xs" "is_subtree t1 x" using Node.prems(1) by auto
then have "max_deg t1 \<le> max_deg x" using Node.IH by force
then show ?thesis using mdeg_ge_child[OF x_def(1)] by simp
qed(simp)
qed
lemma mdeg_gt_0_if_nempty: "xs \<noteq> {||} \<Longrightarrow> max_deg (Node r xs) > 0"
using fcard_fempty by auto
corollary empty_if_mdeg_0: "max_deg (Node r xs) = 0 \<Longrightarrow> xs = {||}"
using mdeg_gt_0_if_nempty by (metis less_numeral_extra(3))
lemma nempty_if_mdeg_n0: "max_deg (Node r xs) \<noteq> 0 \<Longrightarrow> xs \<noteq> {||}"
by auto
corollary empty_iff_mdeg_0: "max_deg (Node r xs) = 0 \<longleftrightarrow> xs = {||}"
using nempty_if_mdeg_n0 empty_if_mdeg_0 by auto
lemma mdeg_root: "max_deg (Node r xs) = max_deg (Node v xs)"
by simp
lemma mdeg_ge_fcard: "fcard xs \<le> max_deg (Node r xs)"
by simp
lemma mdeg_fcard_if_fcard_ge_child:
"\<forall>(t,e) \<in> fset xs. max_deg t \<le> fcard xs \<Longrightarrow> max_deg (Node r xs) = fcard xs"
using mdeg_child_if_gt_fcard[of xs r] mdeg_ge_fcard[of xs r] by fastforce
lemma mdeg_fcard_if_fcard_ge_child':
"\<forall>t \<in> fst ` fset xs. max_deg t \<le> fcard xs \<Longrightarrow> max_deg (Node r xs) = fcard xs"
using mdeg_fcard_if_fcard_ge_child[of xs r] by fastforce
lemma fcard_single_1: "fcard {|x|} = 1"
by (simp add: fcard_finsert)
lemma fcard_single_1_iff: "fcard xs = 1 \<longleftrightarrow> (\<exists>x. xs = {|x|})"
by (metis all_not_fin_conv bot.extremum fcard_seteq fcard_single_1
finsert_fsubset le_numeral_extra(4))
lemma fcard_not0_if_elem: "\<exists>x. x \<in> fset xs \<Longrightarrow> fcard xs \<noteq> 0"
by auto
lemma fcard1_if_le1_elem: "\<lbrakk>fcard xs \<le> 1; x \<in> fset xs\<rbrakk> \<Longrightarrow> fcard xs = 1"
- using fcard_not0_if_elem by fastforce
+ using fcard_not0_if_elem[of xs] by fastforce
lemma singleton_if_fcard_le1_elem: "\<lbrakk>fcard xs \<le> 1; x \<in> fset xs\<rbrakk> \<Longrightarrow> xs = {|x|}"
using fcard_single_1_iff[of xs] fcard1_if_le1_elem by fastforce
lemma singleton_if_mdeg_le1_elem: "\<lbrakk>max_deg (Node r xs) \<le> 1; x \<in> fset xs\<rbrakk> \<Longrightarrow> xs = {|x|}"
using singleton_if_fcard_le1_elem[of xs] mdeg_ge_fcard[of xs] by simp
lemma singleton_if_mdeg_le1_elem_suc: "\<lbrakk>max_deg t \<le> 1; x \<in> fset (sucs t)\<rbrakk> \<Longrightarrow> sucs t = {|x|}"
using singleton_if_mdeg_le1_elem[of "root t" "sucs t"] by simp
lemma fcard0_if_le1_not_singleton: "\<lbrakk>\<forall>x. xs \<noteq> {|x|}; fcard xs \<le> 1\<rbrakk> \<Longrightarrow> fcard xs = 0"
using fcard_single_1_iff[of xs] by fastforce
lemma empty_fset_if_fcard_le1_not_singleton: "\<lbrakk>\<forall>x. xs \<noteq> {|x|}; fcard xs \<le> 1\<rbrakk> \<Longrightarrow> xs = {||}"
using fcard0_if_le1_not_singleton by auto
lemma fcard0_if_mdeg_le1_not_single: "\<lbrakk>\<forall>x. xs \<noteq> {|x|}; max_deg (Node r xs) \<le> 1\<rbrakk> \<Longrightarrow> fcard xs = 0"
using fcard0_if_le1_not_singleton[of xs] mdeg_ge_fcard[of xs] by simp
lemma empty_fset_if_mdeg_le1_not_single: "\<lbrakk>\<forall>x. xs \<noteq> {|x|}; max_deg (Node r xs) \<le> 1\<rbrakk> \<Longrightarrow> xs = {||}"
using fcard0_if_mdeg_le1_not_single by auto
lemma fcard0_if_mdeg_le1_not_single_suc:
"\<lbrakk>\<forall>x. sucs t \<noteq> {|x|}; max_deg t \<le> 1\<rbrakk> \<Longrightarrow> fcard (sucs t) = 0"
using fcard0_if_mdeg_le1_not_single[of "sucs t" "root t"] by simp
lemma empty_fset_if_mdeg_le1_not_single_suc: "\<lbrakk>\<forall>x. sucs t \<noteq> {|x|}; max_deg t \<le> 1\<rbrakk> \<Longrightarrow> sucs t = {||}"
using fcard0_if_mdeg_le1_not_single_suc by auto
lemma mdeg_1_singleton:
assumes "max_deg (Node r xs) = 1"
shows "\<exists>x. xs = {|x|}"
proof -
obtain x where x_def: "x |\<in>| xs"
using assms by (metis all_not_fin_conv empty_iff_mdeg_0 zero_neq_one)
moreover have "fcard xs \<le> 1" using assms mdeg_ge_fcard by metis
ultimately have "xs = {|x|}"
by (metis order_bot_class.bot.not_eq_extremum diff_Suc_1 diff_is_0_eq' fcard_finsert_disjoint
less_nat_zero_code mk_disjoint_finsert pfsubset_fcard_mono)
then show ?thesis by simp
qed
lemma subtree_child_if_dvert_notr_mdeg_le1:
assumes "max_deg (Node r xs) \<le> 1" and "v \<noteq> r" and "v \<in> dverts (Node r xs)"
shows "\<exists>r' e zs. is_subtree (Node r' {|(Node v zs,e)|}) (Node r xs)"
proof -
obtain r' ys zs where zs_def: "is_subtree (Node r' ys) (Node r xs)" "Node v zs \<in> fst ` fset ys"
using subtree_child_if_dvert_notroot[OF assms(2,3)] by blast
have 0: "max_deg (Node r' ys) \<le> 1" using mdeg_ge_sub[OF zs_def(1)] assms(1) by simp
obtain e where "{|(Node v zs,e)|} = ys"
using singleton_if_mdeg_le1_elem[OF 0] zs_def(2) by fastforce
then show ?thesis using zs_def(1) by blast
qed
lemma subtree_child_if_dvert_notroot_mdeg_le1:
"\<lbrakk>max_deg t \<le> 1; v \<noteq> root t; v \<in> dverts t\<rbrakk>
\<Longrightarrow> \<exists>r' e zs. is_subtree (Node r' {|(Node v zs,e)|}) t"
using subtree_child_if_dvert_notr_mdeg_le1[of "root t" "sucs t"] by simp
lemma mdeg_child_sucs_eq_if_gt1:
assumes "max_deg (Node r {|(t,e)|}) > 1"
shows "max_deg (Node r {|(t,e)|}) = max_deg (Node v (sucs t))"
proof -
have "fcard {|(t,e)|} = 1" using fcard_single_1 by fast
then have "max_deg (Node r {|(t,e)|}) = max_deg t" using assms by simp
then show ?thesis using mdeg_root[of "root t" "sucs t" v] dtree.exhaust_sel[of t] by argo
qed
lemma mdeg_child_sucs_le: "max_deg (Node v (sucs t)) \<le> max_deg (Node r {|(t,e)|})"
using mdeg_root[of v "sucs t" "root t"] by simp
lemma mdeg_eq_child_if_singleton_gt1:
"max_deg (Node r {|(t1,e1)|}) > 1 \<Longrightarrow> max_deg (Node r {|(t1,e1)|}) = max_deg t1"
using mdeg_singleton[of r t1] by (auto simp: fcard_single_1 max_def)
lemma fcard_gt1_if_mdeg_gt_child:
assumes "max_deg (Node r xs) > n" and "t1 \<in> fst ` fset xs" and "max_deg t1 \<le> n" and "n\<noteq>0"
shows "fcard xs > 1"
proof(rule ccontr)
assume "\<not>fcard xs > 1"
then have "fcard xs \<le> 1" by simp
then have "\<exists>e1. xs = {|(t1,e1)|}" using assms(2) singleton_if_fcard_le1_elem by fastforce
then show False using mdeg_singleton[of r t1] assms(1,3,4) by (auto simp: fcard_single_1)
qed
lemma fcard_gt1_if_mdeg_gt_suc:
"\<lbrakk>max_deg t2 > n; t1 \<in> fst ` fset (sucs t2); max_deg t1 \<le> n; n\<noteq>0\<rbrakk> \<Longrightarrow> fcard (sucs t2) > 1"
using fcard_gt1_if_mdeg_gt_child[of n "root t2" "sucs t2"] by simp
lemma fcard_gt1_if_mdeg_gt_child1:
"\<lbrakk>max_deg (Node r xs) > 1; t1 \<in> fst ` fset xs; max_deg t1 \<le> 1\<rbrakk> \<Longrightarrow> fcard xs > 1"
using fcard_gt1_if_mdeg_gt_child by auto
lemma fcard_gt1_if_mdeg_gt_suc1:
"\<lbrakk>max_deg t2 > 1; t1 \<in> fst ` fset (sucs t2); max_deg t1 \<le> 1\<rbrakk> \<Longrightarrow> fcard (sucs t2) > 1"
using fcard_gt1_if_mdeg_gt_suc by blast
lemma fcard_lt_non_inj_f:
"\<lbrakk>f a = f b; a \<in> fset xs; b \<in> fset xs; a\<noteq>b\<rbrakk> \<Longrightarrow> fcard (f |`| xs) < fcard xs"
proof(induction xs)
case (insert x xs)
then consider "a \<in> fset xs" "b \<in> fset xs" | "a = x" "b \<in> fset xs" | "a \<in> fset xs" "b = x"
by auto
then show ?case
proof(cases)
case 1
then show ?thesis
using insert.IH insert.prems(1,4) by (simp add: fcard_finsert_if)
next
case 2
then show ?thesis
proof(cases "fcard (f |`| xs) = fcard xs")
case True
then show ?thesis
using 2 insert.hyps insert.prems(1)
- by (metis fcard_finsert_disjoint fimage_finsert finsert_fimage lessI notin_fset)
+ by (metis fcard_finsert_disjoint fimage_finsert finsert_fimage lessI)
next
case False
then have "fcard (f |`| xs) \<le> fcard xs" using fcard_image_le by auto
then have "fcard (f |`| xs) < fcard xs" using False by simp
then show ?thesis
- using 2 insert.prems(1) notin_fset fcard_image_le fcard_mono fimage_finsert less_le_not_le
+ using 2 insert.prems(1) fcard_image_le fcard_mono fimage_finsert less_le_not_le
by (metis order_class.order.not_eq_order_implies_strict finsert_fimage fsubset_finsertI)
qed
next
case 3
then show ?thesis
proof(cases "fcard (f |`| xs) = fcard xs")
case True
then show ?thesis
using 3 insert.hyps insert.prems(1)
- by (metis fcard_finsert_disjoint fimage_finsert finsert_fimage lessI notin_fset)
+ by (metis fcard_finsert_disjoint fimage_finsert finsert_fimage lessI)
next
case False
then have "fcard (f |`| xs) \<le> fcard xs" using fcard_image_le by auto
then have "fcard (f |`| xs) < fcard xs" using False by simp
then show ?thesis
- using 3 insert.prems(1) notin_fset fcard_image_le fcard_mono fimage_finsert less_le_not_le
+ using 3 insert.prems(1) fcard_image_le fcard_mono fimage_finsert less_le_not_le
by (metis order_class.order.not_eq_order_implies_strict finsert_fimage fsubset_finsertI)
qed
qed
qed (simp)
lemma mdeg_img_le:
assumes "\<forall>(t,e) \<in> fset xs. max_deg (fst (f (t,e))) \<le> max_deg t"
shows "max_deg (Node r (f |`| xs)) \<le> max_deg (Node r xs)"
proof(cases "max_deg (Node r (f |`| xs)) = fcard (f |`| xs)")
case True
then show ?thesis using fcard_image_le[of f xs] by auto
next
case False
then have "max_deg (Node r (f |`| xs)) > fcard (f |`| xs)"
using mdeg_ge_fcard[of "f |`| xs"] by simp
then obtain t1 e1 where t1_def:
"(t1,e1) \<in> fset (f |`| xs)" "max_deg t1 = max_deg (Node r (f |`| xs))"
- using mdeg_child_if_gt_fcard[of "f |`| xs" r] by fastforce
+ using mdeg_child_if_gt_fcard[of "f |`| xs" r]
+ by (metis (no_types, opaque_lifting) fst_conv imageE surj_pair)
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "f (t2,e2) = (t1,e1)" by auto
then have "max_deg t2 \<ge> max_deg (Node r (f |`| xs))" using t1_def(2) assms by fastforce
then show ?thesis using mdeg_ge_child[OF t2_def(1)] by simp
qed
lemma mdeg_img_le':
assumes "\<forall>(t,e) \<in> fset xs. max_deg (f t) \<le> max_deg t"
shows "max_deg (Node r ((\<lambda>(t,e). (f t, e)) |`| xs)) \<le> max_deg (Node r xs)"
using mdeg_img_le[of xs "\<lambda>(t,e). (f t, e)"] assms by simp
lemma mdeg_le_if_fcard_and_child_le:
"\<lbrakk>\<forall>(t,e) \<in> fset xs. max_deg t \<le> m; fcard xs \<le> m\<rbrakk> \<Longrightarrow> max_deg (Node r xs) \<le> m"
using mdeg_ge_fcard mdeg_child_if_gt_fcard[of xs r] by fastforce
lemma mdeg_child_if_child_max:
"\<lbrakk>\<forall>(t,e) \<in> fset xs. max_deg t \<le> max_deg t1; fcard xs \<le> max_deg t1; (t1,e1) \<in> fset xs\<rbrakk>
\<Longrightarrow> max_deg (Node r xs) = max_deg t1"
using mdeg_le_if_fcard_and_child_le[of xs "max_deg t1"] mdeg_ge_child[of t1 e1 xs] by simp
corollary mdeg_child_if_child_max':
"\<lbrakk>\<forall>(t,e) \<in> fset xs. max_deg t \<le> max_deg t1; fcard xs \<le> max_deg t1; t1 \<in> fst ` fset xs\<rbrakk>
\<Longrightarrow> max_deg (Node r xs) = max_deg t1"
using mdeg_child_if_child_max[of xs t1] by force
lemma mdeg_img_eq:
assumes "\<forall>(t,e) \<in> fset xs. max_deg (fst (f (t,e))) = max_deg t"
and "fcard (f |`| xs) = fcard xs"
shows "max_deg (Node r (f |`| xs)) = max_deg (Node r xs)"
proof(cases "max_deg (Node r (f |`| xs)) = fcard (f |`| xs)")
case True
then have "\<forall>(t,e) \<in> fset (f |`| xs). max_deg t \<le> fcard (f |`| xs)"
- using mdeg_ge_child by fastforce
+ using mdeg_ge_child
+ by (metis (mono_tags, lifting) case_prodI2)
then have "\<forall>(t,e) \<in> fset xs. max_deg t \<le> fcard xs" using assms by fastforce
then have "max_deg (Node r xs) = fcard xs" using mdeg_fcard_if_fcard_ge_child by fast
then show ?thesis using True assms(2) by simp
next
case False
then have "max_deg (Node r (f |`| xs)) > fcard (f |`| xs)"
using mdeg_ge_fcard[of "f |`| xs"] by simp
then obtain t1 e1 where t1_def:
"(t1,e1) \<in> fset (f |`| xs)" "max_deg t1 = max_deg (Node r (f |`| xs))"
- using mdeg_child_if_gt_fcard[of "f |`| xs" r] by fastforce
+ using mdeg_child_if_gt_fcard[of "f |`| xs" r]
+ by (metis (no_types, opaque_lifting) fst_conv imageE old.prod.exhaust)
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "f (t2,e2) = (t1,e1)" by auto
then have mdeg_t21: "max_deg t2 = max_deg t1" using assms(1) by auto
have "\<forall>(t3,e3) \<in> fset (f |`| xs). max_deg t3 \<le> max_deg t1"
- using t1_def(2) mdeg_ge_child[where xs="f |`| xs"] by force
+ using t1_def(2) mdeg_ge_child[where xs="f |`| xs"]
+ by (metis (no_types, lifting) case_prodI2)
then have "\<forall>(t3,e3) \<in> fset xs. max_deg (fst (f (t3,e3))) \<le> max_deg t1" by auto
then have "\<forall>(t3,e3) \<in> fset xs. max_deg t3 \<le> max_deg t2" using assms(1) mdeg_t21 by fastforce
moreover have "max_deg t2 \<ge> fcard xs" using t1_def(2) assms(2) mdeg_t21 by simp
ultimately have "max_deg (Node r xs) = max_deg t2"
- using t2_def(1) mdeg_child_if_child_max by fast
+ using t2_def(1) mdeg_child_if_child_max by metis
then show ?thesis using t1_def(2) mdeg_t21 by simp
qed
lemma num_leaves_1_if_mdeg_1: "max_deg t \<le> 1 \<Longrightarrow> num_leaves t = 1"
proof(induction t)
case (Node r xs)
then show ?case
proof(cases "max_deg (Node r xs) = 0")
case True
then show ?thesis using empty_iff_mdeg_0 by auto
next
case False
then have "max_deg (Node r xs) = 1" using Node.prems by simp
then obtain t e where t_def: "xs = {|(t,e)|}" "(t,e) \<in> fset xs"
using mdeg_1_singleton by fastforce
then have "max_deg t \<le> 1" using Node.prems mdeg_ge_child by fastforce
then show ?thesis using Node.IH t_def(1) by simp
qed
qed
lemma num_leaves_ge1: "num_leaves t \<ge> 1"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "xs = {||}")
case False
- then obtain t e where t_def: "(t,e) \<in> fset xs" using notin_fset by fast
+ then obtain t e where t_def: "(t,e) \<in> fset xs" by fast
then have "1 \<le> num_leaves t" using Node by simp
then show ?thesis
using fset_sum_ge_elem[OF finite_fset[of xs] t_def, of "\<lambda>(t,e). num_leaves t"] by auto
qed (simp)
qed
lemma num_leaves_ge_card: "num_leaves (Node r xs) \<ge> fcard xs"
proof(cases "xs = {||}")
case False
have "fcard xs = (\<Sum>x\<in> fset xs. 1)" using fcard.rep_eq by auto
also have "\<dots> \<le> (\<Sum>x\<in> fset xs. num_leaves (fst x))" using num_leaves_ge1 sum_mono by metis
finally show ?thesis using False by (simp add: fst_def prod.case_distrib)
qed (simp add: fcard_fempty)
lemma num_leaves_root: "num_leaves (Node r xs) = num_leaves (Node r' xs)"
by simp
lemma num_leaves_singleton: "num_leaves (Node r {|(t,e)|}) = num_leaves t"
by simp
subsection \<open>List Conversions\<close>
function dtree_to_list :: "('a,'b) dtree \<Rightarrow> ('a\<times>'b) list" where
"dtree_to_list (Node r {|(t,e)|}) = (root t,e) # dtree_to_list t"
| "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> dtree_to_list (Node r xs) = []"
by (metis darcs_mset.cases surj_pair) auto
termination by lexicographic_order
fun dtree_from_list :: "'a \<Rightarrow> ('a\<times>'b) list \<Rightarrow> ('a,'b) dtree" where
"dtree_from_list r [] = Node r {||}"
| "dtree_from_list r ((v,e)#xs) = Node r {|(dtree_from_list v xs, e)|}"
fun wf_list_arcs :: "('a\<times>'b) list \<Rightarrow> bool" where
"wf_list_arcs [] = True"
| "wf_list_arcs ((v,e)#xs) = (e \<notin> snd ` set xs \<and> wf_list_arcs xs)"
fun wf_list_verts :: "('a\<times>'b) list \<Rightarrow> bool" where
"wf_list_verts [] = True"
| "wf_list_verts ((v,e)#xs) = (v \<notin> fst ` set xs \<and> wf_list_verts xs)"
lemma dtree_to_list_sub_dverts_ins:
"insert (root t) (fst ` set (dtree_to_list t)) \<subseteq> dverts t"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case False
then obtain t e where t_def: "xs = {|(t,e)|}"
using mdeg_1_singleton by fastforce
then show ?thesis using Node.IH by fastforce
qed (auto)
qed
lemma dtree_to_list_eq_dverts_ins:
"max_deg t \<le> 1 \<Longrightarrow> insert (root t) (fst ` set (dtree_to_list t)) = dverts t"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "max_deg (Node r xs) = 0")
case True
then have "xs = {||}" using empty_iff_mdeg_0 by auto
moreover from this have "\<forall>x. xs \<noteq> {|x|}" by blast
ultimately show ?thesis by simp
next
case False
then have "max_deg (Node r xs) = 1" using Node.prems by simp
then obtain t e where t_def: "xs = {|(t,e)|}" "(t,e) \<in> fset xs"
using mdeg_1_singleton by fastforce
then have "max_deg t \<le> 1" using Node.prems mdeg_ge_child by fastforce
then have "insert (root t) (fst ` set (dtree_to_list t)) = dverts t"
using Node.IH t_def(2) by auto
then show ?thesis using Node.prems(1) t_def(1) by simp
qed
qed
lemma dtree_to_list_eq_dverts_sucs:
"max_deg t \<le> 1 \<Longrightarrow> fst ` set (dtree_to_list t) = (\<Union>x \<in> fset (sucs t). dverts (fst x))"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "max_deg (Node r xs) = 0")
case True
then have "xs = {||}" using empty_iff_mdeg_0 by auto
moreover from this have "\<forall>x. xs \<noteq> {|x|}" by blast
ultimately show ?thesis by simp
next
case False
then have "max_deg (Node r xs) = 1" using Node.prems by simp
then obtain t e where t_def: "xs = {|(t,e)|}" "(t,e) \<in> fset xs"
using mdeg_1_singleton by fastforce
then have "max_deg t \<le> 1" using Node.prems mdeg_ge_child by fastforce
then have "fst ` set (dtree_to_list t) = (\<Union>x \<in> fset (sucs t). dverts (fst x))"
using Node.IH t_def(2) by auto
moreover from this have "dverts t = insert (root t) (\<Union>x \<in> fset (sucs t). dverts (fst x))"
using \<open>max_deg t \<le> 1\<close> dtree_to_list_eq_dverts_ins by fastforce
ultimately show ?thesis using Node.prems(1) t_def(1) by force
qed
qed
lemma dtree_to_list_sub_dverts:
"wf_dverts t \<Longrightarrow> fst ` set (dtree_to_list t) \<subseteq> dverts t - {root t}"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case False
then obtain t e where t_def: "xs = {|(t,e)|}"
using mdeg_1_singleton by fastforce
then have "wf_dverts t" using Node.prems mdeg_ge_child by fastforce
then have "fst ` set (dtree_to_list t) \<subseteq> dverts t - {root t}" using Node.IH t_def(1) by auto
then have "fst ` set (dtree_to_list (Node r xs)) \<subseteq> dverts t"
using t_def(1) dtree.set_sel(1) by auto
then show ?thesis using Node.prems(1) t_def(1) by (simp add: wf_dverts_iff_dverts')
qed (auto)
qed
lemma dtree_to_list_eq_dverts:
"\<lbrakk>wf_dverts t; max_deg t \<le> 1\<rbrakk> \<Longrightarrow> fst ` set (dtree_to_list t) = dverts t - {root t}"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "max_deg (Node r xs) = 0")
case True
then have "xs = {||}" using empty_iff_mdeg_0 by auto
moreover from this have "\<forall>x. xs \<noteq> {|x|}" by blast
ultimately show ?thesis by simp
next
case False
then have "max_deg (Node r xs) = 1" using Node.prems by simp
then obtain t e where t_def: "xs = {|(t,e)|}" "(t,e) \<in> fset xs"
using mdeg_1_singleton by fastforce
then have "max_deg t \<le> 1 \<and> wf_dverts t" using Node.prems mdeg_ge_child by fastforce
then have "fst ` set (dtree_to_list t) = dverts t - {root t}" using Node.IH t_def(2) by auto
then have "fst ` set (dtree_to_list (Node r xs)) = dverts t"
using t_def(1) dtree.set_sel(1) by auto
then show ?thesis using Node.prems(1) t_def(1) by (simp add: wf_dverts_iff_dverts')
qed
qed
lemma dtree_to_list_eq_dverts_single:
"\<lbrakk>max_deg t \<le> 1; sucs t = {|(t1,e1)|}\<rbrakk> \<Longrightarrow> fst ` set (dtree_to_list t) = dverts t1"
by (simp add: dtree_to_list_eq_dverts_sucs)
lemma dtree_to_list_sub_darcs: "snd ` set (dtree_to_list t) \<subseteq> darcs t"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case False
then obtain t e where "xs = {|(t,e)|}"
using mdeg_1_singleton by fastforce
then show ?thesis using Node.IH by fastforce
qed (auto)
qed
lemma dtree_to_list_eq_darcs: "max_deg t \<le> 1 \<Longrightarrow> snd ` set (dtree_to_list t) = darcs t"
proof(induction t)
case (Node r xs)
show ?case
proof(cases "max_deg (Node r xs) = 0")
case True
then have "xs = {||}" using empty_iff_mdeg_0 by auto
moreover from this have "\<forall>x. xs \<noteq> {|x|}" by blast
ultimately show ?thesis by simp
next
case False
then have "max_deg (Node r xs) = 1" using Node.prems by simp
then obtain t e where t_def: "xs = {|(t,e)|}" "(t,e) \<in> fset xs"
using mdeg_1_singleton by fastforce
then have "max_deg t \<le> 1" using Node.prems mdeg_ge_child by fastforce
then have "snd ` set (dtree_to_list t) = darcs t" using Node.IH t_def(2) by auto
then show ?thesis using t_def(1) by simp
qed
qed
lemma dtree_from_list_eq_dverts: "dverts (dtree_from_list r xs) = insert r (fst ` set xs)"
by(induction xs arbitrary: r) force+
lemma dtree_from_list_eq_darcs: "darcs (dtree_from_list r xs) = snd ` set xs"
by(induction xs arbitrary: r) force+
lemma dtree_from_list_root_r[simp]: "root (dtree_from_list r xs) = r"
using dtree.sel(1) dtree_from_list.elims by metis
lemma dtree_from_list_v_eq_r:
"Node r xs = dtree_from_list v ys \<Longrightarrow> r = v"
using dtree.sel(1)[of r xs] by simp
lemma dtree_from_list_fcard0_empty: "fcard (sucs (dtree_from_list r [])) = 0"
by simp
lemma dtree_from_list_fcard0_iff_empty: "fcard (sucs (dtree_from_list r xs)) = 0 \<longleftrightarrow> xs = []"
by(induction xs) auto
lemma dtree_from_list_fcard1_iff_nempty: "fcard (sucs (dtree_from_list r xs)) = 1 \<longleftrightarrow> xs \<noteq> []"
by(induction xs) (auto simp: fcard_single_1 fcard_fempty)
lemma dtree_from_list_fcard_le1: "fcard (sucs (dtree_from_list r xs)) \<le> 1"
by(induction xs) (auto simp: fcard_single_1 fcard_fempty)
lemma dtree_from_empty_deg_0: "max_deg (dtree_from_list r []) = 0"
by simp
lemma dtree_from_list_deg_le_1: "max_deg (dtree_from_list r xs) \<le> 1"
proof(induction xs arbitrary: r)
case Nil
have "max_deg (dtree_from_list r []) = 0" by simp
also have "\<dots> \<le> 1" by blast
finally show ?case by blast
next
case (Cons x xs)
obtain v e where v_def: "x = (v,e)" by force
let ?xs = "{|(dtree_from_list v xs, e)|}"
have "dtree_from_list r (x#xs) = Node r ?xs" by (simp add: v_def)
moreover have "max_deg (dtree_from_list v xs) \<le> 1" using Cons by simp
moreover have "max_deg (Node r ?xs) = max (max_deg (dtree_from_list v xs)) (fcard ?xs)"
using mdeg_singleton by fast
ultimately show ?case by (simp add: fcard_finsert_if max_def)
qed
lemma dtree_from_list_deg_1: "xs \<noteq> [] \<longleftrightarrow> max_deg (dtree_from_list r xs) = 1"
proof (cases xs)
case (Cons x xs)
obtain v e where v_def: "x = (v,e)" by force
let ?xs = "{|(dtree_from_list v xs, e)|}"
have "dtree_from_list r (x#xs) = Node r ?xs" by (simp add: v_def)
moreover have "max_deg (dtree_from_list v xs) \<le> 1" using dtree_from_list_deg_le_1 by fast
moreover have "max_deg (Node r ?xs) = max (max_deg (dtree_from_list v xs)) (fcard ?xs)"
using mdeg_singleton by fast
ultimately show ?thesis using Cons by (simp add: fcard_finsert_if max_def)
qed (metis dtree_from_empty_deg_0 zero_neq_one)
lemma dtree_from_list_singleton: "xs \<noteq> [] \<Longrightarrow> \<exists>t e. dtree_from_list r xs = Node r {|(t,e)|}"
using dtree_from_list.elims[of r xs] by fastforce
lemma dtree_from_to_list_id: "max_deg t \<le> 1 \<Longrightarrow> dtree_from_list (root t) (dtree_to_list t) = t"
proof(induction t)
case (Node r xs)
then show ?case
proof(cases "max_deg (Node r xs) = 0")
case True
then have "xs = {||}" using empty_iff_mdeg_0 by auto
moreover from this have "\<forall>x. xs \<noteq> {|x|}" by blast
ultimately show ?thesis using Node.prems by simp
next
case False
then have "max_deg (Node r xs) = 1" using Node.prems by simp
then obtain t e where t_def: "xs = {|(t,e)|}" "(t,e) \<in> fset xs"
using mdeg_1_singleton by fastforce
then have "max_deg t \<le> 1" using Node.prems mdeg_ge_child by fastforce
then show ?thesis using Node.IH t_def(1) by simp
qed
qed
lemma dtree_to_from_list_id: "dtree_to_list (dtree_from_list r xs) = xs"
proof(induction xs arbitrary: r)
case Nil
then show ?case
using dtree_from_list_deg_1 dtree_from_list_deg_le_1 dtree_from_to_list_id by metis
next
case (Cons x xs)
obtain v e where v_def: "x = (v,e)" by force
then have "dtree_to_list (dtree_from_list r (x#xs)) = (v,e)#dtree_to_list (dtree_from_list v xs)"
by (metis dtree_from_list.elims dtree_to_list.simps(1) dtree.sel(1) dtree_from_list.simps(2))
then show ?case by (simp add: v_def Cons)
qed
lemma dtree_from_list_eq_singleton_hd:
"Node r0 {|(t0,e0)|} = dtree_from_list v1 ys \<Longrightarrow> (\<exists>xs. (root t0, e0) # xs = ys)"
using dtree_to_list.simps(1)[of r0 t0 e0] dtree_to_from_list_id[of v1 ys] by simp
lemma dtree_from_list_eq_singleton:
"Node r0 {|(t0,e0)|} = dtree_from_list v1 ys \<Longrightarrow> r0 = v1 \<and> (\<exists>xs. (root t0, e0) # xs = ys)"
using dtree_from_list_eq_singleton_hd by fastforce
lemma dtree_from_list_uneq_sequence:
"\<lbrakk>is_subtree (Node r0 {|(t0,e0)|}) (dtree_from_list v1 ys);
Node r0 {|(t0,e0)|} \<noteq> dtree_from_list v1 ys\<rbrakk>
\<Longrightarrow> \<exists>e as bs. as @ (r0,e) # (root t0, e0) # bs = ys"
proof(induction v1 ys rule: dtree_from_list.induct)
case (2 r v e xs)
then show ?case
proof(cases "Node r0 {|(t0,e0)|} = dtree_from_list v xs")
case True
then show ?thesis using dtree_from_list_eq_singleton by fast
next
case False
then obtain e1 as bs where "as @ (r0, e1) # (root t0, e0) # bs = xs" using 2 by auto
then have "((v,e)#as) @ (r0, e1) # (root t0, e0) # bs = (v, e) # xs" by simp
then show ?thesis by blast
qed
qed(simp)
lemma dtree_from_list_sequence:
"\<lbrakk>is_subtree (Node r0 {|(t0,e0)|}) (dtree_from_list v1 ys)\<rbrakk>
\<Longrightarrow> \<exists>e as bs. as @ (r0,e) # (root t0, e0) # bs = ((v1,e1)#ys)"
using dtree_from_list_uneq_sequence[of r0 t0 e0] dtree_from_list_eq_singleton append_Cons by fast
lemma dtree_from_list_eq_empty:
"Node r {||} = dtree_from_list v ys \<Longrightarrow> r = v \<and> ys = []"
using dtree_to_from_list_id dtree_from_list_v_eq_r dtree_from_list.simps(1) by metis
lemma dtree_from_list_sucs_cases:
"Node r xs = dtree_from_list v ys \<Longrightarrow> xs = {||} \<or> (\<exists>x. xs = {|x|})"
using dtree.inject dtree_from_list.simps(1) dtree_to_from_list_id dtree_to_list.simps(2) by metis
lemma dtree_from_list_uneq_sequence_xs:
"strict_subtree (Node r0 xs0) (dtree_from_list v1 ys)
\<Longrightarrow> \<exists>e as bs. as @ (r0,e) # bs = ys \<and> Node r0 xs0 = dtree_from_list r0 bs"
proof(induction v1 ys rule: dtree_from_list.induct)
case (2 r v e xs)
then show ?case
proof(cases "Node r0 xs0 = dtree_from_list v xs")
case True
then show ?thesis using dtree_from_list_root_r dtree.sel(1)[of r0 xs0] by fastforce
next
case False
then obtain e1 as bs where 0: "as @ (r0,e1) # bs = xs" "Node r0 xs0 = dtree_from_list r0 bs"
using 2 unfolding strict_subtree_def by auto
then have "((v,e)#as) @ (r0,e1) # bs = (v,e) # xs" by simp
then show ?thesis using 0(2) by blast
qed
qed(simp add: strict_subtree_def)
lemma dtree_from_list_sequence_xs:
"\<lbrakk>is_subtree (Node r xs) (dtree_from_list v1 ys)\<rbrakk>
\<Longrightarrow> \<exists>e as bs. as @ (r,e) # bs = ((v1,e1)#ys) \<and> Node r xs = dtree_from_list r bs"
using dtree_from_list_uneq_sequence_xs[of r xs] dtree_from_list_v_eq_r strict_subtree_def
by (fast intro!: append_Cons)
lemma dtree_from_list_sequence_dverts:
"\<lbrakk>is_subtree (Node r xs) (dtree_from_list v1 ys)\<rbrakk>
\<Longrightarrow> \<exists>e as bs. as @ (r,e) # bs = ((v1,e1)#ys) \<and> dverts (Node r xs) = insert r (fst ` set bs)"
using dtree_from_list_sequence_xs[of r xs v1 ys e1] dtree_from_list_eq_dverts by metis
lemma dtree_from_list_dverts_subset_set:
"set bs \<subseteq> set ds \<Longrightarrow> dverts (dtree_from_list r bs) \<subseteq> dverts (dtree_from_list r ds)"
by (auto simp: dtree_from_list_eq_dverts)
lemma wf_darcs'_iff_wf_list_arcs: "wf_list_arcs xs \<longleftrightarrow> wf_darcs' (dtree_from_list r xs)"
by(induction xs arbitrary: r rule: wf_list_arcs.induct) (auto simp: dtree_from_list_eq_darcs)
lemma wf_darcs_iff_wf_list_arcs: "wf_list_arcs xs \<longleftrightarrow> wf_darcs (dtree_from_list r xs)"
using wf_darcs'_iff_wf_list_arcs wf_darcs_iff_darcs' by fast
lemma wf_dverts_iff_wf_list_verts:
"r \<notin> fst ` set xs \<and> wf_list_verts xs \<longleftrightarrow> wf_dverts (dtree_from_list r xs)"
by (induction xs arbitrary: r rule: wf_list_verts.induct)
(auto simp: dtree_from_list_eq_dverts wf_dverts_iff_dverts')
theorem wf_dtree_iff_wf_list:
"wf_list_arcs xs \<and> r \<notin> fst ` set xs \<and> wf_list_verts xs \<longleftrightarrow> wf_dtree (dtree_from_list r xs)"
using wf_darcs_iff_wf_list_arcs wf_dverts_iff_wf_list_verts unfolding wf_dtree_def by fast
lemma wf_list_arcs_if_wf_darcs: "wf_darcs t \<Longrightarrow> wf_list_arcs (dtree_to_list t)"
proof(induction t)
case (Node r xs)
then show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case True
then show ?thesis using dtree_to_list.simps(2) by simp
next
case False
then obtain t1 e1 where "xs = {|(t1,e1)|}" by auto
then show ?thesis
using Node dtree_to_list_sub_darcs unfolding wf_darcs_iff_darcs' by fastforce
qed
qed
lemma wf_list_verts_if_wf_dverts: "wf_dverts t \<Longrightarrow> wf_list_verts (dtree_to_list t)"
proof(induction t)
case (Node r xs)
then show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case True
then show ?thesis using dtree_to_list.simps(2) by simp
next
case False
then obtain t1 e1 where "xs = {|(t1,e1)|}" by auto
then show ?thesis using Node dtree_to_list_sub_dverts by (fastforce simp: wf_dverts_iff_dverts')
qed
qed
lemma distinct_if_wf_list_arcs: "wf_list_arcs xs \<Longrightarrow> distinct xs"
by (induction xs) force+
lemma distinct_if_wf_list_verts: "wf_list_verts xs \<Longrightarrow> distinct xs"
by (induction xs) force+
lemma wf_list_arcs_alt: "wf_list_arcs xs \<longleftrightarrow> distinct (map snd xs)"
by (induction xs) force+
lemma wf_list_verts_alt: "wf_list_verts xs \<longleftrightarrow> distinct (map fst xs)"
by (induction xs) force+
lemma subtree_from_list_split_eq_if_wfverts:
assumes "wf_list_verts (as@(r,e)#bs)"
and "v \<notin> fst ` set (as@(r,e)#bs)"
and "is_subtree (Node r xs) (dtree_from_list v (as@(r,e)#bs))"
shows "Node r xs = dtree_from_list r bs"
proof -
have 0: "wf_list_verts ((v,e)#as@(r,e)#bs)" using assms(1,2) by simp
obtain as' e' bs' where as'_def:
"as'@(r,e')#bs' = (v,e)#as@(r,e)#bs" "Node r xs = dtree_from_list r bs'"
using assms(3) dtree_from_list_sequence_xs[of r xs] by blast
then have 0: "wf_list_verts (as'@(r,e')#bs')" using assms(1,2) by simp
have r_as': "r \<notin> fst ` set as'" using 0 unfolding wf_list_verts_alt by simp
moreover have r_bs': "r \<notin> fst ` set bs'" using 0 unfolding wf_list_verts_alt by simp
moreover have "(r,e) \<in> set (as'@(r,e')#bs')" using as'_def(1) by simp
ultimately have "(r,e')= (r,e)" by force
then show ?thesis
using r_as' r_bs' as'_def append_Cons_eq_iff[of "(r,e)" as' bs' "(v,e)#as" bs] by force
qed
lemma subtree_from_list_split_eq_if_wfdverts:
"\<lbrakk>wf_dverts (dtree_from_list v (as@(r,e)#bs));
is_subtree (Node r xs) (dtree_from_list v (as@(r,e)#bs))\<rbrakk>
\<Longrightarrow> Node r xs = dtree_from_list r bs"
using subtree_from_list_split_eq_if_wfverts wf_dverts_iff_wf_list_verts by fast
lemma dtree_from_list_dverts_subset_wfdverts:
assumes "set bs \<subseteq> set ds"
and "wf_dverts (dtree_from_list v (as@(r,e1)#bs))"
and "wf_dverts (dtree_from_list v (cs@(r,e2)#ds))"
and "is_subtree (Node r xs) (dtree_from_list v (as@(r,e1)#bs))"
and "is_subtree (Node r ys) (dtree_from_list v (cs@(r,e2)#ds))"
shows "dverts (Node r xs) \<subseteq> dverts (Node r ys)"
using dtree_from_list_dverts_subset_set[OF assms(1)]
subtree_from_list_split_eq_if_wfdverts[OF assms(2,4)]
subtree_from_list_split_eq_if_wfdverts[OF assms(3,5)]
by simp
lemma dtree_from_list_dverts_subset_wfdverts':
assumes "wf_dverts (dtree_from_list v as)"
and "wf_dverts (dtree_from_list v cs)"
and "is_subtree (Node r xs) (dtree_from_list v as)"
and "is_subtree (Node r ys) (dtree_from_list v cs)"
and "\<exists>as' e1 bs cs' e2 ds. as'@(r,e1)#bs = as \<and> cs'@(r,e2)#ds = cs \<and> set bs \<subseteq> set ds"
shows "dverts (Node r xs) \<subseteq> dverts (Node r ys)"
using dtree_from_list_dverts_subset_wfdverts assms by metis
lemma dtree_to_list_sequence_subtree:
"\<lbrakk>max_deg t \<le> 1; strict_subtree (Node r xs) t\<rbrakk>
\<Longrightarrow> \<exists>as e bs. dtree_to_list t = as@(r,e)#bs \<and> Node r xs = dtree_from_list r bs"
by (metis dtree_from_list_uneq_sequence_xs dtree_from_to_list_id)
lemma dtree_to_list_sequence_subtree':
"\<lbrakk>max_deg t \<le> 1; strict_subtree (Node r xs) t\<rbrakk>
\<Longrightarrow> \<exists>as e bs. dtree_to_list t = as@(r,e)#bs \<and> dtree_to_list (Node r xs) = bs"
using dtree_to_from_list_id[of r] dtree_to_list_sequence_subtree[of t r xs] by fastforce
lemma dtree_to_list_subtree_dverts_eq_fsts:
"\<lbrakk>max_deg t \<le> 1; strict_subtree (Node r xs) t\<rbrakk>
\<Longrightarrow> \<exists>as e bs. dtree_to_list t = as@(r,e)#bs \<and> insert r (fst ` set bs) = dverts (Node r xs)"
by (metis dtree_from_list_eq_dverts dtree_to_list_sequence_subtree)
lemma dtree_to_list_subtree_dverts_eq_fsts':
"\<lbrakk>max_deg t \<le> 1; strict_subtree (Node r xs) t\<rbrakk>
\<Longrightarrow> \<exists>as e bs. dtree_to_list t = as@(r,e)#bs \<and> (fst ` set ((r,e)#bs)) = dverts (Node r xs)"
using dtree_to_list_subtree_dverts_eq_fsts by fastforce
lemma dtree_to_list_split_subtree:
assumes "as@(r,e)#bs = dtree_to_list t"
shows "\<exists>xs. strict_subtree (Node r xs) t \<and> dtree_to_list (Node r xs) = bs"
using assms proof(induction t arbitrary: as rule: dtree_to_list.induct)
case (1 r1 t1 e1)
show ?case
proof(cases as)
case Nil
then have "dtree_to_list (Node r (sucs t1)) = bs" using "1.prems" by auto
moreover have "is_subtree (Node r (sucs t1)) (Node r1 {|(t1, e1)|})"
using subtree_if_child[of t1 "{|(t1, e1)|}"] "1.prems" Nil by simp
moreover have "Node r1 {|(t1, e1)|} \<noteq> (Node r (sucs t1))" by (blast intro!: singleton_uneq')
ultimately show ?thesis unfolding strict_subtree_def by blast
next
case (Cons a as')
then show ?thesis using 1 unfolding strict_subtree_def by fastforce
qed
qed(simp)
lemma dtree_to_list_split_subtree_dverts_eq_fsts:
assumes "max_deg t \<le> 1" and "as@(r,e)#bs = dtree_to_list t"
shows "\<exists>xs. strict_subtree (Node r xs) t \<and> dverts (Node r xs) = insert r (fst`set bs)"
proof -
obtain xs where xs_def:
"is_subtree (Node r xs) t" "Node r xs \<noteq> t" "dtree_to_list (Node r xs) = bs"
using dtree_to_list_split_subtree[OF assms(2)] unfolding strict_subtree_def by blast
have "max_deg (Node r xs) \<le> 1" using mdeg_ge_sub[OF xs_def(1)] assms(1) by simp
then show ?thesis
using dtree_to_list_eq_dverts_ins[of "Node r xs"] xs_def strict_subtree_def by auto
qed
lemma dtree_to_list_split_subtree_dverts_eq_fsts':
assumes "max_deg t \<le> 1" and "as@(r,e)#bs = dtree_to_list t"
shows "\<exists>xs. strict_subtree (Node r xs) t \<and> dverts (Node r xs) = (fst ` set ((r,e)#bs))"
using dtree_to_list_split_subtree_dverts_eq_fsts[OF assms] by simp
lemma dtree_from_list_dverts_subset_wfdverts1:
assumes "dverts t1 \<subseteq> fst ` set ((r,e2)#bs)"
and "wf_dverts (dtree_from_list v (as@(r,e2)#bs))"
and "is_subtree (Node r ys) (dtree_from_list v (as@(r,e2)#bs))"
shows "dverts t1 \<subseteq> dverts (Node r ys)"
using subtree_from_list_split_eq_if_wfdverts[OF assms(2,3)] assms(1) dtree_from_list_eq_dverts
by fastforce
lemma dtree_from_list_dverts_subset_wfdverts1':
assumes "wf_dverts (dtree_from_list v cs)"
and "is_subtree (Node r ys) (dtree_from_list v cs)"
and "\<exists>as e bs. as@(r,e)#bs = cs \<and> dverts t1 \<subseteq> fst ` set ((r,e)#bs)"
shows "dverts t1 \<subseteq> dverts (Node r ys)"
using dtree_from_list_dverts_subset_wfdverts1 assms by fast
lemma dtree_from_list_1_leaf: "num_leaves (dtree_from_list r xs) = 1"
using num_leaves_1_if_mdeg_1 dtree_from_list_deg_le_1 by fast
subsection \<open>Inserting in Dtrees\<close>
abbreviation insert_before ::
"'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> (('a,'b) dtree \<times> 'b) fset \<Rightarrow> (('a,'b) dtree \<times> 'b) fset" where
"insert_before v e y xs \<equiv> ffold (\<lambda>(t1,e1).
finsert (if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))) {||} xs"
fun insert_between :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a,'b) dtree \<Rightarrow> ('a,'b) dtree" where
"insert_between v e x y (Node r xs) = (if x=r \<and> (\<exists>t. t \<in> fst ` fset xs \<and> root t = y)
then Node r (insert_before v e y xs)
else if x=r then Node r (finsert (Node v {||},e) xs)
else Node r ((\<lambda>(t,e1). (insert_between v e x y t,e1)) |`| xs))"
lemma insert_between_id_if_notin: "x \<notin> dverts t \<Longrightarrow> insert_between v e x y t = t"
proof(induction t)
case (Node r xs)
have "\<forall>(t,e) \<in> fset xs. x \<notin> dverts t" using Node.prems by force
then have "\<forall>(t,e1) \<in> fset xs. (\<lambda>(t,e1). (insert_between v e x y t,e1)) (t,e1) = (t,e1)"
using Node.IH by auto
then have "((\<lambda>(t,e1). (insert_between v e x y t,e1)) |`| xs) = xs"
by (smt (verit, ccfv_threshold) fset.map_cong0 case_prodE fimage_ident)
then show ?case using Node.prems by simp
qed
context wf_dtree
begin
lemma insert_before_commute_aux:
assumes "f = (\<lambda>(t1,e1). finsert (if root t1 = y1 then (Node v {|(t1,e1)|},e) else (t1,e1)))"
shows "(f y \<circ> f x) z = (f x \<circ> f y) z"
proof -
obtain t1 e1 where y_def: "y = (t1, e1)" by fastforce
obtain t2 e2 where "x = (t2, e2)" by fastforce
then show ?thesis using assms y_def by auto
qed
lemma insert_before_commute:
"comp_fun_commute (\<lambda>(t1,e1). finsert (if root t1 = y1 then (Node v {|(t1,e1)|},e) else (t1,e1)))"
using comp_fun_commute_def insert_before_commute_aux by fastforce
interpretation Comm:
comp_fun_commute "\<lambda>(t1,e1). finsert (if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
by (rule insert_before_commute)
lemma root_not_new_in_orig:
"\<lbrakk>(t1,e1) \<in> fset (insert_before v e y xs); root t1 \<noteq> v\<rbrakk> \<Longrightarrow> (t1,e1) \<in> fset xs"
proof(induction xs)
case empty
then show ?case by simp
next
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
show ?case
proof(cases "(t1,e1) \<in> fset (insert_before v e y xs)")
case True
then show ?thesis using insert.IH insert.prems(2) by simp
next
case False
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "?f x = (t1,e1)" using False insert.prems(1) by force
then have "x = (t1,e1)"
by (smt (z3) insert.prems(2) dtree.sel(1) old.prod.exhaust prod.inject case_prod_conv)
then show ?thesis by simp
qed
qed
lemma root_not_y_in_new:
"\<lbrakk>(t1,e1) \<in> fset xs; root t1 \<noteq> y\<rbrakk> \<Longrightarrow> (t1,e1) \<in> fset (insert_before v e y xs)"
proof(induction xs)
case empty
then show ?case by simp
next
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
show ?case
proof(cases "(t1,e1) = x")
case True
then show ?thesis using insert by auto
next
case False
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then show ?thesis using insert.IH insert.prems by force
qed
qed
lemma root_noty_if_in_insert_before:
"\<lbrakk>(t1,e1) \<in> fset (insert_before v e y xs); v\<noteq>y\<rbrakk> \<Longrightarrow> root t1 \<noteq> y"
proof(induction xs)
case empty
then show ?case by simp
next
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
show ?case
proof(cases "(t1,e1) \<in> fset (insert_before v e y xs)")
case True
then show ?thesis using insert.IH insert.prems(2) by fast
next
case False
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have 0: "?f x = (t1,e1)" using insert.prems False by simp
then show ?thesis
proof(cases "root t1 = v")
case True
then show ?thesis using insert.prems(2) by simp
next
case False
then show ?thesis by (smt (z3) dtree.sel(1) old.prod.exhaust prod.inject 0 case_prod_conv)
qed
qed
qed
lemma in_insert_before_child_in_orig:
"\<lbrakk>(t1,e1) \<in> fset (insert_before v e y xs); (t1,e1) \<notin> fset xs\<rbrakk>
\<Longrightarrow> \<exists>(t2,e2) \<in> fset xs. (Node v {|(t2,e2)|}) = t1 \<and> root t2 = y \<and> e1=e"
proof(induction xs)
case empty
then show ?case by simp
next
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
show ?case
proof(cases "(t1,e1) \<in> fset (insert_before v e y xs)")
case True
then show ?thesis using insert.IH insert.prems(2) by simp
next
case False
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then show ?thesis
- by (smt (z3) False Pair_inject old.prod.case case_prodI2 finsert_iff insert.prems notin_fset)
+ by (smt (z3) False Pair_inject old.prod.case case_prodI2 finsert_iff insert.prems)
qed
qed
lemma insert_before_not_y_id:
"\<not>(\<exists>t. t \<in> fst ` fset xs \<and> root t = y) \<Longrightarrow> insert_before v e y xs = xs"
proof(induction xs)
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "insert_before v e y (finsert x xs) = finsert x (insert_before v e y xs)"
- using notin_fset insert.prems
+ using insert.prems
by (smt (z3) old.prod.exhaust case_prod_conv finsertCI fst_conv image_eqI)
moreover have "\<not>(\<exists>t. t \<in> fst ` fset xs \<and> root t = y)" using insert.prems by auto
ultimately show ?case using insert.IH by blast
qed (simp)
lemma insert_before_alt:
"insert_before v e y xs
= (\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1)) |`| xs"
by(induction xs) (auto simp: Product_Type.prod.case_distrib)
lemma dverts_insert_before_aux:
"\<exists>t. t \<in> fst ` fset xs \<and> root t = y
\<Longrightarrow> (\<Union>x\<in>fset (insert_before v e y xs). \<Union> (dverts ` Basic_BNFs.fsts x))
= insert v (\<Union>x\<in>fset xs. \<Union> (dverts ` Basic_BNFs.fsts x))"
proof(induction xs)
case empty
then show ?case by simp
next
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
obtain t1 e1 where t1_def: "x = (t1,e1)" by fastforce
then show ?case
proof(cases "root t1 = y")
case True
then have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "insert_before v e y (finsert x xs)
= finsert (Node v {|(t1,e1)|},e) (insert_before v e y xs)"
using t1_def True by simp
then have 0: "(\<Union>x\<in>fset (insert_before v e y (finsert x xs)). \<Union> (dverts ` Basic_BNFs.fsts x))
= insert v (dverts t1) \<union> (\<Union>x\<in>fset (insert_before v e y xs). \<Union> (dverts ` Basic_BNFs.fsts x))"
using t1_def by simp
have 1: "dverts (Node v {|(t1,e1)|}) = insert v (dverts t1)" by simp
show ?thesis
proof(cases "\<exists>t. t \<in> fst ` fset xs \<and> root t = y")
case True
then show ?thesis using t1_def 0 insert.IH by simp
next
case False
then show ?thesis using t1_def 0 insert_before_not_y_id by force
qed
next
case False
then have 0: "\<exists>t. t \<in> fst ` fset xs \<and> root t = y" using insert.prems t1_def by force
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "insert_before v e y (finsert x xs) = finsert x (insert_before v e y xs)"
by (simp add: False t1_def)
then show ?thesis using insert.IH insert.prems 0 by simp
qed
qed
lemma insert_between_add_v_if_x_in:
"x \<in> dverts t \<Longrightarrow> dverts (insert_between v e x y t) = insert v (dverts t)"
using wf_verts proof(induction t)
case (Node r xs)
show ?case
proof(cases "x=r")
case False
then obtain t e1 where t_def: "(t,e1) \<in> fset xs" "x \<in> dverts t" using Node.prems(1) by auto
then have "\<forall>(t2,e2) \<in> fset xs. (t,e1) \<noteq> (t2,e2) \<longrightarrow> x \<notin> dverts t2"
using Node.prems(2) by (fastforce simp: wf_dverts_iff_dverts')
then have "\<forall>(t2,e2) \<in> fset xs. (t,e1) = (t2,e2) \<or> (insert_between v e x y t2) = t2"
using insert_between_id_if_notin by fast
moreover have "(insert_between v e x y t,e1)
\<in> fset ((\<lambda>(t,e1). (insert_between v e x y t,e1)) |`| xs)" using t_def(1) by force
moreover have "dverts (insert_between v e x y t) = insert v (dverts t)"
using Node.IH Node.prems(2) t_def by auto
ultimately show ?thesis using False by force
qed (auto simp: dverts_insert_before_aux)
qed
lemma insert_before_only1_new:
assumes "\<forall>(x,e1) \<in> fset xs. \<forall>(y,e2) \<in> fset xs. (dverts x \<inter> dverts y = {} \<or> (x,e1)=(y,e2))"
and "(t1,e1) \<noteq> (t2,e2)"
and "(t1,e1) \<in> fset (insert_before v e y xs)"
and "(t2,e2) \<in> fset (insert_before v e y xs)"
shows "(t1,e1) \<in> fset xs \<or> (t2,e2) \<in> fset xs"
proof (rule ccontr)
assume "\<not>((t1,e1) \<in> fset xs \<or> (t2,e2) \<in> fset xs)"
then have asm: "(t1,e1) \<notin> fset xs" "(t2,e2) \<notin> fset xs" by auto
obtain t3 e3 where t3_def: "(t3, e3)\<in>fset xs" "Node v {|(t3, e3)|} = t1" "root t3 = y" "e1=e"
using in_insert_before_child_in_orig assms(3) asm(1) by fast
obtain t4 e4 where t4_def: "(t4, e4)\<in>fset xs" "Node v {|(t4, e4)|} = t2" "root t4 = y" "e2=e"
using in_insert_before_child_in_orig assms(4) asm(2) by fast
then have "dverts t3 \<inter> dverts t4 \<noteq> {}" using t3_def(3) dtree.set_sel(1) by force
then have "(t3,e3) = (t4,e4)" using assms(1) t3_def(1) t4_def(1) by fast
then show False using assms(2) t3_def(2,4) t4_def(2,4) by fast
qed
lemma disjoint_dverts_aux1:
assumes "\<forall>(t1,e1) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. (dverts t1 \<inter> dverts t2 = {} \<or> (t1,e1)=(t2,e2))"
and "v \<notin> dverts (Node r xs)"
and "(t1,e1) \<in> fset (insert_before v e y xs)"
and "(t2,e2) \<in> fset (insert_before v e y xs)"
and "(t1,e1) \<noteq> (t2,e2)"
shows "dverts t1 \<inter> dverts t2 = {}"
proof -
consider "(t1,e1) \<in> fset xs" "(t2,e2) \<in> fset xs"
| "(t1,e1) \<notin> fset xs" "(t2,e2) \<in> fset xs"
| "(t1,e1) \<in> fset xs" "(t2,e2) \<notin> fset xs"
using insert_before_only1_new assms(1,3-5) by fast
then show ?thesis
proof(cases)
case 1
then show ?thesis using assms(1,5) by fast
next
case 2
obtain t3 e3 where t3_def: "(t3, e3)\<in>fset xs" "Node v {|(t3, e3)|} = t1" "root t3 = y" "e1=e"
using in_insert_before_child_in_orig assms(3) 2 by fast
then have "y\<noteq>v" using assms(2) dtree.set_sel(1) by force
then have "(t3,e3) \<noteq> (t2,e2)" using assms(4) t3_def(3) root_noty_if_in_insert_before by fast
then have "dverts t3 \<inter> dverts t2 = {}" using assms(1) 2(2) t3_def(1) by fast
then show ?thesis using assms(1,2) t3_def(1,2) 2(2) by force
next
case 3
obtain t3 e3 where t3_def: "(t3, e3)\<in>fset xs" "Node v {|(t3, e3)|} = t2" "root t3 = y" "e2=e"
using in_insert_before_child_in_orig assms(4) 3 by fast
then have "y\<noteq>v" using assms(2) dtree.set_sel(1) by force
then have "(t3,e3) \<noteq> (t1,e1)" using assms(3) t3_def(3) root_noty_if_in_insert_before by fast
then have "dverts t3 \<inter> dverts t1 = {}" using assms(1) 3(1) t3_def(1) by fast
then show ?thesis using assms(2) t3_def(2) 3(1) by force
qed
qed
lemma disjoint_dverts_aux1':
assumes "wf_dverts (Node r xs)" and "v \<notin> dverts (Node r xs)"
shows "\<forall>(x,e1) \<in> fset (insert_before v e y xs). \<forall>(y,e2) \<in> fset (insert_before v e y xs).
dverts x \<inter> dverts y = {} \<or> (x,e1) = (y,e2)"
using assms disjoint_dverts_aux1 disjoint_dverts_if_wf unfolding wf_dverts_iff_dverts' by fast
lemma insert_before_wf_dverts:
"\<lbrakk>\<forall>(t,e1) \<in> fset xs. wf_dverts t; v \<notin> dverts(Node r xs); (t1,e1) \<in> fset (insert_before v e y xs)\<rbrakk>
\<Longrightarrow> wf_dverts t1"
proof(induction xs)
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
show ?case
proof(cases "(t1,e1) \<in> fset (insert_before v e y xs)")
case in_xs: True
then show ?thesis
proof(cases "?f x = (t1,e1)")
case True
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "insert_before v e y (finsert x xs) = insert_before v e y xs"
- using True in_xs notin_fset by fastforce
+ using True in_xs by fastforce
then show ?thesis using insert.IH insert.prems by simp
next
case False
then show ?thesis using in_xs insert.IH insert.prems(1,2) by auto
qed
next
case False
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "?f x = (t1,e1)" using False insert.prems(3) by fastforce
then show ?thesis
proof(cases "root t1 = v")
case True
then have "(t1,e1) \<notin> fset (finsert x xs)" using insert.prems(2) dtree.set_sel(1) by force
then obtain t2 e2 where
t2_def: "(t2, e2)\<in>fset (finsert x xs)" "Node v {|(t2, e2)|} = t1" "root t2 = y" "e1=e"
using in_insert_before_child_in_orig[of t1] insert.prems(3) by blast
then show ?thesis using insert.prems(1,2) by (fastforce simp: wf_dverts_iff_dverts')
next
case False
then have "(t1,e1) = x"
- using insert.prems(1) notin_fset dtree.sel(1) \<open>?f x = (t1,e1)\<close>
+ using insert.prems(1) dtree.sel(1) \<open>?f x = (t1,e1)\<close>
by (smt (verit, ccfv_SIG) Pair_inject old.prod.case case_prodE finsertI1)
then show ?thesis using insert.prems(1) by auto
qed
qed
qed (simp)
lemma insert_before_root_nin_verts:
"\<lbrakk>\<forall>(t,e1)\<in>fset xs. r \<notin> dverts t; v \<notin> dverts (Node r xs); (t1,e1) \<in> fset (insert_before v e y xs)\<rbrakk>
\<Longrightarrow> r \<notin> dverts t1"
proof(induction xs)
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
show ?case
proof(cases "(t1,e1) \<in> fset (insert_before v e y xs)")
case in_xs: True
then show ?thesis
proof(cases "?f x = (t1,e1)")
case True
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "insert_before v e y (finsert x xs) = insert_before v e y xs"
- using True in_xs notin_fset by fastforce
+ using True in_xs by fastforce
then show ?thesis using insert.IH insert.prems by simp
next
case False
then show ?thesis using in_xs insert.IH insert.prems(1,2) by auto
qed
next
case False
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "?f x = (t1,e1)" using False insert.prems(3) by fastforce
then show ?thesis
proof(cases "root t1 = v")
case True
then have "(t1,e1) \<notin> fset (finsert x xs)" using insert.prems(2) dtree.set_sel(1) by force
then obtain t2 e2 where
t2_def: "(t2, e2)\<in>fset (finsert x xs)" "Node v {|(t2, e2)|} = t1" "root t2 = y" "e1=e"
using in_insert_before_child_in_orig[of t1] insert.prems(3) by blast
then show ?thesis using insert.prems(1,2) by fastforce
next
case False
then have "(t1,e1) = x"
- using insert.prems(1) notin_fset dtree.sel(1) \<open>?f x = (t1,e1)\<close>
+ using insert.prems(1) dtree.sel(1) \<open>?f x = (t1,e1)\<close>
by (smt (verit, ccfv_SIG) Pair_inject old.prod.case case_prodE finsertI1)
then show ?thesis using insert.prems(1) by auto
qed
qed
qed (simp)
lemma disjoint_dverts_aux2:
assumes "wf_dverts (Node r xs)" and "v \<notin> dverts (Node r xs)"
shows "\<forall>(x,e1) \<in> fset (finsert (Node v {||},e) xs). \<forall>(y,e2) \<in> fset (finsert (Node v {||},e) xs).
dverts x \<inter> dverts y = {} \<or> (x,e1) = (y,e2)"
using assms by (fastforce simp: wf_dverts_iff_dverts')
lemma disjoint_dverts_aux3:
assumes "(t2,e2) \<in> (\<lambda>(t1,e1). (insert_between v e x y t1, e1)) ` fset xs"
and "(t3,e3) \<in> (\<lambda>(t1,e1). (insert_between v e x y t1, e1)) ` fset xs"
and "(t2,e2)\<noteq>(t3,e3)"
and "(t,e1) \<in> fset xs"
and "x \<in> dverts t"
and "wf_dverts (Node r xs)"
and "v \<notin> dverts (Node r xs)"
shows "dverts t2 \<inter> dverts t3 = {}"
proof -
have "\<forall>(t2,e2) \<in> fset xs. (t,e1)=(t2,e2) \<or> x \<notin> dverts t2"
using assms(4-6) by (fastforce simp: wf_dverts_iff_dverts')
then have nt1_id: "\<forall>(t2,e2) \<in> fset xs. (t,e1) = (t2,e2) \<or> insert_between v e x y t2 = t2"
using insert_between_id_if_notin by fastforce
have dverts_t1: "dverts (insert_between v e x y t) = insert v (dverts t)"
using assms(5-6) by (simp add: insert_between_add_v_if_x_in)
have t1_disj: "\<forall>(t2,e2) \<in> fset xs. (t,e1) = (t2,e2) \<or> dverts t2 \<inter> insert v (dverts t) = {}"
using assms(4-7) by (fastforce simp: wf_dverts_iff_dverts')
consider "(t2,e2) = (insert_between v e x y t,e1)"
| "(t3,e3) = (insert_between v e x y t,e1)"
| "(t2,e2) \<noteq> (insert_between v e x y t,e1)" "(t3,e3) \<noteq> (insert_between v e x y t,e1)"
by fast
then show ?thesis
proof(cases)
case 1
then have "(t3,e3) \<in> fset xs" using assms(2,3) nt1_id by fastforce
moreover have "(t3,e3) \<noteq> (t,e1)" using assms(2,3) 1 nt1_id by fastforce
ultimately show ?thesis using 1 t1_disj dverts_t1 by fastforce
next
case 2
then have "(t2,e2) \<in> fset xs" using assms(1,3) nt1_id by fastforce
moreover have "(t2,e2) \<noteq> (t,e1)" using assms(1,3) 2 nt1_id by auto
ultimately show ?thesis using 2 t1_disj dverts_t1 by fastforce
next
case 3
then have "(t2,e2) \<in> fset xs" using assms(1) nt1_id by fastforce
moreover have "(t3,e3) \<in> fset xs" using assms(2) 3(2) nt1_id by auto
ultimately show ?thesis using assms(3,6) by (fastforce simp: wf_dverts_iff_dverts')
qed
qed
lemma insert_between_wf_dverts: "v \<notin> dverts t \<Longrightarrow> wf_dverts (insert_between v e x y t)"
using wf_dtree_axioms proof(induction t)
case (Node r xs)
then interpret wf_dtree "Node r xs" by blast
consider "x=r" "\<exists>t. t \<in> fst ` fset xs \<and> root t = y"
| "x=r" "\<not>(\<exists>t. t \<in> fst ` fset xs \<and> root t = y)" | "x\<noteq>r" by fast
then show ?case
proof(cases)
case 1
then have "insert_between v e x y (Node r xs) = Node r (insert_before v e y xs)" by simp
moreover have "\<forall>(x,e1) \<in> fset (insert_before v e y xs). r \<notin> dverts x"
using insert_before_root_nin_verts wf_verts Node.prems(1)
by (fastforce simp: wf_dverts_iff_dverts')
moreover have "\<forall>(x,e1) \<in> fset (insert_before v e y xs). wf_dverts x"
using insert_before_wf_dverts Node.prems(1) wf_verts by fastforce
moreover have "\<forall>(x, e1)\<in>fset (insert_before v e y xs).
\<forall>(y, e2)\<in>fset (insert_before v e y xs). dverts x \<inter> dverts y = {} \<or> (x, e1) = (y, e2)"
using disjoint_dverts_aux1' Node.prems(1) wf_verts unfolding wf_dverts_iff_dverts' by fast
ultimately show ?thesis by (fastforce simp: wf_dverts_iff_dverts')
next
case 2
then have "insert_between v e x y (Node r xs) = Node r (finsert (Node v {||},e) xs)" by simp
then show ?thesis
using disjoint_dverts_aux2[of r xs v] Node.prems(1) wf_verts
by (fastforce simp: wf_dverts_iff_dverts')
next
case 3
let ?f = "\<lambda>(t1,e1). (insert_between v e x y t1, e1)"
show ?thesis
proof(cases "\<exists>(t1,e1) \<in> fset xs. x \<in> dverts t1")
case True
then obtain t1 e1 where t1_def: "(t1,e1) \<in> fset xs" " x \<in> dverts t1" by blast
then interpret T: wf_dtree t1 using wf_dtree_rec by blast
have "\<forall>(t2,e2) \<in> ?f ` fset xs. \<forall>(t3,e3) \<in> ?f ` fset xs.
(t2,e2) = (t3,e3) \<or> dverts t2 \<inter> dverts t3 = {}"
using T.disjoint_dverts_aux3 Node.prems(1) t1_def wf_verts by blast
moreover have "\<And>t2 e2. (t2,e2) \<in> ?f ` fset xs \<longrightarrow> r \<notin> dverts t2 \<and> wf_dverts t2"
proof
fix t2 e2
assume asm: "(t2,e2) \<in> ?f ` fset xs"
then show "r \<notin> dverts t2 \<and> wf_dverts t2"
proof(cases "(t2,e2) = (insert_between v e x y t1,e1)")
case True
then have "wf_dverts (insert_between v e x y t1)"
using Node.IH Node.prems(1) T.wf_dtree_axioms t1_def(1) by auto
then show ?thesis
using Node.prems(1) wf_verts True T.insert_between_add_v_if_x_in t1_def
by (auto simp: wf_dverts_iff_dverts')
next
case False
have "\<forall>(t2,e2) \<in> fset xs. (t1,e1)=(t2,e2) \<or> x \<notin> dverts t2"
using wf_verts t1_def by (fastforce simp: wf_dverts_iff_dverts')
then have "\<forall>(t2,e2) \<in> fset xs. (t1,e1) = (t2,e2) \<or> insert_between v e x y t2 = t2"
using insert_between_id_if_notin by fastforce
then show ?thesis using wf_verts asm False by (fastforce simp: wf_dverts_iff_dverts')
qed
qed
ultimately show ?thesis using 3 by (fastforce simp: wf_dverts_iff_dverts')
next
case False
then show ?thesis
using wf_verts 3 insert_between_id_if_notin fst_conv
by (smt (verit, ccfv_threshold) fsts.cases dtree.inject dtree.set_cases(1) case_prodI2)
qed
qed
qed
lemma darcs_insert_before_aux:
"\<exists>t. t \<in> fst ` fset xs \<and> root t = y
\<Longrightarrow> (\<Union>x\<in>fset (insert_before v e y xs). \<Union> (darcs ` Basic_BNFs.fsts x) \<union> Basic_BNFs.snds x)
= insert e (\<Union>x\<in>fset xs. \<Union> (darcs ` Basic_BNFs.fsts x) \<union> Basic_BNFs.snds x)"
proof(induction xs)
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
let ?xs = "insert_before v e y (finsert x xs)"
obtain t1 e1 where t1_def: "x = (t1,e1)" by fastforce
then show ?case
proof(cases "root t1 = y")
case True
then have "?xs = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "?xs = finsert (Node v {|(t1,e1)|},e) (insert_before v e y xs)"
using t1_def True by simp
then have 0: "(\<Union>x\<in>fset ?xs. \<Union> (darcs ` Basic_BNFs.fsts x) \<union> Basic_BNFs.snds x)
= (\<Union> (darcs ` {Node v {|(t1,e1)|}}) \<union> {e})
\<union> (\<Union>x\<in>fset (insert_before v e y xs). \<Union> (darcs ` Basic_BNFs.fsts x) \<union> Basic_BNFs.snds x)"
using t1_def by simp
have 1: "dverts (Node v {|(t1,e1)|}) = insert v (dverts t1)" by simp
show ?thesis
proof(cases "\<exists>t. t \<in> fst ` fset xs \<and> root t = y")
case True
then show ?thesis using t1_def 0 insert.IH by simp
next
case False
then show ?thesis using t1_def 0 insert_before_not_y_id by force
qed
next
case False
then have 0: "\<exists>t. t \<in> fst ` fset xs \<and> root t = y" using insert.prems t1_def by force
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "insert_before v e y (finsert x xs) = finsert x (insert_before v e y xs)"
by (simp add: False t1_def)
then show ?thesis using insert.IH insert.prems 0 by simp
qed
qed (simp)
lemma insert_between_add_e_if_x_in:
"x \<in> dverts t \<Longrightarrow> darcs (insert_between v e x y t) = insert e (darcs t)"
using wf_verts proof(induction t)
case (Node r xs)
show ?case
proof(cases "x=r")
case False
then obtain t e1 where t_def: "(t,e1) \<in> fset xs" "x \<in> dverts t" using Node.prems(1) by auto
then have "\<forall>(t2,e2) \<in> fset xs. (t,e1) \<noteq> (t2,e2) \<longrightarrow> x \<notin> dverts t2"
using Node.prems(2) by (fastforce simp: wf_dverts_iff_dverts')
then have "\<forall>(t2,e2) \<in> fset xs. (t,e1) = (t2,e2) \<or> (insert_between v e x y t2) = t2"
using insert_between_id_if_notin by fast
moreover have "(insert_between v e x y t,e1)
\<in> fset ((\<lambda>(t,e1). (insert_between v e x y t,e1)) |`| xs)" using t_def(1) by force
moreover have "darcs (insert_between v e x y t) = insert e (darcs t)"
using Node.IH Node.prems(2) t_def by auto
ultimately show ?thesis using False by force
qed (auto simp: darcs_insert_before_aux)
qed
lemma disjoint_darcs_aux1_aux1:
assumes "disjoint_darcs xs"
and "wf_dverts (Node r xs)"
and "v \<notin> dverts (Node r xs)"
and "e \<notin> darcs (Node r xs)"
and "(t1,e1) \<in> fset (insert_before v e y xs)"
and "(t2,e2) \<in> fset (insert_before v e y xs)"
and "(t1,e1) \<noteq> (t2,e2)"
shows "(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
proof -
consider "(t1,e1) \<in> fset xs" "(t2,e2) \<in> fset xs"
| "(t1,e1) \<notin> fset xs" "(t2,e2) \<in> fset xs"
| "(t1,e1) \<in> fset xs" "(t2,e2) \<notin> fset xs"
using insert_before_only1_new assms(2,5-7) by (fastforce simp: wf_dverts_iff_dverts')
then show ?thesis
proof(cases)
case 1
then show ?thesis using assms(1,7) by fast
next
case 2
obtain t3 e3 where t3_def: "(t3, e3)\<in>fset xs" "Node v {|(t3, e3)|} = t1" "root t3 = y" "e1=e"
using in_insert_before_child_in_orig assms(5) 2 by fast
then have "v\<noteq>y" using assms(3) dtree.set_sel(1) by force
then have "(t3,e3) \<noteq> (t2,e2)" using assms(6) t3_def(3) root_noty_if_in_insert_before by fast
then have "(darcs t3 \<union> {e3}) \<inter> (darcs t2 \<union> {e2}) = {}" using assms(1) 2(2) t3_def(1) by fast
then show ?thesis using assms(4) t3_def(4) 2(2) t3_def(2) by force
next
case 3
obtain t3 e3 where t3_def: "(t3, e3)\<in>fset xs" "Node v {|(t3, e3)|} = t2" "root t3 = y" "e2=e"
using in_insert_before_child_in_orig assms(6) 3 by fast
then have "v\<noteq>y" using assms(3) dtree.set_sel(1) by force
then have "(t3,e3) \<noteq> (t1,e1)" using assms(5) t3_def(3) root_noty_if_in_insert_before by fast
then have "(darcs t3 \<union> {e3}) \<inter> (darcs t1 \<union> {e1}) = {}" using assms(1) 3(1) t3_def(1) by fast
then show ?thesis using assms(4) t3_def(4) 3(1) t3_def(2) by force
qed
qed
lemma disjoint_darcs_aux1_aux2:
assumes "disjoint_darcs xs"
and "e \<notin> darcs (Node r xs)"
and "(t1,e1) \<in> fset (insert_before v e y xs)"
shows "e1 \<notin> darcs t1"
proof(cases "(t1,e1) \<in> fset xs")
case True
then show ?thesis using assms(1) by fast
next
case False
then obtain t3 e3 where "(t3, e3)\<in>fset xs" "Node v {|(t3, e3)|} = t1" "e1=e"
using in_insert_before_child_in_orig assms(3) by fast
then show ?thesis using assms(2) by auto
qed
lemma disjoint_darcs_aux1:
assumes "wf_dverts (Node r xs)" and "v \<notin> dverts (Node r xs)"
and "wf_darcs (Node r xs)" and "e \<notin> darcs (Node r xs)"
shows "disjoint_darcs (insert_before v e y xs)" (is "disjoint_darcs ?xs")
proof -
have 0: "disjoint_darcs xs" using assms(3) disjoint_darcs_if_wf_xs by simp
then have "\<forall>(t1,e1) \<in> fset ?xs. e1 \<notin> darcs t1"
using disjoint_darcs_aux1_aux2[of xs] assms(4) by fast
moreover have "\<forall>(t1,e1) \<in> fset ?xs. \<forall>(t2,e2) \<in> fset ?xs.
(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {} \<or> (t1,e1) = (t2,e2)"
using disjoint_darcs_aux1_aux1[of xs] assms(1,2,4) 0 by blast
ultimately show ?thesis by fast
qed
lemma insert_before_wf_darcs:
"\<lbrakk>wf_darcs (Node r xs); e \<notin> darcs (Node r xs); (t1,e1) \<in> fset (insert_before v e y xs)\<rbrakk>
\<Longrightarrow> wf_darcs t1"
proof(induction xs)
case (insert x xs)
let ?f = "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1))"
show ?case
proof(cases "(t1,e1) \<in> fset (insert_before v e y xs)")
case in_xs: True
then show ?thesis
proof(cases "?f x = (t1,e1)")
case True
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have "insert_before v e y (finsert x xs) = insert_before v e y xs"
- using True in_xs notin_fset by fastforce
+ using True in_xs by fastforce
moreover have "disjoint_darcs xs"
using disjoint_darcs_insert[OF disjoint_darcs_if_wf_xs[OF insert.prems(1)]] .
ultimately show ?thesis
using insert.IH insert.prems unfolding wf_darcs_iff_darcs' by force
next
case False
have "disjoint_darcs xs"
using disjoint_darcs_insert[OF disjoint_darcs_if_wf_xs[OF insert.prems(1)]] .
then show ?thesis
using in_xs False insert.IH insert.prems(1,2) by (simp add: wf_darcs_iff_darcs')
qed
next
case False
have "insert_before v e y (finsert x xs) = finsert (?f x) (insert_before v e y xs)"
by (simp add: insert.hyps prod.case_distrib)
then have 0: "?f x = (t1,e1)" using False insert.prems(3) by fastforce
then show ?thesis
proof(cases "e1=e")
case True
then have "(t1,e1) \<notin> fset (finsert x xs)" using insert.prems(2) dtree.set_sel(1) by force
then obtain t2 e2 where
t2_def: "(t2, e2)\<in>fset (finsert x xs)" "Node v {|(t2, e2)|} = t1" "root t2 = y" "e1=e"
using in_insert_before_child_in_orig[of t1] insert.prems(3) by blast
then show ?thesis
using insert.prems(1) t2_def by (fastforce simp: wf_darcs_iff_darcs')
next
case False
then have "(t1,e1) = x"
by (smt (z3) 0 old.prod.exhaust prod.inject case_prod_Pair_iden case_prod_conv)
then show ?thesis using insert.prems(1) by auto
qed
qed
qed (simp)
lemma disjoint_darcs_aux2:
assumes "wf_darcs (Node r xs)" and "e \<notin> darcs (Node r xs)"
shows "disjoint_darcs (finsert (Node v {||},e) xs)"
using assms unfolding wf_darcs_iff_darcs' by fastforce
lemma disjoint_darcs_aux3_aux1:
assumes "(t,e1) \<in> fset xs"
and "x \<in> dverts t"
and "wf_darcs (Node r xs)"
and "e \<notin> darcs (Node r xs)"
and "(t2,e2) \<in> (\<lambda>(t1,e1). (insert_between v e x y t1, e1)) ` fset xs"
and "(t3,e3) \<in> (\<lambda>(t1,e1). (insert_between v e x y t1, e1)) ` fset xs"
and "(t2,e2)\<noteq>(t3,e3)"
and "wf_dverts (Node r xs)"
shows "(darcs t2 \<union> {e2}) \<inter> (darcs t3 \<union> {e3}) = {}"
proof -
have "\<forall>(t2,e2) \<in> fset xs. (t,e1)=(t2,e2) \<or> x \<notin> dverts t2"
using assms(1,2,8) by (fastforce simp: wf_dverts_iff_dverts')
then have nt1_id: "\<forall>(t2,e2) \<in> fset xs. (t,e1) = (t2,e2) \<or> insert_between v e x y t2 = t2"
using insert_between_id_if_notin by fastforce
have darcs_t: "darcs (insert_between v e x y t) = insert e (darcs t)"
using assms(2,3) by (simp add: insert_between_add_e_if_x_in)
consider "(t2,e2) = (insert_between v e x y t,e1)"
| "(t3,e3) = (insert_between v e x y t,e1)"
| "(t2,e2) \<noteq> (insert_between v e x y t,e1)" "(t3,e3) \<noteq> (insert_between v e x y t,e1)"
by fast
then show ?thesis
proof(cases)
case 1
then have "(t3,e3) \<in> fset xs" using assms(6,7) nt1_id by fastforce
moreover have "(t3,e3) \<noteq> (t,e1)" using assms(6,7) 1 nt1_id by fastforce
ultimately have "(darcs t \<union> {e1,e}) \<inter> (darcs t3 \<union> {e3}) = {}"
using assms(1,3,4) unfolding wf_darcs_iff_darcs' by fastforce
then show ?thesis using 1 darcs_t by auto
next
case 2
then have "(t2,e2) \<in> fset xs" using assms(5,7) nt1_id by fastforce
moreover have "(t2,e2) \<noteq> (t,e1)" using assms(5,7) 2 nt1_id by auto
ultimately have "(darcs t \<union> {e1,e}) \<inter> (darcs t2 \<union> {e2}) = {}"
using assms(1,3,4) unfolding wf_darcs_iff_darcs' by fastforce
then show ?thesis using 2 darcs_t by force
next
case 3
then have "(t2,e2) \<in> fset xs" using assms(5) nt1_id by fastforce
moreover have "(t3,e3) \<in> fset xs" using assms(6) 3(2) nt1_id by auto
ultimately show ?thesis using assms(3,7) unfolding wf_darcs_iff_darcs' by fastforce
qed
qed
lemma disjoint_darcs_aux3_aux2:
assumes "(t,e1) \<in> fset xs"
and "x \<in> dverts t"
and "wf_darcs (Node r xs)"
and "e \<notin> darcs (Node r xs)"
and "(t2,e2) \<in> (\<lambda>(t1,e1). (insert_between v e x y t1, e1)) ` fset xs"
and "wf_dverts (Node r xs)"
shows "e2 \<notin> darcs t2"
proof(cases "(t2,e2) \<in> fset xs")
case True
then show ?thesis using assms(3) unfolding wf_darcs_iff_darcs' by auto
next
case False
obtain t1 where t1_def: "insert_between v e x y t1 = t2" "(t1,e2) \<in> fset xs"
using assms(5) by fast
then have "x \<in> dverts t1" using insert_between_id_if_notin False by fastforce
then have "t = t1" using assms(1,2,6) t1_def(2) by (fastforce simp: wf_dverts_iff_dverts')
then have darcs_t: "darcs t2 = insert e (darcs t1)"
using insert_between_add_e_if_x_in assms(2) t1_def(1) by force
then show ?thesis using assms(3,4) t1_def(2) unfolding wf_darcs_iff_darcs' by fastforce
qed
lemma disjoint_darcs_aux3:
assumes "(t,e1) \<in> fset xs"
and "x \<in> dverts t"
and "wf_darcs (Node r xs)"
and "e \<notin> darcs (Node r xs)"
and "wf_dverts (Node r xs)"
shows "disjoint_darcs ((\<lambda>(t1,e1). (insert_between v e x y t1, e1)) |`| xs)"
proof -
let ?xs = "(\<lambda>(t1,e1). (insert_between v e x y t1, e1)) |`| xs"
let ?xs' = "(\<lambda>(t1,e1). (insert_between v e x y t1, e1)) ` fset xs"
have 0: "fset ?xs = ?xs'" by simp
then have "\<forall>(t1,e1) \<in> fset ?xs. e1 \<notin> darcs t1"
using disjoint_darcs_aux3_aux2 assms by blast
moreover have "\<forall>(t1,e1) \<in> ?xs'. \<forall>(t2,e2) \<in> ?xs'.
(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {} \<or> (t1,e1) = (t2,e2)"
using disjoint_darcs_aux3_aux1 assms by blast
ultimately show ?thesis using 0 by fastforce
qed
lemma insert_between_wf_darcs:
"\<lbrakk>e \<notin> darcs t; v \<notin> dverts t \<rbrakk> \<Longrightarrow> wf_darcs (insert_between v e x y t)"
using wf_dtree_axioms proof(induction t)
case (Node r xs)
then interpret wf_dtree "Node r xs" by blast
consider "x=r" "\<exists>t. t \<in> fst ` fset xs \<and> root t = y"
| "x=r" "\<not>(\<exists>t. t \<in> fst ` fset xs \<and> root t = y)" | "x\<noteq>r" by fast
then show ?case
proof(cases)
case 1
then have "insert_between v e x y (Node r xs) = Node r (insert_before v e y xs)" by simp
moreover have "\<forall>(x,e1) \<in> fset (insert_before v e y xs). wf_darcs x"
using insert_before_wf_darcs Node.prems(1) wf_arcs by fast
moreover have "disjoint_darcs (insert_before v e y xs)"
using disjoint_darcs_aux1[OF wf_verts Node.prems(2) wf_arcs Node.prems(1)] .
ultimately show ?thesis by (simp add: wf_darcs_if_darcs'_aux)
next
case 2
then have "insert_between v e x y (Node r xs) = Node r (finsert (Node v {||},e) xs)" by simp
then show ?thesis
using disjoint_darcs_aux2 Node.prems(1) wf_arcs by (simp add: wf_darcs_iff_darcs')
next
case 3
let ?f = "\<lambda>(t1,e1). (insert_between v e x y t1, e1)"
show ?thesis
proof(cases "\<exists>(t1,e1) \<in> fset xs. x \<in> dverts t1")
case True
then obtain t1 e1 where t1_def: "(t1,e1) \<in> fset xs" " x \<in> dverts t1" by blast
then interpret T: wf_dtree t1 using wf_dtree_rec by blast
have "\<And>t2 e2. (t2,e2) \<in> fset (?f |`| xs) \<longrightarrow> wf_darcs t2"
proof
fix t2 e2
assume asm: "(t2,e2) \<in> fset (?f |`| xs)"
then show "wf_darcs t2"
proof(cases "(t2,e2) = (insert_between v e x y t1,e1)")
case True
then have "wf_darcs (insert_between v e x y t1)"
using Node t1_def(1) T.wf_dtree_axioms
by (metis dtree.set_intros(2) dtree.set_intros(3) insertI1 prod_set_simps(1))
then show ?thesis using True by blast
next
case False
have "\<forall>(t2,e2) \<in> fset xs. (t1,e1)=(t2,e2) \<or> x \<notin> dverts t2"
using wf_verts t1_def by (fastforce simp: wf_dverts_iff_dverts')
then have "\<forall>(t2,e2) \<in> fset xs. (t1,e1) = (t2,e2) \<or> insert_between v e x y t2 = t2"
using insert_between_id_if_notin by fastforce
then show ?thesis using wf_arcs asm False by fastforce
qed
qed
moreover have "disjoint_darcs (?f |`| xs)"
using T.disjoint_darcs_aux3 Node.prems(1) t1_def wf_arcs wf_verts by presburger
ultimately show ?thesis using 3 by (fastforce simp: wf_darcs_iff_darcs')
next
case False
then show ?thesis
using wf_arcs 3 insert_between_id_if_notin fst_conv
by (smt (verit, ccfv_threshold) fsts.cases dtree.inject dtree.set_cases(1) case_prodI2)
qed
qed
qed
theorem insert_between_wf_dtree:
"\<lbrakk>e \<notin> darcs t; v \<notin> dverts t \<rbrakk> \<Longrightarrow> wf_dtree (insert_between v e x y t)"
by (simp add: insert_between_wf_dverts insert_between_wf_darcs wf_dtree_def)
lemma snds_neq_card_eq_card_snd:
"\<forall>(t,e) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. e\<noteq>e2 \<or> (t,e) = (t2,e2) \<Longrightarrow> fcard xs = fcard (snd |`| xs)"
proof(induction xs)
case empty
then have "(snd |`| {||}) = {||}" by blast
then show ?case by (simp add: fcard_fempty)
next
case (insert x xs)
have "fcard xs = fcard (snd |`| xs)" using insert.IH insert.prems by fastforce
moreover have "snd x |\<notin>| snd |`| xs"
proof
assume asm: "snd x |\<in>| snd |`| xs"
then obtain t e where t_def: "x = (t,e)" by fastforce
then obtain t2 where t2_def: "(t2,e) |\<in>| xs" using asm by auto
then have "(t,e)\<noteq>(t2,e)" using insert.hyps t_def by blast
- moreover have "(t,e) \<in> fset (finsert x xs)" using t_def notin_fset by simp
- moreover have "(t2,e) \<in> fset (finsert x xs)" using t2_def notin_fset by fastforce
+ moreover have "(t,e) \<in> fset (finsert x xs)" using t_def by simp
+ moreover have "(t2,e) \<in> fset (finsert x xs)" using t2_def by fastforce
ultimately show False using insert.prems by fast
qed
ultimately show ?case by (simp add: fcard_finsert_disjoint local.insert.hyps)
qed
lemma snds_neq_img_snds_neq:
assumes "\<forall>(t,e) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. e\<noteq>e2 \<or> (t,e) = (t2,e2)"
shows "\<forall>(t1,e1) \<in> fset ((\<lambda>(t1,e1). (f t1, e1)) |`| xs).
\<forall>(t2,e2) \<in> fset ((\<lambda>(t1,e1). (f t1, e1)) |`| xs). e1\<noteq>e2 \<or> (t1,e1) = (t2,e2)"
using assms by auto
lemma snds_neq_if_disjoint_darcs:
assumes "disjoint_darcs xs"
shows "\<forall>(t,e) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. e\<noteq>e2 \<or> (t,e) = (t2,e2)"
using assms by fast
lemma snds_neq_img_card_eq:
assumes "\<forall>(t,e) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. e\<noteq>e2 \<or> (t,e) = (t2,e2)"
shows "fcard ((\<lambda>(t1,e1). (f t1, e1)) |`| xs) = fcard xs"
proof -
let ?f = "\<lambda>(t1,e1). (f t1, e1)"
have "\<forall>(t,e) \<in> fset (?f |`| xs). \<forall>(t2,e2) \<in> fset (?f |`| xs). e\<noteq>e2 \<or> (t,e) = (t2,e2)"
using assms snds_neq_img_snds_neq by auto
then have "fcard (?f |`| xs) = fcard (snd |`| (?f |`| xs))"
using snds_neq_card_eq_card_snd by blast
moreover have "snd |`| (?f |`| xs) = snd |`| xs" by force
moreover have "fcard xs = fcard (snd |`| xs)" using snds_neq_card_eq_card_snd assms by blast
ultimately show ?thesis by simp
qed
lemma fst_neq_img_card_eq:
assumes "\<forall>(t,e) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. f t \<noteq> f t2 \<or> (t,e) = (t2,e2)"
shows "fcard ((\<lambda>(t1,e1). (f t1, e1)) |`| xs) = fcard xs"
using assms proof(induction xs)
case empty
then have "(snd |`| {||}) = {||}" by blast
then show ?case by (simp add: fcard_fempty)
next
case (insert x xs)
have "fcard xs = fcard ((\<lambda>(t1,e1). (f t1, e1)) |`| xs)" using insert by fastforce
moreover have "(\<lambda>(t1,e1). (f t1, e1)) x |\<notin>| (\<lambda>(t1,e1). (f t1, e1)) |`| xs"
proof
assume asm: "(\<lambda>(t1,e1). (f t1, e1)) x |\<in>| (\<lambda>(t1,e1). (f t1, e1)) |`| xs"
then obtain t e where t_def: "x = (t,e)" by fastforce
then obtain t2 e2 where t2_def:
"(t2,e2) |\<in>| xs" "(\<lambda>(t1,e1). (f t1, e1)) (t2,e2) = (\<lambda>(t1,e1). (f t1, e1)) (t,e)"
using asm by auto
then have "(t,e)\<noteq>(t2,e)" using insert.hyps t_def by fast
- moreover have "(t,e) \<in> fset (finsert x xs)" using t_def notin_fset by simp
- moreover have "(t2,e2) \<in> fset (finsert x xs)" using t2_def(1) notin_fset by fastforce
+ moreover have "(t,e) \<in> fset (finsert x xs)" using t_def by simp
+ moreover have "(t2,e2) \<in> fset (finsert x xs)" using t2_def(1) by fastforce
ultimately show False using insert.prems t2_def(2) by fast
qed
ultimately show ?case by (simp add: fcard_finsert_disjoint local.insert.hyps)
qed
lemma x_notin_insert_before:
assumes "x |\<notin>| xs" and "wf_dverts (Node r (finsert x xs))"
shows "(\<lambda>(t1,e1). if root t1 = y then (Node v {|(t1,e1)|},e) else (t1,e1)) x
|\<notin>| (insert_before v e y xs)" (is "?f x |\<notin>|_")
proof (cases "root (fst x) = y")
case True
then obtain t1 e1 where t1_def: "x = (t1,e1)" "root t1 = y" by fastforce
then have 0: "\<forall>(t2,e2) \<in> fset xs. dverts t1 \<inter> dverts t2 = {}"
- using assms notin_fset disjoint_dverts_if_wf_aux by fastforce
+ using assms disjoint_dverts_if_wf_aux by fastforce
then have "\<forall>(t2,e2) \<in> fset xs. root t2 \<noteq> y"
by (smt (verit, del_insts) dtree.set_sel(1) t1_def(2) case_prodD case_prodI2 disjoint_iff)
+ hence "\<nexists>t. t \<in> fst ` fset xs \<and> dtree.root t = y"
+ by fastforce
then have 1: "(insert_before v e y xs) = xs" using insert_before_not_y_id by fastforce
have "?f x = (Node v {|(t1,e1)|},e)" using t1_def by simp
then have "\<forall>(t2,e2) \<in> fset xs. (fst (?f x)) \<noteq> t2" using 0 dtree.set_sel(1) by fastforce
then have "\<forall>(t2,e2) \<in> fset (insert_before v e y xs). ?f x \<noteq> (t2,e2)" using 1 by fastforce
- then show ?thesis using notin_fset by fast
+ then show ?thesis by fast
next
case False
then have x_id: "?f x = x" by (smt (verit) old.prod.exhaust case_prod_conv fst_conv)
then show ?thesis
proof(cases "\<exists>t1. t1 \<in> fst ` fset xs \<and> root t1 = y")
case True
then obtain t1 e1 where t1_def: "(t1,e1) \<in> fset xs" "root t1 = y" by force
then have "(t1,e1) \<in> fset (finsert x xs)" by auto
then have 0: "\<forall>(t2,e2) \<in> fset (finsert x xs). (t1,e1) = (t2,e2) \<or> dverts t1 \<inter> dverts t2 = {}"
using assms(2) disjoint_dverts_if_wf_aux by fast
then have "\<forall>(t2,e2) \<in> fset (finsert x xs). (t1,e1) = (t2,e2) \<or> root t2 \<noteq> y"
using dtree.set_sel(1) t1_def(2) insert_not_empty
by (smt (verit, ccfv_threshold) Int_insert_right_if1 prod.case_eq_if insert_absorb)
then have "\<nexists>t. t \<in> fst ` fset (xs |-| {|(t1,e1)|}) \<and> root t = y" by fastforce
then have 1: "?f |`| (xs |-| {|(t1,e1)|}) = (xs |-| {|(t1,e1)|})"
using insert_before_not_y_id[of "xs |-| {|(t1,e1)|}"] by (simp add: insert_before_alt)
have "?f (t1,e1) = (Node v {|(t1,e1)|},e)" using t1_def by simp
then have "?f |`| xs = finsert (Node v {|(t1,e1)|},e) (?f |`| (xs |-| {|(t1,e1)|}))"
- using t1_def(1) notin_fset by (metis (no_types, lifting) fimage_finsert finsert_fminus)
+ using t1_def(1) by (metis (no_types, lifting) fimage_finsert finsert_fminus)
then have "?f |`| xs = finsert (Node v {|(t1,e1)|},e) (xs |-| {|(t1,e1)|})"
using 1 by simp
then have 2: "insert_before v e y xs = finsert (Node v {|(t1,e1)|},e) (xs |-| {|(t1,e1)|})"
by (simp add: insert_before_alt)
- have "dverts t1 \<inter> dverts (fst x) = {}" using 0 assms(1) notin_fset t1_def(1) by fastforce
+ have "dverts t1 \<inter> dverts (fst x) = {}" using 0 assms(1) t1_def(1) by fastforce
then have "(Node v {|(t1,e1)|},e) \<noteq> x" using dtree.set_sel(1) by fastforce
then show ?thesis using 2 assms(1) x_id by auto
next
case False
then have "(insert_before v e y xs) = xs" using insert_before_not_y_id by fastforce
then show ?thesis using assms(1) x_id by simp
qed
qed
end
end
\ No newline at end of file
diff --git a/thys/Query_Optimization/IKKBZ.thy b/thys/Query_Optimization/IKKBZ.thy
--- a/thys/Query_Optimization/IKKBZ.thy
+++ b/thys/Query_Optimization/IKKBZ.thy
@@ -1,2977 +1,2993 @@
(* Author: Bernhard Stöckl *)
theory IKKBZ
imports Complex_Main "CostFunctions" "QueryGraph" "List_Dtree" "HOL-Library.Sorting_Algorithms"
begin
section \<open>IKKBZ\<close>
subsection \<open>Additional Proofs for Merging Lists\<close>
lemma merge_comm_if_not_equiv: "\<forall>x \<in> set xs. \<forall>y \<in> set ys. compare cmp x y \<noteq> Equiv \<Longrightarrow>
Sorting_Algorithms.merge cmp xs ys = Sorting_Algorithms.merge cmp ys xs"
apply(induction xs ys rule: Sorting_Algorithms.merge.induct)
by(auto intro: compare.quasisym_not_greater simp: compare.asym_greater)
lemma set_merge: "set xs \<union> set ys = set (Sorting_Algorithms.merge cmp xs ys)"
using mset_merge set_mset_mset set_mset_union by metis
lemma input_empty_if_merge_empty: "Sorting_Algorithms.merge cmp xs ys = [] \<Longrightarrow> xs = [] \<and> ys = []"
using Un_empty set_empty2 set_merge by metis
lemma merge_assoc:
"Sorting_Algorithms.merge cmp xs (Sorting_Algorithms.merge cmp ys zs)
= Sorting_Algorithms.merge cmp (Sorting_Algorithms.merge cmp xs ys) zs"
(is "?merge _ xs (?merge cmp _ zs) = _")
proof(induction xs "?merge cmp ys zs" arbitrary: ys zs taking: cmp rule: Sorting_Algorithms.merge.induct)
case (2 cmp v vs)
show ?case using input_empty_if_merge_empty[OF 2[symmetric]] by simp
next
case ind: (3 x xs r rs)
then show ?case
proof(induction ys zs taking: cmp rule: Sorting_Algorithms.merge.induct)
case (3 y ys z zs)
then show ?case
using ind compare.asym_greater
by (smt (verit, best) compare.trans_not_greater list.inject merge.simps(3))
qed (auto)
qed (simp)
lemma merge_comp_commute:
assumes "\<forall>x \<in> set xs. \<forall>y \<in> set ys. compare cmp x y \<noteq> Equiv"
shows "Sorting_Algorithms.merge cmp xs (Sorting_Algorithms.merge cmp ys zs)
= Sorting_Algorithms.merge cmp ys (Sorting_Algorithms.merge cmp xs zs)"
using assms merge_assoc merge_comm_if_not_equiv by metis
lemma wf_list_arcs_merge:
"\<lbrakk>wf_list_arcs xs; wf_list_arcs ys; snd ` set xs \<inter> snd ` set ys = {}\<rbrakk>
\<Longrightarrow> wf_list_arcs (Sorting_Algorithms.merge cmp xs ys)"
proof(induction xs ys taking: cmp rule: Sorting_Algorithms.merge.induct)
case (3 x xs y ys)
obtain v1 e1 where v1_def[simp]: "x = (v1,e1)" by force
obtain v2 e2 where v2_def[simp]: "y = (v2,e2)" by force
show ?case
proof(cases "compare cmp x y = Greater")
case True
have "e2 \<notin> snd ` set (x#xs)" using "3.prems"(3) by auto
moreover have "e2 \<notin> snd ` set ys" using "3.prems"(2) by simp
ultimately have "e2 \<notin> snd ` set (Sorting_Algorithms.merge cmp (x#xs) ys)"
using set_merge by fast
then show ?thesis using True 3 by force
next
case False
have "e1 \<notin> snd `set (y#ys)" using "3.prems"(3) by auto
moreover have "e1 \<notin> snd ` set xs" using "3.prems"(1) by simp
ultimately have "e1 \<notin> snd `set (Sorting_Algorithms.merge cmp xs (y#ys))"
using set_merge by fast
then show ?thesis using False 3 by force
qed
qed (auto)
lemma wf_list_lverts_merge:
"\<lbrakk>wf_list_lverts xs; wf_list_lverts ys;
\<forall>v1 \<in> fst ` set xs. \<forall>v2 \<in> fst ` set ys. set v1 \<inter> set v2 = {}\<rbrakk>
\<Longrightarrow> wf_list_lverts (Sorting_Algorithms.merge cmp xs ys)"
proof(induction xs ys taking: cmp rule: Sorting_Algorithms.merge.induct)
case (3 x xs y ys)
obtain v1 e1 where v1_def[simp]: "x = (v1,e1)" by force
obtain v2 e2 where v2_def[simp]: "y = (v2,e2)" by force
show ?case
proof(cases "compare cmp x y = Greater")
case True
have "\<forall>v \<in> fst ` set (x#xs). set v2 \<inter> set v = {}" using "3.prems"(3) by auto
moreover have "\<forall>v \<in> fst ` set ys. set v2 \<inter> set v = {}" using "3.prems"(2) by simp
ultimately have "\<forall>v \<in> fst ` set (Sorting_Algorithms.merge cmp (x#xs) ys). set v2 \<inter> set v = {}"
using set_merge[of "x#xs"] by blast
then show ?thesis using True 3 by force
next
case False
have "\<forall>v \<in> fst ` set (y#ys). set v1 \<inter> set v = {}" using "3.prems"(3) by auto
moreover have "\<forall>v \<in> fst ` set xs. set v1 \<inter> set v = {}" using "3.prems"(1) by simp
ultimately have "\<forall>v \<in> fst ` set (Sorting_Algorithms.merge cmp xs (y#ys)). set v1 \<inter> set v = {}"
using set_merge[of xs] by auto
then show ?thesis using False 3 by force
qed
qed (auto)
lemma merge_hd_exists_preserv:
"\<lbrakk>\<exists>(t1,e1) \<in> fset xs. hd as = (root t1,e1); \<exists>(t1,e1) \<in> fset xs. hd bs = (root t1,e1)\<rbrakk>
\<Longrightarrow> \<exists>(t1,e1) \<in> fset xs. hd (Sorting_Algorithms.merge cmp as bs) = (root t1,e1)"
by(induction as bs rule: Sorting_Algorithms.merge.induct) auto
lemma merge_split_supset:
assumes "as@r#bs = (Sorting_Algorithms.merge cmp xs ys)"
shows "\<exists>bs' as'. set bs' \<subseteq> set bs \<and> (as'@r#bs' = xs \<or> as'@r#bs' = ys)"
using assms proof(induction xs ys arbitrary: as taking: cmp rule: Sorting_Algorithms.merge.induct)
case (3 x xs y ys)
let ?merge = "Sorting_Algorithms.merge cmp"
show ?case
proof(cases "compare cmp x y = Greater")
case True
then show ?thesis
proof(cases as)
case Nil
have "set ys \<subseteq> set (?merge (x#xs) ys)" using set_merge by fast
then show ?thesis using Nil True "3.prems" by auto
next
case (Cons c cs)
then have "cs@r#bs = ?merge (x#xs) ys" using True "3.prems" by simp
then obtain as' bs' where as_def: "set bs' \<subseteq> set bs" "as'@r#bs' = x#xs \<or> as'@r#bs' = ys"
using "3.IH"(1)[OF True] by blast
have "as'@r#bs' = x#xs \<or> (y#as')@r#bs' = y#ys" using as_def(2) by simp
then show ?thesis using as_def(1) by blast
qed
next
case False
then show ?thesis
proof(cases as)
case Nil
have "set xs \<subseteq> set (?merge xs (y#ys))" using set_merge by fast
then show ?thesis using Nil False "3.prems" by auto
next
case (Cons c cs)
then have "cs@r#bs = ?merge xs (y#ys)" using False "3.prems" by simp
then obtain as' bs' where as_def: "set bs' \<subseteq> set bs" "as'@r#bs' = xs \<or> as'@r#bs' = y#ys"
using "3.IH"(2)[OF False] by blast
have "(x#as')@r#bs' = x#xs \<or> as'@r#bs' = y#ys" using as_def(2) by simp
then show ?thesis using as_def(1) by blast
qed
qed
qed(auto)
lemma merge_split_supset_fst:
assumes "as@(r,e)#bs = (Sorting_Algorithms.merge cmp xs ys)"
shows "\<exists>as' bs'. set bs' \<subseteq> set bs \<and> (as'@(r,e)#bs' = xs \<or> as'@(r,e)#bs' = ys)"
using merge_split_supset[OF assms] by blast
lemma merge_split_supset':
assumes "r \<in> set (Sorting_Algorithms.merge cmp xs ys)"
shows "\<exists>as bs as' bs'. as@r#bs = (Sorting_Algorithms.merge cmp xs ys)
\<and> set bs' \<subseteq> set bs \<and> (as'@r#bs' = xs \<or> as'@r#bs' = ys)"
using merge_split_supset split_list[OF assms] by metis
lemma merge_split_supset_fst':
assumes "r \<in> fst ` set (Sorting_Algorithms.merge cmp xs ys)"
shows "\<exists>as e bs as' bs'. as@(r,e)#bs = (Sorting_Algorithms.merge cmp xs ys)
\<and> set bs' \<subseteq> set bs \<and> (as'@(r,e)#bs' = xs \<or> as'@(r,e)#bs' = ys)"
proof -
obtain e where "(r,e) \<in> set (Sorting_Algorithms.merge cmp xs ys)" using assms by auto
then show ?thesis using merge_split_supset'[of "(r,e)"] by blast
qed
lemma merge_split_supset_subtree:
assumes "\<forall>as bs. as@(r,e)#bs = xs \<longrightarrow>
(\<exists>zs. is_subtree (Node r zs) t \<and> dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs))"
and "\<forall>as bs. as@(r,e)#bs = ys \<longrightarrow>
(\<exists>zs. is_subtree (Node r zs) t \<and> dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs))"
and "as@(r,e)#bs = (Sorting_Algorithms.merge cmp xs ys)"
shows "\<exists>zs. is_subtree (Node r zs) t \<and> dverts (Node r zs) \<subseteq> (fst ` set ((r,e)#bs))"
proof -
obtain as' bs' where bs'_def: "set bs' \<subseteq> set bs" "as'@(r,e)#bs' = xs \<or> as'@(r,e)#bs' = ys"
using merge_split_supset[OF assms(3)] by blast
obtain zs where zs_def: "is_subtree (Node r zs) t" "dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs')"
using assms(1,2) bs'_def(2) by blast
then have "dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs)" using bs'_def(1) by auto
then show ?thesis using zs_def(1) by blast
qed
lemma merge_split_supset_strict_subtree:
assumes "\<forall>as bs. as@(r,e)#bs = xs \<longrightarrow> (\<exists>zs. strict_subtree (Node r zs) t
\<and> dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs))"
and "\<forall>as bs. as@(r,e)#bs = ys \<longrightarrow> (\<exists>zs. strict_subtree (Node r zs) t
\<and> dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs))"
and "as@(r,e)#bs = (Sorting_Algorithms.merge cmp xs ys)"
shows "\<exists>zs. strict_subtree (Node r zs) t
\<and> dverts (Node r zs) \<subseteq> (fst ` set ((r,e)#bs))"
proof -
obtain as' bs' where bs'_def: "set bs' \<subseteq> set bs" "as'@(r,e)#bs' = xs \<or> as'@(r,e)#bs' = ys"
using merge_split_supset[OF assms(3)] by blast
obtain zs where zs_def:
"strict_subtree (Node r zs) t" "dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs')"
using assms(1,2) bs'_def(2) by blast
then have "dverts (Node r zs) \<subseteq> fst ` set ((r,e)#bs)" using bs'_def(1) by auto
then show ?thesis using zs_def(1,2) by blast
qed
lemma sorted_app_l: "sorted cmp (xs@ys) \<Longrightarrow> sorted cmp xs"
by(induction xs rule: sorted.induct) auto
lemma sorted_app_r: "sorted cmp (xs@ys) \<Longrightarrow> sorted cmp ys"
by(induction xs) (auto simp: sorted_Cons_imp_sorted)
subsection \<open>Merging Subtrees of Ranked Dtrees\<close>
locale ranked_dtree = list_dtree t for t :: "('a list,'b) dtree" +
fixes rank :: "'a list \<Rightarrow> real"
fixes cmp :: "('a list\<times>'b) comparator"
assumes cmp_antisym:
"\<lbrakk>v1 \<noteq> []; v2 \<noteq> []; compare cmp (v1,e1) (v2,e2) = Equiv\<rbrakk> \<Longrightarrow> set v1 \<inter> set v2 \<noteq> {} \<or> e1=e2"
begin
lemma ranked_dtree_rec: "\<lbrakk>Node r xs = t; (x,e) \<in> fset xs\<rbrakk> \<Longrightarrow> ranked_dtree x cmp"
using wf_arcs wf_lverts by(unfold_locales) (auto dest: cmp_antisym)
lemma ranked_dtree_rec_suc: "(x,e) \<in> fset (sucs t) \<Longrightarrow> ranked_dtree x cmp"
using ranked_dtree_rec[of "root t"] by force
lemma ranked_dtree_subtree: "is_subtree x t \<Longrightarrow> ranked_dtree x cmp"
using ranked_dtree_axioms proof(induction t)
case (Node r xs)
then interpret ranked_dtree "Node r xs" by blast
show ?case using Node ranked_dtree_rec by (cases "x = Node r xs") auto
qed
subsubsection \<open>Definitions\<close>
lift_definition cmp' :: "('a list\<times>'b) comparator" is
"(\<lambda>x y. if rank (rev (fst x)) < rank (rev (fst y)) then Less
else if rank (rev (fst x)) > rank (rev (fst y)) then Greater
else compare cmp x y)"
by (smt (z3) comp.distinct(3) compare.less_iff_sym_greater compare.refl compare.trans_equiv
compare.trans_less comparator_def)
abbreviation disjoint_sets :: "(('a list, 'b) dtree \<times> 'b) fset \<Rightarrow> bool" where
"disjoint_sets xs \<equiv> disjoint_darcs xs \<and> disjoint_dlverts xs \<and> (\<forall>(t,e) \<in> fset xs. [] \<notin> dverts t)"
abbreviation merge_f :: "'a list \<Rightarrow> (('a list, 'b) dtree \<times> 'b) fset
\<Rightarrow> ('a list, 'b) dtree \<times> 'b \<Rightarrow> ('a list \<times> 'b) list \<Rightarrow> ('a list \<times> 'b) list" where
"merge_f r xs \<equiv> \<lambda>(t,e) b. if (t,e) \<in> fset xs \<and> list_dtree (Node r xs)
\<and> (\<forall>(v,e') \<in> set b. set v \<inter> dlverts t = {} \<and> v \<noteq> [] \<and> e' \<notin> darcs t \<union> {e})
then Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t,e)|})) b else b"
definition merge :: "('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"merge t1 \<equiv> dtree_from_list (root t1) (ffold (merge_f (root t1) (sucs t1)) [] (sucs t1))"
subsubsection \<open>Commutativity Proofs\<close>
lemma cmp_sets_not_dsjnt_if_equiv:
"\<lbrakk>v1 \<noteq> []; v2 \<noteq> []\<rbrakk> \<Longrightarrow> compare cmp' (v1,e1) (v2,e2) = Equiv \<Longrightarrow> set v1 \<inter> set v2 \<noteq> {} \<or> e1=e2"
by(auto simp: cmp'.rep_eq dest: cmp_antisym split: if_splits)
lemma dtree_to_list_x_in_dverts:
"x \<in> fst ` set (dtree_to_list (Node r {|(t1,e1)|})) \<Longrightarrow> x \<in> dverts t1"
using dtree_to_list_sub_dverts_ins by auto
lemma dtree_to_list_x_in_dlverts:
"x \<in> fst ` set (dtree_to_list (Node r {|(t1,e1)|})) \<Longrightarrow> set x \<subseteq> dlverts t1"
using dtree_to_list_x_in_dverts lverts_if_in_verts by fast
lemma dtree_to_list_x1_disjoint:
"dlverts t1 \<inter> dlverts t2 = {}
\<Longrightarrow> \<forall>x1 \<in> fst ` set (dtree_to_list (Node r {|(t1,e1)|})). set x1 \<inter> dlverts t2 = {}"
using dtree_to_list_x_in_dlverts by fast
lemma dtree_to_list_xs_disjoint:
"dlverts t1 \<inter> dlverts t2 = {}
\<Longrightarrow> \<forall>x1 \<in> fst ` set (dtree_to_list (Node r {|(t1,e1)|})).
\<forall>x2 \<in> fst ` set (dtree_to_list (Node r' {|(t2,e2)|})). set x1 \<inter> set x2 = {}"
using dtree_to_list_x_in_dlverts by (metis inf_mono subset_empty)
lemma dtree_to_list_e_in_darcs:
"e \<in> snd ` set (dtree_to_list (Node r {|(t1,e1)|})) \<Longrightarrow> e \<in> darcs t1 \<union> {e1}"
using dtree_to_list_sub_darcs by fastforce
lemma dtree_to_list_e_disjoint:
"(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}
\<Longrightarrow> \<forall>e \<in> snd ` set (dtree_to_list (Node r {|(t1,e1)|})). e \<notin> darcs t2 \<union> {e2}"
using dtree_to_list_e_in_darcs by fast
lemma dtree_to_list_es_disjoint:
"(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}
\<Longrightarrow> \<forall>e3 \<in> snd ` set (dtree_to_list (Node r {|(t1,e1)|})).
\<forall>e4 \<in> snd ` set (dtree_to_list (Node r' {|(t2,e2)|})). e3 \<noteq> e4"
using dtree_to_list_e_disjoint dtree_to_list_e_in_darcs by fast
lemma dtree_to_list_xs_not_equiv:
assumes "dlverts t1 \<inter> dlverts t2 = {}"
and "(darcs t1 \<union> {e3}) \<inter> (darcs t2 \<union> {e4}) = {}"
and "(x1,e1) \<in> set (dtree_to_list (Node r {|(t1,e3)|}))" and "x1 \<noteq> []"
and "(x2,e2) \<in> set (dtree_to_list (Node r' {|(t2,e4)|}))" and "x2 \<noteq> []"
shows "compare cmp' (x1,e1) (x2,e2) \<noteq> Equiv"
using dtree_to_list_xs_disjoint[OF assms(1)] cmp_sets_not_dsjnt_if_equiv[of x1 x2 e1 e2]
dtree_to_list_es_disjoint[OF assms(2)] assms(3-6) by fastforce
lemma merge_dtree1_not_equiv:
assumes "dlverts t1 \<inter> dlverts t2 = {}"
and "(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
and "[] \<notin> dverts t1"
and "[] \<notin> dverts t2"
and "xs = dtree_to_list (Node r {|(t1,e1)|})"
and "ys = dtree_to_list (Node r' {|(t2,e2)|})"
shows "\<forall>(x1,e1)\<in>set xs. \<forall>(x2,e2)\<in>set ys. compare cmp' (x1,e1) (x2,e2) \<noteq> Equiv"
proof -
have "\<forall>(x1,e1)\<in>set xs. x1 \<noteq> []"
using assms(3,5) dtree_to_list_x_in_dverts
by (smt (verit) case_prod_conv case_prod_eta fst_conv pair_imageI surj_pair)
moreover have "\<forall>(x1,e1)\<in>set ys. x1 \<noteq> []"
using assms(4,6) dtree_to_list_x_in_dverts
by (smt (verit) case_prod_conv case_prod_eta fst_conv pair_imageI surj_pair)
ultimately show ?thesis using dtree_to_list_xs_not_equiv[of t1 t2] assms(1,2,5,6) by fast
qed
lemma merge_commute_aux1:
assumes "dlverts t1 \<inter> dlverts t2 = {}"
and "(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
and "[] \<notin> dverts t1"
and "[] \<notin> dverts t2"
and "xs = dtree_to_list (Node r {|(t1,e1)|})"
and "ys = dtree_to_list (Node r' {|(t2,e2)|})"
shows "Sorting_Algorithms.merge cmp' xs ys = Sorting_Algorithms.merge cmp' ys xs"
using merge_dtree1_not_equiv merge_comm_if_not_equiv assms by fast
lemma dtree_to_list_x1_list_disjoint:
"set x2 \<inter> dlverts t1 = {}
\<Longrightarrow> \<forall>x1 \<in> fst ` set (dtree_to_list (Node r {|(t1,e1)|})). set x1 \<inter> set x2 = {}"
using dtree_to_list_x_in_dlverts by fast
lemma dtree_to_list_e1_list_disjoint':
"set x2 \<inter> darcs t1 \<union> {e1} = {}
\<Longrightarrow> \<forall>x1 \<in> snd ` set (dtree_to_list (Node r {|(t1,e1)|})). x1 \<notin> set x2"
using dtree_to_list_e_in_darcs by blast
lemma dtree_to_list_e1_list_disjoint:
"e2 \<notin> darcs t1 \<union> {e1}
\<Longrightarrow> \<forall>x1 \<in> snd ` set (dtree_to_list (Node r {|(t1,e1)|})). x1 \<noteq> e2"
using dtree_to_list_e_in_darcs by fast
lemma dtree_to_list_xs_list_not_equiv:
assumes "(x1,e1) \<in> set (dtree_to_list (Node r {|(t1,e3)|}))"
and "x1 \<noteq> []"
and "\<forall>(v,e) \<in> set ys. set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e3}"
and "(x2,e2) \<in> set ys"
shows "compare cmp' (x1,e1) (x2,e2) \<noteq> Equiv"
proof -
have "set x1 \<inter> set x2 = {}" using dtree_to_list_x1_list_disjoint assms(1,3,4) by fastforce
moreover have "e1 \<noteq> e2" using dtree_to_list_e1_list_disjoint assms(1,3,4) by fastforce
ultimately show ?thesis using cmp_sets_not_dsjnt_if_equiv assms(2-4) by auto
qed
lemma merge_commute_aux2:
assumes "[] \<notin> dverts t1"
and "xs = dtree_to_list (Node r {|(t1,e1)|})"
and "\<forall>(v,e) \<in> set ys. set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1}"
shows "Sorting_Algorithms.merge cmp' xs ys = Sorting_Algorithms.merge cmp' ys xs"
proof -
have "\<forall>(x1,e1)\<in>set xs. x1 \<noteq> []"
using assms(1,2) dtree_to_list_x_in_dverts
by (smt (verit) case_prod_conv case_prod_eta fst_conv pair_imageI surj_pair)
then have "\<forall>(x1,e1)\<in>set xs. \<forall>(x2,e2)\<in>set ys. compare cmp' (x1,e1) (x2,e2) \<noteq> Equiv"
using assms(2,3) dtree_to_list_xs_list_not_equiv by force
then show ?thesis using merge_comm_if_not_equiv by fast
qed
lemma merge_inter_preserv':
assumes "f = (merge_f r xs)"
and "\<not>(\<forall>(v,_) \<in> set z. set v \<inter> dlverts t1 = {})"
shows "\<not>(\<forall>(v,_) \<in> set (f (t2,e2) z). set v \<inter> dlverts t1 = {})"
proof(cases "f (t2,e2) z = z")
case False
then have "f (t2,e2) z = Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) z"
by(simp add: assms(1)) meson
then show ?thesis using assms(2) set_merge by force
qed (simp add: assms(2))
lemma merge_inter_preserv:
assumes "f = (merge_f r xs)"
and "\<not>(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> e \<notin> darcs t1 \<union> {e1})"
shows "\<not>(\<forall>(v,e) \<in> set (f (t2,e2) z). set v \<inter> dlverts t1 = {} \<and> e \<notin> darcs t1 \<union> {e1})"
proof(cases "f (t2,e2) z = z")
case True
then show ?thesis using assms(2) by simp
next
case False
then have "f (t2,e2) z = Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) z"
by(simp add: assms(1)) meson
then show ?thesis
using assms(2) set_merge[of "dtree_to_list (Node r {|(t2,e2)|})"] by simp blast
qed
lemma merge_f_eq_z_if_inter':
"\<not>(\<forall>(v,_) \<in> set z. set v \<inter> dlverts t1 = {}) \<Longrightarrow> (merge_f r xs) (t1,e1) z = z"
by auto
lemma merge_f_eq_z_if_inter:
"\<not>(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> e \<notin> darcs t1 \<union> {e1})
\<Longrightarrow> (merge_f r xs) (t1,e1) z = z"
by auto
lemma merge_empty_inter_preserv_aux:
assumes "f = (merge_f r xs)"
and "(t2,e2) \<in> fset xs"
and "\<forall>(v,e) \<in> set z. set v \<inter> dlverts t2 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t2 \<union> {e2}"
and "list_dtree (Node r xs)"
and "(t1,e1) \<in> fset xs"
and "(t1,e1) \<noteq> (t2,e2)"
and "\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1}"
shows "\<forall>(v,e) \<in> set (f (t2,e2) z). set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1}"
proof -
have 0: "f (t2,e2) z = Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) z"
using assms(1-6) by simp
let ?ys = "dtree_to_list (Node r {|(t2,e2)|})"
interpret list_dtree "Node r xs" using assms(4) .
have "disjoint_dlverts xs" using wf_lverts by simp
then have "\<forall>v\<in>fst ` set ?ys. set v \<inter> dlverts t1 = {}"
using dtree_to_list_x1_disjoint assms(2,5,6) by fast
then have 1: "\<forall>v\<in>fst ` set (Sorting_Algorithms.merge cmp' ?ys z). set v \<inter> dlverts t1 = {}"
using assms(7) set_merge[of ?ys] by fastforce
have "disjoint_darcs xs" using disjoint_darcs_if_wf_xs[OF wf_arcs] .
then have 2: "(darcs t2 \<union> {e2}) \<inter> (darcs t1 \<union> {e1}) = {}" using assms(2,5,6) by fast
have "\<forall>e\<in>snd ` set ?ys. e \<notin> darcs t1 \<union> {e1}" using dtree_to_list_e_disjoint[OF 2] by blast
then have 2: "\<forall>e\<in>snd ` set (Sorting_Algorithms.merge cmp' ?ys z). e \<notin> darcs t1 \<union> {e1}"
using assms(7) set_merge[of ?ys] by fastforce
have "[] \<notin> dverts t2" using assms(2) empty_notin_wf_dlverts wf_lverts by fastforce
then have "\<forall>v\<in>fst ` set ?ys. v \<noteq> []" by (metis dtree_to_list_x_in_dverts)
then have "\<forall>v\<in>fst ` set (Sorting_Algorithms.merge cmp' ?ys z). v \<noteq> []"
using assms(7) set_merge[of ?ys] by fastforce
then show ?thesis using 0 1 2 by fastforce
qed
lemma merge_empty_inter_preserv:
assumes "f = (merge_f r xs)"
and "\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1}"
and "(t1,e1) \<in> fset xs"
and "(t1,e1) \<noteq> (t2,e2)"
shows "\<forall>(v,e) \<in> set (f (t2,e2) z). set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1}"
proof(cases "f (t2,e2) z = z")
case True
then show ?thesis using assms(2) by simp
next
case False
have "(t2,e2) \<in> fset xs" using False assms(1) by simp argo
moreover have "list_dtree (Node r xs)" using False assms(1) by simp argo
moreover have "\<forall>(v,e) \<in> set z. set v \<inter> dlverts t2 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t2 \<union> {e2}"
using False assms(1) by simp argo
ultimately show ?thesis using merge_empty_inter_preserv_aux assms by presburger
qed
lemma merge_commute_aux3:
assumes "f = (merge_f r xs)"
and "list_dtree (Node r xs)"
and "(t1,e1) \<noteq> (t2,e2)"
and "(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
and "(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t2 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t2 \<union> {e2})"
and "(t1,e1) \<in> fset xs"
and "(t2,e2) \<in> fset xs"
shows "(f (t2, e2) \<circ> f (t1, e1)) z = (f (t1, e1) \<circ> f (t2, e2)) z"
proof -
let ?merge = "Sorting_Algorithms.merge"
let ?xs = "dtree_to_list (Node r {|(t1, e1)|})"
let ?ys = "dtree_to_list (Node r {|(t2, e2)|})"
interpret list_dtree "Node r xs" using assms(2) .
have disj: "dlverts t1 \<inter> dlverts t2 = {}" "[] \<notin> dverts t1" "[] \<notin> dverts t2"
using assms(3,6,7) disjoint_dlverts_if_wf[OF wf_lverts] empty_notin_wf_dlverts[OF wf_lverts]
by fastforce+
have disj2: "(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
using assms(2,3,6,7) disjoint_darcs_if_wf_aux5[OF wf_arcs] by blast
have "f (t2, e2) z = Sorting_Algorithms.merge cmp' ?ys z" using assms(1,2,5,7) by simp
moreover have "\<forall>(v,e)\<in>set (f (t2,e2) z). set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1}"
using merge_empty_inter_preserv[OF assms(1)] assms(3,4,6) by simp
ultimately have 2: "(f (t1, e1) \<circ> f (t2, e2)) z = ?merge cmp' ?xs (?merge cmp' ?ys z)"
using assms(1-2,6) by auto
have "f (t1, e1) z = Sorting_Algorithms.merge cmp' ?xs z" using assms(1-2,4,6) by simp
moreover have "\<forall>(v,e)\<in>set (f (t1, e1) z). set v \<inter> dlverts t2 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t2 \<union> {e2}"
using merge_empty_inter_preserv[OF assms(1)] assms(3,5,7) by presburger
ultimately have 3: "(f (t2, e2) \<circ> f (t1,e1)) z = ?merge cmp' ?ys (?merge cmp' ?xs z)"
using assms(1-2,7) by simp
have "\<forall>x\<in>set ?xs. \<forall>y\<in>set ?ys. compare cmp' x y \<noteq> Equiv"
using merge_dtree1_not_equiv[OF disj(1) disj2] disj(2,3) by fast
then have "?merge cmp' ?xs (?merge cmp' ?ys z) = ?merge cmp' ?ys (?merge cmp' ?xs z)"
using merge_comp_commute by blast
then show ?thesis using 2 3 by simp
qed
lemma merge_commute_aux:
assumes "f = (merge_f r xs)"
shows "(f y \<circ> f x) z = (f x \<circ> f y) z"
proof -
obtain t1 e1 where y_def[simp]: "x = (t1, e1)" by fastforce
obtain t2 e2 where x_def[simp]: "y = (t2, e2)" by fastforce
show ?thesis
proof(cases "(t1,e1) \<in> fset xs \<and> (t2,e2) \<in> fset xs")
case True
then consider "list_dtree (Node r xs)" "(t1,e1) \<noteq> (t2,e2)"
"(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
"(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t2 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t2 \<union> {e2})"
| "(t1,e1) = (t2,e2)"
| "\<not>list_dtree (Node r xs)"
| "\<not>(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> e \<notin> darcs t1 \<union> {e1})"
| "\<not>(\<forall>(v,e) \<in> set z. set v \<inter> dlverts t2 = {} \<and> e \<notin> darcs t2 \<union> {e2})"
| "\<not>(\<forall>(v,_) \<in> set z. v \<noteq> [])"
by fast
then show ?thesis
proof(cases)
case 1
then show ?thesis using merge_commute_aux3[OF assms] True by simp
next
case 4
then have "f x z = z" by(auto simp: assms)
then have 0: "(f y \<circ> f x) z = f y z" by simp
have "\<not>(\<forall>(v,e) \<in> set (f y z). set v \<inter> dlverts t1 = {} \<and> e \<notin> darcs t1 \<union> {e1})"
using merge_inter_preserv[OF assms 4] by simp
then have "(f x \<circ> f y) z = f y z" using assms merge_f_eq_z_if_inter by auto
then show ?thesis using 0 by simp
next
case 5
then have "f y z = z" by(auto simp: assms)
then have 0: "(f x \<circ> f y) z = f x z" by simp
have "\<not>(\<forall>(v,e) \<in> set (f x z). set v \<inter> dlverts t2 = {} \<and> e \<notin> darcs t2 \<union> {e2})"
using merge_inter_preserv[OF assms 5] by simp
then have "(f y \<circ> f x) z = f x z" using assms merge_f_eq_z_if_inter by simp
then show ?thesis using 0 by simp
next
case 6
then have "(f x \<circ> f y) z = z" by(auto simp: assms)
also have "z = (f y \<circ> f x) z" using 6 by(auto simp: assms)
finally show ?thesis by simp
qed(auto simp: assms)
next
case False
then have "(\<forall>z. f x z = z) \<or> (\<forall>z. f y z = z)" by(auto simp: assms)
then show ?thesis by force
qed
qed
lemma merge_commute: "comp_fun_commute (merge_f r xs)"
using comp_fun_commute_def merge_commute_aux by blast
interpretation Comm: comp_fun_commute "merge_f r xs" by (rule merge_commute)
subsubsection \<open>Merging Preserves Arcs and Verts\<close>
lemma empty_list_valid_merge:
"(\<forall>(v,e) \<in> set []. set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
by simp
lemma disjoint_sets_sucs: "disjoint_sets (sucs t)"
using empty_notin_wf_dlverts list_dtree.wf_lverts list_dtree_rec dtree.collapse
disjoint_dlverts_if_wf[OF wf_lverts] disjoint_darcs_if_wf[OF wf_arcs] by blast
lemma empty_not_elem_subset:
"\<lbrakk>xs |\<subseteq>| ys; \<forall>(t,e) \<in> fset ys. [] \<notin> dverts t\<rbrakk> \<Longrightarrow> \<forall>(t,e) \<in> fset xs. [] \<notin> dverts t"
by (meson less_eq_fset.rep_eq subset_iff)
lemma disjoint_sets_subset:
assumes "xs |\<subseteq>| ys" and "disjoint_sets ys"
shows " disjoint_sets xs"
using disjoint_darcs_subset[OF assms(1)] disjoint_dlverts_subset[OF assms(1)]
empty_not_elem_subset[OF assms(1)] assms by fast
lemma merge_mdeg_le_1: "max_deg (merge t1) \<le> 1"
unfolding merge_def by (rule dtree_from_list_deg_le_1)
lemma merge_mdeg_le1_sub: "is_subtree t1 (merge t2) \<Longrightarrow> max_deg t1 \<le> 1"
using merge_mdeg_le_1 le_trans mdeg_ge_sub by fast
lemma merge_fcard_le1: "fcard (sucs (merge t1)) \<le> 1"
unfolding merge_def by (rule dtree_from_list_fcard_le1)
lemma merge_fcard_le1_sub: "is_subtree t1 (merge t2) \<Longrightarrow> fcard (sucs t1) \<le> 1"
using merge_mdeg_le1_sub mdeg_ge_fcard[of "sucs t1" "root t1"] by force
lemma merge_f_alt:
assumes "P = (\<lambda>xs. list_dtree (Node r xs))"
and "Q = (\<lambda>(t,e) b. (\<forall>(v,e') \<in> set b. set v \<inter> dlverts t = {} \<and> v\<noteq>[] \<and> e' \<notin> darcs t \<union> {e}))"
and "R = (\<lambda>(t,e) b. Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t,e)|})) b)"
shows "merge_f r xs = (\<lambda>a b. if a \<notin> fset xs \<or> \<not> Q a b \<or> \<not> P xs then b else R a b)"
using assms by force
lemma merge_f_alt_commute:
assumes "P = (\<lambda>xs. list_dtree (Node r xs))"
and "Q = (\<lambda>(t,e) b. (\<forall>(v,e') \<in> set b. set v \<inter> dlverts t = {} \<and> v \<noteq> [] \<and> e' \<notin> darcs t \<union> {e}))"
and "R = (\<lambda>(t,e) b. Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t,e)|})) b)"
shows "comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not> Q a b \<or> \<not> P xs then b else R a b)"
proof -
have "comp_fun_commute (merge_f r xs)" using merge_commute by fast
then show ?thesis using merge_f_alt[OF assms] by simp
qed
lemma merge_ffold_supset:
assumes "xs |\<subseteq>| ys" and "list_dtree (Node r ys)"
shows "ffold (merge_f r ys) acc xs = ffold (merge_f r xs) acc xs"
proof -
let ?P = "\<lambda>xs. list_dtree (Node r xs)"
let ?Q = "\<lambda>(t,e) b. (\<forall>(v,e') \<in> set b. set v \<inter> dlverts t = {} \<and> v \<noteq> [] \<and> e' \<notin> darcs t \<union> {e})"
let ?R = "\<lambda>(t,e) b. Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t,e)|})) b"
have 0: "\<And>xs. comp_fun_commute (\<lambda>a b. if a \<notin> fset xs \<or> \<not> ?Q a b \<or> \<not> ?P xs then b else ?R a b)"
using merge_f_alt_commute by blast
have "ffold (\<lambda>a b. if a \<notin> fset ys \<or> \<not> ?Q a b \<or> \<not> ?P ys then b else ?R a b) acc xs
= ffold (\<lambda>a b. if a \<notin> fset xs \<or> \<not> ?Q a b \<or> \<not> ?P xs then b else ?R a b) acc xs"
using ffold_commute_supset[OF assms(1), of ?P ?Q ?R, OF assms(2) list_dtree_subset 0] by auto
then show ?thesis using merge_f_alt by presburger
qed
lemma merge_f_merge_if_not_snd:
"merge_f r xs (t1,e1) z \<noteq> z \<Longrightarrow>
merge_f r xs (t1,e1) z = Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t1,e1)|})) z"
by(simp) meson
lemma merge_f_merge_if_conds:
"\<lbrakk>list_dtree (Node r xs); \<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1};
(t1,e1) \<in> fset xs\<rbrakk>
\<Longrightarrow> merge_f r xs (t1,e1) z = Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t1,e1)|})) z"
by force
lemma merge_f_merge_if_conds_empty:
"\<lbrakk>list_dtree (Node r xs); (t1,e1) \<in> fset xs\<rbrakk>
\<Longrightarrow> merge_f r xs (t1,e1) []
= Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|(t1,e1)|})) []"
using merge_f_merge_if_conds by simp
lemma merge_ffold_empty_inter_preserv:
"\<lbrakk>list_dtree (Node r ys); xs |\<subseteq>| ys;
\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1};
(t1,e1) \<in> fset ys; (t1,e1) \<notin> fset xs; (v,e) \<in> set (ffold (merge_f r xs) z xs)\<rbrakk>
\<Longrightarrow> set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1}"
proof(induction xs)
case (insert x xs)
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
let ?merge = "Sorting_Algorithms.merge"
interpret list_dtree "Node r ys" using insert.prems(1) .
have 0: "list_dtree (Node r (finsert x xs))" using list_dtree_subset insert.prems(1,2) by blast
show ?case
proof(cases "ffold ?f z (finsert x xs) = ffold ?f' z xs")
case True
then have "(v,e) \<in> set (ffold ?f' z xs)" using insert.prems(6) by argo
then show ?thesis using insert.IH insert.prems by force
next
case not_right: False
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
show ?thesis
proof(cases "(v,e) \<in> set (dtree_to_list (Node r {|(t2,e2)|}))")
case True
have uneq: "(t2,e2) \<noteq> (t1,e1)" using insert.prems(5) t2_def by fastforce
- moreover have 1: "(t2,e2) \<in> fset ys" using insert.prems(2) notin_fset by fastforce
+ moreover have 1: "(t2,e2) \<in> fset ys" using insert.prems(2) by fastforce
ultimately have "dlverts t1 \<inter> dlverts t2 = {}" using insert.prems(4) wf_lverts by fastforce
then have 2: "\<forall>x1\<in>fst ` set (dtree_to_list (Node r {|(t2, e2)|})). set x1 \<inter> dlverts t1 = {}"
using dtree_to_list_x1_disjoint by fast
have "(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
using insert.prems(4) uneq 1 disjoint_darcs_if_wf_aux5 wf_arcs by fast
then have 3: "\<forall>e\<in>snd ` set (dtree_to_list (Node r {|(t2, e2)|})). e \<notin> darcs t1 \<union> {e1}"
using dtree_to_list_e_disjoint by fast
have "[] \<notin> dverts t2" using 1 wf_lverts empty_notin_wf_dlverts by auto
then have "\<forall>x1\<in>fst ` set (dtree_to_list (Node r {|(t2, e2)|})). x1 \<noteq> []"
using 1 dtree_to_list_x_in_dverts by metis
then show ?thesis using True 2 3 by fastforce
next
case False
have "xs |\<subseteq>| finsert x xs" by blast
then have f_xs: "ffold ?f z xs = ffold ?f' z xs"
using merge_ffold_supset 0 by presburger
have "ffold ?f z (finsert x xs) = ?f x (ffold ?f z xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
then have 0: "ffold ?f z (finsert x xs) = ?f x (ffold ?f' z xs)" using f_xs by argo
then have "?f x (ffold ?f' z xs) \<noteq> ffold ?f' z xs" using not_right by argo
then have "?f (t2,e2) (ffold ?f' z xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using merge_f_merge_if_not_snd t2_def by blast
then have "ffold ?f z (finsert x xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using 0 t2_def by argo
then have "(v,e) \<in> set (?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs))"
using insert.prems(6) by argo
then have "(v,e) \<in> set (ffold ?f' z xs)" using set_merge False by fast
then show ?thesis using insert.IH insert.prems by force
qed
qed
-qed(auto)
+qed (auto simp: ffold.rep_eq)
lemma merge_ffold_empty_inter_preserv':
"\<lbrakk>list_dtree (Node r (finsert x xs));
\<forall>(v,e) \<in> set z. set v \<inter> dlverts t1 = {} \<and> v\<noteq>[] \<and> e \<notin> darcs t1 \<union> {e1};
(t1,e1) \<in> fset (finsert x xs); (t1,e1) \<notin> fset xs; (v,e) \<in> set (ffold (merge_f r xs) z xs)\<rbrakk>
\<Longrightarrow> set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1}"
using merge_ffold_empty_inter_preserv[of r "finsert x xs" xs z t1 e1 v e] by fast
lemma merge_ffold_set_sub_union:
"list_dtree (Node r xs)
\<Longrightarrow> set (ffold (merge_f r xs) [] xs) \<subseteq> (\<Union>x\<in>fset xs. set (dtree_to_list (Node r {|x|})))"
proof(induction xs)
case (insert x xs)
obtain t1 e1 where t1_def[simp]: "x = (t1,e1)" by fastforce
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
have "(t1, e1) \<in> fset (finsert x xs)" by simp
- moreover have "(t1, e1) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t1, e1) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold ?f' [] xs). set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
using merge_ffold_empty_inter_preserv'[OF insert.prems empty_list_valid_merge] by blast
have 0: "list_dtree (Node r xs)" using list_dtree_subset insert.prems by blast
have "ffold ?f [] (finsert x xs) = ?f x (ffold ?f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
also have "\<dots> = ?f x (ffold ?f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems by fastforce
finally have "ffold ?f [] (finsert x xs)
= Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs)"
using merge_f_merge_if_conds[OF insert.prems xs_val] by simp
then have "set (ffold ?f [] (finsert x xs))
= set (Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs))"
by argo
then have "set (ffold ?f [] (finsert x xs))
= (set (dtree_to_list (Node r {|x|})) \<union> set (ffold ?f' [] xs))" using set_merge by fast
then show ?case using 0 insert.IH insert.prems by auto
-qed (simp)
+qed (simp add: ffold.rep_eq)
lemma merge_ffold_nempty:
"\<lbrakk>list_dtree (Node r xs); xs \<noteq> {||}\<rbrakk> \<Longrightarrow> ffold (merge_f r xs) [] xs \<noteq> []"
proof(induction xs)
case (insert x xs)
define f where "f = merge_f r (finsert x xs)"
define f' where "f' = merge_f r xs"
let ?merge = "Sorting_Algorithms.merge cmp'"
have 0: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(1) by blast
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
have "(t2, e2) \<in> fset (finsert x xs)" by simp
- moreover have "(t2, e2) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t2, e2) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold f' [] xs). set v \<inter> dlverts t2 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t2 \<union> {e2})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(1) empty_list_valid_merge] f'_def
by blast
have "ffold f [] (finsert x xs) = f x (ffold f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] f_def by blast
also have "\<dots> = f x (ffold f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(1) f_def f'_def by fastforce
finally have "ffold f [] (finsert x xs) = ?merge (dtree_to_list (Node r {|x|})) (ffold f' [] xs)"
using xs_val insert.prems f_def by simp
then have merge: "ffold f [] (finsert x xs)
= ?merge (dtree_to_list (Node r {|(t2,e2)|})) (ffold f'[] xs)"
using t2_def by blast
then show ?case
using input_empty_if_merge_empty[of cmp' "dtree_to_list (Node r {|(t2,e2)|})"] f_def by auto
qed(simp)
lemma merge_f_ndisjoint_sets_aux:
"\<not>disjoint_sets xs
\<Longrightarrow> \<not>((t,e) \<in> fset xs \<and> disjoint_sets xs \<and> (\<forall>(v,_) \<in> set b. set v \<inter> dlverts t = {} \<and> v \<noteq> []))"
by blast
lemma merge_f_not_list_dtree: "\<not>list_dtree (Node r xs) \<Longrightarrow> (merge_f r xs) a b = b"
using merge_f_alt by simp
lemma merge_ffold_empty_if_nwf: "\<not>list_dtree (Node r ys) \<Longrightarrow> ffold (merge_f r ys) [] xs = []"
proof(induction xs)
case (insert x xs)
define f where "f = merge_f r ys"
let ?f = "merge_f r ys"
let ?merge = "Sorting_Algorithms.merge cmp'"
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
have "ffold f [] (finsert x xs) = ?f x (ffold f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] f_def by blast
then have "ffold f [] (finsert x xs) = ffold f [] xs"
using insert.prems merge_f_not_list_dtree by force
then show ?case using insert f_def by argo
-qed(simp)
+qed (simp add: ffold.rep_eq)
lemma merge_empty_if_nwf: "\<not>list_dtree (Node r xs) \<Longrightarrow> merge (Node r xs) = Node r {||}"
unfolding merge_def using merge_ffold_empty_if_nwf by simp
lemma merge_empty_if_nwf_sucs: "\<not>list_dtree t1 \<Longrightarrow> merge t1 = Node (root t1) {||}"
using merge_empty_if_nwf[of "root t1" "sucs t1"] by simp
lemma merge_empty: "merge (Node r {||}) = Node r {||}"
- unfolding merge_def by simp
-
-lemma merge_empty_sucs: "sucs t1 = {||} \<Longrightarrow> merge t1 = Node (root t1) {||}"
- unfolding merge_def by simp
+proof -
+ have "comp_fun_commute (\<lambda>(t, e) b. b)"
+ by (simp add: comp_fun_commute_const cond_case_prod_eta)
+ hence "dtree_from_list r (ffold (\<lambda>(t, e) b. b) [] {||}) = Node r {||}"
+ using comp_fun_commute.ffold_empty
+ by (smt (verit, best) dtree_from_list.simps(1))
+ thus ?thesis
+ unfolding merge_def by simp
+qed
+
+lemma merge_empty_sucs:
+ assumes "sucs t1 = {||}"
+ shows "merge t1 = Node (root t1) {||}"
+proof -
+ have "dtree_from_list (dtree.root t1) (ffold (\<lambda>(t, e) b. b) [] {||}) = Node (dtree.root t1) {||}"
+ by (simp add: ffold.rep_eq)
+ with assms show ?thesis
+ unfolding merge_def by simp
+qed
lemma merge_singleton_sucs:
assumes "list_dtree (Node (root t1) (sucs t1))" and "sucs t1 \<noteq> {||}"
shows "\<exists>t e. merge t1 = Node (root t1) {|(t,e)|}"
unfolding merge_def using merge_ffold_nempty[OF assms] dtree_from_list_singleton by fast
lemma merge_singleton:
assumes "list_dtree (Node r xs)" and "xs \<noteq> {||}"
shows "\<exists>t e. merge (Node r xs) = Node r {|(t,e)|}"
unfolding merge_def dtree.sel(1) using merge_ffold_nempty[OF assms] dtree_from_list_singleton
by fastforce
lemma merge_cases: "\<exists>t e. merge (Node r xs) = Node r {|(t,e)|} \<or> merge (Node r xs) = Node r {||}"
using merge_singleton merge_empty_if_nwf merge_empty by blast
lemma merge_cases_sucs:
"\<exists>t e. merge t1 = Node (root t1) {|(t,e)|} \<or> merge t1 = Node (root t1) {||}"
using merge_singleton_sucs[of t1] merge_empty_if_nwf_sucs merge_empty_sucs by auto
lemma merge_single_root:
"(t2,e2) \<in> fset (sucs (merge (Node r xs))) \<Longrightarrow> merge (Node r xs) = Node r {|(t2,e2)|}"
using merge_cases[of r xs] by fastforce
lemma merge_single_root_sucs:
"(t2,e2) \<in> fset (sucs (merge t1)) \<Longrightarrow> merge t1 = Node (root t1) {|(t2,e2)|}"
using merge_cases_sucs[of t1] by auto
lemma merge_single_root1:
"t2 \<in> fst ` fset (sucs (merge (Node r xs))) \<Longrightarrow> \<exists>e2. merge (Node r xs) = Node r {|(t2,e2)|}"
using merge_single_root by fastforce
lemma merge_single_root1_sucs:
"t2 \<in> fst ` fset (sucs (merge t1)) \<Longrightarrow> \<exists>e2. merge t1 = Node (root t1) {|(t2,e2)|}"
using merge_single_root_sucs by fastforce
lemma merge_nempty_sucs: "\<lbrakk>list_dtree t1; sucs t1 \<noteq> {||}\<rbrakk> \<Longrightarrow> sucs (merge t1) \<noteq> {||}"
using merge_singleton_sucs by fastforce
lemma merge_nempty: "\<lbrakk>list_dtree (Node r xs); xs \<noteq> {||}\<rbrakk> \<Longrightarrow> sucs (merge (Node r xs)) \<noteq> {||}"
using merge_singleton by fastforce
lemma merge_xs: "merge (Node r xs) = dtree_from_list r (ffold (merge_f r xs) [] xs)"
unfolding merge_def dtree.sel(1) dtree.sel(2) by blast
lemma merge_root_eq[simp]: "root (merge t1) = root t1"
unfolding merge_def by simp
lemma merge_ffold_fsts_in_childverts:
"\<lbrakk>list_dtree (Node r xs); y \<in> fst ` set (ffold (merge_f r xs) [] xs)\<rbrakk>
\<Longrightarrow> \<exists>t1 \<in> fst ` fset xs. y \<in> dverts t1"
proof(induction xs)
case (insert x xs)
obtain t1 e1 where t1_def[simp]: "x = (t1,e1)" by fastforce
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
have "(t1, e1) \<in> fset (finsert x xs)" by simp
- moreover have "(t1, e1) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t1, e1) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold ?f' [] xs). set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(1) empty_list_valid_merge] by blast
have 0: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(1) by blast
then show ?case
proof(cases "y \<in> fst ` set (ffold (merge_f r xs) [] xs)")
case True
then show ?thesis using insert.IH[OF 0] by simp
next
case False
have "ffold ?f [] (finsert x xs) = ?f x (ffold ?f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
also have "\<dots> = ?f x (ffold ?f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(1) by fastforce
finally have "ffold ?f [] (finsert x xs)
= Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs)"
using xs_val insert.prems by simp
then have "set (ffold ?f [] (finsert x xs))
= set (Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs))"
by argo
then have "set (ffold ?f [] (finsert x xs))
= (set (dtree_to_list (Node r {|x|})) \<union> set (ffold ?f' [] xs))"
using set_merge by fast
then have "y \<in> fst ` set (dtree_to_list (Node r {|x|}))" using False insert.prems by fast
then show ?thesis by (simp add: dtree_to_list_x_in_dverts)
qed
-qed (simp)
+qed (simp add: ffold.rep_eq)
lemma verts_child_if_merge_child:
assumes "t1 \<in> fst ` fset (sucs (merge t0))" and "x \<in> dverts t1"
shows "\<exists>t2 \<in> fst ` fset (sucs t0). x \<in> dverts t2"
proof -
have 0: "list_dtree t0" using assms(1) merge_empty_if_nwf_sucs by fastforce
have "merge t0 \<noteq> Node (root t0) {||}" using assms(1) by force
then obtain e1 where e1_def: "merge t0 = Node (root t0) {|(t1,e1)|}"
using assms(1) merge_single_root1_sucs by blast
then obtain ys where ys_def:
"(root t1, e1) # ys = ffold (merge_f (root t0) (sucs t0)) [] (sucs t0)"
unfolding merge_def by (metis (no_types, lifting) dtree_to_list.simps(1) dtree_to_from_list_id)
then have "merge t0 = dtree_from_list (root t0) ((root t1, e1) # ys)" unfolding merge_def by simp
then have "t1 = dtree_from_list (root t1) ys" using e1_def by simp
then have "dverts t1 = (fst ` set ((root t1, e1) # ys))"
using dtree_from_list_eq_dverts[of "root t1" ys] by simp
then have "x \<in> fst ` set (ffold (merge_f (root t0) (sucs t0)) [] (sucs t0))"
using assms(2) ys_def by simp
then show ?thesis using merge_ffold_fsts_in_childverts[of "root t0"] 0 by simp
qed
lemma sucs_dverts_eq_dtree_list:
assumes "(t1,e1) \<in> fset (sucs t)" and "max_deg t1 \<le> 1"
shows "dverts (Node (root t) {|(t1,e1)|}) - {root t}
= fst ` set (dtree_to_list (Node (root t) {|(t1,e1)|}))"
proof -
- have "{|(t1,e1)|} |\<subseteq>| sucs t" using assms(1) notin_fset by fast
+ have "{|(t1,e1)|} |\<subseteq>| sucs t" using assms(1) by fast
then have wf: "wf_dverts (Node (root t) {|(t1,e1)|})"
using wf_verts wf_dverts_sub by (metis dtree.exhaust_sel)
have "\<forall>(t1,e1) \<in> fset (sucs t) . fcard {|(t1,e1)|} = 1" using fcard_single_1 by fast
moreover have "max_deg (Node (root t) {|(t1,e1)|}) = max (max_deg t1) (fcard {|(t1,e1)|})"
using mdeg_singleton by fast
ultimately have "max_deg (Node (root t) {|(t1,e1)|}) \<le> 1"
using assms by fastforce
then show ?thesis using dtree_to_list_eq_dverts[OF wf] by simp
qed
lemma merge_ffold_set_eq_union:
"list_dtree (Node r xs)
\<Longrightarrow> set (ffold (merge_f r xs) [] xs) = (\<Union>x\<in>fset xs. set (dtree_to_list (Node r {|x|})))"
proof(induction xs)
case (insert x xs)
obtain t1 e1 where t1_def[simp]: "x = (t1,e1)" by fastforce
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
have "(t1, e1) \<in> fset (finsert x xs)" by simp
- moreover have "(t1, e1) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t1, e1) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold ?f' [] xs). set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(1) empty_list_valid_merge] by blast
have 1: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(1) by blast
have "ffold ?f [] (finsert x xs) = ?f x (ffold ?f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
also have "\<dots> = ?f x (ffold ?f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(1) by fastforce
finally have "ffold ?f [] (finsert x xs)
= Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs)"
using xs_val insert.prems by simp
then have "set (ffold ?f [] (finsert x xs))
= set (Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs))"
by argo
then have "set (ffold ?f [] (finsert x xs))
= (set (dtree_to_list (Node r {|x|})) \<union> set (ffold ?f' [] xs))" using set_merge by fast
then show ?case using 1 insert.IH by simp
-qed (simp)
+qed (simp add: ffold.rep_eq)
lemma sucs_dverts_no_root:
"(t1,e1) \<in> fset (sucs t) \<Longrightarrow> dverts (Node (root t) {|(t1,e1)|}) - {root t} = dverts t1"
using wf_verts wf_dverts'.simps unfolding wf_dverts_iff_dverts' by fastforce
lemma dverts_merge_sub:
assumes "\<forall>t \<in> fst ` fset (sucs t0). max_deg t \<le> 1"
shows "dverts (merge t0) \<subseteq> dverts t0"
proof
fix x
assume asm: "x \<in> dverts (merge t0)"
show "x \<in> dverts t0"
proof(cases "x = root (merge t0)")
case True
then show ?thesis by (simp add: dtree.set_sel(1))
next
case False
then obtain t1 e1 where t1_def: "merge t0 = Node (root t0) ({|(t1,e1)|})"
using merge_cases_sucs asm by fastforce
then have 0: "list_dtree (Node (root t0) (sucs t0))"
using merge_empty_if_nwf_sucs by fastforce
have "x \<in> fst ` set (ffold (merge_f (root t0) (sucs t0)) [] (sucs t0))"
using t1_def unfolding merge_def using False asm t1_def
dtree_from_list_eq_dverts[of "root t0" "ffold (merge_f (root t0) (sucs t0)) [] (sucs t0)"]
by auto
then obtain t2 e2 where t2_def:
"(t2,e2) \<in> fset (sucs t0)" "x \<in> fst ` set (dtree_to_list (Node (root t0) {|(t2,e2)|}))"
using merge_ffold_set_sub_union[OF 0] by fast
then have "x \<in> dverts t2" by (simp add: dtree_to_list_x_in_dverts)
then show ?thesis using t2_def(1) dtree.set_sel(2) by fastforce
qed
qed
lemma dverts_merge_eq[simp]:
assumes "\<forall>t \<in> fst ` fset (sucs t). max_deg t \<le> 1"
shows "dverts (merge t) = dverts t"
proof -
have "\<forall>(t1,e1) \<in> fset (sucs t). dverts (Node (root t) {|(t1,e1)|}) - {root t}
= fst ` set (dtree_to_list (Node (root t) {|(t1,e1)|}))"
using sucs_dverts_eq_dtree_list assms
by (smt (verit, ccfv_threshold) case_prodI2 fst_conv image_iff)
then have "\<forall>(t1,e1) \<in> fset (sucs t). dverts t1
= fst ` set (dtree_to_list (Node (root t) {|(t1,e1)|}))"
by (metis (mono_tags, lifting) sucs_dverts_no_root case_prodD case_prodI2)
then have "(\<Union>x\<in>fset (sucs t). \<Union> (dverts ` Basic_BNFs.fsts x))
= (\<Union>x\<in>fset (sucs t). fst ` set (dtree_to_list (Node (root t) {|x|})))"
by force
then have "dverts t
= insert (root t) (\<Union>x\<in>fset (sucs t). fst ` set (dtree_to_list (Node (root t) {|x|})))"
using dtree.simps(6)[of "root t" "sucs t"] by auto
also have "\<dots> = insert (root t) (fst ` set (ffold (merge_f (root t) (sucs t)) [] (sucs t)))"
using merge_ffold_set_eq_union[of "root t" "sucs t"] list_dtree_axioms by auto
also have "\<dots> = dverts (dtree_from_list (root t) (ffold (merge_f (root t) (sucs t)) [] (sucs t)))"
using dtree_from_list_eq_dverts[of "root t"] by blast
finally show ?thesis unfolding merge_def by blast
qed
lemma dlverts_merge_eq[simp]:
assumes "\<forall>t \<in> fst ` fset (sucs t). max_deg t \<le> 1"
shows "dlverts (merge t) = dlverts t"
using dverts_merge_eq[OF assms] by (simp add: dlverts_eq_dverts_union)
lemma sucs_darcs_eq_dtree_list:
assumes "(t1,e1) \<in> fset (sucs t)" and "max_deg t1 \<le> 1"
shows "darcs (Node (root t) {|(t1,e1)|}) = snd ` set (dtree_to_list (Node (root t) {|(t1,e1)|}))"
proof -
have "\<forall>(t1,e1) \<in> fset (sucs t) . fcard {|(t1,e1)|} = 1" using fcard_single_1 by fast
moreover have "max_deg (Node (root t) {|(t1,e1)|}) = max (max_deg t1) (fcard {|(t1,e1)|})"
using mdeg_singleton by fast
ultimately have "max_deg (Node (root t) {|(t1,e1)|}) \<le> 1"
using assms by fastforce
then show ?thesis using dtree_to_list_eq_darcs by blast
qed
lemma darcs_merge_eq[simp]:
assumes "\<forall>t \<in> fst ` fset (sucs t). max_deg t \<le> 1"
shows "darcs (merge t) = darcs t"
proof -
have 0: "list_dtree (Node (root t) (sucs t))" using list_dtree_axioms by simp
have "\<forall>(t1,e1) \<in> fset (sucs t). darcs (Node (root t) {|(t1,e1)|})
= snd ` set (dtree_to_list (Node (root t) {|(t1,e1)|}))"
using sucs_darcs_eq_dtree_list assms
by (smt (verit, ccfv_threshold) case_prodI2 fst_conv image_iff)
then have "\<forall>(t1,e1) \<in> fset (sucs t). darcs t1 \<union> {e1}
= snd ` set (dtree_to_list (Node (root t) {|(t1,e1)|}))"
by simp
moreover have "darcs t = (\<Union>(t1,e1)\<in>fset (sucs t). darcs t1 \<union> {e1})"
using dtree.simps(7)[of "root t" "sucs t"] by force
ultimately have "darcs t
= (\<Union>(t1,e1)\<in>fset (sucs t). snd ` set (dtree_to_list (Node (root t) {|(t1,e1)|})))"
by (smt (verit, best) Sup.SUP_cong case_prodE case_prod_conv)
also have "\<dots> = (snd ` set (ffold (merge_f (root t) (sucs t)) [] (sucs t)))"
using merge_ffold_set_eq_union[OF 0] by blast
also have "\<dots> = darcs (dtree_from_list (root t) (ffold (merge_f (root t) (sucs t)) [] (sucs t)))"
using dtree_from_list_eq_darcs[of "root t"] by fast
finally show ?thesis unfolding merge_def by blast
qed
subsubsection \<open>Merging Preserves Well-Formedness\<close>
lemma dtree_to_list_x_in_darcs:
"x \<in> snd ` set (dtree_to_list (Node r {|(t1,e1)|})) \<Longrightarrow> x \<in> (darcs t1 \<union> {e1})"
using dtree_to_list_sub_darcs by fastforce
lemma dtree_to_list_snds_disjoint:
"(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}
\<Longrightarrow> snd ` set (dtree_to_list (Node r {|(t1,e1)|})) \<inter> (darcs t2 \<union> {e2}) = {}"
using dtree_to_list_x_in_darcs by fast
lemma dtree_to_list_snds_disjoint2:
"(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}
\<Longrightarrow> snd ` set (dtree_to_list (Node r {|(t1,e1)|}))
\<inter> snd ` set (dtree_to_list (Node r {|(t2,e2)|})) = {}"
using disjoint_iff dtree_to_list_x_in_darcs by metis
lemma merge_ffold_arc_inter_preserv:
"\<lbrakk>list_dtree (Node r ys); xs |\<subseteq>| ys; (darcs t1 \<union> {e1}) \<inter> (snd ` set z) = {};
(t1,e1) \<in> fset ys; (t1,e1) \<notin> fset xs\<rbrakk>
\<Longrightarrow> (darcs t1 \<union> {e1}) \<inter> (snd ` set (ffold (merge_f r xs) z xs)) = {}"
proof(induction xs)
case (insert x xs)
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
let ?merge = "Sorting_Algorithms.merge"
show ?case
proof(cases "ffold ?f z (finsert x xs) = ffold ?f' z xs")
case True
then show ?thesis using insert.IH insert.prems by auto
next
case False
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
have 0: "list_dtree (Node r (finsert x xs))" using list_dtree_subset insert.prems(1,2) by blast
have "(t2,e2) \<noteq> (t1,e1)" using insert.prems(5) t2_def by fastforce
- moreover have "(t2,e2) \<in> fset ys" using insert.prems(2) notin_fset by fastforce
+ moreover have "(t2,e2) \<in> fset ys" using insert.prems(2) by fastforce
moreover have "disjoint_darcs ys"
using disjoint_darcs_if_wf[OF list_dtree.wf_arcs [OF insert.prems(1)]] by simp
ultimately have "(darcs t1 \<union> {e1}) \<inter> (darcs t2 \<union> {e2}) = {}"
using insert.prems(4) by fast
then have 1: "(darcs t1 \<union> {e1}) \<inter> snd ` set (dtree_to_list (Node r {|(t2, e2)|})) = {}"
using dtree_to_list_snds_disjoint by fast
have 2: "(darcs t1 \<union> {e1}) \<inter> snd ` set (ffold ?f' z xs) = {}"
using insert.IH insert.prems by simp
have "xs |\<subseteq>| finsert x xs" by blast
then have f_xs: "ffold ?f z xs = ffold ?f' z xs"
using merge_ffold_supset 0 by presburger
have "ffold ?f z (finsert x xs) = ?f x (ffold ?f z xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
then have 0: "ffold ?f z (finsert x xs) = ?f x (ffold ?f' z xs)" using f_xs by argo
then have "?f x (ffold ?f' z xs) \<noteq> ffold ?f' z xs" using False by argo
then have "?f (t2,e2) (ffold ?f' z xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using merge_f_merge_if_not_snd t2_def by blast
then have "ffold ?f z (finsert x xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using 0 t2_def by argo
then have "set (ffold ?f z (finsert x xs))
= set (dtree_to_list (Node r {|(t2,e2)|})) \<union> set (ffold ?f' z xs)"
using set_merge[of "dtree_to_list (Node r {|(t2,e2)|})"] by presburger
then show ?thesis using 1 2 by fast
qed
-qed (auto)
+qed (auto simp: ffold.rep_eq)
lemma merge_ffold_wf_list_arcs:
"\<lbrakk>\<And>x. x \<in> fset xs \<Longrightarrow> wf_darcs (Node r {|x|}); list_dtree (Node r xs)\<rbrakk>
\<Longrightarrow> wf_list_arcs (ffold (merge_f r xs) [] xs)"
proof(induction xs)
case (insert x xs)
obtain t1 e1 where t1_def[simp]: "x = (t1,e1)" by fastforce
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
have 0: "(t1, e1) \<in> fset (finsert x xs)" by simp
- moreover have t1_not_xs: "(t1, e1) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have t1_not_xs: "(t1, e1) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold ?f' [] xs). set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(2) empty_list_valid_merge] by blast
have 1: "wf_list_arcs (dtree_to_list (Node r {|x|}))"
using insert.prems(1) 0 t1_def wf_list_arcs_if_wf_darcs by fast
have "list_dtree (Node r xs)" using list_dtree_subset insert.prems(2) by blast
then have 2: "wf_list_arcs (ffold ?f' [] xs)" using insert.IH insert.prems by auto
have "darcs (Node r {|x|}) \<inter> snd ` set (ffold ?f' [] xs) = {}"
using merge_ffold_arc_inter_preserv[OF insert.prems(2), of xs t1 e1 "[]"] t1_not_xs by auto
then have 3: "snd ` set (dtree_to_list (Node r {|x|})) \<inter> snd ` set (ffold ?f' [] xs) = {}"
using dtree_to_list_sub_darcs by fast
have "ffold ?f [] (finsert x xs) = ?f x (ffold ?f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
also have "\<dots> = ?f x (ffold ?f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(2) by fastforce
finally have "ffold ?f [] (finsert x xs)
= Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs)"
using xs_val insert.prems by simp
then show ?case using wf_list_arcs_merge[OF 1 2 3] by presburger
-qed (simp)
+qed (simp add: ffold.rep_eq)
lemma merge_wf_darcs: "wf_darcs (merge t)"
proof -
have "wf_list_arcs (ffold (merge_f (root t) (sucs t)) [] (sucs t))"
using merge_ffold_wf_list_arcs[OF wf_darcs_sucs[OF wf_arcs]] list_dtree_axioms by simp
then show ?thesis using wf_darcs_iff_wf_list_arcs merge_def by fastforce
qed
lemma merge_ffold_wf_list_lverts:
"\<lbrakk>\<And>x. x \<in> fset xs \<Longrightarrow> wf_dlverts (Node r {|x|}); list_dtree (Node r xs)\<rbrakk>
\<Longrightarrow> wf_list_lverts (ffold (merge_f r xs) [] xs)"
proof(induction xs)
case (insert x xs)
obtain t1 e1 where t1_def[simp]: "x = (t1,e1)" by fastforce
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
have 0: "(t1, e1) \<in> fset (finsert x xs)" by simp
- moreover have "(t1, e1) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t1, e1) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold ?f' [] xs). set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(2) empty_list_valid_merge] by blast
have 1: "wf_list_lverts (dtree_to_list (Node r {|x|}))"
using insert.prems(1) 0 t1_def wf_list_lverts_if_wf_dlverts by fast
have "list_dtree (Node r xs)" using list_dtree_subset insert.prems(2) by blast
then have 2: "wf_list_lverts (ffold ?f' [] xs)" using insert.IH insert.prems by auto
have "\<forall>v2\<in>fst ` set (ffold ?f' [] xs). set v2 \<inter> dlverts t1 = {}"
using xs_val by fastforce
then have 3: "\<forall>v1\<in>fst ` set (dtree_to_list (Node r {|x|})). \<forall>v2\<in>fst ` set (ffold ?f' [] xs).
set v1 \<inter> set v2 = {}"
using dtree_to_list_x1_list_disjoint t1_def by fast
have "ffold ?f [] (finsert x xs) = ?f x (ffold ?f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
also have "\<dots> = ?f x (ffold ?f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(2) by fastforce
finally have "ffold ?f [] (finsert x xs)
= Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold ?f' [] xs)"
using xs_val insert.prems by simp
then show ?case using wf_list_lverts_merge[OF 1 2 3] by presburger
-qed (simp)
+qed (simp add: ffold.rep_eq)
lemma merge_ffold_root_inter_preserv:
"\<lbrakk>list_dtree (Node r xs); \<forall>t1 \<in> fst ` fset xs. set r' \<inter> dlverts t1 = {};
\<forall>v1 \<in> fst ` set z. set r' \<inter> set v1 = {}; (v,e) \<in> set (ffold (merge_f r xs) z xs)\<rbrakk>
\<Longrightarrow> set r' \<inter> set v = {}"
proof(induction xs)
case (insert x xs)
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
let ?merge = "Sorting_Algorithms.merge"
have 0: "list_dtree (Node r xs)" using insert.prems(1) list_dtree_subset by blast
show ?case
proof(cases "ffold ?f z (finsert x xs) = ffold ?f' z xs")
case True
then show ?thesis using insert.IH[OF 0] insert.prems(2-4) by simp
next
case not_right: False
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
show ?thesis
proof(cases "(v,e) \<in> set (dtree_to_list (Node r {|(t2,e2)|}))")
case True
then show ?thesis using dtree_to_list_x1_list_disjoint insert.prems(2) by fastforce
next
case False
have "xs |\<subseteq>| finsert x xs" by blast
then have f_xs: "ffold ?f z xs = ffold ?f' z xs"
using merge_ffold_supset[of xs "finsert x xs"] insert.prems(1) by blast
have "ffold ?f z (finsert x xs) = ?f x (ffold ?f z xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
then have 1: "ffold ?f z (finsert x xs) = ?f x (ffold ?f' z xs)" using f_xs by argo
then have "?f x (ffold ?f' z xs) \<noteq> ffold ?f' z xs" using not_right by argo
then have "?f (t2,e2) (ffold ?f' z xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using merge_f_merge_if_not_snd t2_def by blast
then have "ffold ?f z (finsert x xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using 1 t2_def by argo
then have "(v,e) \<in> set (?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs))"
using insert.prems(4) by argo
then have "(v,e) \<in> set (ffold ?f' z xs)" using set_merge False by fast
then show ?thesis using insert.IH insert.prems(2-3) 0 by auto
qed
qed
-qed (fastforce)
+qed (fastforce simp: ffold.rep_eq)
lemma merge_wf_dlverts: "wf_dlverts (merge t)"
proof -
have 0: "list_dtree (Node (root t) (sucs t))" using list_dtree_axioms by simp
have 1: "\<forall>t1\<in>fst ` fset (sucs t). set (root t) \<inter> dlverts t1 = {}"
using wf_lverts wf_dlverts.simps[of "root t"] by fastforce
have "\<forall>v\<in>fst ` set (ffold (merge_f (root t) (sucs t)) [] (sucs t)). set (root t) \<inter> set v = {}"
using wf_lverts merge_ffold_root_inter_preserv[OF 0 1] by force
moreover have "wf_list_lverts (ffold (merge_f (root t) (sucs t)) [] (sucs t))"
using merge_ffold_wf_list_lverts[OF wf_dlverts_sucs[OF wf_lverts] 0] by simp
moreover have "root t \<noteq> []" using wf_lverts wf_dlverts.elims(2) by fastforce
ultimately show ?thesis unfolding merge_def using wf_dlverts_iff_wf_list_lverts by blast
qed
theorem merge_list_dtree: "list_dtree (merge t)"
using merge_wf_dlverts merge_wf_darcs list_dtree_def by blast
corollary merge_ranked_dtree: "ranked_dtree (merge t) cmp"
using merge_list_dtree ranked_dtree_def ranked_dtree_axioms by auto
subsubsection \<open>Additional Merging Properties\<close>
lemma merge_ffold_distinct:
"\<lbrakk>list_dtree (Node r xs); \<forall>t1 \<in> fst ` fset xs. \<forall>v\<in>dverts t1. distinct v;
\<forall>v1 \<in> fst ` set z. distinct v1; v \<in> fst ` set (ffold (merge_f r xs) z xs)\<rbrakk>
\<Longrightarrow> distinct v"
proof(induction xs)
case (insert x xs)
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
let ?merge = "Sorting_Algorithms.merge"
have 0: "list_dtree (Node r xs)" using insert.prems(1) list_dtree_subset by blast
show ?case
proof(cases "ffold ?f z (finsert x xs) = ffold ?f' z xs")
case True
then show ?thesis using insert.IH[OF 0] insert.prems(2-4) by simp
next
case not_right: False
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
show ?thesis
proof(cases "v \<in> fst ` set (dtree_to_list (Node r {|(t2,e2)|}))")
case True
have "\<forall>v\<in>dverts t2. distinct v" using insert.prems(2) by simp
then have 2: "\<forall>v\<in>fst ` set (dtree_to_list (Node r {|(t2,e2)|})). distinct v"
by (simp add: dtree_to_list_x_in_dverts)
then show ?thesis using True by auto
next
case False
have "xs |\<subseteq>| finsert x xs" by blast
then have f_xs: "ffold ?f z xs = ffold ?f' z xs"
using merge_ffold_supset insert.prems(1) by presburger
have "ffold ?f z (finsert x xs) = ?f x (ffold ?f z xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
then have 1: "ffold ?f z (finsert x xs) = ?f x (ffold ?f' z xs)" using f_xs by argo
then have "?f x (ffold ?f' z xs) \<noteq> ffold ?f' z xs" using not_right by argo
then have "?f (t2,e2) (ffold ?f' z xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using merge_f_merge_if_not_snd t2_def by blast
then have "ffold ?f z (finsert x xs)
= ?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs)"
using 1 t2_def by argo
then have "v \<in> fst ` set (?merge cmp' (dtree_to_list (Node r {|(t2,e2)|})) (ffold ?f' z xs))"
using insert.prems(4) by argo
then have "v \<in> fst ` set (ffold ?f' z xs)" using set_merge False by fast
then show ?thesis using insert.IH[OF 0] insert.prems(2-3) by simp
qed
qed
-qed (fastforce)
+qed (fastforce simp: ffold.rep_eq)
lemma distinct_merge:
assumes "\<forall>v\<in>dverts t. distinct v" and "v\<in>dverts (merge t)"
shows "distinct v"
proof(cases "v = root t")
case True
then show ?thesis by (simp add: dtree.set_sel(1) assms(1))
next
case False
then have 0: "v \<in> fst ` set (ffold (merge_f (root t) (sucs t)) [] (sucs t))"
using merge_def assms(2) dtree_from_list_eq_dverts[of "root t"] by auto
moreover have "\<forall>t1\<in>fst ` fset (sucs t). \<forall>v\<in>dverts t1. distinct v"
using assms(1) dverts_child_subset[of "root t" "sucs t"] by auto
moreover have "\<forall>v1\<in>fst ` set []. distinct v1" by simp
moreover have 0: "list_dtree (Node (root t) (sucs t))" using list_dtree_axioms by simp
ultimately show ?thesis using merge_ffold_distinct by fast
qed
lemma merge_hd_root_eq[simp]: "hd (root (merge t1)) = hd (root t1)"
unfolding merge_def by auto
lemma merge_ffold_hd_is_child:
"\<lbrakk>list_dtree (Node r xs); xs \<noteq> {||}\<rbrakk>
\<Longrightarrow> \<exists>(t1,e1) \<in> fset xs. hd (ffold (merge_f r xs) [] xs) = (root t1,e1)"
proof(induction xs)
case (insert x xs)
interpret Comm: comp_fun_commute "merge_f r (finsert x xs)" by (rule merge_commute)
define f where "f = merge_f r (finsert x xs)"
define f' where "f' = merge_f r xs"
let ?merge = "Sorting_Algorithms.merge cmp'"
have 0: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(1) by blast
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
have i1: "\<exists>(t1, e1)\<in>fset (finsert x xs). hd (dtree_to_list (Node r {|(t2,e2)|})) = (root t1, e1)"
by simp
have "(t2, e2) \<in> fset (finsert x xs)" by simp
- moreover have "(t2, e2) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t2, e2) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold f' [] xs). set v \<inter> dlverts t2 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t2 \<union> {e2})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(1) empty_list_valid_merge] f'_def
by blast
have "ffold f [] (finsert x xs) = f x (ffold f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] f_def by blast
also have "\<dots> = f x (ffold f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(1) f_def f'_def by fastforce
finally have "ffold f [] (finsert x xs) = ?merge (dtree_to_list (Node r {|x|})) (ffold f' [] xs)"
using xs_val insert.prems f_def by simp
then have merge: "ffold f [] (finsert x xs)
= ?merge (dtree_to_list (Node r {|(t2,e2)|})) (ffold f'[] xs)"
using t2_def by blast
show ?case
proof(cases "xs = {||}")
case True
- then show ?thesis using merge i1 f_def by auto
+ then show ?thesis using merge i1 f_def by (auto simp: ffold.rep_eq)
next
case False
then have i2: "\<exists>(t1,e1) \<in> fset (finsert x xs). hd (ffold f' [] xs) = (root t1,e1)"
using insert.IH[OF 0] f'_def by simp
show ?thesis using merge_hd_exists_preserv[OF i1 i2] merge f_def by simp
qed
qed(simp)
lemma merge_ffold_nempty_if_child:
assumes "(t1,e1) \<in> fset (sucs (merge t0))"
shows "ffold (merge_f (root t0) (sucs t0)) [] (sucs t0) \<noteq> []"
using assms unfolding merge_def by auto
lemma merge_ffold_hd_eq_child:
assumes "(t1,e1) \<in> fset (sucs (merge t0))"
shows "hd (ffold (merge_f (root t0) (sucs t0)) [] (sucs t0)) = (root t1,e1)"
proof -
have "merge t0 = (dtree_from_list (root t0) (ffold (merge_f (root t0) (sucs t0)) [] (sucs t0)))"
unfolding merge_def by blast
have "merge t0 = (Node (root t0) {|(t1,e1)|})" using merge_cases_sucs[of t0] assms by auto
have 0: "(Node (root t0) {|(t1,e1)|})
= (dtree_from_list (root t0) (ffold (merge_f (root t0) (sucs t0)) [] (sucs t0)))"
using merge_cases_sucs[of t0] assms unfolding merge_def by fastforce
then obtain ys where "(root t1, e1) # ys = ffold (merge_f (root t0) (sucs t0)) [] (sucs t0)"
using dtree_from_list_eq_singleton[OF 0] by blast
then show ?thesis using list.sel(1)[of "(root t1, e1)" ys] by simp
qed
lemma merge_child_in_orig:
assumes "(t1,e1) \<in> fset (sucs (merge t0))"
shows "\<exists>(t2,e2) \<in> fset (sucs t0). (root t2,e2) = (root t1,e1)"
proof -
have 0: "list_dtree (Node (root t0) (sucs t0))" using assms merge_empty_if_nwf_sucs by fastforce
have "sucs t0 \<noteq> {||}" using assms merge_empty_sucs by fastforce
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset (sucs t0)"
"hd (ffold (merge_f (root t0) (sucs t0)) [] (sucs t0)) = (root t2,e2)"
using merge_ffold_hd_is_child[OF 0] by blast
then show ?thesis using merge_ffold_hd_eq_child[OF assms] by auto
qed
lemma ffold_singleton: "comp_fun_commute f \<Longrightarrow> ffold f z {|x|} = f x z"
- using comp_fun_commute.ffold_finsert by fastforce
+ using comp_fun_commute.ffold_finsert
+ by (metis comp_fun_commute.ffold_empty finsert_absorb finsert_not_fempty)
lemma ffold_singleton1:
"\<lbrakk>comp_fun_commute (\<lambda>a b. if P a b then Q a b else R a b); P x z\<rbrakk>
\<Longrightarrow> ffold (\<lambda>a b. if P a b then Q a b else R a b) z {|x|} = Q x z"
using ffold_singleton by fastforce
lemma ffold_singleton2:
"\<lbrakk>comp_fun_commute (\<lambda>a b. if P a b then Q a b else R a b); \<not>P x z\<rbrakk>
\<Longrightarrow> ffold (\<lambda>a b. if P a b then Q a b else R a b) z {|x|} = R x z"
using ffold_singleton by fastforce
lemma merge_ffold_singleton_if_wf:
assumes "list_dtree (Node r {|(t1,e1)|})"
shows "ffold (merge_f r {|(t1,e1)|}) [] {|(t1,e1)|} = dtree_to_list (Node r {|(t1,e1)|})"
proof -
interpret Comm: comp_fun_commute "merge_f r {|(t1,e1)|}" by (rule merge_commute)
define f where "f = merge_f r {|(t1,e1)|}"
have "ffold f [] {|(t1,e1)|} = f (t1,e1) (ffold f [] {||})"
using Comm.ffold_finsert f_def by blast
- then show ?thesis using f_def assms by simp
+ then show ?thesis using f_def assms by (simp add: ffold.rep_eq)
qed
lemma merge_singleton_if_wf:
assumes "list_dtree (Node r {|(t1,e1)|})"
shows "merge (Node r {|(t1,e1)|}) = dtree_from_list r (dtree_to_list (Node r {|(t1,e1)|}))"
using merge_ffold_singleton_if_wf[OF assms] merge_xs by simp
lemma merge_disjoint_if_child:
"merge (Node r {|(t1,e1)|}) = Node r {|(t2,e2)|} \<Longrightarrow> list_dtree (Node r {|(t1,e1)|})"
using merge_empty_if_nwf by fastforce
lemma merge_root_child_eq:
"merge (Node r {|(t1,e1)|}) = Node r {|(t2,e2)|} \<Longrightarrow> root t1 = root t2"
using merge_singleton_if_wf[OF merge_disjoint_if_child] by fastforce
lemma merge_ffold_split_subtree:
"\<lbrakk>\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1; list_dtree (Node r xs);
as@(v,e)#bs = ffold (merge_f r xs) [] xs\<rbrakk>
\<Longrightarrow> \<exists>ys. strict_subtree (Node v ys) (Node r xs) \<and> dverts (Node v ys) \<subseteq> (fst ` set ((v,e)#bs))"
proof(induction xs arbitrary: as bs)
case (insert x xs)
obtain t1 e1 where t1_def[simp]: "x = (t1,e1)" by fastforce
define f' where "f' = merge_f r xs"
let ?f = "merge_f r (finsert x xs)"
let ?f' = "merge_f r xs"
have "(t1, e1) \<in> fset (finsert x xs)" by simp
- moreover have "(t1, e1) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t1, e1) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold ?f' [] xs). set v \<inter> dlverts t1 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t1 \<union> {e1})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(2) empty_list_valid_merge] by blast
have 0: "\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1" using insert.prems(1) by simp
have 1: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(2) by blast
have "ffold ?f [] (finsert x xs) = ?f x (ffold ?f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] by blast
also have "\<dots> = ?f x (ffold ?f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(2) by fastforce
finally have ind: "ffold ?f [] (finsert x xs)
= Sorting_Algorithms.merge cmp' (dtree_to_list (Node r {|x|})) (ffold f' [] xs)"
using insert.prems(2) xs_val f'_def by simp
have "max_deg (fst x) \<le> 1" using insert.prems(1) by simp
then have "max_deg (Node r {|x|}) \<le> 1"
using mdeg_child_sucs_eq_if_gt1[of r "fst x" "snd x" "root (fst x)"] by fastforce
then have "\<forall>as bs. as@(v,e)#bs = dtree_to_list (Node r {|x|}) \<longrightarrow>
(\<exists>zs. strict_subtree (Node v zs) (Node r {|x|})
\<and> dverts (Node v zs) \<subseteq> fst ` set ((v,e)#bs))"
using dtree_to_list_split_subtree_dverts_eq_fsts' by fast
then have left: "\<forall>as bs. as@(v,e)#bs = dtree_to_list (Node r {|x|}) \<longrightarrow>
(\<exists>zs. strict_subtree (Node v zs) (Node r (finsert x xs))
\<and> dverts (Node v zs) \<subseteq> fst ` set ((v,e)#bs))"
using strict_subtree_singleton[where xs="finsert x xs"] by blast
have "\<forall>as bs. as@(v,e)#bs = ffold f' [] xs \<longrightarrow>
(\<exists>zs. strict_subtree (Node v zs) (Node r xs)
\<and> dverts (Node v zs) \<subseteq> fst ` set ((v,e)#bs))"
using insert.IH[OF 0 1] f'_def by blast
then have right: "\<forall>as bs. as@(v,e)#bs = ffold f' [] xs \<longrightarrow>
(\<exists>zs. strict_subtree (Node v zs) (Node r (finsert x xs))
\<and> dverts (Node v zs) \<subseteq> fst ` set ((v,e)#bs))"
using strict_subtree_subset[where r=r and xs=xs and ys="finsert x xs"] by fast
then show ?case using merge_split_supset_strict_subtree[OF left right] ind insert.prems(3) by simp
-qed (simp)
+qed (simp add: ffold.rep_eq)
lemma merge_strict_subtree_dverts_sup:
assumes "\<forall>t \<in> fst ` fset (sucs t). max_deg t \<le> 1"
and "strict_subtree (Node r xs) (merge t)"
shows "\<exists>ys. is_subtree (Node r ys) t \<and> dverts (Node r ys) \<subseteq> dverts (Node r xs)"
proof -
have 0: "list_dtree (Node (root t) (sucs t))" using list_dtree_axioms by simp
have "\<forall>as r e bs. as@(r,e)#bs = ffold (merge_f (root t) (sucs t)) [] (sucs t)
\<longrightarrow> (\<exists>ys. strict_subtree (Node r ys) (Node (root t) (sucs t))
\<and> dverts (Node r ys) \<subseteq> fst ` set ((r,e)#bs))"
using merge_ffold_split_subtree[OF assms(1) 0] by blast
then have "\<forall>as r e bs. as@(r,e)#bs = ffold (merge_f (root t) (sucs t)) [] (sucs t) \<longrightarrow>
(\<exists>ys. strict_subtree (Node r ys) t \<and> dverts (Node r ys) \<subseteq> fst ` set ((r,e)#bs))"
by simp
obtain as e bs where bs_def: "as@(r,e)#bs = ffold (merge_f (root t) (sucs t)) [] (sucs t)"
using assms(2) dtree_from_list_uneq_sequence_xs[of r] unfolding merge_def by blast
have "wf_dverts (merge t)" by (simp add: merge_wf_dlverts wf_dverts_if_wf_dlverts)
then have wf: "wf_dverts (dtree_from_list (root t) (as@(r,e)#bs))"
unfolding merge_def bs_def .
moreover obtain ys where
"strict_subtree (Node r ys) t" "dverts (Node r ys) \<subseteq> fst ` set ((r,e)#bs)"
using merge_ffold_split_subtree[OF assms(1) 0 bs_def] by auto
moreover have "strict_subtree (Node r xs) (dtree_from_list (root t) (as@(r,e)#bs))"
using assms(2) unfolding bs_def merge_def .
ultimately show ?thesis
using dtree_from_list_dverts_subset_wfdverts1 unfolding strict_subtree_def by fast
qed
lemma merge_subtree_dverts_supset:
assumes "\<forall>t\<in>fst ` fset (sucs t). max_deg t \<le> 1" and "is_subtree (Node r xs) (merge t)"
shows "\<exists>ys. is_subtree (Node r ys) t \<and> dverts (Node r ys) \<subseteq> dverts (Node r xs)"
proof(cases "Node r xs = merge t")
case True
then obtain ys where "t = Node r ys" using merge_root_eq dtree.exhaust_sel dtree.sel(1) by metis
then show ?thesis using dverts_merge_eq[OF assms(1)] True by auto
next
case False
then show ?thesis using merge_strict_subtree_dverts_sup assms strict_subtree_def by blast
qed
lemma merge_subtree_dlverts_supset:
assumes "\<forall>t\<in>fst ` fset (sucs t). max_deg t \<le> 1" and "is_subtree (Node r xs) (merge t)"
shows "\<exists>ys. is_subtree (Node r ys) t \<and> dlverts (Node r ys) \<subseteq> dlverts (Node r xs)"
proof -
obtain ys where "is_subtree (Node r ys) t" "dverts (Node r ys) \<subseteq> dverts (Node r xs)"
using merge_subtree_dverts_supset[OF assms] by blast
then show ?thesis using dlverts_eq_dverts_union[of "Node r ys"] dlverts_eq_dverts_union by fast
qed
end
subsection \<open>Normalizing Dtrees\<close>
context ranked_dtree
begin
subsubsection \<open>Definitions\<close>
function normalize1 :: "('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"normalize1 (Node r {|(t1,e)|}) =
(if rank (rev (root t1)) < rank (rev r) then Node (r@root t1) (sucs t1)
else Node r {|(normalize1 t1,e)|})"
| "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> normalize1 (Node r xs) = Node r ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)"
by (metis darcs_mset.cases old.prod.exhaust) fast+
termination by lexicographic_order
lemma normalize1_size_decr[termination_simp]:
"normalize1 t1 \<noteq> t1 \<Longrightarrow> size (normalize1 t1) < size t1"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then show ?thesis using dtree_size_eq_root[of "root t" "sucs t"] by simp
next
case False
then show ?thesis using dtree_size_img_le 1 by auto
qed
next
case (2 xs r)
then have 0: "\<forall>t \<in> fst ` fset xs. size (normalize1 t) \<le> size t" by fastforce
moreover have "\<exists>t \<in> fst ` fset xs. size (normalize1 t) < size t"
using elem_neq_if_fset_neq[of normalize1 xs] 2 by fastforce
ultimately show ?case using dtree_size_img_lt "2.hyps" by auto
qed
lemma normalize1_size_le: "size (normalize1 t1) \<le> size t1"
by(cases "normalize1 t1=t1") (auto dest: normalize1_size_decr)
fun normalize :: "('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"normalize t1 = (let t2 = normalize1 t1 in if t1 = t2 then t2 else normalize t2)"
subsubsection \<open>Basic Proofs\<close>
lemma root_normalize1_eq1:
"\<not>rank (rev (root t1)) < rank (rev r) \<Longrightarrow> root (normalize1 (Node r {|(t1,e1)|})) = r"
by simp
lemma root_normalize1_eq1':
"\<not>rank (rev (root t1)) \<le> rank (rev r) \<Longrightarrow> root (normalize1 (Node r {|(t1,e1)|})) = r"
by simp
lemma root_normalize1_eq2: "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> root (normalize1 (Node r xs)) = r"
by simp
lemma fset_img_eq: "\<forall>x \<in> fset xs. f x = x \<Longrightarrow> f |`| xs = xs"
using fset_inject[of xs "f |`| xs"] by simp
lemma fset_img_uneq: "f |`| xs \<noteq> xs \<Longrightarrow> \<exists>x \<in> fset xs. f x \<noteq> x"
using fset_img_eq by fastforce
lemma fset_img_uneq_prod: "(\<lambda>(t,e). (f t, e)) |`| xs \<noteq> xs \<Longrightarrow> \<exists>(t,e) \<in> fset xs. f t \<noteq> t"
using fset_img_uneq[of "\<lambda>(t,e). (f t, e)" xs] by auto
lemma contr_if_normalize1_uneq:
"normalize1 t1 \<noteq> t1
\<Longrightarrow> \<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) t1 \<and> rank (rev (root t2)) < rank (rev v)"
proof(induction t1 rule: normalize1.induct)
case (2 xs r)
then show ?case using fset_img_uneq_prod[of normalize1 xs] by fastforce
qed(fastforce)
lemma contr_before_normalize1:
"\<lbrakk>is_subtree (Node v {|(t1,e1)|}) (normalize1 t3); rank (rev (root t1)) < rank (rev v)\<rbrakk>
\<Longrightarrow> \<exists>v' t2 e2. is_subtree (Node v' {|(t2,e2)|}) t3 \<and> rank (rev (root t2)) < rank (rev v')"
using contr_if_normalize1_uneq by force
subsubsection \<open>Normalizing Preserves Well-Formedness\<close>
lemma normalize1_darcs_sub: "darcs (normalize1 t1) \<subseteq> darcs t1"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then have "darcs (normalize1 (Node r {|(t,e)|})) = darcs (Node (r@root t) (sucs t))" by simp
also have "\<dots> = darcs (Node (root t) (sucs t))" using darcs_sub_if_children_sub by fast
finally show ?thesis by auto
next
case False
then show ?thesis using 1 by auto
qed
qed (fastforce)
lemma disjoint_darcs_normalize1:
"wf_darcs t1 \<Longrightarrow> disjoint_darcs ((\<lambda>(t,e). (normalize1 t,e)) |`| (sucs t1))"
using disjoint_darcs_img[OF disjoint_darcs_if_wf, of t1 normalize1]
by (simp add: normalize1_darcs_sub)
lemma wf_darcs_normalize1: "wf_darcs t1 \<Longrightarrow> wf_darcs (normalize1 t1)"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then show ?thesis
using "1.prems" dtree.collapse singletonI finsert.rep_eq case_prodD
unfolding wf_darcs_iff_darcs'
by (metis (no_types, lifting) wf_darcs'.simps bot_fset.rep_eq normalize1.simps(1))
next
case False
have "disjoint_darcs {|(normalize1 t,e)|}"
using normalize1_darcs_sub disjoint_darcs_if_wf_xs[OF "1.prems"] by auto
then show ?thesis using 1 False unfolding wf_darcs_iff_darcs' by force
qed
next
case (2 xs r)
then show ?case
using disjoint_darcs_normalize1[OF "2.prems"]
by (fastforce simp: wf_darcs_iff_darcs')
qed
lemma normalize1_dlverts_eq[simp]: "dlverts (normalize1 t1) = dlverts t1"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then show ?thesis using dlverts.simps[of "root t" "sucs t"] by force
next
case False
then show ?thesis using 1 by auto
qed
qed (fastforce)
lemma normalize1_dverts_contr_subtree:
"\<lbrakk>v \<in> dverts (normalize1 t1); v \<notin> dverts t1\<rbrakk>
\<Longrightarrow> \<exists>v2 t2 e2. is_subtree (Node v2 {|(t2,e2)|}) t1
\<and> v2 @ root t2 = v \<and> rank (rev (root t2)) < rank (rev v2)"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then show ?thesis using "1.prems" dverts_suc_subseteq by fastforce
next
case False
then show ?thesis using 1 by auto
qed
qed(fastforce)
lemma normalize1_dverts_app_contr:
"\<lbrakk>v \<in> dverts (normalize1 t1); v \<notin> dverts t1\<rbrakk>
\<Longrightarrow> \<exists>v1\<in>dverts t1. \<exists>v2\<in>dverts t1. v1 @ v2 = v \<and> rank (rev v2) < rank (rev v1)"
using normalize1_dverts_contr_subtree
by (fastforce simp: single_subtree_root_dverts single_subtree_child_root_dverts)
lemma disjoint_dlverts_img:
assumes "disjoint_dlverts xs" and "\<forall>(t,e) \<in> fset xs. dlverts (f t) \<subseteq> dlverts t"
shows "disjoint_dlverts ((\<lambda>(t,e). (f t,e)) |`| xs)" (is "disjoint_dlverts ?xs")
proof (rule ccontr)
assume "\<not> disjoint_dlverts ?xs"
then obtain x1 e1 y1 e2 where asm: "(x1,e1) \<in> fset ?xs" "(y1,e2) \<in> fset ?xs"
"dlverts x1 \<inter> dlverts y1 \<noteq> {} \<and> (x1,e1)\<noteq>(y1,e2)" by blast
then obtain x2 where x2_def: "f x2 = x1" "(x2,e1) \<in> fset xs" by auto
obtain y2 where y2_def: "f y2 = y1" "(y2,e2) \<in> fset xs" using asm(2) by auto
have "dlverts x1 \<subseteq> dlverts x2" using assms(2) x2_def by fast
moreover have "dlverts y1 \<subseteq> dlverts y2" using assms(2) y2_def by fast
ultimately have "\<not> disjoint_dlverts xs" using asm(3) x2_def y2_def by blast
then show False using assms(1) by blast
qed
lemma disjoint_dlverts_normalize1:
"disjoint_dlverts xs \<Longrightarrow> disjoint_dlverts ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)"
using disjoint_dlverts_img[of xs] by simp
lemma disjoint_dlverts_normalize1_sucs:
"disjoint_dlverts (sucs t1) \<Longrightarrow> disjoint_dlverts ((\<lambda>(t,e). (normalize1 t,e)) |`| (sucs t1))"
using disjoint_dlverts_img[of "sucs t1"] by simp
lemma disjoint_dlverts_normalize1_wf:
"wf_dlverts t1 \<Longrightarrow> disjoint_dlverts ((\<lambda>(t,e). (normalize1 t,e)) |`| (sucs t1))"
using disjoint_dlverts_img[OF disjoint_dlverts_if_wf, of t1] by simp
lemma disjoint_dlverts_normalize1_wf':
"wf_dlverts (Node r xs) \<Longrightarrow> disjoint_dlverts ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)"
using disjoint_dlverts_img[OF disjoint_dlverts_if_wf, of "Node r xs"] by simp
lemma root_empty_inter_dlverts_normalize1:
assumes "wf_dlverts t1" and "(x1,e1) \<in> fset ((\<lambda>(t,e). (normalize1 t,e)) |`| (sucs t1))"
shows "set (root t1) \<inter> dlverts x1 = {}"
proof (rule ccontr)
assume asm: "set (root t1) \<inter> dlverts x1 \<noteq> {}"
obtain x2 where x2_def: "normalize1 x2 = x1" "(x2,e1) \<in> fset (sucs t1)" using assms(2) by auto
have "set (root t1) \<inter> dlverts x2 \<noteq> {}" using x2_def(1) asm by force
then show False using x2_def(2) assms(1) wf_dlverts.simps[of "root t1" "sucs t1"] by auto
qed
lemma wf_dlverts_normalize1: "wf_dlverts t1 \<Longrightarrow> wf_dlverts (normalize1 t1)"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
have 0: "\<forall>(t1,e1) \<in> fset (sucs t). wf_dlverts t1"
using "1.prems" wf_dlverts.simps[of "root t" "sucs t"] by auto
have "\<forall>(t1,e1) \<in> fset (sucs t). set (root t) \<inter> dlverts t1 = {}"
using "1.prems" wf_dlverts.simps[of "root t"] by fastforce
then have "\<forall>(t1,e1) \<in> fset (sucs t). set (r@root t) \<inter> dlverts t1 = {}"
using suc_in_dlverts "1.prems" by fastforce
then show ?thesis using True 0 disjoint_dlverts_if_wf[of t] "1.prems" by auto
next
case False
then show ?thesis
using root_empty_inter_dlverts_normalize1[OF "1.prems"] disjoint_dlverts_normalize1 1 by auto
qed
next
case (2 xs r)
have "\<forall>(t1,e1) \<in> fset ((\<lambda>(t, e). (normalize1 t, e)) |`| xs). set r \<inter> dlverts t1 = {}"
using root_empty_inter_dlverts_normalize1[OF "2.prems"] by force
then show ?case using disjoint_dlverts_normalize1 2 by auto
qed
corollary list_dtree_normalize1: "list_dtree (normalize1 t)"
using wf_dlverts_normalize1[OF wf_lverts] wf_darcs_normalize1[OF wf_arcs] list_dtree_def by blast
corollary ranked_dtree_normalize1: "ranked_dtree (normalize1 t) cmp"
using list_dtree_normalize1 ranked_dtree_def ranked_dtree_axioms by blast
lemma normalize_darcs_sub: "darcs (normalize t1) \<subseteq> darcs t1"
apply(induction t1 rule: normalize.induct)
by (smt (verit) normalize1_darcs_sub normalize.simps subset_trans)
lemma normalize_dlverts_eq: "dlverts (normalize t1) = dlverts t1"
by(induction t1 rule: normalize.induct) (metis (full_types) normalize.elims normalize1_dlverts_eq)
theorem ranked_dtree_normalize: "ranked_dtree (normalize t) cmp"
using ranked_dtree_axioms apply(induction t rule: normalize.induct)
by (smt (verit) ranked_dtree.normalize.elims ranked_dtree.ranked_dtree_normalize1)
subsubsection \<open>Distinctness and hd preserved\<close>
lemma distinct_normalize1: "\<lbrakk>\<forall>v\<in>dverts t. distinct v; v\<in>dverts (normalize1 t)\<rbrakk> \<Longrightarrow> distinct v"
using ranked_dtree_axioms proof(induction t rule: normalize1.induct)
case (1 r t e)
then interpret R: ranked_dtree "Node r {|(t, e)|}" rank by blast
show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
interpret T: ranked_dtree t rank using R.ranked_dtree_rec by auto
have "set r \<inter> set (root t) = {}"
using R.wf_lverts dlverts.simps[of "root t" "sucs t"] by auto
then have "distinct (r@root t)" by (auto simp: dtree.set_sel(1) "1.prems"(1))
moreover have "\<forall>v \<in> (\<Union>(t, e)\<in>fset (sucs t). dverts t). distinct v"
using "1.prems"(1) dtree.set(1)[of "root t" "sucs t"] by fastforce
ultimately show ?thesis using dverts_root_or_child "1.prems"(2) True by auto
next
case False
then show ?thesis using R.ranked_dtree_rec 1 by auto
qed
next
case (2 xs r)
then interpret R: ranked_dtree "Node r xs" rank by blast
show ?case using R.ranked_dtree_rec 2 by fastforce
qed
lemma distinct_normalize: "\<forall>v\<in>dverts t. distinct v \<Longrightarrow> \<forall>v\<in>dverts (normalize t). distinct v"
using ranked_dtree_axioms proof(induction t rule: normalize.induct)
case (1 t)
then interpret T1: ranked_dtree "t" rank by blast
interpret T2: ranked_dtree "normalize1 t" rank by (simp add: T1.ranked_dtree_normalize1)
show ?case
by (smt (verit, del_insts) 1 T1.distinct_normalize1 T2.ranked_dtree_axioms normalize.simps)
qed
lemma normalize1_hd_root_eq[simp]:
assumes "root t1 \<noteq> []"
shows "hd (root (normalize1 t1)) = hd (root t1)"
proof(cases "\<forall>x. sucs t1 \<noteq> {|x|}")
case True
then show ?thesis using normalize1.simps(2)[of "sucs t1" "root t1"] by simp
next
case False
then obtain t e where "{|(t, e)|} = sucs t1" by auto
then show ?thesis using normalize1.simps(1)[of "root t1" t e] assms by simp
qed
corollary normalize1_hd_root_eq':
"wf_dlverts t1 \<Longrightarrow> hd (root (normalize1 t1)) = hd (root t1)"
using normalize1_hd_root_eq[of t1] wf_dlverts.simps[of "root t1" "sucs t1"] by simp
lemma normalize1_root_nempty:
assumes "root t1 \<noteq> []"
shows "root (normalize1 t1) \<noteq> []"
proof(cases "\<forall>x. sucs t1 \<noteq> {|x|}")
case True
then show ?thesis using normalize1.simps(2)[of "sucs t1" "root t1"] assms by simp
next
case False
then obtain t e where "{|(t, e)|} = sucs t1" by auto
then show ?thesis using normalize1.simps(1)[of "root t1" t e] assms by simp
qed
lemma normalize_hd_root_eq[simp]: "root t1 \<noteq> [] \<Longrightarrow> hd (root (normalize t1)) = hd (root t1)"
using ranked_dtree_axioms proof(induction t1 rule: normalize.induct)
case (1 t)
then show ?case
proof(cases "t = normalize1 t")
case False
then have "normalize t = normalize (normalize1 t)" by (simp add: Let_def)
then show ?thesis using 1 normalize1_root_nempty by force
qed(simp)
qed
corollary normalize_hd_root_eq'[simp]: "wf_dlverts t1 \<Longrightarrow> hd (root (normalize t1)) = hd (root t1)"
using normalize_hd_root_eq wf_dlverts.simps[of "root t1" "sucs t1"] by simp
subsubsection \<open>Normalize and Sorting\<close>
lemma normalize1_uneq_if_contr:
"\<lbrakk>is_subtree (Node r1 {|(t1,e1)|}) t2; rank (rev (root t1)) < rank (rev r1); wf_darcs t2\<rbrakk>
\<Longrightarrow> t2 \<noteq> normalize1 t2"
proof(induction t2 rule: normalize1.induct)
case (1 r t e)
then show ?case
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then show ?thesis using combine_uneq by fastforce
next
case False
then show ?thesis using 1 by auto
qed
next
case (2 xs r)
then obtain t e where t_def: "(t,e) \<in> fset xs" "is_subtree (Node r1 {|(t1,e1)|}) t" by auto
then have "t \<noteq> normalize1 t" using 2 by fastforce
then have "(normalize1 t, e) \<notin> fset xs"
using "2.prems"(3) t_def(1) by (auto simp: wf_darcs_iff_darcs')
moreover have "(normalize1 t, e) \<in> fset ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)"
using t_def(1) by auto
ultimately have "(\<lambda>(t,e). (normalize1 t,e)) |`| xs \<noteq> xs" using t_def(1) by fastforce
then show ?case using "2.hyps" by simp
qed
lemma sorted_ranks_if_normalize1_eq:
"\<lbrakk>wf_darcs t2; is_subtree (Node r1 {|(t1,e1)|}) t2; t2 = normalize1 t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
using normalize1_uneq_if_contr by fastforce
lemma normalize_sorted_ranks:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) (normalize t)\<rbrakk> \<Longrightarrow> rank (rev r) \<le> rank (rev (root t1))"
using ranked_dtree_axioms proof(induction t rule: normalize.induct)
case (1 t)
then interpret T: ranked_dtree t by blast
show ?case
using 1 sorted_ranks_if_normalize1_eq[OF T.wf_arcs]
by (smt (verit, ccfv_SIG) T.ranked_dtree_normalize1 normalize.simps)
qed
lift_definition cmp'' :: "('a list\<times>'b) comparator" is
"(\<lambda>x y. if rank (rev (fst x)) < rank (rev (fst y)) then Less
else if rank (rev (fst x)) > rank (rev (fst y)) then Greater
else Equiv)"
by (simp add: comparator_def)
lemma dtree_to_list_sorted_if_no_contr:
"\<lbrakk>\<And>r1 t1 e1. is_subtree (Node r1 {|(t1,e1)|}) t2 \<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))\<rbrakk>
\<Longrightarrow> sorted cmp'' (dtree_to_list (Node r {|(t2,e2)|}))"
proof(induction cmp'' "dtree_to_list (Node r {|(t2,e2)|})" arbitrary: r t2 e2 rule: sorted.induct)
case (2 x)
then show ?case using sorted_single[of cmp'' x] by simp
next
case (3 y x xs)
then obtain r1 t1 e1 where r1_def: "t2 = Node r1 {|(t1,e1)|}"
using dtree_to_list.elims[of t2] by fastforce
have "y = (root t2,e2)" using "3.hyps"(2) r1_def by simp
moreover have "x = (root t1,e1)" using "3.hyps"(2) r1_def by simp
moreover have "rank (rev (root t2)) \<le> rank (rev (root t1))" using "3.prems" r1_def by auto
ultimately have "compare cmp'' y x \<noteq> Greater" using cmp''.rep_eq by simp
moreover have "sorted cmp'' (dtree_to_list t2)" using 3 r1_def by auto
ultimately show ?case using 3 r1_def by simp
qed(simp)
lemma dtree_to_list_sorted_if_no_contr':
"\<lbrakk>\<And>r1 t1 e1. is_subtree (Node r1 {|(t1,e1)|}) t2 \<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))\<rbrakk>
\<Longrightarrow> sorted cmp'' (dtree_to_list t2)"
using dtree_to_list_sorted_if_no_contr[of t2] sorted_Cons_imp_sorted by fastforce
lemma dtree_to_list_sorted_if_subtree:
"\<lbrakk>is_subtree t1 t2;
\<And>r1 t1 e1. is_subtree (Node r1 {|(t1,e1)|}) t2 \<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))\<rbrakk>
\<Longrightarrow> sorted cmp'' (dtree_to_list (Node r {|(t1,e1)|}))"
using dtree_to_list_sorted_if_no_contr subtree_trans by blast
lemma dtree_to_list_sorted_if_subtree':
"\<lbrakk>is_subtree t1 t2;
\<And>r1 t1 e1. is_subtree (Node r1 {|(t1,e1)|}) t2 \<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))\<rbrakk>
\<Longrightarrow> sorted cmp'' (dtree_to_list t1)"
using dtree_to_list_sorted_if_no_contr' subtree_trans by blast
lemma normalize_dtree_to_list_sorted:
"is_subtree t1 (normalize t) \<Longrightarrow> sorted cmp'' (dtree_to_list (Node r {|(t1,e1)|}))"
using dtree_to_list_sorted_if_subtree normalize_sorted_ranks by blast
lemma normalize_dtree_to_list_sorted':
"is_subtree t1 (normalize t) \<Longrightarrow> sorted cmp'' (dtree_to_list t1)"
using dtree_to_list_sorted_if_subtree' normalize_sorted_ranks by blast
lemma gt_if_rank_contr: "rank (rev r0) < rank (rev r) \<Longrightarrow> compare cmp'' (r, e) (r0, e0) = Greater"
by (auto simp: cmp''.rep_eq)
lemma rank_le_if_ngt: "compare cmp'' (r, e) (r0, e0) \<noteq> Greater \<Longrightarrow> rank (rev r) \<le> rank (rev r0)"
using gt_if_rank_contr by force
lemma rank_le_if_sorted_from_list:
assumes "sorted cmp'' ((v1,e1)#ys)" and "is_subtree (Node r0 {|(t0,e0)|}) (dtree_from_list v1 ys)"
shows "rank (rev r0) \<le> rank (rev (root t0))"
proof -
obtain e as bs where e_def: "as @ (r0, e) # (root t0, e0) # bs = ((v1,e1)#ys)"
using dtree_from_list_sequence[OF assms(2)] by blast
then have "sorted cmp'' (as @ (r0, e) # (root t0, e0) # bs)" using assms(1) by simp
then have "sorted cmp'' ((r0, e) # (root t0, e0) # bs)" using sorted_app_r by blast
then show ?thesis using rank_le_if_ngt by auto
qed
lemma cmp'_gt_if_cmp''_gt: "compare cmp'' x y = Greater \<Longrightarrow> compare cmp' x y = Greater"
by (auto simp: cmp'.rep_eq cmp''.rep_eq split: if_splits)
lemma cmp'_lt_if_cmp''_lt: "compare cmp'' x y = Less \<Longrightarrow> compare cmp' x y = Less"
by (auto simp: cmp'.rep_eq cmp''.rep_eq)
lemma cmp''_ge_if_cmp'_gt:
"compare cmp' x y = Greater \<Longrightarrow> compare cmp'' x y = Greater \<or> compare cmp'' x y = Equiv"
by (auto simp: cmp'.rep_eq cmp''.rep_eq split: if_splits)
lemma cmp''_nlt_if_cmp'_gt: "compare cmp' x y = Greater \<Longrightarrow> compare cmp'' y x \<noteq> Greater"
by (auto simp: cmp'.rep_eq cmp''.rep_eq)
interpretation Comm: comp_fun_commute "merge_f r xs" by (rule merge_commute)
lemma sorted_cmp''_merge:
"\<lbrakk>sorted cmp'' xs; sorted cmp'' ys\<rbrakk> \<Longrightarrow> sorted cmp'' (Sorting_Algorithms.merge cmp' xs ys)"
proof(induction xs ys taking: cmp' rule: Sorting_Algorithms.merge.induct)
case (3 x xs y ys)
let ?merge = "Sorting_Algorithms.merge cmp'"
show ?case
proof(cases "compare cmp' x y = Greater")
case True
have "?merge (x # xs) (y#ys) = y # (?merge (x # xs) ys)" using True by simp
moreover have "sorted cmp'' (?merge (x # xs) ys)" using 3 True sorted_Cons_imp_sorted by fast
ultimately show ?thesis
using cmp''_nlt_if_cmp'_gt[OF True] "3.prems" sorted_rec[of cmp'' y]
merge.elims[of cmp' "x#xs" ys "?merge (x # xs) ys"]
by metis
next
case False
have "?merge (x#xs) (y#ys) = x # (?merge xs (y#ys))" using False by simp
moreover have "sorted cmp'' (?merge xs (y#ys))" using 3 False sorted_Cons_imp_sorted by fast
ultimately show ?thesis
using cmp'_gt_if_cmp''_gt False "3.prems" sorted_rec[of cmp'' x]
merge.elims[of cmp' xs "y#ys" "?merge xs (y#ys)"]
by metis
qed
qed(auto)
lemma merge_ffold_sorted:
"\<lbrakk>list_dtree (Node r xs); \<And>t2 r1 t1 e1. \<lbrakk>t2 \<in> fst ` fset xs; is_subtree (Node r1 {|(t1,e1)|}) t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))\<rbrakk>
\<Longrightarrow> sorted cmp'' (ffold (merge_f r xs) [] xs)"
proof(induction xs)
case (insert x xs)
interpret Comm: comp_fun_commute "merge_f r (finsert x xs)" by (rule merge_commute)
define f where "f = merge_f r (finsert x xs)"
define f' where "f' = merge_f r xs"
let ?merge = "Sorting_Algorithms.merge cmp'"
have 0: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(1) by blast
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
have ind1: "sorted cmp'' (dtree_to_list (Node r {|(t2,e2)|}))"
using dtree_to_list_sorted_if_no_contr insert.prems(2) by fastforce
have "\<And>t2 r1 t1 e1. \<lbrakk>t2 \<in> fst ` fset xs; is_subtree (Node r1 {|(t1, e1)|}) t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
using insert.prems(2) by fastforce
then have ind2: "sorted cmp'' (ffold f' [] xs)" using insert.IH[OF 0] f'_def by blast
have "(t2, e2) \<in> fset (finsert x xs)" by simp
- moreover have "(t2, e2) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t2, e2) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold f' [] xs). set v \<inter> dlverts t2 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t2 \<union> {e2})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(1) empty_list_valid_merge] f'_def
by blast
have "ffold f [] (finsert x xs) = f x (ffold f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] f_def by blast
also have "\<dots> = f x (ffold f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(1) f_def f'_def by fastforce
finally have "ffold f [] (finsert x xs) = ?merge (dtree_to_list (Node r {|x|})) (ffold f' [] xs)"
using xs_val insert.prems f_def by simp
then have merge: "ffold f [] (finsert x xs)
= ?merge (dtree_to_list (Node r {|(t2,e2)|})) (ffold f'[] xs)"
using t2_def by blast
then show ?case using sorted_cmp''_merge[OF ind1 ind2] f_def by auto
-qed(simp)
+qed (simp add: ffold.rep_eq)
lemma not_single_subtree_if_nwf:
"\<not>list_dtree (Node r xs) \<Longrightarrow> \<not>is_subtree (Node r1 {|(t1,e1)|}) (merge (Node r xs))"
using merge_empty_if_nwf by simp
lemma not_single_subtree_if_nwf_sucs:
"\<not>list_dtree t2 \<Longrightarrow> \<not>is_subtree (Node r1 {|(t1,e1)|}) (merge t2)"
using merge_empty_if_nwf_sucs by simp
lemma merge_strict_subtree_nocontr:
assumes "\<And>t2 r1 t1 e1. \<lbrakk>t2 \<in> fst ` fset xs; is_subtree (Node r1 {|(t1,e1)|}) t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
and "strict_subtree (Node r1 {|(t1,e1)|}) (merge (Node r xs))"
shows "rank (rev r1) \<le> rank (rev (root t1))"
proof(cases "list_dtree (Node r xs)")
case True
obtain e as bs where e_def: "as @ (r1, e) # (root t1, e1) # bs = ffold (merge_f r xs) [] xs"
using dtree_from_list_uneq_sequence assms(2) unfolding merge_def dtree.sel strict_subtree_def
by fast
have "sorted cmp'' (ffold (merge_f r xs) [] xs)"
using merge_ffold_sorted[OF True assms(1)] by simp
then have "sorted cmp'' ((r1, e) # (root t1, e1) # bs)"
using e_def sorted_app_r[of cmp'' as "(r1, e) # (root t1, e1) # bs"] by simp
then show ?thesis using rank_le_if_sorted_from_list by fastforce
next
case False
then show ?thesis using not_single_subtree_if_nwf assms(2) by (simp add: strict_subtree_def)
qed
lemma merge_strict_subtree_nocontr2:
assumes "\<And>r1 t1 e1. is_subtree (Node r1 {|(t1,e1)|}) (Node r xs)
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
and "strict_subtree (Node r1 {|(t1,e1)|}) (merge (Node r xs))"
shows "rank (rev r1) \<le> rank (rev (root t1))"
using merge_strict_subtree_nocontr[OF assms] by fastforce
lemma merge_strict_subtree_nocontr_sucs:
assumes "\<And>t2 r1 t1 e1. \<lbrakk>t2 \<in> fst ` fset (sucs t0); is_subtree (Node r1 {|(t1,e1)|}) t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
and "strict_subtree (Node r1 {|(t1,e1)|}) (merge t0)"
shows "rank (rev r1) \<le> rank (rev (root t1))"
using merge_strict_subtree_nocontr[of "sucs t0" r1 t1 e1 "root t0"] assms by simp
lemma merge_strict_subtree_nocontr_sucs2:
assumes "\<And>r1 t1 e1. is_subtree (Node r1 {|(t1,e1)|}) t2 \<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
and "strict_subtree (Node r1 {|(t1,e1)|}) (merge t2)"
shows "rank (rev r1) \<le> rank (rev (root t1))"
using merge_strict_subtree_nocontr2[of "root t2" "sucs t2"] assms by auto
lemma no_contr_imp_parent:
"\<lbrakk>is_subtree (Node r1 {|(t1,e1)|}) (Node r xs) \<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1));
t2 \<in> fst ` fset xs; is_subtree (Node r1 {|(t1,e1)|}) t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
using subtree_if_child subtree_trans by fast
lemma no_contr_imp_subtree:
"\<lbrakk>\<And>t2 r1 t1 e1. \<lbrakk>t2 \<in> fst ` fset xs; is_subtree (Node r1 {|(t1,e1)|}) t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1));
is_subtree (Node r1 {|(t1,e1)|}) (Node r xs); \<forall>x. xs \<noteq> {|x|}\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
by fastforce
lemma no_contr_imp_subtree_fcard:
"\<lbrakk>\<And>t2 r1 t1 e1. \<lbrakk>t2 \<in> fst ` fset xs; is_subtree (Node r1 {|(t1,e1)|}) t2\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1));
is_subtree (Node r1 {|(t1,e1)|}) (Node r xs); fcard xs \<noteq> 1\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (root t1))"
using fcard_single_1_iff[of xs] by fastforce
end
subsection \<open>Removing Wedges\<close>
context ranked_dtree
begin
fun merge1 :: "('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"merge1 (Node r xs) = (
if fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1) then merge (Node r xs)
else Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs))"
lemma merge1_dverts_eq[simp]: "dverts (merge1 t) = dverts t"
using ranked_dtree_axioms proof(induction t)
case (Node r xs)
then interpret R: ranked_dtree "Node r xs" rank by blast
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then show ?thesis by simp
next
case False
then show ?thesis using Node.IH R.ranked_dtree_rec by auto
qed
qed
lemma merge1_dlverts_eq[simp]: "dlverts (merge1 t) = dlverts t"
using ranked_dtree_axioms proof(induction t)
case (Node r xs)
then interpret R: ranked_dtree "Node r xs" rank by blast
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then show ?thesis by simp
next
case False
then show ?thesis using Node.IH R.ranked_dtree_rec by auto
qed
qed
lemma dverts_merge1_img_sub:
"\<forall>(t2,e2) \<in> fset xs. dverts (merge1 t2) \<subseteq> dverts t2
\<Longrightarrow> dverts (Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs)) \<subseteq> dverts (Node r xs)"
by fastforce
lemma merge1_dverts_sub: "dverts (merge1 t1) \<subseteq> dverts t1"
proof(induction t1)
case (Node r xs)
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then show ?thesis using dverts_merge_sub by force
next
case False
then have "\<forall>(t2,e2) \<in> fset xs. dverts (merge1 t2) \<subseteq> dverts t2" using Node by fastforce
then show ?thesis using False dverts_merge1_img_sub by auto
qed
qed
lemma disjoint_dlverts_merge1: "disjoint_dlverts ((\<lambda>(t,e). (merge1 t,e)) |`| (sucs t))"
proof -
have "\<forall>(t, e)\<in>fset (sucs t). dlverts (merge1 t) \<subseteq> dlverts t"
using ranked_dtree.merge1_dlverts_eq ranked_dtree_rec[of "root t"] by force
then show ?thesis using disjoint_dlverts_img[OF disjoint_dlverts_if_wf[OF wf_lverts]] by simp
qed
lemma root_empty_inter_dlverts_merge1:
assumes "(x1,e1) \<in> fset ((\<lambda>(t,e). (merge1 t,e)) |`| (sucs t))"
shows "set (root t) \<inter> dlverts x1 = {}"
proof (rule ccontr)
assume asm: "set (root t) \<inter> dlverts x1 \<noteq> {}"
obtain x2 where x2_def: "merge1 x2 = x1" "(x2,e1) \<in> fset (sucs t)" using assms by auto
then interpret X: ranked_dtree x2 using ranked_dtree_rec dtree.collapse by blast
have "set (root t) \<inter> dlverts x2 \<noteq> {}" using X.merge1_dlverts_eq x2_def(1) asm by argo
then show False using x2_def(2) wf_lverts wf_dlverts.simps[of "root t" "sucs t"] by auto
qed
lemma wf_dlverts_merge1: "wf_dlverts (merge1 t)"
using ranked_dtree_axioms proof(induction t)
case (Node r xs)
then interpret R: ranked_dtree "Node r xs" rank by blast
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then show ?thesis using R.merge_wf_dlverts by simp
next
case False
have "(\<forall>(t,e) \<in> fset ((\<lambda>(t,e). (merge1 t,e)) |`| xs). set r \<inter> dlverts t = {} \<and> wf_dlverts t)"
using R.ranked_dtree_rec Node.IH R.root_empty_inter_dlverts_merge1 by fastforce
then show ?thesis using R.disjoint_dlverts_merge1 R.wf_lverts False by auto
qed
qed
lemma merge1_darcs_eq[simp]: "darcs (merge1 t) = darcs t"
using ranked_dtree_axioms proof(induction t)
case (Node r xs)
then interpret R: ranked_dtree "Node r xs" rank by blast
show ?case using Node.IH R.ranked_dtree_rec by auto
qed
lemma disjoint_darcs_merge1: "disjoint_darcs ((\<lambda>(t,e). (merge1 t,e)) |`| (sucs t))"
proof -
have "\<forall>(t, e)\<in>fset (sucs t). darcs (merge1 t) \<subseteq> darcs t"
using ranked_dtree.merge1_darcs_eq ranked_dtree_rec[of "root t"] by force
then show ?thesis using disjoint_darcs_img[OF disjoint_darcs_if_wf[OF wf_arcs]] by simp
qed
lemma wf_darcs_merge1: "wf_darcs (merge1 t)"
using ranked_dtree_axioms proof(induction t)
case (Node r xs)
then interpret R: ranked_dtree "Node r xs" rank by blast
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then show ?thesis using R.merge_wf_darcs by simp
next
case False
then show ?thesis
using R.disjoint_darcs_merge1 R.ranked_dtree_rec Node.IH
by (auto simp: wf_darcs_iff_darcs')
qed
qed
theorem ranked_dtree_merge1: "ranked_dtree (merge1 t) cmp"
by(unfold_locales) (auto simp: wf_darcs_merge1 wf_dlverts_merge1 dest: cmp_antisym)
lemma distinct_merge1:
"\<lbrakk>\<forall>v\<in>dverts t. distinct v; v\<in>dverts (merge1 t)\<rbrakk> \<Longrightarrow> distinct v"
using ranked_dtree_axioms proof(induction t arbitrary: v rule: merge1.induct)
case (1 r xs)
then interpret R: ranked_dtree "Node r xs" rank by blast
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then show ?thesis using R.distinct_merge[OF "1.prems"(1)] "1.prems"(2) by simp
next
case ind: False
then show ?thesis
proof(cases "v = r")
case False
have "v\<in>dverts (merge1 (Node r xs)) \<longleftrightarrow> v \<in> dverts (Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs))"
using ind by auto
then obtain t e where t_def: "(t,e) \<in> fset xs" "v \<in> dverts (merge1 t)"
using False "1.prems"(2) by auto
then have "\<forall>v\<in>dverts t. distinct v" using "1.prems"(1) by force
then show ?thesis using "1.IH"[OF ind] t_def R.ranked_dtree_rec by fast
qed(simp add: "1.prems"(1))
qed
qed
lemma merge1_root_eq[simp]: "root (merge1 t1) = root t1"
by(induction t1) simp
lemma merge1_hd_root_eq[simp]: "hd (root (merge1 t1)) = hd (root t1)"
by simp
lemma merge1_mdeg_le: "max_deg (merge1 t1) \<le> max_deg t1"
proof(induction t1)
case (Node r xs)
then show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then have "max_deg (merge1 (Node r xs)) \<le> 1" using merge_mdeg_le_1 by simp
then show ?thesis using mdeg_ge_fcard[of xs] True by simp
next
case False
have 0: "\<forall>(t,e) \<in> fset xs. max_deg (merge1 t) \<le> max_deg t" using Node by force
have "merge1 (Node r xs) = (Node r ((\<lambda>(t, e). (merge1 t, e)) |`| xs))"
using False by auto
then show ?thesis using mdeg_img_le'[OF 0] by simp
qed
qed
lemma merge1_childdeg_gt1_if_fcard_gt1:
"fcard (sucs (merge1 t1)) > 1 \<Longrightarrow> \<exists>t \<in> fst ` fset (sucs t1). max_deg t > 1"
proof(induction t1)
case (Node r xs)
have 0: "\<not>(fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1))"
using merge_fcard_le1[of "Node r xs"] Node.prems(1) by fastforce
then have "fcard (sucs (merge1 (Node r xs))) \<le> fcard xs" using fcard_image_le by auto
then show ?case using 0 Node.prems(1) by fastforce
qed
lemma merge1_fcard_le: "fcard (sucs (merge1 (Node r xs))) \<le> fcard xs"
using fcard_image_le merge_fcard_le1[of "Node r xs"] by auto
lemma merge1_subtree_if_fcard_gt1:
"\<lbrakk>is_subtree (Node r xs) (merge1 t1); fcard xs > 1\<rbrakk>
\<Longrightarrow> \<exists>ys. merge1 (Node r ys) = Node r xs \<and> is_subtree (Node r ys) t1 \<and> fcard xs \<le> fcard ys"
proof(induction t1)
case (Node r1 xs1)
have 0: "\<not>(fcard xs1 > 1 \<and> (\<forall>t \<in> fst ` fset xs1. max_deg t \<le> 1))"
using merge_fcard_le1_sub Node.prems by fastforce
then have eq: "merge1 (Node r1 xs1) = Node r1 ((\<lambda>(t,e). (merge1 t,e)) |`| xs1)" by auto
show ?case
proof(cases "Node r xs = merge1 (Node r1 xs1)")
case True
moreover have "r = r1" using True eq by auto
moreover have "fcard xs \<le> fcard xs1" using merge1_fcard_le True dtree.sel(2)[of r xs] by auto
ultimately show ?thesis using self_subtree Node.prems(2) by auto
next
case False
then obtain t2 e2 where "(t2,e2) \<in> fset xs1" "is_subtree (Node r xs) (merge1 t2)"
using eq Node.prems(1) by auto
then show ?thesis using Node.IH[of "(t2,e2)" t2] Node.prems(2) by fastforce
qed
qed
lemma merge1_childdeg_gt1_if_fcard_gt1_sub:
"\<lbrakk>is_subtree (Node r xs) (merge1 t1); fcard xs > 1\<rbrakk>
\<Longrightarrow> \<exists>ys. merge1 (Node r ys) = Node r xs \<and> is_subtree (Node r ys) t1
\<and> (\<exists>t \<in> fst ` fset ys. max_deg t > 1)"
using merge1_subtree_if_fcard_gt1 merge1_childdeg_gt1_if_fcard_gt1 dtree.sel(2) by metis
lemma merge1_img_eq: "\<forall>(t2,e2) \<in> fset xs. merge1 t2 = t2 \<Longrightarrow> ((\<lambda>(t,e). (merge1 t,e)) |`| xs) = xs"
using fset_img_eq[of xs "\<lambda>(t,e). (merge1 t,e)"] by force
lemma merge1_wedge_if_uneq:
"merge1 t1 \<noteq> t1
\<Longrightarrow> \<exists>r xs. is_subtree (Node r xs) t1 \<and> fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)"
proof(induction t1)
case (Node r xs)
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then show ?thesis by auto
next
case False
then have "merge1 (Node r xs) = Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs)" by auto
then obtain t2 e2 where "(t2,e2) \<in> fset xs" "merge1 t2 \<noteq> t2"
using Node.prems merge1_img_eq[of xs] by auto
then show ?thesis using Node.IH[of "(t2,e2)"] by auto
qed
qed
lemma merge1_mdeg_gt1_if_uneq:
assumes "merge1 t1 \<noteq> t1"
shows "max_deg t1 > 1"
proof -
obtain r xs where r_def: "is_subtree (Node r xs) t1" "1 < fcard xs"
using merge1_wedge_if_uneq[OF assms] by fast
then show ?thesis using mdeg_ge_fcard[of xs] mdeg_ge_sub by force
qed
corollary merge1_eq_if_mdeg_le1: "max_deg t1 \<le> 1 \<Longrightarrow> merge1 t1 = t1"
using merge1_mdeg_gt1_if_uneq by fastforce
lemma merge1_not_merge_if_fcard_gt1:
"\<lbrakk>merge1 (Node r ys) = Node r xs; fcard xs > 1\<rbrakk> \<Longrightarrow> merge (Node r ys) \<noteq> Node r xs"
using merge_fcard_le1[of "Node r ys"] by auto
lemma merge1_img_if_not_merge:
"merge1 (Node r xs) \<noteq> merge (Node r xs)
\<Longrightarrow> merge1 (Node r xs) = Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs)"
by auto
lemma merge1_img_if_fcard_gt1:
"\<lbrakk>merge1 (Node r ys) = Node r xs; fcard xs > 1\<rbrakk>
\<Longrightarrow> merge1 (Node r ys) = Node r ((\<lambda>(t,e). (merge1 t,e)) |`| ys)"
using merge1_img_if_not_merge merge1_not_merge_if_fcard_gt1[of r ys] by simp
lemma merge1_elem_in_img_if_fcard_gt1:
"\<lbrakk>merge1 (Node r ys) = Node r xs; fcard xs > 1; (t2,e2) \<in> fset xs\<rbrakk>
\<Longrightarrow> \<exists>t1. (t1,e2) \<in> fset ys \<and> merge1 t1 = t2"
using merge1_img_if_fcard_gt1 by fastforce
lemma child_mdeg_gt1_if_sub_fcard_gt1:
"\<lbrakk>is_subtree (Node r xs) (Node v ys); Node r xs \<noteq> Node v ys; fcard xs > 1\<rbrakk>
\<Longrightarrow> \<exists>t1 e2. (t1,e2) \<in> fset ys \<and> max_deg t1 > 1"
using mdeg_ge_fcard[of xs] mdeg_ge_sub by force
lemma merge1_subtree_if_mdeg_gt1:
"\<lbrakk>is_subtree (Node r xs) (merge1 t1); max_deg (Node r xs) > 1\<rbrakk>
\<Longrightarrow> \<exists>ys. merge1 (Node r ys) = Node r xs \<and> is_subtree (Node r ys) t1"
proof(induction t1)
case (Node r1 xs1)
then have 0: "\<not>(fcard xs1 > 1 \<and> (\<forall>t \<in> fst ` fset xs1. max_deg t \<le> 1))"
using merge_mdeg_le1_sub by fastforce
then have eq: "merge1 (Node r1 xs1) = Node r1 ((\<lambda>(t,e). (merge1 t,e)) |`| xs1)" by auto
show ?case
proof(cases "Node r xs = merge1 (Node r1 xs1)")
case True
moreover have "r = r1" using True eq by auto
moreover have "fcard xs \<le> fcard xs1" using merge1_fcard_le True dtree.sel(2)[of r xs] by auto
ultimately show ?thesis using self_subtree Node.prems(2) by auto
next
case False
then obtain t2 e2 where "(t2,e2) \<in> fset xs1" "is_subtree (Node r xs) (merge1 t2)"
using eq Node.prems(1) by auto
then show ?thesis using Node.IH[of "(t2,e2)" t2] Node.prems(2) by fastforce
qed
qed
lemma merge1_child_in_orig:
assumes "merge1 (Node r ys) = Node r xs" and "(t1,e1) \<in> fset xs"
shows "\<exists>t2. (t2,e1) \<in> fset ys \<and> root t2 = root t1"
proof(cases "fcard ys > 1 \<and> (\<forall>t \<in> fst ` fset ys. max_deg t \<le> 1)")
case True
then show ?thesis using merge_child_in_orig[of t1 e1 "Node r ys"] assms by auto
next
case False
then have "merge1 (Node r ys) = Node r ((\<lambda>(t,e). (merge1 t,e)) |`| ys)" by auto
then show ?thesis using assms by fastforce
qed
lemma dverts_if_subtree_merge1:
"is_subtree (Node r xs) (merge1 t1) \<Longrightarrow> r \<in> dverts t1"
using merge1_dverts_sub dverts_subtree_subset by fastforce
lemma subtree_merge1_orig:
"is_subtree (Node r xs) (merge1 t1) \<Longrightarrow> \<exists>ys. is_subtree (Node r ys) t1"
using dverts_if_subtree_merge1 subtree_root_if_dverts by fast
lemma merge1_subtree_dlverts_supset:
"is_subtree (Node r xs) (merge1 t)
\<Longrightarrow> \<exists>ys. is_subtree (Node r ys) t \<and> dlverts (Node r ys) \<subseteq> dlverts (Node r xs)"
using ranked_dtree_axioms proof(induction t)
case (Node r1 xs1)
then interpret R: ranked_dtree "Node r1 xs1" by simp
show ?case
proof(cases "Node r xs = merge1 (Node r1 xs1)")
case True
then have "dlverts (Node r1 xs1) \<subseteq> dlverts (Node r xs)" using R.merge1_dlverts_eq by simp
moreover have "r = r1" using True dtree.sel(1)[of r xs] by auto
ultimately show ?thesis by auto
next
case uneq: False
show ?thesis
proof(cases "fcard xs1 > 1 \<and> (\<forall>t \<in> fst ` fset xs1. max_deg t \<le> 1)")
case True
then show ?thesis using R.merge_subtree_dlverts_supset Node.prems by simp
next
case False
then have eq: "merge1 (Node r1 xs1) = Node r1 ((\<lambda>(t,e). (merge1 t,e)) |`| xs1)" by auto
then obtain t2 e2 where "(t2,e2) \<in> fset xs1" "is_subtree (Node r xs) (merge1 t2)"
using Node.prems(1) uneq by auto
then show ?thesis using Node.IH[of "(t2,e2)"] R.ranked_dtree_rec by auto
qed
qed
qed
end
subsection \<open>IKKBZ-Sub\<close>
function denormalize :: "('a list, 'b) dtree \<Rightarrow> 'a list" where
"denormalize (Node r {|(t,e)|}) = r @ denormalize t"
| "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> denormalize (Node r xs) = r"
using dtree_to_list.cases by blast+
termination by lexicographic_order
lemma denormalize_set_eq_dlverts: "max_deg t1 \<le> 1 \<Longrightarrow> set (denormalize t1) = dlverts t1"
proof(induction t1 rule: denormalize.induct)
case (1 r t e)
then show ?case using mdeg_ge_child[of t e "{|(t, e)|}"] by force
next
case (2 xs r)
then have "max_deg (Node r xs) = 0" using mdeg_1_singleton[of r xs] by fastforce
then have "xs = {||}" by (auto intro!: empty_if_mdeg_0)
then show ?case using 2 by auto
qed
lemma denormalize_set_sub_dlverts: "set (denormalize t1) \<subseteq> dlverts t1"
by(induction t1 rule: denormalize.induct) auto
lemma denormalize_distinct:
"\<lbrakk>\<forall>v \<in> dverts t1. distinct v; wf_dlverts t1\<rbrakk> \<Longrightarrow> distinct (denormalize t1)"
proof(induction t1 rule: denormalize.induct)
case (1 r t e)
then have "set r \<inter> set (denormalize t) = {}" using denormalize_set_sub_dlverts by fastforce
then show ?case using 1 by auto
next
case (2 xs r)
then show ?case by simp
qed
lemma denormalize_hd_root:
assumes "root t \<noteq> []"
shows "hd (denormalize t) = hd (root t)"
proof(cases "\<forall>x. sucs t \<noteq> {|x|}")
case True
then show ?thesis using denormalize.simps(2)[of "sucs t" "root t"] by simp
next
case False
then obtain t1 e where "{|(t1, e)|} = sucs t" by auto
then show ?thesis using denormalize.simps(1)[of "root t" t1 e] assms by simp
qed
lemma denormalize_hd_root_wf: "wf_dlverts t \<Longrightarrow> hd (denormalize t) = hd (root t)"
using denormalize_hd_root empty_notin_wf_dlverts dtree.set_sel(1)[of t] by force
lemma denormalize_nempty_if_wf: "wf_dlverts t \<Longrightarrow> denormalize t \<noteq> []"
by (induction t rule: denormalize.induct) auto
context ranked_dtree
begin
lemma fcard_normalize_img_if_disjoint:
"disjoint_darcs xs \<Longrightarrow> fcard ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) = fcard xs"
using snds_neq_img_card_eq[of xs] by fast
lemma fcard_merge1_img_if_disjoint:
"disjoint_darcs xs \<Longrightarrow> fcard ((\<lambda>(t,e). (merge1 t,e)) |`| xs) = fcard xs"
using snds_neq_img_card_eq[of xs] by fast
lemma fsts_uneq_if_disjoint_lverts_nempty:
"\<lbrakk>disjoint_dlverts xs; \<forall>(t, e)\<in>fset xs. dlverts t \<noteq> {}\<rbrakk>
\<Longrightarrow> \<forall>(t, e)\<in>fset xs. \<forall>(t2, e2)\<in>fset xs. t \<noteq> t2 \<or> (t, e) = (t2, e2)"
by fast
lemma normalize1_dlverts_nempty:
"\<forall>(t, e)\<in>fset xs. dlverts t \<noteq> {}
\<Longrightarrow> \<forall>(t, e)\<in>fset ((\<lambda>(t, e). (normalize1 t, e)) |`| xs). dlverts t \<noteq> {}"
by auto
lemma normalize1_fsts_uneq:
assumes "disjoint_dlverts xs" and "\<forall>(t, e)\<in>fset xs. dlverts t \<noteq> {}"
shows "\<forall>(t, e)\<in>fset xs. \<forall>(t2, e2)\<in>fset xs. normalize1 t \<noteq> normalize1 t2 \<or> (t,e) = (t2,e2)"
by (smt (verit) assms Int_absorb case_prodD case_prodI2 normalize1_dlverts_eq)
lemma fcard_normalize_img_if_disjoint_lverts:
"\<lbrakk>disjoint_dlverts xs; \<forall>(t, e)\<in>fset xs. dlverts t \<noteq> {}\<rbrakk>
\<Longrightarrow> fcard ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) = fcard xs"
using fst_neq_img_card_eq[of xs normalize1] normalize1_fsts_uneq by auto
lemma fcard_normalize_img_if_wf_dlverts:
"wf_dlverts (Node r xs) \<Longrightarrow> fcard ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) = fcard xs"
using dlverts_nempty_if_wf fcard_normalize_img_if_disjoint_lverts[of xs] by force
lemma fcard_normalize_img_if_wf_dlverts_sucs:
"wf_dlverts t1 \<Longrightarrow> fcard ((\<lambda>(t,e). (normalize1 t,e)) |`| (sucs t1)) = fcard (sucs t1)"
using fcard_normalize_img_if_wf_dlverts[of "root t1" "sucs t1"] by simp
lemma singleton_normalize1:
assumes "disjoint_darcs xs" and "\<forall>x. xs \<noteq> {|x|}"
shows "\<forall>x. (\<lambda>(t,e). (normalize1 t,e)) |`| xs \<noteq> {|x|}"
proof (rule ccontr)
assume "\<not>(\<forall>x. (\<lambda>(t,e). (normalize1 t,e)) |`| xs \<noteq> {|x|})"
then obtain x where "(\<lambda>(t,e). (normalize1 t,e)) |`| xs = {|x|}" by blast
then have "fcard ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) = 1" using fcard_single_1 by force
then have "fcard xs = 1" using fcard_normalize_img_if_disjoint[OF assms(1)] by simp
then have "\<exists>x. xs = {|x|}" using fcard_single_1_iff by fast
then show False using assms(2) by simp
qed
lemma num_leaves_normalize1_eq[simp]: "wf_darcs t1 \<Longrightarrow> num_leaves (normalize1 t1) = num_leaves t1"
proof(induction t1)
case (Node r xs)
then show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case True
have "fcard ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) = fcard xs"
using fcard_normalize_img_if_disjoint Node.prems
by (auto simp: wf_darcs_iff_darcs')
moreover have "\<forall>t\<in>fst ` fset xs. num_leaves (normalize1 t) = num_leaves t"
using Node by fastforce
ultimately show ?thesis using Node sum_img_eq[of xs] True by force
next
case False
then obtain t e where t_def: "xs = {|(t,e)|}" by auto
show ?thesis
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then show ?thesis
using t_def num_leaves_singleton num_leaves_root[of "root t" "sucs t"] by simp
next
case False
then show ?thesis
using num_leaves_singleton t_def Node by (simp add: wf_darcs_iff_darcs')
qed
qed
qed
lemma num_leaves_normalize_eq[simp]: "wf_darcs t1 \<Longrightarrow> num_leaves (normalize t1) = num_leaves t1"
proof(induction t1 rule: normalize.induct)
case (1 t)
then have "num_leaves (normalize1 t) = num_leaves t" using num_leaves_normalize1_eq by blast
then show ?case using 1 wf_darcs_normalize1 by (smt (verit, best) normalize.simps)
qed
lemma num_leaves_normalize1_le: "num_leaves (normalize1 t1) \<le> num_leaves t1"
proof(induction t1)
case (Node r xs)
then show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case True
have fcard_le: "fcard ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) \<le> fcard xs"
by (simp add: fcard_image_le)
moreover have xs_le: "\<forall>t\<in>fst ` fset xs. num_leaves (normalize1 t) \<le> num_leaves t"
using Node by fastforce
ultimately show ?thesis using Node sum_img_le[of xs] xs_le \<open>\<forall>x. xs \<noteq> {|x|}\<close> by simp
next
case False
then obtain t e where t_def: "xs = {|(t,e)|}" by auto
show ?thesis
proof(cases "rank (rev (root t)) < rank (rev r)")
case True
then show ?thesis
using t_def num_leaves_singleton num_leaves_root[of "root t" "sucs t"] by simp
next
case False
then show ?thesis using num_leaves_singleton t_def Node by simp
qed
qed
qed
lemma num_leaves_normalize_le: "num_leaves (normalize t1) \<le> num_leaves t1"
proof(induction t1 rule: normalize.induct)
case (1 t)
then have "num_leaves (normalize1 t) \<le> num_leaves t" using num_leaves_normalize1_le by blast
then show ?case using 1 by (smt (verit) le_trans normalize.simps)
qed
lemma num_leaves_merge1_le: "num_leaves (merge1 t1) \<le> num_leaves t1"
proof(induction t1)
case (Node r xs)
then show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then have "merge1 (Node r xs) = merge (Node r xs)" by simp
then have "num_leaves (merge1 (Node r xs)) = 1"
unfolding merge_def using dtree_from_list_1_leaf by fastforce
also have "\<dots> < fcard xs" using True by blast
also have "\<dots> \<le> num_leaves (Node r xs)" using num_leaves_ge_card by fast
finally show ?thesis by simp
next
case False
have "\<forall>t \<in> fst ` fset xs. num_leaves (merge1 t) \<le> num_leaves t" using Node by force
then show ?thesis using sum_img_le False by auto
qed
qed
lemma num_leaves_merge1_lt: "max_deg t1 > 1 \<Longrightarrow> num_leaves (merge1 t1) < num_leaves t1"
proof(induction t1)
case (Node r xs)
show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then have "merge1 (Node r xs) = merge (Node r xs)" by simp
then have "num_leaves (merge1 (Node r xs)) = 1"
unfolding merge_def using dtree_from_list_1_leaf by fastforce
also have "\<dots> < fcard xs" using True by blast
finally show ?thesis using num_leaves_ge_card less_le_trans by fast
next
case False
have 0: "xs \<noteq> {||}" using Node.prems by (metis nempty_if_mdeg_n0 not_one_less_zero)
have 1: "\<forall>t \<in> fst ` fset xs. num_leaves (merge1 t) \<le> num_leaves t"
using num_leaves_merge1_le by blast
have "\<exists>t \<in> fst ` fset xs. max_deg t > 1" using Node.prems False mdeg_child_if_wedge by auto
then have 2: "\<exists>t \<in> fst ` fset xs. num_leaves (merge1 t) < num_leaves t" using Node.IH by force
have 3: "\<forall>t\<in>fst ` fset xs. 0 < num_leaves t"
using num_leaves_ge1 by (metis neq0_conv not_one_le_zero)
from False have "merge1 (Node r xs) = Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs)" by auto
then have "num_leaves (merge1 (Node r xs))
= (\<Sum>(t,e)\<in> fset ((\<lambda>(t,e). (merge1 t,e)) |`| xs). num_leaves t)" using 0 by auto
then show ?thesis using 0 sum_img_lt[OF 1 2 3] by simp
qed
qed
lemma ikkbz_num_leaves_decr:
"max_deg t1 > 1 \<Longrightarrow> num_leaves (merge1 (normalize t1)) < num_leaves t1"
using num_leaves_merge1_lt num_leaves_normalize_le num_leaves_1_if_mdeg_1 num_leaves_ge1
by (metis antisym_conv2 dual_order.antisym dual_order.trans not_le_imp_less num_leaves_merge1_le)
function ikkbz_sub :: "('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"ikkbz_sub t1 = (if max_deg t1 \<le> 1 then t1 else ikkbz_sub (merge1 (normalize t1)))"
by auto
termination using ikkbz_num_leaves_decr by(relation "measure (\<lambda>t. num_leaves t)") auto
lemma ikkbz_sub_darcs_sub: "darcs (ikkbz_sub t) \<subseteq> darcs t"
using ranked_dtree_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
show ?case
proof(cases "max_deg t \<le> 1")
case False
have "darcs (merge1 (normalize t)) = darcs (normalize t)"
using ranked_dtree.merge1_darcs_eq ranked_dtree.ranked_dtree_normalize "1.prems" by blast
moreover have "ranked_dtree (merge1 (normalize t)) cmp"
using ranked_dtree.ranked_dtree_normalize "1.prems" ranked_dtree.ranked_dtree_merge1 by blast
moreover have "\<not> (max_deg t \<le> 1 \<or> \<not> list_dtree t)" using False ranked_dtree_def "1.prems" by blast
ultimately show ?thesis using "1.IH" normalize_darcs_sub by force
qed(simp)
qed
lemma ikkbz_sub_dlverts_eq[simp]: "dlverts (ikkbz_sub t) = dlverts t"
using ranked_dtree_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
show ?case
proof(cases "max_deg t \<le> 1")
case True
then show ?thesis by simp
next
case False
then show ?thesis
using 1 ranked_dtree.merge1_dlverts_eq[of "normalize t"] normalize_dlverts_eq
ranked_dtree.ranked_dtree_normalize ranked_dtree.ranked_dtree_merge1 ikkbz_sub.elims by metis
qed
qed
lemma ikkbz_sub_wf_darcs: "wf_darcs (ikkbz_sub t)"
using ranked_dtree_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
then show ?case
proof(cases "max_deg t \<le> 1")
case True
then show ?thesis using "1.prems" list_dtree_def ranked_dtree_def by auto
next
case False
then show ?thesis
using 1 ranked_dtree.ranked_dtree_normalize ranked_dtree.ranked_dtree_merge1
by (metis ikkbz_sub.simps)
qed
qed
lemma ikkbz_sub_wf_dlverts: "wf_dlverts (ikkbz_sub t)"
using ranked_dtree_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
then show ?case
proof(cases "max_deg t \<le> 1")
case True
then show ?thesis using "1.prems" list_dtree_def ranked_dtree_def by auto
next
case False
then show ?thesis
using 1 ranked_dtree.ranked_dtree_normalize ranked_dtree.ranked_dtree_merge1
by (metis ikkbz_sub.simps)
qed
qed
theorem ikkbz_sub_list_dtree: "list_dtree (ikkbz_sub t)"
using ikkbz_sub_wf_darcs ikkbz_sub_wf_dlverts list_dtree_def by blast
corollary ikkbz_sub_ranked_dtree: "ranked_dtree (ikkbz_sub t) cmp"
using ikkbz_sub_list_dtree ranked_dtree_def ranked_dtree_axioms by blast
lemma ikkbz_sub_mdeg_le1: "max_deg (ikkbz_sub t1) \<le> 1"
by (induction t1 rule: ikkbz_sub.induct) simp
corollary denormalize_ikkbz_eq_dlverts: "set (denormalize (ikkbz_sub t)) = dlverts t"
using denormalize_set_eq_dlverts ikkbz_sub_mdeg_le1 ikkbz_sub_dlverts_eq by blast
lemma distinct_ikkbz_sub: "\<lbrakk>\<forall>v\<in>dverts t. distinct v; v\<in>dverts (ikkbz_sub t)\<rbrakk> \<Longrightarrow> distinct v"
using list_dtree_axioms proof(induction t arbitrary: v rule: ikkbz_sub.induct)
case (1 t)
then interpret T1: ranked_dtree t rank cmp
using ranked_dtree_axioms by (simp add: ranked_dtree_def)
show ?case
using 1 T1.ranked_dtree_normalize T1.distinct_normalize ranked_dtree.merge1_dverts_eq
ranked_dtree.wf_dlverts_merge1 ranked_dtree.wf_darcs_merge1
by (metis ikkbz_sub.elims list_dtree_def)
qed
corollary distinct_denormalize_ikkbz_sub:
"\<forall>v\<in>dverts t. distinct v \<Longrightarrow> distinct (denormalize (ikkbz_sub t))"
using distinct_ikkbz_sub ikkbz_sub_wf_dlverts denormalize_distinct by blast
lemma ikkbz_sub_hd_root[simp]: "hd (root (ikkbz_sub t)) = hd (root t)"
using list_dtree_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
then interpret T1: ranked_dtree t rank cmp
using ranked_dtree_axioms by (simp add: ranked_dtree_def)
show ?case
using 1 merge1_hd_root_eq ranked_dtree.axioms(1) ranked_dtree.ranked_dtree_merge1
by (metis T1.ranked_dtree_normalize T1.wf_lverts ikkbz_sub.simps normalize_hd_root_eq')
qed
corollary denormalize_ikkbz_sub_hd_root[simp]: "hd (denormalize (ikkbz_sub t)) = hd (root t)"
using ikkbz_sub_hd_root denormalize_hd_root
by (metis dtree.set_sel(1) empty_notin_wf_dlverts ikkbz_sub_wf_dlverts)
end
locale precedence_graph = finite_directed_tree +
fixes rank :: "'a list \<Rightarrow> real"
fixes cost :: "'a list \<Rightarrow> real"
fixes cmp :: "('a list\<times>'b) comparator"
assumes asi_rank: "asi rank root cost"
and cmp_antisym:
"\<lbrakk>v1 \<noteq> []; v2 \<noteq> []; compare cmp (v1,e1) (v2,e2) = Equiv\<rbrakk> \<Longrightarrow> set v1 \<inter> set v2 \<noteq> {} \<or> e1=e2"
begin
definition to_list_dtree :: "('a list, 'b) dtree" where
"to_list_dtree = finite_directed_tree.to_dtree to_list_tree [root]"
lemma to_list_dtree_single: "v \<in> dverts to_list_dtree \<Longrightarrow> \<exists>x. v = [x] \<and> x \<in> verts T"
unfolding to_list_dtree_def using to_list_tree_single
by (simp add: finite_directed_tree.dverts_eq_verts to_list_tree_finite_directed_tree)
lemma to_list_dtree_wf_dverts: "wf_dverts to_list_dtree"
using finite_directed_tree.wf_dverts_to_dtree[OF to_list_tree_finite_directed_tree]
by(simp add: to_list_dtree_def)
lemma to_list_dtree_wf_dlverts: "wf_dlverts to_list_dtree"
unfolding to_list_dtree_def
by (simp add: to_list_tree_fin_list_directed_tree fin_list_directed_tree.wf_dlverts_to_dtree)
lemma to_list_dtree_wf_darcs: "wf_darcs to_list_dtree"
using finite_directed_tree.wf_darcs_to_dtree[OF to_list_tree_finite_directed_tree]
by(simp add: to_list_dtree_def)
lemma to_list_dtree_list_dtree: "list_dtree to_list_dtree"
by(simp add: list_dtree_def to_list_dtree_wf_dlverts to_list_dtree_wf_darcs)
lemma to_list_dtree_ranked_dtree: "ranked_dtree to_list_dtree cmp"
by(auto simp: ranked_dtree_def to_list_dtree_list_dtree ranked_dtree_axioms_def dest: cmp_antisym)
interpretation t: ranked_dtree to_list_dtree by (rule to_list_dtree_ranked_dtree)
definition ikkbz_sub :: "'a list" where
"ikkbz_sub = denormalize (t.ikkbz_sub to_list_dtree)"
lemma dverts_eq_verts_to_list_tree: "dverts to_list_dtree = pre_digraph.verts to_list_tree"
unfolding to_list_dtree_def
by (simp add: finite_directed_tree.dverts_eq_verts to_list_tree_finite_directed_tree)
lemma dverts_eq_verts_img: "dverts to_list_dtree = (\<lambda>x. [x]) ` verts T"
by (simp add: dverts_eq_verts_to_list_tree to_list_tree_def)
lemma dlverts_eq_verts: "dlverts to_list_dtree = verts T"
by (simp add: dverts_eq_verts_img dlverts_eq_dverts_union)
theorem ikkbz_set_eq_verts: "set ikkbz_sub = verts T"
using dlverts_eq_verts ikkbz_sub_def t.denormalize_ikkbz_eq_dlverts by simp
lemma distinct_to_list_tree: "\<forall>v\<in>verts to_list_tree. distinct v"
unfolding to_list_tree_def by simp
lemma distinct_to_list_dtree: "\<forall>v\<in>dverts to_list_dtree. distinct v"
using distinct_to_list_tree dverts_eq_verts_to_list_tree by blast
theorem distinct_ikkbz_sub: "distinct ikkbz_sub"
unfolding ikkbz_sub_def
using distinct_to_list_dtree t.distinct_denormalize_ikkbz_sub by blast
lemma to_list_dtree_root_eq_root: "Dtree.root (to_list_dtree) = [root]"
unfolding to_list_dtree_def
by (simp add: finite_directed_tree.to_dtree_root_eq_root to_list_tree_finite_directed_tree)
lemma to_list_dtree_hd_root_eq_root[simp]: "hd (Dtree.root to_list_dtree) = root"
by (simp add: to_list_dtree_root_eq_root)
theorem ikkbz_sub_hd_eq_root[simp]: "hd ikkbz_sub = root"
unfolding ikkbz_sub_def using t.denormalize_ikkbz_sub_hd_root to_list_dtree_root_eq_root by simp
end
subsection \<open>Full IKKBZ\<close>
locale tree_query_graph = undir_tree_todir G + query_graph G for G
locale cmp_tree_query_graph = tree_query_graph +
fixes cmp :: "('a list\<times>'b) comparator"
assumes cmp_antisym:
"\<lbrakk>v1 \<noteq> []; v2 \<noteq> []; compare cmp (v1,e1) (v2,e2) = Equiv\<rbrakk> \<Longrightarrow> set v1 \<inter> set v2 \<noteq> {} \<or> e1=e2"
locale ikkbz_query_graph = cmp_tree_query_graph +
fixes cost :: "'a joinTree \<Rightarrow> real"
fixes cost_r :: "'a \<Rightarrow> ('a list \<Rightarrow> real)"
fixes rank_r :: "'a \<Rightarrow> ('a list \<Rightarrow> real)"
assumes asi_rank: "r \<in> verts G \<Longrightarrow> asi (rank_r r) r (cost_r r)"
and cost_correct:
"\<lbrakk>valid_tree t; no_cross_products t; left_deep t\<rbrakk>
\<Longrightarrow> cost_r (first_node t) (revorder t) = cost t"
begin
abbreviation ikkbz_sub :: "'a \<Rightarrow> 'a list" where
"ikkbz_sub r \<equiv> precedence_graph.ikkbz_sub (dir_tree_r r) r (rank_r r) cmp"
abbreviation cost_l :: "'a list \<Rightarrow> real" where
"cost_l xs \<equiv> cost (create_ldeep xs)"
lemma precedence_graph_r:
"r \<in> verts G \<Longrightarrow> precedence_graph (dir_tree_r r) r (rank_r r) (cost_r r) cmp"
using fin_directed_tree_r cmp_antisym
by (simp add: precedence_graph_def precedence_graph_axioms_def asi_rank)
lemma nempty_if_set_eq_verts: "set xs = verts G \<Longrightarrow> xs \<noteq> []"
using verts_nempty by force
lemma revorder_if_set_eq_verts: "set xs = verts G \<Longrightarrow> revorder (create_ldeep xs) = rev xs"
using nempty_if_set_eq_verts create_ldeep_order unfolding revorder_eq_rev_inorder by blast
lemma cost_correct':
"\<lbrakk>set xs = verts G; distinct xs; no_cross_products (create_ldeep xs)\<rbrakk>
\<Longrightarrow> cost_r (hd xs) (rev xs) = cost_l xs"
using cost_correct[of "create_ldeep xs"] revorder_if_set_eq_verts create_ldeep_ldeep[of xs]
unfolding valid_tree_def distinct_relations_def
by (simp add: create_ldeep_order create_ldeep_relations first_node_eq_hd nempty_if_set_eq_verts)
lemma ikkbz_sub_verts_eq: "r \<in> verts G \<Longrightarrow> set (ikkbz_sub r) = verts G"
using precedence_graph.ikkbz_set_eq_verts precedence_graph_r verts_dir_tree_r_eq by fast
lemma ikkbz_sub_distinct: "r \<in> verts G \<Longrightarrow> distinct (ikkbz_sub r)"
using precedence_graph.distinct_ikkbz_sub precedence_graph_r by fast
lemma ikkbz_sub_hd_eq_root: "r \<in> verts G \<Longrightarrow> hd (ikkbz_sub r) = r"
using precedence_graph.ikkbz_sub_hd_eq_root precedence_graph_r by fast
definition ikkbz :: "'a list" where
"ikkbz \<equiv> arg_min_on cost_l {ikkbz_sub r|r. r \<in> verts G}"
lemma ikkbz_sub_set_fin: "finite {ikkbz_sub r|r. r \<in> verts G}"
by simp
lemma ikkbz_sub_set_nempty: "{ikkbz_sub r|r. r \<in> verts G} \<noteq> {}"
by (simp add: verts_nempty)
lemma ikkbz_in_ikkbz_sub_set: "ikkbz \<in> {ikkbz_sub r|r. r \<in> verts G}"
unfolding ikkbz_def using ikkbz_sub_set_fin ikkbz_sub_set_nempty arg_min_if_finite by blast
lemma ikkbz_eq_ikkbz_sub: "\<exists>r \<in> verts G. ikkbz = ikkbz_sub r"
using ikkbz_in_ikkbz_sub_set by blast
lemma ikkbz_min_ikkbz_sub: "r \<in> verts G \<Longrightarrow> cost_l ikkbz \<le> cost_l (ikkbz_sub r)"
unfolding ikkbz_def using ikkbz_sub_set_fin arg_min_least by fast
lemma ikkbz_distinct: "distinct ikkbz"
using ikkbz_eq_ikkbz_sub ikkbz_sub_distinct by fastforce
lemma ikkbz_set_eq_verts: "set ikkbz = verts G"
using ikkbz_eq_ikkbz_sub ikkbz_sub_verts_eq by force
lemma ikkbz_nempty: "ikkbz \<noteq> []"
using ikkbz_set_eq_verts verts_nempty by fastforce
lemma ikkbz_hd_in_verts: "hd ikkbz \<in> verts G"
using ikkbz_nempty ikkbz_set_eq_verts by fastforce
lemma inorder_ikkbz: "inorder (create_ldeep ikkbz) = ikkbz"
using create_ldeep_order ikkbz_nempty by blast
lemma inorder_ikkbz_distinct: "distinct (inorder (create_ldeep ikkbz))"
using ikkbz_distinct inorder_ikkbz by simp
lemma inorder_relations_eq_verts: "relations (create_ldeep ikkbz) = verts G"
using ikkbz_set_eq_verts create_ldeep_relations ikkbz_nempty by blast
theorem ikkbz_valid_tree: "valid_tree (create_ldeep ikkbz)"
unfolding valid_tree_def distinct_relations_def
using inorder_ikkbz_distinct inorder_relations_eq_verts by blast
end
(* non commutative merging based on inserting (INCOMPLETE) *)
locale old = list_dtree t for t :: "('a list,'b) dtree" +
fixes rank :: "'a list \<Rightarrow> real"
begin
function find_pos_aux :: "'a list \<Rightarrow> 'a list \<Rightarrow> ('a list,'b) dtree \<Rightarrow> ('a list \<times> 'a list)" where
"find_pos_aux v p (Node r {|(t1,_)|}) =
(if rank (rev v) \<le> rank (rev r) then (p,r) else find_pos_aux v r t1)"
| "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> find_pos_aux v p (Node r xs) =
(if rank (rev v) \<le> rank (rev r) then (p,r) else (r,r))"
by (metis combine.cases old.prod.exhaust) auto
termination by lexicographic_order
function find_pos :: "'a list \<Rightarrow> ('a list,'b) dtree \<Rightarrow> ('a list \<times> 'a list)" where
"find_pos v (Node r {|(t1,_)|}) = find_pos_aux v r t1"
| "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> find_pos v (Node r xs) = (r,r)"
by (metis dtree.exhaust surj_pair) auto
termination by lexicographic_order
abbreviation insert_chain :: "('a list\<times>'b) list \<Rightarrow> ('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"insert_chain xs t1 \<equiv>
foldr (\<lambda>(v,e) t2. case find_pos v t2 of (x,y) \<Rightarrow> insert_between v e x y t2) xs t1"
fun merge :: "('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"merge (Node r xs) = ffold (\<lambda>(t,e) b. case b of Node r xs \<Rightarrow>
if xs = {||} then Node r {|(t,e)|} else insert_chain (dtree_to_list t) b)
(Node r {||}) xs"
lemma ffold_if_False_eq_acc:
"\<lbrakk>\<forall>a. \<not>P a; comp_fun_commute (\<lambda>a b. if \<not>P a then b else Q a b)\<rbrakk>
\<Longrightarrow> ffold (\<lambda>a b. if \<not>P a then b else Q a b) acc xs = acc"
proof(induction xs)
case (insert x xs)
let ?f = "\<lambda>a b. if \<not>P a then b else Q a b"
have "ffold ?f acc (finsert x xs) = ?f x (ffold ?f acc xs)"
using insert.hyps by (simp add: comp_fun_commute.ffold_finsert insert.prems(2))
then have "ffold ?f acc (finsert x xs) = ffold ?f acc xs" using insert.prems by simp
then show ?case using insert.IH insert.prems by simp
qed(simp add: comp_fun_commute.ffold_empty)
lemma find_pos_rank_less: "rank (rev v) \<le> rank (rev r) \<Longrightarrow> find_pos_aux v p (Node r xs) = (p,r)"
by(cases "\<exists>x. xs = {|x|}") auto
lemma find_pos_y_in_dverts: "(x,y) = find_pos_aux v p t1 \<Longrightarrow> y \<in> dverts t1"
proof(induction t1 arbitrary: p)
case (Node r xs)
then show ?case
proof(cases "rank (rev v) \<le> rank (rev r)")
case True
then show ?thesis using Node.prems by(cases "\<exists>x. xs = {|x|}") auto
next
case False
then show ?thesis using Node by(cases "\<exists>x. xs = {|x|}") fastforce+
qed
qed
lemma find_pos_x_in_dverts: "(x,y) = find_pos_aux v p t1 \<Longrightarrow> x \<in> dverts t1 \<or> p=x"
proof(induction t1 arbitrary: p)
case (Node r xs)
then show ?case
proof(cases "rank (rev v) \<le> rank (rev r)")
case True
then show ?thesis using Node.prems by(cases "\<exists>x. xs = {|x|}") auto
next
case False
then show ?thesis using Node by(cases "\<exists>x. xs = {|x|}") fastforce+
qed
qed
end
end
\ No newline at end of file
diff --git a/thys/Query_Optimization/IKKBZ_Optimality.thy b/thys/Query_Optimization/IKKBZ_Optimality.thy
--- a/thys/Query_Optimization/IKKBZ_Optimality.thy
+++ b/thys/Query_Optimization/IKKBZ_Optimality.thy
@@ -1,6239 +1,6266 @@
(* Author: Bernhard Stöckl *)
theory IKKBZ_Optimality
imports Complex_Main "CostFunctions" "QueryGraph" "IKKBZ" "HOL-Library.Sublist"
begin
section \<open>Optimality of IKKBZ\<close>
context directed_tree
begin
fun forward_arcs :: "'a list \<Rightarrow> bool" where
"forward_arcs [] = True"
| "forward_arcs [x] = True"
| "forward_arcs (x#xs) = ((\<exists>y \<in> set xs. y \<rightarrow>\<^bsub>T\<^esub> x) \<and> forward_arcs xs)"
fun no_back_arcs :: "'a list \<Rightarrow> bool" where
"no_back_arcs [] = True"
| "no_back_arcs (x#xs) = ((\<nexists>y. y \<in> set xs \<and> y \<rightarrow>\<^bsub>T\<^esub> x) \<and> no_back_arcs xs)"
definition forward :: "'a list \<Rightarrow> bool" where
"forward xs = (\<forall>i \<in> {1..(length xs - 1)}. \<exists>j < i. xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i)"
definition no_back :: "'a list \<Rightarrow> bool" where
"no_back xs = (\<nexists>i j. i < j \<and> j < length xs \<and> xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i)"
definition seq_conform :: "'a list \<Rightarrow> bool" where
"seq_conform xs \<equiv> forward_arcs (rev xs) \<and> no_back_arcs xs"
definition before :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"before s1 s2 \<equiv> seq_conform s1 \<and> seq_conform s2 \<and> set s1 \<inter> set s2 = {}
\<and> (\<exists>x \<in> set s1. \<exists>y \<in> set s2. x \<rightarrow>\<^bsub>T\<^esub> y)"
definition before2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"before2 s1 s2 \<equiv> seq_conform s1 \<and> seq_conform s2 \<and> set s1 \<inter> set s2 = {}
\<and> (\<exists>x \<in> set s1. \<exists>y \<in> set s2. x \<rightarrow>\<^bsub>T\<^esub> y)
\<and> (\<forall>x \<in> set s1. \<forall>v \<in> verts T - set s1 - set s2. \<not> x \<rightarrow>\<^bsub>T\<^esub> v)"
lemma before_alt1:
"(\<exists>i < length s1. \<exists>j < length s2. s1!i \<rightarrow>\<^bsub>T\<^esub> s2!j) \<longleftrightarrow> (\<exists>x \<in> set s1. \<exists>y \<in> set s2. x \<rightarrow>\<^bsub>T\<^esub> y)"
using in_set_conv_nth by metis
lemma before_alt2:
"(\<forall>i < length s1. \<forall>v \<in> verts T - set s1 - set s2. \<not> s1!i \<rightarrow>\<^bsub>T\<^esub> v)
\<longleftrightarrow> (\<forall>x \<in> set s1. \<forall>v \<in> verts T - set s1 - set s2. \<not> x \<rightarrow>\<^bsub>T\<^esub> v)"
using in_set_conv_nth by metis
lemma no_back_alt_aux: "(\<forall>i j. i \<ge> j \<or> j \<ge> length xs \<or> \<not>(xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i)) \<Longrightarrow> no_back xs"
using less_le_not_le no_back_def by auto
lemma no_back_alt: "(\<forall>i j. i \<ge> j \<or> j \<ge> length xs \<or> \<not>(xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i)) \<longleftrightarrow> no_back xs"
using no_back_alt_aux by (auto simp: no_back_def)
lemma no_back_arcs_alt_aux1: "\<lbrakk>no_back_arcs xs; i < j; j < length xs\<rbrakk> \<Longrightarrow> \<not>(xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i)"
proof(induction xs arbitrary: i j)
case (Cons x xs)
then show ?case
proof(cases "i = 0")
case True
then show ?thesis using Cons.prems by simp
next
case False
then show ?thesis using Cons by auto
qed
qed(simp)
lemma no_back_insert_aux:
"(\<forall>i j. i \<ge> j \<or> j \<ge> length (x#xs) \<or> \<not>((x#xs)!j \<rightarrow>\<^bsub>T\<^esub> (x#xs)!i))
\<Longrightarrow> (\<forall>i j. i \<ge> j \<or> j \<ge> length xs \<or> \<not>(xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i))"
by force
lemma no_back_insert: "no_back (x#xs) \<Longrightarrow> no_back xs"
using no_back_alt no_back_insert_aux by blast
lemma no_arc_fst_if_no_back:
assumes "no_back (x#xs)" and "y \<in> set xs"
shows "\<not> y \<rightarrow>\<^bsub>T\<^esub> x"
proof -
have 0: "(x#xs)!0 = x" by simp
obtain j where "xs!j = y" "j < length xs" using assms(2) by (auto simp: in_set_conv_nth)
then have "(x#xs)!(Suc j) = y \<and> Suc j < length (x#xs)" by simp
then show ?thesis using assms(1) 0 by (metis no_back_def zero_less_Suc)
qed
lemma no_back_arcs_alt_aux2: "no_back xs \<Longrightarrow> no_back_arcs xs"
by(induction xs) (auto simp: no_back_insert no_arc_fst_if_no_back)
lemma no_back_arcs_alt: "no_back xs \<longleftrightarrow> no_back_arcs xs"
using no_back_arcs_alt_aux1 no_back_arcs_alt_aux2 no_back_alt by fastforce
lemma forward_arcs_alt_aux1:
"\<lbrakk>forward_arcs xs; i \<in> {1..(length (rev xs) - 1)}\<rbrakk> \<Longrightarrow> \<exists>j < i. (rev xs)!j \<rightarrow>\<^bsub>T\<^esub> (rev xs)!i"
proof(induction xs rule: forward_arcs.induct)
case (3 x x' xs)
then show ?case
proof(cases "i = length (rev (x#x'#xs)) - 1")
case True
then have i: "(rev (x#x'#xs))!i = x" by (simp add: nth_append)
then obtain y where y_def: "y\<in>set (x'#xs)" "y \<rightarrow>\<^bsub>T\<^esub> x" using "3.prems" by auto
then obtain j where j_def: "rev (x'#xs)!j = y" "j < length (rev (x'#xs))"
using in_set_conv_nth[of y] by fastforce
then have "rev (x#x'#xs)!j = y" by (auto simp: nth_append)
then show ?thesis using y_def(2) i j_def(2) True by auto
next
case False
then obtain j where j_def: "j < i" "rev (x' # xs)!j \<rightarrow>\<^bsub>T\<^esub> rev (x' # xs)!i" using 3 by auto
then have "rev (x#x'#xs)!j = rev (x'#xs)!j" using "3.prems"(2) by (auto simp: nth_append)
moreover have "rev (x#x'#xs)!i = rev (x'#xs)!i"
using "3.prems"(2) False by (auto simp: nth_append)
ultimately show ?thesis using j_def by auto
qed
qed(auto)
lemma forward_split_aux:
assumes "forward (xs@ys)" and "i\<in>{1..length xs - 1}"
shows "\<exists>j<i. xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i"
proof -
obtain j where "j < i \<and> (xs@ys)!j \<rightarrow>\<^bsub>T\<^esub> (xs@ys)!i" using assms forward_def by force
moreover have "i < length xs" using assms(2) by auto
ultimately show ?thesis by (auto simp: nth_append)
qed
lemma forward_split: "forward (xs@ys) \<Longrightarrow> forward xs"
using forward_split_aux forward_def by blast
lemma forward_cons:
"forward (rev (x#xs)) \<Longrightarrow> forward (rev xs)"
using forward_split by simp
lemma arc_to_lst_if_forward:
assumes "forward (rev (x#xs))" and "xs = y#ys"
shows "\<exists>y \<in> set xs. y \<rightarrow>\<^bsub>T\<^esub> x"
proof -
have "(x#xs)!0 = x" by simp
have "(rev xs@[x])!(length xs) = (xs@[x])!(length xs)" by (metis length_rev nth_append_length)
then have i: "rev (x#xs)!(length xs) = x" by simp
have "length xs \<in> {1..(length (rev (x#xs)) - 1)}" using assms(2) by simp
then obtain j where j_def: "j < length xs \<and> (rev (x#xs))!j \<rightarrow>\<^bsub>T\<^esub> (rev (x#xs))!length xs"
using assms(1) forward_def[of "rev (x#xs)"] by blast
then have "rev xs!j \<in> set xs" using length_rev nth_mem set_rev by metis
then have "rev (x#xs)!j \<in> set xs" by (auto simp: j_def nth_append)
then show ?thesis using i j_def by auto
qed
lemma forward_arcs_alt_aux2: "forward (rev xs) \<Longrightarrow> forward_arcs xs"
proof(induction xs rule: forward_arcs.induct)
case (3 x y xs)
then have "forward_arcs (y # xs)" using forward_cons by blast
then show ?case using arc_to_lst_if_forward "3.prems" by simp
qed(auto)
lemma forward_arcs_alt: "forward xs \<longleftrightarrow> forward_arcs (rev xs)"
using forward_arcs_alt_aux1 forward_arcs_alt_aux2 forward_def by fastforce
corollary forward_arcs_alt': "forward (rev xs) \<longleftrightarrow> forward_arcs xs"
using forward_arcs_alt by simp
corollary forward_arcs_split: "forward_arcs (ys@xs) \<Longrightarrow> forward_arcs xs"
using forward_split[of "rev xs" "rev ys"] forward_arcs_alt by simp
lemma seq_conform_alt: "seq_conform xs \<longleftrightarrow> forward xs \<and> no_back xs"
using forward_arcs_alt no_back_arcs_alt seq_conform_def by simp
lemma forward_app_aux:
assumes "forward s1" "forward s2" "\<exists>x\<in>set s1. x \<rightarrow>\<^bsub>T\<^esub> hd s2" "i\<in>{1..length (s1@s2) - 1}"
shows "\<exists>j<i. (s1@s2)!j \<rightarrow>\<^bsub>T\<^esub> (s1@s2)!i"
proof -
consider "i\<in>{1..length s1 - 1}" | "i = length s1" | "i\<in>{length s1 + 1..length s1 + length s2 - 1}"
using assms(4) by fastforce
then show ?thesis
proof(cases)
case 1
then obtain j where j_def: "j < i" "s1!j \<rightarrow>\<^bsub>T\<^esub> s1!i" using assms(1) forward_def by blast
moreover have "(s1@s2)!i = s1!i" using 1 by (auto simp: nth_append)
moreover have "(s1@s2)!j = s1!j" using 1 j_def(1) by (auto simp: nth_append)
ultimately show ?thesis by auto
next
case 2
then have "s2 \<noteq> []" using assms(4) by force
then have "(s1@s2)!i = hd s2" using 2 assms(4) by (simp add: hd_conv_nth nth_append)
then obtain x where x_def: "x\<in>set s1" "x \<rightarrow>\<^bsub>T\<^esub> (s1@s2)!i" using assms(3) by force
then obtain j where "s1!j = x" "j < length s1" by (auto simp: in_set_conv_nth)
then show ?thesis using x_def(2) 2 by (auto simp: nth_append)
next
case 3
then have "i-length s1 \<in> {1..length s2 - 1}" by fastforce
then obtain j where j_def: "j < (i-length s1)" "s2!j \<rightarrow>\<^bsub>T\<^esub> s2!(i-length s1)"
using assms(2) forward_def by blast
moreover have "(s1@s2)!i = s2!(i-length s1)" using 3 by (auto simp: nth_append)
moreover have "(s1@s2)!(j+length s1) = s2!j" using 3 j_def(1) by (auto simp: nth_append)
ultimately have "(j+length s1) < i \<and> (s1@s2)!(j+length s1) \<rightarrow>\<^bsub>T\<^esub> (s1@s2)!i" by force
then show ?thesis by blast
qed
qed
lemma forward_app: "\<lbrakk>forward s1; forward s2; \<exists>x\<in>set s1. x \<rightarrow>\<^bsub>T\<^esub> hd s2\<rbrakk> \<Longrightarrow> forward (s1@s2)"
by (simp add: forward_def forward_app_aux)
lemma before_conform1I: "before s1 s2 \<Longrightarrow> seq_conform s1"
unfolding before_def by blast
lemma before_forward1I: "before s1 s2 \<Longrightarrow> forward s1"
unfolding before_def seq_conform_alt by blast
lemma before_no_back1I: "before s1 s2 \<Longrightarrow> no_back s1"
unfolding before_def seq_conform_alt by blast
lemma before_ArcI: "before s1 s2 \<Longrightarrow> \<exists>x \<in> set s1. \<exists>y \<in> set s2. x \<rightarrow>\<^bsub>T\<^esub> y"
unfolding before_def by blast
lemma before_conform2I: "before s1 s2 \<Longrightarrow> seq_conform s2"
unfolding before_def by blast
lemma before_forward2I: "before s1 s2 \<Longrightarrow> forward s2"
unfolding before_def seq_conform_alt by blast
lemma before_no_back2I: "before s1 s2 \<Longrightarrow> no_back s2"
unfolding before_def seq_conform_alt by blast
lemma hd_reach_all_forward_arcs:
"\<lbrakk>hd (rev xs) \<in> verts T; forward_arcs xs; x \<in> set xs\<rbrakk> \<Longrightarrow> hd (rev xs) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x"
proof(induction xs arbitrary: x rule: forward_arcs.induct)
case (3 z y ys)
then have 0: "(\<exists>y \<in> set (y#ys). y \<rightarrow>\<^bsub>T\<^esub> z)" "forward_arcs (y#ys)" by auto
have hd_eq: "hd (rev (z # y # ys)) = hd (rev (y # ys))"
using hd_rev[of "y#ys"] by (auto simp: last_ConsR)
then show ?case
proof(cases "x = z")
case True
then obtain x' where x'_def: "x' \<in> set (y#ys)" "x' \<rightarrow>\<^bsub>T\<^esub> x" using "3.prems"(2) by auto
then have "hd (rev (z # y # ys)) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x'" using 3 hd_eq by simp
then show ?thesis using x'_def(2) reachable_adj_trans by blast
next
case False
then show ?thesis using 3 hd_eq by simp
qed
qed(auto)
lemma hd_reach_all_forward:
"\<lbrakk>hd xs \<in> verts T; forward xs; x \<in> set xs\<rbrakk> \<Longrightarrow> hd xs \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x"
using hd_reach_all_forward_arcs[of "rev xs"] by (simp add: forward_arcs_alt)
lemma hd_in_verts_if_forward: "forward (x#y#xs) \<Longrightarrow> hd (x#y#xs) \<in> verts T"
unfolding forward_def by fastforce
lemma two_elems_if_length_gt1: "length xs > 1 \<Longrightarrow> \<exists>x y ys. x#y#ys=xs"
by (metis create_ldeep_rev.cases list.size(3) One_nat_def length_Cons less_asym zero_less_Suc)
lemma hd_in_verts_if_forward': "\<lbrakk>length xs > 1; forward xs\<rbrakk> \<Longrightarrow> hd xs \<in> verts T"
using two_elems_if_length_gt1 hd_in_verts_if_forward by blast
lemma hd_reach_all_forward':
"\<lbrakk>length xs > 1; forward xs; x \<in> set xs\<rbrakk> \<Longrightarrow> hd xs \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x"
by (simp add: hd_in_verts_if_forward' hd_reach_all_forward)
lemma hd_reach_all_forward'':
"\<lbrakk>forward (x#y#xs); z \<in> set (x#y#xs)\<rbrakk> \<Longrightarrow> hd (x#y#xs) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> z"
using hd_in_verts_if_forward hd_reach_all_forward by blast
lemma no_back_if_distinct_forward: "\<lbrakk>forward xs; distinct xs\<rbrakk> \<Longrightarrow> no_back xs"
unfolding no_back_def proof
assume "\<exists>i j. i < j \<and> j < length xs \<and> xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i" and assms: "forward xs" "distinct xs"
then obtain i j where i_def: "i < j" "j < length xs" "xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i" by blast
show False
proof(cases "i=0")
case True
then have "xs!i = hd xs" using i_def(1,2) hd_conv_nth[of xs] by fastforce
then have "xs!i \<rightarrow>\<^sup>*\<^bsub>T\<^esub> xs!j" using i_def(1,2) assms(1) hd_reach_all_forward' by simp
then have "xs!i \<rightarrow>\<^sup>+\<^bsub>T\<^esub> xs!j" using reachable_neq_reachable1 i_def(3) by force
then show ?thesis using i_def(3) reachable1_not_reverse by blast
next
case False
then have "i \<in> {1 .. length xs - 1}" using i_def(1,2) by simp
then obtain j' where j'_def: "j' < i" "xs!j' \<rightarrow>\<^bsub>T\<^esub> xs!i"
using assms(1) unfolding forward_def by blast
have "xs!j' = xs!j" using i_def(3) j'_def(2) two_in_arcs_contr by fastforce
moreover have "xs!j' \<noteq> xs!j"
using j'_def(1) i_def(1,2) assms(2) nth_eq_iff_index_eq by fastforce
ultimately show ?thesis by blast
qed
qed
corollary seq_conform_if_dstnct_fwd: "\<lbrakk>forward xs; distinct xs\<rbrakk> \<Longrightarrow> seq_conform xs"
using no_back_if_distinct_forward seq_conform_def forward_arcs_alt no_back_arcs_alt by blast
lemma forward_arcs_single: "forward_arcs [x]"
by simp
lemma forward_single: "forward [x]"
unfolding forward_def by simp
lemma no_back_arcs_single: "no_back_arcs [x]"
by simp
lemma no_back_single: "no_back [x]"
unfolding no_back_def by simp
lemma seq_conform_single: "seq_conform [x]"
unfolding seq_conform_def by simp
lemma forward_arc_to_head':
assumes "forward ys" and "x \<notin> set ys" and "y \<in> set ys" and "x \<rightarrow>\<^bsub>T\<^esub> y"
shows "y = hd ys"
proof (rule ccontr)
assume asm: "y \<noteq> hd ys"
obtain i where i_def: "i < length ys" "ys!i = y" using assms(3) by (auto simp: in_set_conv_nth)
then have "i \<noteq> 0" using asm by (metis drop0 hd_drop_conv_nth)
then have "i \<in> {1..(length ys - 1)}" using i_def(1) by simp
then obtain j where j_def: "j < i" "ys!j \<rightarrow>\<^bsub>T\<^esub> ys!i"
using assms(1) forward_def by blast
then show False using assms(4,2) j_def(2) i_def two_in_arcs_contr by fastforce
qed
corollary forward_arc_to_head:
"\<lbrakk>forward ys; set xs \<inter> set ys = {}; x \<in> set xs; y \<in> set ys; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> y = hd ys"
using forward_arc_to_head' by blast
lemma forward_app':
"\<lbrakk>forward s1; forward s2; set s1 \<inter> set s2 = {}; \<exists>x\<in>set s1. \<exists>y\<in>set s2. x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> forward (s1@s2)"
using forward_app[of s1 s2] forward_arc_to_head by blast
lemma reachable1_from_outside_dom:
"\<lbrakk>x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; x \<notin> set ys; y \<in> set ys\<rbrakk> \<Longrightarrow> \<exists>x'. \<exists>y' \<in> set ys. x' \<notin> set ys \<and> x' \<rightarrow>\<^bsub>T\<^esub> y'"
by (induction x y rule: trancl.induct) auto
lemma hd_reachable1_from_outside':
"\<lbrakk>x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; forward ys; x \<notin> set ys; y \<in> set ys\<rbrakk> \<Longrightarrow> \<exists>y' \<in> set ys. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> hd ys"
apply(induction x y rule: trancl.induct)
using forward_arc_to_head' by force+
lemma hd_reachable1_from_outside:
"\<lbrakk>x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; forward ys; set xs \<inter> set ys = {}; x \<in> set xs; y \<in> set ys\<rbrakk>
\<Longrightarrow> \<exists>y' \<in> set ys. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> hd ys"
using hd_reachable1_from_outside' by blast
lemma reachable1_append_old_if_arc:
assumes "\<exists>x\<in>set xs. \<exists>y\<in>set ys. x \<rightarrow>\<^bsub>T\<^esub> y"
and "z \<notin> set xs"
and "forward xs"
and "y\<in>set (xs @ ys)"
and "z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
shows "\<exists>y\<in>set ys. z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof(cases "y \<in> set ys")
case True
then show ?thesis using assms(5) by blast
next
case False
then have "y \<in> set xs" using assms(4) by simp
then have 0: "z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> hd xs" using hd_reachable1_from_outside'[OF assms(5,3,2)] by blast
then have 1: "hd xs \<in> verts T" using reachable1_in_verts(2) by auto
obtain x y where x_def: "x\<in>set xs" "y\<in>set ys" "x \<rightarrow>\<^bsub>T\<^esub> y" using assms(1) by blast
then have "hd xs \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x" using hd_reach_all_forward[OF 1 assms(3)] by simp
then have "hd xs \<rightarrow>\<^sup>*\<^bsub>T\<^esub> y" using x_def(3) by force
then show ?thesis using reachable1_reachable_trans[OF 0] x_def(2) by blast
qed
lemma reachable1_append_old_if_arcU:
"\<lbrakk>\<exists>x\<in>set xs. \<exists>y\<in>set ys. x \<rightarrow>\<^bsub>T\<^esub> y; set U \<inter> set xs = {}; z \<in> set U;
forward xs; y\<in>set (xs @ ys); z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> \<exists>y\<in>set ys. z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using reachable1_append_old_if_arc[of xs ys] by auto
lemma before_arc_to_hd: "before xs ys \<Longrightarrow> \<exists>x \<in> set xs. x \<rightarrow>\<^bsub>T\<^esub> hd ys"
using forward_arc_to_head before_def seq_conform_alt by auto
lemma no_back_backarc_app1:
"\<lbrakk>j < length (xs@ys); j \<ge> length xs; i < j; no_back ys; (xs@ys)!j \<rightarrow>\<^bsub>T\<^esub> (xs@ys)!i\<rbrakk>
\<Longrightarrow> i < length xs"
by (rule ccontr) (auto simp add: no_back_def nth_append)
lemma no_back_backarc_app2: "\<lbrakk>no_back xs; i < j; (xs@ys)!j \<rightarrow>\<^bsub>T\<^esub> (xs@ys)!i\<rbrakk> \<Longrightarrow> j \<ge> length xs"
by (rule ccontr) (auto simp add: no_back_def nth_append)
lemma no_back_backarc_i_in_xs:
"\<lbrakk>no_back ys; j < length (xs@ys); i < j; (xs@ys)!j \<rightarrow>\<^bsub>T\<^esub> (xs@ys)!i\<rbrakk>
\<Longrightarrow> xs!i \<in> set xs \<and> (xs@ys)!i = xs!i"
by (auto simp add: no_back_def nth_append)
lemma no_back_backarc_j_in_ys:
"\<lbrakk>no_back xs; j < length (xs@ys); i < j; (xs@ys)!j \<rightarrow>\<^bsub>T\<^esub> (xs@ys)!i\<rbrakk>
\<Longrightarrow> ys!(j-length xs) \<in> set ys \<and> (xs@ys)!j = ys!(j-length xs)"
by (auto simp add: no_back_def nth_append)
lemma no_back_backarc_difsets:
assumes "no_back xs" and "no_back ys"
and "i < j" and "j < length (xs @ ys)" and "(xs @ ys) ! j \<rightarrow>\<^bsub>T\<^esub> (xs @ ys) ! i"
shows "\<exists>x \<in> set xs. \<exists>y \<in> set ys. y \<rightarrow>\<^bsub>T\<^esub> x"
using no_back_backarc_i_in_xs[OF assms(2,4,3)] no_back_backarc_j_in_ys[OF assms(1,4,3)] assms(5)
by auto
lemma no_back_backarc_difsets':
"\<lbrakk>no_back xs; no_back ys; \<exists>i j. i < j \<and> j < length (xs@ys) \<and> (xs@ys)!j \<rightarrow>\<^bsub>T\<^esub> (xs@ys)!i\<rbrakk>
\<Longrightarrow> \<exists>x \<in> set xs. \<exists>y \<in> set ys. y \<rightarrow>\<^bsub>T\<^esub> x"
using no_back_backarc_difsets by blast
lemma no_back_before_aux:
assumes "seq_conform xs" and "seq_conform ys"
and "set xs \<inter> set ys = {}" and "(\<exists>x\<in>set xs. \<exists>y\<in>set ys. x \<rightarrow>\<^bsub>T\<^esub> y)"
shows "no_back (xs @ ys)"
unfolding no_back_def by (metis assms adj_in_verts(2) forward_arc_to_head hd_reach_all_forward
inf_commute reachable1_not_reverse reachable_rtranclI rtrancl_into_trancl1 seq_conform_alt
no_back_backarc_difsets')
lemma no_back_before: "before xs ys \<Longrightarrow> no_back (xs@ys)"
using before_def no_back_before_aux by simp
lemma seq_conform_if_before: "before xs ys \<Longrightarrow> seq_conform (xs@ys)"
using no_back_before before_def seq_conform_alt forward_app before_arc_to_hd by simp
lemma no_back_arc_if_fwd_dstct:
assumes "forward (as@bs)" and "distinct (as@bs)"
shows "\<not>(\<exists>x\<in>set bs. \<exists>y\<in>set as. x \<rightarrow>\<^bsub>T\<^esub> y)"
proof
assume "\<exists>x\<in>set bs. \<exists>y\<in>set as. x \<rightarrow>\<^bsub>T\<^esub> y"
then obtain x y where x_def: "x\<in>set bs" "y\<in>set as" "x \<rightarrow>\<^bsub>T\<^esub> y" by blast
then obtain i where i_def: "as!i = y" "i < length as" by (auto simp: in_set_conv_nth)
obtain j where j_def: "bs!j = x" "j < length bs" using x_def(1) by (auto simp: in_set_conv_nth)
then have "(as@bs)!(j+length as) = x" by (simp add: nth_append)
moreover have "(as@bs)!i = y" using i_def by (simp add: nth_append)
moreover have "i < (j+length as)" using i_def(2) by simp
moreover have "(j+length as) < length (as @ bs)" using j_def by simp
ultimately show False
using no_back_if_distinct_forward[OF assms] x_def(3) unfolding no_back_def by blast
qed
lemma no_back_reach1_if_fwd_dstct:
assumes "forward (as@bs)" and "distinct (as@bs)"
shows "\<not>(\<exists>x\<in>set bs. \<exists>y\<in>set as. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
proof
assume "\<exists>x\<in>set bs. \<exists>y\<in>set as. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
then obtain x y where x_def: "x\<in>set bs" "y\<in>set as" "x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y" by blast
have fwd_as: "forward as" using forward_split[OF assms(1)] by blast
have x_as: "x \<notin> set as" using x_def(1) assms(2) by auto
show False
using assms(1) x_def append.assoc list.distinct(1) Nil_is_append_conv append_Nil2[of "as@bs"]
append_eq_append_conv2[of "as@bs" "as@bs" bs as] forward_arc_to_head' hd_append2
hd_reach_all_forward hd_reachable1_from_outside'[OF x_def(3) fwd_as x_as x_def(2)]
in_set_conv_decomp_first[of y as] in_set_conv_decomp_last reachable1_from_outside_dom
reachable1_in_verts(2) reachable1_not_reverse reachable1_reachable_trans
by metis
qed
lemma split_length_i: "i \<le> length bs \<Longrightarrow> \<exists>xs ys. xs@ys = bs \<and> length xs = i"
using length_take append_take_drop_id min_absorb2 by metis
lemma split_length_i_prefix:
assumes "length as \<le> i" "i < length (as@bs)"
shows "\<exists>xs ys. xs@ys = bs \<and> length (as@xs) = i"
proof -
obtain n where n_def: "n + length as = i"
using assms(1) ab_semigroup_add_class.add.commute le_Suc_ex by blast
then have "n \<le> length bs" using assms(2) by simp
then show ?thesis using split_length_i n_def by fastforce
qed
lemma forward_alt_aux1:
assumes "i \<in> {1..length xs - 1}" and "j<i" and "xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i"
shows "\<exists>as bs. as@bs = xs \<and> length as = i \<and> (\<exists>x \<in> set as. x \<rightarrow>\<^bsub>T\<^esub> xs!i)"
proof -
obtain as bs where "as@bs = xs \<and> length as = i"
using assms(1) atLeastAtMost_iff diff_le_self le_trans split_length_i[of i xs] by metis
then show ?thesis using assms(2,3) nth_append[of as bs j] by force
qed
lemma forward_alt_aux1':
"forward xs
\<Longrightarrow> \<forall>i \<in> {1..length xs - 1}. \<exists>as bs. as@bs = xs \<and> length as = i \<and> (\<exists>x \<in> set as. x \<rightarrow>\<^bsub>T\<^esub> xs!i)"
using forward_alt_aux1 unfolding forward_def by fastforce
lemma forward_alt_aux2:
"\<lbrakk>as@bs = xs; length as = i; \<exists>x \<in> set as. x \<rightarrow>\<^bsub>T\<^esub> xs!i\<rbrakk> \<Longrightarrow> \<exists>j<i. xs!j \<rightarrow>\<^bsub>T\<^esub> xs!i"
by (auto simp add: nth_append in_set_conv_nth)
lemma forward_alt_aux2':
"\<forall>i \<in> {1..length xs - 1}. \<exists>as bs. as@bs = xs \<and> length as = i \<and> (\<exists>x \<in> set as. x \<rightarrow>\<^bsub>T\<^esub> xs!i)
\<Longrightarrow> forward xs"
using forward_alt_aux2 unfolding forward_def by blast
corollary forward_alt:
"\<forall>i \<in> {1..length xs - 1}. \<exists>as bs. as@bs = xs \<and> length as = i \<and> (\<exists>x \<in> set as. x \<rightarrow>\<^bsub>T\<^esub> xs!i)
\<longleftrightarrow> forward xs"
using forward_alt_aux1'[of xs] forward_alt_aux2' by blast
lemma move_mid_forward_if_noarc_aux:
assumes "as \<noteq> []"
and "\<not>(\<exists>x \<in> set U. \<exists>y \<in> set bs. x \<rightarrow>\<^bsub>T\<^esub> y)"
and "forward (as@U@bs@cs)"
and "i \<in> {1..length (as@bs@U@cs) - 1}"
shows "\<exists>j<i. (as@bs@U@cs) ! j \<rightarrow>\<^bsub>T\<^esub> (as@bs@U@cs) ! i"
proof -
have 0: "i \<in> {1..length (as@U@bs@cs) - 1}" using assms(4) by auto
consider "i < length as" | "i \<in> {length as..length (as@bs) - 1}"
| "i \<in> {length (as@bs)..length (as@bs@U) - 1}"
| "i \<ge> length (as@bs@U)"
by fastforce
then show ?thesis
proof(cases)
case 1
then have "(as@U@bs@cs)!i = (as@bs@U@cs)!i" by (simp add: nth_append)
then obtain j where j_def: "j<i" "(as@U@bs@cs)!j \<rightarrow>\<^bsub>T\<^esub> ((as@bs)@U@cs)!i"
using assms(3) 0 unfolding forward_def by fastforce
then have "(as@U@bs@cs)!j = ((as@bs)@U@cs)!j" using 1 by (simp add: nth_append)
then show ?thesis using j_def by auto
next
case 2
have "((as@bs)@U@cs)!i = bs!(i - length as)"
using 2 assms(4) nth_append root_in_T directed_tree_axioms in_degree_root_zero
by (metis directed_tree.in_deg_one_imp_not_root atLeastAtMost_iff diff_diff_cancel
diff_is_0_eq diff_le_self diff_less_mono neq0_conv zero_less_diff)
then have i_in_bs: "((as@bs)@U@cs)!i \<in> set bs" using assms(4) 2 by auto
have "(i - length as) < length bs" using 2 assms(4) by force
then have "((as@bs)@U@cs)!i = (as@U@bs@cs)!(i + length U)"
using 2 by (auto simp: nth_append)
moreover have "(i + length U) \<in> {1.. length (as@U@bs@cs) - 1}" using 2 0 by force
ultimately obtain j where j_def:
"j < (i + length U)" "(as@U@bs@cs)!j \<rightarrow>\<^bsub>T\<^esub> ((as@bs)@U@cs)!i"
using assms(3) unfolding forward_def by fastforce
have "i < length (as@bs)" using \<open>i - length as < length bs\<close> by force
moreover have "length as \<le> i" using 2 by simp
ultimately obtain xs ys where xs_def: "bs = xs@ys" "length (as@xs) = i"
using split_length_i_prefix by blast
then have "j < (length (as@U@xs))" using 2 j_def(1) by simp
then have "(as@U@bs@cs)!j \<in> set (as@U@xs)" by (auto simp: xs_def(1) nth_append)
then have "(as@U@bs@cs)!j \<in> set (as@xs)" using assms(2) j_def(2) i_in_bs by auto
then obtain j' where j'_def: "j' < length (as@xs)" "(as@xs)!j' = (as@U@bs@cs)!j"
using in_set_conv_nth[of "(as@U@bs@cs)!j"] nth_append by blast
then have "((as@bs)@U@cs)!j' = (as@U@bs@cs)!j"
using nth_append[of "as@xs"] xs_def(1) by simp
then show ?thesis using j_def(2) j'_def(1) xs_def(2) by force
next
case 3
then have i_len_U: "i - length (as@bs) < length U" using assms(4) by fastforce
have i_len_asU: "i - length bs < length (as@U)" using 3 assms(4) by force
have "((as@bs)@U@cs)!i = (U@cs)!(i - length (as@bs))"
using 3 by (auto simp: nth_append)
also have "\<dots> = (as@U)!(i - length bs)"
using 3 i_len_U by (auto simp: ab_semigroup_add_class.add.commute nth_append)
also have "\<dots> = (as@U@bs@cs)!(i - length bs)"
using i_len_asU nth_append[of "as@U"] by simp
finally have 1: "((as@bs)@U@cs)!i = (as@U@bs@cs)!(i - length bs)" .
have "(i - length bs) \<ge> length as" using 3 by auto
then have "(i - length bs) \<ge> 1" using assms(1) length_0_conv[of as] by force
then have "(i - length bs) \<in> {1.. length (as@U@bs@cs) - 1}" using 0 by auto
then obtain j where j_def: "j < (i - length bs)" "(as@U@bs@cs)!j \<rightarrow>\<^bsub>T\<^esub> ((as@bs)@U@cs)!i"
using assms(3) 1 unfolding forward_def by fastforce
have "length as \<le> (i - length bs)" using 3 by auto
then obtain xs ys where xs_def: "U = xs@ys" "length (as@xs) = (i - length bs)"
using split_length_i_prefix[of as] i_len_asU by blast
then have "j < (length (as@xs))" using 3 j_def(1) by simp
then have "(as@U@bs@cs)!j \<in> set (as@bs@xs)" by (auto simp: xs_def(1) nth_append)
then obtain j' where j'_def: "j' < length (as@bs@xs)" "(as@bs@xs)!j' = (as@U@bs@cs)!j"
using in_set_conv_nth[of "(as@U@bs@cs)!j"] by blast
then have "((as@bs)@U@cs)!j' = (as@U@bs@cs)!j"
using nth_append[of "as@bs@xs"] xs_def(1) by simp
moreover have "j' < i" using j'_def(1) xs_def(2) 3 by auto
ultimately show ?thesis using j_def(2) by force
next
case 4
have len_eq: "length (as@U@bs) = length (as@bs@U)" by simp
have "((as@bs)@U@cs)!i = cs!(i - length (as@bs@U))"
using 4 nth_append[of "as@bs@U"] by simp
also have "\<dots> = cs!(i - length (as@U@bs))" using len_eq by argo
finally have "((as@bs)@U@cs)!i = ((as@U@bs)@cs)!i" using 4 nth_append[of "as@U@bs"] by simp
then obtain j where j_def: "j < i" "(as@U@bs@cs)!j \<rightarrow>\<^bsub>T\<^esub> ((as@bs)@U@cs)!i"
using assms(3) 0 unfolding forward_def by fastforce
have "length (as@U@bs) \<le> i" using 4 by auto
moreover have "i < length ((as@U@bs)@cs)" using 0 by auto
ultimately obtain xs ys where xs_def: "xs@ys = cs" "length ((as@U@bs) @ xs) = i"
using split_length_i_prefix[of "as@U@bs" i] by blast
then have "j < (length (as@U@bs@xs))" using 4 j_def(1) by simp
then have "(as@U@bs@cs)!j \<in> set (as@bs@U@xs)" by (auto simp: xs_def(1)[symmetric] nth_append)
then obtain j' where j'_def: "j' < length (as@bs@U@xs)" "(as@bs@U@xs)!j' = (as@U@bs@cs)!j"
using in_set_conv_nth[of "(as@U@bs@cs)!j"] by blast
then have "((as@bs)@U@cs)!j' = (as@U@bs@cs)!j"
using nth_append[of "as@bs@U@xs"] xs_def(1)[symmetric] by simp
moreover have "j' < i" using j'_def(1) xs_def(2) 4 by auto
ultimately show ?thesis using j_def(2) by auto
qed
qed
lemma move_mid_forward_if_noarc:
"\<lbrakk>as \<noteq> []; \<not>(\<exists>x \<in> set U. \<exists>y \<in> set bs. x \<rightarrow>\<^bsub>T\<^esub> y); forward (as@U@bs@cs)\<rbrakk>
\<Longrightarrow> forward (as@bs@U@cs)"
using move_mid_forward_if_noarc_aux unfolding forward_def by blast
lemma move_mid_backward_if_noarc_aux:
assumes "\<exists>x\<in>set U. x \<rightarrow>\<^bsub>T\<^esub> hd V"
and "forward V"
and "forward (as@U@bs@V@cs)"
and "i \<in> {1..length (as@U@V@bs@cs) - 1}"
shows "\<exists>j<i. (as@U@V@bs@cs) ! j \<rightarrow>\<^bsub>T\<^esub> (as@U@V@bs@cs) ! i"
proof -
have 0: "i \<in> {1..length (as@U@bs@V@cs) - 1}" using assms(4) by auto
consider "i < length (as@U)" | "i = length (as@U)" "i \<le> length (as@U@V) - 1"
| "i \<in> {length (as@U) + 1..length (as@U@V) - 1}"
| "i \<in> {length (as@U@V)..length (as@U@V@bs) - 1}"
| "i \<ge> length (as@U@V@bs)"
by fastforce
then show ?thesis
proof(cases)
case 1
then have "(as@U@bs@V@cs)!i = (as@U@V@bs@cs)!i" by (simp add: nth_append)
then obtain j where j_def: "j<i" "(as@U@bs@V@cs)!j \<rightarrow>\<^bsub>T\<^esub> (as@U@V@bs@cs)!i"
using assms(3) 0 unfolding forward_def by fastforce
then have "(as@U@V@bs@cs)!j = (as@U@bs@V@cs)!j" using 1 by (simp add: nth_append)
then show ?thesis using j_def by auto
next
case 2
have "(as@U@V@bs@cs)!i = (V@bs@cs)!0" using 2(1) by (auto simp: nth_append)
then have "(as@U@V@bs@cs)!i = hd V"
using 2 assms(4) hd_append hd_conv_nth Suc_n_not_le_n atLeastAtMost_iff le_diff_conv2
by (metis ab_semigroup_add_class.add.commute append.right_neutral Suc_eq_plus1_left)
then obtain x where x_def: "x \<in> set U" "x \<rightarrow>\<^bsub>T\<^esub> (as@U@V@bs@cs)!i" using assms(1) by auto
then obtain j where j_def: "(as@U)!j = x" "j < i" using in_set_conv_nth[of x] 2 by fastforce
then have "(as@U@V@bs@cs)!j = x" using 2(1) by (auto simp: nth_append)
then show ?thesis using j_def(2) x_def(2) by blast
next
case 3
have "i - length (as@U) \<in> {1 .. length V - 1}" using 3 by force
then obtain j where j_def: "j < (i - length (as@U))" "V!j \<rightarrow>\<^bsub>T\<^esub> V!(i - length (as@U))"
using assms(2) unfolding forward_def by blast
then have "(as@U@V@bs@cs)!(j+length (as@U)) = V!j"
using 3 nth_append[of "as@U"] nth_append[of V] by auto
moreover have "(as@U@V@bs@cs)!i = V!(i - length (as@U))"
using 3 nth_append[of "as@U"] nth_append[of V] by auto
moreover have "j+length (as@U) < i" using j_def(1) by simp
ultimately show ?thesis using j_def(2) by auto
next
case 4
have "(as@U@V@bs@cs)!i = (bs@cs)!(i - length (as@U@V))" using 4 nth_append[of "as@U@V"] by simp
also have "\<dots> = bs!(i - length (as@U@V))" using 4 assms(4) by (auto simp: nth_append)
also have "\<dots> = (as@U@bs)!(i - length (as@U@V) + length (as@U))" by (simp add: nth_append)
also have "\<dots> = (as@U@bs)!(i - length V)" using 4 by simp
finally have 1: "(as@U@V@bs@cs)!i = (as@U@bs@V@cs)!(i - length V)"
using 4 assms(4) nth_append[of "as@U@bs"] by auto
have "(i - length V) \<ge> length (as@U)" using 4 by auto
then have "(i - length V) \<ge> 1" using assms(1) length_0_conv by fastforce
then have "(i - length V) \<in> {1.. length (as@U@bs@V@cs) - 1}" using 0 by auto
then obtain j where j_def: "j < i - length V" "(as@U@bs@V@cs)!j \<rightarrow>\<^bsub>T\<^esub> (as@U@V@bs@cs)!i"
using assms(3) 1 unfolding forward_def by fastforce
have "length (as@U) \<le> (i - length V)" using 4 by fastforce
moreover have "(i - length V) < length ((as@U)@bs)" using 4 assms(4) by auto
ultimately obtain xs ys where xs_def: "xs@ys = bs" "length ((as@U)@ xs) = i - length V"
using split_length_i_prefix[of "as@U"] by blast
then have "j < (length (as@U@xs))" using 4 j_def(1) by simp
then have "(as@U@bs@V@cs)!j \<in> set (as@U@V@xs)" by (auto simp: xs_def(1)[symmetric] nth_append)
then obtain j' where j'_def: "j' < length (as@U@V@xs)" "(as@U@V@xs)!j' = (as@U@bs@V@cs)!j"
using in_set_conv_nth[of "(as@U@bs@V@cs)!j"] by blast
then have "(as@U@V@bs@cs)!j' = (as@U@bs@V@cs)!j"
using nth_append[of "as@U@V@xs"] xs_def(1) by auto
moreover have "j' < i" using j'_def(1) xs_def(2) 4 by auto
ultimately show ?thesis using j_def(2) by auto
next
case 5
have len_eq: "length (as@U@bs@V) = length (as@U@V@bs)" by simp
have "(as@U@V@bs@cs)!i = cs!(i - length (as@U@V@bs))"
using 5 nth_append[of "as@U@V@bs"] by auto
also have "\<dots> = cs!(i - length (as@U@bs@V))" using len_eq by argo
finally have "(as@U@V@bs@cs)!i = ((as@U@bs@V)@cs)!i"
using 5 nth_append[of "as@U@bs@V"] by simp
then obtain j where j_def: "j < i" "(as@U@bs@V@cs)!j \<rightarrow>\<^bsub>T\<^esub> (as@U@V@bs@cs)!i"
using assms(3) 0 unfolding forward_def by fastforce
have "length (as@U@bs@V) \<le> i" using 5 by auto
moreover have "i < length ((as@U@bs@V)@cs)" using 0 by auto
ultimately obtain xs ys where xs_def: "xs@ys = cs" "length ((as@U@bs@V) @ xs) = i"
using split_length_i_prefix[of "as@U@bs@V" i] by blast
then have "j < (length (as@U@bs@V@xs))" using 5 j_def(1) by simp
then have "(as@U@bs@V@cs)!j \<in> set (as@U@V@bs@xs)"
by (auto simp: xs_def(1)[symmetric] nth_append)
then obtain j' where j'_def: "j' < length (as@U@V@bs@xs)" "(as@U@V@bs@xs)!j' = (as@U@bs@V@cs)!j"
using in_set_conv_nth[of "(as@U@bs@V@cs)!j"] by blast
then have "(as@U@V@bs@cs)!j' = (as@U@bs@V@cs)!j"
using nth_append[of "as@U@V@bs@xs"] xs_def(1) by force
moreover have "j' < i" using j'_def(1) xs_def(2) 5 by auto
ultimately show ?thesis using j_def(2) by auto
qed
qed
lemma move_mid_backward_if_noarc:
"\<lbrakk>before U V; forward (as@U@bs@V@cs)\<rbrakk> \<Longrightarrow> forward (as@U@V@bs@cs)"
using before_forward2I
by (simp add: forward_def before_arc_to_hd move_mid_backward_if_noarc_aux)
lemma move_mid_backward_if_noarc':
"\<lbrakk>\<exists>x\<in>set U. \<exists>y\<in>set V. x \<rightarrow>\<^bsub>T\<^esub> y; forward V; set U \<inter> set V = {}; forward (as@U@bs@V@cs)\<rbrakk>
\<Longrightarrow> forward (as@U@V@bs@cs)"
using move_mid_backward_if_noarc_aux[of U V as bs cs] forward_arc_to_head[of V U] forward_def
by blast
end
subsection \<open>Sublist Additions\<close>
lemma fst_sublist_if_not_snd_sublist:
"\<lbrakk>xs@ys=A@B; \<not> sublist B ys\<rbrakk> \<Longrightarrow> \<exists>as bs. as @ bs = xs \<and> bs @ ys = B"
by (metis suffix_append suffix_def suffix_imp_sublist)
lemma sublist_before_if_mid:
assumes "sublist U (A@V)" and "A @ V @ B = xs" and "set U \<inter> set V = {}" and "U\<noteq>[]"
shows "\<exists>as bs cs. as @ U @ bs @ V @ cs = xs"
proof -
obtain C D where C_def: "(C @ U) @ D = A @ V" using assms(1) by (auto simp: sublist_def)
have "sublist V D"
using assms(3,4) fst_sublist_if_not_snd_sublist[OF C_def] disjoint_iff_not_equal last_appendR
by (metis Int_iff Un_Int_eq(1) append_Nil2 append_self_conv2 set_append last_in_set sublist_def)
then show ?thesis using assms(2) C_def sublist_def append.assoc by metis
qed
lemma list_empty_if_subset_dsjnt: "\<lbrakk>set xs \<subseteq> set ys; set xs \<inter> set ys = {}\<rbrakk> \<Longrightarrow> xs = []"
using semilattice_inf_class.inf.orderE by fastforce
lemma empty_if_sublist_dsjnt: "\<lbrakk>sublist xs ys; set xs \<inter> set ys = {}\<rbrakk> \<Longrightarrow> xs = []"
using set_mono_sublist list_empty_if_subset_dsjnt by fast
lemma sublist_snd_if_fst_dsjnt:
assumes "sublist U (V@B)" and "set U \<inter> set V = {}"
shows "sublist U B"
proof -
consider "sublist U V" | "sublist U B" | "(\<exists>xs1 xs2. U = xs1@xs2 \<and> suffix xs1 V \<and> prefix xs2 B)"
using assms(1) sublist_append by blast
then show ?thesis
proof(cases)
case 1
then show ?thesis using assms(2) empty_if_sublist_dsjnt by blast
next
case 2
then show ?thesis by simp
next
case 3
then obtain xs ys where xs_def: "U = xs@ys" "suffix xs V" "prefix ys B" by blast
then have "set xs \<subseteq> set V" by (simp add: set_mono_suffix)
then have "xs = []" using xs_def(1) assms(2) list_empty_if_subset_dsjnt by fastforce
then show ?thesis using xs_def(1,3) by simp
qed
qed
lemma sublist_fst_if_snd_dsjnt:
assumes "sublist U (B@V)" and "set U \<inter> set V = {}"
shows "sublist U B"
proof -
consider "sublist U V" | "sublist U B" | "(\<exists>xs1 xs2. U = xs1@xs2 \<and> suffix xs1 B \<and> prefix xs2 V)"
using assms(1) sublist_append by blast
then show ?thesis
proof(cases)
case 1
then show ?thesis using assms(2) empty_if_sublist_dsjnt by blast
next
case 2
then show ?thesis by simp
next
case 3
then obtain xs ys where xs_def: "U = xs@ys" "suffix xs B" "prefix ys V" by blast
then have "set ys \<subseteq> set V" by (simp add: set_mono_prefix)
then have "ys = []" using xs_def(1) assms(2) list_empty_if_subset_dsjnt by fastforce
then show ?thesis using xs_def(1,2) by simp
qed
qed
lemma sublist_app: "sublist (A @ B) C \<Longrightarrow> sublist A C \<and> sublist B C"
using sublist_order.dual_order.trans by blast
lemma sublist_Cons: "sublist (A # B) C \<Longrightarrow> sublist [A] C \<and> sublist B C"
using sublist_app[of "[A]"] by simp
lemma sublist_set_elem: "\<lbrakk>sublist xs (A@B); x \<in> set xs\<rbrakk> \<Longrightarrow> x \<in> set A \<or> x \<in> set B"
using set_mono_sublist by fastforce
lemma subset_snd_if_hd_notin_fst:
assumes "sublist ys (V @ B)" and "hd ys \<notin> set V" and "ys \<noteq> []"
shows "set ys \<subseteq> set B"
proof -
have "\<not> sublist ys V" using assms(2,3) by(auto simp: sublist_def)
then consider "sublist ys B" | "(\<exists>xs1 xs2. ys = xs1@xs2 \<and> suffix xs1 V \<and> prefix xs2 B)"
using assms(1) sublist_append by blast
then show ?thesis
proof(cases)
case 1
then show ?thesis using set_mono_sublist by blast
next
case 2
then obtain xs zs where xs_def: "ys = xs@zs" "suffix xs V" "prefix zs B" by blast
then have "set xs \<subseteq> set V" by (simp add: set_mono_suffix)
then have "xs = []" using xs_def(1) assms(2,3) hd_append hd_in_set subsetD by fastforce
then show ?thesis using xs_def(1,3) by (simp add: set_mono_prefix)
qed
qed
lemma suffix_ndjsnt_snd_if_nempty: "\<lbrakk>suffix xs (A@V); V \<noteq> []; xs \<noteq> []\<rbrakk> \<Longrightarrow> set xs \<inter> set V \<noteq> {}"
using empty_if_sublist_dsjnt disjoint_iff
by (metis sublist_append_leftI suffix_append suffix_imp_sublist)
lemma sublist_not_mid:
assumes "sublist U ((A @ V) @ B)" and "set U \<inter> set V = {}" and "V \<noteq> []"
shows "sublist U A \<or> sublist U B"
proof -
consider "sublist U A" | "sublist U V" | "(\<exists>xs1 xs2. U = xs1@xs2 \<and> suffix xs1 A \<and> prefix xs2 V)"
| "sublist U B" | "(\<exists>xs1 xs2. U = xs1@xs2 \<and> suffix xs1 (A@V) \<and> prefix xs2 B)"
using assms(1) sublist_append by metis
then show ?thesis
proof(cases)
case 2
then show ?thesis using assms(2) empty_if_sublist_dsjnt by blast
next
case 3
then show ?thesis using assms(2) sublist_append sublist_fst_if_snd_dsjnt by blast
next
case 5
then obtain xs ys where xs_def: "U = xs@ys" "suffix xs (A@V)" "prefix ys B" by blast
then have "set xs \<inter> set V \<noteq> {} \<or> xs = []" using suffix_ndjsnt_snd_if_nempty assms(3) by blast
then have "xs = []" using xs_def(1) assms(2) by auto
then show ?thesis using xs_def(1,3) by simp
qed(auto)
qed
lemma sublist_Y_cases_UV:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "U \<in> Y"
and "V \<in> Y"
and "U \<noteq> []"
and "V \<noteq> []"
and "(\<forall>xs \<in> Y. sublist xs (as@U@bs@V@cs))"
and "xs \<in> Y"
shows "sublist xs as \<or> sublist xs bs \<or> sublist xs cs \<or> U = xs \<or> V = xs"
using assms append_assoc sublist_not_mid by metis
lemma sublist_behind_if_nbefore:
assumes "sublist U xs" "sublist V xs" "\<nexists>as bs cs. as @ U @ bs @ V @ cs = xs" "set U \<inter> set V = {}"
shows "\<exists>as bs cs. as @ V @ bs @ U @ cs = xs"
proof -
have "V \<noteq> []" using assms(1,3) unfolding sublist_def by blast
obtain A B where A_def: "A @ V @ B = xs" using assms(2) by (auto simp: sublist_def)
then have "\<not>sublist U A" unfolding sublist_def using assms(3) by fastforce
moreover have "sublist U ((A @ V) @ B)" using assms(1) A_def by simp
ultimately have "sublist U B" using assms(4) sublist_not_mid \<open>V\<noteq>[]\<close> by blast
then show ?thesis unfolding sublist_def using A_def by blast
qed
lemma sublists_preserv_move_U:
"\<lbrakk>set xs \<inter> set U = {}; set xs \<inter> set V = {}; V\<noteq>[]; sublist xs (as@U@bs@V@cs)\<rbrakk>
\<Longrightarrow> sublist xs (as@bs@U@V@cs)"
using append_assoc self_append_conv2 sublist_def sublist_not_mid by metis
lemma sublists_preserv_move_UY:
"\<lbrakk>\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}; xs \<in> Y; U \<in> Y; V \<in> Y;
V \<noteq> []; sublist xs (as@U@bs@V@cs)\<rbrakk>
\<Longrightarrow> sublist xs (as@bs@U@V@cs)"
using sublists_preserv_move_U append_assoc sublist_appendI by metis
lemma sublists_preserv_move_UY_all:
"\<lbrakk>\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}; U \<in> Y; V \<in> Y;
V \<noteq> []; \<forall>xs \<in> Y. sublist xs (as@U@bs@V@cs)\<rbrakk>
\<Longrightarrow> \<forall>xs \<in> Y. sublist xs (as@bs@U@V@cs)"
using sublists_preserv_move_UY[of Y] by simp
lemma sublists_preserv_move_V:
"\<lbrakk>set xs \<inter> set U = {}; set xs \<inter> set V = {}; U\<noteq>[]; sublist xs (as@U@bs@V@cs)\<rbrakk>
\<Longrightarrow> sublist xs (as@U@V@bs@cs)"
using append_assoc self_append_conv2 sublist_def sublist_not_mid by metis
lemma sublists_preserv_move_VY:
"\<lbrakk>\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}; xs \<in> Y; U \<in> Y; V \<in> Y;
U \<noteq> []; sublist xs (as@U@bs@V@cs)\<rbrakk>
\<Longrightarrow> sublist xs (as@U@V@bs@cs)"
using sublists_preserv_move_V append_assoc sublist_appendI by metis
lemma sublists_preserv_move_VY_all:
"\<lbrakk>\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}; U \<in> Y; V \<in> Y;
U \<noteq> []; \<forall>xs \<in> Y. sublist xs (as@U@bs@V@cs)\<rbrakk>
\<Longrightarrow> \<forall>xs \<in> Y. sublist xs (as@U@V@bs@cs)"
using sublists_preserv_move_VY[of Y] by simp
lemma distinct_sublist_first:
"\<lbrakk>sublist as (x#xs); distinct (x#xs); x \<in> set as\<rbrakk> \<Longrightarrow> take (length as) (x#xs) = as"
unfolding sublist_def using distinct_app_trans_l distinct_ys_not_xs hd_in_set
by (metis list.sel(1) append_assoc append_eq_conv_conj append_self_conv2 hd_append2)
lemma distinct_sublist_first_remainder:
"\<lbrakk>sublist as (x#xs); distinct (x#xs); x \<in> set as\<rbrakk> \<Longrightarrow> as @ drop (length as) (x#xs) = x#xs"
using distinct_sublist_first append_take_drop_id[of "length as" "x#xs"] by fastforce
lemma distinct_set_diff: "distinct (xs@ys) \<Longrightarrow> set ys = set (xs@ys) - set xs"
by auto
lemma list_of_sublist_concat_eq:
assumes "\<forall>as \<in> Y. \<forall>bs \<in> Y. as = bs \<or> set as \<inter> set bs = {}"
and "\<forall>as \<in> Y. sublist as xs"
and "distinct xs"
and "set xs = \<Union>(set ` Y)"
and "finite Y"
shows "\<exists>ys. set ys = Y \<and> concat ys = xs \<and> distinct ys"
using assms proof(induction "Finite_Set.card Y" arbitrary: Y xs)
case (Suc n)
show ?case
proof(cases xs)
case Nil
then have "Y = {[]} \<or> Y = {}" using Suc.prems(4) by auto
then have "set [[]] = Y \<and> concat [[]] = xs \<and> distinct [[]]" using Nil Suc.hyps(2) by auto
then show ?thesis by blast
next
case (Cons x xs')
then obtain as where as_def: "x \<in> set as" "as \<in> Y" using Suc.prems(4) by auto
then have 0: "as @ (drop (length as) xs) = xs"
using Suc.prems(2,3) distinct_sublist_first_remainder Cons by fast
then have "\<forall>bs \<in> (Y - {as}). sublist bs (drop (length as) xs)"
using Suc.prems(1,2) as_def(2) by (metis DiffE insertI1 sublist_snd_if_fst_dsjnt)
moreover have "\<forall>cs \<in> (Y - {as}). \<forall>bs \<in> (Y - {as}). cs = bs \<or> set cs \<inter> set bs = {}"
using Suc.prems(1) by simp
moreover have "distinct (drop (length as) xs)" using Suc.prems(3) by simp
moreover have "set (drop (length as) xs) = \<Union> (set ` (Y-{as}))"
using Suc.prems(1,3,4) distinct_set_diff[of as "drop (length as) xs"] as_def(2) 0 by auto
moreover have "n = Finite_Set.card (Y-{as})" using Suc.hyps(2) as_def(2) Suc.prems(5) by simp
ultimately obtain ys where ys_def:
"set ys = (Y-{as})" "concat ys = drop (length as) xs" "distinct ys"
using Suc.hyps(1) Suc.prems(5) by blast
then have "set (as#ys) = Y \<and> concat (as#ys) = xs \<and> distinct (as#ys)" using 0 as_def(2) by auto
then show ?thesis by blast
qed
qed(auto)
lemma extract_length_decr[termination_simp]:
"List.extract P xs = Some (as,x,bs) \<Longrightarrow> length bs < length xs"
by (simp add: extract_Some_iff)
fun separate_P :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list" where
"separate_P P acc xs = (case List.extract P xs of
None \<Rightarrow> (acc,xs)
| Some (as,x,bs) \<Rightarrow> (case separate_P P (x#acc) bs of (acc',xs') \<Rightarrow> (acc', as@xs')))"
lemma separate_not_P_snd: "separate_P P acc xs = (as,bs) \<Longrightarrow> \<forall>x \<in> set bs. \<not>P x"
proof(induction P acc xs arbitrary: as bs rule: separate_P.induct)
case (1 P acc xs)
then show ?case
proof(cases "List.extract P xs")
case None
then have "bs = xs" using "1.prems" by simp
then show ?thesis using None by (simp add: extract_None_iff)
next
case (Some a)
then obtain cs x ds where x_def[simp]: "a = (cs,x,ds)" by(cases a) auto
then obtain acc' xs' where acc'_def: "separate_P P (x#acc) ds = (acc',xs')" by fastforce
then have "(acc', cs@xs') = (as,bs)" using "1.prems" Some by simp
moreover have "\<forall>x \<in> set xs'. \<not>P x" using "1.IH" acc'_def Some x_def by blast
ultimately show ?thesis using Some by (auto simp: extract_Some_iff)
qed
qed
lemma separate_input_impl_none: "separate_P P acc xs = (acc,xs) \<Longrightarrow> List.extract P xs = None"
using extract_None_iff separate_not_P_snd by fast
lemma separate_input_iff_none: "List.extract P xs = None \<longleftrightarrow> separate_P P acc xs = (acc,xs)"
using separate_input_impl_none by auto
lemma separate_P_fst_acc:
"separate_P P acc xs = (as,bs) \<Longrightarrow> \<exists>as'. as = as'@acc \<and> (\<forall>x \<in> set as'. P x)"
proof(induction P acc xs arbitrary: as bs rule: separate_P.induct)
case (1 P acc xs)
then show ?case
proof(cases "List.extract P xs")
case None
then show ?thesis using "1.prems" by simp
next
case (Some a)
then obtain cs x ds where x_def[simp]: "a = (cs,x,ds)" by(cases a) auto
then obtain acc' xs' where acc'_def: "separate_P P (x#acc) ds = (acc',xs')" by fastforce
then have "(acc', cs@xs') = (as,bs)" using "1.prems" Some by simp
then have "\<exists>as'. as = as'@(x#acc) \<and> (\<forall>x \<in> set as'. P x)"
using "1.IH" acc'_def Some x_def by blast
then show ?thesis using Some by (auto simp: extract_Some_iff)
qed
qed
lemma separate_P_fst: "separate_P P [] xs = (as,bs) \<Longrightarrow> \<forall>x \<in> set as. P x"
using separate_P_fst_acc by fastforce
subsection \<open>Optimal Solution for Lists of Fixed Sets\<close>
lemma distinct_seteq_set_length_eq:
"x \<in> {ys. set ys = xs \<and> distinct ys} \<Longrightarrow> length x = Finite_Set.card xs"
using distinct_card by fastforce
lemma distinct_seteq_set_Cons:
"\<lbrakk>Finite_Set.card xs = Suc n; x \<in> {ys. set ys = xs \<and> distinct ys}\<rbrakk>
\<Longrightarrow> \<exists>y ys. y # ys = x \<and> length ys = n \<and> distinct ys \<and> finite (set ys)"
using distinct_seteq_set_length_eq[of x] Suc_length_conv[of n x] by force
lemma distinct_seteq_set_Cons':
"\<lbrakk>Finite_Set.card xs = Suc n; x \<in> {ys. set ys = xs \<and> distinct ys}\<rbrakk>
\<Longrightarrow> \<exists>y ys zs. y # ys = x \<and> Finite_Set.card zs = n \<and> distinct ys \<and> set ys = zs"
using distinct_seteq_set_length_eq[of x] Suc_length_conv[of n x] by force
lemma distinct_seteq_set_Cons'':
"\<lbrakk>Finite_Set.card xs = Suc n; x \<in> {ys. set ys = xs \<and> distinct ys}\<rbrakk>
\<Longrightarrow> \<exists>y ys zs. y # ys = x \<and> y \<in> xs
\<and> set ys = zs \<and> Finite_Set.card zs = n \<and> distinct ys \<and> finite zs"
using distinct_seteq_set_Cons by fastforce
lemma distinct_seteq_set_Cons_in_set:
"\<lbrakk>Finite_Set.card xs = Suc n; x \<in> {ys. set ys = xs \<and> distinct ys}\<rbrakk>
\<Longrightarrow> \<exists>y ys zs. y#ys = x \<and> y \<in> xs \<and> Finite_Set.card zs = n \<and> ys\<in>{ys. set ys = zs \<and> distinct ys}"
using distinct_seteq_set_Cons'' by auto
lemma distinct_seteq_set_Cons_in_set':
"\<lbrakk>Finite_Set.card xs = Suc n; x \<in> {ys. set ys = xs \<and> distinct ys}\<rbrakk>
\<Longrightarrow> \<exists>y ys. x = y#ys \<and> y \<in> xs \<and> ys\<in>{ys. set ys = (xs - {y}) \<and> distinct ys}"
using distinct_seteq_set_Cons'' by fastforce
lemma distinct_seteq_eq_set_union:
"Finite_Set.card xs = Suc n
\<Longrightarrow> {ys. set ys = xs \<and> distinct ys}
= {y # ys |y ys. y \<in> xs \<and> ys \<in> {as. set as = (xs - {y}) \<and> distinct as}}"
using distinct_seteq_set_Cons_in_set' by force
lemma distinct_seteq_sub_set_union:
"Finite_Set.card xs = Suc n
\<Longrightarrow> {ys. set ys = xs \<and> distinct ys}
\<subseteq> {y # ys |y ys. y \<in> xs \<and> ys \<in> {as. \<exists>a \<in> xs. set as = (xs - {a}) \<and> distinct as}}"
using distinct_seteq_set_Cons_in_set' by fast
lemma finite_set_union: "\<lbrakk>finite ys; \<forall>y \<in> ys. finite y\<rbrakk> \<Longrightarrow> finite (\<Union>y \<in> ys. y)"
by simp
lemma Cons_set_eq_union_set:
"{x # y | x y y'. x \<in> xs \<and> y \<in> y' \<and> y' \<in> ys} = {x # y | x y. x \<in> xs \<and> y \<in> (\<Union>y \<in> ys. y)}"
by blast
lemma finite_set_Cons_union_finite:
"\<lbrakk>finite xs; finite ys; \<forall>y \<in> ys. finite y\<rbrakk>
\<Longrightarrow> finite {x # y | x y. x \<in> xs \<and> y \<in> (\<Union>y \<in> ys. y)}"
by (simp add: finite_image_set2)
lemma finite_set_Cons_finite:
"\<lbrakk>finite xs; finite ys; \<forall>y \<in> ys. finite y\<rbrakk>
\<Longrightarrow> finite {x # y | x y y'. x \<in> xs \<and> y \<in> y' \<and> y' \<in> ys}"
using Cons_set_eq_union_set[of xs] by (simp add: finite_image_set2)
lemma finite_set_Cons_finite':
"\<lbrakk>finite xs; finite ys\<rbrakk> \<Longrightarrow> finite {x # y |x y. x \<in> xs \<and> y \<in> ys}"
by (auto simp add: finite_image_set2)
lemma Cons_set_alt: "{x # y |x y. x \<in> xs \<and> y \<in> ys} = {zs. \<exists>x y. x # y = zs \<and> x \<in> xs \<and> y \<in> ys}"
by blast
lemma Cons_set_sub:
assumes "Finite_Set.card xs = Suc n"
shows "{ys. set ys = xs \<and> distinct ys}
\<subseteq> {x # y |x y. x \<in> xs \<and> y \<in> (\<Union>y \<in> xs. {as. set as = xs - {y} \<and> distinct as})}"
using distinct_seteq_eq_set_union[OF assms] by auto
lemma distinct_seteq_finite: "finite xs \<Longrightarrow> finite {ys. set ys = xs \<and> distinct ys}"
proof(induction "Finite_Set.card xs" arbitrary: xs)
case (Suc n)
have "finite (\<Union>y \<in> xs. {as. set as = xs - {y} \<and> distinct as})" using Suc by simp
then have "finite {x # y |x y. x \<in> xs \<and> y \<in> (\<Union>y \<in> xs. {as. set as = xs - {y} \<and> distinct as})}"
using finite_set_Cons_finite'[OF Suc.prems] by blast
then show ?case using finite_subset[OF Cons_set_sub] Suc.hyps(2)[symmetric] by blast
qed(simp)
lemma distinct_setsub_split:
"{ys. set ys \<subseteq> xs \<and> distinct ys}
= {ys. set ys = xs \<and> distinct ys} \<union> (\<Union>y \<in> xs. {ys. set ys \<subseteq> (xs-{y}) \<and> distinct ys})"
by blast
lemma distinct_setsub_finite: "finite xs \<Longrightarrow> finite {ys. set ys \<subseteq> xs \<and> distinct ys}"
proof(induction "Finite_Set.card xs" arbitrary: xs)
case (Suc x)
then show ?case using distinct_seteq_finite distinct_setsub_split[of xs] by auto
qed(simp)
lemma valid_UV_lists_finite:
"finite xs \<Longrightarrow> finite {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x}"
using distinct_seteq_finite by force
lemma valid_UV_lists_r_subset:
"{x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x \<and> take 1 x = [r]}
\<subseteq> {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x}"
by blast
lemma valid_UV_lists_r_finite:
"finite xs \<Longrightarrow> finite {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x \<and> take 1 x = [r]}"
using valid_UV_lists_finite finite_subset[OF valid_UV_lists_r_subset] by fast
lemma valid_UV_lists_arg_min_ex_aux:
"\<lbrakk>finite ys; ys \<noteq> {}; ys = {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using arg_min_if_finite(1)[of ys f] arg_min_least[of ys, where ?f = f] by auto
lemma valid_UV_lists_arg_min_ex:
"\<lbrakk>finite xs; ys \<noteq> {}; ys = {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using valid_UV_lists_finite valid_UV_lists_arg_min_ex_aux[of ys] by blast
lemma valid_UV_lists_arg_min_r_ex_aux:
"\<lbrakk>finite ys; ys \<noteq> {};
ys = {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x \<and> take 1 x = [r]}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using arg_min_if_finite(1)[of ys f] arg_min_least[of ys, where ?f = f] by auto
lemma valid_UV_lists_arg_min_r_ex:
"\<lbrakk>finite xs; ys \<noteq> {};
ys = {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x \<and> take 1 x = [r]}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using valid_UV_lists_r_finite[of xs] valid_UV_lists_arg_min_r_ex_aux[of ys] by blast
lemma valid_UV_lists_nemtpy:
assumes "finite xs" "set (U@V) \<subseteq> xs" "distinct (U@V)"
shows "{x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x} \<noteq> {}"
proof -
obtain cs where "set cs = xs - set (U@V) \<and> distinct cs"
using assms(1) finite_distinct_list[of "xs - set (U@V)"] by blast
then have "[]@U@[]@V@cs = U@V@cs" "set (U@V@cs) = xs" "distinct (U@V@cs)" using assms by auto
then show ?thesis by blast
qed
lemma valid_UV_lists_nemtpy':
"\<lbrakk>finite xs; set U \<inter> set V = {}; set U \<subseteq> xs; set V \<subseteq> xs; distinct U; distinct V\<rbrakk>
\<Longrightarrow> {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x} \<noteq> {}"
using valid_UV_lists_nemtpy[of xs] by simp
lemma valid_UV_lists_nemtpy_r:
assumes "finite xs" and "set (U@V) \<subseteq> xs" and "distinct (U@V)"
and "take 1 U = [r] \<or> r \<notin> set U \<union> set V" and "r \<in> xs"
shows "{x. (\<exists>as bs cs. as@U@bs@V@cs = x) \<and> set x = xs \<and> distinct x \<and> take 1 x = [r]} \<noteq> {}"
proof(cases "take 1 U = [r]")
case True
obtain cs where "set cs = xs - set (U@V) \<and> distinct cs"
using assms(1) finite_distinct_list by auto
then have "[]@U@[]@V@cs = U@V@cs" "set (U@V@cs) = xs" "distinct (U@V@cs)" using assms by auto
then show ?thesis using True take1_singleton_app by fast
next
case False
obtain cs where cs_def: "set cs = xs - ({r} \<union> set (U@V)) \<and> distinct cs"
using assms(1) finite_distinct_list by auto
then have "[r]@U@[]@V@cs = [r]@U@V@cs" "set ([r]@U@V@cs) = xs" "distinct ([r]@U@V@cs)"
"take 1 ([r]@U@V@cs) = [r]"
using assms False by auto
then show ?thesis by (smt (verit, del_insts) empty_Collect_eq)
qed
lemma valid_UV_lists_nemtpy_r':
"\<lbrakk>finite xs; set U \<inter> set V = {}; set U \<subseteq> xs; set V \<subseteq> xs; distinct U; distinct V;
take 1 U = [r] \<or> r \<notin> set U \<union> set V; r \<in> xs\<rbrakk>
\<Longrightarrow> {x. \<exists>as bs cs. as@U@bs@V@cs = x \<and> set x = xs \<and> distinct x \<and> take 1 x = [r]} \<noteq> {}"
using valid_UV_lists_nemtpy_r[of xs] by simp
lemma valid_UV_lists_arg_min_ex':
"\<lbrakk>finite xs; set U \<inter> set V = {}; set U \<subseteq> xs; set V \<subseteq> xs; distinct U; distinct V;
ys = {x. (\<exists>as bs cs. as@U@bs@V@cs = x) \<and> set x = xs \<and> distinct x}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using valid_UV_lists_arg_min_ex[of xs] valid_UV_lists_nemtpy'[of xs] by simp
lemma valid_UV_lists_arg_min_r_ex':
"\<lbrakk>finite xs; set U \<inter> set V = {}; set U \<subseteq> xs; set V \<subseteq> xs; distinct U; distinct V;
take 1 U = [r] \<or> r \<notin> set U \<union> set V; r \<in> xs;
ys = {x. (\<exists>as bs cs. as@U@bs@V@cs = x) \<and> set x = xs \<and> distinct x \<and> take 1 x = [r]}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using valid_UV_lists_arg_min_r_ex[of xs] valid_UV_lists_nemtpy_r'[of xs] by simp
lemma valid_UV_lists_alt:
assumes "P = (\<lambda>x. (\<exists>as bs cs. as@U@bs@V@cs = x) \<and> set x = xs \<and> distinct x)"
shows "{x. (\<exists>as bs cs. as@U@bs@V@cs = x) \<and> set x = xs \<and> distinct x} = {ys. P ys}"
using assms by simp
lemma valid_UV_lists_argmin_ex:
fixes cost :: "'a list \<Rightarrow> real"
assumes "P = (\<lambda>x. (\<exists>as bs cs. as@U@bs@V@cs = x) \<and> set x = xs \<and> distinct x)"
and "finite xs"
and "set U \<inter> set V = {}"
and "set U \<subseteq> xs"
and "set V \<subseteq> xs"
and "distinct U"
and "distinct V"
shows "\<exists>as' bs' cs'. P (as'@U@bs'@V@cs') \<and>
(\<forall>as bs cs. P (as@U@bs@V@cs) \<longrightarrow> cost (as'@U@bs'@V@cs') \<le> cost (as@U@bs@V@cs))"
proof -
obtain y where "y \<in> {ys. P ys} \<and> (\<forall>z \<in> {ys. P ys}. cost y \<le> cost z)"
using valid_UV_lists_arg_min_ex'[OF assms(2-7)] assms(1) by fastforce
then show ?thesis using assms(1) by blast
qed
lemma valid_UV_lists_argmin_ex_noP:
fixes cost :: "'a list \<Rightarrow> real"
assumes "finite xs"
and "set U \<inter> set V = {}"
and "set U \<subseteq> xs"
and "set V \<subseteq> xs"
and "distinct U"
and "distinct V"
shows "\<exists>as' bs' cs'. set (as' @ U @ bs' @ V @ cs') = xs \<and> distinct (as' @ U @ bs' @ V @ cs')
\<and> (\<forall>as bs cs. set (as @ U @ bs @ V @ cs) = xs \<and> distinct (as @ U @ bs @ V @ cs)
\<longrightarrow> cost (as' @ U @ bs' @ V @ cs') \<le> cost (as @ U @ bs @ V @ cs))"
using valid_UV_lists_argmin_ex[OF refl assms] by metis
lemma valid_UV_lists_argmin_r_ex:
fixes cost :: "'a list \<Rightarrow> real"
assumes "P = (\<lambda>x. (\<exists>as bs cs. as@U@bs@V@cs = x) \<and> set x = xs \<and> distinct x \<and> take 1 x = [r])"
and "finite xs"
and "set U \<inter> set V = {}"
and "set U \<subseteq> xs"
and "set V \<subseteq> xs"
and "distinct U"
and "distinct V"
and "take 1 U = [r] \<or> r \<notin> set U \<union> set V"
and "r \<in> xs"
shows "\<exists>as' bs' cs'. P (as'@U@bs'@V@cs') \<and>
(\<forall>as bs cs. P (as@U@bs@V@cs) \<longrightarrow> cost (as'@U@bs'@V@cs') \<le> cost (as@U@bs@V@cs))"
proof -
obtain y where "y \<in> {ys. P ys} \<and> (\<forall>z \<in> {ys. P ys}. cost y \<le> cost z)"
using valid_UV_lists_arg_min_r_ex'[OF assms(2-9)] assms(1) by fastforce
then show ?thesis using assms(1) by blast
qed
lemma valid_UV_lists_argmin_r_ex_noP:
fixes cost :: "'a list \<Rightarrow> real"
assumes "finite xs"
and "set U \<inter> set V = {}"
and "set U \<subseteq> xs"
and "set V \<subseteq> xs"
and "distinct U"
and "distinct V"
and "take 1 U = [r] \<or> r \<notin> set U \<union> set V"
and "r \<in> xs"
shows "\<exists>as' bs' cs'. set (as' @ U @ bs' @ V @ cs') = xs
\<and> distinct (as' @ U @ bs' @ V @ cs') \<and> take 1 (as' @ U @ bs' @ V @ cs') = [r]
\<and> (\<forall>as bs cs. set (as @ U @ bs @ V @ cs) = xs
\<and> distinct (as @ U @ bs @ V @ cs) \<and> take 1 (as @ U @ bs @ V @ cs) = [r]
\<longrightarrow> cost (as' @ U @ bs' @ V @ cs') \<le> cost (as @ U @ bs @ V @ cs))"
using valid_UV_lists_argmin_r_ex[OF refl assms] by metis
lemma valid_UV_lists_argmin_r_ex_noP':
fixes cost :: "'a list \<Rightarrow> real"
assumes "finite xs"
and "set U \<inter> set V = {}"
and "set U \<subseteq> xs"
and "set V \<subseteq> xs"
and "distinct U"
and "distinct V"
and "take 1 U = [r] \<or> r \<notin> set U \<union> set V"
and "r \<in> xs"
shows "\<exists>as' bs' cs'. set (as' @ U @ bs' @ V @ cs') = xs
\<and> distinct (as' @ U @ bs' @ V @ cs') \<and> take 1 (as' @ U @ bs' @ V @ cs') = [r]
\<and> (\<forall>as bs cs. set (as @ U @ bs @ V @ cs) = xs
\<and> distinct (as @ U @ bs @ V @ cs) \<and> take 1 (as @ U @ bs @ V @ cs) = [r]
\<longrightarrow> cost (rev (as' @ U @ bs' @ V @ cs')) \<le> cost (rev (as @ U @ bs @ V @ cs)))"
using valid_UV_lists_argmin_r_ex_noP[OF assms] by meson
lemma take1_split_nempty: "ys \<noteq> [] \<Longrightarrow> take 1 (xs@ys@zs) = take 1 (xs@ys)"
by (metis append.assoc append_Nil2 gr_zeroI length_0_conv less_one same_append_eq
take_append take_eq_Nil zero_less_diff)
lemma take1_elem: "\<lbrakk>take 1 (xs@ys) = [r]; r \<in> set xs\<rbrakk> \<Longrightarrow> take 1 xs = [r]"
using in_set_conv_decomp_last[of r xs] by auto
lemma take1_nelem: "\<lbrakk>take 1 (xs@ys) = [r]; r \<notin> set ys\<rbrakk> \<Longrightarrow> take 1 xs = [r]"
using take1_elem[of xs ys r] append_self_conv2[of xs] hd_in_set[of ys]
by (fastforce dest: hd_eq_take1)
lemma take1_split_nelem_nempty: "\<lbrakk>take 1 (xs@ys@zs) = [r]; ys \<noteq> []; r \<notin> set ys\<rbrakk> \<Longrightarrow> take 1 xs = [r]"
using take1_split_nempty take1_nelem by fastforce
lemma take1_empty_if_nelem: "\<lbrakk>take 1 (as@bs@cs) = [r]; r \<notin> set as\<rbrakk> \<Longrightarrow> as = []"
using take1_split_nelem_nempty[of "[]" as "bs@cs"] by auto
lemma take1_empty_if_mid: "\<lbrakk>take 1 (as@bs@cs) = [r]; r \<in> set bs; distinct (as@bs@cs)\<rbrakk> \<Longrightarrow> as = []"
using take1_empty_if_nelem by fastforce
lemma take1_mid_if_elem:
"\<lbrakk>take 1 (as@bs@cs) = [r]; r \<in> set bs; distinct (as@bs@cs)\<rbrakk> \<Longrightarrow> take 1 bs = [r]"
using take1_empty_if_mid[of as bs cs] by (fastforce intro: take1_elem)
lemma contr_optimal_nogap_no_r:
assumes "asi rank r cost"
and "rank (rev V) \<le> rank (rev U)"
and "finite xs"
and "set U \<inter> set V = {}"
and "set U \<subseteq> xs"
and "set V \<subseteq> xs"
and "distinct U"
and "distinct V"
and "r \<notin> set U \<union> set V"
and "r \<in> xs"
shows "\<exists>as' cs'. distinct (as' @ U @ V @ cs') \<and> take 1 (as' @ U @ V @ cs') = [r]
\<and> set (as' @ U @ V @ cs') = xs \<and> (\<forall>as bs cs. set (as @ U @ bs @ V @ cs) = xs
\<and> distinct (as @ U @ bs @ V @ cs) \<and> take 1 (as @ U @ bs @ V @ cs) = [r]
\<longrightarrow> cost (rev (as' @ U @ V @ cs')) \<le> cost (rev (as @ U @ bs @ V @ cs)))"
proof -
define P where "P ys \<equiv> set ys = xs \<and> distinct ys \<and> take 1 ys = [r]" for ys
obtain as' bs' cs' where bs'_def:
"set (as'@U@bs'@V@cs') = xs" "distinct (as'@U@bs'@V@cs')" "take 1 (as'@U@bs'@V@cs') = [r]"
"\<forall>as bs cs. P (as @ U @ bs @ V @ cs) \<longrightarrow>
cost (rev (as' @ U @ bs' @ V @ cs')) \<le> cost (rev (as @ U @ bs @ V @ cs))"
using valid_UV_lists_argmin_r_ex_noP'[OF assms(3-8)] assms(9,10) unfolding P_def by blast
then consider "U = []" | "V = [] \<or> bs' = []"
| "rank (rev bs') \<le> rank (rev U)" "U \<noteq> []" "bs' \<noteq> []"
| "rank (rev U) \<le> rank (rev bs')" "U \<noteq> []" "V \<noteq> []" "bs' \<noteq> []"
by fastforce
then show ?thesis
proof(cases)
case 1
then have "\<forall>as bs cs. P (as @ U @ bs @ V @ cs) \<longrightarrow>
cost (rev ((as'@bs')@U@V@cs')) \<le> cost (rev (as @ U @ bs @ V @ cs))"
using bs'_def(4) by simp
moreover have "set ((as'@bs')@U@V@cs') = xs" using bs'_def(1) by auto
moreover have "distinct ((as'@bs')@U@V@cs')" using bs'_def(2) by auto
moreover have "take 1 ((as'@bs')@U@V@cs') = [r]" using bs'_def(3) 1 by auto
ultimately show ?thesis unfolding P_def by blast
next
case 2
then have "\<forall>as bs cs. P (as @ U @ bs @ V @ cs) \<longrightarrow>
cost (rev (as'@U@V@bs'@cs')) \<le> cost (rev (as @ U @ bs @ V @ cs))" using bs'_def(4) by auto
moreover have "set (as'@U@V@bs'@cs') = xs" using bs'_def(1) by auto
moreover have "distinct (as'@U@V@bs'@cs')" using bs'_def(2) by auto
moreover have "take 1 (as'@U@V@bs'@cs') = [r]" using bs'_def(3) 2 by auto
ultimately show ?thesis unfolding P_def by blast
next
case 3
have 0: "distinct (as'@bs'@U@V@cs')" using bs'_def(2) by auto
have 1: "take 1 (as'@bs'@U@V@cs') = [r]"
using bs'_def(3) assms(9) 3(2) take1_split_nelem_nempty[of as' U "bs'@V@cs'"] by simp
then have "cost (rev (as'@bs'@U@V@cs')) \<le> cost (rev (as'@U@bs'@V@cs'))"
using asi_le_rfst[OF assms(1) 3(1,3,2) 0] bs'_def(3) by blast
then have "\<forall>as bs cs. P (as @ U @ bs @ V @ cs) \<longrightarrow>
cost (rev ((as'@bs')@U@V@cs')) \<le> cost (rev (as @ U @ bs @ V @ cs))"
using bs'_def(4) by fastforce
moreover have "set ((as'@bs')@U@V@cs') = xs" using bs'_def(1) by auto
moreover have "distinct ((as'@bs')@U@V@cs')" using 0 by simp
moreover have "take 1 ((as'@bs')@U@V@cs') = [r]" using 1 by simp
ultimately show ?thesis using P_def by blast
next
case 4
then have 3: "rank (rev V) \<le> rank (rev bs')" using assms(2) by simp
have 0: "distinct ((as'@U)@V@bs'@cs')" using bs'_def(2) by auto
have 1: "take 1 (as'@U@V@bs'@cs') = [r]"
using bs'_def(3) assms(9) 4(2) take1_split_nelem_nempty[of as' U "bs'@V@cs'"] by simp
then have "cost (rev (as'@U@V@bs'@cs')) \<le> cost (rev ((as'@U)@bs'@V@cs'))"
using asi_le_rfst[OF assms(1) 3 4(3,4) 0] bs'_def(3) by simp
then have "\<forall>as bs cs. P (as @ U @ bs @ V @ cs) \<longrightarrow>
cost (rev (as'@U@V@bs'@cs')) \<le> cost (rev (as @ U @ bs @ V @ cs))"
using bs'_def(4) by fastforce
moreover have "set (as'@U@V@bs'@cs') = xs" using bs'_def(1) by auto
moreover have "distinct (as'@U@V@bs'@cs')" using 0 by simp
ultimately show ?thesis using P_def 1 by blast
qed
qed
fun combine_lists_P :: "('a list \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list list \<Rightarrow> 'a list list" where
"combine_lists_P _ y [] = [y]"
| "combine_lists_P P y (x#xs) = (if P (x@y) then combine_lists_P P (x@y) xs else (x@y)#xs)"
fun make_list_P :: "('a list \<Rightarrow> bool) \<Rightarrow> 'a list list \<Rightarrow> 'a list list \<Rightarrow> 'a list list" where
"make_list_P P acc xs = (case List.extract P xs of
None \<Rightarrow> rev acc @ xs
| Some (as,y,bs) \<Rightarrow> make_list_P P (combine_lists_P P y (rev as @ acc)) bs)"
lemma combine_lists_concat_rev_eq: "concat (rev (combine_lists_P P y xs)) = concat (rev xs) @ y"
by (induction P y xs rule: combine_lists_P.induct) auto
lemma make_list_concat_rev_eq: "concat (make_list_P P acc xs) = concat (rev acc) @ concat xs"
proof(induction P acc xs rule: make_list_P.induct)
case (1 P acc xs)
then show ?case
proof(cases "List.extract P xs")
case (Some a)
then obtain as x bs where x_def[simp]: "a = (as,x,bs)" by(cases a) auto
then have "concat (make_list_P P acc xs)
= concat (rev (combine_lists_P P x (rev as @ acc))) @ concat bs"
using 1 Some by simp
also have "\<dots> = concat (rev acc) @ concat (as@x#bs)"
using combine_lists_concat_rev_eq[of P] by simp
finally show ?thesis using Some extract_SomeE by force
qed(simp)
qed
lemma combine_lists_sublists:
"\<exists>x \<in> {y} \<union> set xs. sublist as x \<Longrightarrow> \<exists>x \<in> set (combine_lists_P P y xs). sublist as x"
proof (induction P y xs rule: combine_lists_P.induct)
case (2 P y x xs)
then show ?case
proof(cases "sublist as x \<or> sublist as y")
case True
then have "sublist as (x@y)" using sublist_order.dual_order.trans by blast
then show ?thesis using 2 by force
next
case False
then show ?thesis using 2 by simp
qed
qed(simp)
lemma make_list_sublists:
"\<exists>x \<in> set acc \<union> set xs. sublist cs x \<Longrightarrow> \<exists>x \<in> set (make_list_P P acc xs). sublist cs x"
proof(induction P acc xs rule: make_list_P.induct)
case (1 P acc xs)
then show ?case
proof(cases "List.extract P xs")
case (Some a)
then obtain as x bs where x_def[simp]: "a = (as,x,bs)" by(cases a) auto
then have "make_list_P P acc xs = make_list_P P (combine_lists_P P x (rev as @ acc)) bs"
using Some by simp
then have "\<exists>a \<in> set (combine_lists_P P x (rev as @ acc)) \<union> set bs. sublist cs a"
using Some combine_lists_sublists[of x "rev as @ acc" cs] "1.prems"
by (auto simp: extract_Some_iff)
then show ?thesis using 1 Some by simp
qed(simp)
qed
lemma combine_lists_nempty: "\<lbrakk>[] \<notin> set xs; y \<noteq> []\<rbrakk> \<Longrightarrow> [] \<notin> set (combine_lists_P P y xs)"
by (induction P y xs rule: combine_lists_P.induct) auto
lemma make_list_nempty:
"\<lbrakk>[] \<notin> set acc; [] \<notin> set xs\<rbrakk> \<Longrightarrow> [] \<notin> set (make_list_P P acc xs)"
proof (induction P acc xs rule: make_list_P.induct)
case (1 P acc xs)
show ?case
proof(cases "List.extract P xs")
case None
then show ?thesis using 1 by simp
next
case (Some a)
then show ?thesis using 1 by (auto simp: extract_Some_iff combine_lists_nempty)
qed
qed
lemma combine_lists_notP:
"\<forall>x\<in>set xs. \<not>P x \<Longrightarrow> (\<exists>x. combine_lists_P P y xs = [x]) \<or> (\<forall>x\<in>set (combine_lists_P P y xs). \<not>P x)"
by (induction P y xs rule: combine_lists_P.induct) auto
lemma combine_lists_single: "xs = [x] \<Longrightarrow> combine_lists_P P y xs = [x@y]"
by auto
lemma combine_lists_lastP:
"P (last xs) \<Longrightarrow> (\<exists>x. combine_lists_P P y xs = [x]) \<or> (P (last (combine_lists_P P y xs)))"
by (induction P y xs rule: combine_lists_P.induct) auto
lemma make_list_notP:
"\<lbrakk>(\<forall>x \<in> set acc. \<not>P x) \<or> P (last acc)\<rbrakk>
\<Longrightarrow> (\<forall>x\<in>set (make_list_P P acc xs). \<not>P x) \<or> (\<exists>y ys. make_list_P P acc xs = y # ys \<and> P y)"
proof(induction P acc xs rule: make_list_P.induct)
case (1 P acc xs)
then show ?case
proof(cases "List.extract P xs")
case None
then show ?thesis
proof(cases "\<forall>x \<in> set acc. \<not>P x")
case True
from None have "\<forall>x \<in> set xs. \<not> P x" by (simp add: extract_None_iff)
then show ?thesis using True "1.prems" None by auto
next
case False
then have "acc \<noteq> []" by auto
then have "make_list_P P acc xs = last acc # rev (butlast acc) @ xs" using None by simp
then show ?thesis using False "1.prems" by blast
qed
next
case (Some a)
then obtain as x bs where x_def[simp]: "a = (as,x,bs)" by(cases a) auto
show ?thesis
proof(cases "\<forall>x \<in> set acc. \<not>P x")
case True
then have "\<forall>x \<in> set (rev as @ acc). \<not>P x" using Some by (auto simp: extract_Some_iff)
then have "(\<forall>x\<in>set (combine_lists_P P x (rev as @ acc)). \<not> P x)
\<or> P (last (combine_lists_P P x (rev as @ acc)))"
using combine_lists_notP[of "rev as @ acc" P] by force
then show ?thesis using "1.IH" Some by simp
next
case False
then have "P (last acc) \<and> acc \<noteq> []" using "1.prems" by auto
then have "P (last (rev as @ acc))" using "1.prems" by simp
then have "(\<forall>x\<in>set (combine_lists_P P x (rev as @ acc)). \<not> P x)
\<or> P (last (combine_lists_P P x (rev as @ acc)))"
using combine_lists_lastP[of P] by force
then show ?thesis using "1.IH" Some by simp
qed
qed
qed
corollary make_list_notP_empty_acc:
"(\<forall>x\<in>set (make_list_P P [] xs). \<not>P x) \<or> (\<exists>y ys. make_list_P P [] xs = y # ys \<and> P y)"
using make_list_notP[of "[]"] by auto
definition unique_set_r :: "'a \<Rightarrow> 'a list set \<Rightarrow> 'a list \<Rightarrow> bool" where
"unique_set_r r Y ys \<longleftrightarrow> set ys = \<Union>(set ` Y) \<and> distinct ys \<and> take 1 ys = [r]"
context directed_tree
begin
definition fwd_sub :: "'a \<Rightarrow> 'a list set \<Rightarrow> 'a list \<Rightarrow> bool" where
"fwd_sub r Y ys \<longleftrightarrow> unique_set_r r Y ys \<and> forward ys \<and> (\<forall>xs \<in> Y. sublist xs ys)"
lemma distinct_mid_unique1: "\<lbrakk>distinct (xs@U@ys); U\<noteq>[]; xs@U@ys = as@U@bs\<rbrakk> \<Longrightarrow> as = xs"
using distinct_app_trans_r distinct_ys_not_xs[of xs "U@ys"] hd_append2[of U] append_is_Nil_conv[of U]
by (metis append_Cons_eq_iff distinct.simps(2) list.exhaust_sel list.set_sel(1))
lemma distinct_mid_unique2: "\<lbrakk>distinct (xs@U@ys); U\<noteq>[]; xs@U@ys = as@U@bs\<rbrakk> \<Longrightarrow> ys = bs"
using distinct_mid_unique1 by blast
lemma concat_all_sublist: "\<forall>x \<in> set xs. sublist x (concat xs)"
using split_list by force
lemma concat_all_sublist_rev: "\<forall>x \<in> set xs. sublist x (concat (rev xs))"
using split_list by force
lemma concat_all_sublist1:
assumes "distinct (as@U@bs)"
and "concat cs @ U @ concat ds = as@U@bs"
and "U \<noteq> []"
and "set (cs@U#ds) = Y"
shows "\<exists>X. X \<subseteq> Y \<and> set as = \<Union>(set ` X) \<and> (\<forall>xs \<in> X. sublist xs as)"
proof -
have eq: "concat cs = as"
using distinct_mid_unique1[of "concat cs" U "concat ds"] assms(1-3) by simp
then have "\<forall>xs \<in> set cs. sublist xs as" using concat_all_sublist by blast
then show ?thesis using eq assms(4) by fastforce
qed
lemma concat_all_sublist2:
assumes "distinct (as@U@bs)"
and "concat cs @ U @ concat ds = as@U@bs"
and "U \<noteq> []"
and "set (cs@U#ds) = Y"
shows "\<exists>X. X \<subseteq> Y \<and> set bs = \<Union>(set ` X) \<and> (\<forall>xs \<in> X. sublist xs bs)"
proof -
have eq: "concat ds = bs"
using distinct_mid_unique1[of "concat cs" U "concat ds"] assms(1-3) by simp
then have "\<forall>xs \<in> set ds. sublist xs bs" using concat_all_sublist by blast
then show ?thesis using eq assms(4) by fastforce
qed
lemma concat_split_mid:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "finite Y"
and "U \<in> Y"
and "distinct (as@U@bs)"
and "set (as@U@bs) = \<Union>(set ` Y)"
and "\<forall>xs \<in> Y. sublist xs (as@U@bs)"
and "U \<noteq> []"
shows "\<exists>cs ds. concat cs = as \<and> concat ds = bs \<and> set (cs@U#ds) = Y \<and> distinct (cs@U#ds)"
proof -
obtain ys where ys_def: "set ys = Y" "concat ys = as@U@bs" "distinct ys"
using list_of_sublist_concat_eq[OF assms(1,6,4,5,2)] by blast
then obtain cs ds where cs_def: "cs@U#ds = ys"
using assms(3) in_set_conv_decomp_first[of U ys] by blast
then have "List.extract ((=) U) ys = Some (cs,U,ds)"
using extract_Some_iff[of "(=) U"] ys_def(3) by auto
then have "concat cs @ U @ concat ds = as@U@bs" using ys_def(2) cs_def by auto
then have "concat cs = as \<and> concat ds = bs"
using distinct_mid_unique1[of "concat cs" U] assms(4,7) by auto
then show ?thesis using ys_def(1,3) cs_def by blast
qed
lemma mid_all_sublists_set1:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "finite Y"
and "U \<in> Y"
and "distinct (as@U@bs)"
and "set (as@U@bs) = \<Union>(set ` Y)"
and "\<forall>xs \<in> Y. sublist xs (as@U@bs)"
and "U \<noteq> []"
shows "\<exists>X. X \<subseteq> Y \<and> set as = \<Union>(set ` X) \<and> (\<forall>xs \<in> X. sublist xs as)"
proof -
obtain ys where ys_def: "set ys = Y" "concat ys = as@U@bs" "distinct ys"
using list_of_sublist_concat_eq[OF assms(1,6,4,5,2)] by blast
then obtain cs ds where cs_def: "cs@U#ds = ys"
using assms(3) in_set_conv_decomp_first[of U ys] by blast
then have "List.extract ((=) U) ys = Some (cs,U,ds)"
using extract_Some_iff[of "(=) U"] ys_def(3) by auto
then have "concat cs @ U @ concat ds = as@U@bs" using ys_def(2) cs_def by auto
then show ?thesis using cs_def ys_def(1) concat_all_sublist1[OF assms(4)] assms(7) by force
qed
lemma mid_all_sublists_set2:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "finite Y"
and "U \<in> Y"
and "distinct (as@U@bs)"
and "set (as@U@bs) = \<Union>(set ` Y)"
and "\<forall>xs \<in> Y. sublist xs (as@U@bs)"
and "U \<noteq> []"
shows "\<exists>X. X \<subseteq> Y \<and> set bs = \<Union>(set ` X) \<and> (\<forall>xs \<in> X. sublist xs bs)"
proof -
obtain ys where ys_def: "set ys = Y" "concat ys = as@U@bs" "distinct ys"
using list_of_sublist_concat_eq[OF assms(1,6,4,5,2)] by blast
then obtain cs ds where cs_def: "cs@U#ds = ys"
using assms(3) in_set_conv_decomp_first[of U ys] by blast
then have "List.extract ((=) U) ys = Some (cs,U,ds)"
using extract_Some_iff[of "(=) U"] ys_def(3) by auto
then have "concat cs @ U @ concat ds = as@U@bs" using ys_def(2) cs_def by auto
then show ?thesis using cs_def ys_def(1) concat_all_sublist2[OF assms(4)] assms(7) by force
qed
lemma nonempty_notin_distinct_prefix:
assumes "distinct (as@bs@V@cs)" and "concat as' = as" and "V \<noteq> []"
shows "V \<notin> set as'"
proof
assume "V \<in> set as'"
then have "set V \<subseteq> set as" using assms(2) by auto
then have "set as \<inter> set V \<noteq> {}" using assms(3) by (simp add: Int_absorb1)
then show False using assms(1) by auto
qed
lemma concat_split_UV:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "finite Y"
and "U \<in> Y"
and "V \<in> Y"
and "distinct (as@U@bs@V@cs)"
and "set (as@U@bs@V@cs) = \<Union>(set ` Y)"
and "\<forall>xs \<in> Y. sublist xs (as@U@bs@V@cs)"
and "U \<noteq> []"
and "V \<noteq> []"
shows "\<exists>as' bs' cs'. concat as' = as \<and> concat bs' = bs \<and> concat cs' = cs
\<and> set (as'@U#bs'@V#cs') = Y \<and> distinct (as'@U#bs'@V#cs')"
proof -
obtain as' ds where as'_def:
"concat as' = as" "concat ds = bs@V@cs" "set (as'@U#ds) = Y" "distinct (as'@U#ds)"
using concat_split_mid[OF assms(1-3,5-8)] by auto
have 0: "distinct (bs@V@cs)" using assms(5) by simp
have "V \<notin> set as'"
using assms(5,9) as'_def(1) nonempty_notin_distinct_prefix[of as "U@bs"] by auto
moreover have "V \<noteq> U" using assms(5,8,9) empty_if_sublist_dsjnt[of U] by auto
ultimately have "V \<in> set ds" using as'_def(3) assms(4) by auto
then show ?thesis
using as'_def 0 assms(9) concat_append distinct_mid_unique1
by (metis concat.simps(2) distinct_mid_unique2 split_list)
qed
lemma cost_decr_if_noarc_lessrank:
assumes "asi rank r cost"
and "b \<noteq> []"
and "r \<notin> set U"
and "U \<noteq> []"
and "set (as@U@bs@cs) = \<Union>(set ` Y)"
and "distinct (as@U@bs@cs)"
and "take 1 (as@U@bs@cs) = [r]"
and "forward (as@U@bs@cs)"
and "concat (b#bs') = bs"
and "(\<forall>xs \<in> Y. sublist xs as \<or> sublist xs U
\<or> (\<exists>x \<in> set (b#bs'). sublist xs x) \<or> sublist xs cs)"
and "\<not>(\<exists>x \<in> set U. \<exists>y \<in> set b. x \<rightarrow>\<^bsub>T\<^esub> y)"
and "rank (rev b) < rank (rev U)"
shows "fwd_sub r Y (as@b@U@concat bs'@cs)
\<and> cost (rev (as@b@U@concat bs'@cs)) < cost (rev (as@U@bs@cs))"
proof -
have rank_yU: "rank (rev b) < rank (rev U)" using assms(12) by simp
have 0: "take 1 (as@b@U@concat bs'@cs) = [r]"
using take1_singleton_app take1_split_nelem_nempty[OF assms(7,4,3)] by fast
have 1: "distinct (as@b@U@ concat bs'@cs)" using assms(6,9) by force
have "take 1 (as@U@b@concat bs'@cs) = [r]" using assms(7,9) by force
then have cost_lt: "cost (rev (as@b@U@concat bs'@cs)) < cost (rev (as@U@bs@cs))"
using asi_lt_rfst[OF assms(1) rank_yU assms(2,4) 1 0] assms(9) by fastforce
have P: "set (as@b@U@concat bs'@cs) = \<Union>(set ` Y)" using assms(5,9) by fastforce
then have P: "unique_set_r r Y (as@b@U@concat bs'@cs)"
using 0 1 unfolding unique_set_r_def by blast
have "(\<forall>xs \<in> Y. sublist xs as \<or> sublist xs U \<or> sublist xs b
\<or> sublist xs (concat bs') \<or> sublist xs cs)"
using assms(10) concat_all_sublist[of bs']
sublist_order.dual_order.trans[where a = "concat bs'"] by auto
then have all_sub: "\<forall>xs \<in> Y. sublist xs (as@b@U@concat bs'@cs)"
by (metis sublist_order.order.trans sublist_append_leftI sublist_append_rightI)
have "as \<noteq> []" using take1_split_nelem_nempty[OF assms(7,4,3)] by force
then have "forward (as@b@U@concat bs'@cs)"
using move_mid_forward_if_noarc assms(8,9,11) by auto
then show ?thesis using assms(12) P all_sub cost_lt fwd_sub_def by blast
qed
lemma cost_decr_if_noarc_lessrank':
assumes "asi rank r cost"
and "b \<noteq> []"
and "r \<notin> set U"
and "U \<noteq> []"
and "set (as@U@bs@cs) = \<Union>(set ` Y)"
and "distinct (as@U@bs@cs)"
and "take 1 (as@U@bs@cs) = [r]"
and "forward (as@U@bs@cs)"
and "concat (b#bs') = bs"
and "(\<forall>xs \<in> Y. sublist xs as \<or> sublist xs U
\<or> (\<exists>x \<in> set (b#bs'). sublist xs x) \<or> sublist xs cs)"
and "\<not>(\<exists>x \<in> set U. \<exists>y \<in> set b. x \<rightarrow>\<^bsub>T\<^esub> y)"
and "rank (rev b) < rank (rev V)"
and "rank (rev V) \<le> rank (rev U)"
shows "fwd_sub r Y (as@b@U@concat bs'@cs)
\<and> cost (rev (as@b@U@concat bs'@cs)) < cost (rev (as@U@bs@cs))"
using cost_decr_if_noarc_lessrank[OF assms(1-11)] assms(12,13) by simp
lemma sublist_exists_append:
"\<exists>a\<in>set ((x # xs) @ [b]). sublist ys a \<Longrightarrow> \<exists>a\<in>set(xs @ [x@b]). sublist ys a"
using sublist_order.dual_order.trans by auto
lemma sublist_set_concat_cases:
"\<exists>a\<in>set ((x # xs) @ [b]). sublist ys a \<Longrightarrow> sublist ys (concat (rev xs)) \<or> sublist ys x \<or> sublist ys b"
using sublist_order.dual_order.trans concat_all_sublist_rev[of xs] by auto
lemma sublist_set_concat_or_cases_aux1:
"sublist ys as \<or> sublist ys U \<or> sublist ys cs
\<Longrightarrow> sublist ys (as @ U @ concat (rev xs)) \<or> sublist ys cs"
using sublist_order.dual_order.trans by blast
lemma sublist_set_concat_or_cases_aux2:
"\<exists>a\<in>set ((x # xs) @ [b]). sublist ys a
\<Longrightarrow> sublist ys (as @ U @ concat (rev xs)) \<or> sublist ys x \<or> sublist ys b"
using sublist_set_concat_cases[of x xs b ys] sublist_order.dual_order.trans by blast
lemma sublist_set_concat_or_cases:
"sublist ys as \<or> sublist ys U \<or> (\<exists>a\<in>set ((x#xs) @ [b]). sublist ys a) \<or> sublist ys cs \<Longrightarrow>
sublist ys (as@U@ concat (rev xs)) \<or> sublist ys x \<or> (\<exists>a\<in>set [b]. sublist ys a) \<or> sublist ys cs"
using sublist_set_concat_or_cases_aux1[of ys as U cs] sublist_set_concat_or_cases_aux2[of x xs b ys]
by auto
corollary not_reachable1_append_if_not_old:
"\<lbrakk>\<not> (\<exists>z\<in>set U. \<exists>y\<in>set b. z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); set U \<inter> set x = {}; forward x;
\<exists>z\<in>set x. \<exists>y\<in>set b. z \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> \<not> (\<exists>z\<in>set U. \<exists>y\<in>set (x@b). z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
using reachable1_append_old_if_arcU[of x b U] by auto
lemma combine_lists_notP:
assumes "asi rank r cost"
and "b \<noteq> []"
and "r \<notin> set U"
and "U \<noteq> []"
and "set (as@U@bs@cs) = \<Union>(set ` Y)"
and "distinct (as@U@bs@cs)"
and "take 1 (as@U@bs@cs) = [r]"
and "forward (as@U@bs@cs)"
and "concat (rev ys @ [b]) = bs"
and "(\<forall>xs \<in> Y. sublist xs as \<or> sublist xs U
\<or> (\<exists>x \<in> set (ys @ [b]). sublist xs x) \<or> sublist xs cs)"
and "rank (rev V) \<le> rank (rev U)"
and "\<not>(\<exists>x \<in> set U. \<exists>y \<in> set b. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
and "rank (rev b) < rank (rev V)"
and "P = (\<lambda>x. rank (rev x) < rank (rev V))"
and "\<forall>x\<in>set ys. \<not>P x"
and "\<forall>xs. fwd_sub r Y xs \<longrightarrow> cost (rev (as@U@bs@cs)) \<le> cost (rev xs)"
and "\<forall>x \<in> set ys. x \<noteq> []"
and "\<forall>x \<in> set ys. forward x"
and "forward b"
shows "\<forall>x\<in>set (combine_lists_P P b ys). \<not>P x \<and> forward x"
using assms proof(induction P b ys rule: combine_lists_P.induct)
case (1 P b)
have 0: "concat (b#[]) = bs" using "1.prems"(9) by simp
have 2: "(\<forall>xs \<in> Y. sublist xs as \<or> sublist xs U
\<or> (\<exists>x \<in> set ([b]). sublist xs x) \<or> sublist xs cs)" using "1.prems"(10) by simp
have 3: "\<not> (\<exists>x\<in>set U. \<exists>y\<in>set b. x \<rightarrow>\<^bsub>T\<^esub> y)" using "1.prems"(12) by blast
show ?case
using cost_decr_if_noarc_lessrank'[OF 1(1-8) 0 2 3 1(13,11)] 1(16) by auto
next
case (2 P b x xs)
have "take 1 as = [r]" using "2.prems"(3,4,7) take1_split_nelem_nempty by fast
then have "r \<in> set as" using in_set_takeD[of r 1] by simp
then have "r \<notin> set x" using "2.prems"(6,9) by force
then have "x \<noteq> []" using "2.prems"(17) by simp
text \<open>Arc between x and b otherwise not optimal.\<close>
have 4: "as@U@bs@cs = (as@U@concat (rev xs)) @ x @ b @ cs" using "2.prems"(9) by simp
have set: "set ((as@U@concat (rev xs)) @ x @ b @ cs) = \<Union> (set ` Y)"
using "2.prems"(5) 4 by simp
have dst: "distinct ((as@U@concat (rev xs)) @ x @ b @ cs)" using "2.prems"(6) 4 by simp
have tk1: "take 1 ((as@U@concat (rev xs)) @ x @ b @ cs) = [r]" using "2.prems"(7) 4 by simp
have fwd: "forward ((as@U@concat (rev xs)) @ x @ b @ cs)" using "2.prems"(8) 4 by simp
have cnct: "concat (b # []) = b" by simp
have sblst: "\<forall>xs' \<in> Y. sublist xs' (as @ U @ concat (rev xs)) \<or> sublist xs' x
\<or> (\<exists>a\<in>set [b]. sublist xs' a) \<or> sublist xs' cs"
using "2.prems"(10) sublist_set_concat_or_cases[where as = as] by simp
have "rank (rev b) < rank (rev x)" using "2.prems"(13-15) by simp
then have arc_xb: "\<exists>z\<in>set x. \<exists>y\<in>set b. z \<rightarrow>\<^bsub>T\<^esub> y"
using "2.prems"(16) 4
cost_decr_if_noarc_lessrank[OF 2(2,3) \<open>r\<notin>set x\<close> \<open>x\<noteq>[]\<close> set dst tk1 fwd cnct sblst]
by fastforce
have "set x \<inter> set b = {}" using dst by auto
then have fwd: "forward (x@b)" using forward_app' arc_xb "2.prems"(18,19) by simp
show ?case
proof(cases "P (x @ b)")
case True
have 0: "x @ b \<noteq> []" using "2.prems"(2) by blast
have 1: "concat (rev xs @ [x @ b]) = bs" using "2.prems"(9) by simp
have 3: "\<forall>xs' \<in> Y. sublist xs' as \<or> sublist xs' U
\<or> (\<exists>a\<in>set (xs @ [x @ b]). sublist xs' a) \<or> sublist xs' cs"
using "2.prems"(10) sublist_exists_append by fast
have "set U \<inter> set x = {}" using 4 "2.prems"(6) by force
then have 4: "\<not> (\<exists>z\<in>set U. \<exists>y\<in>set (x @ b). z \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
using not_reachable1_append_if_not_old[OF "2.prems"(12)] "2.prems"(18) arc_xb by simp
have 5: "rank (rev (x @ b)) < rank (rev V)" using True "2.prems"(14) by simp
show ?thesis
using "2.IH"[OF True 2(2) 0 2(4-9) 1 3 2(12) 4 5 2(15)] 2(16-19) fwd by auto
next
case False
then show ?thesis using "2.prems"(15,18) fwd by simp
qed
qed
lemma sublist_app_l: "sublist ys cs \<Longrightarrow> sublist ys (xs @ cs)"
using sublist_order.dual_order.trans by blast
lemma sublist_split_concat:
assumes "a \<in> set (acc @ (as@x#bs))" and "sublist ys a"
shows "(\<exists>a\<in>set (rev acc @ as @ [x]). sublist ys a) \<or> sublist ys (concat bs @ cs)"
proof(cases "a \<in> set (rev acc @ as @ [x])")
case True
then show ?thesis using assms(2) by blast
next
case False
then have "a \<in> set bs" using assms(1) by simp
then show ?thesis
using assms(2) concat_all_sublist[of bs]
sublist_order.dual_order.trans[where c = ys, where b = "concat bs"]
by fastforce
qed
lemma sublist_split_concat':
"\<exists>a \<in> set (acc @ (as@x#bs)). sublist ys a \<or> sublist ys cs
\<Longrightarrow> (\<exists>a\<in>set (rev acc @ as @ [x]). sublist ys a) \<or> sublist ys (concat bs @ cs)"
using sublist_split_concat sublist_app_l[of ys cs] by blast
lemma make_list_notP:
assumes "asi rank r cost"
and "r \<notin> set U"
and "U \<noteq> []"
and "set (as@U@bs@cs) = \<Union>(set ` Y)"
and "distinct (as@U@bs@cs)"
and "take 1 (as@U@bs@cs) = [r]"
and "forward (as@U@bs@cs)"
and "concat (rev acc @ ys) = bs"
and "(\<forall>xs \<in> Y. sublist xs as \<or> sublist xs U
\<or> (\<exists>x \<in> set (acc @ ys). sublist xs x) \<or> sublist xs cs)"
and "rank (rev V) \<le> rank (rev U)"
and "\<And>xs. \<lbrakk>xs \<in> set ys; \<exists>x \<in> set U. \<exists>y \<in> set xs. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
and "P = (\<lambda>x. rank (rev x) < rank (rev V))"
and "\<forall>xs. fwd_sub r Y xs \<longrightarrow> cost (rev (as@U@bs@cs)) \<le> cost (rev xs)"
and "\<forall>x \<in> set ys. x \<noteq> []"
and "\<forall>x \<in> set ys. forward x"
and "\<forall>x \<in> set acc. x \<noteq> []"
and "\<forall>x \<in> set acc. forward x"
and "\<forall>x \<in> set acc. \<not>P x"
shows "\<forall>x\<in>set (make_list_P P acc ys). \<not>P x"
using assms proof(induction P acc ys rule: make_list_P.induct)
case (1 P acc xs)
then show ?case
proof(cases "List.extract P xs")
case None
then have "\<forall>x \<in> set xs. \<not> P x" by (simp add: extract_None_iff)
then show ?thesis using "1.prems"(18) None by auto
next
case (Some a)
then obtain as' x bs' where x_def[simp]: "a = (as',x,bs')" by(cases a) auto
then have x: "\<forall>x \<in> set (rev as' @ acc). \<not>P x" "xs = as'@x#bs'" "rank (rev x) < rank (rev V)"
using Some "1.prems"(12,18) by (auto simp: extract_Some_iff)
have "x \<noteq> []" using "1.prems"(14) Some by (simp add: extract_Some_iff)
have eq: "as@U@bs@cs = as@U@(concat (rev acc @ as' @ [x])) @ (concat bs' @ cs)"
using "1.prems"(8) Some by (simp add: extract_Some_iff)
then have 0: "set (as@U@(concat (rev acc @ as' @ [x])) @ (concat bs' @ cs)) = \<Union> (set ` Y)"
using "1.prems"(4) by argo
have 2: "distinct (as@U@(concat (rev acc @ as' @ [x])) @ (concat bs' @ cs))"
using "1.prems"(5) eq by argo
have 3: "take 1 (as@U@(concat (rev acc @ as' @ [x])) @ (concat bs' @ cs)) = [r]"
using "1.prems"(6) eq by argo
have 4: "forward (as@U@(concat (rev acc @ as' @ [x])) @ (concat bs' @ cs))"
using "1.prems"(7) eq by argo
have 5: "concat (rev (rev as' @ acc) @ [x]) = concat (rev acc @ as' @ [x])" by simp
have 6: "\<forall>xs\<in>Y. sublist xs as \<or> sublist xs U
\<or> (\<exists>x\<in>set ((rev as' @ acc) @ [x]). sublist xs x) \<or> sublist xs (concat bs' @ cs)"
using "1.prems"(9) x(2) sublist_split_concat'[of acc as' x bs', where cs = cs]
by auto
have 7: "\<not> (\<exists>x'\<in>set U. \<exists>y\<in>set x. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)" using "1.prems"(11) x(2,3) by fastforce
have 8: "\<forall>xs. fwd_sub r Y xs
\<longrightarrow> cost (rev (as@U@concat(rev acc@as'@[x])@concat bs'@cs)) \<le> cost (rev xs)"
using "1.prems"(13) eq by simp
have notP: "\<forall>x\<in>set (combine_lists_P P x (rev as' @ acc)). \<not> P x \<and> forward x"
using "1.prems"(14-17) x(2)
combine_lists_notP[OF 1(2) \<open>x\<noteq>[]\<close> 1(3,4) 0 2 3 4 5 6 1(11) 7 x(3) 1(13) x(1) 8]
by auto
have cnct: "concat (rev (combine_lists_P P x (rev as' @ acc)) @ bs') = bs"
using "1.prems"(8) combine_lists_concat_rev_eq[of P] x(2) by simp
have sblst: "\<forall>xs\<in>Y. sublist xs as \<or> sublist xs U
\<or> (\<exists>a\<in>set (combine_lists_P P x (rev as' @ acc) @ bs'). sublist xs a) \<or> sublist xs cs"
using "1.prems"(9) x(2) combine_lists_sublists[of x "rev as'@acc", where P=P] by auto
have "\<forall>x\<in>set (combine_lists_P P x (rev as' @ acc)). x \<noteq> []"
using combine_lists_nempty[of "rev as' @ acc"] "1.prems"(14,16) x(2) by auto
then have "\<forall>x\<in>set (make_list_P P (combine_lists_P P x (rev as' @ acc)) bs'). \<not> P x"
using "1.IH"[OF Some x_def[symmetric] refl 1(2-8) cnct sblst 1(11-14)] notP x(2) 1(15,16)
by simp
then show ?thesis using Some by simp
qed
qed
lemma no_back_reach1_if_fwd_dstct_bs:
"\<lbrakk>forward (as@concat bs@V@cs); distinct (as@concat bs@V@cs); xs \<in> set bs\<rbrakk>
\<Longrightarrow> \<not>(\<exists>x'\<in>set V. \<exists>y\<in>set xs. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
using no_back_reach1_if_fwd_dstct[of "as@concat bs" "V@cs"] by auto
lemma mid_ranks_ge_if_reach1:
assumes "[] \<notin> Y"
and "U \<in> Y"
and "distinct (as@U@bs@V@cs)"
and "forward (as@U@bs@V@cs)"
and "concat bs' = bs"
and "concat cs' = cs"
and "set (as'@U#bs'@V#cs') = Y"
and "\<And>xs. \<lbrakk>xs \<in> Y; \<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
shows "\<forall>xs \<in> set bs'. (\<exists>x\<in>set U. \<exists>y\<in>set xs. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<longrightarrow> rank (rev V) \<le> rank (rev xs)"
proof -
have "\<forall>xs \<in> set bs'. \<forall>y\<in>set xs. \<not>(\<exists>x\<in>set V. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
using assms(3-6) no_back_reach1_if_fwd_dstct_bs[of "as@U"] by fastforce
then have 0: "\<forall>xs \<in> set bs'. (\<exists>y\<in>set xs. \<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)
\<longrightarrow> (\<exists>y\<in>set xs. \<exists>x\<in>set U. \<not> (\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
by blast
have "\<forall>xs \<in> set bs'. xs \<noteq> U"
using assms(1-3,5) concat_all_sublist empty_if_sublist_dsjnt[of U U] by fastforce
then have "\<And>xs. \<lbrakk>xs \<in> set bs'; \<exists>y\<in>set xs. \<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> xs \<noteq> U \<and> (\<exists>y\<in>set xs. \<exists>x\<in>set U. \<not> (\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> xs \<in> Y"
using 0 assms(7) by auto
then show ?thesis using assms(8) by blast
qed
lemma bs_ranks_only_ge:
assumes "asi rank r cost"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "r \<notin> set U"
and "U \<in> Y"
and "set (as@U@bs@V@cs) = \<Union>(set ` Y)"
and "distinct (as@U@bs@V@cs)"
and "take 1 (as@U@bs@V@cs) = [r]"
and "forward (as@U@bs@V@cs)"
and "concat as' = as"
and "concat bs' = bs"
and "concat cs' = cs"
and "set (as'@U#bs'@V#cs') = Y"
and "rank (rev V) \<le> rank (rev U)"
and "\<forall>zs. fwd_sub r Y zs \<longrightarrow> cost (rev (as@U@bs@V@cs)) \<le> cost (rev zs)"
and "\<And>xs. \<lbrakk>xs \<in> Y; \<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
shows "\<exists>zs. concat zs = bs \<and> (\<forall>z \<in> set zs. rank (rev V) \<le> rank (rev z)) \<and> [] \<notin> set zs"
proof -
let ?P = "\<lambda>x. rank (rev x) < rank (rev V)"
have "U \<noteq> []" using assms(3,5) by blast
have cnct: "concat (rev [] @ bs') = bs" using assms(11) by simp
have "\<forall>xs\<in>Y. sublist xs as \<or> xs = U \<or> xs = V
\<or> (\<exists>x\<in>set ([] @ bs'). sublist xs x) \<or> sublist xs cs"
using assms(10,12,13) concat_all_sublist by auto
then have sblst:
"\<forall>xs\<in>Y. sublist xs as \<or> sublist xs U \<or> (\<exists>x\<in>set ([] @ bs'). sublist xs x) \<or> sublist xs (V@cs)"
using sublist_app_l by fast
have 0: "\<And>xs. \<lbrakk>xs \<in> set bs'; \<exists>x\<in>set U. \<exists>y\<in>set xs. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
using mid_ranks_ge_if_reach1[OF assms(3,5,7,9,11-13)] assms(16) by blast
have "\<forall>x\<in>set bs'. x \<noteq> []" using assms(3,13) by auto
moreover have 2: "\<forall>x\<in>set bs'. forward x" using assms(2,13) by auto
ultimately have "(\<forall>x\<in>set (make_list_P ?P [] bs'). rank (rev V) \<le> rank (rev x))"
using assms(15)
make_list_notP[OF assms(1,4) \<open>U\<noteq>[]\<close> assms(6-9) cnct sblst assms(14) 0 refl]
by fastforce
then show ?thesis
using assms(3,11,13) make_list_concat_rev_eq[of ?P "[]"] make_list_nempty[of "[]" bs'] by auto
qed
lemma cost_ge_if_all_bs_ge:
assumes "asi rank r cost"
and "V \<noteq> []"
and "distinct (as@ds@concat bs@V@cs)"
and "take 1 as = [r]"
and "forward V"
and "\<forall>z\<in>set bs. rank (rev V) \<le> rank (rev z)"
and "[] \<notin> set bs"
shows "cost (rev (as@ds@V@concat bs@cs)) \<le> cost (rev (as@ds@concat bs@V@cs))"
using assms proof(induction bs arbitrary: ds)
case (Cons b bs)
have 0: "distinct (as@(ds@b)@concat bs@V@cs)" using Cons.prems(3) by simp
have r_b: "rank (rev V) \<le> rank (rev b)" using Cons.prems(6) by simp
have "b \<noteq> []" using Cons.prems(7) by auto
have dst: "distinct ((as@ds)@V@b@concat bs@cs)" using Cons.prems(3) by auto
have "take 1 ((as@ds)@V@b@concat bs@cs) = [r]"
using Cons.prems(4) take1_singleton_app by metis
moreover have "take 1 ((as@ds)@b@V@concat bs@cs) = [r]"
using Cons.prems(4) take1_singleton_app by metis
ultimately have "cost (rev (as@ds@V@b@concat bs@cs)) \<le> cost (rev (as@ds@b@V@concat bs@cs))"
using asi_le_rfst[OF Cons.prems(1) r_b Cons.prems(2) \<open>b\<noteq>[]\<close> dst] by simp
then show ?case using Cons.IH[OF Cons.prems(1,2) 0] Cons.prems(4-7) by simp
qed(simp)
lemma bs_ge_if_all_ge:
assumes "asi rank r cost"
and "V \<noteq> []"
and "distinct (as@bs@V@cs)"
and "take 1 as = [r]"
and "forward V"
and "concat bs' = bs"
and "\<forall>z\<in>set bs'. rank (rev V) \<le> rank (rev z)"
and "[] \<notin> set bs'"
and "bs \<noteq> []"
shows "rank (rev V) \<le> rank (rev bs)"
proof -
have dst: "distinct (as@[]@concat bs'@V@cs)" using assms(3,6) by simp
then have cost_le: "cost (rev (as@V@bs@cs)) \<le> cost (rev (as@bs@V@cs))"
using cost_ge_if_all_bs_ge[OF assms(1,2) dst] assms(3-9) by simp
have tk1: "take 1 ((as)@bs@V@cs) = [r]" using assms(4) take1_singleton_app by metis
have tk1': "take 1 ((as)@V@bs@cs) = [r]" using assms(4) take1_singleton_app by metis
have dst: "distinct ((as)@V@bs@cs)" using assms(3) by auto
show ?thesis using asi_le_iff_rfst[OF assms(1,2,9) tk1' tk1 dst] cost_le by simp
qed
lemma bs_ge_if_optimal:
assumes "asi rank r cost"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "finite Y"
and "r \<notin> set U"
and "U \<in> Y"
and "V \<in> Y"
and "distinct (as@U@bs@V@cs)"
and "set (as@U@bs@V@cs) = \<Union>(set ` Y)"
and "\<forall>xs \<in> Y. sublist xs (as@U@bs@V@cs)"
and "take 1 (as@U@bs@V@cs) = [r]"
and "forward (as@U@bs@V@cs)"
and "bs \<noteq> []"
and "rank (rev V) \<le> rank (rev U)"
and "\<forall>zs. fwd_sub r Y zs \<longrightarrow> cost (rev (as@U@bs@V@cs)) \<le> cost (rev zs)"
and "\<And>xs. \<lbrakk>xs \<in> Y; \<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
shows "rank (rev V) \<le> rank (rev bs)"
proof -
obtain as' bs' cs' where bs'_def: "concat as' = as" "concat bs' = bs" "concat cs' = cs"
"set (as'@U#bs'@V#cs') = Y"
using concat_split_UV[OF assms(2,5,7-11)] assms(4,7,8) by blast
obtain bs2 where bs2_def:
"concat bs2 = bs" "(\<forall>z\<in>set bs2. rank (rev V) \<le> rank (rev z))" "[] \<notin> set bs2"
using bs_ranks_only_ge[OF assms(1,3,4,6,7,10,9,12,13) bs'_def assms(15-17)] by blast
have "V \<noteq> []" using assms(4,8) by blast
have "take 1 as = [r]" using take1_split_nelem_nempty[OF assms(12)] assms(4,6,7) by blast
then have "take 1 (as@U) = [r]" using take1_singleton_app by fast
then show ?thesis
using bs_ge_if_all_ge[OF assms(1) \<open>V\<noteq>[]\<close>, of "as@U"] bs2_def assms(3,8,9,14) by auto
qed
lemma bs_ranks_only_ge_r:
assumes "[] \<notin> Y"
and "distinct (as@U@bs@V@cs)"
and "forward (as@U@bs@V@cs)"
and "as = []"
and "concat bs' = bs"
and "concat cs' = cs"
and "set (U#bs'@V#cs') = Y"
and "\<And>xs. \<lbrakk>xs \<in> Y; \<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
shows "\<forall>z \<in> set bs'. rank (rev V) \<le> rank (rev z)"
proof -
have "U \<in> Y" using assms(7) by auto
then have "U \<noteq> []" using assms(1) by blast
have "V \<noteq> []" using assms(1,7) by auto
have 0: "\<And>xs. \<lbrakk>xs \<in> set bs'; \<exists>x\<in>set U. \<exists>y\<in>set xs. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
using mid_ranks_ge_if_reach1[OF assms(1) \<open>U\<in>Y\<close> assms(2,3,5,6), of "[]"] assms(7,8) by auto
have "\<exists>x y ys. x#y#ys= as@U@bs@V@cs"
using \<open>U\<noteq>[]\<close> \<open>V\<noteq>[]\<close> append_Cons append.left_neutral list.exhaust by metis
then have hd_T: "hd (as@U@bs@V@cs) \<in> verts T" using hd_in_verts_if_forward assms(3) by metis
moreover have "\<forall>x\<in>set bs'. \<forall>y\<in>set x. y \<in> set (as@U@bs@V@cs)" using assms(5) by auto
ultimately have "\<forall>x\<in>set bs'. \<forall>y\<in>set x. hd (U@bs@V@cs) \<rightarrow>\<^sup>*\<^bsub>T\<^esub> y"
using hd_reach_all_forward assms(3,4) by auto
then have 1: "\<forall>x\<in>set bs'. \<forall>y\<in>set x. hd U \<rightarrow>\<^sup>*\<^bsub>T\<^esub> y" using assms(1,7) by auto
have "\<forall>x\<in>set bs'. \<forall>y\<in>set x. y \<notin> set U" using assms(2,5) by auto
then have "\<forall>x\<in>set bs'. \<forall>y\<in>set x. y \<noteq> hd U" using assms(1,7) by fastforce
then have "\<forall>x\<in>set bs'. \<forall>y\<in>set x. hd U \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y" using 1 by blast
then have "\<forall>x\<in>set bs'. \<exists>y\<in>set x. hd U \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y" using assms(1,7) by auto
then show ?thesis using 0 \<open>U \<noteq> []\<close> hd_in_set by blast
qed
lemma bs_ge_if_rU:
assumes "asi rank r cost"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "finite Y"
and "r \<in> set U"
and "U \<in> Y"
and "V \<in> Y"
and "distinct (as@U@bs@V@cs)"
and "set (as@U@bs@V@cs) = \<Union>(set ` Y)"
and "\<forall>xs \<in> Y. sublist xs (as@U@bs@V@cs)"
and "take 1 (as@U@bs@V@cs) = [r]"
and "forward (as@U@bs@V@cs)"
and "bs \<noteq> []"
and "\<And>xs. \<lbrakk>xs \<in> Y; \<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
shows "rank (rev V) \<le> rank (rev bs)"
proof -
obtain as' bs' cs' where bs'_def: "concat as' = as" "concat bs' = bs" "concat cs' = cs"
"set (as'@U#bs'@V#cs') = Y"
using concat_split_UV[OF assms(2,5,7-11)] assms(4,7,8) by blast
have "take 1 U = [r]" using take1_mid_if_elem[OF assms(12,6,9)] .
moreover have "as = []" using take1_empty_if_mid[OF assms(12,6,9)] .
ultimately have tk1: "take 1 (as@U) = [r]" by simp
then have "set (U#bs'@V#cs') = Y" using bs'_def(1,4) assms(4) \<open>as=[]\<close> by auto
then have 0: "(\<forall>z\<in>set bs'. rank (rev V) \<le> rank (rev z))"
using bs_ranks_only_ge_r[OF assms(4,9,13) \<open>as=[]\<close> bs'_def(2,3)] assms(15) by blast
have "V \<noteq> []" using assms(4,8) by blast
have "[] \<notin> set bs'" using assms(4) bs'_def(2,4) by auto
then show ?thesis
using bs_ge_if_all_ge[OF assms(1) \<open>V\<noteq>[]\<close>, of "as@U"] 0 bs'_def(2) tk1 assms(3,8,9,14) by auto
qed
lemma sublist_before_if_before:
assumes "hd xs = root" and "forward xs" and "distinct xs"
and "sublist U xs" and "sublist V xs" and "before U V"
shows "\<exists>as bs cs. as @ U @ bs @ V @ cs = xs"
proof (rule ccontr)
assume "\<nexists>as bs cs. as @ U @ bs @ V @ cs = xs"
then obtain as bs cs where V_bf_U: "xs = as @ V @ bs @ U @ cs"
using sublist_behind_if_nbefore[OF assms(4,5)] assms(6) before_def by blast
obtain x y where x_def: "x \<in> set U" "y \<in> set V" "x \<rightarrow>\<^bsub>T\<^esub> y"
using assms(6) before_def by auto
then obtain i where i_def: "V!i = y" "i < length V" by (auto simp: in_set_conv_nth)
then have i_xs: "(as@V@bs@U@cs)!(i + length as) = y" by (simp add: nth_append)
have "root \<noteq> y" using x_def(3) dominated_not_root by auto
then have "i + length as > 0" using i_def(2) i_xs assms(1,5) V_bf_U hd_conv_nth[of xs] by force
then have "i + length as \<ge> 1" by linarith
then have "i + length as \<in> {1..length (as@V@bs@U@cs) - 1}" using i_def(2) by simp
then obtain j where j_def: "j < i + length as" "(as@V@bs@U@cs)!j \<rightarrow>\<^bsub>T\<^esub> y"
using assms(2) V_bf_U i_xs unfolding forward_def by blast
then have "(as@V@bs@U@cs)!j = (as@V)!j" using i_def(2) by (auto simp: nth_append)
then have "(as@V@bs@U@cs)!j \<in> set (as@V)" using i_def(2) j_def(1) nth_mem[of "j" "as@V"] by simp
then have "(as@V@bs@U@cs)!j \<noteq> x" using assms(3) V_bf_U x_def(1) by auto
then show False using j_def(2) x_def(3) two_in_arcs_contr by fastforce
qed
lemma forward_UV_lists_subset:
"{x. set x = X \<and> distinct x \<and> take 1 x = [r] \<and> forward x \<and> (\<forall>xs \<in> Y. sublist xs x)}
\<subseteq> {x. set x = X \<and> distinct x}"
by blast
lemma forward_UV_lists_finite:
"finite xs
\<Longrightarrow> finite {x. set x = xs \<and> distinct x \<and> take 1 x = [r] \<and> forward x \<and> (\<forall>xs \<in> Y. sublist xs x)}"
using distinct_seteq_finite finite_subset[OF forward_UV_lists_subset] by auto
lemma forward_UV_lists_arg_min_ex_aux:
"\<lbrakk>finite ys; ys \<noteq> {};
ys = {x. set x = xs \<and> distinct x \<and> take 1 x = [r] \<and> forward x \<and> (\<forall>xs \<in> Y. sublist xs x)}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using arg_min_if_finite(1)[of ys f] arg_min_least[of ys, where ?f = f] by auto
lemma forward_UV_lists_arg_min_ex:
"\<lbrakk>finite xs; ys \<noteq> {};
ys = {x. set x = xs \<and> distinct x \<and> take 1 x = [r] \<and> forward x \<and> (\<forall>xs \<in> Y. sublist xs x)}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using forward_UV_lists_finite forward_UV_lists_arg_min_ex_aux by auto
lemma forward_UV_lists_argmin_ex':
fixes f :: "'a list \<Rightarrow> real"
assumes "P = (\<lambda>x. set x = X \<and> distinct x \<and> take 1 x = [r])"
and "Q = (\<lambda>ys. P ys \<and> forward ys \<and> (\<forall>xs \<in> Y. sublist xs ys))"
and "\<exists>x. Q x"
shows "\<exists>zs. Q zs \<and> (\<forall>as. Q as \<longrightarrow> f zs \<le> f as)"
using forward_UV_lists_arg_min_ex[of X "{x. Q x}"] using assms by fastforce
lemma forward_UV_lists_argmin_ex:
fixes f :: "'a list \<Rightarrow> real"
assumes "\<exists>x. fwd_sub r Y x"
shows "\<exists>zs. fwd_sub r Y zs \<and> (\<forall>as. fwd_sub r Y as \<longrightarrow> f zs \<le> f as)"
using forward_UV_lists_argmin_ex' assms unfolding fwd_sub_def unique_set_r_def by simp
lemma no_gap_if_contr_seq_fwd:
assumes "asi rank root cost"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "finite Y"
and "U \<in> Y"
and "V \<in> Y"
and "before U V"
and "rank (rev V) \<le> rank (rev U)"
and "\<And>xs. \<lbrakk>xs \<in> Y; \<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
and "\<exists>x. fwd_sub root Y x"
shows "\<exists>zs. fwd_sub root Y zs \<and> sublist (U@V) zs
\<and> (\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
proof -
obtain zs where zs_def:
"set zs = \<Union>(set ` Y)" "distinct zs" "take 1 zs = [root]" "forward zs"
"(\<forall>xs \<in> Y. sublist xs zs)" "(\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using forward_UV_lists_argmin_ex[OF assms(11), of "\<lambda>xs. cost (rev xs)"]
unfolding unique_set_r_def fwd_sub_def by blast
then have "hd zs = root" using hd_eq_take1 by fast
then obtain as bs cs where bs_def: "as @ U @ bs @ V @ cs = zs"
using sublist_before_if_before zs_def(2,4,5) assms(6-8) by blast
then have bs_prems: "distinct (as@U@bs@V@cs)" "set (as@U@bs@V@cs) = \<Union>(set ` Y)"
"\<forall>xs\<in>Y. sublist xs (as@U@bs@V@cs)" "take 1 (as@U@bs@V@cs) = [root]" "forward (as@U@bs@V@cs)"
using zs_def(1-5) by auto
show ?thesis
proof(cases "bs = []")
case True
then have "sublist (U@V) zs" using bs_def sublist_def by force
then show ?thesis using zs_def unfolding unique_set_r_def fwd_sub_def by blast
next
case bs_nempty: False
then have rank_le: "rank (rev V) \<le> rank (rev bs)"
proof(cases "root \<in> set U")
case True
then show ?thesis
using bs_ge_if_rU[OF assms(1-5) True assms(6,7) bs_prems bs_nempty assms(10)]
by blast
next
case False
have "\<forall>zs. fwd_sub root Y zs \<longrightarrow> cost (rev (as@U@bs@V@cs)) \<le> cost (rev zs)"
using zs_def(6) bs_def by blast
then show ?thesis
using bs_ge_if_optimal[OF assms(1-5)] bs_nempty bs_prems False assms(6,7,9,10)
by blast
qed
have 0: "distinct ((as@U)@V@bs@cs)" using bs_def zs_def(2) by auto
have "take 1 (as@U) = [root]"
using bs_def assms(4,6) take1_split_nempty[of U as] zs_def(3) by fastforce
then have 1: "take 1 (as@U@V@bs@cs) = [root]"
using take1_singleton_app[of "as@U" root "V@bs@cs"] by simp
have 2: "\<forall>xs\<in>Y. sublist xs (as@U@V@bs@cs)"
using zs_def(5) bs_def sublists_preserv_move_VY_all[OF assms(2,6,7)] assms(4,6) by blast
have "V \<noteq> []" using assms(4,7) by blast
have "cost (rev (as@U@V@bs@cs)) \<le> cost (rev zs)"
using asi_le_rfst[OF assms(1) rank_le \<open>V\<noteq>[]\<close> bs_nempty 0] 1 zs_def(3) bs_def by simp
then have cost_le: "\<forall>ys. fwd_sub root Y ys \<longrightarrow> cost (rev (as@U@V@bs@cs)) \<le> cost (rev ys)"
using zs_def(6) by fastforce
have "forward (as@U@V@bs@cs)"
using move_mid_backward_if_noarc assms(8) zs_def(4) bs_def by blast
moreover have "set (as@U@V@bs@cs) = \<Union> (set ` Y)"
unfolding zs_def(1)[symmetric] bs_def[symmetric] by force
ultimately have "fwd_sub root Y (as@U@V@bs@cs)"
unfolding unique_set_r_def fwd_sub_def using 0 1 2 by fastforce
moreover have "sublist (U@V) (as@U@V@bs@cs)" unfolding sublist_def by fastforce
ultimately show ?thesis using cost_le by blast
qed
qed
lemma combine_union_sets_alt:
fixes X Y
defines "Z \<equiv> X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}"
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> X. \<forall>ys \<in> X. xs = ys \<or> set xs \<inter> set ys = {}"
shows "Z = X \<union> (Y - {x. set x \<inter> \<Union>(set ` X) \<noteq> {}})"
unfolding assms(1) using assms(2,3) by fast
lemma combine_union_sets_disjoint:
fixes X Y
defines "Z \<equiv> X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}"
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> X. \<forall>ys \<in> X. xs = ys \<or> set xs \<inter> set ys = {}"
shows "\<forall>xs \<in> Z. \<forall>ys \<in> Z. xs = ys \<or> set xs \<inter> set ys = {}"
unfolding Z_def using assms(2,3) by force
lemma combine_union_sets_set_sub1_aux:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>ys \<in> X. \<exists>U \<in> Y. \<exists>V \<in> Y. U@V = ys"
and "x \<in> \<Union>(set ` Y)"
shows "x \<in> \<Union>(set ` (X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}))"
proof -
let ?Z = "X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}"
obtain ys where ys_def: "x \<in> set ys" "ys \<in> Y" using assms(3) by blast
then show ?thesis
proof(cases "ys \<in> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}")
case True
then show ?thesis using ys_def(1) by auto
next
case False
then obtain U V where U_def: "U \<in> Y" "V \<in> Y" "U@V \<in> X" "set ys \<inter> set (U@V) \<noteq> {}"
using ys_def(2) assms(2) by fast
then consider "set ys \<inter> set U \<noteq> {}" | "set ys \<inter> set V \<noteq> {}" by fastforce
then show ?thesis
proof(cases)
case 1
then have "U = ys" using assms(1) U_def(1) ys_def(2) by blast
then show ?thesis using ys_def(1) U_def(3) by fastforce
next
case 2
then have "V = ys" using assms(1) U_def(2) ys_def(2) by blast
then show ?thesis using ys_def(1) U_def(3) by fastforce
qed
qed
qed
lemma combine_union_sets_set_sub1:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>ys \<in> X. \<exists>U \<in> Y. \<exists>V \<in> Y. U@V = ys"
shows "\<Union>(set ` Y) \<subseteq> \<Union>(set ` (X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}))"
using combine_union_sets_set_sub1_aux[OF assms] by blast
lemma combine_union_sets_set_sub2:
assumes "\<forall>ys \<in> X. \<exists>U \<in> Y. \<exists>V \<in> Y. U@V = ys"
shows "\<Union>(set ` (X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}})) \<subseteq> \<Union>(set ` Y)"
using assms by fastforce
lemma combine_union_sets_set_eq:
assumes "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>ys \<in> X. \<exists>U \<in> Y. \<exists>V \<in> Y. U@V = ys"
shows "\<Union>(set ` (X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}})) = \<Union>(set ` Y)"
using combine_union_sets_set_sub1[OF assms] combine_union_sets_set_sub2[OF assms(2)] by blast
lemma combine_union_sets_sublists:
assumes "sublist x ys"
and "\<forall>xs \<in> X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}. sublist xs ys"
and "xs \<in> insert x X \<union> {xs. xs \<in> Y \<and> set xs \<inter> \<Union>(set ` (insert x X)) = {}}"
shows "sublist xs ys"
using assms by auto
lemma combine_union_sets_optimal_cost:
assumes "asi rank root cost"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "finite Y"
and "\<exists>x. fwd_sub root Y x"
and "\<forall>ys \<in> X. \<exists>U \<in> Y. \<exists>V \<in> Y. U@V = ys \<and> before U V \<and> rank (rev V) \<le> rank (rev U)
\<and> (\<forall>xs \<in> Y. (\<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> xs \<noteq> U)
\<longrightarrow> rank (rev V) \<le> rank (rev xs))"
and "\<forall>xs \<in> X. \<forall>ys \<in> X. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> X. \<forall>ys \<in> X. xs = ys \<or> \<not>(\<exists>x\<in>set xs. \<exists>y\<in>set ys. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
and "finite X"
shows "\<exists>zs. fwd_sub root (X \<union> {x. x \<in> Y \<and> set x \<inter> \<Union>(set ` X) = {}}) zs
\<and> (\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using assms(10,1-9) proof(induction X rule: finite_induct)
case empty
then show ?case using forward_UV_lists_argmin_ex by simp
next
case (insert x X)
let ?Y = "X \<union> {xs. xs \<in> Y \<and> set xs \<inter> \<Union>(set ` X) = {}}"
let ?X = "insert x X \<union> {xs. xs \<in> Y \<and> set xs \<inter> \<Union>(set ` (insert x X)) = {}}"
obtain zs where zs_def:
"fwd_sub root ?Y zs" "(\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using insert.IH[OF insert(4-9)] insert.prems(7,8,9) by auto
obtain U V where U_def: "U \<in> Y" "V \<in> Y" "U@V = x" "before U V" "rank (rev V) \<le> rank (rev U)"
"\<forall>xs \<in> Y. (\<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> xs \<noteq> U)
\<longrightarrow> rank (rev V) \<le> rank (rev xs)"
using insert.prems(7) by auto
then have U: "U \<in> ?Y" using insert.prems(2,8) insert.hyps(2) by fastforce
have V: "V \<in> ?Y" using U_def(2,3) insert.prems(8) insert.hyps(2) by fastforce
have disj: "\<forall>xs \<in> ?Y. \<forall>ys \<in> ?Y. xs = ys \<or> set xs \<inter> set ys = {}"
using combine_union_sets_disjoint[of Y X] insert.prems(2,8) by blast
have fwd: "\<forall>xs \<in> ?Y. forward xs"
using insert.prems(3,7) seq_conform_alt seq_conform_if_before by fastforce
have nempty: "[] \<notin> ?Y" using insert.prems(4,7) by blast
have fin: "finite ?Y" using insert.prems(5) insert.hyps(1) by simp
have 0: "\<And>xs. \<lbrakk>xs \<in> ?Y; \<exists>y\<in>set xs. \<not> (\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
using U_def(3,6) insert.prems(9) insert.hyps(2) by auto
then have "\<exists>zs. fwd_sub root ?Y zs \<and> sublist (U@V) zs
\<and> (\<forall>as. fwd_sub root ?Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using no_gap_if_contr_seq_fwd[OF insert.prems(1) disj fwd nempty fin U V U_def(4,5)] zs_def(1)
unfolding fwd_sub_def unique_set_r_def by blast
then obtain xs where xs_def:
"fwd_sub root ?Y xs" "sublist (U@V) xs"
"(\<forall>as. fwd_sub root ?Y as \<longrightarrow> cost (rev xs) \<le> cost (rev as))"
by blast
then have cost: "(\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev xs) \<le> cost (rev as))"
using zs_def by fastforce
have 0: "\<forall>ys \<in> (insert x X). \<exists>U \<in> Y. \<exists>V \<in> Y. U@V = ys" using insert.prems(7) by fastforce
then have "\<forall>ys \<in> X. \<exists>U \<in> Y. \<exists>V \<in> Y. U@V = ys" by simp
then have "\<Union>(set ` ?Y) = \<Union>(set ` Y)"
using combine_union_sets_set_eq[OF insert.prems(2)] by simp
then have "\<Union>(set ` ?X) = \<Union>(set ` ?Y)"
using combine_union_sets_set_eq[OF insert.prems(2) 0] by simp
then have P_eq: "unique_set_r root ?X = unique_set_r root ?Y" unfolding unique_set_r_def by simp
have "\<And>ys. \<lbrakk>sublist (U@V) ys; (\<forall>xs \<in> ?Y. sublist xs ys)\<rbrakk> \<Longrightarrow> (\<forall>xs \<in> ?X. sublist xs ys)"
using combine_union_sets_sublists[of x, where Y=Y and X=X] U_def(3) by blast
then have "\<And>ys. \<lbrakk>sublist (U@V) ys; fwd_sub root ?Y ys\<rbrakk> \<Longrightarrow> fwd_sub root ?X ys"
unfolding P_eq fwd_sub_def by blast
then show ?case using xs_def(1,2) cost by blast
qed
lemma bs_ge_if_geV:
assumes "asi rank r cost"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "finite Y"
and "U \<in> Y"
and "V \<in> Y"
and "distinct (as@U@bs@V@cs)"
and "set (as@U@bs@V@cs) = \<Union>(set ` Y)"
and "\<forall>xs \<in> Y. sublist xs (as@U@bs@V@cs)"
and "take 1 (as@U@bs@V@cs) = [r]"
and "bs \<noteq> []"
and "\<forall>xs \<in> Y. xs \<noteq> U \<longrightarrow> rank (rev V) \<le> rank (rev xs)"
shows "rank (rev V) \<le> rank (rev bs)"
proof -
obtain as' bs' cs' where bs'_def: "concat as' = as" "concat bs' = bs" "concat cs' = cs"
"set (as'@U#bs'@V#cs') = Y"
using concat_split_UV[OF assms(2,5-10)] assms(4,6,7) by blast
have tk1: "take 1 (as@U) = [r]"
using take1_split_nempty[of U as] assms(4,6,11) by force
have "\<forall>z\<in>set bs'. z \<noteq> U"
using bs'_def(2) assms(4,6,8) concat_all_sublist by (fastforce dest!: empty_if_sublist_dsjnt)
then have 0: "\<forall>z\<in>set bs'. rank (rev V) \<le> rank (rev z)"
using assms(13) bs'_def(4) by auto
have "V \<noteq> []" using assms(4,7) by blast
have "[] \<notin> set bs'" using assms(4) bs'_def(2,4) by auto
then show ?thesis
using bs_ge_if_all_ge[OF assms(1) \<open>V\<noteq>[]\<close>, of "as@U"] 0 bs'_def(2) tk1 assms(3,7,8,12) by auto
qed
lemma no_gap_if_geV:
assumes "asi rank root cost"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "finite Y"
and "U \<in> Y"
and "V \<in> Y"
and "before U V"
and "\<forall>xs \<in> Y. xs \<noteq> U \<longrightarrow> rank (rev V) \<le> rank (rev xs)"
and "\<exists>x. fwd_sub root Y x"
shows "\<exists>zs. fwd_sub root Y zs \<and> sublist (U@V) zs
\<and> (\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
proof -
obtain zs where zs_def:
"set zs = \<Union>(set ` Y)" "distinct zs" "take 1 zs = [root]" "forward zs"
"(\<forall>xs \<in> Y. sublist xs zs)" "(\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using forward_UV_lists_argmin_ex[OF assms(10), of "\<lambda>x. cost (rev x)"]
unfolding fwd_sub_def unique_set_r_def by blast
then have "hd zs = root" using hd_eq_take1 by fast
then obtain as bs cs where bs_def: "as @ U @ bs @ V @ cs = zs"
using sublist_before_if_before zs_def(2,4,5) assms(6-8) by blast
then have bs_prems: "distinct (as@U@bs@V@cs)" "set (as@U@bs@V@cs) = \<Union>(set ` Y)"
"\<forall>xs\<in>Y. sublist xs (as@U@bs@V@cs)" "take 1 (as@U@bs@V@cs) = [root]"
using zs_def(1-5) by auto
show ?thesis
proof(cases "bs = []")
case True
then have "sublist (U@V) zs" using bs_def sublist_def by force
then show ?thesis using zs_def unfolding fwd_sub_def unique_set_r_def by blast
next
case False
then have rank_le: "rank (rev V) \<le> rank (rev bs)"
using bs_ge_if_geV[OF assms(1-7) bs_prems False assms(9)] by blast
have 0: "distinct ((as@U)@V@bs@cs)" using bs_def zs_def(2) by auto
have "take 1 (as@U) = [root]"
using bs_def assms(4,6) take1_split_nempty[of U as] zs_def(3) by fastforce
then have 1: "take 1 (as@U@V@bs@cs) = [root]"
using take1_singleton_app[of "as@U" root "V@bs@cs"] by simp
have 2: "\<forall>xs\<in>Y. sublist xs (as@U@V@bs@cs)"
using zs_def(5) bs_def sublists_preserv_move_VY_all[OF assms(2,6,7)] assms(4,6) by blast
have "V \<noteq> []" using assms(4,7) by blast
have "cost (rev (as@U@V@bs@cs)) \<le> cost (rev zs)"
using asi_le_rfst[OF assms(1) rank_le \<open>V\<noteq>[]\<close> False 0] 1 zs_def(3) bs_def by simp
then have cost_le: "\<forall>ys. fwd_sub root Y ys \<longrightarrow> cost (rev (as@U@V@bs@cs)) \<le> cost (rev ys)"
using zs_def(6) by fastforce
have "forward (as@U@V@bs@cs)"
using move_mid_backward_if_noarc assms(8) zs_def(4) bs_def by blast
moreover have "set (as@U@V@bs@cs) = \<Union>(set ` Y)" using bs_def zs_def(1) by fastforce
ultimately have "fwd_sub root Y (as@U@V@bs@cs)"
unfolding fwd_sub_def unique_set_r_def using 0 1 2 by auto
moreover have "sublist (U@V) (as@U@V@bs@cs)" unfolding sublist_def by fastforce
ultimately show ?thesis using cost_le by blast
qed
qed
lemma app_UV_set_optimal_cost:
assumes "asi rank root cost"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. forward xs"
and "[] \<notin> Y"
and "finite Y"
and "U \<in> Y"
and "V \<in> Y"
and "before U V"
and "\<forall>xs \<in> Y. xs \<noteq> U \<longrightarrow> rank (rev V) \<le> rank (rev xs)"
and "\<exists>x. fwd_sub root Y x"
shows "\<exists>zs. fwd_sub root ({U@V} \<union> {x. x \<in> Y \<and> x \<noteq> U \<and> x \<noteq> V}) zs
\<and> (\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
proof -
have P_eq: "unique_set_r root Y = unique_set_r root ({U@V} \<union> {x. x \<in> Y \<and> x \<noteq> U \<and> x \<noteq> V})"
unfolding unique_set_r_def using assms(6,7) by auto
have "\<exists>zs. fwd_sub root Y zs \<and> sublist (U@V) zs
\<and> (\<forall>as. fwd_sub root Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using no_gap_if_geV[OF assms(1-10)] by blast
then show ?thesis unfolding P_eq fwd_sub_def by blast
qed
end
context tree_query_graph
begin
lemma no_cross_ldeep_rev_if_forward:
assumes "xs \<noteq> []" and "r \<in> verts G" and "directed_tree.forward (dir_tree_r r) (rev xs)"
shows "no_cross_products (create_ldeep_rev xs)"
using assms proof(induction xs rule: create_ldeep_rev.induct)
case (3 x y ys)
then interpret T: directed_tree "dir_tree_r r" r using directed_tree_r by blast
have split: "create_ldeep_rev (x#y#ys) = Join (create_ldeep_rev (y#ys)) (Relation x)" by simp
have "rev (x#y#ys) ! (length (y#ys)) = x" using nth_append_length[of "rev (y#ys)"] by simp
moreover have "length (y#ys) \<in> {1..length (rev (x#y#ys)) - 1}" by simp
ultimately obtain j where j_def: "j < (length (y#ys))" "rev (x#y#ys)!j \<rightarrow>\<^bsub>dir_tree_r r\<^esub> x"
using "3.prems"(3) unfolding T.forward_def by fastforce
then have "rev (x#y#ys)!j \<in> set (y#ys)"
using nth_mem[of j "rev (y#ys)"] by (auto simp add: nth_append)
then have "\<exists>x'\<in>relations (create_ldeep_rev (y#ys)). x' \<rightarrow>\<^bsub>dir_tree_r r\<^esub> x"
using j_def(2) create_ldeep_rev_relations[of "y#ys"] by blast
then have 1: "\<exists>x'\<in>relations (create_ldeep_rev (y#ys)). x' \<rightarrow>\<^bsub>G\<^esub>x"
using assms(2) dir_tree_r_dom_in_G by blast
have "T.forward (rev (y#ys))" using "3.prems"(3) T.forward_cons by blast
then show ?case using 1 3 by simp
qed(auto)
lemma no_cross_ldeep_if_forward:
"\<lbrakk>xs \<noteq> []; r \<in> verts G; directed_tree.forward (dir_tree_r r) xs\<rbrakk>
\<Longrightarrow> no_cross_products (create_ldeep xs)"
unfolding create_ldeep_def using no_cross_ldeep_rev_if_forward by simp
lemma no_cross_ldeep_if_forward':
"\<lbrakk>set xs = verts G; r \<in> verts G; directed_tree.forward (dir_tree_r r) xs\<rbrakk>
\<Longrightarrow> no_cross_products (create_ldeep xs)"
using no_cross_ldeep_if_forward[of xs] by fastforce
lemma forward_if_ldeep_rev_no_cross:
assumes "r \<in> verts G" and "no_cross_products (create_ldeep_rev xs)"
and "hd (rev xs) = r" and "distinct xs"
shows "directed_tree.forward_arcs (dir_tree_r r) xs"
using assms proof(induction xs rule: create_ldeep_rev.induct)
case 1
then show ?case using directed_tree_r directed_tree.forward_arcs.simps(1) by fast
next
case (2 x)
then show ?case using directed_tree_r directed_tree.forward_arcs.simps(2) by fast
next
case (3 x y ys)
then interpret T: directed_tree "dir_tree_r r" r using directed_tree_r by blast
have "hd (rev (y # ys)) = r" using "3.prems"(3) hd_append2[of "rev (y#ys)" "[x]"] by simp
then have ind: "T.forward_arcs (y#ys)" using 3 by fastforce
have matching: "matching_rels (create_ldeep_rev (x#y#ys))"
using matching_rels_if_no_cross "3.prems"(2) by simp
have "r \<in> relations (create_ldeep_rev (x#y#ys))" using "3.prems"(3)
using create_ldeep_rev_relations[of "x#y#ys"] hd_rev[of "x#y#ys"] by simp
then obtain p' where p'_def:
"awalk r p' x \<and> set (awalk_verts r p') \<subseteq> relations (create_ldeep_rev (x#y#ys))"
using no_cross_awalk[OF matching "3.prems"(2)] by force
then obtain p where p_def:
"apath r p x" "set (awalk_verts r p) \<subseteq> relations (create_ldeep_rev (x#y#ys))"
using apath_awalk_to_apath awalk_to_apath_verts_subset by blast
then have "pre_digraph.apath (dir_tree_r r) r p x" using apath_in_dir_if_apath_G by blast
moreover have "r \<noteq> x"
using "3.prems"(3,4) T.no_back_arcs.cases[of "rev (x#y#ys)"] distinct_first_uneq_last[of x]
by fastforce
ultimately obtain u where u_def:
"u \<rightarrow>\<^bsub>dir_tree_r r\<^esub> x" "u \<in> set (pre_digraph.awalk_verts (dir_tree_r r) r p)"
using p_def(2) T.awalk_verts_dom_if_uneq T.awalkI_apath by blast
then have "u \<in> relations (create_ldeep_rev (x#y#ys))"
using awalk_verts_G_T "3.prems"(1) p_def(2) by auto
then have "u \<in> set (x#y#ys)" by (simp add: create_ldeep_rev_relations)
then show ?case using u_def(1) ind T.forward_arcs.simps(3) T.loopfree.adj_not_same by auto
qed
lemma forward_if_ldeep_no_cross:
"\<lbrakk>r \<in> verts G; no_cross_products (create_ldeep xs); hd xs = r; distinct xs\<rbrakk>
\<Longrightarrow> directed_tree.forward (dir_tree_r r) xs"
using forward_if_ldeep_rev_no_cross directed_tree.forward_arcs_alt directed_tree_r
by (fastforce simp: create_ldeep_def)
lemma no_cross_ldeep_iff_forward:
"\<lbrakk>xs \<noteq> []; r \<in> verts G; hd xs = r; distinct xs\<rbrakk>
\<Longrightarrow> no_cross_products (create_ldeep xs) \<longleftrightarrow> directed_tree.forward (dir_tree_r r) xs"
using forward_if_ldeep_no_cross no_cross_ldeep_if_forward by blast
lemma no_cross_if_fwd_ldeep:
"\<lbrakk>r \<in> verts G; left_deep t; directed_tree.forward (dir_tree_r r) (inorder t)\<rbrakk>
\<Longrightarrow> no_cross_products t"
using no_cross_ldeep_if_forward[OF inorder_nempty] by fastforce
lemma forward_if_ldeep_no_cross':
"\<lbrakk>first_node t \<in> verts G; distinct_relations t; left_deep t; no_cross_products t\<rbrakk>
\<Longrightarrow> directed_tree.forward (dir_tree_r (first_node t)) (inorder t)"
using forward_if_ldeep_no_cross by (simp add: first_node_eq_hd distinct_relations_def)
lemma no_cross_iff_forward_ldeep:
"\<lbrakk>first_node t \<in> verts G; distinct_relations t; left_deep t\<rbrakk>
\<Longrightarrow> no_cross_products t \<longleftrightarrow> directed_tree.forward (dir_tree_r (first_node t)) (inorder t)"
using no_cross_if_fwd_ldeep forward_if_ldeep_no_cross' by blast
lemma sublist_before_if_before:
assumes "hd xs = r" and "no_cross_products (create_ldeep xs)" and "r \<in> verts G" and "distinct xs"
and "sublist U xs" and "sublist V xs" and "directed_tree.before (dir_tree_r r) U V"
shows "\<exists>as bs cs. as @ U @ bs @ V @ cs = xs"
using directed_tree.sublist_before_if_before[OF directed_tree_r] forward_if_ldeep_no_cross assms
by blast
lemma nocross_UV_lists_subset:
"{x. set x = X \<and> distinct x \<and> take 1 x = [r]
\<and> no_cross_products (create_ldeep x) \<and> (\<forall>xs \<in> Y. sublist xs x)}
\<subseteq> {x. set x = X \<and> distinct x}"
by blast
lemma nocross_UV_lists_finite:
"finite xs
\<Longrightarrow> finite {x. set x = xs \<and> distinct x \<and> take 1 x = [r]
\<and> no_cross_products (create_ldeep x) \<and> (\<forall>xs \<in> Y. sublist xs x)}"
using distinct_seteq_finite finite_subset[OF nocross_UV_lists_subset] by auto
lemma nocross_UV_lists_arg_min_ex_aux:
"\<lbrakk>finite ys; ys \<noteq> {};
ys = {x. set x = xs \<and> distinct x \<and> take 1 x = [r]
\<and> no_cross_products (create_ldeep x) \<and> (\<forall>xs \<in> Y. sublist xs x)}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using arg_min_if_finite(1)[of ys f] arg_min_least[of ys, where ?f = f] by auto
lemma nocross_UV_lists_arg_min_ex:
"\<lbrakk>finite xs; ys \<noteq> {};
ys = {x. set x = xs \<and> distinct x \<and> take 1 x = [r]
\<and> no_cross_products (create_ldeep x) \<and> (\<forall>xs \<in> Y. sublist xs x)}\<rbrakk>
\<Longrightarrow> \<exists>y \<in> ys. \<forall>z \<in> ys. (f :: 'a list \<Rightarrow> real) y \<le> f z"
using nocross_UV_lists_finite nocross_UV_lists_arg_min_ex_aux by auto
lemma nocross_UV_lists_argmin_ex:
fixes f :: "'a list \<Rightarrow> real"
assumes "P = (\<lambda>x. set x = X \<and> distinct x \<and> take 1 x = [r])"
and "Q = (\<lambda>ys. P ys \<and> no_cross_products (create_ldeep ys) \<and> (\<forall>xs \<in> Y. sublist xs ys))"
and "\<exists>x. Q x"
shows "\<exists>zs. Q zs \<and> (\<forall>as. Q as \<longrightarrow> f zs \<le> f as)"
using nocross_UV_lists_arg_min_ex[of X "{x. Q x}"] using assms by fastforce
lemma no_gap_if_contr_seq:
fixes Y r
defines "X \<equiv> \<Union>(set ` Y)"
defines "P \<equiv> (\<lambda>ys. set ys = X \<and> distinct ys \<and> take 1 ys = [r])"
defines "Q \<equiv> (\<lambda>ys. P ys \<and> no_cross_products (create_ldeep ys) \<and> (\<forall>xs \<in> Y. sublist xs ys))"
assumes "asi rank r c"
and "\<forall>xs \<in> Y. \<forall>ys \<in> Y. xs = ys \<or> set xs \<inter> set ys = {}"
and "\<forall>xs \<in> Y. directed_tree.forward (dir_tree_r r) xs"
and "[] \<notin> Y"
and "finite Y"
and "U \<in> Y"
and "V \<in> Y"
and "r \<in> verts G"
and "directed_tree.before (dir_tree_r r) U V"
and "rank (rev V) \<le> rank (rev U)"
and "\<And>xs. \<lbrakk>xs \<in> Y; \<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>dir_tree_r r\<^esub> y)
\<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>dir_tree_r r\<^esub> y); xs \<noteq> U\<rbrakk>
\<Longrightarrow> rank (rev V) \<le> rank (rev xs)"
and "\<exists>x. Q x"
shows "\<exists>zs. Q zs \<and> sublist (U@V) zs \<and> (\<forall>as. Q as \<longrightarrow> c (rev zs) \<le> c (rev as))"
proof -
interpret T: directed_tree "dir_tree_r r" r using assms(11) directed_tree_r by auto
let ?Q = "(\<lambda>ys. P ys \<and> T.forward ys \<and> (\<forall>xs \<in> Y. sublist xs ys))"
have "?Q = Q"
using no_cross_ldeep_iff_forward assms(11,2,3) hd_eq_take1 nempty_if_take1[where r=r] by fast
then show ?thesis
using T.no_gap_if_contr_seq_fwd[OF assms(4-10,12-14)] assms(15,1,2)
unfolding T.fwd_sub_def unique_set_r_def by auto
qed
end
subsection "Arc Invariants"
function path_lverts :: "('a list,'b) dtree \<Rightarrow> 'a \<Rightarrow> 'a set" where
"path_lverts (Node r {|(t,e)|}) x = (if x \<in> set r then {} else set r \<union> path_lverts t x)"
| "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> path_lverts (Node r xs) x = (if x \<in> set r then {} else set r)"
by (metis darcs_mset.cases old.prod.exhaust) fast+
termination by lexicographic_order
definition path_lverts_list :: "('a list \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'a set" where
"path_lverts_list xs x = (\<Union>(t,e)\<in> set (takeWhile (\<lambda>(t,e). x \<notin> set t) xs). set t)"
definition dom_children :: "('a list,'b) dtree \<Rightarrow> ('a,'b) pre_digraph \<Rightarrow> bool" where
"dom_children t1 T = (\<forall>t \<in> fst ` fset (sucs t1). \<forall>x \<in> dverts t.
\<exists>r \<in> set (root t1) \<union> path_lverts t (hd x). r \<rightarrow>\<^bsub>T\<^esub> hd x)"
abbreviation children_deg1 :: "(('a,'b) dtree \<times> 'b) fset \<Rightarrow> (('a,'b) dtree \<times> 'b) set" where
"children_deg1 xs \<equiv> {(t,e). (t,e) \<in> fset xs \<and> max_deg t \<le> 1}"
lemma path_lverts_subset_dlverts: "path_lverts t x \<subseteq> dlverts t"
by(induction t x rule: path_lverts.induct) auto
lemma path_lverts_to_list_eq:
"path_lverts t x = path_lverts_list (dtree_to_list (Node r0 {|(t,e)|})) x"
by (induction t rule: dtree_to_list.induct) (auto simp: path_lverts_list_def)
lemma path_lverts_from_list_eq:
"path_lverts (dtree_from_list r0 ys) x = path_lverts_list ((r0,e0)#ys) x"
unfolding path_lverts_list_def using path_lverts.simps(2)[of "{||}"]
by (induction ys rule: dtree_from_list.induct) (force, cases "x \<in> set r0", auto)
lemma path_lverts_child_union_root_sub:
assumes "t2 \<in> fst ` fset (sucs t1)"
shows "path_lverts t1 x \<subseteq> set (root t1) \<union> path_lverts t2 x"
proof(cases "\<forall>x. sucs t1 \<noteq> {|x|}")
case True
then show ?thesis using path_lverts.simps(2)[of "sucs t1" "root t1"] by simp
next
case False
then obtain e2 where "sucs t1 = {|(t2,e2)|}" using assms by fastforce
then show ?thesis
using path_lverts.simps(1)[of "root t1" t2 e2] dtree.collapse[of t1]
by(cases "x \<in> set (root t1)") fastforce+
qed
lemma path_lverts_simps1_sucs:
"\<lbrakk>x \<notin> set (root t1); sucs t1 = {|(t2,e2)|}\<rbrakk>
\<Longrightarrow> set (root t1) \<union> path_lverts t2 x = path_lverts t1 x"
using path_lverts.simps(1)[of "root t1" t2 e2 x] dtree.exhaust_sel[of t1] by argo
lemma subtree_path_lverts_sub:
"\<lbrakk>wf_dlverts t1; max_deg t1 \<le> 1; is_subtree (Node r xs) t1; t2 \<in> fst ` fset xs; x\<in>set (root t2)\<rbrakk>
\<Longrightarrow> set r \<subseteq> path_lverts t1 x"
proof(induction t1)
case (Node r1 xs1)
then have "xs1 \<noteq> {||}" by force
then have "max_deg (Node r1 xs1) = 1"
using Node.prems(2) empty_if_mdeg_0[of r1 xs1] by fastforce
then obtain t e where t_def: "xs1 = {|(t,e)|}" using mdeg_1_singleton by fastforce
have x_t2: "x \<in> dlverts t2" using Node.prems(5) lverts_if_in_verts dtree.set_sel(1) by fast
show ?case
proof(cases "Node r1 xs1 = Node r xs")
case True
then show ?thesis using Node.prems(1,4) x_t2 t_def by force
next
case False
then have 0: "is_subtree (Node r xs) t" using t_def Node.prems(3) by force
moreover have "max_deg t \<le> 1" using t_def Node.prems(2) mdeg_ge_child[of t e xs1] by simp
moreover have "x \<notin> set r1" using t_def x_t2 Node.prems(1,4) 0 subtree_in_dlverts by force
ultimately show ?thesis using Node.IH t_def Node.prems(1,4,5) by auto
qed
qed
lemma path_lverts_empty_if_roothd:
assumes "root t \<noteq> []"
shows "path_lverts t (hd (root t)) = {}"
proof(cases "\<forall>x. sucs t \<noteq> {|x|}")
case True
then show ?thesis using path_lverts.simps(2)[of "sucs t" "root t"] by force
next
case False
then obtain t1 e1 where t1_def: "sucs t = {|(t1, e1)|}" by auto
then have "path_lverts t (hd (root t)) =
(if hd (root t) \<in> set (root t) then {} else set (root t) \<union> path_lverts t1 (hd (root t)))"
using path_lverts.simps(1) dtree.collapse by metis
then show ?thesis using assms by simp
qed
lemma path_lverts_subset_root_if_childhd:
assumes "t1 \<in> fst ` fset (sucs t)" and "root t1 \<noteq> []"
shows "path_lverts t (hd (root t1)) \<subseteq> set (root t)"
proof(cases "\<forall>x. sucs t \<noteq> {|x|}")
case True
then show ?thesis using path_lverts.simps(2)[of "sucs t" "root t"] by simp
next
case False
then obtain e1 where "sucs t = {|(t1, e1)|}" using assms(1) by fastforce
then have "path_lverts t (hd (root t1)) =
(if hd (root t1) \<in> set (root t) then {} else set (root t) \<union> path_lverts t1 (hd (root t1)))"
using path_lverts.simps(1) dtree.collapse by metis
then show ?thesis using path_lverts_empty_if_roothd[OF assms(2)] by auto
qed
lemma path_lverts_list_merge_supset_xs_notin:
"\<forall>v \<in> fst ` set ys. a \<notin> set v
\<Longrightarrow> path_lverts_list xs a \<subseteq> path_lverts_list (Sorting_Algorithms.merge cmp xs ys) a"
proof(induction xs ys taking: cmp rule: Sorting_Algorithms.merge.induct)
case (3 x xs y ys)
obtain v1 e1 where v1_def[simp]: "x = (v1,e1)" by force
obtain v2 e2 where "y = (v2,e2)" by force
then show ?case using 3 by (auto simp: path_lverts_list_def)
qed (auto simp: path_lverts_list_def)
lemma path_lverts_list_merge_supset_ys_notin:
"\<forall>v \<in> fst ` set xs. a \<notin> set v
\<Longrightarrow> path_lverts_list ys a \<subseteq> path_lverts_list (Sorting_Algorithms.merge cmp xs ys) a"
proof(induction xs ys taking: cmp rule: Sorting_Algorithms.merge.induct)
case (3 x xs y ys)
obtain v1 e1 where v1_def[simp]: "x = (v1,e1)" by force
obtain v2 e2 where "y = (v2,e2)" by force
then show ?case using 3 by (auto simp: path_lverts_list_def)
qed (auto simp: path_lverts_list_def)
lemma path_lverts_list_merge_supset_xs:
"\<lbrakk>\<exists>v \<in> fst ` set xs. a \<in> set v; \<forall>v1 \<in> fst ` set xs. \<forall>v2 \<in> fst ` set ys. set v1 \<inter> set v2 = {}\<rbrakk>
\<Longrightarrow> path_lverts_list xs a \<subseteq> path_lverts_list (Sorting_Algorithms.merge cmp xs ys) a"
using path_lverts_list_merge_supset_xs_notin by fast
lemma path_lverts_list_merge_supset_ys:
"\<lbrakk>\<exists>v \<in> fst ` set ys. a \<in> set v; \<forall>v1 \<in> fst ` set xs. \<forall>v2 \<in> fst ` set ys. set v1 \<inter> set v2 = {}\<rbrakk>
\<Longrightarrow> path_lverts_list ys a \<subseteq> path_lverts_list (Sorting_Algorithms.merge cmp xs ys) a"
using path_lverts_list_merge_supset_ys_notin by fast
lemma dom_children_if_all_singletons:
"\<forall>(t1,e1) \<in> fset xs. dom_children (Node r {|(t1, e1)|}) T \<Longrightarrow> dom_children (Node r xs) T"
by (auto simp: dom_children_def)
lemma dom_children_all_singletons:
"\<lbrakk>dom_children (Node r xs) T; (t1,e1) \<in> fset xs\<rbrakk> \<Longrightarrow> dom_children (Node r {|(t1, e1)|}) T"
by (auto simp: dom_children_def)
lemma dom_children_all_singletons':
"\<lbrakk>dom_children (Node r xs) T; t1\<in> fst ` fset xs\<rbrakk> \<Longrightarrow> dom_children (Node r {|(t1, e1)|}) T"
by (auto simp: dom_children_def)
lemma root_arc_if_dom_root_child_nempty:
"\<lbrakk>dom_children (Node r xs) T; t1 \<in> fst ` fset xs; root t1 \<noteq> []\<rbrakk>
\<Longrightarrow> \<exists>x\<in>set r. \<exists>y\<in>set (root t1). x \<rightarrow>\<^bsub>T\<^esub> y"
unfolding dom_children_def using dtree.set_sel(1) path_lverts_empty_if_roothd[of t1]
by fastforce
lemma root_arc_if_dom_root_child_wfdlverts:
"\<lbrakk>dom_children (Node r xs) T; t1 \<in> fst ` fset xs; wf_dlverts t1\<rbrakk>
\<Longrightarrow> \<exists>x\<in>set r. \<exists>y\<in>set (root t1). x \<rightarrow>\<^bsub>T\<^esub> y"
using root_arc_if_dom_root_child_nempty dtree.set_sel(1)[of t1] empty_notin_wf_dlverts
by fastforce
lemma root_arc_if_dom_wfdlverts:
"\<lbrakk>dom_children (Node r xs) T; t1 \<in> fst ` fset xs; wf_dlverts (Node r xs)\<rbrakk>
\<Longrightarrow> \<exists>x\<in>set r. \<exists>y\<in>set (root t1). x \<rightarrow>\<^bsub>T\<^esub> y"
using root_arc_if_dom_root_child_wfdlverts[of r xs T t1] by fastforce
lemma children_deg1_sub_xs: "{(t,e). (t,e) \<in> fset xs \<and> max_deg t \<le> 1} \<subseteq> (fset xs)"
by blast
lemma finite_children_deg1: "finite {(t,e). (t,e) \<in> fset xs \<and> max_deg t \<le> 1}"
using children_deg1_sub_xs[of xs] by (simp add: finite_subset)
lemma finite_children_deg1': "{(t,e). (t,e) \<in> fset xs \<and> max_deg t \<le> 1} \<in> {A. finite A}"
using finite_children_deg1 by blast
lemma children_deg1_fset_id[simp]: "fset (Abs_fset (children_deg1 xs)) = children_deg1 xs"
using Abs_fset_inverse[OF finite_children_deg1'] by auto
lemma xs_sub_children_deg1: "\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1 \<Longrightarrow> (fset xs) \<subseteq> children_deg1 xs"
by auto
lemma children_deg1_full:
"\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1 \<Longrightarrow> (Abs_fset (children_deg1 xs)) = xs"
using xs_sub_children_deg1[of xs] children_deg1_sub_xs[of xs] by (simp add: fset_inverse)
locale ranked_dtree_with_orig = ranked_dtree t rank cmp + directed_tree T root
for t :: "('a list, 'b) dtree" and rank cost cmp and T :: "('a, 'b) pre_digraph" and root +
assumes asi_rank: "asi rank root cost"
and dom_mdeg_gt1:
"\<lbrakk>is_subtree (Node r xs) t; t1 \<in> fst ` fset xs; max_deg (Node r xs) > 1\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
and dom_sub_contr:
"\<lbrakk>is_subtree (Node r xs) t; t1 \<in> fst ` fset xs;
\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) (Node r xs) \<and> rank (rev (Dtree.root t2)) < rank (rev v)\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
and dom_contr:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; rank (rev (Dtree.root t1)) < rank (rev r);
max_deg (Node r {|(t1,e1)|}) = 1\<rbrakk>
\<Longrightarrow> dom_children (Node r {|(t1,e1)|}) T"
and dom_wedge:
"\<lbrakk>is_subtree (Node r xs) t; fcard xs > 1\<rbrakk>
\<Longrightarrow> dom_children (Node r (Abs_fset (children_deg1 xs))) T"
and arc_in_dlverts:
"\<lbrakk>is_subtree (Node r xs) t; x \<in> set r; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts (Node r xs)"
and verts_conform: "v \<in> dverts t \<Longrightarrow> seq_conform v"
and verts_distinct: "v \<in> dverts t \<Longrightarrow> distinct v"
begin
lemma dom_contr':
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; rank (rev (Dtree.root t1)) < rank (rev r);
max_deg (Node r {|(t1,e1)|}) \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r {|(t1,e1)|}) T"
using dom_contr mdeg_ge_sub mdeg_singleton[of r t1] by (simp add: fcard_single_1)
lemma dom_self_contr:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; rank (rev (Dtree.root t1)) < rank (rev r)\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using dom_sub_contr by fastforce
lemma dom_wedge_full:
"\<lbrakk>is_subtree (Node r xs) t; fcard xs > 1; \<forall>t \<in> fst ` fset xs. max_deg t \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r xs) T"
using dom_wedge children_deg1_full by fastforce
lemma dom_wedge_singleton:
"\<lbrakk>is_subtree (Node r xs) t; fcard xs > 1; t1 \<in> fst ` fset xs; max_deg t1 \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r {|(t1,e1)|}) T"
using dom_children_all_singletons' dom_wedge children_deg1_fset_id by fastforce
lemma arc_to_dverts_in_subtree:
"\<lbrakk>is_subtree (Node r xs) t; x \<in> set r; x \<rightarrow>\<^bsub>T\<^esub> y; y \<in> set v; v \<in> dverts t\<rbrakk>
\<Longrightarrow> v \<in> dverts (Node r xs)"
using list_in_verts_if_lverts[OF arc_in_dlverts] dverts_same_if_set_wf[OF wf_lverts]
dverts_subtree_subset by blast
lemma dlverts_arc_in_dlverts:
"\<lbrakk>is_subtree t1 t; x \<rightarrow>\<^bsub>T\<^esub> y; x \<in> dlverts t1\<rbrakk> \<Longrightarrow> y \<in> dlverts t1"
proof(induction t1)
case (Node r xs)
then show ?case
proof(cases "x \<in> set r")
case True
then show ?thesis using arc_in_dlverts Node.prems(1,2) by blast
next
case False
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "x \<in> dlverts t2"
using Node.prems(3) by auto
- then have "is_subtree t2 (Node r xs)" using subtree_if_child by fastforce
+ then have "is_subtree t2 (Node r xs)" using subtree_if_child
+ by (metis image_iff prod.sel(1))
then have "is_subtree t2 t" using Node.prems(1) subtree_trans by blast
then show ?thesis using Node.IH Node.prems(2) t2_def by fastforce
qed
qed
lemma dverts_arc_in_dlverts:
"\<lbrakk>is_subtree t1 t; v1 \<in> dverts t1; x \<in> set v1; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts t1"
using dlverts_arc_in_dlverts by (simp add: lverts_if_in_verts)
lemma dverts_arc_in_dverts:
assumes "is_subtree t1 t"
and "v1 \<in> dverts t1"
and "x \<in> set v1"
and "x \<rightarrow>\<^bsub>T\<^esub> y"
and "y \<in> set v2"
and "v2 \<in> dverts t"
shows "v2 \<in> dverts t1"
proof -
have "x \<in> dlverts t1" using assms(2,3) lverts_if_in_verts by fast
then obtain v where v_def: "v\<in>dverts t1" "y \<in> set v"
using list_in_verts_if_lverts[OF dlverts_arc_in_dlverts] assms(1-4) lverts_if_in_verts by blast
then show ?thesis
using dverts_same_if_set_wf[OF wf_lverts] assms(1,5,6) dverts_subtree_subset by blast
qed
lemma dlverts_reach1_in_dlverts:
"\<lbrakk>x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; is_subtree t1 t; x \<in> dlverts t1\<rbrakk> \<Longrightarrow> y \<in> dlverts t1"
by(induction x y rule: trancl.induct) (auto simp: dlverts_arc_in_dlverts)
lemma dlverts_reach_in_dlverts:
"\<lbrakk>x \<rightarrow>\<^sup>*\<^bsub>T\<^esub> y; is_subtree t1 t; x \<in> dlverts t1\<rbrakk> \<Longrightarrow> y \<in> dlverts t1"
using dlverts_reach1_in_dlverts by blast
lemma dverts_reach1_in_dlverts:
"\<lbrakk>is_subtree t1 t; v1 \<in> dverts t1; x \<in> set v1; x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts t1"
using dlverts_reach1_in_dlverts by (simp add: lverts_if_in_verts)
lemma dverts_reach_in_dlverts:
"\<lbrakk>is_subtree t1 t; v1 \<in> dverts t1; x \<in> set v1; x \<rightarrow>\<^sup>*\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts t1"
using list_in_verts_iff_lverts dverts_reach1_in_dlverts by (cases "x=y",fastforce,blast)
lemma dverts_reach1_in_dverts:
"\<lbrakk>is_subtree t1 t; v1 \<in> dverts t1; x \<in> set v1; x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; y \<in> set v2; v2 \<in> dverts t\<rbrakk>
\<Longrightarrow> v2 \<in> dverts t1"
by (meson dverts_reach1_in_dlverts dverts_arc_in_dverts list_in_verts_if_lverts tranclE)
lemma dverts_same_if_set_subtree:
"\<lbrakk>is_subtree t1 t; v1 \<in> dverts t1; x \<in> set v1; x \<in> set v2; v2 \<in> dverts t\<rbrakk> \<Longrightarrow> v1 = v2"
using dverts_same_if_set_wf[OF wf_lverts] dverts_subtree_subset by blast
lemma dverts_reach_in_dverts:
"\<lbrakk>is_subtree t1 t; v1 \<in> dverts t1; x \<in> set v1; x \<rightarrow>\<^sup>*\<^bsub>T\<^esub> y; y \<in> set v2; v2 \<in> dverts t\<rbrakk>
\<Longrightarrow> v2 \<in> dverts t1"
using dverts_same_if_set_subtree dverts_reach1_in_dverts by blast
lemma dverts_reach1_in_dverts_root:
"\<lbrakk>is_subtree t1 t; v \<in> dverts t; \<exists>x\<in>set (Dtree.root t1). \<exists>y\<in>set v. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> v \<in> dverts t1"
using dverts_reach1_in_dverts dtree.set_sel(1) by blast
lemma dverts_reach1_in_dverts_r:
"\<lbrakk>is_subtree (Node r xs) t; v \<in> dverts t; \<exists>x\<in>set r. \<exists>y\<in>set v. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> v \<in> dverts (Node r xs)"
using dverts_reach1_in_dverts[of "Node r xs"] by (auto intro: dtree.set_intros(1))
lemma dom_mdeg_gt1_subtree:
"\<lbrakk>is_subtree tn t; is_subtree (Node r xs) tn; t1 \<in> fst ` fset xs; max_deg (Node r xs) > 1\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using dom_mdeg_gt1 subtree_trans by blast
lemma dom_sub_contr_subtree:
"\<lbrakk>is_subtree tn t; is_subtree (Node r xs) tn; t1 \<in> fst ` fset xs;
\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) (Node r xs) \<and> rank (rev (Dtree.root t2)) < rank (rev v)\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using dom_sub_contr subtree_trans by blast
lemma dom_contr_subtree:
"\<lbrakk>is_subtree tn t; is_subtree (Node r {|(t1,e1)|}) tn; rank (rev (Dtree.root t1)) < rank (rev r);
max_deg (Node r {|(t1,e1)|}) = 1\<rbrakk>
\<Longrightarrow> dom_children (Node r {|(t1,e1)|}) T"
using dom_contr subtree_trans by blast
lemma dom_wedge_subtree:
"\<lbrakk>is_subtree tn t; is_subtree (Node r xs) tn; fcard xs > 1\<rbrakk>
\<Longrightarrow> dom_children (Node r (Abs_fset (children_deg1 xs))) T"
using dom_wedge subtree_trans by blast
corollary dom_wedge_subtree':
"is_subtree tn t \<Longrightarrow>\<forall>r xs. is_subtree (Node r xs) tn \<longrightarrow> fcard xs > 1
\<longrightarrow> dom_children (Node r (Abs_fset {(t, e). (t, e) \<in> fset xs \<and> max_deg t \<le> Suc 0})) T"
by (auto simp only: dom_wedge_subtree One_nat_def[symmetric])
lemma dom_wedge_full_subtree:
"\<lbrakk>is_subtree tn t; is_subtree (Node r xs) tn; fcard xs > 1; \<forall>t \<in> fst ` fset xs. max_deg t \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r xs) T"
using dom_wedge_full subtree_trans by fast
lemma arc_in_dlverts_subtree:
"\<lbrakk>is_subtree tn t; is_subtree (Node r xs) tn; x \<in> set r; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts (Node r xs)"
using arc_in_dlverts subtree_trans by blast
corollary arc_in_dlverts_subtree':
"is_subtree tn t \<Longrightarrow> \<forall>r xs. is_subtree (Node r xs) tn \<longrightarrow> (\<forall>x. x \<in> set r
\<longrightarrow> (\<forall>y. x \<rightarrow>\<^bsub>T\<^esub> y \<longrightarrow> y \<in> set r \<or> (\<exists>c\<in>fset xs. y \<in> dlverts (fst c))))"
using arc_in_dlverts_subtree by simp
lemma verts_conform_subtree: "\<lbrakk>is_subtree tn t; v \<in> dverts tn\<rbrakk> \<Longrightarrow> seq_conform v"
using verts_conform dverts_subtree_subset by blast
lemma verts_distinct_subtree: "\<lbrakk>is_subtree tn t; v \<in> dverts tn\<rbrakk> \<Longrightarrow> distinct v"
using verts_distinct dverts_subtree_subset by blast
lemma ranked_dtree_orig_subtree: "is_subtree x t \<Longrightarrow> ranked_dtree_with_orig x rank cost cmp T root"
unfolding ranked_dtree_with_orig_def ranked_dtree_with_orig_axioms_def
by (simp add: ranked_dtree_subtree directed_tree_axioms dom_mdeg_gt1_subtree dom_contr_subtree
dom_sub_contr_subtree dom_wedge_subtree' arc_in_dlverts_subtree'
verts_conform_subtree verts_distinct_subtree asi_rank)
corollary ranked_dtree_orig_rec:
"\<lbrakk>Node r xs = t; (x,e) \<in> fset xs\<rbrakk> \<Longrightarrow> ranked_dtree_with_orig x rank cost cmp T root"
using ranked_dtree_orig_subtree[of x] subtree_if_child[of x xs] by force
lemma child_disjoint_root:
"\<lbrakk>is_subtree (Node r xs) t; t1 \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> set r \<inter> set (Dtree.root t1) = {}"
using wf_dlverts_subtree[OF wf_lverts] dlverts_eq_dverts_union dtree.set_sel(1) by fastforce
lemma distint_verts_subtree:
assumes "is_subtree (Node r xs) t" and "t1 \<in> fst ` fset xs"
shows "distinct (r @ Dtree.root t1)"
proof -
have "(Dtree.root t1) \<in> dverts t" using dtree.set_sel(1) assms dverts_subtree_subset by fastforce
then show ?thesis
using verts_distinct assms(1) dverts_subtree_subset child_disjoint_root[OF assms] by force
qed
corollary distint_verts_singleton_subtree:
"is_subtree (Node r {|(t1,e1)|}) t \<Longrightarrow> distinct (r @ Dtree.root t1)"
using distint_verts_subtree by simp
lemma dom_between_child_roots:
assumes "is_subtree (Node r {|(t1,e1)|}) t" and "rank (rev (Dtree.root t1)) < rank (rev r)"
shows "\<exists>x\<in>set r. \<exists>y\<in>set (Dtree.root t1). x \<rightarrow>\<^bsub>T\<^esub> y"
using dom_self_contr[OF assms] wf_dlverts_subtree[OF wf_lverts assms(1)]
hd_in_set[of "Dtree.root t1"] dtree.set_sel(1)[of t1] empty_notin_wf_dlverts[of t1] by fastforce
lemma contr_before:
assumes "is_subtree (Node r {|(t1,e1)|}) t" and "rank (rev (Dtree.root t1)) < rank (rev r)"
shows "before r (Dtree.root t1)"
proof -
have "(Dtree.root t1) \<in> dverts t" using dtree.set_sel(1) assms(1) dverts_subtree_subset by fastforce
then have "seq_conform (Dtree.root t1)" using verts_conform by simp
moreover have "seq_conform r" using verts_conform assms(1) dverts_subtree_subset by force
ultimately show ?thesis
using before_def dom_between_child_roots[OF assms] child_disjoint_root[OF assms(1)] by auto
qed
lemma contr_forward:
assumes "is_subtree (Node r {|(t1,e1)|}) t" and "rank (rev (Dtree.root t1)) < rank (rev r)"
shows "forward (r@Dtree.root t1)"
proof -
have "(Dtree.root t1) \<in> dverts t" using dtree.set_sel(1) assms(1) dverts_subtree_subset by fastforce
then have "seq_conform (Dtree.root t1)" using verts_conform by simp
moreover have "seq_conform r" using verts_conform assms(1) dverts_subtree_subset by force
ultimately show ?thesis
using seq_conform_def forward_arcs_alt dom_self_contr assms forward_app by simp
qed
lemma contr_seq_conform:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; rank (rev (Dtree.root t1)) < rank (rev r)\<rbrakk>
\<Longrightarrow> seq_conform (r @ Dtree.root t1)"
using seq_conform_if_before contr_before by simp
lemma verts_forward: "\<forall>v \<in> dverts t. forward v"
using seq_conform_alt verts_conform by simp
lemma dverts_reachable1_if_dom_children_aux_root:
assumes "\<forall>v\<in>dverts (Node r xs). \<exists>x\<in>set r0 \<union> X \<union> path_lverts (Node r xs) (hd v). x \<rightarrow>\<^bsub>T\<^esub> hd v"
and "\<forall>y\<in>X. \<exists>x\<in>set r0. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
and "forward r"
shows "\<forall>y\<in>set r. \<exists>x\<in>set r0. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof(cases "r = []")
case False
then have "path_lverts (Node r xs) (hd r) = {}"
using path_lverts_empty_if_roothd[of "Node r xs"] by simp
then obtain x where x_def: "x\<in>set r0 \<union> X" "x \<rightarrow>\<^bsub>T\<^esub> hd r" using assms(1) by auto
then have "hd r \<in> verts T" using adj_in_verts(2) by auto
then have "\<forall>y\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using hd_reach_all_forward x_def(2) assms(3) reachable1_reachable_trans by blast
moreover obtain y where "y \<in> set r0" "y \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x" using assms(2) x_def by auto
ultimately show ?thesis using reachable_reachable1_trans by blast
qed(simp)
lemma dverts_reachable1_if_dom_children_aux:
"\<lbrakk>\<forall>v\<in>dverts t1. \<exists>x\<in>set r0 \<union> X \<union> path_lverts t1 (hd v). x \<rightarrow>\<^bsub>T\<^esub> hd v;
\<forall>y\<in>X. \<exists>x\<in>set r0. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; \<forall>v\<in>dverts t1. forward v; v\<in>dverts t1\<rbrakk>
\<Longrightarrow> \<forall>y\<in>set v. \<exists>x\<in>set r0. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof(induction t1 arbitrary: X rule: dtree_to_list.induct)
case (1 r t e)
have r_reachable1: "\<forall>y\<in>set r. \<exists>x\<in>set r0. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using dverts_reachable1_if_dom_children_aux_root[OF "1.prems"(1,2)] "1.prems"(3) by simp
then show ?case
proof(cases "r = v")
case True
then show ?thesis using r_reachable1 by simp
next
case False
have r_reach1: "\<forall>y\<in>set r \<union> X. \<exists>x\<in>set r0. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y" using "1.prems"(2) r_reachable1 by blast
have "\<forall>x. path_lverts (Node r {|(t, e)|}) x \<subseteq> set r \<union> path_lverts t x"
by simp
then have 0: "\<forall>v\<in>dverts t. \<exists>x\<in>set r0 \<union> (set r \<union> X) \<union> (path_lverts t (hd v)). x \<rightarrow>\<^bsub>T\<^esub> hd v"
using "1.prems"(1) by fastforce
then show ?thesis using "1.IH"[OF 0 r_reach1] "1.prems"(3,4) False by simp
qed
next
case (2 xs r)
then show ?case
proof(cases "\<exists>x\<in>set r0 \<union> X. x \<rightarrow>\<^bsub>T\<^esub> hd v")
case True
then obtain x where x_def: "x\<in>set r0 \<union> X" "x \<rightarrow>\<^bsub>T\<^esub> hd v" using "2.prems"(1,4) by blast
then have "hd v \<in> verts T" using x_def(2) adj_in_verts(2) by auto
moreover have "forward v" using "2.prems"(3,4) by blast
ultimately have v_reach1: "\<forall>y\<in>set v. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using hd_reach_all_forward x_def(2) reachable1_reachable_trans by blast
then show ?thesis using "2.prems"(2) x_def(1) reachable_reachable1_trans by blast
next
case False
then obtain x where x_def: "x \<in> path_lverts (Node r xs) (hd v)" "x \<rightarrow>\<^bsub>T\<^esub> hd v"
using "2.prems"(1,4) by blast
then have "x \<in> set r" using path_lverts.simps(2)[OF "2.hyps"] empty_iff by metis
then obtain x' where x'_def: "x'\<in>set r0" "x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> x"
using dverts_reachable1_if_dom_children_aux_root[OF "2.prems"(1,2)] "2.prems"(3) by auto
then have x'_v: "x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> hd v" using x_def(2) by simp
then have "hd v \<in> verts T" using x_def(2) adj_in_verts(2) by auto
moreover have "forward v" using "2.prems"(3,4) by blast
ultimately have v_reach1: "\<forall>y\<in>set v. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using hd_reach_all_forward x'_v reachable1_reachable_trans by blast
then show ?thesis using x'_def(1) by blast
qed
qed
lemma dlverts_reachable1_if_dom_children_aux:
"\<lbrakk>\<forall>v\<in>dverts t1. \<exists>x\<in>set r \<union> X \<union> path_lverts t1 (hd v). x \<rightarrow>\<^bsub>T\<^esub> hd v;
\<forall>y\<in>X. \<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; \<forall>v\<in>dverts t1. forward v; y\<in>dlverts t1\<rbrakk>
\<Longrightarrow> \<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using dverts_reachable1_if_dom_children_aux list_in_verts_iff_lverts[of y t1] by blast
lemma dverts_reachable1_if_dom_children:
assumes "dom_children t1 T" and "v \<in> dverts t1" and "v \<noteq> Dtree.root t1" and "\<forall>v\<in>dverts t1. forward v"
shows "\<forall>y\<in>set v. \<exists>x\<in>set (Dtree.root t1). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof -
obtain t2 where t2_def: "t2 \<in> fst ` fset (sucs t1)" "v \<in> dverts t2"
using assms(2,3) dverts_root_or_suc by force
then have 0: "\<forall>v\<in>dverts t2. \<exists>x\<in>set (Dtree.root t1) \<union> {} \<union> path_lverts t2 (hd v). x \<rightarrow>\<^bsub>T\<^esub> hd v"
using assms(1) unfolding dom_children_def by blast
moreover have "\<forall>v\<in>dverts t2. forward v" using assms(4) t2_def(1) dverts_suc_subseteq by blast
ultimately show ?thesis using dverts_reachable1_if_dom_children_aux t2_def(2) by blast
qed
lemma subtree_dverts_reachable1_if_mdeg_gt1:
"\<lbrakk>is_subtree t1 t; max_deg t1 > 1; v \<in> dverts t1; v \<noteq> Dtree.root t1\<rbrakk>
\<Longrightarrow> \<forall>y\<in>set v. \<exists>x\<in>set (Dtree.root t1). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof(induction t1)
case (Node r xs)
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "v \<in> dverts t2" by auto
then obtain x where x_def: "x\<in>set r" "x \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t2)"
using dom_mdeg_gt1 Node.prems(1,2) by fastforce
then have t2_T: "hd (Dtree.root t2) \<in> verts T" using adj_in_verts(2) by simp
have "is_subtree t2 (Node r xs)" using subtree_if_child[of t2 xs r] t2_def(1) by force
then have subt2: "is_subtree t2 t" using subtree_trans Node.prems(1) by blast
have "Dtree.root t2 \<in> dverts t"
using subt2 dverts_subtree_subset by (fastforce simp: dtree.set_sel(1))
then have fwd_t2: "forward (Dtree.root t2)" by (simp add: verts_forward)
then have t2_reach1: "\<forall>y\<in>set (Dtree.root t2). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using hd_reach_all_forward[OF t2_T fwd_t2] x_def(2) reachable1_reachable_trans by blast
then consider "Dtree.root t2 = v" | "Dtree.root t2 \<noteq> v" "max_deg t2 > 1" | "Dtree.root t2 \<noteq> v" "max_deg t2 \<le> 1"
by fastforce
then show ?case
proof(cases)
case 1
then show ?thesis using t2_reach1 x_def(1) by auto
next
case 2
then have "\<forall>y\<in>set v. \<exists>x\<in>set (Dtree.root t2). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y" using Node.IH subt2 t2_def by simp
then show ?thesis
using t2_reach1 x_def(1) reachable1_reachable reachable1_reachable_trans
unfolding dtree.sel(1) by blast
next
case 3
then have "fcard xs > 1" using Node.prems(2) t2_def(1) fcard_gt1_if_mdeg_gt_child1 by fastforce
then have dom: "dom_children (Node r {|(t2,e2)|}) T"
using dom_wedge_singleton[OF Node.prems(1)] t2_def(1) 3(2) by fastforce
have "\<forall>v \<in> dverts (Node r xs). forward v"
using Node.prems(1) seq_conform_alt verts_conform_subtree by blast
then have "\<forall>v \<in> dverts (Node r {|(t2, e2)|}). forward v" using t2_def(1) by simp
then show ?thesis
using dverts_reachable1_if_dom_children[OF dom] t2_def(2) Node.prems(4)
unfolding dtree.sel(1) by simp
qed
qed
lemma subtree_dverts_reachable1_if_mdeg_gt1_singleton:
assumes "is_subtree (Node r {|(t1,e1)|}) t"
and "max_deg (Node r {|(t1,e1)|}) > 1"
and "v \<in> dverts t1"
and "v \<noteq> Dtree.root t1"
shows "\<forall>y\<in>set v. \<exists>x\<in>set (Dtree.root t1). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof -
have "is_subtree t1 t" using subtree_trans[OF subtree_if_child assms(1)] by simp
then show ?thesis
using assms(2-4) mdeg_eq_child_if_singleton_gt1[OF assms(2)]
subtree_dverts_reachable1_if_mdeg_gt1 by simp
qed
lemma subtree_dverts_reachable1_if_mdeg_le1_subcontr:
"\<lbrakk>is_subtree t1 t; max_deg t1 \<le> 1; is_subtree (Node v2 {|(t2,e2)|}) t1;
rank (rev (Dtree.root t2)) < rank (rev v2); v \<in> dverts t1; v \<noteq> Dtree.root t1\<rbrakk>
\<Longrightarrow> \<forall>y\<in>set v. \<exists>x\<in>set (Dtree.root t1). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof(induction t1)
case (Node r xs)
then show ?case
proof(cases "Node v2 {|(t2,e2)|} = Node r xs")
case True
then have "dom_children (Node r xs) T" using dom_contr' Node.prems(1,2,4) by blast
moreover have "\<forall>v \<in> dverts (Node r xs). forward v"
using Node.prems(1) seq_conform_alt verts_conform_subtree by blast
ultimately show ?thesis using dverts_reachable1_if_dom_children Node.prems(5,6) by blast
next
case False
then obtain t3 e3 where t3_def: "(t3,e3) \<in> fset xs" "is_subtree (Node v2 {|(t2,e2)|}) t3"
using Node.prems(3) by auto
then have t3_xs: "xs = {|(t3,e3)|}"
using Node.prems(2) by (simp add: singleton_if_mdeg_le1_elem)
then have v_t3: "v \<in> dverts t3" using Node.prems(5,6) by simp
then have t3_dom: "\<exists>x\<in>set r. x \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t3)"
using dom_sub_contr Node.prems(1,3,4) t3_xs by fastforce
then have t3_T: "hd (Dtree.root t3) \<in> verts T" using adj_in_verts(2) by blast
have "is_subtree t3 (Node r xs)" using subtree_if_child[of t3 xs] t3_xs by simp
then have sub_t3: "is_subtree t3 t" using subtree_trans Node.prems(1) by blast
then have "Dtree.root t3 \<in> dverts t"
using dverts_subtree_subset by (fastforce simp: dtree.set_sel(1))
then have "forward (Dtree.root t3)" by (simp add: verts_forward)
then have t3_reach1: "\<exists>x\<in>set r. \<forall>y\<in>set(Dtree.root t3). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using hd_reach_all_forward[OF t3_T] t3_dom reachable1_reachable_trans by blast
show ?thesis
proof(cases "v = Dtree.root t3")
case True
then show ?thesis using t3_reach1 by auto
next
case False
moreover have "max_deg t3 \<le> 1" using Node.prems(2) t3_def(1) mdeg_ge_child by fastforce
ultimately have "\<forall>y\<in>set v. \<exists>x\<in>set (Dtree.root t3). x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
using Node.IH sub_t3 t3_def Node.prems(4) v_t3 by simp
then show ?thesis
using t3_reach1 reachable1_reachable_trans reachable1_reachable unfolding dtree.sel(1)
by blast
qed
qed
qed
lemma subtree_y_reach_if_mdeg_gt1_notroot_reach:
assumes "is_subtree (Node r {|(t1,e1)|}) t"
and "max_deg (Node r {|(t1,e1)|}) > 1"
and "v \<noteq> r"
and "v \<in> dverts t"
and "v \<noteq> Dtree.root t1"
and "y \<in> set v"
and "\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
shows "\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof -
have "v \<in> dverts (Node r {|(t1,e1)|})" using dverts_reach1_in_dverts_r assms(1,4,6,7) by blast
then show ?thesis using subtree_dverts_reachable1_if_mdeg_gt1_singleton assms(1-3,5,6) by simp
qed
lemma subtree_eqroot_if_mdeg_gt1_reach:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; max_deg (Node r {|(t1,e1)|}) > 1; v \<in> dverts t;
\<exists>y\<in>set v. \<not>(\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); v \<noteq> r\<rbrakk>
\<Longrightarrow> Dtree.root t1 = v"
using subtree_y_reach_if_mdeg_gt1_notroot_reach by blast
lemma subtree_rank_ge_if_mdeg_gt1_reach:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; max_deg (Node r {|(t1,e1)|}) > 1; v \<in> dverts t;
\<exists>y\<in>set v. \<not>(\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y); v \<noteq> r\<rbrakk>
\<Longrightarrow> rank (rev (Dtree.root t1)) \<le> rank (rev v)"
using subtree_eqroot_if_mdeg_gt1_reach by blast
lemma subtree_y_reach_if_mdeg_le1_notroot_subcontr:
assumes "is_subtree (Node r {|(t1,e1)|}) t"
and "max_deg (Node r {|(t1,e1)|}) \<le> 1"
and "is_subtree (Node v2 {|(t2,e2)|}) t1"
and "rank (rev (Dtree.root t2)) < rank (rev v2)"
and "v \<noteq> r"
and "v \<in> dverts t"
and "v \<noteq> Dtree.root t1"
and "y \<in> set v"
and "\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
shows "\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
proof -
have 0: "is_subtree t1 (Node r {|(t1,e1)|})" using subtree_if_child[of t1 "{|(t1,e1)|}"] by simp
then have subt1: "is_subtree t1 t" using assms(1) subtree_trans by blast
have "v \<in> dverts (Node r {|(t1,e1)|})"
using dverts_reach1_in_dverts_r assms(1,6,8,9) by blast
then have "v \<in> dverts t1" using assms(5) by simp
moreover have "max_deg t1 \<le> 1" using assms(2) mdeg_ge_sub[OF 0] by simp
ultimately show ?thesis
using subtree_dverts_reachable1_if_mdeg_le1_subcontr[OF subt1] assms(3,4,7,8) by blast
qed
lemma rank_ge_if_mdeg_le1_dvert_nocontr:
assumes "max_deg t1 \<le> 1"
and "\<nexists>v2 t2 e2. is_subtree (Node v2 {|(t2,e2)|}) t1 \<and> rank (rev (Dtree.root t2)) < rank (rev v2)"
and "v \<in> dverts t1"
shows "rank (rev (Dtree.root t1)) \<le> rank (rev v)"
using assms proof(induction t1)
case (Node r xs)
then show ?case
proof(cases "v = r")
case False
then obtain t2 e2 where t2_def: "xs = {|(t2,e2)|}" "v \<in> dverts t2"
using Node.prems(1,3) singleton_if_mdeg_le1_elem by fastforce
have "max_deg t2 \<le> 1" using Node.prems(1) mdeg_ge_child[of t2 e2 xs] t2_def(1) by simp
then have "rank (rev (Dtree.root t2)) \<le> rank (rev v)"
using Node.IH t2_def Node.prems(2) by fastforce
then show ?thesis using Node.prems(2) t2_def(1) by fastforce
qed(simp)
qed
lemma subtree_rank_ge_if_mdeg_le1_nocontr:
assumes "is_subtree (Node r {|(t1,e1)|}) t"
and "max_deg (Node r {|(t1,e1)|}) \<le> 1"
and "\<nexists>v2 t2 e2. is_subtree (Node v2 {|(t2,e2)|}) t1 \<and> rank (rev (Dtree.root t2)) < rank (rev v2)"
and "v \<noteq> r"
and "v \<in> dverts t"
and "y \<in> set v"
and "\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y"
shows "rank (rev (Dtree.root t1)) \<le> rank (rev v)"
proof -
have 0: "is_subtree t1 (Node r {|(t1,e1)|})" using subtree_if_child[of t1 "{|(t1,e1)|}"] by simp
then have 0: "max_deg t1 \<le> 1" using assms(2) mdeg_ge_sub[OF 0] by simp
have "v \<in> dverts (Node r {|(t1,e1)|})" using dverts_reach1_in_dverts_r assms(1,5-7) by blast
then have "v \<in> dverts t1" using assms(4) by simp
then show ?thesis using rank_ge_if_mdeg_le1_dvert_nocontr 0 assms(3) by blast
qed
lemma subtree_rank_ge_if_mdeg_le1':
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; max_deg (Node r {|(t1,e1)|}) \<le> 1; v \<noteq> r;
v \<in> dverts t; y \<in> set v; \<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y; \<not>(\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)\<rbrakk>
\<Longrightarrow> rank (rev (Dtree.root t1)) \<le> rank (rev v)"
using subtree_y_reach_if_mdeg_le1_notroot_subcontr subtree_rank_ge_if_mdeg_le1_nocontr
apply(cases "\<exists>v2 t2 e2. is_subtree (Node v2 {|(t2,e2)|}) t1 \<and> rank (rev (Dtree.root t2))<rank (rev v2)")
by blast+
lemma subtree_rank_ge_if_mdeg_le1:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; max_deg (Node r {|(t1,e1)|}) \<le> 1; v \<noteq> r;
v \<in> dverts t; \<exists>y \<in> set v. \<not>(\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)\<rbrakk>
\<Longrightarrow> rank (rev (Dtree.root t1)) \<le> rank (rev v)"
using subtree_y_reach_if_mdeg_le1_notroot_subcontr subtree_rank_ge_if_mdeg_le1_nocontr
apply(cases "\<exists>v2 t2 e2. is_subtree (Node v2 {|(t2,e2)|}) t1 \<and> rank (rev (Dtree.root t2))<rank (rev v2)")
by blast+
lemma subtree_rank_ge_if_reach:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) t; v \<noteq> r; v \<in> dverts t;
\<exists>y \<in> set v. \<not>(\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)\<rbrakk>
\<Longrightarrow> rank (rev (Dtree.root t1)) \<le> rank (rev v)"
using subtree_rank_ge_if_mdeg_le1 subtree_rank_ge_if_mdeg_gt1_reach
by (cases "max_deg (Node r {|(t1,e1)|}) \<le> 1") (auto simp del: max_deg.simps)
lemma subtree_rank_ge_if_reach':
"is_subtree (Node r {|(t1,e1)|}) t \<Longrightarrow> \<forall>v \<in> dverts t.
(\<exists>y\<in>set v. \<not> (\<exists>x'\<in>set (Dtree.root t1). x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set r. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> v \<noteq> r)
\<longrightarrow> rank (rev (Dtree.root t1)) \<le> rank (rev v)"
using subtree_rank_ge_if_reach by blast
subsubsection \<open>Normalizing preserves Arc Invariants\<close>
lemma normalize1_mdeg_le: "max_deg (normalize1 t1) \<le> max_deg t1"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
then show ?thesis using mdeg_child_sucs_le by fastforce
next
case False
then have "max_deg (normalize1 (Node r {|(t, e)|}))
= max (max_deg (normalize1 t)) (fcard {|(normalize1 t, e)|})"
using mdeg_singleton by force
then show ?thesis using mdeg_singleton[of r t] 1 False by (simp add: fcard_single_1)
qed
next
case (2 xs r)
then have 0: "\<forall>(t,e) \<in> fset xs. max_deg (normalize1 t) \<le> max_deg t" by fastforce
have "max_deg (normalize1 (Node r xs)) = max_deg (Node r ((\<lambda>(t,e). (normalize1 t,e)) |`| xs))"
using "2.hyps" by simp
then show ?case using mdeg_img_le'[OF 0] by simp
qed
lemma normalize1_mdeg_eq:
"wf_darcs t1
\<Longrightarrow> max_deg (normalize1 t1) = max_deg t1 \<or> (max_deg (normalize1 t1) = 0 \<and> max_deg t1 = 1)"
proof(induction t1 rule: normalize1.induct)
case ind: (1 r t e)
then have 0: "max_deg (Node r {|(t, e)|}) \<ge> 1"
using mdeg_ge_fcard[of "{|(t, e)|}"] by (simp add: fcard_single_1)
then consider "rank (rev (Dtree.root t)) < rank (rev r)"
| "\<not>rank (rev (Dtree.root t)) < rank (rev r)" "max_deg (normalize1 t) \<le> 1"
| "\<not>rank (rev (Dtree.root t)) < rank (rev r)" "max_deg (normalize1 t) > 1" by linarith
then show ?case
proof(cases)
case 1
then show ?thesis
using mdeg_singleton mdeg_root fcard_single_1
by (metis max_def nle_le dtree.exhaust_sel leI less_one normalize1.simps(1))
next
case 2
then have "max_deg (normalize1 (Node r {|(t, e)|})) = 1"
using mdeg_singleton[of r "normalize1 t"] by (auto simp: fcard_single_1)
moreover have "max_deg (Node r {|(t, e)|}) = 1 "
using mdeg_singleton[of r t] ind 2
by (auto simp: fcard_single_1 wf_darcs_iff_darcs')
ultimately show ?thesis by simp
next
case 3
then show ?thesis
using mdeg_singleton[of r t] mdeg_singleton[of r "normalize1 t"] ind
by (auto simp: fcard_single_1)
qed
next
case ind: (2 xs r)
then consider "max_deg (Node r xs) \<le> 1"
| "max_deg (Node r xs) > 1" "max_deg (Node r xs) = fcard xs"
| "max_deg (Node r xs) > 1" "fcard xs < max_deg (Node r xs)"
using mdeg_ge_fcard[of xs] by fastforce
then show ?case
proof(cases)
case 1
then show ?thesis using normalize1_mdeg_le[of "Node r xs"] by fastforce
next
case 2
then have "max_deg (Node r xs) \<le> max_deg (normalize1 (Node r xs))"
using mdeg_ge_fcard[of "(\<lambda>(t, e). (normalize1 t, e)) |`| xs"] ind
by (simp add: fcard_normalize_img_if_disjoint wf_darcs_iff_darcs')
then show ?thesis using normalize1_mdeg_le[of "Node r xs"] by simp
next
case 3
then obtain t e where t_def: "(t,e) \<in> fset xs" "max_deg (Node r xs) = max_deg t"
using mdeg_child_if_gt_fcard by fastforce
have "max_deg (normalize1 t) \<le> max_deg (Node r ((\<lambda>(t,e). (normalize1 t,e)) |`| xs))"
using mdeg_ge_child[of "normalize1 t" e "(\<lambda>(t,e). (normalize1 t,e)) |`| xs" r] t_def(1)
by fastforce
then have "max_deg (Node r xs) \<le> max_deg (normalize1 (Node r xs))"
using ind.hyps ind.IH[OF t_def(1) refl] ind.prems 3(1) t_def
by (fastforce simp: wf_darcs_iff_darcs')
then show ?thesis using normalize1_mdeg_le[of "Node r xs"] by simp
qed
qed
lemma normalize1_mdeg_eq':
"wf_dlverts t1
\<Longrightarrow> max_deg (normalize1 t1) = max_deg t1 \<or> (max_deg (normalize1 t1) = 0 \<and> max_deg t1 = 1)"
proof(induction t1 rule: normalize1.induct)
case ind: (1 r t e)
then have 0: "max_deg (Node r {|(t, e)|}) \<ge> 1"
using mdeg_ge_fcard[of "{|(t, e)|}"] by (simp add: fcard_single_1)
then consider "rank (rev (Dtree.root t)) < rank (rev r)"
| "\<not>rank (rev (Dtree.root t)) < rank (rev r)" "max_deg (normalize1 t) \<le> 1"
| "\<not>rank (rev (Dtree.root t)) < rank (rev r)" "max_deg (normalize1 t) > 1" by linarith
then show ?case
proof(cases)
case 1
then show ?thesis
using mdeg_singleton[of r t] mdeg_root[of "Dtree.root t" "sucs t"]
by (auto simp: fcard_single_1 simp del: max_deg.simps)
next
case 2
then have "max_deg (normalize1 (Node r {|(t, e)|})) = 1"
using mdeg_singleton[of r "normalize1 t"] by (auto simp: fcard_single_1)
moreover have "max_deg (Node r {|(t, e)|}) = 1 "
using mdeg_singleton[of r t] ind 2 by (auto simp: fcard_single_1)
ultimately show ?thesis by simp
next
case 3
then show ?thesis
using mdeg_singleton[of r t] mdeg_singleton[of r "normalize1 t"] ind
by (auto simp: fcard_single_1)
qed
next
case ind: (2 xs r)
consider "max_deg (Node r xs) \<le> 1"
| "max_deg (Node r xs) > 1" "max_deg (Node r xs) = fcard xs"
| "max_deg (Node r xs) > 1" "fcard xs < max_deg (Node r xs)"
using mdeg_ge_fcard[of xs] by fastforce
then show ?case
proof(cases)
case 1
then show ?thesis using normalize1_mdeg_le[of "Node r xs"] by (auto simp del: max_deg.simps)
next
case 2
have 0: "\<forall>(t, e)\<in>fset xs. dlverts t \<noteq> {}" using dlverts_nempty_if_wf ind.prems by auto
then have "max_deg (Node r xs) \<le> max_deg (normalize1 (Node r xs))"
using mdeg_ge_fcard[of "(\<lambda>(t, e). (normalize1 t, e)) |`| xs"] ind 2
by (simp add: fcard_normalize_img_if_disjoint_lverts)
then show ?thesis using normalize1_mdeg_le[of "Node r xs"] by simp
next
case 3
then obtain t e where t_def: "(t,e) \<in> fset xs" "max_deg (Node r xs) = max_deg t"
using mdeg_child_if_gt_fcard by fastforce
have "max_deg (normalize1 t) \<le> max_deg (Node r ((\<lambda>(t,e). (normalize1 t,e)) |`| xs))"
using mdeg_ge_child[of "normalize1 t" e "(\<lambda>(t,e). (normalize1 t,e)) |`| xs"] t_def(1)
by (force simp del: max_deg.simps)
then have "max_deg (Node r xs) \<le> max_deg (normalize1 (Node r xs))"
using ind 3(1) t_def by (fastforce simp del: max_deg.simps)
then show ?thesis using normalize1_mdeg_le[of "Node r xs"] by simp
qed
qed
lemma normalize1_dom_mdeg_gt1:
"\<lbrakk>is_subtree (Node r xs) (normalize1 t); t1 \<in> fst ` fset xs; max_deg (Node r xs) > 1\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using ranked_dtree_with_orig_axioms proof(induction t rule: normalize1.induct)
case (1 r1 t e)
then interpret R: ranked_dtree_with_orig "Node r1 {|(t,e)|}" by blast
have sub_t: "is_subtree t (Node r1 {|(t,e)|})" using subtree_if_child[of t "{|(t,e)|}"] by simp
show ?case
proof(cases "Node r xs = normalize1 (Node r1 {|(t,e)|})")
case eq: True
then have 0: "max_deg (Node r1 {|(t,e)|}) > 1"
by (metis normalize1_mdeg_le "1.prems"(3) less_le_trans)
then have max_t: "max_deg t > 1" by (metis dtree.exhaust_sel mdeg_child_sucs_eq_if_gt1)
then show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r1)")
case True
then have eq: "Node r xs = Node (r1@Dtree.root t) (sucs t)" using eq by simp
then have "t1 \<in> fst ` fset (sucs t)" using "1.prems"(2) by simp
then obtain v where "v \<in> set (Dtree.root t)" "v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using R.dom_mdeg_gt1[of "Dtree.root t" "sucs t"] sub_t max_t by auto
then show ?thesis using eq by auto
next
case False
obtain v where v_def: "v \<in> set r1" "v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t)"
using max_t R.dom_mdeg_gt1[of r1 "{|(t, e)|}"] 0 by auto
interpret T: ranked_dtree_with_orig t using R.ranked_dtree_orig_rec by simp
have eq: "Node r xs = Node r1 {|(normalize1 t, e)|}" using False eq by simp
then have "t1 = normalize1 t" using "1.prems"(2) by simp
moreover have "Dtree.root t \<noteq> []"
using empty_notin_wf_dlverts[OF T.wf_lverts] dtree.set_sel(1)[of t] by auto
ultimately have "hd (Dtree.root t1) = hd (Dtree.root t)" using normalize1_hd_root_eq by blast
then show ?thesis using v_def eq by auto
qed
next
case uneq: False
show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r1)")
case True
then have "normalize1 (Node r1 {|(t,e)|}) = Node (r1@Dtree.root t) (sucs t)" by simp
then obtain t2 where t2_def: "t2 \<in> fst ` fset (sucs t)" "is_subtree (Node r xs) t2"
using uneq "1.prems"(1) by fastforce
then have "is_subtree t2 t" using subtree_if_suc by blast
then have "is_subtree (Node r xs) (Node r1 {|(t,e)|})"
using subtree_trans subtree_if_suc t2_def(2) by auto
then show ?thesis using R.dom_mdeg_gt1 "1.prems" by blast
next
case False
then have "normalize1 (Node r1 {|(t,e)|}) = Node r1 {|(normalize1 t, e)|}" by simp
then have "is_subtree (Node r xs) (normalize1 t)" using uneq "1.prems"(1) by auto
then show ?thesis using "1.IH" False "1.prems"(2,3) R.ranked_dtree_orig_rec by simp
qed
qed
next
case (2 xs1 r1)
then interpret R: ranked_dtree_with_orig "Node r1 xs1" by blast
show ?case
proof(cases "Node r xs = normalize1 (Node r1 xs1)")
case True
then have 0: "max_deg (Node r1 xs1) > 1"
using normalize1_mdeg_le "2.prems"(3) less_le_trans by (fastforce simp del: max_deg.simps)
then obtain t where t_def: "t \<in> fst ` fset xs1" "normalize1 t = t1"
using "2.prems"(2) "2.hyps" True by fastforce
then have sub_t: "is_subtree t (Node r1 xs1)" using subtree_if_child by fast
then obtain v where v_def: "v \<in> set r1" "v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t)"
using R.dom_mdeg_gt1[of r1] t_def(1) 0 by auto
interpret T: ranked_dtree_with_orig t using R.ranked_dtree_orig_rec t_def(1) by force
have "Dtree.root t \<noteq> []"
using empty_notin_wf_dlverts[OF T.wf_lverts] dtree.set_sel(1)[of t] by auto
then have "hd (Dtree.root t1) = hd (Dtree.root t)" using normalize1_hd_root_eq t_def(2) by blast
then show ?thesis using v_def "2.hyps" True by auto
next
case False
then show ?thesis using 2 R.ranked_dtree_orig_rec by auto
qed
qed
lemma child_contr_if_new_contr:
assumes "\<not>rank (rev (Dtree.root t1)) < rank (rev r)"
and "rank (rev (Dtree.root (normalize1 t1))) < rank (rev r)"
shows "\<exists>t2 e2. sucs t1 = {|(t2,e2)|} \<and> rank (rev (Dtree.root t2)) < rank (rev (Dtree.root t1))"
proof -
obtain t2 e2 where t2_def: "sucs t1 = {|(t2,e2)|}"
using root_normalize1_eq2[of "sucs t1" "Dtree.root t1"] assms by fastforce
then show ?thesis
using root_normalize1_eq1[of t2 "Dtree.root t1" e2] assms dtree.collapse[of t1] by fastforce
qed
lemma sub_contr_if_new_contr:
assumes "\<not>rank (rev (Dtree.root t1)) < rank (rev r)"
and "rank (rev (Dtree.root (normalize1 t1))) < rank (rev r)"
shows "\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) t1 \<and> rank (rev (Dtree.root t2)) < rank (rev v)"
proof -
obtain t2 e2 where t2_def: "sucs t1 = {|(t2,e2)|}" "rank (rev (Dtree.root t2)) < rank (rev (Dtree.root t1))"
using child_contr_if_new_contr[OF assms] by blast
then have "is_subtree (Node (Dtree.root t1) {|(t2,e2)|}) t1"
using is_subtree.simps[of "Node (Dtree.root t1) {|(t2,e2)|}" "Dtree.root t1" "sucs t1"] by fastforce
then show ?thesis using t2_def(2) by blast
qed
lemma normalize1_subtree_same_hd:
"\<lbrakk>is_subtree (Node v {|(t1,e1)|}) (normalize1 t)\<rbrakk>
\<Longrightarrow> \<exists>t3 e3. (is_subtree (Node v {|(t3,e3)|}) t \<and> hd (Dtree.root t1) = hd (Dtree.root t3))
\<or> (\<exists>v2. v = v2 @ Dtree.root t3 \<and> sucs t3 = {|(t1,e1)|}
\<and> is_subtree (Node v2 {|(t3,e3)|}) t \<and> rank (rev (Dtree.root t3)) < rank (rev v2))"
using wf_lverts wf_arcs proof(induction t rule: normalize1.induct)
case (1 r t e)
show ?case
proof(cases "Node v {|(t1,e1)|} = normalize1 (Node r {|(t,e)|})")
case eq: True
then show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
then show ?thesis using 1 eq by auto
next
case False
then have eq: "Node v {|(t1,e1)|} = Node r {|(normalize1 t,e)|}" using eq by simp
then show ?thesis using normalize1_hd_root_eq' "1.prems"(2) by auto
qed
next
case uneq: False
then show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
then obtain t2 e2 where "(t2,e2) \<in> fset (sucs t)" "is_subtree (Node v {|(t1,e1)|}) t2"
using "1.prems"(1) uneq by auto
then show ?thesis using is_subtree.simps[of "Node v {|(t1,e1)|}" "Dtree.root t" "sucs t"] by auto
next
case False
then have "is_subtree (Node v {|(t1,e1)|}) (normalize1 t)" using "1.prems"(1) uneq by auto
then show ?thesis
using "1.IH" "1.prems"(2,3) False by (auto simp: wf_darcs_iff_darcs')
qed
qed
next
case (2 xs r)
then have "\<forall>x. ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) \<noteq> {|x|}"
using singleton_normalize1 by (simp add: wf_darcs_iff_darcs')
then have "Node v {|(t1,e1)|} \<noteq> Node r ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)" by auto
then obtain t2 e2 where "(t2,e2) \<in> fset xs \<and> is_subtree (Node v {|(t1,e1)|}) (normalize1 t2)"
using "2.prems"(1) "2.hyps" by auto
then show ?case using "2.IH" "2.prems"(2,3) by (fastforce simp: wf_darcs_iff_darcs')
qed
lemma normalize1_dom_sub_contr:
"\<lbrakk>is_subtree (Node r xs) (normalize1 t); t1 \<in> fst ` fset xs;
\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) (Node r xs) \<and> rank (rev (Dtree.root t2)) < rank (rev v)\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using ranked_dtree_with_orig_axioms proof(induction t rule: normalize1.induct)
case (1 r1 t e)
then interpret R: ranked_dtree_with_orig "Node r1 {|(t,e)|}" by blast
interpret T: ranked_dtree_with_orig t using R.ranked_dtree_orig_rec by simp
have sub_t: "is_subtree (Node (Dtree.root t) (sucs t)) (Node r1 {|(t,e)|})"
using subtree_if_child[of t "{|(t,e)|}"] by simp
obtain v t2 e2 where v_def:
"is_subtree (Node v {|(t2,e2)|}) (Node r xs)" "rank (rev (Dtree.root t2)) < rank (rev v)"
using "1.prems"(3) by blast
show ?case
proof(cases "Node r xs = normalize1 (Node r1 {|(t,e)|})")
case eq: True
then show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r1)")
case True
then have eq: "Node r xs = Node (r1@Dtree.root t) (sucs t)" using eq by simp
then consider "Node r xs = Node v {|(t2,e2)|}" "max_deg (Node r xs) \<le> 1"
| "Node r xs \<noteq> Node v {|(t2,e2)|}" | "max_deg (Node r xs) > 1"
by linarith
then show ?thesis
proof(cases)
case 1
then have "max_deg (Node (r1@Dtree.root t) (sucs t)) \<le> 1" using eq by blast
then have "max_deg t \<le> 1" using mdeg_root[of "Dtree.root t" "sucs t"] by simp
then have "max_deg (Node r1 {|(t,e)|}) = 1"
using mdeg_singleton[of r1 t] by (simp add: fcard_single_1)
then have dom: "dom_children (Node r1 {|(t, e)|}) T" using R.dom_contr True by auto
have 0: "t1 \<in> fst ` fset (sucs t)" using eq "1.prems"(2) by blast
then have "Dtree.root t1 \<in> dverts t"
using dtree.set_sel(1) T.dverts_child_subset dtree.exhaust_sel psubsetD by metis
then obtain r2 where r2_def:
"r2 \<in> set r1 \<union> path_lverts t (hd (Dtree.root t1))" "r2 \<rightarrow>\<^bsub>T\<^esub> (hd (Dtree.root t1))"
using dom unfolding dom_children_def by auto
have "Dtree.root t1 \<noteq> []"
using empty_notin_wf_dlverts T.wf_lverts 0 T.dverts_child_subset
by (metis dtree.exhaust_sel dtree.set_sel(1) psubsetD)
then have "r2 \<in> set r1 \<union> set (Dtree.root t)"
using path_lverts_subset_root_if_childhd[OF 0] r2_def(1) by fast
then show ?thesis using r2_def(2) eq by auto
next
case 2
then obtain t3 e3 where t3_def:
"(t3,e3) \<in> fset (sucs t)" "is_subtree (Node v {|(t2,e2)|}) t3"
using eq v_def(1) by auto
have "is_subtree t3 t" using t3_def(1) subtree_if_suc by fastforce
then have "is_subtree (Node v {|(t2,e2)|}) (Node (Dtree.root t) (sucs t))"
using t3_def(2) subtree_trans by auto
moreover have "t1 \<in> fst ` fset (sucs t)" using eq "1.prems"(2) by blast
ultimately obtain v where v_def: "v \<in> set (Dtree.root t) \<and> v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using R.dom_sub_contr[OF sub_t] v_def(2) eq by blast
then show ?thesis using eq by auto
next
case 3
then show ?thesis using R.normalize1_dom_mdeg_gt1 "1.prems"(1,2) by blast
qed
next
case False
then have eq: "Node r xs = Node r1 {|(normalize1 t, e)|}" using eq by simp
have hd: "hd (Dtree.root (normalize1 t)) = hd (Dtree.root t)"
using normalize1_hd_root_eq' T.wf_lverts by blast
have "\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) t \<and> rank (rev (Dtree.root t2)) < rank (rev v)"
using contr_before_normalize1 eq v_def sub_contr_if_new_contr False by auto
then show ?thesis using R.dom_sub_contr[of r1 "{|(t,e)|}"] eq "1.prems"(2) hd by auto
qed
next
case uneq: False
show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r1)")
case True
then have "normalize1 (Node r1 {|(t,e)|}) = Node (r1@Dtree.root t) (sucs t)" by simp
then obtain t2 where t2_def: "t2 \<in> fst ` fset (sucs t)" "is_subtree (Node r xs) t2"
using uneq "1.prems"(1) by fastforce
then have "is_subtree t2 t" using subtree_if_suc by blast
then have "is_subtree (Node r xs) (Node r1 {|(t,e)|})"
using subtree_trans subtree_if_child t2_def(2) by auto
then show ?thesis using R.dom_sub_contr "1.prems"(2,3) by fast
next
case False
then have "normalize1 (Node r1 {|(t,e)|}) = Node r1 {|(normalize1 t, e)|}" by simp
then have "is_subtree (Node r xs) (normalize1 t)" using uneq "1.prems"(1) by auto
then show ?thesis using "1.IH" False "1.prems"(2,3) R.ranked_dtree_orig_rec by simp
qed
qed
next
case (2 xs1 r1)
then interpret R: ranked_dtree_with_orig "Node r1 xs1" by blast
show ?case
proof(cases "Node r xs = normalize1 (Node r1 xs1)")
case True
then have eq: "Node r xs = Node r1 ((\<lambda>(t,e). (normalize1 t,e)) |`| xs1)" using "2.hyps" by simp
obtain v t2 e2 where v_def:
"is_subtree (Node v {|(t2,e2)|}) (Node r xs)" "rank (rev (Dtree.root t2)) < rank (rev v)"
using "2.prems"(3) by blast
obtain t where t_def: "t \<in> fst ` fset xs1" "normalize1 t = t1" using "2.prems"(2) eq by force
then interpret T: ranked_dtree_with_orig t using R.ranked_dtree_orig_rec by force
have "\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) (Node r1 xs1)
\<and> rank (rev (Dtree.root t2)) < rank (rev v)"
using True contr_before_normalize1 v_def by presburger
moreover have "hd (Dtree.root t1) = hd (Dtree.root t)"
using normalize1_hd_root_eq' T.wf_lverts t_def(2) by blast
ultimately show ?thesis using R.dom_sub_contr[of r1 xs1] t_def(1) eq by auto
next
case False
then obtain t e where "(t,e) \<in> fset xs1 \<and> is_subtree (Node r xs) (normalize1 t)"
using "2.prems"(1) "2.hyps" by auto
then show ?thesis using "2.IH" "2.prems"(2,3) R.ranked_dtree_orig_rec by fast
qed
qed
lemma dom_children_combine_aux:
assumes "dom_children (Node r {|(t1, e1)|}) T"
and "t2 \<in> fst ` fset (sucs t1)"
and "x \<in> dverts t2"
shows "\<exists>v \<in> set (r @ Dtree.root t1) \<union> path_lverts t2 (hd x). v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
using path_lverts_child_union_root_sub[OF assms(2)] assms dtree.set_sel(2)
unfolding dom_children_def by fastforce
lemma dom_children_combine:
"dom_children (Node r {|(t1, e1)|}) T \<Longrightarrow> dom_children (Node (r@Dtree.root t1) (sucs t1)) T"
using dom_children_combine_aux by (simp add: dom_children_def)
lemma path_lverts_normalize1_sub:
"\<lbrakk>wf_dlverts t1; x \<in> dverts (normalize1 t1); max_deg (normalize1 t1) \<le> 1\<rbrakk>
\<Longrightarrow> path_lverts t1 (hd x) \<subseteq> path_lverts (normalize1 t1) (hd x)"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
then have eq: "normalize1 (Node r {|(t, e)|}) = Node (r@Dtree.root t) (sucs t)" by simp
then show ?thesis
proof(cases "x = r@Dtree.root t")
case True
then show ?thesis using 1 by auto
next
case False
then obtain t1 e1 where t1_def: "(t1,e1) \<in> fset (sucs t)" "x \<in> dverts t1"
using "1.prems"(2) eq by auto
then have 0: "hd x \<in> dlverts t1"
using hd_in_lverts_if_wf "1.prems"(1) wf_dlverts_sucs by force
then have "hd x \<in> dlverts t" using t1_def(1) suc_in_dlverts by fast
then have 2: "hd x \<notin> set r" using "1.prems"(1) by auto
have "wf_dlverts t" using "1.prems"(1) by simp
then have "hd x \<notin> set (Dtree.root t)" using 0 t1_def(1) wf_dlverts.simps[of "Dtree.root t"] by fastforce
then have hd_nin: "hd x \<notin> set (r @ Dtree.root t)" using 2 by auto
then obtain t2 e2 where "sucs t = {|(t2,e2)|}"
using "1.prems"(3) \<open>hd x \<in> dlverts t\<close> \<open>hd x \<notin> set (Dtree.root t)\<close> mdeg_root eq
by (metis dtree.collapse denormalize.simps(2) denormalize_set_eq_dlverts surj_pair)
then show ?thesis using eq hd_nin path_lverts_simps1_sucs by fastforce
qed
next
case uneq: False
then have "normalize1 (Node r {|(t, e)|}) = Node r {|(normalize1 t, e)|}" by simp
then have "max_deg (normalize1 t) \<le> 1"
using "1.prems"(3) mdeg_singleton[of r "normalize1 t"] fcard_single_1 max_def by auto
then show ?thesis using uneq 1 by auto
qed
next
case (2 xs r)
then have "max_deg (normalize1 (Node r xs)) = max_deg (Node r xs) \<or> max_deg (Node r xs) = 1"
using normalize1_mdeg_eq' by blast
then have "max_deg (Node r xs) \<le> 1" using "2.prems"(3) by (auto simp del: max_deg.simps)
then have "fcard xs = 0"
using mdeg_ge_fcard[of xs r] fcard_single_1_iff[of xs] "2.hyps" by fastforce
then show ?case using 2 by simp
qed
lemma dom_children_normalize1_aux_1:
assumes "dom_children (Node r {|(t1, e1)|}) T"
and "sucs t1 = {|(t2,e2)|}"
and "wf_dlverts t1"
and "normalize1 t1 = Node (Dtree.root t1 @ Dtree.root t2) (sucs t2)"
and "max_deg t1 = 1"
and "x \<in> dverts (normalize1 t1)"
shows "\<exists>v \<in> set r \<union> path_lverts (normalize1 t1) (hd x). v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
proof(cases "x = Dtree.root t1 @ Dtree.root t2")
case True
then have 0: "hd x = hd (Dtree.root t1)" using assms(3,4) normalize1_hd_root_eq' by fastforce
then obtain v where v_def: "v \<in> set r \<union> path_lverts t1 (hd x)" "v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
using assms(1) dtree.set_sel(1) unfolding dom_children_def by auto
have "Dtree.root t1 \<noteq> []" using assms(3) wf_dlverts.simps[of "Dtree.root t1" "sucs t1"] by simp
then show ?thesis using v_def 0 path_lverts_empty_if_roothd by auto
next
case False
then obtain t3 e3 where t3_def: "(t3,e3) \<in> fset (sucs t2)" "x \<in> dverts t3"
using assms(2,4,6) by auto
then have "x \<in> dverts t2" using dtree.set(1)[of "Dtree.root t2" "sucs t2"] by fastforce
then have "x \<in> dverts (Node (Dtree.root t1) {|(t2,e2)|})" by auto
then have "x \<in> dverts t1" using assms(2) dtree.exhaust_sel by metis
then obtain v where v_def: "v \<in> set r \<union> path_lverts t1 (hd x)" "v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
using assms(1) dtree.set_sel(1) unfolding dom_children_def by auto
have "path_lverts t1 (hd x) \<subseteq> path_lverts (Node (Dtree.root t1 @ Dtree.root t2) (sucs t2)) (hd x)"
using assms(3-6) normalize1_mdeg_le path_lverts_normalize1_sub by metis
then show ?thesis using v_def assms(4) by auto
qed
lemma dom_children_normalize1_1:
"\<lbrakk>dom_children (Node r {|(t1, e1)|}) T; sucs t1 = {|(t2,e2)|}; wf_dlverts t1;
normalize1 t1 = Node (Dtree.root t1 @ Dtree.root t2) (sucs t2); max_deg t1 = 1\<rbrakk>
\<Longrightarrow> dom_children (Node r {|(normalize1 t1, e1)|}) T"
using dom_children_normalize1_aux_1 by (simp add: dom_children_def)
lemma dom_children_normalize1_aux:
assumes "\<forall>x\<in>dverts t1. \<exists>v \<in> set r0 \<union> path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
and "wf_dlverts t1"
and "max_deg t1 \<le> 1"
and "x \<in> dverts (normalize1 t1)"
shows "\<exists>v \<in> set r0 \<union> path_lverts (normalize1 t1) (hd x). v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
using assms proof(induction t1 arbitrary: r0 rule: normalize1.induct)
case (1 r t e)
have deg1: "max_deg (Node r {|(t, e)|}) = 1"
using "1.prems"(3) mdeg_ge_fcard[of "{|(t, e)|}"] by (simp add: fcard_single_1)
then show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
have 0: "dom_children (Node r0 {|(Node r {|(t, e)|}, e)|}) T"
using "1.prems"(1) unfolding dom_children_def by simp
show ?thesis using dom_children_normalize1_aux_1[OF 0] "1.prems"(1,2,4) deg1 True by auto
next
case ncontr: False
show ?thesis
proof(cases "x = r")
case True
then show ?thesis using "1.prems"(1,2) by auto
next
case False
have "wf_dlverts (normalize1 t)" using "1.prems"(2) wf_dlverts_normalize1 by auto
then have "hd x \<in> dlverts (normalize1 t)"
using hd_in_lverts_if_wf False ncontr "1.prems"(1,4) by fastforce
then have hd: "hd x \<notin> set r" using "1.prems"(2) ncontr wf_dlverts_normalize1 by fastforce
then have eq: "path_lverts (Node r {|(t, e)|}) (hd x) = set r \<union> path_lverts t (hd x)" by simp
then have eq1: "path_lverts (Node r {|(normalize1 t, e)|}) (hd x)
= set r \<union> path_lverts (normalize1 t) (hd x)" by auto
have "\<forall>x\<in>dverts t. path_lverts (Node r {|(t, e)|}) (hd x) \<subseteq> set r \<union> path_lverts t (hd x)"
using path_lverts_child_union_root_sub by simp
then have 2: "\<forall>x\<in>dverts t. \<exists>v\<in>set (r0@r) \<union> path_lverts t (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
using "1.prems"(1) by fastforce
have "max_deg t \<le> 1" using "1.prems"(3) mdeg_ge_child[of t e "{|(t, e)|}"] by simp
then show ?thesis using "1.IH"[OF ncontr 2] "1.prems"(2,4) ncontr hd by auto
qed
qed
next
case (2 xs r)
then have "fcard xs \<le> 1" using mdeg_ge_fcard[of xs] by simp
then have "fcard xs = 0" using "2.hyps" fcard_single_1_iff[of xs] by fastforce
then show ?case using 2 by auto
qed
lemma dom_children_normalize1:
"\<lbrakk>dom_children (Node r0 {|(t1,e1)|}) T; wf_dlverts t1; max_deg t1 \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r0 {|(normalize1 t1,e1)|}) T"
using dom_children_normalize1_aux by (simp add: dom_children_def)
lemma dom_children_child_self_aux:
assumes "dom_children t1 T"
and "sucs t1 = {|(t2, e2)|}"
and "rank (rev (Dtree.root t2)) < rank (rev (Dtree.root t1))"
and "t = Node r {|(t1, e1)|}"
and "x \<in> dverts t1"
shows "\<exists>v \<in> set r \<union> path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
proof(cases "x = Dtree.root t1")
case True
have "is_subtree (Node (Dtree.root t1) {|(t2, e2)|}) (Node r {|(t1, e1)|})"
using subtree_if_child[of "t1" "{|(t1, e1)|}"] assms(2) dtree.collapse[of t1] by simp
then show ?thesis using dom_sub_contr[of r "{|(t1, e1)|}"] assms(3,4) True by auto
next
case False
then have "x \<in> (\<Union>y\<in>fset (sucs t1). \<Union> (dverts ` Basic_BNFs.fsts y))"
using assms(5) dtree.set(1)[of "Dtree.root t1" "sucs t1"] by auto
then have "x \<in> dverts t2" using assms(2) by auto
then obtain v where v_def: "v \<in> set (Dtree.root t1) \<union> path_lverts t2 (hd x)" "v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
using assms(1,2) dtree.set_sel(1) unfolding dom_children_def by auto
interpret T1: list_dtree t1 using list_dtree_rec assms(4) by simp
interpret T2: list_dtree t2 using T1.list_dtree_rec_suc assms(2) by simp
have "hd x \<in> dlverts t2" using \<open>x \<in> dverts t2\<close> by (simp add: hd_in_lverts_if_wf T2.wf_lverts)
then have "hd x \<notin> set (Dtree.root t1)"
using T1.wf_lverts wf_dlverts.simps[of "Dtree.root t1" "sucs t1"] assms(2) by fastforce
then have "path_lverts t1 (hd x) = set (Dtree.root t1) \<union> path_lverts t2 (hd x)"
using assms(2) by (simp add: path_lverts_simps1_sucs)
then show ?thesis using v_def by auto
qed
lemma dom_children_child_self:
assumes "dom_children t1 T"
and "sucs t1 = {|(t2, e2)|}"
and "rank (rev (Dtree.root t2)) < rank (rev (Dtree.root t1))"
and "t = Node r {|(t1, e1)|}"
shows "dom_children (Node r {|(t1, e1)|}) T"
using dom_children_child_self_aux[OF assms] by (simp add: dom_children_def)
lemma normalize1_dom_contr:
"\<lbrakk>is_subtree (Node r {|(t1,e1)|}) (normalize1 t); rank (rev (Dtree.root t1)) < rank (rev r);
max_deg (Node r {|(t1,e1)|}) = 1\<rbrakk>
\<Longrightarrow> dom_children (Node r {|(t1,e1)|}) T"
using ranked_dtree_with_orig_axioms proof(induction t rule: normalize1.induct)
case (1 r1 t e)
then interpret R: ranked_dtree_with_orig "Node r1 {|(t,e)|}" by blast
interpret T: ranked_dtree_with_orig t using R.ranked_dtree_orig_rec by simp
have sub_t: "is_subtree (Node (Dtree.root t) (sucs t)) (Node r1 {|(t,e)|})"
using subtree_if_child[of t "{|(t,e)|}"] by simp
show ?case
proof(cases "Node r {|(t1,e1)|} = normalize1 (Node r1 {|(t,e)|})")
case eq: True
then show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r1)")
case True
then have eq: "Node r {|(t1,e1)|} = Node (r1@Dtree.root t) (sucs t)" using eq by simp
then have "max_deg t = 1" using mdeg_root[of "Dtree.root t" "sucs t"] 1 by simp
then have "max_deg (Node r1 {|(t,e)|}) = 1"
using mdeg_singleton[of r1 t] by (simp add: fcard_single_1)
then have "dom_children (Node r1 {|(t, e)|}) T" using R.dom_contr[of r1 t e] True by simp
then show ?thesis using dom_children_combine eq by simp
next
case False
then have eq: "Node r {|(t1,e1)|} = Node r1 {|(normalize1 t, e)|}" using eq by simp
then obtain t2 e2 where t2_def:
"sucs t = {|(t2, e2)|}" "rank (rev (Dtree.root t2)) < rank (rev (Dtree.root t))"
using child_contr_if_new_contr False "1.prems"(2) by blast
then have "is_subtree (Node (Dtree.root t) {|(t2, e2)|}) (Node r1 {|(t, e)|})" using sub_t by simp
have "max_deg t = 1"
using "1.prems"(3) eq mdeg_singleton mdeg_root t2_def
by (metis dtree.collapse fcard_single_1 normalize1.simps(1))
then have "max_deg (Node (Dtree.root t) {|(t2, e2)|}) = 1"
using t2_def(1) dtree.collapse[of t] by simp
then have "dom_children (Node (Dtree.root t) (sucs t)) T"
using R.dom_contr sub_t t2_def "1.prems"(3) by simp
then have "dom_children t T" using dtree.exhaust_sel by simp
then have "dom_children (Node r1 {|(t,e)|}) T"
using R.dom_children_child_self t2_def by simp
then show ?thesis using dom_children_normalize1 \<open>max_deg t = 1\<close> T.wf_lverts eq by auto
qed
next
case uneq: False
show ?thesis
proof(cases "rank (rev (Dtree.root t)) < rank (rev r1)")
case True
then have "normalize1 (Node r1 {|(t,e)|}) = Node (r1@Dtree.root t) (sucs t)" by simp
then obtain t2 where t2_def: "t2 \<in> fst ` fset (sucs t)" "is_subtree (Node r {|(t1,e1)|}) t2"
using uneq "1.prems"(1) by fastforce
then have "is_subtree t2 t" using subtree_if_suc by blast
then have "is_subtree (Node r {|(t1,e1)|}) (Node r1 {|(t,e)|})"
using subtree_trans subtree_if_child t2_def(2) by auto
then show ?thesis using R.dom_contr "1.prems"(2,3) by blast
next
case False
then have "normalize1 (Node r1 {|(t,e)|}) = Node r1 {|(normalize1 t, e)|}" by simp
then have "is_subtree (Node r {|(t1,e1)|}) (normalize1 t)" using uneq "1.prems"(1) by auto
then show ?thesis using "1.IH" False "1.prems"(2,3) R.ranked_dtree_orig_rec by simp
qed
qed
next
case (2 xs r1)
then have eq: "normalize1 (Node r1 xs) = Node r1 ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)"
using "2.hyps" by simp
interpret R: ranked_dtree_with_orig "Node r1 xs" using "2.prems"(4) by blast
have "\<forall>x. ((\<lambda>(t,e). (normalize1 t,e)) |`| xs) \<noteq> {|x|}"
using singleton_normalize1 "2.hyps" disjoint_darcs_if_wf_xs[OF R.wf_arcs] by auto
then have "Node r {|(t1,e1)|} \<noteq> Node r1 ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)" by auto
then obtain t3 e3 where t3_def:
"(t3,e3) \<in> fset xs" "is_subtree (Node r {|(t1, e1)|}) (normalize1 t3)"
using "2.prems"(1) eq by auto
then show ?case using "2.IH" "2.prems"(2,3) R.ranked_dtree_orig_rec by simp
qed
lemma dom_children_normalize1_img_full:
assumes "dom_children (Node r xs) T"
and "\<forall>(t1,e1) \<in> fset xs. wf_dlverts t1"
and "\<forall>(t1,e1) \<in> fset xs. max_deg t1 \<le> 1"
shows "dom_children (Node r ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs)) T"
proof -
have "\<forall>(t1, e1) \<in> fset xs. dom_children (Node r {|(t1, e1)|}) T"
using dom_children_all_singletons[OF assms(1)] by blast
then have "\<forall>(t1, e1) \<in> fset xs. dom_children (Node r {|(normalize1 t1, e1)|}) T"
using dom_children_normalize1 assms(2,3) by fast
then show ?thesis
using dom_children_if_all_singletons[of "(\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs"] by fastforce
qed
lemma children_deg1_normalize1_sub:
"(\<lambda>(t1,e1). (normalize1 t1,e1)) ` children_deg1 xs
\<subseteq> children_deg1 ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs)"
using normalize1_mdeg_le order_trans by auto
lemma normalize1_children_deg1_sub_if_wfarcs:
"\<forall>(t1,e1)\<in>fset xs. wf_darcs t1
\<Longrightarrow> children_deg1 ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs)
\<subseteq> (\<lambda>(t1,e1). (normalize1 t1,e1)) ` children_deg1 xs"
using normalize1_mdeg_eq by fastforce
lemma normalize1_children_deg1_eq_if_wfarcs:
"\<forall>(t1,e1)\<in>fset xs. wf_darcs t1
\<Longrightarrow> (\<lambda>(t1,e1). (normalize1 t1,e1)) ` children_deg1 xs
= children_deg1 ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs)"
- using children_deg1_normalize1_sub normalize1_children_deg1_sub_if_wfarcs by fast
+ using children_deg1_normalize1_sub normalize1_children_deg1_sub_if_wfarcs
+ by (meson subset_antisym)
lemma normalize1_children_deg1_sub_if_wflverts:
"\<forall>(t1,e1)\<in>fset xs. wf_dlverts t1
\<Longrightarrow> children_deg1 ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs)
\<subseteq> (\<lambda>(t1,e1). (normalize1 t1,e1)) ` children_deg1 xs"
using normalize1_mdeg_eq' by fastforce
lemma normalize1_children_deg1_eq_if_wflverts:
"\<forall>(t1,e1)\<in>fset xs. wf_dlverts t1
\<Longrightarrow> (\<lambda>(t1,e1). (normalize1 t1,e1)) ` children_deg1 xs
= children_deg1 ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs)"
- using children_deg1_normalize1_sub normalize1_children_deg1_sub_if_wflverts by fast
+ using children_deg1_normalize1_sub normalize1_children_deg1_sub_if_wflverts
+ by (meson subset_antisym)
lemma dom_children_normalize1_img:
assumes "dom_children (Node r (Abs_fset (children_deg1 xs))) T"
and "\<forall>(t1,e1) \<in> fset xs. wf_dlverts t1"
shows "dom_children (Node r (Abs_fset (children_deg1 ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs)))) T"
proof -
have "\<forall>(t1, e1) \<in> children_deg1 xs. dom_children (Node r {|(t1, e1)|}) T"
using dom_children_all_singletons[OF assms(1)] children_deg1_fset_id by blast
then have "\<forall>(t2, e2) \<in> (\<lambda>(t1,e1). (normalize1 t1,e1)) ` children_deg1 xs.
dom_children (Node r {|(t2, e2)|}) T"
using dom_children_normalize1 assms(2) by fast
then have "\<forall>(t2, e2) \<in> children_deg1 ((\<lambda>(t1,e1). (normalize1 t1,e1)) |`| xs).
dom_children (Node r {|(t2, e2)|}) T"
using normalize1_children_deg1_eq_if_wflverts[of xs] assms(2) by blast
- then show ?thesis using dom_children_if_all_singletons children_deg1_fset_id by fast
+ then show ?thesis using dom_children_if_all_singletons children_deg1_fset_id
+ proof -
+ have "\<forall>f as p. \<exists>pa. (dom_children (Node (as::'a list) f) p \<or> pa |\<in>| f) \<and> (\<not> (case pa of (d, b::'b) \<Rightarrow> dom_children (Node as {|(d, b)|}) p) \<or> dom_children (Node as f) p)"
+ using dom_children_if_all_singletons by blast
+ then obtain pp :: "(('a list, 'b) Dtree.dtree \<times> 'b) fset \<Rightarrow> 'a list \<Rightarrow> ('a, 'b) pre_digraph \<Rightarrow> ('a list, 'b) Dtree.dtree \<times> 'b" where
+ f1: "\<And>as f p. (dom_children (Node as f) p \<or> pp f as p |\<in>| f) \<and> (\<not> (case pp f as p of (d, b) \<Rightarrow> dom_children (Node as {|(d, b)|}) p) \<or> dom_children (Node as f) p)"
+ by metis
+ moreover
+ { assume "\<not> (case pp (Abs_fset (children_deg1 ((\<lambda>(d, y). (normalize1 d, y)) |`| xs))) r T of (d, b) \<Rightarrow> dom_children (Node r {|(d, b)|}) T)"
+ then have "pp (Abs_fset (children_deg1 ((\<lambda>(d, y). (normalize1 d, y)) |`| xs))) r T \<notin> children_deg1 ((\<lambda>(d, y). (normalize1 d, y)) |`| xs)"
+ by (smt (z3) \<open>\<forall>(t2, e2) \<in>children_deg1 ((\<lambda>(t1, e1). (normalize1 t1, e1)) |`| xs). dom_children (Node r {|(t2, e2)|}) T\<close>)
+ then have "pp (Abs_fset (children_deg1 ((\<lambda>(d, y). (normalize1 d, y)) |`| xs))) r T |\<notin>| Abs_fset (children_deg1 ((\<lambda>(d, y). (normalize1 d, y)) |`| xs))"
+ by (metis (no_types) children_deg1_fset_id)
+ then have ?thesis
+ using f1 by blast }
+ ultimately show ?thesis
+ by meson
+ qed
qed
lemma normalize1_dom_wedge:
"\<lbrakk>is_subtree (Node r xs) (normalize1 t); fcard xs > 1\<rbrakk>
\<Longrightarrow> dom_children (Node r (Abs_fset (children_deg1 xs))) T"
using ranked_dtree_with_orig_axioms proof(induction t rule: normalize1.induct)
case (1 r1 t e)
then interpret R: ranked_dtree_with_orig "Node r1 {|(t,e)|}" by blast
have sub_t: "is_subtree (Node (Dtree.root t) (sucs t)) (Node r1 {|(t,e)|})"
using subtree_if_child[of t "{|(t,e)|}"] by simp
show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r1)")
case True
then have eq: "normalize1 (Node r1 {|(t,e)|}) = Node (r1@Dtree.root t) (sucs t)" by simp
then show ?thesis
proof(cases "Node r xs = normalize1 (Node r1 {|(t,e)|})")
case True
then have "Node r xs = Node (r1@Dtree.root t) (sucs t)" using eq by simp
then show ?thesis using R.dom_wedge[OF sub_t] "1.prems"(2) unfolding dom_children_def by auto
next
case False
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset (sucs t)" "is_subtree (Node r xs) t2"
using "1.prems"(1) eq by auto
then have "is_subtree (Node r xs) t" using subtree_if_suc subtree_trans by fastforce
then show ?thesis using R.dom_wedge sub_t "1.prems"(2) by simp
qed
next
case False
then show ?thesis using 1 R.ranked_dtree_orig_rec by (auto simp: fcard_single_1)
qed
next
case (2 xs1 r1)
then have eq: "normalize1 (Node r1 xs1) = Node r1 ((\<lambda>(t,e). (normalize1 t,e)) |`| xs1)"
using "2.hyps" by simp
interpret R: ranked_dtree_with_orig "Node r1 xs1" using "2.prems"(3) by blast
have "\<forall>x. ((\<lambda>(t,e). (normalize1 t,e)) |`| xs1) \<noteq> {|x|}"
using singleton_normalize1 "2.hyps" disjoint_darcs_if_wf_xs[OF R.wf_arcs] by auto
then show ?case
proof(cases "Node r xs = normalize1 (Node r1 xs1)")
case True
then have "1 < fcard xs1" using eq "2.prems"(2) fcard_image_le less_le_trans by fastforce
then have "dom_children (Node r1 (Abs_fset (children_deg1 xs1))) T" using R.dom_wedge by simp
then show ?thesis using dom_children_normalize1_img eq R.wf_lverts True by fastforce
next
case False
then show ?thesis using 2 R.ranked_dtree_orig_rec by fastforce
qed
qed
corollary normalize1_dom_wedge':
"\<forall>r xs. is_subtree (Node r xs) (normalize1 t) \<longrightarrow> fcard xs > 1
\<longrightarrow> dom_children (Node r (Abs_fset {(t, e). (t, e) \<in> fset xs \<and> max_deg t \<le> Suc 0})) T"
by (auto simp only: normalize1_dom_wedge One_nat_def[symmetric])
lemma normalize1_verts_conform: "v \<in> dverts (normalize1 t) \<Longrightarrow> seq_conform v"
using ranked_dtree_with_orig_axioms proof(induction t rule: normalize1.induct)
case ind: (1 r t e)
then interpret R: ranked_dtree_with_orig "Node r {|(t, e)|}" by blast
consider "rank (rev (Dtree.root t)) < rank (rev r)" "v = r@Dtree.root t"
| "rank (rev (Dtree.root t)) < rank (rev r)" "v \<noteq> r@Dtree.root t"
| "\<not>rank (rev (Dtree.root t)) < rank (rev r)"
by blast
then show ?case
proof(cases)
case 1
then show ?thesis using R.contr_seq_conform by auto
next
case 2
then have "v \<in> dverts (Node r {|(t, e)|})" using dverts_suc_subseteq ind.prems by fastforce
then show ?thesis using R.verts_conform by blast
next
case 3
then show ?thesis using R.verts_conform ind R.ranked_dtree_orig_rec by auto
qed
next
case (2 xs r)
then interpret R: ranked_dtree_with_orig "Node r xs" by blast
show ?case using R.verts_conform 2 R.ranked_dtree_orig_rec by auto
qed
corollary normalize1_verts_distinct: "v \<in> dverts (normalize1 t) \<Longrightarrow> distinct v"
using distinct_normalize1 verts_distinct by auto
lemma dom_mdeg_le1_aux:
assumes "max_deg t \<le> 1"
and "is_subtree (Node v {|(t2, e2)|}) t"
and "rank (rev (Dtree.root t2)) < rank (rev v)"
and "t1 \<in> fst ` fset (sucs t)"
and "x \<in> dverts t1"
shows "\<exists>r\<in>set (Dtree.root t) \<union> path_lverts t1 (hd x). r \<rightarrow>\<^bsub>T\<^esub> hd x"
using assms ranked_dtree_with_orig_axioms proof(induction t arbitrary: t1)
case (Node r xs)
then interpret R: ranked_dtree_with_orig "Node r xs" by blast
interpret T1: ranked_dtree_with_orig t1 using Node.prems(4) R.ranked_dtree_orig_rec by force
have "fcard xs > 0" using Node.prems(4) fcard_seteq by fastforce
then have "fcard xs = 1" using mdeg_ge_fcard[of xs] Node.prems(1) by simp
then obtain e1 where e1_def: "xs = {|(t1,e1)|}"
using Node.prems(4) fcard_single_1_iff[of xs] by auto
have mdeg1: "max_deg (Node r xs) = 1"
using Node.prems(1) mdeg_ge_fcard[of xs] \<open>fcard xs = 1\<close> by simp
show ?case
proof(cases "Node v {|(t2, e2)|} = Node r xs")
case True
then have "dom_children (Node r xs) T"
using mdeg1 Node.prems(2,3) R.dom_contr_subtree by blast
then show ?thesis unfolding dom_children_def using e1_def Node.prems(5) by simp
next
case False
then have sub_t1: "is_subtree (Node v {|(t2, e2)|}) t1"
using Node.prems(2) e1_def is_subtree.simps[of "Node v {|(t2, e2)|}"] by force
show ?thesis
proof(cases "x = Dtree.root t1")
case True
then show ?thesis using R.dom_sub_contr[OF self_subtree] Node.prems(3) e1_def sub_t1 by auto
next
case False
then obtain t3 where t3_def: "t3 \<in> fst ` fset (sucs t1)" "x \<in> dverts t3"
using Node.prems(5) dverts_root_or_child[of x "Dtree.root t1" "sucs t1"] by fastforce
have mdeg_t1: "max_deg t1 \<le> 1" using mdeg_ge_child[of t1 e1 xs] e1_def mdeg1 by simp
moreover have "fcard (sucs t1) > 0" using t3_def fcard_seteq by fastforce
ultimately have "fcard (sucs t1) = 1" using mdeg_ge_fcard[of "sucs t1" "Dtree.root t1"] by simp
then obtain e3 where e3_def: "sucs t1 = {|(t3, e3)|}"
using t3_def fcard_single_1_iff[of "sucs t1"] by fastforce
have ind: "\<exists>r\<in>set (Dtree.root t1) \<union> path_lverts t3 (hd x). r \<rightarrow>\<^bsub>T\<^esub> hd x"
using Node.IH mdeg_t1 e1_def sub_t1 Node.prems(3) t3_def T1.ranked_dtree_with_orig_axioms
by auto
have "hd x \<in> dlverts t3" using t3_def hd_in_lverts_if_wf T1.wf_lverts wf_dlverts_suc by blast
then have "hd x \<notin> set (Dtree.root t1)"
using t3_def dlverts_notin_root_sucs[OF T1.wf_lverts] by blast
then have "path_lverts t1 (hd x) = set (Dtree.root t1) \<union> path_lverts t3 (hd x)"
using path_lverts_simps1_sucs e3_def by fastforce
then show ?thesis using ind by blast
qed
qed
qed
lemma dom_mdeg_le1:
assumes "max_deg t \<le> 1"
and "is_subtree (Node v {|(t2, e2)|}) t"
and "rank (rev (Dtree.root t2)) < rank (rev v)"
shows "dom_children t T"
using dom_mdeg_le1_aux[OF assms] unfolding dom_children_def by blast
lemma dom_children_normalize1_preserv:
assumes "max_deg (normalize1 t1) \<le> 1" and "dom_children t1 T" and "wf_dlverts t1"
shows "dom_children (normalize1 t1) T"
using assms proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
then show ?thesis using 1 dom_children_combine by force
next
case False
then have "max_deg (normalize1 t) \<le> 1"
using "1.prems"(1) mdeg_ge_child[of "normalize1 t" e "{|(normalize1 t,e)|}"] by simp
then have "max_deg t \<le> 1" using normalize1_mdeg_eq' "1.prems"(3) by fastforce
then show ?thesis using dom_children_normalize1 False "1.prems"(2,3) by simp
qed
next
case (2 xs r)
have "max_deg (Node r xs) \<le> 1"
using normalize1_mdeg_eq'[OF "2.prems"(3)] "2.prems"(1) by fastforce
then have "fcard xs \<le> 1" using mdeg_ge_fcard[of xs] by simp
then have "fcard xs = 0" using fcard_single_1_iff[of xs] "2.hyps" by fastforce
then have "normalize1 (Node r xs) = Node r xs" using "2.hyps" by simp
then show ?case using "2.prems"(2) by simp
qed
lemma dom_mdeg_le1_normalize1:
assumes "max_deg (normalize1 t) \<le> 1" and "normalize1 t \<noteq> t"
shows "dom_children (normalize1 t) T"
proof -
obtain v t2 e2 where "is_subtree (Node v {|(t2, e2)|}) t" "rank (rev (Dtree.root t2)) < rank (rev v)"
using contr_if_normalize1_uneq assms(2) by blast
moreover have "max_deg t \<le> 1" using assms(1) normalize1_mdeg_eq wf_arcs by fastforce
ultimately show ?thesis
using dom_mdeg_le1 dom_children_normalize1_preserv assms(1) wf_lverts by blast
qed
lemma normalize_mdeg_eq:
"wf_darcs t1
\<Longrightarrow> max_deg (normalize t1) = max_deg t1 \<or> (max_deg (normalize t1) = 0 \<and> max_deg t1 = 1)"
apply (induction t1 rule: normalize.induct)
by (smt (verit, ccfv_threshold) normalize1_mdeg_eq wf_darcs_normalize1 normalize.simps)
lemma normalize_mdeg_eq':
"wf_dlverts t1
\<Longrightarrow> max_deg (normalize t1) = max_deg t1 \<or> (max_deg (normalize t1) = 0 \<and> max_deg t1 = 1)"
apply (induction t1 rule: normalize.induct)
by (smt (verit, ccfv_threshold) normalize1_mdeg_eq' wf_dlverts_normalize1 normalize.simps)
corollary mdeg_le1_normalize:
"\<lbrakk>max_deg (normalize t1) \<le> 1; wf_dlverts t1\<rbrakk> \<Longrightarrow> max_deg t1 \<le> 1"
using normalize_mdeg_eq' by fastforce
lemma dom_children_normalize_preserv:
assumes "max_deg (normalize t1) \<le> 1" and "dom_children t1 T" and "wf_dlverts t1"
shows "dom_children (normalize t1) T"
using assms proof(induction t1 rule: normalize.induct)
case (1 t1)
then show ?case
proof(cases "t1 = normalize1 t1")
case True
then show ?thesis using "1.prems" dom_children_normalize1_preserv by simp
next
case False
have "max_deg t1 \<le> 1" using mdeg_le1_normalize "1.prems"(1,3) by blast
then have "max_deg (normalize1 t1) \<le> 1"
using normalize1_mdeg_eq' "1.prems"(3) by fastforce
then have "dom_children (normalize1 t1) T"
using dom_children_normalize1_preserv "1.prems"(2,3) by blast
then show ?thesis using 1 False by (simp add: Let_def wf_dlverts_normalize1)
qed
qed
lemma dom_mdeg_le1_normalize:
assumes "max_deg (normalize t) \<le> 1" and "normalize t \<noteq> t"
shows "dom_children (normalize t) T"
using assms ranked_dtree_with_orig_axioms proof(induction t rule: normalize.induct)
case (1 t)
then interpret T: ranked_dtree_with_orig t by blast
show ?case
using 1 T.dom_mdeg_le1_normalize1 T.wf_lverts wf_dlverts_normalize1
by (smt (verit) dom_children_normalize_preserv normalize.elims mdeg_le1_normalize)
qed
lemma normalize1_arc_in_dlverts:
"\<lbrakk>is_subtree (Node v ys) (normalize1 t); x \<in> set v; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts (Node v ys)"
using ranked_dtree_with_orig_axioms proof(induction t rule: normalize1.induct)
case ind: (1 r t e)
then interpret R: ranked_dtree_with_orig "Node r {|(t, e)|}" by blast
show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
then have eq: "normalize1 (Node r {|(t, e)|}) = Node (r@Dtree.root t) (sucs t)" by simp
then show ?thesis
proof(cases "Node v ys = Node (r@Dtree.root t) (sucs t)")
case True
then consider "x \<in> set r" | "x \<in> set (Dtree.root t)" using ind.prems(2) by auto
then show ?thesis
proof(cases)
case 1
then have "y \<in> dlverts (Node r {|(t, e)|})"
using R.arc_in_dlverts ind.prems(3) by fastforce
then show ?thesis using eq normalize1_dlverts_eq[of "Node r {|(t, e)|}"] True by simp
next
case 2
then have "y \<in> dlverts t"
using R.arc_in_dlverts[of "Dtree.root t" "sucs t"] ind.prems(3)
subtree_if_child[of t "{|(t, e)|}"] by simp
then show ?thesis using eq normalize1_dlverts_eq[of "Node r {|(t, e)|}"] True by simp
qed
next
case False
then obtain t2 where t2_def: "t2 \<in> fst ` fset (sucs t)" "is_subtree (Node v ys) t2"
using ind.prems(1) eq by force
then have "is_subtree (Node v ys) (Node r {|(t, e)|})"
using subtree_trans[OF t2_def(2)] subtree_if_suc by auto
then show ?thesis using R.arc_in_dlverts ind.prems(2,3) by blast
qed
next
case nocontr: False
then show ?thesis
proof(cases "Node v ys = Node r {|(normalize1 t, e)|}")
case True
then have "y \<in> dlverts (Node r {|(t, e)|})"
using R.arc_in_dlverts ind.prems(2,3) by fastforce
then show ?thesis using nocontr True by simp
next
case False
then have "is_subtree (Node v ys) (normalize1 t)" using ind.prems(1) nocontr by auto
then show ?thesis using ind.IH[OF nocontr] ind.prems(2,3) R.ranked_dtree_orig_rec by simp
qed
qed
next
case (2 xs r)
then interpret R: ranked_dtree_with_orig "Node r xs" by blast
have eq: "normalize1 (Node r xs) = Node r ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)"
using "2.hyps" by simp
show ?case
proof(cases "Node v ys = normalize1 (Node r xs)")
case True
then have "y \<in> dlverts (Node r xs)" using R.arc_in_dlverts "2.hyps" "2.prems"(2,3) by simp
then show ?thesis using True by simp
next
case False
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "is_subtree (Node v ys) (normalize1 t2)"
using "2.hyps" "2.prems"(1) by auto
then show ?thesis using "2.IH" "2.prems"(2,3) R.ranked_dtree_orig_rec by simp
qed
qed
lemma normalize1_arc_in_dlverts':
"\<forall>r xs. is_subtree (Node r xs) (normalize1 t) \<longrightarrow> (\<forall>x. x \<in> set r
\<longrightarrow> (\<forall>y. x \<rightarrow>\<^bsub>T\<^esub> y \<longrightarrow> y \<in> set r \<or> (\<exists>x\<in>fset xs. y \<in> dlverts (fst x))))"
using normalize1_arc_in_dlverts by simp
theorem ranked_dtree_orig_normalize1: "ranked_dtree_with_orig (normalize1 t) rank cost cmp T root"
by (simp add: ranked_dtree_with_orig_def ranked_dtree_with_orig_axioms_def asi_rank
normalize1_dom_contr normalize1_dom_mdeg_gt1 normalize1_dom_sub_contr
normalize1_dom_wedge' directed_tree_axioms normalize1_arc_in_dlverts'
ranked_dtree_normalize1 normalize1_verts_conform normalize1_verts_distinct)
theorem ranked_dtree_orig_normalize: "ranked_dtree_with_orig (normalize t) rank cost cmp T root"
using ranked_dtree_with_orig_axioms proof(induction t rule: normalize.induct)
case (1 t)
then interpret T: ranked_dtree_with_orig t by blast
show ?case using "1.IH" T.ranked_dtree_orig_normalize1 by(auto simp: Let_def)
qed
subsubsection \<open>Merging preserves Arc Invariants\<close>
interpretation Comm: comp_fun_commute "merge_f r xs" by (rule merge_commute)
lemma path_lverts_supset_z:
"\<lbrakk>list_dtree (Node r xs); \<forall>t1 \<in> fst ` fset xs. a \<notin> dlverts t1\<rbrakk>
\<Longrightarrow> path_lverts_list z a \<subseteq> path_lverts_list (ffold (merge_f r xs) z xs) a"
proof(induction xs)
case (insert x xs)
interpret Comm: comp_fun_commute "merge_f r (finsert x xs)" by (rule merge_commute)
define f where "f = merge_f r (finsert x xs)"
define f' where "f' = merge_f r xs"
let ?merge = "Sorting_Algorithms.merge cmp'"
have 0: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(1) by blast
show ?case
proof(cases "ffold f z (finsert x xs) = ffold f' z xs")
case True
then show ?thesis using insert.IH 0 insert.prems(2) f_def f'_def by auto
next
case False
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
have 1: "\<forall>v\<in>fst ` set (dtree_to_list (Node r {|(t2, e2)|})). a \<notin> set v"
using insert.prems(2) dtree_to_list_x_in_dlverts by auto
have "xs |\<subseteq>| finsert x xs" by blast
then have f_xs: "ffold f z xs = ffold f' z xs"
using merge_ffold_supset insert.prems(1) f_def f'_def by presburger
have "ffold f z (finsert x xs) = f x (ffold f z xs)"
using Comm.ffold_finsert[OF insert.hyps] f_def by blast
then have 2: "ffold f z (finsert x xs) = f x (ffold f' z xs)" using f_xs by argo
then have "f x (ffold f' z xs) \<noteq> ffold f' z xs" using False f_def f'_def by argo
then have "f (t2,e2) (ffold f' z xs)
= ?merge (dtree_to_list (Node r {|(t2,e2)|})) (ffold f' z xs)"
using merge_f_merge_if_not_snd t2_def f_def by blast
then have "ffold f z (finsert x xs)
= ?merge (dtree_to_list (Node r {|(t2,e2)|})) (ffold f' z xs)"
using 2 t2_def by argo
then have "path_lverts_list (ffold f' z xs) a \<subseteq> path_lverts_list (ffold f z (finsert x xs)) a"
using path_lverts_list_merge_supset_ys_notin[OF 1] by presburger
then show ?thesis using insert.IH 0 insert.prems(2) f_def f'_def by auto
qed
-qed(simp)
+qed (simp add: ffold.rep_eq)
lemma path_lverts_merge_ffold_sup:
"\<lbrakk>list_dtree (Node r xs); t1 \<in> fst ` fset xs; a \<in> dlverts t1\<rbrakk>
\<Longrightarrow> path_lverts t1 a \<subseteq> path_lverts_list (ffold (merge_f r xs) [] xs) a"
proof(induction xs)
case (insert x xs)
interpret Comm: comp_fun_commute "merge_f r (finsert x xs)" by (rule merge_commute)
define f where "f = merge_f r (finsert x xs)"
define f' where "f' = merge_f r xs"
let ?merge = "Sorting_Algorithms.merge cmp'"
have 0: "list_dtree (Node r xs)" using list_dtree_subset insert.prems(1) by blast
obtain t2 e2 where t2_def[simp]: "x = (t2,e2)" by fastforce
have "(t2, e2) \<in> fset (finsert x xs)" by simp
- moreover have "(t2, e2) \<notin> fset xs" using insert.hyps notin_fset by fastforce
+ moreover have "(t2, e2) \<notin> fset xs" using insert.hyps by fastforce
ultimately have xs_val:
"(\<forall>(v,e) \<in> set (ffold f' [] xs). set v \<inter> dlverts t2 = {} \<and> v \<noteq> [] \<and> e \<notin> darcs t2 \<union> {e2})"
using merge_ffold_empty_inter_preserv'[OF insert.prems(1) empty_list_valid_merge] f'_def
by blast
have "ffold f [] (finsert x xs) = f x (ffold f [] xs)"
using Comm.ffold_finsert[OF insert.hyps] f_def by blast
also have "\<dots> = f x (ffold f' [] xs)"
using merge_ffold_supset[of xs "finsert x xs" r "[]"] insert.prems(1) f_def f'_def by fastforce
finally have "ffold f [] (finsert x xs) = ?merge (dtree_to_list (Node r {|x|})) (ffold f' [] xs)"
using merge_f_merge_if_conds xs_val insert.prems f_def by simp
then have merge: "ffold f [] (finsert x xs)
= ?merge (dtree_to_list (Node r {|(t2,e2)|})) (ffold f'[] xs)"
using t2_def by blast
show ?case
proof(cases "t1 = t2")
case True
then have "\<forall>v\<in>fst ` set (ffold f' [] xs). a \<notin> set v"
using insert.prems(3) xs_val by fastforce
then have "path_lverts_list (dtree_to_list (Node r {|(t2,e2)|})) a
\<subseteq> path_lverts_list (ffold f [] (finsert x xs)) a"
using merge path_lverts_list_merge_supset_xs_notin by fastforce
then show ?thesis using True f_def path_lverts_to_list_eq by force
next
case False
then have "a \<notin> dlverts t2" using insert.prems list_dtree.wf_lverts by fastforce
then have 1: "\<forall>v\<in>fst ` set (dtree_to_list (Node r {|(t2, e2)|})). a \<notin> set v"
using dtree_to_list_x_in_dlverts by fast
have "path_lverts t1 a \<subseteq> path_lverts_list (ffold f' [] xs) a"
using insert.IH[OF 0] insert.prems(2,3) False f'_def by simp
then show ?thesis using f_def merge path_lverts_list_merge_supset_ys_notin[OF 1] by auto
qed
qed(simp)
lemma path_lverts_merge_sup_aux:
assumes "list_dtree (Node r xs)" and "t1 \<in> fst ` fset xs" and "a \<in> dlverts t1"
and "ffold (merge_f r xs) [] xs = (v1, e1) # ys"
shows "path_lverts t1 a \<subseteq> path_lverts (dtree_from_list v1 ys) a"
proof -
have "xs \<noteq> {||}" using assms(2) by auto
have "path_lverts t1 a \<subseteq> path_lverts_list (ffold (merge_f r xs) [] xs) a"
using path_lverts_merge_ffold_sup[OF assms(1-3)] .
then show ?thesis using path_lverts_from_list_eq assms(4) by fastforce
qed
lemma path_lverts_merge_sup:
assumes "list_dtree (Node r xs)" and "t1 \<in> fst ` fset xs" and "a \<in> dlverts t1"
shows "\<exists>t2 e2. merge (Node r xs) = Node r {|(t2,e2)|}
\<and> path_lverts t1 a \<subseteq> path_lverts t2 a"
proof -
have "xs \<noteq> {||}" using assms(2) by auto
then obtain t2 e2 where t2_def: "merge (Node r xs) = Node r {|(t2,e2)|}"
using merge_singleton[OF assms(1)] by blast
obtain y ys where y_def: "ffold (merge_f r xs) [] xs = y # ys"
using merge_ffold_nempty[OF assms(1) \<open>xs \<noteq> {||}\<close>] list.exhaust_sel by blast
obtain v1 e1 where "y = (v1,e1)" by fastforce
then show ?thesis using merge_xs path_lverts_merge_sup_aux[OF assms] t2_def y_def by fastforce
qed
lemma path_lverts_merge_sup_sucs:
assumes "list_dtree t0" and "t1 \<in> fst ` fset (sucs t0)" and "a \<in> dlverts t1"
shows "\<exists>t2 e2. merge t0 = Node (Dtree.root t0) {|(t2,e2)|}
\<and> path_lverts t1 a \<subseteq> path_lverts t2 a"
using path_lverts_merge_sup[of "Dtree.root t0" "sucs t0"] assms by simp
lemma merge_dom_children_aux:
assumes "list_dtree t0"
and "\<forall>x\<in>dverts t1. \<exists>v \<in> set (Dtree.root t0) \<union> path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
and "t1 \<in> fst ` fset (sucs t0)"
and "wf_dlverts t1"
and "x \<in> dverts t1"
shows "\<exists>!t2 \<in> fst ` fset (sucs (merge t0)).
\<exists>v \<in> set (Dtree.root (merge t0)) \<union> path_lverts t2 (hd x). v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
proof -
have "hd x \<in> dlverts t1" using assms(4,5) by (simp add: hd_in_lverts_if_wf)
then obtain t2 e2 where t2_def:
"merge t0 = Node (Dtree.root t0) {|(t2,e2)|}" "path_lverts t1 (hd x) \<subseteq> path_lverts t2 (hd x)"
using path_lverts_merge_sup_sucs[OF assms(1,3)] by blast
then show ?thesis using assms(2,5) by force
qed
lemma merge_dom_children_aux':
assumes "dom_children t0 T"
and "\<forall>t1 \<in> fst ` fset (sucs t0). wf_dlverts t1"
and "t2 \<in> fst ` fset (sucs (merge t0))"
and "x \<in> dverts t2"
shows "\<exists>v\<in>set (Dtree.root (merge t0)) \<union> path_lverts t2 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
proof -
have disj: "list_dtree t0"
using assms(3) merge_empty_if_nwf_sucs[of t0] by fastforce
obtain t1 where t1_def: "t1 \<in> fst ` fset (sucs t0)" "x \<in> dverts t1"
using verts_child_if_merge_child[OF assms(3,4)] by blast
then have 0: "\<forall>x\<in>dverts t1. \<exists>v\<in>set (Dtree.root t0) \<union> path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
using assms(1) unfolding dom_children_def by blast
then have "wf_dlverts t1" using t1_def(1) assms(2) by blast
then obtain t3 where t3_def: "t3 \<in> fst ` fset (sucs (merge t0))"
"(\<exists>v\<in>set (Dtree.root (merge t0)) \<union> path_lverts t3 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x)"
using merge_dom_children_aux[OF disj 0] t1_def by blast
then have "t3 = t2" using assms(3) merge_single_root1_sucs by fastforce
then show ?thesis using t3_def(2) by blast
qed
lemma merge_dom_children_sucs:
assumes "dom_children t0 T" and "\<forall>t1 \<in> fst ` fset (sucs t0). wf_dlverts t1"
shows "dom_children (merge t0) T"
using merge_dom_children_aux'[OF assms] dom_children_def by fast
lemma merge_dom_children:
"\<lbrakk>dom_children (Node r xs) T; \<forall>t1 \<in> fst ` fset xs. wf_dlverts t1\<rbrakk>
\<Longrightarrow> dom_children (merge (Node r xs)) T"
using merge_dom_children_sucs by auto
lemma merge_dom_children_if_ndisjoint:
"\<not>list_dtree (Node r xs) \<Longrightarrow> dom_children (merge (Node r xs)) T"
using merge_empty_if_nwf unfolding dom_children_def by simp
lemma merge_subtree_fcard_le1: "is_subtree (Node r xs) (merge t1) \<Longrightarrow> fcard xs \<le> 1"
using merge_mdeg_le1_sub le_trans mdeg_ge_fcard by fast
lemma merge_dom_mdeg_gt1:
"\<lbrakk>is_subtree (Node r xs) (merge t2); t1 \<in> fst ` fset xs; max_deg (Node r xs) > 1\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
using merge_mdeg_le1_sub by fastforce
lemma merge_root_if_contr:
"\<lbrakk>\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t1 \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2));
is_subtree (Node v {|(t2,e2)|}) (merge t1); rank (rev (Dtree.root t2)) < rank (rev v)\<rbrakk>
\<Longrightarrow> Node v {|(t2,e2)|} = merge t1"
using merge_strict_subtree_nocontr_sucs2[of t1 v] strict_subtree_def by fastforce
lemma merge_new_contr_fcard_gt1:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t1 \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "Node v {|(t2,e2)|} = (merge t1)"
and "rank (rev (Dtree.root t2)) < rank (rev v)"
shows "fcard (sucs t1) > 1"
proof -
have t_v: "Dtree.root t1 = v" using assms(2) dtree.sel(1)[of v "{|(t2,e2)|}"] by simp
have "\<forall>t2 e2. Node v {|(t2,e2)|} \<noteq> t1"
using assms merge_root_child_eq self_subtree less_le_not_le by metis
then have "\<forall>x. sucs t1 \<noteq> {|x|}" using t_v dtree.collapse[of t1] by force
moreover have "sucs t1 \<noteq> {||}" using assms(2) merge_empty_sucs by force
ultimately show ?thesis using fcard_single_1_iff[of "sucs t1"] fcard_0_eq[of "sucs t1"] by force
qed
lemma merge_dom_sub_contr_if_nocontr:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "is_subtree (Node r xs) (merge t)"
and "t1 \<in> fst ` fset xs"
and "\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) (Node r xs)
\<and> rank (rev (Dtree.root t2)) < rank (rev v)"
shows "\<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
proof -
obtain v t2 e2 where t2_def:
"is_subtree (Node v {|(t2,e2)|}) (Node r xs)" "rank (rev (Dtree.root t2)) < rank (rev v)"
using assms(4) by blast
then have "is_subtree (Node v {|(t2,e2)|}) (merge t)" using assms(2) subtree_trans by blast
then have eq: "Node v {|(t2,e2)|} = merge t" using merge_root_if_contr assms(1) t2_def(2) by blast
then have t_v: "Dtree.root t = v" using dtree.sel(1)[of v "{|(t2,e2)|}"] by simp
have eq2: "Node v {|(t2,e2)|} = Node r xs"
using eq assms(2) t2_def(1) subtree_antisym[of "Node v {|(t2, e2)|}"] by simp
have "fcard (sucs t) > 1" using merge_new_contr_fcard_gt1[OF assms(1) eq t2_def(2)] by simp
then have mdeg: "max_deg t > 1" using mdeg_ge_fcard[of "sucs t" "Dtree.root t"] by simp
have sub: "is_subtree (Node (Dtree.root t) (sucs t)) t" using self_subtree[of t] by simp
obtain e1 where e1_def: "(t1, e1)\<in>fset (sucs (merge t))"
using assms(3) eq eq2 dtree.sel(2)[of r xs] by force
then obtain t3 where t3_def: "(t3, e1)\<in>fset (sucs t)" "Dtree.root t3 = Dtree.root t1"
using merge_child_in_orig[OF e1_def] by blast
then have "\<exists>v\<in>set (Dtree.root t). v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)" using dom_mdeg_gt1 sub mdeg by fastforce
then show ?thesis using t_v eq2 by blast
qed
lemma merge_dom_contr_if_nocontr_mdeg_le1:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "is_subtree (Node r {|(t1,e1)|}) (merge t)"
and "rank (rev (Dtree.root t1)) < rank (rev r)"
and "\<forall>t \<in> fst ` fset (sucs t). max_deg t \<le> 1"
shows "dom_children (Node r {|(t1,e1)|}) T"
proof -
have eq: "Node r {|(t1,e1)|} = merge t" using merge_root_if_contr[OF assms(1-3)] .
have 0: "\<forall>t1\<in>fst ` fset (sucs t). wf_dlverts t1" using wf_lverts wf_dlverts_suc by auto
have "fcard (sucs t) > 1" using merge_new_contr_fcard_gt1[OF assms(1) eq assms(3)] by simp
then have "dom_children t T" using dom_wedge_full[of "Dtree.root t"] assms(4) self_subtree by force
then show ?thesis using merge_dom_children_sucs 0 eq by simp
qed
lemma merge_dom_wedge:
"\<lbrakk>is_subtree (Node r xs) (merge t1); fcard xs > 1; \<forall>t \<in> fst ` fset xs. max_deg t \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r xs) T"
using merge_subtree_fcard_le1 by fastforce
subsubsection \<open>Merge1 preserves Arc Invariants\<close>
lemma merge1_dom_mdeg_gt1:
assumes "is_subtree (Node r xs) (merge1 t)" and "t1 \<in> fst ` fset xs" and "max_deg (Node r xs) > 1"
shows "\<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
proof -
obtain ys where ys_def: "merge1 (Node r ys) = Node r xs" "is_subtree (Node r ys) t"
using merge1_subtree_if_mdeg_gt1[OF assms(1,3)] by blast
then obtain t3 where t3_def: "t3 \<in> fst ` fset ys" "Dtree.root t3 = Dtree.root t1"
using assms(2) merge1_child_in_orig by fastforce
have "max_deg (Node r ys) > 1" using merge1_mdeg_le[of "Node r ys"] ys_def(1) assms(3) by simp
then show ?thesis using dom_mdeg_gt1[OF ys_def(2) t3_def(1)] t3_def by simp
qed
lemma max_deg1_gt_1_if_new_contr:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t0 \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "is_subtree (Node r {|(t1,e1)|}) (merge1 t0)"
and "rank (rev (Dtree.root t1)) < rank (rev r)"
shows "max_deg t0 > 1"
using assms merge1_mdeg_gt1_if_uneq by force
lemma merge1_subtree_if_new_contr:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t0 \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "is_subtree (Node r xs) (merge1 t0)"
and "is_subtree (Node v {|(t1,e1)|}) (Node r xs)"
and "rank (rev (Dtree.root t1)) < rank (rev v)"
shows "\<exists>ys. is_subtree (Node r ys) t0 \<and> merge1 (Node r ys) = Node r xs"
using assms proof(induction t0)
case (Node r' ys)
then consider "fcard ys > 1" "(\<forall>t \<in> fst ` fset ys. max_deg t \<le> 1)"
| "\<not>(fcard ys > 1 \<and> (\<forall>t \<in> fst ` fset ys. max_deg t \<le> 1))" "Node r xs = merge1 (Node r' ys)"
| "\<not>(fcard ys > 1 \<and> (\<forall>t \<in> fst ` fset ys. max_deg t \<le> 1))" "Node r xs \<noteq> merge1 (Node r' ys)"
by blast
then show ?case
proof(cases)
case 1
then have "is_subtree (Node v {|(t1, e1)|}) (merge (Node r' ys))"
using subtree_trans[OF Node.prems(3,2)] by force
then have "Node v {|(t1, e1)|} = merge (Node r' ys)"
using merge_root_if_contr Node.prems(1,4) by blast
then have "Node r xs = merge1 (Node r' ys)"
using Node.prems(2,3) 1 subtree_eq_if_trans_eq1 by fastforce
then show ?thesis using 1 dtree.sel(1)[of r xs] by auto
next
case 2
then have "r = r'" using dtree.sel(1)[of r xs] by force
then show ?thesis using 2(2) by auto
next
case 3
then have "merge1 (Node r' ys) = Node r' ((\<lambda>(t,e). (merge1 t,e)) |`| ys)" by auto
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset ys" "is_subtree (Node r xs) (merge1 t2)"
using Node.prems(2) 3(2) by auto
- then have subt2: "is_subtree t2 (Node r' ys)" using subtree_if_child by fastforce
+ then have subt2: "is_subtree t2 (Node r' ys)" using subtree_if_child
+ by (metis fstI image_eqI)
then have "\<And>r1 t3 e3. is_subtree (Node r1 {|(t3, e3)|}) t2
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t3))"
using Node.prems(1) subtree_trans by blast
then obtain ys' where ys_def: "is_subtree (Node r ys') t2" "merge1 (Node r ys') = Node r xs"
using Node.IH[OF t2_def(1)] Node.prems(3,4) t2_def(2) by auto
then show ?thesis using subtree_trans subt2 by blast
qed
qed
lemma merge1_dom_sub_contr:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "is_subtree (Node r xs) (merge1 t)"
and "t1 \<in> fst ` fset xs"
and "\<exists>v t2 e2. is_subtree (Node v {|(t2,e2)|}) (Node r xs)\<and>rank (rev (Dtree.root t2))<rank (rev v)"
shows "\<exists>v \<in> set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)"
proof -
obtain ys where ys_def: "is_subtree (Node r ys) t" "merge1 (Node r ys) = Node r xs"
using merge1_subtree_if_new_contr assms(1,2,4) by blast
then interpret R: ranked_dtree_with_orig "Node r ys" using ranked_dtree_orig_subtree by blast
obtain v t2 e2 where v_def:
"is_subtree (Node v {|(t2,e2)|}) (Node r xs)" "rank (rev (Dtree.root t2)) < rank (rev v)"
using assms(4) by blast
then have "is_subtree (Node v {|(t2,e2)|}) (merge1 (Node r ys))" using ys_def by simp
then have mdeg_gt1: "max_deg (Node r ys) > 1"
using max_deg1_gt_1_if_new_contr assms(1) v_def(2) subtree_trans ys_def(1) by blast
obtain t3 where t3_def: "t3 \<in> fst ` fset ys" "Dtree.root t3 = Dtree.root t1"
using ys_def(2) assms(3) merge1_child_in_orig by fastforce
then show ?thesis using R.dom_mdeg_gt1[OF self_subtree] mdeg_gt1 by fastforce
qed
lemma merge1_merge_point_if_new_contr:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t0 \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "wf_darcs t0"
and "is_subtree (Node r {|(t1,e1)|}) (merge1 t0)"
and "rank (rev (Dtree.root t1)) < rank (rev r)"
shows "\<exists>ys. is_subtree (Node r ys) t0 \<and> fcard ys > 1 \<and> (\<forall>t\<in> fst ` fset ys. max_deg t \<le> 1)
\<and> merge1 (Node r ys) = Node r {|(t1,e1)|}"
using assms proof(induction t0)
case (Node v xs)
then consider "fcard xs > 1" "(\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)"
| "fcard xs \<le> 1" | "fcard xs > 1" "\<not>(\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)"
by linarith
then show ?case
proof(cases)
case 1
then have "is_subtree (Node r {|(t1, e1)|}) (merge (Node v xs))" using Node.prems(3) by simp
then have "Node r {|(t1, e1)|} = merge (Node v xs)"
using merge_root_if_contr Node.prems(1,4) by blast
then show ?thesis using 1 dtree.sel(1)[of r "{|(t1, e1)|}"] by auto
next
case 2
then have "merge1 (Node v xs) = Node v ((\<lambda>(t,e). (merge1 t,e)) |`| xs)" by auto
then have "xs \<noteq> {||}" using Node.prems(3) by force
then have "fcard xs = 1" using 2 le_Suc_eq by auto
then obtain t2 e2 where t2_def: "xs = {|(t2,e2)|}" using fcard_single_1_iff[of xs] by fast
then have "Node r {|(t1, e1)|} \<noteq> merge1 (Node v {|(t2,e2)|})" using Node.prems(1,4) 2 by force
then have "is_subtree (Node r {|(t1, e1)|}) (merge1 t2)" using Node.prems(3) t2_def 2 by auto
moreover have "\<And>r1 t3 e3. is_subtree (Node r1 {|(t3, e3)|}) t2
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t3))"
using Node.prems(1) t2_def by fastforce
ultimately show ?thesis using Node.IH[of "(t2,e2)"] Node.prems(2,4) t2_def by fastforce
next
case 3
then have "fcard ((\<lambda>(t,e). (merge1 t,e)) |`| xs) > 1"
using fcard_merge1_img_if_disjoint disjoint_darcs_if_wf_xs[OF Node.prems(2)] by simp
then have "Node r {|(t1,e1)|} \<noteq> merge1 (Node v xs)"
using fcard_single_1_iff[of "(\<lambda>(t,e). (merge1 t,e)) |`| xs"] 3(2) by auto
moreover have "merge1 (Node v xs) = Node v ((\<lambda>(t,e). (merge1 t,e)) |`| xs)" using 3(2) by auto
ultimately obtain t2 e2 where t2_def:
"(t2,e2) \<in> fset xs" "is_subtree (Node r {|(t1, e1)|}) (merge1 t2)"
using Node.prems(3) by auto
- then have "is_subtree t2 (Node v xs)" using subtree_if_child by fastforce
+ then have "is_subtree t2 (Node v xs)" using subtree_if_child
+ by (metis fst_conv image_eqI)
then have "\<And>r1 t3 e3. is_subtree (Node r1 {|(t3, e3)|}) t2
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t3))"
using Node.prems(1) subtree_trans by blast
then obtain ys where ys_def: "is_subtree (Node r ys) t2" "1 < fcard ys"
"(\<forall>t\<in>fst ` fset ys. max_deg t \<le> 1)" "merge1 (Node r ys) = Node r {|(t1, e1)|}"
using Node.IH[OF t2_def(1)] Node.prems(2,4) t2_def by fastforce
then show ?thesis using t2_def(1) by auto
qed
qed
lemma merge1_dom_contr:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "is_subtree (Node r {|(t1,e1)|}) (merge1 t)"
and "rank (rev (Dtree.root t1)) < rank (rev r)"
and "max_deg (Node r {|(t1,e1)|}) = 1"
shows "dom_children (Node r {|(t1,e1)|}) T"
proof -
obtain ys where ys_def: "is_subtree (Node r ys) t" "fcard ys > 1"
"\<forall>t\<in>fst ` fset ys. max_deg t \<le> 1" "merge1 (Node r ys) = Node r {|(t1,e1)|}"
using merge1_merge_point_if_new_contr wf_arcs assms(1-3) by blast
have "\<forall>t1\<in>fst ` fset ys. wf_dlverts t1"
using ys_def(1) list_dtree.wf_lverts list_dtree_sub by fastforce
then show ?thesis using merge_dom_children_sucs[OF dom_wedge_full] ys_def by fastforce
qed
lemma merge1_dom_children_merge_sub_aux:
assumes "merge1 t = t2"
and "is_subtree (Node r' xs') t"
and "fcard xs' > 1"
and "(\<forall>t\<in>fst ` fset xs'. max_deg t \<le> 1)"
and "max_deg t2 \<le> 1"
and "x \<in> dverts t2"
and "x \<noteq> Dtree.root t2"
shows "\<exists>v \<in> path_lverts t2 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
using assms ranked_dtree_with_orig_axioms proof(induction t arbitrary: t2)
case (Node r xs)
then interpret R: ranked_dtree_with_orig "Node r xs" by blast
obtain t1 e1 where t1_def: "(t1,e1) \<in> fset (sucs t2)" "x \<in> dverts t1"
by (metis Node.prems(6,7) fsts.simps dtree.sel dtree.set_cases(1) fst_conv surj_pair)
then have t2_sucs: "sucs t2 = {|(t1,e1)|}"
using Node.prems(5) empty_iff_mdeg_0[of "Dtree.root t2" "sucs t2"]
mdeg_1_singleton[of "Dtree.root t2" "sucs t2"] by auto
have wf_t2: "wf_dlverts t2" using Node.prems(1) R.wf_dlverts_merge1 by blast
then have "wf_dlverts t1" using t1_def(1) wf_dlverts_suc by fastforce
then have "hd x \<in> dlverts t1" using t1_def(2) hd_in_lverts_if_wf by blast
then have "hd x \<notin> set (Dtree.root t2)" using dlverts_notin_root_sucs[OF wf_t2] t1_def(1) by fastforce
then have path_t2: "path_lverts t2 (hd x) = set (Dtree.root t2) \<union> path_lverts t1 (hd x)"
using path_lverts_simps1_sucs t2_sucs by fastforce
show ?case
proof(cases "Node r xs = Node r' xs'")
case True
then have "merge (Node r' xs') = t2" using Node.prems(1,3,4) by simp
then have "dom_children t2 T"
using R.dom_wedge_full[OF Node.prems(2-4)] merge_dom_children R.wf_lverts True by fastforce
then have "\<exists>v\<in>set (Dtree.root t2) \<union> path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
using t1_def unfolding dom_children_def by auto
then show ?thesis using path_t2 by blast
next
case False
then have "\<not>(fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1))"
using Node.prems(3,4) child_mdeg_gt1_if_sub_fcard_gt1[OF Node.prems(2)] by force
then have eq: "merge1 (Node r xs) = Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs)" by auto
then obtain t3 e3 where t3_def: "(t3,e3) \<in> fset xs" "is_subtree (Node r' xs') t3"
using Node.prems(2) False by auto
have "fcard ((\<lambda>(t,e). (merge1 t,e)) |`| xs) = 1"
using Node.prems(1) eq t2_sucs fcard_single_1 by fastforce
then have "fcard xs = 1"
using fcard_merge1_img_if_disjoint disjoint_darcs_if_wf_xs[OF R.wf_arcs] by simp
then have "xs = {|(t3,e3)|}" using fcard_single_1_iff[of xs] t3_def(1) by auto
then have t13: "merge1 t3 = t1" using t2_sucs eq Node.prems(1) by force
then have mdegt3: "max_deg t1 \<le> 1"
using Node.prems(5) mdeg_ge_child[of t1 e1 "sucs t2" "Dtree.root t2"] t2_sucs by fastforce
have mdeg_gt1: "max_deg (Node r xs) > 1"
using mdeg_ge_fcard[of xs' r'] Node.prems(2,3) mdeg_ge_sub[of "Node r' xs'" "Node r xs"]
by simp
show ?thesis
proof(cases "x = Dtree.root t1")
case True
then have "\<exists>v\<in>set r. v \<rightarrow>\<^bsub>T\<^esub> hd x"
using R.dom_mdeg_gt1[of r xs] t3_def(1) mdeg_gt1 t13 by fastforce
then show ?thesis using path_t2 Node.prems(1) by auto
next
case False
then have "\<exists>v\<in>path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
using Node.IH t1_def(2) t3_def t13 assms(3,4) mdegt3 R.ranked_dtree_orig_rec by simp
then show ?thesis using path_t2 by blast
qed
qed
qed
lemma merge1_dom_children_fcard_gt1_aux:
assumes "dom_children (Node r (Abs_fset (children_deg1 ys))) T"
and "is_subtree (Node r ys) t"
and "merge1 (Node r ys) = Node r xs"
and "fcard xs > 1"
and "max_deg t2 \<le> 1"
and "t2 \<in> fst ` fset xs"
and "x \<in> dverts t2"
shows "\<exists>v\<in>set r \<union> path_lverts t2 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
proof -
obtain t1 where t1_def: "t1 \<in> fst ` fset ys" "merge1 t1 = t2"
using merge1_elem_in_img_if_fcard_gt1[OF assms(3,4)] assms(6) by fastforce
then have x_t: "x \<in> dverts t1" using merge1_dverts_sub assms(7) by blast
show ?thesis
proof(cases "max_deg t1 \<le> 1")
case True
then have "t1 \<in> fst ` fset (sucs (Node r (Abs_fset (children_deg1 ys))))"
using t1_def(1) children_deg1_fset_id by force
then have "\<exists>v\<in>set r \<union> path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> hd x"
using assms(1) x_t unfolding dom_children_def by auto
then show ?thesis using t1_def(2) merge1_mdeg_gt1_if_uneq[of t1] True by force
next
case False
then obtain r' xs' where r'_def:
"is_subtree (Node r' xs') t1" "1 < fcard xs'" "(\<forall>t\<in>fst ` fset xs'. max_deg t \<le> 1)"
using merge1_wedge_if_uneq[of t1] assms(5) t1_def(2) by fastforce
interpret R: ranked_dtree_with_orig "Node r ys" using ranked_dtree_orig_subtree assms(2) .
interpret T: ranked_dtree_with_orig t1 using R.ranked_dtree_orig_rec t1_def(1) by force
have "max_deg (Node r ys) > 1"
using assms(3,4) merge1_fcard_le[of r ys] mdeg_ge_fcard[of ys] by simp
show ?thesis
proof (cases "x = Dtree.root t2")
case True
have "max_deg (Node r ys) > 1"
using assms(3,4) merge1_fcard_le[of r ys] mdeg_ge_fcard[of ys] by simp
then show ?thesis using dom_mdeg_gt1[OF assms(2) t1_def(1)] True t1_def(2) by auto
next
case False
then show ?thesis
using T.merge1_dom_children_merge_sub_aux[OF t1_def(2) r'_def assms(5,7)] by blast
qed
qed
qed
lemma merge1_dom_children_fcard_gt1:
assumes "dom_children (Node r (Abs_fset (children_deg1 ys))) T"
and "is_subtree (Node r ys) t"
and "merge1 (Node r ys) = Node r xs"
and "fcard xs > 1"
shows "dom_children (Node r (Abs_fset (children_deg1 xs))) T"
unfolding dom_children_def
using merge1_dom_children_fcard_gt1_aux[OF assms] children_deg1_fset_id[of xs] by fastforce
lemma merge1_dom_wedge:
assumes "is_subtree (Node r xs) (merge1 t)" and "fcard xs > 1"
shows "dom_children (Node r (Abs_fset (children_deg1 xs))) T"
proof -
obtain ys where ys_def:
"merge1 (Node r ys) = Node r xs" "is_subtree (Node r ys) t" "fcard xs \<le> fcard ys"
using merge1_subtree_if_fcard_gt1[OF assms] by blast
have "dom_children (Node r (Abs_fset (children_deg1 ys))) T"
using dom_wedge ys_def(2,3) assms(2) by simp
then show ?thesis using merge1_dom_children_fcard_gt1 ys_def(2,1) assms(2) by blast
qed
corollary merge1_dom_wedge':
"\<forall>r xs. is_subtree (Node r xs) (merge1 t) \<longrightarrow> fcard xs > 1
\<longrightarrow> dom_children (Node r (Abs_fset {(t, e). (t, e) \<in> fset xs \<and> max_deg t \<le> Suc 0})) T"
by (auto simp only: merge1_dom_wedge One_nat_def[symmetric])
corollary merge1_verts_conform: "v \<in> dverts (merge1 t) \<Longrightarrow> seq_conform v"
by (simp add: verts_conform)
corollary merge1_verts_distinct: "\<lbrakk>v \<in> dverts (merge1 t)\<rbrakk> \<Longrightarrow> distinct v"
using distinct_merge1 verts_distinct by auto
lemma merge1_mdeg_le1_wedge_if_fcard_gt1:
assumes "max_deg (merge1 t1) \<le> 1"
and "wf_darcs t1"
and "is_subtree (Node v ys) t1"
and "fcard ys > 1"
shows "(\<forall>t \<in> fst ` fset ys. max_deg t \<le> 1)"
using assms proof(induction t1 rule: merge1.induct)
case (1 r xs)
then show ?case
proof(cases "fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)")
case True
then have "Node v ys = Node r xs"
using "1.prems"(3,4) mdeg_ge_sub mdeg_ge_fcard[of ys] by fastforce
then show ?thesis using True by simp
next
case False
then have eq: "merge1 (Node r xs) = Node r ((\<lambda>(t, e). (merge1 t, e)) |`| xs)" by auto
have "fcard ((\<lambda>(t, e). (merge1 t, e)) |`| xs) = fcard xs"
using fcard_merge1_img_if_disjoint disjoint_darcs_if_wf_xs[OF "1.prems"(2)] by simp
then have "fcard xs \<le> 1"
by (metis "1.prems"(1) False merge1.simps num_leaves_1_if_mdeg_1 num_leaves_ge_card)
then have "Node v ys \<noteq> Node r xs" using "1.prems"(4) by auto
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "is_subtree (Node v ys) t2"
using "1.prems"(3) by auto
then have "max_deg (merge1 t2) \<le> 1"
using "1.prems"(1) False eq
mdeg_ge_child[of "merge1 t2" e2 "(\<lambda>(t, e). (merge1 t, e)) |`| xs"]
by fastforce
then show ?thesis using "1.IH"[OF False t2_def(1) refl] t2_def "1.prems"(2,4) by fastforce
qed
qed
lemma dom_mdeg_le1_merge1_aux:
assumes "max_deg (merge1 t) \<le> 1"
and "merge1 t \<noteq> t"
and "t1 \<in> fst ` fset (sucs (merge1 t))"
and "x \<in> dverts t1"
shows "\<exists>r\<in>set (Dtree.root (merge1 t)) \<union> path_lverts t1 (hd x). r \<rightarrow>\<^bsub>T\<^esub> hd x"
using assms ranked_dtree_with_orig_axioms proof(induction t arbitrary: t1 rule: merge1.induct)
case (1 r xs)
then interpret R: ranked_dtree_with_orig "Node r xs" by blast
show ?case
proof(cases "fcard xs > 1")
case True
then have 0: "(\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1)"
using merge1_mdeg_le1_wedge_if_fcard_gt1[OF "1.prems"(1) R.wf_arcs] by auto
then have "dom_children (merge (Node r xs)) T"
using True merge_dom_children_sucs R.dom_wedge_full R.wf_lverts self_subtree wf_dlverts_suc
by fast
then show ?thesis unfolding dom_children_def using "1.prems"(3,4) 0 True by auto
next
case False
then have rec: "\<not>(fcard xs > 1 \<and> (\<forall>t \<in> fst ` fset xs. max_deg t \<le> 1))" by simp
then have eq: "merge1 (Node r xs) = Node r ((\<lambda>(t,e). (merge1 t,e)) |`| xs)" by auto
obtain t2 e2 where t2_def: "xs = {|(t2,e2)|}" "merge1 t2 = t1"
using "1.prems"(3) False singleton_if_fcard_le1_elem[of xs] by fastforce
show ?thesis
proof(cases "x = Dtree.root t1")
case True
have "max_deg (Node r xs) > 1" using merge1_mdeg_gt1_if_uneq "1.prems"(2) by blast
then show ?thesis using True R.dom_mdeg_gt1[OF self_subtree] t2_def by auto
next
case False
then obtain t3 where t3_def: "t3 \<in> fst ` fset (sucs (merge1 t2))" "x \<in> dverts t3"
using "1.prems"(4) t2_def(2) dverts_root_or_suc by fastforce
have mdeg1: "max_deg (merge1 t2) \<le> 1"
using "1.prems"(1) mdeg_ge_child[of t1 e2 "(\<lambda>(t,e). (merge1 t,e)) |`| xs"] eq t2_def
by simp
then have 0: "\<exists>r\<in>set (Dtree.root (merge1 t2)) \<union> path_lverts t3 (hd x). r \<rightarrow>\<^bsub>T\<^esub> hd x"
using "1.IH" rec mdeg1 t3_def "1.prems"(2) eq t2_def R.ranked_dtree_orig_rec by auto
obtain e3 where e3_def: "sucs t1 = {|(t3, e3)|}"
using t3_def singleton_if_mdeg_le1_elem_suc mdeg1 t2_def(2) by fastforce
have "wf_dlverts t1" using wf_dlverts_suc "1.prems"(3) R.wf_dlverts_merge1 by blast
then have "hd x \<in> dlverts t3"
using t3_def(2) "1.prems"(4) list_in_verts_iff_lverts hd_in_set[of x] empty_notin_wf_dlverts
by fast
then have "hd x \<notin> set (Dtree.root t1)"
using t3_def(1) dlverts_notin_root_sucs[OF \<open>wf_dlverts t1\<close>] t2_def(2) by blast
then show ?thesis using 0 path_lverts_simps1_sucs[of "hd x" t1] e3_def t2_def(2) by blast
qed
qed
qed
lemma dom_mdeg_le1_merge1:
"\<lbrakk>max_deg (merge1 t) \<le> 1; merge1 t \<noteq> t\<rbrakk> \<Longrightarrow> dom_children (merge1 t) T"
unfolding dom_children_def using dom_mdeg_le1_merge1_aux by blast
lemma merge1_arc_in_dlverts:
"\<lbrakk>is_subtree (Node r xs) (merge1 t); x \<in> set r; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts (Node r xs)"
using merge1_subtree_dlverts_supset arc_in_dlverts by blast
theorem merge1_ranked_dtree_orig:
assumes "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
shows "ranked_dtree_with_orig (merge1 t) rank cost cmp T root"
using assms merge1_arc_in_dlverts
unfolding ranked_dtree_with_orig_def ranked_dtree_with_orig_axioms_def
by(simp add: directed_tree_axioms ranked_dtree_merge1 merge1_verts_distinct merge1_verts_conform
merge1_dom_mdeg_gt1 merge1_dom_contr merge1_dom_sub_contr merge1_dom_wedge' asi_rank)
theorem merge1_normalize_ranked_dtree_orig:
"ranked_dtree_with_orig (merge1 (normalize t)) rank cost cmp T root"
using ranked_dtree_with_orig.merge1_ranked_dtree_orig[OF ranked_dtree_orig_normalize]
by (simp add: normalize_sorted_ranks)
theorem ikkbz_sub_ranked_dtree_orig: "ranked_dtree_with_orig (ikkbz_sub t) rank cost cmp T root"
using ranked_dtree_with_orig_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
then show ?case
proof(cases "max_deg t \<le> 1")
case True
then show ?thesis using "1.prems" by auto
next
case False
then show ?thesis
by (metis 1 ranked_dtree_with_orig.merge1_normalize_ranked_dtree_orig ikkbz_sub.simps)
qed
qed
subsection \<open>Optimality of IKKBZ-Sub result constrained to Invariants\<close>
lemma dtree_size_skip_decr[termination_simp]: "size (Node r (sucs t1)) < size (Node v {|(t1,e1)|})"
using dtree_size_eq_root[of "Dtree.root t1" "sucs t1"] by auto
lemma dtree_size_skip_decr1: "size (Node (r @ Dtree.root t1) (sucs t1)) < size (Node r {|(t1,e1)|})"
using dtree_size_skip_decr by auto
function normalize_full :: "('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"normalize_full (Node r {|(t1,e1)|}) = normalize_full (Node (r@Dtree.root t1) (sucs t1))"
| "\<forall>x. xs \<noteq> {|x|} \<Longrightarrow> normalize_full (Node r xs) = Node r xs"
using dtree_to_list.cases by blast+
termination using dtree_size_skip_decr "termination" in_measure wf_measure by metis
subsubsection \<open>Result fulfills the requirements\<close>
lemma ikkbz_sub_eq_if_mdeg_le1: "max_deg t1 \<le> 1 \<Longrightarrow> ikkbz_sub t1 = t1"
by simp
lemma ikkbz_sub_eq_iff_mdeg_le1: "max_deg t1 \<le> 1 \<longleftrightarrow> ikkbz_sub t1 = t1"
using ikkbz_sub_mdeg_le1[of t1] by fastforce
lemma dom_mdeg_le1_ikkbz_sub: "ikkbz_sub t \<noteq> t \<Longrightarrow> dom_children (ikkbz_sub t) T"
using ranked_dtree_with_orig_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
then interpret T: ranked_dtree_with_orig t by simp
interpret NT: ranked_dtree_with_orig "normalize t"
using T.ranked_dtree_orig_normalize by blast
interpret MT: ranked_dtree_with_orig "merge1 (normalize t)"
using T.merge1_normalize_ranked_dtree_orig by blast
show ?case
proof(cases "max_deg t \<le> 1")
case True
then show ?thesis using "1.prems" by auto
next
case False
then show ?thesis
proof(cases "max_deg (merge1 (normalize t)) \<le> 1")
case True
then show ?thesis
using NT.dom_mdeg_le1_merge1 T.dom_mdeg_le1_normalize T.list_dtree_axioms False
by force
next
case False
then have "ikkbz_sub (merge1 (normalize t)) \<noteq> (merge1 (normalize t))"
using ikkbz_sub_mdeg_le1[of "merge1 (normalize t)"] by force
then show ?thesis using 1 MT.ranked_dtree_with_orig_axioms by auto
qed
qed
qed
lemma combine_denormalize_eq:
"denormalize (Node r {|(t1,e1)|}) = denormalize (Node (r@Dtree.root t1) (sucs t1))"
by (induction t1 rule: denormalize.induct) auto
lemma normalize1_denormalize_eq: "wf_dlverts t1 \<Longrightarrow> denormalize (normalize1 t1) = denormalize t1"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case using combine_denormalize_eq[of r t] by simp
next
case (2 xs r)
then show ?case
using fcard_single_1_iff[of "(\<lambda>(t,e). (normalize1 t,e)) |`| xs"] fcard_single_1_iff[of xs]
by (auto simp: fcard_normalize_img_if_wf_dlverts)
qed
lemma normalize1_denormalize_eq': "wf_darcs t1 \<Longrightarrow> denormalize (normalize1 t1) = denormalize t1"
proof(induction t1 rule: normalize1.induct)
case (1 r t e)
then show ?case using combine_denormalize_eq[of r t] by (auto simp: wf_darcs_iff_darcs')
next
case (2 xs r)
then show ?case
using fcard_single_1_iff[of "(\<lambda>(t,e). (normalize1 t,e)) |`| xs"] fcard_single_1_iff[of xs]
by (auto simp: fcard_normalize_img_if_disjoint wf_darcs_iff_darcs')
qed
lemma normalize_denormalize_eq: "wf_dlverts t1 \<Longrightarrow> denormalize (normalize t1) = denormalize t1"
apply (induction t1 rule: normalize.induct)
by (smt (verit) normalize1_denormalize_eq normalize.simps wf_dlverts_normalize1)
lemma normalize_denormalize_eq': "wf_darcs t1 \<Longrightarrow> denormalize (normalize t1) = denormalize t1"
apply (induction t1 rule: normalize.induct)
by (smt (verit) normalize1_denormalize_eq' normalize.simps wf_darcs_normalize1)
lemma normalize_full_denormalize_eq[simp]: "denormalize (normalize_full t1) = denormalize t1"
proof(induction t1 rule: normalize_full.induct)
case (1 r t e)
then show ?case using combine_denormalize_eq[of r t] by simp
qed(simp)
lemma combine_dlverts_eq: "dlverts (Node r {|(t1,e1)|}) = dlverts (Node (r@Dtree.root t1) (sucs t1))"
using dlverts.simps[of "Dtree.root t1" "sucs t1"] by auto
lemma normalize_full_dlverts_eq[simp]: "dlverts (normalize_full t1) = dlverts t1"
using combine_dlverts_eq by(induction t1 rule: normalize_full.induct) fastforce+
lemma combine_darcs_sub: "darcs (Node (r@Dtree.root t1) (sucs t1)) \<subseteq> darcs (Node r {|(t1,e1)|})"
using dtree.set(2)[of "Dtree.root t1" "sucs t1"] by auto
lemma normalize_full_darcs_sub: "darcs (normalize_full t1) \<subseteq> darcs t1"
using combine_darcs_sub by(induction t1 rule: normalize_full.induct) fastforce+
lemma combine_nempty_if_wf_dlverts: "wf_dlverts (Node r {|(t1,e1)|}) \<Longrightarrow> r @ Dtree.root t1 \<noteq> []"
by simp
lemma combine_empty_inter_if_wf_dlverts:
assumes "wf_dlverts (Node r {|(t1,e1)|})"
shows "\<forall>(x, e1)\<in>fset (sucs t1). set (r @ Dtree.root t1) \<inter> dlverts x = {} \<and> wf_dlverts x"
proof -
have "\<forall>(x, e1)\<in>fset (sucs t1). set r \<inter> dlverts x = {}" using suc_in_dlverts assms by fastforce
then show ?thesis using wf_dlverts.simps[of "Dtree.root t1" "sucs t1"] assms by auto
qed
lemma combine_disjoint_if_wf_dlverts:
"wf_dlverts (Node r {|(t1,e1)|}) \<Longrightarrow> disjoint_dlverts (sucs t1)"
using wf_dlverts.simps[of "Dtree.root t1" "sucs t1"] by simp
lemma combine_wf_dlverts:
"wf_dlverts (Node r {|(t1,e1)|}) \<Longrightarrow> wf_dlverts (Node (r@Dtree.root t1) (sucs t1))"
using combine_empty_inter_if_wf_dlverts[of r t1] wf_dlverts.simps[of "Dtree.root t1" "sucs t1"]
by force
lemma combine_distinct:
assumes "\<forall>v \<in> dverts (Node r {|(t1,e1)|}). distinct v"
and "wf_dlverts (Node r {|(t1,e1)|})"
and "v \<in> dverts (Node (r@Dtree.root t1) (sucs t1))"
shows "distinct v"
proof(cases "v = r @ Dtree.root t1")
case True
have "(Dtree.root t1) \<in> dverts t1" by (simp add: dtree.set_sel(1))
moreover from this have "set r \<inter> set (Dtree.root t1) = {}"
using assms(2) lverts_if_in_verts by fastforce
ultimately show ?thesis using True assms(1) by simp
next
case False
then show ?thesis using assms(1,3) dverts_suc_subseteq by fastforce
qed
lemma normalize_full_wfdlverts: "wf_dlverts t1 \<Longrightarrow> wf_dlverts (normalize_full t1)"
proof(induction t1 rule: normalize_full.induct)
case (1 r t1 e1)
then show ?case using combine_wf_dlverts[of r t1] by simp
qed(simp)
corollary normalize_full_wfdverts: "wf_dlverts t1 \<Longrightarrow> wf_dverts (normalize_full t1)"
using normalize_full_wfdlverts by (simp add: wf_dverts_if_wf_dlverts)
lemma combine_wf_arcs: "wf_darcs (Node r {|(t1,e1)|}) \<Longrightarrow> wf_darcs (Node (r@Dtree.root t1) (sucs t1))"
using wf_darcs'.simps[of "Dtree.root t1" "sucs t1"] by (simp add: wf_darcs_iff_darcs')
lemma normalize_full_wfdarcs: "wf_darcs t1 \<Longrightarrow> wf_darcs (normalize_full t1)"
using combine_wf_arcs by(induction t1 rule: normalize_full.induct) fastforce+
lemma normalize_full_dom_preserv: "dom_children t1 T \<Longrightarrow> dom_children (normalize_full t1) T"
by (induction t1 rule: normalize_full.induct) (auto simp: dom_children_combine)
lemma combine_forward:
assumes "dom_children (Node r {|(t1,e1)|}) T"
and "\<forall>v \<in> dverts (Node r {|(t1,e1)|}). forward v"
and "wf_dlverts (Node r {|(t1,e1)|})"
and "v \<in> dverts (Node (r@Dtree.root t1) (sucs t1))"
shows "forward v"
proof(cases "v = r @ Dtree.root t1")
case True
have 0: "(Dtree.root t1) \<in> dverts t1" by (simp add: dtree.set_sel(1))
then have fwd_t1: "forward (Dtree.root t1)" using assms(2) by simp
moreover have "set r \<inter> set (Dtree.root t1) = {}" using assms(3) 0 lverts_if_in_verts by fastforce
moreover have "\<exists>x\<in>set r. \<exists>y\<in>set (Dtree.root t1). x \<rightarrow>\<^bsub>T\<^esub> y"
using assms(1,3) root_arc_if_dom_wfdlverts by fastforce
ultimately have "\<exists>x\<in>set r. x \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t1)" using forward_arc_to_head by blast
moreover have fwd_r: "forward r" using assms(2) by simp
ultimately show ?thesis using forward_app fwd_t1 True by simp
next
case False
then show ?thesis using assms(2,4) dverts_suc_subseteq by fastforce
qed
lemma normalize_full_forward:
"\<lbrakk>dom_children t1 T; \<forall>v \<in> dverts t1. forward v; wf_dlverts t1\<rbrakk>
\<Longrightarrow> \<forall>v \<in> dverts (normalize_full t1). forward v"
proof(induction t1 rule: normalize_full.induct)
case (1 r t e)
have "\<forall>v \<in> dverts (Node (r@Dtree.root t) (sucs t)). forward v"
using combine_forward[OF "1.prems"(1,2,3)] by blast
moreover have "dom_children (Node (r@Dtree.root t) (sucs t)) T"
using dom_children_combine "1.prems"(1) by simp
ultimately show ?case using "1.IH" "1.prems"(3) combine_wf_dlverts[of r t e] by fastforce
qed(auto)
lemma normalize_full_max_deg0: "max_deg t1 \<le> 1 \<Longrightarrow> max_deg (normalize_full t1) = 0"
proof(induction t1 rule: normalize_full.induct)
case (1 r t e)
then show ?case using mdeg_child_sucs_le by (fastforce dest: order_trans)
next
case (2 xs r)
then show ?case using empty_fset_if_mdeg_le1_not_single by auto
qed
lemma normalize_full_mdeg_eq: "max_deg t1 > 1 \<Longrightarrow> max_deg (normalize_full t1) = max_deg t1"
proof(induction t1 rule: normalize_full.induct)
case (1 r t e)
then show ?case using mdeg_child_sucs_eq_if_gt1 by force
qed(auto)
lemma normalize_full_empty_sucs: "max_deg t1 \<le> 1 \<Longrightarrow> \<exists>r. normalize_full t1 = Node r {||}"
proof(induction t1 rule: normalize_full.induct)
case (1 r t e)
then show ?case using mdeg_child_sucs_le by (fastforce dest: order_trans)
next
case (2 xs r)
then show ?case using empty_fset_if_mdeg_le1_not_single by auto
qed
lemma normalize_full_forward_singleton:
"\<lbrakk>max_deg t1 \<le> 1; dom_children t1 T; \<forall>v \<in> dverts t1. forward v; wf_dlverts t1\<rbrakk>
\<Longrightarrow> \<exists>r. normalize_full t1 = Node r {||} \<and> forward r"
using normalize_full_empty_sucs normalize_full_forward by fastforce
lemma denormalize_empty_sucs_simp: "denormalize (Node r {||}) = r"
using denormalize.simps(2) by blast
lemma normalize_full_dverts_eq_denormalize:
assumes "max_deg t1 \<le> 1"
shows "dverts (normalize_full t1) = {denormalize t1}"
proof -
obtain r where r_def[simp]: "normalize_full t1 = Node r {||}"
using assms normalize_full_empty_sucs by blast
then have "denormalize (normalize_full t1) = r" by (simp add: denormalize_empty_sucs_simp)
then have "r = denormalize t1" using normalize_full_denormalize_eq by blast
then show ?thesis by simp
qed
lemma normalize_full_normalize_dverts_eq_denormalize:
assumes "wf_dlverts t1" and "max_deg t1 \<le> 1"
shows "dverts (normalize_full (normalize t1)) = {denormalize t1}"
proof -
have "max_deg (normalize t1) \<le> 1" using assms normalize_mdeg_eq' by fastforce
then show ?thesis
using normalize_full_dverts_eq_denormalize normalize_denormalize_eq assms(1) by simp
qed
lemma normalize_full_normalize_dverts_eq_denormalize':
assumes "wf_darcs t1" and "max_deg t1 \<le> 1"
shows "dverts (normalize_full (normalize t1)) = {denormalize t1}"
proof -
have "max_deg (normalize t1) \<le> 1" using assms normalize_mdeg_eq by fastforce
then show ?thesis
using normalize_full_dverts_eq_denormalize normalize_denormalize_eq' assms(1) by simp
qed
lemma denormalize_full_forward:
"\<lbrakk>max_deg t1 \<le> 1; dom_children t1 T; \<forall>v \<in> dverts t1. forward v; wf_dlverts t1\<rbrakk>
\<Longrightarrow> forward (denormalize (normalize_full t1))"
by (metis denormalize_empty_sucs_simp normalize_full_forward_singleton)
lemma denormalize_forward:
"\<lbrakk>max_deg t1 \<le> 1; dom_children t1 T; \<forall>v \<in> dverts t1. forward v; wf_dlverts t1\<rbrakk>
\<Longrightarrow> forward (denormalize t1)"
using denormalize_full_forward by simp
lemma ikkbz_sub_forward_if_uneq: "ikkbz_sub t \<noteq> t \<Longrightarrow> forward (denormalize (ikkbz_sub t))"
using denormalize_forward ikkbz_sub_mdeg_le1 dom_mdeg_le1_ikkbz_sub ikkbz_sub_wf_dlverts
ranked_dtree_with_orig.verts_forward ikkbz_sub_ranked_dtree_orig
by fast
theorem ikkbz_sub_forward:
"\<lbrakk>max_deg t \<le> 1 \<Longrightarrow> dom_children t T\<rbrakk> \<Longrightarrow> forward (denormalize (ikkbz_sub t))"
using ikkbz_sub_forward_if_uneq ikkbz_sub_eq_iff_mdeg_le1[of t]
by (fastforce simp: verts_forward wf_lverts denormalize_forward)
lemma root_arc_singleton:
assumes "dom_children (Node r {|(t1,e1)|}) T" and "wf_dlverts (Node r {|(t1,e1)|})"
shows "\<exists>x\<in>set r. \<exists>y\<in>set (Dtree.root t1). x \<rightarrow>\<^bsub>T\<^esub> y"
using root_arc_if_dom_wfdlverts assms by fastforce
lemma before_if_dom_children_wf_conform:
assumes "dom_children (Node r {|(t1,e1)|}) T"
and "\<forall>v \<in> dverts (Node r {|(t1,e1)|}). seq_conform v"
and "wf_dlverts (Node r {|(t1,e1)|})"
shows "before r (Dtree.root t1)"
proof -
have "seq_conform (Dtree.root t1)" using dtree.set_sel(1) assms(2) by auto
moreover have "seq_conform r" using assms(2) by auto
moreover have "set r \<inter> set (Dtree.root t1) = {}"
using assms(3) dlverts_eq_dverts_union dtree.set_sel(1) by fastforce
ultimately show ?thesis unfolding before_def using root_arc_singleton assms(1,3) by blast
qed
lemma root_arc_singleton':
assumes "Node r {|(t1,e1)|} = t" and "dom_children t T"
shows "\<exists>x\<in>set r. \<exists>y\<in>set (Dtree.root t1). x \<rightarrow>\<^bsub>T\<^esub> y"
using assms root_arc_singleton wf_lverts by blast
lemma root_before_if_dom:
assumes "Node r {|(t1,e1)|} = t" and "dom_children t T"
shows "before r (Dtree.root t1)"
proof -
have "(Dtree.root t1) \<in> dverts t" using dtree.set_sel(1) assms(1) by fastforce
then have "seq_conform (Dtree.root t1)" using verts_conform by simp
moreover have "seq_conform r" using verts_conform assms(1) by auto
ultimately show ?thesis
using before_def child_disjoint_root root_arc_singleton' assms by fastforce
qed
lemma combine_conform:
"\<lbrakk>dom_children (Node r {|(t1,e1)|}) T; \<forall>v \<in> dverts (Node r {|(t1,e1)|}). seq_conform v;
wf_dlverts (Node r {|(t1,e1)|}); v \<in> dverts (Node (r@Dtree.root t1) (sucs t1))\<rbrakk>
\<Longrightarrow> seq_conform v"
apply(cases "v = r@Dtree.root t1")
using before_if_dom_children_wf_conform seq_conform_if_before apply fastforce
using dverts_suc_subseteq by fastforce
lemma denormalize_full_set_eq_dlverts:
"max_deg t1 \<le> 1 \<Longrightarrow> set (denormalize (normalize_full t1)) = dlverts t1"
using denormalize_set_eq_dlverts by auto
lemma denormalize_full_set_eq_dverts_union:
"max_deg t1 \<le> 1 \<Longrightarrow> set (denormalize (normalize_full t1)) = \<Union>(set ` dverts t1)"
using denormalize_full_set_eq_dlverts dlverts_eq_dverts_union by fastforce
corollary hd_eq_denormalize_full:
"wf_dlverts t1 \<Longrightarrow> hd (denormalize (normalize_full t1)) = hd (Dtree.root t1)"
using denormalize_hd_root_wf by auto
corollary denormalize_full_nempty_if_wf:
"wf_dlverts t1 \<Longrightarrow> denormalize (normalize_full t1) \<noteq> []"
using denormalize_nempty_if_wf by auto
lemma take1_eq_denormalize_full:
"wf_dlverts t1 \<Longrightarrow> take 1 (denormalize (normalize_full t1)) = [hd (Dtree.root t1)]"
using hd_eq_denormalize_full take1_eq_hd denormalize_full_nempty_if_wf by fast
lemma P_denormalize_full:
assumes "wf_dlverts t1"
and "\<forall>v \<in> dverts t1. distinct v"
and "hd (Dtree.root t1) = root"
and "max_deg t1 \<le> 1"
shows "unique_set_r root (dverts t1) (denormalize (normalize_full t1))"
using assms unique_set_r_def denormalize_full_set_eq_dverts_union
denormalize_distinct normalize_full_wfdlverts take1_eq_denormalize_full
by fastforce
lemma P_denormalize:
fixes t1 :: "('a list,'b) dtree"
assumes "wf_dlverts t1"
and "\<forall>v \<in> dverts t1. distinct v"
and "hd (Dtree.root t1) = root"
and "max_deg t1 \<le> 1"
shows "unique_set_r root (dverts t1) (denormalize t1)"
using assms P_denormalize_full by auto
lemma denormalize_full_fwd:
assumes "wf_dlverts t1"
and "max_deg t1 \<le> 1"
and "\<forall>xs \<in> (dverts t1). seq_conform xs"
and "dom_children t1 T"
shows "forward (denormalize (normalize_full t1))"
using assms denormalize_forward forward_arcs_alt seq_conform_def by auto
lemma normalize_full_verts_sublist:
"v \<in> dverts t1 \<Longrightarrow> \<exists>v2 \<in> dverts (normalize_full t1). sublist v v2"
proof(induction t1 arbitrary: v rule: normalize_full.induct)
case ind: (1 r t e)
then consider "v = r \<or> v = Dtree.root t" | "\<exists>t1 \<in> fst ` fset (sucs t). v \<in> dverts t1"
using dverts_root_or_suc by fastforce
then show ?case
proof(cases)
case 1
have "\<exists>a\<in>dverts (normalize_full (Node (r @ Dtree.root t) (sucs t))). sublist (r@Dtree.root t) a"
using ind.IH by simp
moreover have "sublist v (r@Dtree.root t)" using 1 by blast
ultimately show ?thesis using sublist_order.dual_order.trans by auto
next
case 2
then show ?thesis using ind.IH[of v] by fastforce
qed
next
case (2 xs r)
then show ?case by fastforce
qed
lemma normalize_full_sublist_preserv:
"\<lbrakk>sublist xs v; v \<in> dverts t1\<rbrakk> \<Longrightarrow> \<exists>v2 \<in> dverts (normalize_full t1). sublist xs v2"
using normalize_full_verts_sublist sublist_order.dual_order.trans by fast
lemma denormalize_full_sublist_preserv:
assumes "sublist xs v" and "v \<in> dverts t1" and "max_deg t1 \<le> 1"
shows "sublist xs (denormalize (normalize_full t1))"
proof -
obtain r where r_def[simp]: "normalize_full t1 = Node r {||}"
using assms(3) normalize_full_empty_sucs by blast
have "sublist xs r" using normalize_full_sublist_preserv[OF assms(1,2)] by simp
then show ?thesis by (simp add: denormalize_empty_sucs_simp)
qed
corollary denormalize_sublist_preserv:
"\<lbrakk>sublist xs v; v \<in> dverts (t1::('a list,'b) dtree); max_deg t1 \<le> 1\<rbrakk>
\<Longrightarrow> sublist xs (denormalize t1)"
using denormalize_full_sublist_preserv by simp
lemma Q_denormalize_full:
assumes "wf_dlverts t1"
and "\<forall>v \<in> dverts t1. distinct v"
and "hd (Dtree.root t1) = root"
and "max_deg t1 \<le> 1"
and "\<forall>xs \<in> (dverts t1). seq_conform xs"
and "dom_children t1 T"
shows "fwd_sub root (dverts t1) (denormalize (normalize_full t1))"
using P_denormalize_full[OF assms(1-4)] assms(1,4-6) denormalize_full_sublist_preserv
by (auto dest: denormalize_full_fwd simp: fwd_sub_def)
corollary Q_denormalize:
assumes "wf_dlverts t1"
and "\<forall>v \<in> dverts t1. distinct v"
and "hd (Dtree.root t1) = root"
and "max_deg t1 \<le> 1"
and "\<forall>xs \<in> (dverts t1). seq_conform xs"
and "dom_children t1 T"
shows "fwd_sub root (dverts t1) (denormalize t1)"
using Q_denormalize_full assms by simp
corollary Q_denormalize_t:
assumes "hd (Dtree.root t) = root"
and "max_deg t \<le> 1"
and "dom_children t T"
shows "fwd_sub root (dverts t) (denormalize t)"
using Q_denormalize wf_lverts assms verts_conform verts_distinct by blast
lemma P_denormalize_ikkbz_sub:
assumes "hd (Dtree.root t) = root"
shows "unique_set_r root (dverts t) (denormalize (ikkbz_sub t))"
proof -
interpret T: ranked_dtree_with_orig "ikkbz_sub t" using ikkbz_sub_ranked_dtree_orig by auto
have "\<forall>v\<in>dverts (ikkbz_sub t). distinct v" using T.verts_distinct by simp
then show ?thesis
using P_denormalize T.wf_lverts ikkbz_sub_mdeg_le1 assms ikkbz_sub_hd_root
unfolding unique_set_r_def denormalize_ikkbz_eq_dlverts dlverts_eq_dverts_union
by blast
qed
lemma merge1_sublist_preserv:
"\<lbrakk>sublist xs v; v \<in> dverts t\<rbrakk> \<Longrightarrow> \<exists>v2 \<in> dverts (merge1 t). sublist xs v2"
using sublist_order.dual_order.trans by auto
lemma normalize1_verts_sublist: "v \<in> dverts t1 \<Longrightarrow> \<exists>v2 \<in> dverts (normalize1 t1). sublist v v2"
proof(induction t1 arbitrary: v rule: normalize1.induct)
case ind: (1 r t e)
show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
consider "v = r \<or> v = Dtree.root t" | "\<exists>t1 \<in> fst ` fset (sucs t). v \<in> dverts t1"
using dverts_root_or_suc using ind.prems by fastforce
then show ?thesis
proof(cases)
case 1
then show ?thesis using True by auto
next
case 2
then show ?thesis using True by fastforce
qed
next
case False
then show ?thesis using ind by auto
qed
next
case (2 xs r)
then show ?case by fastforce
qed
lemma normalize1_sublist_preserv:
"\<lbrakk>sublist xs v; v \<in> dverts t1\<rbrakk> \<Longrightarrow> \<exists>v2 \<in> dverts (normalize1 t1). sublist xs v2"
using normalize1_verts_sublist sublist_order.dual_order.trans by fast
lemma normalize_verts_sublist: "v \<in> dverts t1 \<Longrightarrow> \<exists>v2 \<in> dverts (normalize t1). sublist v v2"
proof(induction t1 arbitrary: v rule: normalize.induct)
case (1 t1)
then show ?case
proof(cases "t1 = normalize1 t1")
case True
then show ?thesis using "1.prems" by auto
next
case False
then have eq: "normalize (normalize1 t1) = normalize t1" by (auto simp: Let_def)
then obtain v2 where v2_def: "v2 \<in> dverts (normalize1 t1)" "sublist v v2"
using normalize1_verts_sublist "1.prems" by blast
then show ?thesis
using "1.IH"[OF refl False v2_def(1)] eq sublist_order.dual_order.trans by auto
qed
qed
lemma normalize_sublist_preserv:
"\<lbrakk>sublist xs v; v \<in> dverts t1\<rbrakk> \<Longrightarrow> \<exists>v2 \<in> dverts (normalize t1). sublist xs v2"
using normalize_verts_sublist sublist_order.dual_order.trans by fast
lemma ikkbz_sub_verts_sublist: "v \<in> dverts t \<Longrightarrow> \<exists>v2 \<in> dverts (ikkbz_sub t). sublist v v2"
using ranked_dtree_with_orig_axioms proof(induction t arbitrary: v rule: ikkbz_sub.induct)
case (1 t)
then interpret T: ranked_dtree_with_orig t by simp
interpret NT: ranked_dtree_with_orig "normalize t"
using T.ranked_dtree_orig_normalize by blast
show ?case
proof(cases "max_deg t \<le> 1")
case True
then show ?thesis using "1.prems"(1) by auto
next
case False
then have 0: "\<not> (max_deg t \<le> 1 \<or> \<not> list_dtree t)" using T.list_dtree_axioms by auto
obtain v1 where v1_def: "v1 \<in> dverts (normalize t)" "sublist v v1"
using normalize_verts_sublist "1.prems"(1) by blast
then have "v1 \<in> dverts (merge1 (normalize t))" using NT.merge1_dverts_eq by blast
then obtain v2 where v2_def: "v2 \<in> dverts (ikkbz_sub t)" "sublist v1 v2"
using 1 0 T.merge1_normalize_ranked_dtree_orig by force
then show ?thesis using v1_def(2) sublist_order.dual_order.trans by blast
qed
qed
lemma ikkbz_sub_sublist_preserv:
"\<lbrakk>sublist xs v; v \<in> dverts t\<rbrakk> \<Longrightarrow> \<exists>v2 \<in> dverts (ikkbz_sub t). sublist xs v2"
using ikkbz_sub_verts_sublist sublist_order.dual_order.trans by fast
lemma denormalize_ikkbz_sub_verts_sublist:
"\<forall>xs \<in> (dverts t). sublist xs (denormalize (ikkbz_sub t))"
using ikkbz_sub_verts_sublist denormalize_sublist_preserv ikkbz_sub_mdeg_le1 by blast
lemma denormalize_ikkbz_sub_sublist_preserv:
"\<lbrakk>sublist xs v; v \<in> dverts t\<rbrakk> \<Longrightarrow> sublist xs (denormalize (ikkbz_sub t))"
using denormalize_ikkbz_sub_verts_sublist sublist_order.dual_order.trans by blast
lemma Q_denormalize_ikkbz_sub:
"\<lbrakk>hd (Dtree.root t) = root; max_deg t \<le> 1 \<Longrightarrow> dom_children t T\<rbrakk>
\<Longrightarrow> fwd_sub root (dverts t) (denormalize (ikkbz_sub t))"
using P_denormalize_ikkbz_sub ikkbz_sub_forward denormalize_ikkbz_sub_verts_sublist fwd_sub_def
by blast
subsubsection \<open>Minimal Cost of the result\<close>
lemma normalize1_dverts_app_before_contr:
"\<lbrakk>v \<in> dverts (normalize1 t); v \<notin> dverts t\<rbrakk>
\<Longrightarrow> \<exists>v1\<in>dverts t. \<exists>v2\<in>dverts t. v1 @ v2 = v \<and> before v1 v2 \<and> rank (rev v2) < rank (rev v1)"
by (fastforce dest: normalize1_dverts_contr_subtree
simp: single_subtree_root_dverts single_subtree_child_root_dverts contr_before)
lemma normalize1_dverts_app_bfr_cntr_rnks:
assumes "v \<in> dverts (normalize1 t)" and "v \<notin> dverts t"
shows "\<exists>U\<in>dverts t. \<exists>V\<in>dverts t. U @ V = v \<and> before U V \<and> rank (rev V) < rank (rev U)
\<and> (\<forall>xs \<in> dverts t. (\<exists>y\<in>set xs. \<not> (\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> xs \<noteq> U)
\<longrightarrow> rank (rev V) \<le> rank (rev xs))"
using normalize1_dverts_contr_subtree[OF assms] subtree_rank_ge_if_reach'
by (fastforce simp: single_subtree_root_dverts single_subtree_child_root_dverts contr_before)
lemma normalize1_dverts_app_bfr_cntr_rnks':
assumes "v \<in> dverts (normalize1 t)" and "v \<notin> dverts t"
shows "\<exists>U\<in>dverts t. \<exists>V\<in>dverts t. U @ V = v \<and> before U V \<and> rank (rev V) \<le> rank (rev U)
\<and> (\<forall>xs \<in> dverts t. (\<exists>y\<in>set xs. \<not> (\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> xs \<noteq> U)
\<longrightarrow> rank (rev V) \<le> rank (rev xs))"
using normalize1_dverts_contr_subtree[OF assms] subtree_rank_ge_if_reach'
by (fastforce simp: single_subtree_root_dverts single_subtree_child_root_dverts contr_before)
lemma normalize1_dverts_split:
"dverts (normalize1 t1)
= {v \<in> dverts (normalize1 t1). v \<notin> dverts t1} \<union> {v \<in> dverts (normalize1 t1). v \<in> dverts t1}"
by blast
lemma normalize1_dlverts_split:
"dlverts (normalize1 t1)
= \<Union>(set ` {v \<in> dverts (normalize1 t1). v \<notin> dverts t1})
\<union> \<Union>(set ` {v \<in> dverts (normalize1 t1). v \<in> dverts t1})"
using dlverts_eq_dverts_union by fastforce
lemma normalize1_dsjnt_in_dverts:
assumes "wf_dlverts t1"
and "v \<in> dverts t1"
and "set v \<inter> \<Union>(set ` {v \<in> dverts (normalize1 t1). v \<notin> dverts t1}) = {}"
shows "v \<in> dverts (normalize1 t1)"
proof -
have "set v \<subseteq> dlverts (normalize1 t1)" using assms(2) lverts_if_in_verts by fastforce
then have sub: "set v \<subseteq> \<Union>(set ` {v \<in> dverts (normalize1 t1). v \<in> dverts t1})"
using normalize1_dlverts_split assms(3) by auto
have "v \<noteq> []" using assms(1,2) empty_notin_wf_dlverts by auto
then obtain x where x_def: "x \<in> set v" by fastforce
then show ?thesis using dverts_same_if_set_wf[OF assms(1,2)] x_def sub by blast
qed
lemma normalize1_dsjnt_subset_split1:
fixes t1
defines "X \<equiv> {v \<in> dverts (normalize1 t1). v \<notin> dverts t1}"
assumes "wf_dlverts t1"
shows "{x. x\<in>dverts t1 \<and> set x \<inter> \<Union>(set ` X) = {}} \<subseteq> {v \<in> dverts (normalize1 t1). v \<in> dverts t1}"
using assms normalize1_dsjnt_in_dverts by blast
lemma normalize1_dsjnt_subset_split2:
fixes t1
defines "X \<equiv> {v \<in> dverts (normalize1 t1). v \<notin> dverts t1}"
assumes "wf_dlverts t1"
shows "{v \<in> dverts (normalize1 t1). v \<in> dverts t1} \<subseteq> {x. x\<in>dverts t1 \<and> set x \<inter> \<Union>(set ` X) = {}}"
using dverts_same_if_set_wf[OF wf_dlverts_normalize1] assms by blast
lemma normalize1_dsjnt_subset_eq_split:
fixes t1
defines "X \<equiv> {v \<in> dverts (normalize1 t1). v \<notin> dverts t1}"
assumes "wf_dlverts t1"
shows "{v \<in> dverts (normalize1 t1). v \<in> dverts t1} = {x. x\<in>dverts t1 \<and> set x \<inter> \<Union>(set ` X) = {}}"
using normalize1_dsjnt_subset_split1 normalize1_dsjnt_subset_split2 assms
by blast
lemma normalize1_dverts_split2:
fixes t1
defines "X \<equiv> {v \<in> dverts (normalize1 t1). v \<notin> dverts t1}"
assumes "wf_dlverts t1"
shows "X \<union> {x. x \<in> dverts t1 \<and> set x \<inter> \<Union>(set ` X) = {}} = dverts (normalize1 t1)"
unfolding assms(1) using normalize1_dsjnt_subset_eq_split[OF assms(2)] by blast
lemma set_subset_if_normalize1_vert: "v1 \<in> dverts (normalize1 t1) \<Longrightarrow> set v1 \<subseteq> dlverts t1"
using lverts_if_in_verts by fastforce
lemma normalize1_new_verts_not_reach1:
assumes "v1 \<in> dverts (normalize1 t)" and "v1 \<notin> dverts t"
and "v2 \<in> dverts (normalize1 t)" and "v2 \<notin> dverts t"
and "v1 \<noteq> v2"
shows "\<not>(\<exists>x\<in>set v1. \<exists>y\<in>set v2. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
using assms ranked_dtree_with_orig_axioms proof(induction t rule: normalize1.induct)
case (1 r t e)
then interpret R: ranked_dtree_with_orig "Node r {|(t, e)|}" by blast
show ?case
proof(cases "rank (rev (Dtree.root t)) < rank (rev r)")
case True
then have eq: "normalize1 (Node r {|(t, e)|}) = Node (r@Dtree.root t) (sucs t)" by simp
have "v1 = r @ Dtree.root t"
using "1.prems"(1,2) dverts_suc_subseteq unfolding eq by fastforce
moreover have "v2 = r @ Dtree.root t"
using "1.prems"(3,4) dverts_suc_subseteq unfolding eq by fastforce
ultimately show ?thesis using "1.prems"(5) by simp
next
case False
then show ?thesis using 1 R.ranked_dtree_orig_rec by simp
qed
next
case (2 xs r)
then interpret R: ranked_dtree_with_orig "Node r xs" by blast
have eq: "normalize1 (Node r xs) = Node r ((\<lambda>(t,e). (normalize1 t,e)) |`| xs)"
using "2.hyps" by simp
obtain t1 e1 where t1_def: "(t1,e1) \<in> fset xs" "v1 \<in> dverts (normalize1 t1)"
using "2.hyps" "2.prems"(1,2) by auto
obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "v2 \<in> dverts (normalize1 t2)"
using "2.hyps" "2.prems"(3,4) by auto
show ?case
proof(cases "t1 = t2")
case True
have "v1 \<notin> dverts t1 \<and> v2 \<notin> dverts t2"
using "2.hyps" "2.prems"(2,4) t1_def(1) t2_def(1) by simp
then show ?thesis using "2.IH" t1_def t2_def True "2.prems"(5) R.ranked_dtree_orig_rec by simp
next
case False
- have sub: "is_subtree t1 (Node r xs)" using t1_def(1) subtree_if_child by fastforce
+ have sub: "is_subtree t1 (Node r xs)" using t1_def(1) subtree_if_child[of t1 xs r] by force
have "set v1 \<subseteq> dlverts t1" using set_subset_if_normalize1_vert t1_def(2) by simp
then have reach_t1: "\<forall>x \<in> set v1. \<forall>y. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y \<longrightarrow> y \<in> dlverts t1"
using R.dlverts_reach1_in_dlverts sub by blast
have "dlverts t1 \<inter> dlverts t2 = {}"
using R.wf_lverts t2_def(1) t1_def(1) wf_dlverts.simps[of r] False by fast
then have "set v2 \<inter> dlverts t1 = {}" using set_subset_if_normalize1_vert t2_def(2) by auto
then show ?thesis using reach_t1 by blast
qed
qed
lemma normalize1_dverts_split_optimal:
defines "X \<equiv> {v \<in> dverts (normalize1 t). v \<notin> dverts t}"
assumes "\<exists>x. fwd_sub root (dverts t) x"
shows "\<exists>zs. fwd_sub root (X \<union> {x. x \<in> dverts t \<and> set x \<inter> \<Union>(set ` X) = {}}) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
proof -
let ?Y = "dverts t"
have dsjt: "\<forall>xs \<in> ?Y. \<forall>ys \<in> ?Y. xs = ys \<or> set xs \<inter> set ys = {}"
using dverts_same_if_set_wf[OF wf_lverts] by blast
have fwd: "\<forall>xs \<in> ?Y. forward xs" by (simp add: verts_forward)
have nempty: "[] \<notin> ?Y" by (simp add: empty_notin_wf_dlverts wf_lverts)
have fin: "finite ?Y" by (simp add: finite_dverts)
have "\<forall>ys \<in> X. \<exists>U \<in> ?Y. \<exists>V \<in> ?Y. U@V = ys \<and> before U V \<and> rank (rev V) \<le> rank (rev U)
\<and> (\<forall>xs \<in> ?Y. (\<exists>y\<in>set xs. \<not>(\<exists>x'\<in>set V. x' \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> (\<exists>x\<in>set U. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y) \<and> xs \<noteq> U)
\<longrightarrow> rank (rev V) \<le> rank (rev xs))"
unfolding X_def using normalize1_dverts_app_bfr_cntr_rnks' by blast
moreover have "\<forall>xs \<in> X. \<forall>ys \<in> X. xs = ys \<or> set xs \<inter> set ys = {}"
unfolding X_def using dverts_same_if_set_wf[OF wf_dlverts_normalize1] wf_lverts by blast
moreover have "\<forall>xs \<in> X. \<forall>ys \<in> X. xs = ys \<or> \<not>(\<exists>x\<in>set xs. \<exists>y\<in>set ys. x \<rightarrow>\<^sup>+\<^bsub>T\<^esub> y)"
unfolding X_def using normalize1_new_verts_not_reach1 by blast
moreover have "finite X" by (simp add: X_def finite_dverts)
ultimately show ?thesis
using combine_union_sets_optimal_cost[OF asi_rank dsjt fwd nempty fin assms(2)] by simp
qed
corollary normalize1_dverts_optimal:
assumes "\<exists>x. fwd_sub root (dverts t) x"
shows "\<exists>zs. fwd_sub root (dverts (normalize1 t)) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using normalize1_dverts_split_optimal assms normalize1_dverts_split2[OF wf_lverts] by simp
lemma normalize_dverts_optimal:
assumes "\<exists>x. fwd_sub root (dverts t) x"
shows "\<exists>zs. fwd_sub root (dverts (normalize t)) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using assms ranked_dtree_with_orig_axioms proof(induction t rule: normalize.induct)
case (1 t)
then interpret T: ranked_dtree_with_orig t by blast
obtain zs where zs_def:
"fwd_sub root (dverts (normalize1 t)) zs"
"\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as)"
using "1.prems" T.normalize1_dverts_optimal by auto
show ?case
proof(cases "t = normalize1 t")
case True
then show ?thesis using zs_def by auto
next
case False
then have eq: "normalize (normalize1 t) = normalize t" by (auto simp: Let_def)
have "\<exists>zs. fwd_sub root (dverts (normalize (normalize1 t))) zs
\<and> (\<forall>as. fwd_sub root (dverts (normalize1 t)) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using "1.IH" False zs_def(1) T.ranked_dtree_orig_normalize1 by blast
then show ?thesis using zs_def eq by force
qed
qed
lemma merge1_dverts_optimal:
assumes "\<exists>x. fwd_sub root (dverts t) x"
shows "\<exists>zs. fwd_sub root (dverts (merge1 t)) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using assms forward_UV_lists_argmin_ex by simp
theorem ikkbz_sub_dverts_optimal:
assumes "\<exists>x. fwd_sub root (dverts t) x"
shows "\<exists>zs. fwd_sub root (dverts (ikkbz_sub t)) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using assms ranked_dtree_with_orig_axioms proof(induction t rule: ikkbz_sub.induct)
case (1 t)
then interpret T: ranked_dtree_with_orig t by simp
interpret NT: ranked_dtree_with_orig "normalize t"
using T.ranked_dtree_orig_normalize by blast
show ?case
proof(cases "max_deg t \<le> 1")
case True
then show ?thesis using "1.prems"(1) forward_UV_lists_argmin_ex by auto
next
case False
then have 0: "\<not> (max_deg t \<le> 1 \<or> \<not> list_dtree t)" using T.list_dtree_axioms by auto
obtain zs where zs_def: "fwd_sub root (dverts (merge1 (normalize t))) zs"
"\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as)"
using "1.prems" T.normalize_dverts_optimal NT.merge1_dverts_eq by auto
have "\<exists>zs. fwd_sub root (dverts (ikkbz_sub (merge1 (normalize t)))) zs
\<and> (\<forall>as. fwd_sub root (dverts (merge1 (normalize t))) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using "1.IH" 0 zs_def(1) T.merge1_normalize_ranked_dtree_orig by blast
then show ?thesis using zs_def 0 by force
qed
qed
lemma ikkbz_sub_dverts_optimal':
assumes "hd (Dtree.root t) = root" and "max_deg t \<le> 1 \<Longrightarrow> dom_children t T"
shows "\<exists>zs. fwd_sub root (dverts (ikkbz_sub t)) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using ikkbz_sub_dverts_optimal Q_denormalize_ikkbz_sub assms by blast
lemma combine_strict_subtree_orig:
assumes "strict_subtree (Node r1 {|(t2,e2)|}) (Node (r@Dtree.root t1) (sucs t1))"
shows "is_subtree (Node r1 {|(t2,e2)|}) (Node r {|(t1,e1)|})"
proof -
obtain t3 where t3_def: "t3 \<in> fst ` fset (sucs t1)" "is_subtree (Node r1 {|(t2,e2)|}) t3"
using assms unfolding strict_subtree_def by force
then show ?thesis using subtree_trans subtree_if_suc[OF t3_def(1)] by auto
qed
lemma combine_subtree_orig_uneq:
assumes "is_subtree (Node r1 {|(t2,e2)|}) (Node (r@Dtree.root t1) (sucs t1))"
shows "Node r1 {|(t2,e2)|} \<noteq> Node r {|(t1,e1)|}"
proof -
have "size (Node r1 {|(t2,e2)|}) \<le> size (Node (r@Dtree.root t1) (sucs t1))"
using assms(1) subtree_size_le by blast
also have "size (Node (r@Dtree.root t1) (sucs t1)) < size (Node r {|(t1,e1)|})"
using dtree_size_skip_decr1 by fast
finally show ?thesis by blast
qed
lemma combine_strict_subtree_ranks_le:
assumes "\<And>r1 t2 e2. strict_subtree (Node r1 {|(t2,e2)|}) (Node r {|(t1,e1)|})
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "strict_subtree (Node r1 {|(t2,e2)|}) (Node (r@Dtree.root t1) (sucs t1))"
shows "rank (rev r1) \<le> rank (rev (Dtree.root t2))"
using combine_strict_subtree_orig assms unfolding strict_subtree_def
by (fast intro!: combine_subtree_orig_uneq )
lemma subtree_child_uneq:
"\<lbrakk>is_subtree t1 t2; t2 \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> t1 \<noteq> Node r xs"
using child_uneq subtree_antisym subtree_if_child by fast
lemma subtree_singleton_child_uneq:
"is_subtree t1 t2 \<Longrightarrow> t1 \<noteq> Node r {|(t2,e2)|}"
using subtree_child_uneq[of t1] by simp
lemma child_subtree_ranks_le_if_strict_subtree:
assumes "\<And>r1 t2 e2. strict_subtree (Node r1 {|(t2,e2)|}) (Node r {|(t1,e1)|})
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "is_subtree (Node r1 {|(t2,e2)|}) t1"
shows "rank (rev r1) \<le> rank (rev (Dtree.root t2))"
using assms subtree_trans subtree_singleton_child_uneq unfolding strict_subtree_def by fastforce
lemma verts_ge_child_if_sorted:
assumes "\<And>r1 t2 e2. strict_subtree (Node r1 {|(t2,e2)|}) (Node r {|(t1,e1)|})
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "max_deg (Node r {|(t1,e1)|}) \<le> 1"
and "v \<in> dverts t1"
shows "rank (rev (Dtree.root t1)) \<le> rank (rev v)"
proof -
have "\<And>r1 t2 e2. is_subtree (Node r1 {|(t2,e2)|}) t1 \<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
using child_subtree_ranks_le_if_strict_subtree[OF assms(1)] by simp
moreover have "max_deg t1 \<le> 1" using mdeg_ge_child[of t1 e1 "{|(t1,e1)|}"] assms(2) by simp
ultimately show ?thesis using rank_ge_if_mdeg_le1_dvert_nocontr assms(3) by fastforce
qed
lemma verts_ge_child_if_sorted':
assumes "\<And>r1 t2 e2. strict_subtree (Node r1 {|(t2,e2)|}) (Node r {|(t1,e1)|})
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "max_deg (Node r {|(t1,e1)|}) \<le> 1"
and "v \<in> dverts (Node r {|(t1,e1)|})"
and "v \<noteq> r"
shows "rank (rev (Dtree.root t1)) \<le> rank (rev v)"
using verts_ge_child_if_sorted[OF assms(1,2)] assms(3,4) by simp
lemma not_combined_sub_dverts_combine:
"{r@Dtree.root t1} \<union> {x. x \<in> dverts (Node r {|(t1,e1)|}) \<and> x \<noteq> r \<and> x \<noteq> Dtree.root t1}
\<subseteq> dverts (Node (r @ Dtree.root t1) (sucs t1))"
using dverts_suc_subseteq dverts_root_or_suc by fastforce
lemma dverts_combine_orig_not_combined:
assumes "wf_dlverts (Node r {|(t1,e1)|})" and "x \<in> dverts (Node (r @ Dtree.root t1) (sucs t1))" and "x \<noteq> r@Dtree.root t1"
shows "x \<in> dverts (Node r {|(t1,e1)|}) \<and> x \<noteq> r \<and> x \<noteq> Dtree.root t1"
proof -
obtain t2 where t2_def: "t2 \<in> fst ` fset (sucs t1)" "x \<in> dverts t2" using assms(2,3) by fastforce
have "set r \<inter> dlverts t2 = {}" using assms(1) suc_in_dlverts'[OF t2_def(1)] by auto
then have "x \<noteq> r" using assms(1) t2_def(2) nempty_inter_notin_dverts by auto
have "Dtree.root t1 \<noteq> []"
using assms(1) empty_notin_wf_dlverts single_subtree_child_root_dverts[OF self_subtree, of t1]
by force
moreover have "set (Dtree.root t1) \<inter> dlverts t2 = {}"
using assms(1) t2_def(1) notin_dlverts_suc_if_wf_in_root by fastforce
ultimately have "x \<noteq> Dtree.root t1" using nempty_inter_notin_dverts t2_def(2) by blast
then show ?thesis using \<open>x \<noteq> r\<close> t2_def dverts_suc_subseteq by auto
qed
lemma dverts_combine_sub_not_combined:
"wf_dlverts (Node r {|(t1,e1)|}) \<Longrightarrow> dverts (Node (r @ Dtree.root t1) (sucs t1))
\<subseteq> {r@Dtree.root t1} \<union> {x. x \<in> dverts (Node r {|(t1,e1)|}) \<and> x \<noteq> r \<and> x \<noteq> Dtree.root t1}"
using dverts_combine_orig_not_combined by fast
lemma dverts_combine_eq_not_combined:
"wf_dlverts (Node r {|(t1,e1)|}) \<Longrightarrow> dverts (Node (r @ Dtree.root t1) (sucs t1))
= {r@Dtree.root t1} \<union> {x. x \<in> dverts (Node r {|(t1,e1)|}) \<and> x \<noteq> r \<and> x \<noteq> Dtree.root t1}"
using dverts_combine_sub_not_combined not_combined_sub_dverts_combine by fast
lemma normalize_full_dverts_optimal_if_sorted:
assumes "asi rank root cost"
and "wf_dlverts t1"
and "\<forall>xs \<in> (dverts t1). distinct xs"
and "\<forall>xs \<in> (dverts t1). seq_conform xs"
and "\<And>r1 t2 e2. strict_subtree (Node r1 {|(t2,e2)|}) t1
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
and "max_deg t1 \<le> 1"
and "hd (Dtree.root t1) = root"
and "dom_children t1 T"
shows "\<exists>zs. fwd_sub root (dverts (normalize_full t1)) zs
\<and> (\<forall>as. fwd_sub root (dverts t1) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using assms proof(induction t1 rule: normalize_full.induct)
case (1 r t e)
let ?Y = "dverts (Node r {|(t,e)|})"
have dsjt: "\<forall>xs \<in> ?Y. \<forall>ys \<in> ?Y. xs = ys \<or> set xs \<inter> set ys = {}"
using dverts_same_if_set_wf[OF "1.prems"(2)] by blast
have fwd: "\<forall>xs \<in> ?Y. forward xs" using "1.prems"(4) seq_conform_alt by blast
have nempty: "[] \<notin> ?Y" using empty_notin_wf_dlverts "1.prems"(2) by blast
have fin: "finite ?Y" by (simp add: finite_dverts)
have U: "r \<in> dverts (Node r {|(t, e)|})" by simp
have V: "Dtree.root t \<in> dverts (Node r {|(t, e)|})"
using single_subtree_child_root_dverts self_subtree by fast
have ge: "\<forall>xs\<in>dverts (Node r {|(t, e)|}). xs \<noteq> r \<longrightarrow> rank (rev (Dtree.root t)) \<le> rank (rev xs)"
using verts_ge_child_if_sorted'[OF "1.prems"(5,6)] by fast
moreover have bfr: "before r (Dtree.root t)"
using before_if_dom_children_wf_conform[OF "1.prems"(8,4,2)].
moreover have Ex: "\<exists>x. fwd_sub root ?Y x" using Q_denormalize_full "1.prems"(1-8) by blast
ultimately obtain zs where zs_def:
"fwd_sub root ({r@Dtree.root t} \<union> {x. x \<in> ?Y \<and> x \<noteq> r \<and> x \<noteq> Dtree.root t}) zs"
"(\<forall>as. fwd_sub root ?Y as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using app_UV_set_optimal_cost[OF "1.prems"(1) dsjt fwd nempty fin U V] by blast
have wf: "wf_dlverts (Node (r @ Dtree.root t) (sucs t))" using "1.prems"(2) combine_wf_dlverts by fast
moreover have dst: "\<forall>v\<in>dverts (Node (r @ Dtree.root t) (sucs t)). distinct v"
using "1.prems"(2,3) combine_distinct by fast
moreover have seq: "\<forall>v\<in>dverts (Node (r @ Dtree.root t) (sucs t)). seq_conform v"
using "1.prems"(2,4,8) combine_conform by blast
moreover have rnk: "\<And>r1 t2 e2. strict_subtree (Node r1 {|(t2,e2)|}) (Node (r @ Dtree.root t) (sucs t))
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
using combine_strict_subtree_ranks_le[OF "1.prems"(5)] by simp
moreover have mdeg: "max_deg (Node (r @ Dtree.root t) (sucs t)) \<le> 1"
using "1.prems"(6) mdeg_child_sucs_le
by (fastforce dest: order_trans simp del: max_deg.simps)
moreover have hd: "hd (Dtree.root (Node (r @ Dtree.root t) (sucs t))) = root"
using "1.prems"(2,7) by simp
moreover have dom: "dom_children (Node (r @ Dtree.root t) (sucs t)) T"
using "1.prems"(8) dom_children_combine by auto
ultimately obtain xs where xs_def:
"fwd_sub root (dverts (normalize_full (Node (r @ Dtree.root t) (sucs t)))) xs"
"(\<forall>as. fwd_sub root (dverts (Node (r @ Dtree.root t) (sucs t))) as
\<longrightarrow> cost (rev xs) \<le> cost (rev as))"
using "1.IH" "1.prems"(1) by blast
then show ?case using dverts_combine_eq_not_combined[OF "1.prems"(2)] zs_def by force
next
case (2 xs r)
have Ex: "\<exists>x. fwd_sub root (dverts (Node r xs)) x"
using Q_denormalize_full "2.prems"(1-8) by blast
then show ?case using "2.hyps"(1) forward_UV_lists_argmin_ex by simp
qed
corollary normalize_full_dverts_optimal_if_sorted':
assumes "max_deg t \<le> 1"
and "hd (Dtree.root t) = root"
and "dom_children t T"
and "\<And>r1 t2 e2. strict_subtree (Node r1 {|(t2,e2)|}) t
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
shows "\<exists>zs. fwd_sub root (dverts (normalize_full t)) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using normalize_full_dverts_optimal_if_sorted asi_rank wf_lverts assms
by (blast intro: verts_distinct verts_conform)
lemma normalize_full_normalize_dverts_optimal:
assumes "max_deg t \<le> 1"
and "hd (Dtree.root t) = root"
and "dom_children t T"
shows "\<exists>zs. fwd_sub root (dverts (normalize_full (normalize t))) zs
\<and> (\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
proof -
interpret NT: ranked_dtree_with_orig "normalize t"
using ranked_dtree_orig_normalize by auto
have mdeg: "max_deg (normalize t) \<le> 1" using assms(1) normalize_mdeg_eq wf_arcs by fastforce
moreover from this have dom: "dom_children (normalize t) T"
using assms(3) dom_mdeg_le1_normalize by fastforce
moreover have hd: "hd (Dtree.root (normalize t)) = root"
using assms(2) normalize_hd_root_eq' wf_lverts by blast
moreover have "\<And>r1 t2 e2. \<lbrakk>is_subtree (Node r1 {|(t2,e2)|}) (normalize t)\<rbrakk>
\<Longrightarrow> rank (rev r1) \<le> rank (rev (Dtree.root t2))"
by (simp add: normalize_sorted_ranks)
ultimately obtain xs where xs_def: "fwd_sub root (dverts (normalize_full (normalize t))) xs"
"(\<forall>as. fwd_sub root (dverts (normalize t)) as \<longrightarrow> cost (rev xs) \<le> cost (rev as))"
using NT.normalize_full_dverts_optimal_if_sorted' strict_subtree_def by blast
obtain zs where zs_def: "fwd_sub root (dverts (normalize t)) zs"
"(\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using normalize_dverts_optimal Q_denormalize_t assms by blast
then show ?thesis using xs_def by force
qed
lemma single_set_distinct_sublist: "\<lbrakk>set ys = set x; distinct ys; sublist x ys\<rbrakk> \<Longrightarrow> x = ys"
unfolding sublist_def
by (metis DiffD2 append.assoc append.left_neutral append.right_neutral list.set_intros(1)
append_Cons distinct_set_diff neq_Nil_conv distinct_app_trans_l)
lemma denormalize_optimal_if_mdeg_le1:
assumes "max_deg t \<le> 1" and "hd (Dtree.root t) = root" and "dom_children t T"
shows "\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev (denormalize t)) \<le> cost (rev as)"
proof -
obtain zs where zs_def: "fwd_sub root (dverts (normalize_full (normalize t))) zs"
"(\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as))"
using normalize_full_normalize_dverts_optimal assms by blast
have "dverts (normalize_full (normalize t)) = {denormalize t}"
using normalize_full_normalize_dverts_eq_denormalize wf_lverts assms(1) by blast
then show ?thesis
using zs_def single_set_distinct_sublist by (auto simp: fwd_sub_def unique_set_r_def)
qed
theorem denormalize_ikkbz_sub_optimal:
assumes "hd (Dtree.root t) = root" and "max_deg t \<le> 1 \<Longrightarrow> dom_children t T"
shows "(\<forall>as. fwd_sub root (dverts t) as
\<longrightarrow> cost (rev (denormalize (ikkbz_sub t))) \<le> cost (rev as))"
proof -
obtain zs where zs_def: "fwd_sub root (dverts (ikkbz_sub t)) zs"
"\<forall>as. fwd_sub root (dverts t) as \<longrightarrow> cost (rev zs) \<le> cost (rev as)"
using ikkbz_sub_dverts_optimal' assms by blast
interpret T: ranked_dtree_with_orig "ikkbz_sub t" using ikkbz_sub_ranked_dtree_orig by simp
have "max_deg (ikkbz_sub t) \<le> 1" using ikkbz_sub_mdeg_le1 by auto
have "hd (Dtree.root (ikkbz_sub t)) = root" using assms(1) ikkbz_sub_hd_root by auto
moreover have "dom_children (ikkbz_sub t) T"
using assms(2) dom_mdeg_le1_ikkbz_sub ikkbz_sub_eq_iff_mdeg_le1 by auto
ultimately have "\<forall>as. fwd_sub root (dverts (ikkbz_sub t)) as
\<longrightarrow> cost (rev (denormalize (ikkbz_sub t))) \<le> cost (rev as)"
using T.denormalize_optimal_if_mdeg_le1[OF ikkbz_sub_mdeg_le1] by blast
then show ?thesis using zs_def order_trans by blast
qed
end
subsection \<open>Arc Invariants hold for Conversion to Dtree\<close>
context precedence_graph
begin
interpretation t: ranked_dtree to_list_dtree by (rule to_list_dtree_ranked_dtree)
lemma subtree_to_list_dtree_tree_dom:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; t \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> r \<rightarrow>\<^bsub>to_list_tree\<^esub> Dtree.root t"
unfolding to_list_dtree_def
using finite_directed_tree.subtree_child_dom to_list_tree_finite_directed_tree by fastforce
lemma subtree_to_list_dtree_dom:
assumes "is_subtree (Node r xs) to_list_dtree" and "t \<in> fst ` fset xs"
shows "hd r \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t)"
proof -
interpret T: directed_tree to_list_tree "[root]" by (rule to_list_tree_directed_tree)
have 0: "r \<rightarrow>\<^bsub>to_list_tree\<^esub> Dtree.root t" using subtree_to_list_dtree_tree_dom assms by blast
then obtain x where x_def: "r = [x] \<and> x \<in> verts T" using to_list_tree_single by force
obtain y where "Dtree.root t = [y]" using 0 to_list_tree_single T.adj_in_verts(2) by blast
then show ?thesis using 0 to_list_tree_def x_def(1) in_arcs_imp_in_arcs_ends by force
qed
lemma to_list_dtree_nempty_root: "is_subtree (Node r xs) to_list_dtree \<Longrightarrow> r \<noteq> []"
using list_dtree.list_dtree_sub list_dtree.wf_lverts to_list_dtree_list_dtree by force
lemma dom_children_aux:
assumes "is_subtree (Node r xs) to_list_dtree"
and "max_deg t1 \<le> 1"
and "(t1,e1) \<in> fset xs"
and "x \<in> dlverts t1"
shows "\<exists>v \<in> set r \<union> path_lverts t1 x. v \<rightarrow>\<^bsub>T\<^esub> x"
proof(cases "x \<in> set (Dtree.root t1)")
case True
have "Dtree.root t1 \<in> dverts to_list_dtree"
using assms(1,3) dverts_subtree_subset dtree.set_sel(1) by fastforce
then have "Dtree.root t1 = [x]" using to_list_dtree_single True by fastforce
then have 0: "hd r \<rightarrow>\<^bsub>T\<^esub> x" using subtree_to_list_dtree_dom assms(1,3) by fastforce
have "r \<in> dverts to_list_dtree" using assms(1) dverts_subtree_subset by force
then have "r = [hd r]" using to_list_dtree_single True by fastforce
then have "hd r \<in> set r" using hd_in_set[of r] by blast
then show ?thesis using 0 by blast
next
case False
obtain t2 where t2_def: "is_subtree t2 t1" "x \<in> set (Dtree.root t2)"
using assms(4) subtree_root_if_dlverts by fastforce
then obtain r1 xs1 where r1_def: "is_subtree (Node r1 xs1) t1" "t2 \<in> fst ` fset xs1"
using subtree_child_if_strict_subtree t2_def False unfolding strict_subtree_def by blast
have "is_subtree (Node r1 xs1) (Node r xs)" using r1_def(1) assms(3) by auto
then have sub_r1: "is_subtree (Node r1 xs1) to_list_dtree" using assms(1) subtree_trans by blast
have sub_t1_r: "is_subtree t1 (Node r xs)"
using subtree_if_child[of t1 xs] assms(3) by force
then have "is_subtree t2 to_list_dtree" using assms(1) subtree_trans t2_def(1) by blast
then have "Dtree.root t2 \<in> dverts to_list_dtree"
using assms(1) dverts_subtree_subset dtree.set_sel(1) by fastforce
then have "Dtree.root t2 = [x]" using to_list_dtree_single t2_def(2) by force
then have 0: "hd r1 \<rightarrow>\<^bsub>T\<^esub> x" using subtree_to_list_dtree_dom[OF sub_r1] r1_def(2) by fastforce
have sub_t1_to: "is_subtree t1 to_list_dtree" using sub_t1_r assms(1) subtree_trans by blast
then have "wf_dlverts t1" using t.wf_lverts list_dtree_def t.list_dtree_sub by blast
moreover have "max_deg t1 \<le> 1" using assms(2) sub_t1_r le_trans mdeg_ge_sub by blast
ultimately have "set r1 \<subseteq> path_lverts t1 x"
using subtree_path_lverts_sub r1_def t2_def(2) by fast
then show ?thesis
using 0 sub_r1 dverts_subtree_subset hd_in_set[of r1] to_list_dtree_single by force
qed
lemma hd_dverts_in_dlverts:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; (t1,e1) \<in> fset xs; x \<in> dverts t1\<rbrakk> \<Longrightarrow> hd x \<in> dlverts t1"
using list_dtree.list_dtree_rec list_dtree.wf_lverts hd_in_lverts_if_wf t.list_dtree_sub
by fastforce
lemma dom_children_aux2:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; max_deg t1 \<le> 1; (t1,e1) \<in> fset xs; x \<in> dverts t1\<rbrakk>
\<Longrightarrow> \<exists>v \<in> set r \<union> path_lverts t1 (hd x). v \<rightarrow>\<^bsub>T\<^esub> (hd x)"
using dom_children_aux hd_dverts_in_dlverts by blast
lemma dom_children_full:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; \<forall>t \<in> fst ` fset xs. max_deg t \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r xs) T"
unfolding dom_children_def using dom_children_aux2 by auto
lemma dom_children':
- "is_subtree (Node r xs) to_list_dtree \<Longrightarrow> dom_children (Node r (Abs_fset (children_deg1 xs))) T"
- unfolding dom_children_def using dom_children_aux2 children_deg1_fset_id by fastforce
+ assumes "is_subtree (Node r xs) to_list_dtree"
+ shows "dom_children (Node r (Abs_fset (children_deg1 xs))) T"
+ unfolding dom_children_def dtree.sel children_deg1_fset_id
+ using dom_children_aux2[OF assms(1)] by fastforce
lemma dom_children_maxdeg_1:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; max_deg (Node r xs) \<le> 1\<rbrakk>
\<Longrightarrow> dom_children (Node r xs) T"
- using dom_children_full mdeg_ge_child by fastforce
+proof (elim dom_children_full)
+ show "max_deg (Node r xs) \<le> 1 \<Longrightarrow> \<forall>t\<in>fst ` fset xs. max_deg t \<le> 1"
+ using mdeg_ge_child by fastforce
+qed
lemma dom_child_subtree:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; t \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> \<exists>v\<in>set r. v \<rightarrow>\<^bsub>T\<^esub> hd (Dtree.root t)"
using subtree_to_list_dtree_dom hd_in_set to_list_dtree_nempty_root by blast
lemma dom_children_maxdeg_1_self:
"max_deg to_list_dtree \<le> 1 \<Longrightarrow> dom_children to_list_dtree T"
using dom_children_maxdeg_1[of "Dtree.root to_list_dtree" "sucs to_list_dtree"] self_subtree by auto
lemma seq_conform_list_tree: "\<forall>v\<in>verts to_list_tree. seq_conform v"
by (simp add: to_list_tree_def seq_conform_single)
lemma conform_list_dtree: "\<forall>v\<in>dverts to_list_dtree. seq_conform v"
using seq_conform_list_tree dverts_eq_verts_to_list_tree by blast
lemma to_list_dtree_vert_single: "\<lbrakk>v \<in> dverts to_list_dtree; x \<in> set v\<rbrakk> \<Longrightarrow> v = [x] \<and> x \<in> verts T"
using to_list_dtree_single by fastforce
lemma to_list_dtree_vert_single_sub:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; x \<in> set r\<rbrakk> \<Longrightarrow> r = [x] \<and> x \<in> verts T"
using to_list_dtree_vert_single dverts_subtree_subset by fastforce
lemma to_list_dtree_child_if_to_list_tree_arc:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; r \<rightarrow>\<^bsub>to_list_tree\<^esub> v\<rbrakk> \<Longrightarrow> \<exists>ys. (Node v ys) \<in> fst ` fset xs"
using finite_directed_tree.child_if_dominated_to_dtree'[OF to_list_tree_finite_directed_tree]
unfolding to_list_dtree_def by simp
lemma to_list_dtree_child_if_arc:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; x \<in> set r; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk>
\<Longrightarrow> \<exists>ys. Node [y] ys \<in> fst ` fset xs"
using to_list_dtree_child_if_to_list_tree_arc to_list_tree_dom_iff to_list_dtree_vert_single_sub
by auto
lemma to_list_dtree_dverts_if_arc:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; x \<in> set r; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> [y] \<in> dverts (Node r xs)"
using to_list_dtree_child_if_arc[of r xs x y] by fastforce
lemma to_list_dtree_dlverts_if_arc:
"\<lbrakk>is_subtree (Node r xs) to_list_dtree; x \<in> set r; x \<rightarrow>\<^bsub>T\<^esub> y\<rbrakk> \<Longrightarrow> y \<in> dlverts (Node r xs)"
using to_list_dtree_child_if_arc[of r xs x y] by fastforce
theorem to_list_dtree_ranked_orig: "ranked_dtree_with_orig to_list_dtree rank cost cmp T root"
using dom_children' to_list_dtree_dlverts_if_arc asi_rank apply(unfold_locales)
by (auto simp: dom_children_maxdeg_1 dom_child_subtree distinct_to_list_dtree conform_list_dtree)
interpretation t: ranked_dtree_with_orig to_list_dtree by (rule to_list_dtree_ranked_orig)
lemma forward_ikkbz_sub: "forward ikkbz_sub"
using ikkbz_sub_def dom_children_maxdeg_1_self t.ikkbz_sub_forward by simp
subsection \<open>Optimality of IKKBZ-Sub\<close>
lemma ikkbz_sub_optimal_Q:
"(\<forall>as. fwd_sub root (verts to_list_tree) as \<longrightarrow> cost (rev ikkbz_sub) \<le> cost (rev as))"
using t.denormalize_ikkbz_sub_optimal to_list_dtree_hd_root_eq_root dom_children_maxdeg_1_self
unfolding dverts_eq_verts_to_list_tree ikkbz_sub_def by blast
lemma to_list_tree_sublist_if_set_eq:
assumes "set ys = \<Union>(set ` verts to_list_tree)" and "xs \<in> verts to_list_tree"
shows "sublist xs ys"
proof -
obtain x where x_def: "xs = [x]" "x \<in> verts T" using to_list_tree_single assms(2) by blast
then have "x \<in> set ys" using assms(1) to_list_tree_def by simp
then show ?thesis using x_def(1) split_list[of x ys] sublist_Cons sublist_append_leftI by fast
qed
lemma hd_eq_tk1_if_set_eq_verts: "set xs = verts T \<Longrightarrow> hd xs = root \<longleftrightarrow> take 1 xs = [root]"
using hd_eq_take1 take1_eq_hd[of xs] non_empty by fastforce
lemma ikkbz_sub_optimal:
"\<lbrakk>set xs = verts T; distinct xs; forward xs; hd xs = root\<rbrakk>
\<Longrightarrow> cost (rev ikkbz_sub) \<le> cost (rev xs)"
using ikkbz_sub_optimal_Q to_list_tree_sublist_if_set_eq
by (simp add: hd_eq_tk1_if_set_eq_verts to_list_tree_union_verts_eq fwd_sub_def unique_set_r_def)
end
subsection \<open>Optimality of IKKBZ\<close>
context ikkbz_query_graph
begin
text \<open>
Optimality only with respect to valid solutions (i.e. contain every relation exactly once).
Furthermore, only join trees without cross products are considered.
\<close>
lemma ikkbz_sub_optimal_cost_r:
"\<lbrakk>set xs = verts G; distinct xs; no_cross_products (create_ldeep xs); hd xs = r; r \<in> verts G\<rbrakk>
\<Longrightarrow> cost_r r (rev (ikkbz_sub r)) \<le> cost_r r (rev xs)"
using precedence_graph.ikkbz_sub_optimal verts_dir_tree_r_eq
by (fast intro: forward_if_ldeep_no_cross precedence_graph_r)
lemma ikkbz_sub_no_cross: "r \<in> verts G \<Longrightarrow> no_cross_products (create_ldeep (ikkbz_sub r))"
using precedence_graph.forward_ikkbz_sub ikkbz_sub_verts_eq
by (fastforce intro: no_cross_ldeep_if_forward' precedence_graph_r)
lemma ikkbz_sub_cost_r_eq_cost:
"r \<in> verts G \<Longrightarrow> cost_r r (rev (ikkbz_sub r)) = cost_l (ikkbz_sub r)"
using ikkbz_sub_verts_eq ikkbz_sub_distinct ikkbz_sub_no_cross ikkbz_sub_hd_eq_root
by (fastforce dest: cost_correct')
corollary ikkbz_sub_optimal:
"\<lbrakk>set xs = verts G; distinct xs; no_cross_products (create_ldeep xs); hd xs = r; r \<in> verts G\<rbrakk>
\<Longrightarrow> cost_l (ikkbz_sub r) \<le> cost_l xs"
using ikkbz_sub_optimal_cost_r cost_correct' ikkbz_sub_cost_r_eq_cost by fastforce
lemma ikkbz_no_cross: "no_cross_products (create_ldeep ikkbz)"
using ikkbz_eq_ikkbz_sub ikkbz_sub_no_cross by force
lemma hd_in_verts_if_set_eq: "set xs = verts G \<Longrightarrow> hd xs \<in> verts G"
using verts_nempty set_empty2[of xs] by force
lemma ikkbz_optimal:
"\<lbrakk>set xs = verts G; distinct xs; no_cross_products (create_ldeep xs)\<rbrakk>
\<Longrightarrow> cost_l ikkbz \<le> cost_l xs"
using ikkbz_min_ikkbz_sub ikkbz_sub_optimal by (fastforce intro: hd_in_verts_if_set_eq)
theorem ikkbz_optimal_tree:
"\<lbrakk>valid_tree t; no_cross_products t; left_deep t\<rbrakk> \<Longrightarrow> cost (create_ldeep ikkbz) \<le> cost t"
using ikkbz_optimal inorder_eq_set by (fastforce simp: distinct_relations_def valid_tree_def)
end
end
\ No newline at end of file
diff --git a/thys/Query_Optimization/List_Dtree.thy b/thys/Query_Optimization/List_Dtree.thy
--- a/thys/Query_Optimization/List_Dtree.thy
+++ b/thys/Query_Optimization/List_Dtree.thy
@@ -1,1035 +1,1037 @@
(* Author: Bernhard Stöckl *)
theory List_Dtree
imports Complex_Main "Graph_Additions" "Dtree"
begin
section \<open>Dtrees of Lists\<close>
subsection \<open>Functions\<close>
abbreviation remove_child :: "'a \<Rightarrow> (('a,'b) dtree \<times> 'b) fset \<Rightarrow> (('a,'b) dtree \<times> 'b) fset" where
"remove_child x xs \<equiv> ffilter (\<lambda>(t,e). root t \<noteq> x) xs"
abbreviation child2 ::
"'a \<Rightarrow> (('a,'b) dtree \<times> 'b) fset \<Rightarrow> (('a,'b) dtree \<times> 'b) fset \<Rightarrow> (('a,'b) dtree \<times> 'b) fset" where
"child2 x zs xs \<equiv> ffold (\<lambda>(t,_) b. case t of Node r ys \<Rightarrow> if r = x then ys |\<union>| b else b) zs xs"
text \<open>Combine children sets to a single set and append element to list.\<close>
fun combine :: "'a list \<Rightarrow> 'a list \<Rightarrow> ('a list,'b) dtree \<Rightarrow> ('a list,'b) dtree" where
"combine x y (Node r xs) = (if x=r \<and> (\<exists>t. t \<in> fst ` fset xs \<and> root t = y)
then Node (r@y) (child2 y (remove_child y xs) xs)
else Node r ((\<lambda>(t,e). (combine x y t,e)) |`| xs))"
text \<open>Basic @{term wf_dverts} property is not strong enough to be preserved in combine operation.\<close>
fun dlverts :: "('a list,'b) dtree \<Rightarrow> 'a set" where
"dlverts (Node r xs) = set r \<union> (\<Union>x\<in>fset xs. dlverts (fst x))"
abbreviation disjoint_dlverts :: "(('a list, 'b) dtree \<times> 'b) fset \<Rightarrow> bool" where
"disjoint_dlverts xs \<equiv>
(\<forall>(x,e1) \<in> fset xs. \<forall>(y,e2) \<in> fset xs. dlverts x \<inter> dlverts y = {} \<or> (x,e1)=(y,e2))"
fun wf_dlverts :: "('a list,'b) dtree \<Rightarrow> bool" where
"wf_dlverts (Node r xs) =
(r \<noteq> [] \<and> (\<forall>(x,e1) \<in> fset xs. set r \<inter> dlverts x = {} \<and> wf_dlverts x) \<and> disjoint_dlverts xs)"
definition wf_dlverts' :: "('a list,'b) dtree \<Rightarrow> bool" where
"wf_dlverts' t \<longleftrightarrow>
wf_dverts t \<and> [] \<notin> dverts t \<and> (\<forall>v1\<in>dverts t. \<forall>v2\<in>dverts t. set v1 \<inter> set v2 = {} \<or> v1=v2)"
fun wf_list_lverts :: "('a list\<times>'b) list \<Rightarrow> bool" where
"wf_list_lverts [] = True"
| "wf_list_lverts ((v,e)#xs) =
(v \<noteq> [] \<and> (\<forall>v2 \<in> fst ` set xs. set v \<inter> set v2 = {}) \<and> wf_list_lverts xs)"
subsection \<open>List Dtrees as Well-Formed Dtrees\<close>
lemma list_in_verts_if_lverts: "x \<in> dlverts t \<Longrightarrow> (\<exists>v \<in> dverts t. x \<in> set v)"
by(induction t) fastforce
lemma list_in_verts_iff_lverts: "x \<in> dlverts t \<longleftrightarrow> (\<exists>v \<in> dverts t. x \<in> set v)"
by(induction t) fastforce
lemma lverts_if_in_verts: "\<lbrakk>v \<in> dverts t; x \<in> set v\<rbrakk> \<Longrightarrow> x \<in> dlverts t"
by(induction t) fastforce
lemma nempty_inter_notin_dverts: "\<lbrakk>v \<noteq> []; set v \<inter> dlverts t = {}\<rbrakk> \<Longrightarrow> v \<notin> dverts t"
using lverts_if_in_verts disjoint_iff_not_equal equals0I set_empty by metis
lemma empty_notin_wf_dlverts: "wf_dlverts t \<Longrightarrow> [] \<notin> dverts t"
by(induction t) auto
lemma wf_dlverts'_rec: "\<lbrakk>wf_dlverts' (Node r xs); t1 \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> wf_dlverts' t1"
unfolding wf_dlverts'_def using wf_dverts_rec[of r xs t1] dverts_child_subseteq[of t1 xs] by blast
lemma wf_dlverts'_suc: "\<lbrakk>wf_dlverts' t; t1 \<in> fst ` fset (sucs t)\<rbrakk> \<Longrightarrow> wf_dlverts' t1"
using wf_dlverts'_rec[of "root t" "sucs t"] by simp
lemma wf_dlverts_suc: "\<lbrakk>wf_dlverts t; t1 \<in> fst ` fset (sucs t)\<rbrakk> \<Longrightarrow> wf_dlverts t1"
using wf_dlverts.simps[of "root t" "sucs t"] by auto
lemma wf_dlverts_subtree: "\<lbrakk>wf_dlverts t; is_subtree t1 t\<rbrakk> \<Longrightarrow> wf_dlverts t1"
by (induction t) auto
lemma dlverts_eq_dverts_union: "dlverts t = \<Union> (set ` dverts t)"
by (induction t) fastforce
lemma dlverts_eq_dverts_union': "dlverts t = (\<Union>x\<in> dverts t. set x)"
using dlverts_eq_dverts_union by simp
lemma dverts_nempty: "dverts t \<noteq> {}"
using dtree.set(1)[of "root t" "sucs t"] by simp
lemma dlverts_nempty_aux: "[] \<notin> dverts t \<Longrightarrow> dlverts t \<noteq> {}"
using dverts_nempty dlverts_eq_dverts_union[of t] by fastforce
lemma dlverts_nempty_if_wf: "wf_dlverts t \<Longrightarrow> dlverts t \<noteq> {}"
using dlverts_nempty_aux empty_notin_wf_dlverts by blast
lemma nempty_root_in_lverts: "root t \<noteq> [] \<Longrightarrow> hd (root t) \<in> dlverts t"
using dtree.set_sel(1) list_in_verts_iff_lverts by fastforce
lemma roothd_in_lverts_if_wf: "wf_dlverts t \<Longrightarrow> hd (root t) \<in> dlverts t"
using wf_dlverts.simps[of "root t" "sucs t"] nempty_root_in_lverts by auto
lemma hd_in_lverts_if_wf: "\<lbrakk>wf_dlverts t; v \<in> dverts t\<rbrakk> \<Longrightarrow> hd v \<in> dlverts t"
using empty_notin_wf_dlverts hd_in_set[of v] lverts_if_in_verts by fast
lemma dlverts_notin_root_sucs:
"\<lbrakk>wf_dlverts t; t1 \<in> fst ` fset (sucs t); x \<in> dlverts t1\<rbrakk> \<Longrightarrow> x \<notin> set (root t)"
using wf_dlverts.simps[of "root t" "sucs t"] by fastforce
lemma dverts_inter_empty_if_verts_inter:
assumes "dlverts x \<inter> dlverts y = {}" and "wf_dlverts x"
shows "dverts x \<inter> dverts y = {}"
proof (rule ccontr)
assume asm: "dverts x \<inter> dverts y \<noteq> {}"
then obtain r where r_def: "r \<in> dverts x" "r \<in> dverts y" by blast
then have "r \<noteq> []" using assms(2) by(auto simp: empty_notin_wf_dlverts)
then obtain v where v_def: "v \<in> set r" by fastforce
then show False using r_def assms(1) lverts_if_in_verts by (metis IntI all_not_in_conv)
qed
lemma disjoint_dlverts_if_wf: "wf_dlverts t \<Longrightarrow> disjoint_dlverts (sucs t)"
using wf_dlverts.simps[of "root t" "sucs t"] by simp
lemma disjoint_dlverts_subset:
assumes "xs |\<subseteq>| ys" and "disjoint_dlverts ys"
shows "disjoint_dlverts xs"
proof (rule ccontr)
assume "\<not> disjoint_dlverts xs"
then obtain x e1 y e2 where x_def: "(x,e1) \<in> fset xs" "(y,e2) \<in> fset xs"
"dlverts x \<inter> dlverts y \<noteq> {} \<and> (x,e1)\<noteq>(y,e2)"
by blast
have "(x,e1) \<in> fset ys" "(y,e2) \<in> fset ys" using x_def(1,2) assms(1) less_eq_fset.rep_eq by fast+
then show False using assms(2) x_def(3) by fast
qed
lemma root_empty_inter_subset:
assumes "xs |\<subseteq>| ys" and "\<forall>(x,e1) \<in> fset ys. set r \<inter> dlverts x = {}"
shows "\<forall>(x,e1) \<in> fset xs. set r \<inter> dlverts x = {}"
using assms less_eq_fset.rep_eq by force
lemma wf_dlverts_sub:
assumes "xs |\<subseteq>| ys" and "wf_dlverts (Node r ys)"
shows "wf_dlverts (Node r xs)"
proof (rule ccontr)
assume asm: "\<not>wf_dlverts (Node r xs)"
have "disjoint_dlverts xs" using assms(2) disjoint_dlverts_subset[OF assms(1)] by simp
moreover have "r \<noteq> []" using assms(2) by simp
moreover have "(\<forall>(x,e1) \<in> fset xs. set r \<inter> dlverts x = {})"
using assms(2) root_empty_inter_subset[OF assms(1)] by fastforce
ultimately obtain x e where x_def: "(x,e) \<in> fset xs" "\<not>wf_dlverts x" using asm by auto
- then have "(x,e) \<in> fset ys" using assms(1) notin_fset fin_mono by metis
+ then have "(x,e) \<in> fset ys" using assms(1) fin_mono by metis
then show False using assms(2) x_def(2) by fastforce
qed
lemma wf_dlverts_sucs: "\<lbrakk>wf_dlverts t; x \<in> fset (sucs t)\<rbrakk> \<Longrightarrow> wf_dlverts (Node (root t) {|x|})"
using wf_dlverts_sub[of "{|x|}" "sucs t" "root t"] by (simp add: less_eq_fset.rep_eq)
lemma wf_dverts_if_wf_dlverts: "wf_dlverts t \<Longrightarrow> wf_dverts t"
proof(induction t)
case (Node r xs)
then have "\<forall>(x,e) \<in> fset xs. wf_dverts x" by auto
moreover have "\<forall>(x,e) \<in> fset xs. r \<notin> dverts x"
using nempty_inter_notin_dverts Node.prems by fastforce
ultimately show ?case
using Node.prems dverts_inter_empty_if_verts_inter wf_dverts_iff_dverts'
by (smt (verit, del_insts) wf_dlverts.simps wf_dverts'.simps case_prodD case_prodI2)
qed
lemma notin_dlverts_child_if_wf_in_root:
"\<lbrakk>wf_dlverts (Node r xs); x \<in> set r; t \<in> fst ` fset xs\<rbrakk> \<Longrightarrow> x \<notin> dlverts t"
by fastforce
lemma notin_dlverts_suc_if_wf_in_root:
"\<lbrakk>wf_dlverts t1; x \<in> set (root t1); t2 \<in> fst ` fset (sucs t1)\<rbrakk> \<Longrightarrow> x \<notin> dlverts t2"
using notin_dlverts_child_if_wf_in_root[of "root t1" "sucs t1"] by simp
lemma root_if_same_lvert_wf:
"\<lbrakk>wf_dlverts (Node r xs); x \<in> set r; v \<in> dverts (Node r xs); x \<in> set v\<rbrakk> \<Longrightarrow> v = r"
by (fastforce simp: lverts_if_in_verts dverts_child_if_not_root notin_dlverts_child_if_wf_in_root)
lemma dverts_same_if_set_wf:
"\<lbrakk>wf_dlverts t; v1 \<in> dverts t; v2 \<in> dverts t; x \<in> set v1; x \<in> set v2\<rbrakk> \<Longrightarrow> v1 = v2"
proof(induction t)
case (Node r xs)
then show ?case
proof(cases "x \<in> set r")
case True
then show ?thesis using Node.prems(2,3,4,5) root_if_same_lvert_wf[OF Node.prems(1)] by blast
next
case False
then obtain t2 e2 where t2_def: "(t2,e2) \<in> fset xs" "x \<in> dlverts t2"
using Node.prems(2,4) lverts_if_in_verts by fastforce
then have "\<forall>(t3,e3)\<in>fset xs. (t3,e3) = (t2,e2) \<or> x \<notin> dlverts t3"
using Node.prems(1) by fastforce
then have "v1 \<in> dverts t2 \<and> v2 \<in> dverts t2"
using Node.prems(2-5) lverts_if_in_verts False by force
then show ?thesis using Node.IH t2_def(1) Node.prems(1,4,5) by auto
qed
qed
lemma dtree_from_list_empty_inter_iff:
"(\<forall>v \<in> fst ` set ((v, e) # xs). set r \<inter> set v = {})
\<longleftrightarrow> (\<forall>(x,e1) \<in> fset {|(dtree_from_list v xs,e)|}. set r \<inter> dlverts x = {})" (is "?P \<longleftrightarrow> ?Q")
proof
assume asm: "?P"
have "dverts (dtree_from_list v xs) = fst ` set ((v,e)#xs)"
by(simp add: dtree_from_list_eq_dverts)
then show ?Q using list_in_verts_if_lverts asm by fastforce
next
assume asm: "?Q"
have "dverts (dtree_from_list v xs) = fst ` set ((v,e)#xs)"
by(simp add: dtree_from_list_eq_dverts)
moreover have "(dtree_from_list v xs,e) \<in> fset {|(dtree_from_list v xs, e)|}" by simp
ultimately show "?P" using asm lverts_if_in_verts by fast
qed
lemma wf_dlverts_iff_wf_list_lverts:
"(\<forall>v \<in> fst ` set xs. set r \<inter> set v = {}) \<and> r \<noteq> [] \<and> wf_list_lverts xs
\<longleftrightarrow> wf_dlverts (dtree_from_list r xs)"
proof(induction xs arbitrary: r rule: wf_list_lverts.induct)
case (2 v e xs)
then show ?case using dtree_from_list_empty_inter_iff[of v e] by auto
qed (simp)
lemma vert_disjoint_if_not_root:
assumes "wf_dlverts t"
and "v \<in> dverts t - {root t}"
shows "set (root t) \<inter> set v = {}"
proof -
obtain t1 e1 where t1_def: "(t1,e1) \<in> fset (sucs t)" "v \<in> dverts t1"
using assms(2) dtree.set_cases(1) by force
then show ?thesis using assms(1) wf_dlverts.simps[of "root t"] lverts_if_in_verts by fastforce
qed
lemma vert_disjoint_if_to_list:
"\<lbrakk>wf_dlverts (Node r {|(t1,e1)|}); v \<in> fst ` set (dtree_to_list t1)\<rbrakk>
\<Longrightarrow> set (root t1) \<inter> set v = {}"
using vert_disjoint_if_not_root dtree_to_list_sub_dverts wf_dverts_if_wf_dlverts by fastforce
lemma wf_list_lverts_if_wf_dlverts: "wf_dlverts t \<Longrightarrow> wf_list_lverts (dtree_to_list t)"
proof(induction t)
case (Node r xs)
then show ?case
proof(cases "\<forall>x. xs \<noteq> {|x|}")
case True
then show ?thesis using dtree_to_list.simps(2) by simp
next
case False
then obtain t1 e1 where t1_def: "xs = {|(t1,e1)|}" by auto
then have "wf_dlverts t1" using Node.prems by simp
then have "root t1 \<noteq> []" using wf_dlverts.simps[of "root t1" "sucs t1"] by simp
then show ?thesis using Node vert_disjoint_if_to_list t1_def by fastforce
qed
qed
lemma child_in_dlverts: "(t1,e) \<in> fset xs \<Longrightarrow> dlverts t1 \<subseteq> dlverts (Node r xs)"
by force
lemma suc_in_dlverts: "(t1,e) \<in> fset (sucs t2) \<Longrightarrow> dlverts t1 \<subseteq> dlverts t2"
using child_in_dlverts[of t1 e "sucs t2" "root t2"] by auto
lemma suc_in_dlverts': "t1 \<in> fst ` fset (sucs t2) \<Longrightarrow> dlverts t1 \<subseteq> dlverts t2"
using suc_in_dlverts by fastforce
lemma subtree_in_dlverts: "is_subtree t1 t2 \<Longrightarrow> dlverts t1 \<subseteq> dlverts t2"
by(induction t2) fastforce
lemma subtree_root_if_dlverts: "x \<in> dlverts t \<Longrightarrow> \<exists>r xs. is_subtree (Node r xs) t \<and> x \<in> set r"
using subtree_root_if_dverts list_in_verts_if_lverts by fast
lemma x_not_root_strict_subtree:
assumes "x \<in> dlverts t" and "x \<notin> set (root t)"
shows "\<exists>r xs t1. is_subtree (Node r xs) t \<and> t1 \<in> fst ` fset xs \<and> x \<in> set (root t1)"
proof -
obtain r xs where r_def: "is_subtree (Node r xs) t" "x \<in> set r"
using subtree_root_if_dlverts[OF assms(1)] by fast
then have sub: "strict_subtree (Node r xs) t" using assms(2) strict_subtree_def by fastforce
then show ?thesis using assms(2) subtree_child_if_strict_subtree[OF sub] r_def(2) by force
qed
lemma dverts_disj_if_wf_dlverts:
"\<lbrakk>wf_dlverts t; v1 \<in> dverts t; v2 \<in> dverts t; v1 \<noteq> v2\<rbrakk> \<Longrightarrow> set v1 \<inter> set v2 = {}"
using dverts_same_if_set_wf by fast
thm empty_notin_wf_dlverts
lemma wf_dlverts'_if_dlverts: "wf_dlverts t \<Longrightarrow> wf_dlverts' t"
using wf_dlverts'_def empty_notin_wf_dlverts dverts_disj_if_wf_dlverts wf_dverts_if_wf_dlverts
by blast
lemma disjoint_dlverts_if_wf'_aux:
assumes "wf_dlverts' (Node r xs)"
and "(t1,e1) \<in> fset xs"
and "(t2,e2) \<in> fset xs"
and "(t1,e1) \<noteq> (t2,e2)"
shows "dlverts t1 \<inter> dlverts t2 = {}"
proof(rule ccontr)
assume "dlverts t1 \<inter> dlverts t2 \<noteq> {}"
then obtain x y where x_def: "x \<in> dverts t1" "y \<in> dverts t2" "set x \<inter> set y \<noteq> {}"
using dlverts_eq_dverts_union[of t1] dlverts_eq_dverts_union[of t2] by auto
then have "x \<in> dverts (Node r xs)" "y \<in> dverts (Node r xs)"
using dverts_child_subseteq assms(2,3) by auto
moreover have "x \<noteq> y"
using assms(1) disjoint_dverts_if_wf_aux[rotated, OF assms(2-4)] x_def(1,2)
unfolding wf_dlverts'_def by blast
ultimately show False using assms(1) x_def(3) unfolding wf_dlverts'_def by blast
qed
lemma disjoint_dlverts_if_wf': "wf_dlverts' (Node r xs) \<Longrightarrow> disjoint_dlverts xs"
using disjoint_dlverts_if_wf'_aux by fast
lemma root_nempty_if_wf': "wf_dlverts' (Node r xs) \<Longrightarrow> r \<noteq> []"
unfolding wf_dlverts'_def by fastforce
lemma disjoint_root_if_wf'_aux:
assumes "wf_dlverts' (Node r xs)"
and "(t1,e1) \<in> fset xs"
shows "set r \<inter> dlverts t1 = {}"
proof(rule ccontr)
assume "set r \<inter> dlverts t1 \<noteq> {}"
then obtain x where x_def: "x \<in> dverts t1" "set x \<inter> set r \<noteq> {}"
using dlverts_eq_dverts_union by fast
then have "x \<in> dverts (Node r xs)" using dverts_child_subseteq assms(2) by auto
moreover have "r \<in> dverts (Node r xs)" by simp
moreover have "x \<noteq> r"
using assms x_def(1) root_not_child_if_wf_dverts unfolding wf_dlverts'_def by fast
ultimately show False using assms(1) x_def(2) unfolding wf_dlverts'_def by blast
qed
lemma disjoint_root_if_wf':
"wf_dlverts' (Node r xs) \<Longrightarrow> \<forall>(t1,e1) \<in> fset xs. set r \<inter> dlverts t1 = {}"
using disjoint_root_if_wf'_aux by fast
lemma wf_dlverts_if_dlverts': "wf_dlverts' t \<Longrightarrow> wf_dlverts t"
proof(induction t)
case (Node r xs)
then have "\<forall>(t1,e1) \<in> fset xs. set r \<inter> dlverts t1 = {}"
using disjoint_root_if_wf' by blast
moreover have "r \<noteq> [] \<and> disjoint_dlverts xs"
using disjoint_dlverts_if_wf' Node.prems root_nempty_if_wf' by fast
moreover have "\<forall>(t1,e1) \<in> fset xs. wf_dlverts t1"
using Node wf_dlverts'_rec by fastforce
ultimately show ?case by auto
qed
lemma wf_dlverts_iff_dlverts': "wf_dlverts t \<longleftrightarrow> wf_dlverts' t"
using wf_dlverts_if_dlverts' wf_dlverts'_if_dlverts by blast
locale list_dtree =
fixes t :: "('a list,'b) dtree"
assumes wf_arcs: "wf_darcs t"
and wf_lverts: "wf_dlverts t"
sublocale list_dtree \<subseteq> wf_dtree
using wf_arcs wf_lverts wf_dverts_if_wf_dlverts by(unfold_locales) auto
theorem list_dtree_iff_wf_list:
"wf_list_arcs xs \<and> (\<forall>v \<in> fst ` set xs. set r \<inter> set v = {}) \<and> r \<noteq> [] \<and> wf_list_lverts xs
\<longleftrightarrow> list_dtree (dtree_from_list r xs)"
using wf_darcs_iff_wf_list_arcs wf_dlverts_iff_wf_list_lverts list_dtree_def by metis
lemma list_dtree_subset:
assumes "xs |\<subseteq>| ys" and "list_dtree (Node r ys)"
shows "list_dtree (Node r xs)"
using wf_dlverts_sub[OF assms(1)] wf_darcs_sub[OF assms(1)] assms(2)
by (unfold_locales) (fast dest: list_dtree.wf_lverts list_dtree.wf_arcs)+
context fin_list_directed_tree
begin
lemma dlverts_disjoint:
assumes "r \<in> verts T" and "(Node r xs) = to_dtree_aux r"
and "(x,e1) \<in> fset xs" and "(y,e2) \<in> fset xs" and "(x,e1)\<noteq>(y,e2)"
shows "dlverts x \<inter> dlverts y = {}"
proof (rule ccontr)
assume "dlverts x \<inter> dlverts y \<noteq> {}"
then obtain v where v_def[simp]: "v \<in> dlverts x" "v \<in> dlverts y" by blast
obtain x1 where x1_def: "v \<in> set x1" "x1 \<in> dverts x" using list_in_verts_if_lverts by force
obtain y1 where y1_def: "v \<in> set y1" "y1 \<in> dverts y" using list_in_verts_if_lverts by force
have 0: "y = to_dtree_aux (Dtree.root y)" using to_dtree_aux_self assms(2,4) by blast
have "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root y"
using assms(2,4) dominated_if_child by (metis (no_types, opaque_lifting) fst_conv image_iff)
then have 1: "Dtree.root y \<in> verts T" using adj_in_verts(2) by simp
have "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root x"
using assms(2,3) dominated_if_child by (metis (no_types, opaque_lifting) fst_conv image_iff)
then have "Dtree.root x \<in> verts T" using adj_in_verts(2) by simp
moreover have "x = to_dtree_aux (Dtree.root x)" using to_dtree_aux_self assms(2,3) by blast
ultimately have "Dtree.root x \<rightarrow>\<^sup>*\<^bsub>T\<^esub> x1" using to_dtree_aux_dverts_reachable x1_def(2) by blast
moreover have "Dtree.root y \<rightarrow>\<^sup>*\<^bsub>T\<^esub> y1" using 0 1 to_dtree_aux_dverts_reachable y1_def(2) by blast
ultimately have "x1 = y1" using disjoint_verts reachable_in_verts(2) x1_def(1) y1_def(1) by auto
then show False using dverts_disjoint[OF assms(2-5)] x1_def(2) y1_def(2) by blast
qed
lemma wf_dlverts_to_dtree_aux: "\<lbrakk>r \<in> verts T; t = to_dtree_aux r\<rbrakk> \<Longrightarrow> wf_dlverts t"
proof(induction t arbitrary: r rule: darcs_mset.induct)
case (1 r' xs)
then have "r = r'" by simp
have "\<forall>(x,e) \<in> fset xs. wf_dlverts x \<and> set r \<inter> dlverts x = {}"
proof (standard, standard, standard)
fix xp x e
assume asm: "xp \<in> fset xs" "xp = (x,e)"
then have 0: "x = to_dtree_aux (Dtree.root x)" using to_dtree_aux_self "1.prems"(2) by simp
have 2: "r \<rightarrow>\<^bsub>T\<^esub> Dtree.root x" using asm "1.prems" \<open>r = r'\<close>
by (metis (no_types, opaque_lifting) dominated_if_child fst_conv image_iff)
then have 3: "Dtree.root x \<in> verts T" using adj_in_verts(2) by simp
then show "wf_dlverts x" using "1.IH" asm 0 by blast
have "r \<notin> dverts x"
proof
assume "r \<in> dverts x"
then have "Dtree.root x \<rightarrow>\<^sup>*\<^bsub>T\<^esub> r" using 0 3 to_dtree_aux_dverts_reachable by blast
then have "r \<rightarrow>\<^sup>+\<^bsub>T\<^esub> r" using 2 by auto
then show False using reachable1_not_reverse by blast
qed
then show "set r \<inter> dlverts x = {}"
using 0 "1.prems"(1) 3 disjoint_iff_not_equal disjoint_verts list_in_verts_if_lverts
by (metis reachable_in_verts(2) to_dtree_aux_dverts_reachable)
qed
moreover have "disjoint_dlverts xs" using dlverts_disjoint "1.prems" by fastforce
ultimately show ?case using \<open>r = r'\<close> by (auto simp add: "1.prems"(1) nempty_verts)
qed
lemma wf_dlverts_to_dtree: "wf_dlverts to_dtree"
using to_dtree_def wf_dlverts_to_dtree_aux root_in_T by blast
theorem list_dtree_to_dtree: "list_dtree to_dtree"
using list_dtree_def wf_dlverts_to_dtree wf_darcs_to_dtree by blast
end
context list_dtree
begin
lemma list_dtree_rec: "\<lbrakk>Node r xs = t; (x,e) \<in> fset xs\<rbrakk> \<Longrightarrow> list_dtree x"
using wf_arcs wf_lverts by(unfold_locales) auto
lemma list_dtree_rec_suc: "(x,e) \<in> fset (sucs t) \<Longrightarrow> list_dtree x"
using list_dtree_rec[of "root t"] by force
lemma list_dtree_sub: "is_subtree x t \<Longrightarrow> list_dtree x"
using list_dtree_axioms proof(induction t rule: darcs_mset.induct)
case (1 r xs)
then interpret list_dtree "Node r xs" by blast
show ?case
proof(cases "x = Node r xs")
case True
then show ?thesis by (simp add: "1.prems")
next
case False
then show ?thesis using "1.IH" list_dtree_rec "1.prems"(1) by auto
qed
qed
theorem from_dtree_fin_list_dir: "fin_list_directed_tree (root t) (from_dtree dt dh t)"
unfolding fin_list_directed_tree_def fin_list_directed_tree_axioms_def
by (auto simp: from_dtree_fin_directed empty_notin_wf_dlverts[OF wf_lverts]
intro: wf_lverts dverts_same_if_set_wf)
subsection \<open>Combining Preserves Well-Formedness\<close>
lemma remove_child_sub: "remove_child x xs |\<subseteq>| xs"
by auto
lemma child2_commute_aux:
assumes "f = (\<lambda>(t,_) b. case t of Node r ys \<Rightarrow> if r = a then ys |\<union>| b else b)"
shows "(f y \<circ> f x) z = (f x \<circ> f y) z"
proof -
obtain r1 ys1 e1 where y_def: "y = (Node r1 ys1, e1)" by (metis dtree.exhaust eq_snd_iff)
obtain r2 ys2 e2 where "x = (Node r2 ys2, e2)" by (metis dtree.exhaust eq_snd_iff)
then show ?thesis by (simp add: assms funion_left_commute y_def)
qed
lemma child2_commute:
"comp_fun_commute (\<lambda>(t,_) b. case t of Node r ys \<Rightarrow> if r = x then ys |\<union>| b else b)"
using comp_fun_commute_def child2_commute_aux by fastforce
interpretation Comm:
comp_fun_commute "\<lambda>(t,_) b. case t of Node r ys \<Rightarrow> if r = x then ys |\<union>| b else b"
by (rule child2_commute)
lemma input_in_child2:
"zs |\<subseteq>| child2 x zs ys"
proof(induction ys)
case empty
then show ?case using Comm.ffold_empty by simp
next
case (insert y ys)
then obtain r xs e where r_def: "(Node r xs,e) = y" by (metis dtree.exhaust surj_pair)
let ?f = "(\<lambda>(t,_) b. case t of Node r ys \<Rightarrow> if r = x then ys |\<union>| b else b)"
show ?case
proof(cases "r=x")
case True
then have "ffold ?f zs (finsert y ys) = xs |\<union>| (ffold ?f zs ys)"
using r_def insert.hyps by force
then show ?thesis using insert.IH by blast
next
case False
then have "ffold ?f zs (finsert y ys) = (ffold ?f zs ys)" using r_def insert.hyps by force
then show ?thesis using insert.IH by blast
qed
qed
lemma child2_subset_if_input1:
"zs' |\<subseteq>| zs \<Longrightarrow> child2 x zs' ys |\<subseteq>| child2 x zs ys"
proof(induction ys)
case (insert y ys)
obtain r xs e where r_def: "(Node r xs, e) = y" by (metis dtree.exhaust surj_pair)
let ?f = "(\<lambda>(t,_) b. case t of Node r ys \<Rightarrow> if r = x then ys |\<union>| b else b)"
show ?case
proof(cases "r=x")
case True
then have "ffold ?f zs (finsert y ys) = xs |\<union>| (ffold ?f zs ys)"
using r_def insert.hyps by force
moreover have "ffold ?f zs' (finsert y ys) = xs |\<union>| (ffold ?f zs' ys)"
using r_def insert.hyps True by force
ultimately show ?thesis using insert by blast
next
case False
then have "ffold ?f zs (finsert y ys) = (ffold ?f zs ys)" using r_def insert.hyps by force
moreover have "ffold ?f zs' (finsert y ys) = (ffold ?f zs' ys)"
using r_def insert.hyps False by force
ultimately show ?thesis using insert by blast
qed
qed (simp)
lemma child2_subset_if_input2:
"ys' |\<subseteq>| ys \<Longrightarrow> child2 x xs ys' |\<subseteq>| child2 x xs ys"
proof(induction "fcard ys" arbitrary: ys)
case (Suc n)
show ?case
proof(cases "ys' = ys")
case False
then obtain z where z_def: "z |\<in>| ys \<and> z |\<notin>| ys'" using Suc.prems by blast
then obtain zs where zs_def: "finsert z zs = ys \<and> z |\<notin>| zs" by blast
then have "ys' |\<subseteq>| zs \<and> fcard zs = n"
using Suc.prems(1) Suc.hyps(2) z_def fcard_finsert_disjoint by fastforce
then have 0: "child2 x xs ys' |\<subseteq>| child2 x xs zs" using Suc.hyps(1) by blast
obtain r rs e where r_def: "(Node r rs, e) = z" by (metis dtree.exhaust surj_pair)
then show ?thesis using 0 zs_def by force
qed (simp)
qed (simp)
lemma darcs_split: "darcs (Node r (xs|\<union>|ys)) = darcs (Node r xs) \<union> darcs (Node r ys)"
by simp
lemma darcs_sub_if_children_sub: "xs |\<subseteq>| ys \<Longrightarrow> darcs (Node r xs) \<subseteq> darcs (Node v ys)"
proof(induction "fcard ys" arbitrary: ys)
case (Suc n)
then show ?case
proof(cases "ys = xs")
case False
then obtain z where z_def: "z |\<in>| ys \<and> z |\<notin>| xs" using Suc.prems by blast
then obtain zs where zs_def: "finsert z zs = ys \<and> z |\<notin>| zs" by blast
then have "xs |\<subseteq>| zs \<and> fcard zs = n"
using Suc.prems(1) Suc.hyps(2) z_def fcard_finsert_disjoint by fastforce
then have "darcs (Node r xs) \<subseteq> darcs (Node v zs)" using Suc.hyps(1) by blast
then show ?thesis using zs_def darcs_split[of v "{|z|}" zs] by auto
qed (simp)
qed (simp)
lemma darc_in_child2_snd_if_nin_fst:
"e \<in> darcs (Node x (child2 a xs ys)) \<Longrightarrow> e \<notin> darcs (Node v ys) \<Longrightarrow> e \<in> darcs (Node r xs)"
proof(induction "ys")
case (insert y ys)
obtain r rs e1 where r_def: "(Node r rs, e1) = y" by (metis dtree.exhaust surj_pair)
then have e_not_rs: "e \<notin> darcs (Node x rs)" using insert.prems(2) by fastforce
show ?case
proof(cases "r = a")
case True
then have "darcs (Node x (child2 a xs (finsert y ys)))
= darcs (Node x (rs |\<union>| (child2 a xs ys)))"
using r_def insert.hyps(1) by force
moreover have "\<dots> = darcs (Node x rs) \<union> darcs (Node x (child2 a xs ys))" by simp
ultimately have "e \<in> darcs (Node x (child2 a xs ys))" using insert.prems(1) e_not_rs by blast
then show ?thesis using insert.IH insert.prems(2) by simp
next
case False
then have "darcs (Node x (child2 a xs (finsert y ys))) = darcs (Node x (child2 a xs ys))"
using r_def insert.hyps(1) by force
then show ?thesis using insert.IH insert.prems by simp
qed
qed (simp)
lemma darc_in_child2_fst_if_nin_snd:
"e \<in> darcs (Node x (child2 a xs ys)) \<Longrightarrow> e \<notin> darcs (Node v xs) \<Longrightarrow> e \<in> darcs (Node r ys)"
using darc_in_child2_snd_if_nin_fst by fast
lemma darcs_child2_sub: "darcs (Node x (child2 y xs ys)) \<subseteq> darcs (Node r xs) \<union> darcs (Node r' ys)"
using darc_in_child2_snd_if_nin_fst by fast
lemma darcs_combine_sub_orig: "darcs (combine x y t1) \<subseteq> darcs t1"
proof(induction t1)
case ind: (Node r xs)
show ?case
proof(cases "x=r \<and> (\<exists>t. t \<in> fst ` fset xs \<and> root t = y)")
case True
then have "darcs (combine x y (Node r xs))
= darcs (Node (x@y) (child2 y (remove_child y xs) xs))" by simp
also have "\<dots> \<subseteq> darcs (Node x (child2 y xs xs))"
using darcs_sub_if_children_sub[of "child2 y (remove_child y xs) xs" "child2 y xs xs"]
child2_subset_if_input1[of "remove_child y xs" xs] remove_child_sub by fast
finally show ?thesis using darcs_child2_sub by fast
next
case False
then have "darcs (combine x y (Node r xs))
= darcs (Node r ((\<lambda>(t,e). (combine x y t,e)) |`| xs))"
by auto
also have "\<dots> \<subseteq> (\<Union>(t,e)\<in>fset xs. \<Union> (darcs ` {t}) \<union> {e})"
using ind.IH wf_dtree_rec by fastforce
finally show ?thesis by force
qed
qed
lemma child2_in_child:
"\<lbrakk>b \<in> fset (child2 a ys xs); b |\<notin>| ys\<rbrakk> \<Longrightarrow> \<exists>rs e. (Node a rs, e) \<in> fset xs \<and> b |\<in>| rs"
proof(induction xs)
case (insert x xs)
obtain r rs e1 where r_def: "(Node r rs, e1) = x" by (metis dtree.exhaust surj_pair)
show ?case
proof(cases "r = a")
case ra: True
then have 0: "child2 a ys (finsert x xs) = rs |\<union>| (child2 a ys xs)"
using r_def insert.hyps(1) by force
show ?thesis
proof(cases "b |\<in>| rs")
case True
then show ?thesis using r_def ra by auto
next
case False
- then have "b \<in> fset (child2 a ys xs)" using insert.prems(1) 0 notin_fset by force
+ then have "b \<in> fset (child2 a ys xs)" using insert.prems(1) 0 by force
then show ?thesis using insert.IH insert.prems(2) by auto
qed
next
case False
then show ?thesis using insert r_def by force
qed
-qed (simp add: notin_fset)
+qed simp
lemma child_in_darcs: "(y,e2) \<in> fset xs \<Longrightarrow> darcs y \<union> {e2} \<subseteq> darcs (Node r xs)"
by force
lemma disjoint_darcs_child2:
assumes "wf_darcs (Node r xs)"
shows "disjoint_darcs (child2 a (remove_child a xs) xs)" (is "disjoint_darcs ?P")
proof (rule ccontr)
assume "\<not> disjoint_darcs ?P"
then obtain x e1 y e2 where asm: "(x,e1) \<in> fset ?P" "(y,e2) \<in> fset ?P" "(e1 \<in> darcs x \<or>
((darcs x \<union> {e1}) \<inter> (darcs y \<union> {e2}) \<noteq> {} \<and> (x,e1)\<noteq>(y,e2)))" by blast
note wf_darcs_iff_darcs'[simp]
consider "(x,e1) \<in> fset (remove_child a xs)" "e1 \<in> darcs x"
| "(x,e1) \<in> fset (remove_child a xs)" "e1 \<notin> darcs x" "(y,e2) \<in> fset (remove_child a xs)"
| "(x,e1) \<in> fset (remove_child a xs)" "e1 \<notin> darcs x" "(y,e2) |\<notin>| (remove_child a xs)"
| "(x,e1) |\<notin>| (remove_child a xs)" "e1 \<in> darcs x"
| "(x,e1) |\<notin>| (remove_child a xs)" "e1 \<notin> darcs x" "(y,e2) \<in> fset (remove_child a xs)"
| "(x,e1) |\<notin>| (remove_child a xs)" "e1 \<notin> darcs x" "(y,e2) |\<notin>| (remove_child a xs)"
- by (auto simp: notin_fset)
+ by auto
then show False
proof(cases)
case 1
then show ?thesis using assms by auto
next
case 2
then show ?thesis using assms asm(3) by fastforce
next
case 3
then have x_xs: "(x,e1) \<in> fset xs" by simp
obtain rs2 re2 where r2_def: "(Node a rs2, re2) \<in> fset xs" "(y,e2) |\<in>| rs2"
using child2_in_child asm(2) 3(3) by fast
- then have "darcs y \<union> {e2} \<subseteq> darcs (Node a rs2)" using child_in_darcs notin_fset by fast
+ then have "darcs y \<union> {e2} \<subseteq> darcs (Node a rs2)" using child_in_darcs by fast
then have "(darcs x \<union> {e1}) \<inter> (darcs (Node a rs2) \<union> {re2}) \<noteq> {}" using 3(2) asm(3) by blast
moreover have "(x,e1)\<noteq>(Node a rs2, re2)" using 3(1) by force
ultimately have "\<not> disjoint_darcs xs" using r2_def(1) x_xs by fast
then show ?thesis using assms by simp
next
case 4
then obtain rs1 re1 where r1_def: "(Node a rs1, re1) \<in> fset xs" "(x,e1) |\<in>| rs1"
using child2_in_child asm(1) by fast
- then have "\<not>disjoint_darcs rs1" using notin_fset 4(2) by fast
+ then have "\<not>disjoint_darcs rs1" using 4(2) by fast
then show ?thesis using assms r1_def(1) by fastforce
next
case 5
then obtain rs1 re1 where r1_def: "(Node a rs1, re1) \<in> fset xs" "(x,e1) |\<in>| rs1"
using child2_in_child asm(1) by fast
have 1: "(darcs (Node a rs1) \<union> {re1}) \<inter> (darcs y \<union> {e2}) \<noteq> {}"
- using r1_def(2) asm(3) 5(2) child_in_darcs notin_fset by fast
+ using r1_def(2) asm(3) 5(2) child_in_darcs by fast
have y_xs: "(y,e2) \<in> fset xs" using 5(3) by simp
then have "(Node a rs1, re1)\<noteq>(y,e2)" using 5(3) by force
then have "\<not> disjoint_darcs xs" using r1_def(1) y_xs 1 by fast
then show ?thesis using assms by simp
next
case 6
then obtain rs1 re1 where r1_def: "(Node a rs1, re1) \<in> fset xs" "(x,e1) |\<in>| rs1"
using child2_in_child asm(1) by fast
then have 1: "(darcs (Node a rs1) \<union> {re1}) \<inter> (darcs y \<union> {e2}) \<noteq> {}"
- using asm(3) 6(2) child_in_darcs notin_fset by fast
+ using asm(3) 6(2) child_in_darcs by fast
obtain rs2 re2 where r2_def: "(Node a rs2, re2) \<in> fset xs" "(y,e2) |\<in>| rs2"
using child2_in_child asm(2) 6(3) by fast
- then have "darcs y \<union> {e2} \<subseteq> darcs (Node a rs2)" using child_in_darcs notin_fset by fast
+ then have "darcs y \<union> {e2} \<subseteq> darcs (Node a rs2)" using child_in_darcs by fast
then have 1: "(darcs (Node a rs1) \<union> {re1}) \<inter> (darcs (Node a rs2) \<union> {re2}) \<noteq> {}"
- using 1 asm(3) 6(2) child_in_darcs notin_fset by blast
+ using 1 asm(3) 6(2) child_in_darcs by blast
then show ?thesis
proof(cases "(Node a rs1, re1) = (Node a rs2, re2)")
case True
then have "(x,e1) \<in> fset rs1 \<and> (y,e2) \<in> fset rs1"
- using r1_def(2) r2_def(2) notin_fset by fast
+ using r1_def(2) r2_def(2) by fast
then show ?thesis using assms r1_def asm(3) 6(2) by fastforce
next
case False
then have "\<not> disjoint_darcs xs" using r1_def(1) r2_def(1) 1 by fast
then show ?thesis using assms by simp
qed
qed
qed
lemma wf_darcs_child2:
assumes "wf_darcs (Node r xs)" and "(x,e) \<in> fset (child2 a (remove_child a xs) xs)"
shows "wf_darcs x"
proof(cases "(x,e) |\<in>| remove_child a xs")
case True
- then show ?thesis using assms(1) notin_fset by (fastforce simp: wf_darcs_iff_darcs')
+ then show ?thesis using assms(1) by (fastforce simp: wf_darcs_iff_darcs')
next
case False
then obtain r rs e1 where "(Node r rs, e1) \<in> fset xs \<and> (x,e) |\<in>| rs \<and> r = a"
using child2_in_child assms(2) by fast
- then show ?thesis using assms notin_fset by (fastforce simp: wf_darcs_iff_darcs')
+ then show ?thesis using assms by (fastforce simp: wf_darcs_iff_darcs')
qed
lemma disjoint_darcs_combine:
assumes "Node r xs = t"
shows "disjoint_darcs ((\<lambda>(t,e). (combine x y t,e)) |`| xs)"
proof -
have "disjoint_darcs xs" using wf_arcs assms by (fastforce simp: wf_darcs_iff_darcs')
then show ?thesis
using disjoint_darcs_img[of xs "combine x y"] by (simp add: darcs_combine_sub_orig)
qed
lemma wf_darcs_combine: "wf_darcs (combine x y t)"
using list_dtree_axioms proof(induction t)
case ind: (Node r xs)
then interpret list_dtree "Node r xs" using ind.prems by blast
show ?case
proof(cases "x=r \<and> (\<exists>t. t \<in> fst ` fset xs \<and> root t = y)")
case True
have "disjoint_darcs (child2 y (remove_child y xs) xs)"
using disjoint_darcs_child2[OF wf_arcs] by simp
moreover have "\<forall>(x,e) \<in> fset (child2 y (remove_child y xs) xs). wf_darcs x"
using wf_darcs_child2 wf_arcs by fast
ultimately show ?thesis using True by (simp add: wf_darcs_iff_darcs')
next
case False
have "disjoint_darcs ((\<lambda>(t,e). (combine x y t, e)) |`| xs)"
using disjoint_darcs_combine ind.prems by simp
moreover have "\<forall>(x,e) \<in> fset xs. list_dtree x" using list_dtree_rec by blast
ultimately show ?thesis using False ind.IH ind.prems by (auto simp: wf_darcs_iff_darcs')
qed
qed
lemma v_in_dlverts_if_in_comb: "v \<in> dlverts (combine x y t) \<Longrightarrow> v \<in> dlverts t"
using list_dtree_axioms proof(induction t)
case ind: (Node r xs)
then interpret list_dtree "Node r xs" using ind.prems by blast
show ?case
proof(cases "x=r \<and> (\<exists>t. t \<in> fst ` fset xs \<and> root t = y)")
case x_and_y: True
show ?thesis
proof(cases "v \<in> set x \<union> set y")
case True
then show ?thesis using x_and_y dtree.set_sel(1) lverts_if_in_verts by fastforce
next
case False
then obtain t e where t_def: "(t,e) \<in> fset (child2 y (remove_child y xs) xs)" "v \<in> dlverts t"
using x_and_y ind.prems by auto
then show ?thesis
proof(cases "(t,e) |\<in>| (remove_child y xs)")
case True
- then have "(t,e) \<in> fset (remove_child y xs)" using notin_fset by fast
+ then have "(t,e) \<in> fset (remove_child y xs)" by fast
then show ?thesis using t_def(2) by force
next
case False
then obtain r1 rs1 re1 where r1_def: "(Node r1 rs1, re1) \<in> fset xs" "(t,e) |\<in>| rs1"
using child2_in_child t_def(1) by fast
- have "is_subtree t (Node r1 rs1)" using subtree_if_child notin_fset r1_def(2) by fastforce
+ have "is_subtree t (Node r1 rs1)" using subtree_if_child r1_def(2)
+ by (metis image_iff prod.sel(1))
moreover have "is_subtree (Node r1 rs1) (Node r xs)"
using subtree_if_child r1_def(1) by fastforce
ultimately have "is_subtree t (Node r xs)" using subtree_trans by blast
then show ?thesis using t_def(2) subtree_in_dlverts by blast
qed
qed
next
case rec: False
then show ?thesis
proof(cases "v \<in> set r")
case False
then have "\<exists>(t,e) \<in> fset xs. v \<in> dlverts (combine x y t)"
using ind.prems list_dtree_rec rec by force
then show ?thesis using ind.IH list_dtree_rec by fastforce
qed (simp)
qed
qed
lemma ex_subtree_if_in_lverts: "v \<in> dlverts t1 \<Longrightarrow> \<exists>t2. is_subtree t2 t1 \<and> v \<in> set (root t2)"
apply(induction t1)
apply(cases)
apply simp
by fastforce
lemma child'_in_child2:
assumes "(Node y rs1,e1) \<in> fset xs" and "(t2,e2) \<in> fset rs1"
shows "(t2,e2) \<in> fset (child2 y ys xs)"
using assms proof(induction xs)
case (insert x xs)
obtain r rs re where r_def: "(Node r rs, re) = x" by (metis dtree.exhaust surj_pair)
show ?case
proof(cases "r = y")
case ry: True
then have 0: "child2 y ys (finsert x xs) = rs |\<union>| (child2 y ys xs)"
using r_def insert.hyps(1) by force
then show ?thesis using insert by fastforce
next
case False
then show ?thesis using insert r_def by force
qed
qed (simp)
lemma v_in_comb_if_in_dlverts: "v \<in> dlverts t \<Longrightarrow> v \<in> dlverts (combine x y t)"
using list_dtree_axioms proof(induction t)
case ind: (Node r xs)
then interpret list_dtree "Node r xs" using ind.prems by blast
show ?case
proof(cases "x=r \<and> (\<exists>t. t \<in> fst ` fset xs \<and> root t = y)")
case x_and_y: True
then have 0: "combine x y (Node r xs) = Node (x@y) (child2 y (remove_child y xs) xs)" by simp
show ?thesis
proof(cases "v \<in> set x \<union> set y")
case True
then show ?thesis using x_and_y dtree.set_sel(1) lverts_if_in_verts by fastforce
next
case False
obtain t where t_def: "is_subtree t (Node r xs)" "v \<in> set (root t)"
using ex_subtree_if_in_lverts ind.prems by fast
then have "Node r xs \<noteq> t" using False x_and_y by fastforce
then obtain t1 e1 where t1_def: "is_subtree t t1" "(t1,e1) \<in> fset xs"
using t_def(1) by force
then show ?thesis
proof(cases "root t1 = y")
case True
then have "t1 \<noteq> t" using False t_def(2) by blast
then obtain rs1 where rs1_def: "t1 = Node y rs1" using True dtree.exhaust_sel by blast
then obtain t2 e2 where t2_def: "is_subtree t t2" "(t2,e2) \<in> fset rs1"
using \<open>t1\<noteq>t\<close> t1_def(1) by auto
have "(t2,e2) \<in> fset (child2 y (remove_child y xs) xs)"
using t2_def(2) rs1_def t1_def(2) child'_in_child2 by fast
- then have "is_subtree t2 (combine x y (Node r xs))" using subtree_if_child 0 by fastforce
+ then have "is_subtree t2 (combine x y (Node r xs))" using subtree_if_child 0
+ using self_subtree by fastforce
then have "is_subtree t (combine x y (Node r xs))" using subtree_trans t2_def(1) by blast
then show ?thesis
using t_def(2) t2_def(1) subtree_in_dlverts dtree.set_sel(1) lverts_if_in_verts by fast
next
case False
then have "(t1,e1) \<in> fset (remove_child y xs)" using t1_def(2) by simp
then have "(t1,e1) \<in> fset (child2 y (remove_child y xs) xs)"
using less_eq_fset.rep_eq input_in_child2 by fast
then have "is_subtree t (combine x y (Node r xs))"
using 0 subtree_if_child subtree_trans t1_def(1) by auto
then show ?thesis
using t_def(2) subtree_in_dlverts dtree.set_sel(1) lverts_if_in_verts by fast
qed
qed
next
case rec: False
then show ?thesis
proof(cases "v \<in> set r")
case False
then obtain t e where t_def: "(t,e) \<in> fset xs" "v \<in> dlverts t" using ind.prems by auto
then have "v \<in> dlverts (combine x y t)" using ind.IH list_dtree_rec by auto
then show ?thesis using rec t_def(1) by force
qed (simp)
qed
qed
lemma dlverts_comb_id[simp]: "dlverts (combine x y t) = dlverts t"
using v_in_comb_if_in_dlverts v_in_dlverts_if_in_comb by blast
lemma wf_dlverts_comb_aux:
assumes "\<forall>(t,e) \<in> fset xs. dlverts (combine x y t) = dlverts t"
and "\<forall>(t1,e1) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)"
and "(t1,e1) \<in> fset ((\<lambda>(t,e). (combine x y t, e)) |`| xs)"
and "(t2,e2) \<in> fset ((\<lambda>(t,e). (combine x y t, e)) |`| xs)"
shows "dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)"
proof -
obtain t1' where t1_def: "combine x y t1' = t1" "(t1',e1) \<in> fset xs" using assms(3) by auto
obtain t2' where t2_def: "combine x y t2' = t2" "(t2',e2) \<in> fset xs" using assms(4) by auto
show ?thesis
proof(cases "dlverts t1' \<inter> dlverts t2' = {}")
case True
then show ?thesis using assms(1) t1_def t2_def by blast
next
case False
then show ?thesis using assms(2) t1_def t2_def by fast
qed
qed
lemma wf_dlverts_child2:
assumes "(t1,e) \<in> fset (child2 y (remove_child y xs) xs)"
and "\<forall>(t,e) \<in> fset xs. wf_dlverts t"
shows "wf_dlverts t1"
proof(cases "(t1,e) |\<in>| (remove_child y xs)")
case True
- then show ?thesis using assms(2) notin_fset by fastforce
+ then show ?thesis using assms(2) by fastforce
next
case False
then obtain rs re where r_def: "(Node y rs, re) \<in> fset xs" "(t1,e)|\<in>| rs"
using child2_in_child assms(1) by fast
- then show ?thesis using assms(2) notin_fset by fastforce
+ then show ?thesis using assms(2) by fastforce
qed
lemma wf_dlverts_child2_aux1:
assumes "(t1,e1) \<in> fset (child2 y (remove_child y xs) xs)"
and "\<exists>t. t \<in> fst ` fset xs \<and> root t = y"
and "wf_dlverts (Node r xs)"
shows "set (r@y) \<inter> dlverts t1 = {}"
proof(cases "(t1,e1) |\<in>| (remove_child y xs)")
case True
- then have t1_def: "root t1 \<noteq> y" "(t1,e1) \<in> fset xs" using notin_fset by fastforce+
+ then have t1_def: "root t1 \<noteq> y" "(t1,e1) \<in> fset xs" by fastforce+
obtain t et where t_def: "(t,et) \<in> fset xs" "root t = y" using assms(2) by force
have "\<forall>y'\<in> set y. y' \<notin> dlverts t1"
proof
fix y'
assume "y' \<in> set y"
then have asm: "y' \<in> dlverts t" using t_def(2) dtree.set_sel(1) lverts_if_in_verts by fastforce
have "dlverts t1 \<inter> dlverts t = {}" using assms(3) t1_def t_def by fastforce
then show "y' \<notin> dlverts t1" using asm by blast
qed
then show ?thesis using assms(3) t1_def(2) by auto
next
case False
then obtain rs1 re1 where r_def: "(Node y rs1, re1) \<in> fset xs" "(t1,e1)|\<in>| rs1"
using child2_in_child assms(1) by fast
- have "\<forall>y'\<in> set y. y' \<notin> dlverts t1" using assms(3) r_def notin_fset by fastforce
- then show ?thesis using assms(3) notin_fset r_def by fastforce
+ have "\<forall>y'\<in> set y. y' \<notin> dlverts t1" using assms(3) r_def by fastforce
+ then show ?thesis using assms(3) r_def by fastforce
qed
lemma wf_dlverts_child2_aux2:
assumes "\<forall>(t1,e1) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs. dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)"
and "\<forall>(t,e) \<in> fset xs. wf_dlverts t"
and "(t1,e1) \<in> fset (child2 y (remove_child y xs) xs)"
and "(t2,e2) \<in> fset (child2 y (remove_child y xs) xs)"
and "(t1,e1)\<noteq>(t2,e2)"
shows "dlverts t1 \<inter> dlverts t2 = {}"
proof(cases "(t1,e1) |\<in>| (remove_child y xs)")
case t1_r: True
then show ?thesis
proof(cases "(t2,e2) |\<in>| (remove_child y xs)")
case True
then show ?thesis
- by (smt (verit, ccfv_threshold) t1_r assms(1,5) Int_iff case_prodD filter_fset notin_fset)
+ by (smt (verit, ccfv_threshold) t1_r assms(1,5) Int_iff case_prodD filter_fset)
next
case False
then obtain rs2 re2 where r_def: "(Node y rs2, re2) \<in> fset xs" "(t2,e2)|\<in>| rs2"
using child2_in_child assms(4) by fast
then show ?thesis
- using t1_r assms(1) notin_fset ffmember_filter inf_assoc inf_bot_right inf_commute
- by (smt (z3) dtree.sel(1) semilattice_inf_class.inf.absorb_iff2 case_prodD child_in_dlverts)
+ using t1_r assms(1) ffmember_filter inf_assoc inf_bot_right inf_commute
+ by (smt (verit) dtree.sel(1) semilattice_inf_class.inf.absorb_iff2 case_prodD child_in_dlverts)
qed
next
case False
then obtain rs1 re1 where r1_def: "(Node y rs1, re1) \<in> fset xs" "(t1,e1)|\<in>| rs1"
using child2_in_child assms(3) by fast
show ?thesis
proof(cases "(t2,e2) |\<in>| (remove_child y xs)")
case True
then show ?thesis
- using r1_def assms(1) notin_fset ffmember_filter inf_assoc inf_bot_right inf_commute
- by (smt (z3) dtree.sel(1) semilattice_inf_class.inf.absorb_iff2 case_prodD child_in_dlverts)
+ using r1_def assms(1) ffmember_filter inf_assoc inf_bot_right inf_commute
+ by (smt (verit) dtree.sel(1) semilattice_inf_class.inf.absorb_iff2 case_prodD child_in_dlverts)
next
case False
then obtain rs2 re2 where r2_def: "(Node y rs2, re2) \<in> fset xs" "(t2,e2) |\<in>| rs2"
using child2_in_child assms(4) by fast
then show ?thesis
proof(cases "rs1=rs2")
case True
have "\<forall>(t1,e1) \<in> fset rs1. \<forall>(t2,e2) \<in> fset rs1.
dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)"
using r1_def(1) assms(2) by fastforce
then show ?thesis
- using r1_def(2) r2_def(2) assms(5) True notin_fset
+ using r1_def(2) r2_def(2) assms(5) True
by (metis (mono_tags, lifting) case_prodD)
next
case False
then have "dlverts (Node y rs1) \<inter> dlverts (Node y rs2) = {}"
using assms(1) r1_def(1) r2_def(1) by fast
then show ?thesis
- using r1_def(2) r2_def(2) child_in_dlverts notin_fset
+ using r1_def(2) r2_def(2) child_in_dlverts
by (metis order_bot_class.bot.extremum_uniqueI inf_mono)
qed
qed
qed
lemma wf_dlverts_combine: "wf_dlverts (combine x y t)"
using list_dtree_axioms proof(induction t)
case ind: (Node r xs)
then interpret list_dtree "Node r xs" using ind.prems by blast
show ?case
proof(cases "x=r \<and> (\<exists>t. t \<in> fst ` fset xs \<and> root t = y)")
case True
let ?xs = "child2 y (remove_child y xs) xs"
have "\<forall>(t1,e1) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs.
dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)" using wf_lverts by fastforce
moreover have "\<forall>(t1,e1) \<in> fset xs. wf_dlverts t1" using wf_lverts by fastforce
ultimately have "\<forall>(t1,e1) \<in> fset ?xs. \<forall>(t2,e2) \<in> fset ?xs.
dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)"
using wf_dlverts_child2_aux2[of xs] by blast
moreover have "\<forall>(x,e) \<in> fset ?xs. wf_dlverts x" using wf_dlverts_child2 wf_lverts by fastforce
moreover have "(x@y) \<noteq> []" using True wf_lverts by simp
moreover have "\<forall>(t1,e1) \<in> fset ?xs. set (x@y) \<inter> dlverts t1 = {}"
using wf_dlverts_child2_aux1 wf_lverts True by fast
ultimately have "wf_dlverts (Node (x@y) ?xs)" by fastforce
moreover have "combine x y (Node r xs) = Node (x@y) ?xs" using True by simp
ultimately show ?thesis by argo
next
case False
let ?xs = "(\<lambda>(t,e). (combine x y t, e)) |`| xs"
have 0: "\<forall>(t,e) \<in> fset xs. dlverts (combine x y t) = dlverts t"
using list_dtree.dlverts_comb_id list_dtree_rec by fast
have 1: "\<forall>(t,e) \<in> fset ?xs. wf_dlverts t" using ind.IH list_dtree_rec by auto
have 2: "\<forall>(t,e) \<in> fset ?xs. set r \<inter> dlverts t = {}" using 0 wf_lverts by fastforce
have "\<forall>(t1,e1) \<in> fset xs. \<forall>(t2,e2) \<in> fset xs.
dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)" using wf_lverts by fastforce
then have 3: "\<forall>(t1,e1) \<in> fset ?xs. \<forall>(t2,e2) \<in> fset ?xs.
dlverts t1 \<inter> dlverts t2 = {} \<or> (t1,e1)=(t2,e2)"
using 0 wf_dlverts_comb_aux[of xs] by blast
have 4: "combine x y (Node r xs) = Node r ?xs" using False by auto
have "r \<noteq> []" using wf_lverts by simp
then show ?thesis using 1 2 3 4 by fastforce
qed
qed
theorem list_dtree_comb: "list_dtree (combine x y t)"
by(unfold_locales) (auto simp: wf_darcs_combine wf_dlverts_combine)
end
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/GTT_Compose.thy b/thys/Regular_Tree_Relations/GTT_Compose.thy
--- a/thys/Regular_Tree_Relations/GTT_Compose.thy
+++ b/thys/Regular_Tree_Relations/GTT_Compose.thy
@@ -1,403 +1,403 @@
theory GTT_Compose
imports GTT
begin
subsection \<open>GTT closure under composition\<close>
inductive_set \<Delta>\<^sub>\<epsilon>_set :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) set" for \<A> \<B> where
\<Delta>\<^sub>\<epsilon>_set_cong: "TA_rule f ps p |\<in>| rules \<A> \<Longrightarrow> TA_rule f qs q |\<in>| rules \<B> \<Longrightarrow> length ps = length qs \<Longrightarrow>
(\<And>i. i < length qs \<Longrightarrow> (ps ! i, qs ! i) \<in> \<Delta>\<^sub>\<epsilon>_set \<A> \<B>) \<Longrightarrow> (p, q) \<in> \<Delta>\<^sub>\<epsilon>_set \<A> \<B>"
| \<Delta>\<^sub>\<epsilon>_set_eps1: "(p, p') |\<in>| eps \<A> \<Longrightarrow> (p, q) \<in> \<Delta>\<^sub>\<epsilon>_set \<A> \<B> \<Longrightarrow> (p', q) \<in> \<Delta>\<^sub>\<epsilon>_set \<A> \<B>"
| \<Delta>\<^sub>\<epsilon>_set_eps2: "(q, q') |\<in>| eps \<B> \<Longrightarrow> (p, q) \<in> \<Delta>\<^sub>\<epsilon>_set \<A> \<B> \<Longrightarrow> (p, q') \<in> \<Delta>\<^sub>\<epsilon>_set \<A> \<B>"
lemma \<Delta>\<^sub>\<epsilon>_states: "\<Delta>\<^sub>\<epsilon>_set \<A> \<B> \<subseteq> fset (\<Q> \<A> |\<times>| \<Q> \<B>)"
proof -
{fix p q assume "(p, q) \<in> \<Delta>\<^sub>\<epsilon>_set \<A> \<B>" then have "(p, q) \<in> fset (\<Q> \<A> |\<times>| \<Q> \<B>)"
- by (induct) (auto dest: rule_statesD eps_statesD simp flip: fmember_iff_member_fset)}
+ by (induct) (auto dest: rule_statesD eps_statesD)}
then show ?thesis by auto
qed
lemma finite_\<Delta>\<^sub>\<epsilon> [simp]: "finite (\<Delta>\<^sub>\<epsilon>_set \<A> \<B>)"
using finite_subset[OF \<Delta>\<^sub>\<epsilon>_states]
by simp
context
includes fset.lifting
begin
lift_definition \<Delta>\<^sub>\<epsilon> :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) fset" is \<Delta>\<^sub>\<epsilon>_set by simp
lemmas \<Delta>\<^sub>\<epsilon>_cong = \<Delta>\<^sub>\<epsilon>_set_cong [Transfer.transferred]
lemmas \<Delta>\<^sub>\<epsilon>_eps1 = \<Delta>\<^sub>\<epsilon>_set_eps1 [Transfer.transferred]
lemmas \<Delta>\<^sub>\<epsilon>_eps2 = \<Delta>\<^sub>\<epsilon>_set_eps2 [Transfer.transferred]
lemmas \<Delta>\<^sub>\<epsilon>_cases = \<Delta>\<^sub>\<epsilon>_set.cases[Transfer.transferred]
lemmas \<Delta>\<^sub>\<epsilon>_induct [consumes 1, case_names \<Delta>\<^sub>\<epsilon>_cong \<Delta>\<^sub>\<epsilon>_eps1 \<Delta>\<^sub>\<epsilon>_eps2] = \<Delta>\<^sub>\<epsilon>_set.induct[Transfer.transferred]
lemmas \<Delta>\<^sub>\<epsilon>_intros = \<Delta>\<^sub>\<epsilon>_set.intros[Transfer.transferred]
lemmas \<Delta>\<^sub>\<epsilon>_simps = \<Delta>\<^sub>\<epsilon>_set.simps[Transfer.transferred]
end
lemma finite_alt_def [simp]:
"finite {(\<alpha>, \<beta>). (\<exists>t. ground t \<and> \<alpha> |\<in>| ta_der \<A> t \<and> \<beta> |\<in>| ta_der \<B> t)}" (is "finite ?S")
- by (auto dest: ground_ta_der_states[THEN fsubsetD] simp flip: fmember_iff_member_fset
+ by (auto dest: ground_ta_der_states[THEN fsubsetD]
intro!: finite_subset[of ?S "fset (\<Q> \<A> |\<times>| \<Q> \<B>)"])
lemma \<Delta>\<^sub>\<epsilon>_def':
"\<Delta>\<^sub>\<epsilon> \<A> \<B> = {|(\<alpha>, \<beta>). (\<exists>t. ground t \<and> \<alpha> |\<in>| ta_der \<A> t \<and> \<beta> |\<in>| ta_der \<B> t)|}"
proof (intro fset_eqI iffI, goal_cases lr rl)
case (lr x) obtain p q where x [simp]: "x = (p, q)" by (cases x)
have "\<exists>t. ground t \<and> p |\<in>| ta_der \<A> t \<and> q |\<in>| ta_der \<B> t" using lr unfolding x
proof (induct rule: \<Delta>\<^sub>\<epsilon>_induct)
case (\<Delta>\<^sub>\<epsilon>_cong f ps p qs q)
obtain ts where ts: "ground (ts i) \<and> ps ! i |\<in>| ta_der \<A> (ts i) \<and> qs ! i |\<in>| ta_der \<B> (ts i)"
if "i < length qs" for i using \<Delta>\<^sub>\<epsilon>_cong(5) by metis
then show ?case using \<Delta>\<^sub>\<epsilon>_cong(1-3)
by (auto intro!: exI[of _ "Fun f (map ts [0..<length qs])"]) blast+
qed (meson ta_der_eps)+
then show ?case by auto
next
case (rl x) obtain p q where x [simp]: "x = (p, q)" by (cases x)
obtain t where "ground t" "p |\<in>| ta_der \<A> t" "q |\<in>| ta_der \<B> t" using rl by auto
then show ?case unfolding x
proof (induct t arbitrary: p q)
case (Fun f ts)
obtain p' ps where p': "TA_rule f ps p' |\<in>| rules \<A>" "p' = p \<or> (p', p) |\<in>| (eps \<A>)|\<^sup>+|" "length ps = length ts"
"\<And>i. i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der \<A> (ts ! i)" using Fun(3) by auto
obtain q' qs where q': "f qs \<rightarrow> q' |\<in>| rules \<B>" "q' = q \<or> (q', q) |\<in>| (eps \<B>)|\<^sup>+|" "length qs = length ts"
"\<And>i. i < length ts \<Longrightarrow> qs ! i |\<in>| ta_der \<B> (ts ! i)" using Fun(4) by auto
have st: "(p', q') |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B>"
using Fun(1)[OF nth_mem _ p'(4) q'(4)] Fun(2) p'(3) q'(3)
by (intro \<Delta>\<^sub>\<epsilon>_cong[OF p'(1) q'(1)]) auto
{assume "(p', p) |\<in>| (eps \<A>)|\<^sup>+|" then have "(p, q') |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B>" using st
by (induct rule: ftrancl_induct) (auto intro: \<Delta>\<^sub>\<epsilon>_eps1)}
from st this p'(2) have st: "(p, q') |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B>" by auto
{assume "(q', q) |\<in>| (eps \<B>)|\<^sup>+|" then have "(p, q) |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B>" using st
by (induct rule: ftrancl_induct) (auto intro: \<Delta>\<^sub>\<epsilon>_eps2)}
from st this q'(2) show "(p, q) |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B>" by auto
qed auto
qed
lemma \<Delta>\<^sub>\<epsilon>_fmember:
"(p, q) |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B> \<longleftrightarrow> (\<exists>t. ground t \<and> p |\<in>| ta_der \<A> t \<and> q |\<in>| ta_der \<B> t)"
by (auto simp: \<Delta>\<^sub>\<epsilon>_def')
definition GTT_comp :: "('q, 'f) gtt \<Rightarrow> ('q, 'f) gtt \<Rightarrow> ('q, 'f) gtt" where
"GTT_comp \<G>\<^sub>1 \<G>\<^sub>2 =
(let \<Delta> = \<Delta>\<^sub>\<epsilon> (snd \<G>\<^sub>1) (fst \<G>\<^sub>2) in
(TA (gtt_rules (fst \<G>\<^sub>1, fst \<G>\<^sub>2)) (eps (fst \<G>\<^sub>1) |\<union>| eps (fst \<G>\<^sub>2) |\<union>| \<Delta>),
TA (gtt_rules (snd \<G>\<^sub>1, snd \<G>\<^sub>2)) (eps (snd \<G>\<^sub>1) |\<union>| eps (snd \<G>\<^sub>2) |\<union>| (\<Delta>|\<inverse>|))))"
lemma gtt_syms_GTT_comp:
"gtt_syms (GTT_comp A B) = gtt_syms A |\<union>| gtt_syms B"
by (auto simp: GTT_comp_def ta_sig_def Let_def)
lemma \<Delta>\<^sub>\<epsilon>_statesD:
"(p, q) |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B> \<Longrightarrow> p |\<in>| \<Q> \<A>"
"(p, q) |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B> \<Longrightarrow> q |\<in>| \<Q> \<B>"
using subsetD[OF \<Delta>\<^sub>\<epsilon>_states, of "(p, q)" \<A> \<B>]
- by (auto simp flip: \<Delta>\<^sub>\<epsilon>.rep_eq fmember_iff_member_fset)
+ by (auto simp flip: \<Delta>\<^sub>\<epsilon>.rep_eq)
lemma \<Delta>\<^sub>\<epsilon>_statesD':
"q |\<in>| eps_states (\<Delta>\<^sub>\<epsilon> \<A> \<B>) \<Longrightarrow> q |\<in>| \<Q> \<A> |\<union>| \<Q> \<B>"
- by (auto simp: eps_states_def fmember.abs_eq dest: \<Delta>\<^sub>\<epsilon>_statesD)
+ by (auto simp: eps_states_def dest: \<Delta>\<^sub>\<epsilon>_statesD)
lemma \<Delta>\<^sub>\<epsilon>_swap:
"prod.swap p |\<in>| \<Delta>\<^sub>\<epsilon> \<A> \<B> \<longleftrightarrow> p |\<in>| \<Delta>\<^sub>\<epsilon> \<B> \<A>"
by (auto simp: \<Delta>\<^sub>\<epsilon>_def')
lemma \<Delta>\<^sub>\<epsilon>_inverse [simp]:
"(\<Delta>\<^sub>\<epsilon> \<A> \<B>)|\<inverse>| = \<Delta>\<^sub>\<epsilon> \<B> \<A>"
by (auto simp: \<Delta>\<^sub>\<epsilon>_def')
lemma gtt_states_comp_union:
"gtt_states (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2) |\<subseteq>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2"
proof (intro fsubsetI, goal_cases lr)
case (lr q) then show ?case
by (auto simp: GTT_comp_def gtt_states_def \<Q>_def dest: \<Delta>\<^sub>\<epsilon>_statesD')
qed
lemma GTT_comp_swap [simp]:
"GTT_comp (prod.swap \<G>\<^sub>2) (prod.swap \<G>\<^sub>1) = prod.swap (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)"
by (simp add: GTT_comp_def ac_simps)
lemma gtt_comp_complete_semi:
assumes s: "q |\<in>| gta_der (fst \<G>\<^sub>1) s" and u: "q |\<in>| gta_der (snd \<G>\<^sub>1) u" and ut: "gtt_accept \<G>\<^sub>2 u t"
shows "q |\<in>| gta_der (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) s" "q |\<in>| gta_der (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) t"
proof (goal_cases L R)
let ?\<G> = "GTT_comp \<G>\<^sub>1 \<G>\<^sub>2"
have sub1l: "rules (fst \<G>\<^sub>1) |\<subseteq>| rules (fst ?\<G>)" "eps (fst \<G>\<^sub>1) |\<subseteq>| eps (fst ?\<G>)"
and sub1r: "rules (snd \<G>\<^sub>1) |\<subseteq>| rules (snd ?\<G>)" "eps (snd \<G>\<^sub>1) |\<subseteq>| eps (snd ?\<G>)"
and sub2r: "rules (snd \<G>\<^sub>2) |\<subseteq>| rules (snd ?\<G>)" "eps (snd \<G>\<^sub>2) |\<subseteq>| eps (snd ?\<G>)"
by (auto simp: GTT_comp_def)
{ case L then show ?case using s ta_der_mono[OF sub1l]
by (auto simp: gta_der_def)
next
case R then show ?case using ut u unfolding gtt_accept_def
proof (induct arbitrary: q s)
case (base s t)
from base(1) obtain p where p: "p |\<in>| gta_der (fst \<G>\<^sub>2) s" "p |\<in>| gta_der (snd \<G>\<^sub>2) t"
by (auto simp: agtt_lang_def)
then have "(p, q) |\<in>| eps (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2))"
using \<Delta>\<^sub>\<epsilon>_fmember[of p q "fst \<G>\<^sub>2" "snd \<G>\<^sub>1"] base(2)
by (auto simp: GTT_comp_def gta_der_def)
from ta_der_eps[OF this] show ?case using p ta_der_mono[OF sub2r]
by (auto simp add: gta_der_def)
next
case (step ss ts f)
from step(1, 4) obtain ps p where "TA_rule f ps p |\<in>| rules (snd \<G>\<^sub>1)" "p = q \<or> (p, q) |\<in>| (eps (snd \<G>\<^sub>1))|\<^sup>+|"
"length ps = length ts" "\<And>i. i < length ts \<Longrightarrow> ps ! i |\<in>| gta_der (snd \<G>\<^sub>1) (ss ! i)"
unfolding gta_der_def by auto
then show ?case using step(1, 2) sub1r(1) ftrancl_mono[OF _ sub1r(2)]
by (auto simp: gta_der_def intro!: exI[of _ p] exI[of _ ps])
qed}
qed
lemmas gtt_comp_complete_semi' = gtt_comp_complete_semi[of _ "prod.swap \<G>\<^sub>2" _ _ "prod.swap \<G>\<^sub>1" for \<G>\<^sub>1 \<G>\<^sub>2,
unfolded fst_swap snd_swap GTT_comp_swap gtt_accept_swap]
lemma gtt_comp_acomplete:
"gcomp_rel UNIV (agtt_lang \<G>\<^sub>1) (agtt_lang \<G>\<^sub>2) \<subseteq> agtt_lang (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)"
proof (intro subrelI, goal_cases LR)
case (LR s t)
then consider
q u where "q |\<in>| gta_der (fst \<G>\<^sub>1) s" "q |\<in>| gta_der (snd \<G>\<^sub>1) u" "gtt_accept \<G>\<^sub>2 u t"
| q u where "q |\<in>| gta_der (snd \<G>\<^sub>2) t" "q |\<in>| gta_der (fst \<G>\<^sub>2) u" "gtt_accept \<G>\<^sub>1 s u"
by (auto simp: gcomp_rel_def gtt_accept_def elim!: agtt_langE)
then show ?case
proof (cases)
case 1 show ?thesis using gtt_comp_complete_semi[OF 1]
by (auto simp: agtt_lang_def gta_der_def)
next
case 2 show ?thesis using gtt_comp_complete_semi'[OF 2]
by (auto simp: agtt_lang_def gta_der_def)
qed
qed
lemma \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>2:
assumes "(q, q') |\<in>| (eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+|" "q |\<in>| gtt_states \<G>\<^sub>2"
"gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
shows "(q, q') |\<in>| (eps (fst \<G>\<^sub>2))|\<^sup>+| \<and> q' |\<in>| gtt_states \<G>\<^sub>2"
using assms(1-2)
proof (induct rule: converse_ftrancl_induct)
case (Base y)
then show ?case using assms(3)
- by (fastforce simp: GTT_comp_def gtt_states_def fmember.abs_eq dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD(1))
+ by (fastforce simp: GTT_comp_def gtt_states_def dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD(1))
next
case (Step q p)
have "(q, p) |\<in>| (eps (fst \<G>\<^sub>2))|\<^sup>+|" "p |\<in>| gtt_states \<G>\<^sub>2"
using Step(1, 4) assms(3)
- by (auto simp: GTT_comp_def gtt_states_def fmember.abs_eq dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD(1))
+ by (auto simp: GTT_comp_def gtt_states_def dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD(1))
then show ?case using Step(3)
by (auto intro: ftrancl_trans)
qed
lemma \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>1:
assumes "(p, r) |\<in>| (eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+|" "p |\<in>| gtt_states \<G>\<^sub>1"
"gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
obtains "r |\<in>| gtt_states \<G>\<^sub>1" "(p, r) |\<in>| (eps (fst \<G>\<^sub>1))|\<^sup>+|"
| q p' where "r |\<in>| gtt_states \<G>\<^sub>2" "p = p' \<or> (p, p') |\<in>| (eps (fst \<G>\<^sub>1))|\<^sup>+|" "(p', q) |\<in>| \<Delta>\<^sub>\<epsilon> (snd \<G>\<^sub>1) (fst \<G>\<^sub>2)"
"q = r \<or> (q, r) |\<in>| (eps (fst \<G>\<^sub>2))|\<^sup>+|"
using assms(1,2)
proof (induct arbitrary: thesis rule: converse_ftrancl_induct)
case (Base p)
from Base(1) consider (a) "(p, r) |\<in>| eps (fst \<G>\<^sub>1)" | (b) "(p, r) |\<in>| eps (fst \<G>\<^sub>2)" |
(c) "(p, r) |\<in>| (\<Delta>\<^sub>\<epsilon> (snd \<G>\<^sub>1) (fst \<G>\<^sub>2))"
- by (auto simp: GTT_comp_def fmember.abs_eq)
+ by (auto simp: GTT_comp_def)
then show ?case using assms(3) Base
- by cases (auto simp: GTT_comp_def gtt_states_def fmember.abs_eq dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD)
+ by cases (auto simp: GTT_comp_def gtt_states_def dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD)
next
case (Step q p)
consider "(q, p) |\<in>| (eps (fst \<G>\<^sub>1))|\<^sup>+|" "p |\<in>| gtt_states \<G>\<^sub>1"
| "(q, p) |\<in>| \<Delta>\<^sub>\<epsilon> (snd \<G>\<^sub>1) (fst \<G>\<^sub>2)" "p |\<in>| gtt_states \<G>\<^sub>2" using assms(3) Step(1, 6)
- by (auto simp: GTT_comp_def gtt_states_def fmember.abs_eq dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD)
+ by (auto simp: GTT_comp_def gtt_states_def dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD)
then show ?case
proof (cases)
case 1 note a = 1 show ?thesis
proof (cases rule: Step(3))
case (2 p' q)
then show ?thesis using assms a
by (auto intro: Step(5) ftrancl_trans)
qed (auto simp: a(2) intro: Step(4) ftrancl_trans[OF a(1)])
next
case 2 show ?thesis using \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>2[OF Step(2) 2(2) assms(3)] Step(5)[OF _ _ 2(1)] by auto
qed
qed
lemma \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>1_\<G>\<^sub>2:
assumes "(q, q') |\<in>| (eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+|" "q |\<in>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2"
"gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
obtains "q |\<in>| gtt_states \<G>\<^sub>1" "q' |\<in>| gtt_states \<G>\<^sub>1" "(q, q') |\<in>| (eps (fst \<G>\<^sub>1))|\<^sup>+|"
| p p' where "q |\<in>| gtt_states \<G>\<^sub>1" "q' |\<in>| gtt_states \<G>\<^sub>2" "q = p \<or> (q, p) |\<in>| (eps (fst \<G>\<^sub>1))|\<^sup>+|"
"(p, p') |\<in>| \<Delta>\<^sub>\<epsilon> (snd \<G>\<^sub>1) (fst \<G>\<^sub>2)" "p' = q' \<or> (p', q') |\<in>| (eps (fst \<G>\<^sub>2))|\<^sup>+|"
| "q |\<in>| gtt_states \<G>\<^sub>2" "(q, q') |\<in>| (eps (fst \<G>\<^sub>2))|\<^sup>+| \<and> q' |\<in>| gtt_states \<G>\<^sub>2"
using assms \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>1 \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>2
by (metis funion_iff)
lemma GTT_comp_eps_fst_statesD:
"(p, q) |\<in>| eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) \<Longrightarrow> p |\<in>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2"
"(p, q) |\<in>| eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) \<Longrightarrow> q |\<in>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2"
- by (auto simp: GTT_comp_def gtt_states_def fmember.abs_eq dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD)
+ by (auto simp: GTT_comp_def gtt_states_def dest: eps_statesD \<Delta>\<^sub>\<epsilon>_statesD)
lemma GTT_comp_eps_ftrancl_fst_statesD:
"(p, q) |\<in>| (eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+| \<Longrightarrow> p |\<in>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2"
"(p, q) |\<in>| (eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+| \<Longrightarrow> q |\<in>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2"
using GTT_comp_eps_fst_statesD[of _ _ \<G>\<^sub>1 \<G>\<^sub>2]
by (meson converse_ftranclE ftranclE)+
lemma GTT_comp_first:
assumes "q |\<in>| ta_der (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) t" "q |\<in>| gtt_states \<G>\<^sub>1"
"gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
shows "q |\<in>| ta_der (fst \<G>\<^sub>1) t"
using assms(1,2)
proof (induct t arbitrary: q)
case (Var q')
have "q \<noteq> q' \<Longrightarrow> q' |\<in>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2" using Var
by (auto dest: GTT_comp_eps_ftrancl_fst_statesD)
then show ?case using Var assms(3)
by (auto elim: \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>1_\<G>\<^sub>2)
next
case (Fun f ts)
obtain q' qs where q': "TA_rule f qs q' |\<in>| rules (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2))"
"q' = q \<or> (q', q) |\<in>| (eps (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+|" "length qs = length ts"
"\<And>i. i < length ts \<Longrightarrow> qs ! i |\<in>| ta_der (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) (ts ! i)"
using Fun(2) by auto
have "q' |\<in>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2" using q'(1)
by (auto simp: GTT_comp_def gtt_states_def dest: rule_statesD)
then have st: "q' |\<in>| gtt_states \<G>\<^sub>1" and eps:"q' = q \<or> (q', q) |\<in>| (eps (fst \<G>\<^sub>1))|\<^sup>+|"
using q'(2) Fun(3) assms(3)
by (auto elim!: \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>1_\<G>\<^sub>2)
from st have rule: "TA_rule f qs q' |\<in>| rules (fst \<G>\<^sub>1)" using assms(3) q'(1)
by (auto simp: GTT_comp_def gtt_states_def dest: rule_statesD)
have "i < length ts \<Longrightarrow> qs ! i |\<in>| ta_der (fst \<G>\<^sub>1) (ts ! i)" for i
using rule q'(3, 4)
by (intro Fun(1)[OF nth_mem]) (auto simp: gtt_states_def dest!: rule_statesD(4))
then show ?case using q'(3) rule eps
by auto
qed
lemma GTT_comp_second:
assumes "gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}" "q |\<in>| gtt_states \<G>\<^sub>2"
"q |\<in>| ta_der (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) t"
shows "q |\<in>| ta_der (snd \<G>\<^sub>2) t"
using assms GTT_comp_first[of q "prod.swap \<G>\<^sub>2" "prod.swap \<G>\<^sub>1"]
by (auto simp: gtt_states_def)
lemma gtt_comp_sound_semi:
fixes \<G>\<^sub>1 \<G>\<^sub>2 :: "('f, 'q) gtt"
assumes as2: "gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
and 1: "q |\<in>| gta_der (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) s" "q |\<in>| gta_der (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) t" "q |\<in>| gtt_states \<G>\<^sub>1"
shows "\<exists>u. q |\<in>| gta_der (snd \<G>\<^sub>1) u \<and> gtt_accept \<G>\<^sub>2 u t" using 1(2,3) unfolding gta_der_def
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
show ?case
proof (cases "p |\<in>| gtt_states \<G>\<^sub>1")
case True
then have *: "TA_rule f ps p |\<in>| rules (snd \<G>\<^sub>1)" using GFun(1, 6) as2
by (auto simp: GTT_comp_def gtt_states_def dest: rule_statesD)
moreover have st: "i < length ps \<Longrightarrow> ps ! i |\<in>| gtt_states \<G>\<^sub>1" for i using *
by (force simp: gtt_states_def dest: rule_statesD)
moreover have "i < length ps \<Longrightarrow> \<exists>u. ps ! i |\<in>| ta_der (snd \<G>\<^sub>1) (term_of_gterm u) \<and> gtt_accept \<G>\<^sub>2 u (ts ! i)" for i
using st GFun(2) by (intro GFun(5)) simp
then obtain us where
"\<And>i. i < length ps \<Longrightarrow> ps ! i |\<in>| ta_der (snd \<G>\<^sub>1) (term_of_gterm (us i)) \<and> gtt_accept \<G>\<^sub>2 (us i) (ts ! i)"
by metis
moreover have "p = q \<or> (p, q) |\<in>| (eps (snd \<G>\<^sub>1))|\<^sup>+|" using GFun(3, 6) True as2
by (auto simp: gtt_states_def elim!: \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>1_\<G>\<^sub>2[of p q "prod.swap \<G>\<^sub>2" "prod.swap \<G>\<^sub>1", simplified])
ultimately show ?thesis using GFun(2)
by (intro exI[of _ "GFun f (map us [0..<length ts])"])
(auto simp: gtt_accept_def intro!: exI[of _ ps] exI[of _ p])
next
case False note nt_st = this
then have False: "p \<noteq> q" using GFun(6) by auto
then have eps: "(p, q) |\<in>| (eps (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)))|\<^sup>+|" using GFun(3) by simp
show ?thesis using \<Delta>\<^sub>\<epsilon>_steps_from_\<G>\<^sub>1_\<G>\<^sub>2[of p q "prod.swap \<G>\<^sub>2" "prod.swap \<G>\<^sub>1", simplified, OF eps]
proof (cases, goal_cases)
case 1 then show ?case using False GFun(3)
by (metis GTT_comp_eps_ftrancl_fst_statesD(1) GTT_comp_swap fst_swap funion_iff)
next
case 2 then show ?case using as2 by (auto simp: gtt_states_def)
next
case 3 then show ?case using as2 GFun(6) by (auto simp: gtt_states_def)
next
case (4 r p')
have meet: "r |\<in>| ta_der (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) (Fun f (map term_of_gterm ts))"
using GFun(1 - 4) 4(3) False
by (auto simp: GTT_comp_def in_ftrancl_UnI intro!: exI[ of _ ps] exI[ of _ p])
then obtain u where wit: "ground u" "p' |\<in>| ta_der (snd \<G>\<^sub>1) u" "r |\<in>| ta_der (fst \<G>\<^sub>2) u"
using 4(4-) unfolding \<Delta>\<^sub>\<epsilon>_def' by blast
from wit(1, 3) have "gtt_accept \<G>\<^sub>2 (gterm_of_term u) (GFun f ts)"
using GTT_comp_second[OF as2 _ meet] unfolding gtt_accept_def
by (intro gmctxt_cl.base agtt_langI[of r])
(auto simp add: gta_der_def gtt_states_def simp del: ta_der_Fun dest: ground_ta_der_states)
then show ?case using 4(5) wit(1, 2)
by (intro exI[of _ "gterm_of_term u"]) (auto simp add: ta_der_trancl_eps)
next
case 5
then show ?case using nt_st as2
by (simp add: gtt_states_def)
qed
qed
qed
lemma gtt_comp_asound:
assumes "gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
shows "agtt_lang (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2) \<subseteq> gcomp_rel UNIV (agtt_lang \<G>\<^sub>1) (agtt_lang \<G>\<^sub>2)"
proof (intro subrelI, goal_cases LR)
case (LR s t)
obtain q where q: "q |\<in>| gta_der (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) s" "q |\<in>| gta_der (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) t"
using LR by (auto simp: agtt_lang_def)
{ (* prepare symmetric cases: q |\<in>| gtt_states \<G>\<^sub>1 and q |\<in>| gtt_states \<G>\<^sub>2 *)
fix \<G>\<^sub>1 \<G>\<^sub>2 s t assume as2: "gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
and 1: "q |\<in>| ta_der (fst (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) (term_of_gterm s)"
"q |\<in>| ta_der (snd (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)) (term_of_gterm t)" "q |\<in>| gtt_states \<G>\<^sub>1"
note st = GTT_comp_first[OF 1(1,3) as2]
obtain u where u: "q |\<in>| ta_der (snd \<G>\<^sub>1) (term_of_gterm u)" "gtt_accept \<G>\<^sub>2 u t"
using gtt_comp_sound_semi[OF as2 1[folded gta_der_def]] by (auto simp: gta_der_def)
have "(s, u) \<in> agtt_lang \<G>\<^sub>1" using st u(1)
by (auto simp: agtt_lang_def gta_der_def)
moreover have "(u, t) \<in> gtt_lang \<G>\<^sub>2" using u(2)
by (auto simp: gtt_accept_def)
ultimately have "(s, t) \<in> agtt_lang \<G>\<^sub>1 O gmctxt_cl UNIV (agtt_lang \<G>\<^sub>2)"
by auto}
note base = this
consider "q |\<in>| gtt_states \<G>\<^sub>1" | "q |\<in>| gtt_states \<G>\<^sub>2" | "q |\<notin>| gtt_states \<G>\<^sub>1 |\<union>| gtt_states \<G>\<^sub>2" by blast
then show ?case using q assms
proof (cases, goal_cases)
case 1 then show ?case using base[of \<G>\<^sub>1 \<G>\<^sub>2 s t]
by (auto simp: gcomp_rel_def gta_der_def)
next
case 2 then show ?case using base[of "prod.swap \<G>\<^sub>2" "prod.swap \<G>\<^sub>1" t s, THEN converseI]
by (auto simp: gcomp_rel_def converse_relcomp converse_agtt_lang gta_der_def gtt_states_def)
(simp add: finter_commute funion_commute gtt_lang_swap prod.swap_def)+
next
case 3 then show ?case using fsubsetD[OF gtt_states_comp_union[of \<G>\<^sub>1 \<G>\<^sub>2], of q]
by (auto simp: gta_der_def gtt_states_def)
qed
qed
lemma gtt_comp_lang_complete:
shows "gtt_lang \<G>\<^sub>1 O gtt_lang \<G>\<^sub>2 \<subseteq> gtt_lang (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2)"
using gmctxt_cl_mono_rel[OF gtt_comp_acomplete, of UNIV \<G>\<^sub>1 \<G>\<^sub>2]
by (simp only: gcomp_rel[symmetric])
lemma gtt_comp_alang:
assumes "gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
shows "agtt_lang (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2) = gcomp_rel UNIV (agtt_lang \<G>\<^sub>1) (agtt_lang \<G>\<^sub>2)"
by (intro equalityI gtt_comp_asound[OF assms] gtt_comp_acomplete)
lemma gtt_comp_lang:
assumes "gtt_states \<G>\<^sub>1 |\<inter>| gtt_states \<G>\<^sub>2 = {||}"
shows "gtt_lang (GTT_comp \<G>\<^sub>1 \<G>\<^sub>2) = gtt_lang \<G>\<^sub>1 O gtt_lang \<G>\<^sub>2"
by (simp only: arg_cong[OF gtt_comp_alang[OF assms], of "gmctxt_cl UNIV"] gcomp_rel)
abbreviation GTT_comp' where
"GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2 \<equiv> GTT_comp (fmap_states_gtt Inl \<G>\<^sub>1) (fmap_states_gtt Inr \<G>\<^sub>2)"
lemma gtt_comp'_alang:
shows "agtt_lang (GTT_comp' \<G>\<^sub>1 \<G>\<^sub>2) = gcomp_rel UNIV (agtt_lang \<G>\<^sub>1) (agtt_lang \<G>\<^sub>2)"
proof -
have [simp]: "finj_on Inl (gtt_states \<G>\<^sub>1)" "finj_on Inr (gtt_states \<G>\<^sub>2)"
by (auto simp add: finj_on.rep_eq)
then show ?thesis
by (subst gtt_comp_alang) (auto simp: agtt_lang_fmap_states_gtt)
qed
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/GTT_Transitive_Closure.thy b/thys/Regular_Tree_Relations/GTT_Transitive_Closure.thy
--- a/thys/Regular_Tree_Relations/GTT_Transitive_Closure.thy
+++ b/thys/Regular_Tree_Relations/GTT_Transitive_Closure.thy
@@ -1,286 +1,286 @@
theory GTT_Transitive_Closure
imports GTT_Compose
begin
subsection \<open>GTT closure under transitivity\<close>
inductive_set \<Delta>_trancl_set :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) set" for A B where
\<Delta>_set_cong: "TA_rule f ps p |\<in>| rules A \<Longrightarrow> TA_rule f qs q |\<in>| rules B \<Longrightarrow> length ps = length qs \<Longrightarrow>
(\<And>i. i < length qs \<Longrightarrow> (ps ! i, qs ! i) \<in> \<Delta>_trancl_set A B) \<Longrightarrow> (p, q) \<in> \<Delta>_trancl_set A B"
| \<Delta>_set_eps1: "(p, p') |\<in>| eps A \<Longrightarrow> (p, q) \<in> \<Delta>_trancl_set A B \<Longrightarrow> (p', q) \<in> \<Delta>_trancl_set A B"
| \<Delta>_set_eps2: "(q, q') |\<in>| eps B \<Longrightarrow> (p, q) \<in> \<Delta>_trancl_set A B \<Longrightarrow> (p, q') \<in> \<Delta>_trancl_set A B"
| \<Delta>_set_trans: "(p, q) \<in> \<Delta>_trancl_set A B \<Longrightarrow> (q, r) \<in> \<Delta>_trancl_set A B \<Longrightarrow> (p, r) \<in> \<Delta>_trancl_set A B"
lemma \<Delta>_trancl_set_states: "\<Delta>_trancl_set \<A> \<B> \<subseteq> fset (\<Q> \<A> |\<times>| \<Q> \<B>)"
proof -
{fix p q assume "(p, q) \<in> \<Delta>_trancl_set \<A> \<B>" then have "(p, q) \<in> fset (\<Q> \<A> |\<times>| \<Q> \<B>)"
- by (induct) (auto dest: rule_statesD eps_statesD simp flip: fmember_iff_member_fset)}
+ by (induct) (auto dest: rule_statesD eps_statesD)}
then show ?thesis by auto
qed
lemma finite_\<Delta>_trancl_set [simp]: "finite (\<Delta>_trancl_set \<A> \<B>)"
using finite_subset[OF \<Delta>_trancl_set_states]
by simp
context
includes fset.lifting
begin
lift_definition \<Delta>_trancl :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) fset" is \<Delta>_trancl_set by simp
lemmas \<Delta>_trancl_cong = \<Delta>_set_cong [Transfer.transferred]
lemmas \<Delta>_trancl_eps1 = \<Delta>_set_eps1 [Transfer.transferred]
lemmas \<Delta>_trancl_eps2 = \<Delta>_set_eps2 [Transfer.transferred]
lemmas \<Delta>_trancl_cases = \<Delta>_trancl_set.cases[Transfer.transferred]
lemmas \<Delta>_trancl_induct [consumes 1, case_names \<Delta>_cong \<Delta>_eps1 \<Delta>_eps2 \<Delta>_trans] = \<Delta>_trancl_set.induct[Transfer.transferred]
lemmas \<Delta>_trancl_intros = \<Delta>_trancl_set.intros[Transfer.transferred]
lemmas \<Delta>_trancl_simps = \<Delta>_trancl_set.simps[Transfer.transferred]
end
lemma \<Delta>_trancl_cl [simp]:
"(\<Delta>_trancl A B)|\<^sup>+| = \<Delta>_trancl A B"
proof -
{fix s t assume "(s, t) |\<in>| (\<Delta>_trancl A B)|\<^sup>+|" then have "(s, t) |\<in>| \<Delta>_trancl A B"
by (induct rule: ftrancl_induct) (auto intro: \<Delta>_trancl_intros)}
then show ?thesis by auto
qed
lemma \<Delta>_trancl_states: "\<Delta>_trancl \<A> \<B> |\<subseteq>| (\<Q> \<A> |\<times>| \<Q> \<B>)"
using \<Delta>_trancl_set_states
by (metis \<Delta>_trancl.rep_eq fSigma_cong less_eq_fset.rep_eq)
definition GTT_trancl where
"GTT_trancl G =
(let \<Delta> = \<Delta>_trancl (snd G) (fst G) in
(TA (rules (fst G)) (eps (fst G) |\<union>| \<Delta>),
TA (rules (snd G)) (eps (snd G) |\<union>| (\<Delta>|\<inverse>|))))"
lemma \<Delta>_trancl_inv:
"(\<Delta>_trancl A B)|\<inverse>| = \<Delta>_trancl B A"
proof -
have [dest]: "(p, q) |\<in>| \<Delta>_trancl A B \<Longrightarrow> (q, p) |\<in>| \<Delta>_trancl B A" for p q A B
by (induct rule: \<Delta>_trancl_induct) (auto intro: \<Delta>_trancl_intros)
show ?thesis by auto
qed
lemma gtt_states_GTT_trancl:
"gtt_states (GTT_trancl G) |\<subseteq>| gtt_states G"
unfolding GTT_trancl_def
- by (auto simp: gtt_states_def \<Q>_def \<Delta>_trancl_inv dest!: fsubsetD[OF \<Delta>_trancl_states] simp flip: fmember_iff_member_fset)
+ by (auto simp: gtt_states_def \<Q>_def \<Delta>_trancl_inv dest!: fsubsetD[OF \<Delta>_trancl_states])
lemma gtt_syms_GTT_trancl:
"gtt_syms (GTT_trancl G) = gtt_syms G"
by (auto simp: GTT_trancl_def ta_sig_def \<Delta>_trancl_inv)
lemma GTT_trancl_base:
"gtt_lang G \<subseteq> gtt_lang (GTT_trancl G)"
using gtt_lang_mono[of G "GTT_trancl G"] by (auto simp: \<Delta>_trancl_inv GTT_trancl_def)
lemma GTT_trancl_trans:
"gtt_lang (GTT_comp (GTT_trancl G) (GTT_trancl G)) \<subseteq> gtt_lang (GTT_trancl G)"
proof -
have [dest]: "(p, q) |\<in>| \<Delta>\<^sub>\<epsilon> (TA (rules A) (eps A |\<union>| (\<Delta>_trancl B A)))
(TA (rules B) (eps B |\<union>| (\<Delta>_trancl A B))) \<Longrightarrow> (p, q) |\<in>| \<Delta>_trancl A B" for p q A B
by (induct rule: \<Delta>\<^sub>\<epsilon>_induct) (auto intro: \<Delta>_trancl_intros simp: \<Delta>_trancl_inv[of B A, symmetric])
show ?thesis
by (intro gtt_lang_mono[of "GTT_comp (GTT_trancl G) (GTT_trancl G)" "GTT_trancl G"])
- (auto simp: GTT_comp_def GTT_trancl_def fmember.abs_eq \<Delta>_trancl_inv)
+ (auto simp: GTT_comp_def GTT_trancl_def \<Delta>_trancl_inv)
qed
lemma agtt_lang_base:
"agtt_lang G \<subseteq> agtt_lang (GTT_trancl G)"
by (rule agtt_lang_mono) (auto simp: GTT_trancl_def \<Delta>_trancl_inv)
lemma \<Delta>\<^sub>\<epsilon>_tr_incl:
"\<Delta>\<^sub>\<epsilon> (TA (rules A) (eps A |\<union>| \<Delta>_trancl B A)) (TA (rules B) (eps B |\<union>| \<Delta>_trancl A B)) = \<Delta>_trancl A B"
(is "?LS = ?RS")
proof -
{fix p q assume "(p, q) |\<in>| ?LS" then have "(p, q) |\<in>| ?RS"
by (induct rule: \<Delta>\<^sub>\<epsilon>_induct)
(auto simp: \<Delta>_trancl_inv[of B A, symmetric] intro: \<Delta>_trancl_intros)}
moreover
{fix p q assume "(p, q) |\<in>| ?RS" then have "(p, q) |\<in>| ?LS"
by (induct rule: \<Delta>_trancl_induct)
(auto simp: \<Delta>_trancl_inv[of B A, symmetric] intro: \<Delta>\<^sub>\<epsilon>_intros)}
ultimately show ?thesis
by auto
qed
lemma agtt_lang_trans:
"gcomp_rel UNIV (agtt_lang (GTT_trancl G)) (agtt_lang (GTT_trancl G)) \<subseteq> agtt_lang (GTT_trancl G)"
proof -
have [intro!, dest]: "(p, q) |\<in>| \<Delta>\<^sub>\<epsilon> (TA (rules A) (eps A |\<union>| (\<Delta>_trancl B A)))
(TA (rules B) (eps B |\<union>| (\<Delta>_trancl A B))) \<Longrightarrow> (p, q) |\<in>| \<Delta>_trancl A B" for p q A B
by (induct rule: \<Delta>\<^sub>\<epsilon>_induct) (auto intro: \<Delta>_trancl_intros simp: \<Delta>_trancl_inv[of B A, symmetric])
show ?thesis
by (rule subset_trans[OF gtt_comp_acomplete agtt_lang_mono])
(auto simp: GTT_comp_def GTT_trancl_def \<Delta>_trancl_inv)
qed
lemma GTT_trancl_acomplete:
"gtrancl_rel UNIV (agtt_lang G) \<subseteq> agtt_lang (GTT_trancl G)"
unfolding gtrancl_rel_def
using agtt_lang_base[of G] gmctxt_cl_mono_rel[OF agtt_lang_base[of G], of UNIV]
using agtt_lang_trans[of G]
unfolding gcomp_rel_def
by (intro kleene_trancl_induct) blast+
lemma Restr_rtrancl_gtt_lang_eq_trancl_gtt_lang:
"(gtt_lang G)\<^sup>* = (gtt_lang G)\<^sup>+"
by (auto simp: rtrancl_trancl_reflcl simp del: reflcl_trancl dest: tranclD tranclD2 intro: gmctxt_cl_refl)
lemma GTT_trancl_complete:
"(gtt_lang G)\<^sup>+ \<subseteq> gtt_lang (GTT_trancl G)"
using GTT_trancl_base subset_trans[OF gtt_comp_lang_complete GTT_trancl_trans]
by (metis trancl_id trancl_mono_set trans_O_iff)
lemma trancl_gtt_lang_arg_closed:
assumes "length ss = length ts" "\<forall>i < length ts. (ss ! i, ts ! i) \<in> (gtt_lang \<G>)\<^sup>+"
shows "(GFun f ss, GFun f ts) \<in> (gtt_lang \<G>)\<^sup>+" (is "?e \<in> _")
proof -
have "all_ctxt_closed UNIV ((gtt_lang \<G>)\<^sup>+)" by (intro all_ctxt_closed_trancl) auto
from all_ctxt_closedD[OF this _ assms] show ?thesis
by auto
qed
lemma \<Delta>_trancl_sound:
assumes "(p, q) |\<in>| \<Delta>_trancl A B"
obtains s t where "(s, t) \<in> (gtt_lang (B, A))\<^sup>+" "p |\<in>| gta_der A s" "q |\<in>| gta_der B t"
using assms
proof (induct arbitrary: thesis rule: \<Delta>_trancl_induct)
case (\<Delta>_cong f ps p qs q)
have "\<exists>si ti. (si, ti) \<in> (gtt_lang (B, A))\<^sup>+ \<and> ps ! i |\<in>| gta_der A (si) \<and>
qs ! i |\<in>| gta_der B (ti)" if "i < length qs" for i
using \<Delta>_cong(5)[OF that] by metis
then obtain ss ts where
"\<And>i. i < length qs \<Longrightarrow> (ss i, ts i) \<in> (gtt_lang (B, A))\<^sup>+ \<and> ps ! i |\<in>| gta_der A (ss i) \<and> qs ! i |\<in>| gta_der B (ts i)" by metis
then show ?case using \<Delta>_cong(1-5)
by (intro \<Delta>_cong(6)[of "GFun f (map ss [0..<length ps])" "GFun f (map ts [0..<length qs])"])
(auto simp: gta_der_def intro!: trancl_gtt_lang_arg_closed)
next
case (\<Delta>_eps1 p p' q) then show ?case
by (metis gta_der_def ta_der_eps)
next
case (\<Delta>_eps2 q q' p) then show ?case
by (metis gta_der_def ta_der_eps)
next
case (\<Delta>_trans p q r)
obtain s1 t1 where "(s1, t1) \<in> (gtt_lang (B, A))\<^sup>+" "p |\<in>| gta_der A s1" "q |\<in>| gta_der B t1"
using \<Delta>_trans(2) .note 1 = this
obtain s2 t2 where "(s2, t2) \<in> (gtt_lang (B, A))\<^sup>+" "q |\<in>| gta_der A s2" "r |\<in>| gta_der B t2"
using \<Delta>_trans(4) . note 2 = this
have "(t1, s2) \<in> gtt_lang (B, A)" using 1(1,3) 2(1,2)
by (auto simp: Restr_rtrancl_gtt_lang_eq_trancl_gtt_lang[symmetric] gtt_lang_join)
then have "(s1, t2) \<in> (gtt_lang (B, A))\<^sup>+" using 1(1) 2(1)
by (meson trancl.trancl_into_trancl trancl_trans)
then show ?case using 1(2) 2(3) by (auto intro: \<Delta>_trans(5)[of s1 t2])
qed
lemma GTT_trancl_sound_aux:
assumes "p |\<in>| gta_der (TA (rules A) (eps A |\<union>| (\<Delta>_trancl B A))) s"
shows "\<exists>t. (s, t) \<in> (gtt_lang (A, B))\<^sup>+ \<and> p |\<in>| gta_der A t"
using assms
proof (induct s arbitrary: p)
case (GFun f ss)
let ?eps = "eps A |\<union>| \<Delta>_trancl B A"
obtain qs q where q: "TA_rule f qs q |\<in>| rules A" "q = p \<or> (q, p) |\<in>| ?eps|\<^sup>+|" "length qs = length ss"
"\<And>i. i < length ss \<Longrightarrow> qs ! i |\<in>| gta_der (TA (rules A) ?eps) (ss ! i)"
using GFun(2) by (auto simp: gta_der_def)
have "\<And>i. i < length ss \<Longrightarrow> \<exists>ti. (ss ! i, ti) \<in> (gtt_lang (A, B))\<^sup>+ \<and> qs ! i |\<in>| gta_der A (ti)"
using GFun(1)[OF nth_mem q(4)] unfolding gta_der_def by fastforce
then obtain ts where ts: "\<And>i. i < length ss \<Longrightarrow> (ss ! i, ts i) \<in> (gtt_lang (A, B))\<^sup>+ \<and> qs ! i |\<in>| gta_der A (ts i)"
by metis
then have q': "q |\<in>| gta_der A (GFun f (map ts [0..<length ss]))"
"(GFun f ss, GFun f (map ts [0..<length ss])) \<in> (gtt_lang (A, B))\<^sup>+" using q(1, 3)
by (auto simp: gta_der_def intro!: exI[of _ qs] exI[of _ q] trancl_gtt_lang_arg_closed)
{fix p q u assume ass: "(p, q) |\<in>| \<Delta>_trancl B A" "(GFun f ss, u) \<in> (gtt_lang (A, B))\<^sup>+ \<and> p |\<in>| gta_der A u"
from \<Delta>_trancl_sound[OF this(1)] obtain s t
where "(s, t) \<in> (gtt_lang (A, B))\<^sup>+" "p |\<in>| gta_der B s" "q |\<in>| gta_der A t" . note st = this
have "(u, s) \<in> gtt_lang (A, B)" using st conjunct2[OF ass(2)]
by (auto simp: Restr_rtrancl_gtt_lang_eq_trancl_gtt_lang[symmetric] gtt_lang_join)
then have "(GFun f ss, t) \<in> (gtt_lang (A, B))\<^sup>+"
using ass st(1) by (meson trancl_into_trancl2 trancl_trans)
then have "\<exists> s t. (GFun f ss, t) \<in> (gtt_lang (A, B))\<^sup>+ \<and> q |\<in>| gta_der A t" using st by blast}
note trancl_step = this
show ?case
proof (cases "q = p")
case True
then show ?thesis using ts q(1, 3)
by (auto simp: gta_der_def intro!: exI[of _"GFun f (map ts [0..< length ss])"] trancl_gtt_lang_arg_closed) blast
next
case False
then have "(q, p) |\<in>| ?eps|\<^sup>+|" using q(2) by simp
then show ?thesis using q(1) q'
proof (induct rule: ftrancl_induct)
case (Base q p) from Base(1) show ?case
proof
assume "(q, p) |\<in>| eps A" then show ?thesis using Base(2) ts q(3)
by (auto simp: gta_der_def intro!: exI[of _"GFun f (map ts [0..< length ss])"]
trancl_gtt_lang_arg_closed exI[of _ qs] exI[of _ q])
next
assume "(q, p) |\<in>| (\<Delta>_trancl B A)"
- then have "(q, p) |\<in>| \<Delta>_trancl B A" by (simp add: fmember.abs_eq)
+ then have "(q, p) |\<in>| \<Delta>_trancl B A" by simp
from trancl_step[OF this] show ?thesis using Base(3, 4)
by auto
qed
next
case (Step p q r)
from Step(2, 4-) obtain s' where s': "(GFun f ss, s') \<in> (gtt_lang (A, B))\<^sup>+ \<and> q |\<in>| gta_der A s'" by auto
show ?case using Step(3)
proof
assume "(q, r) |\<in>| eps A" then show ?thesis using s'
by (auto simp: gta_der_def ta_der_eps intro!: exI[of _ s'])
next
assume "(q, r) |\<in>| \<Delta>_trancl B A"
- then have "(q, r) |\<in>| \<Delta>_trancl B A" by (simp add: fmember.abs_eq)
+ then have "(q, r) |\<in>| \<Delta>_trancl B A" by simp
from trancl_step[OF this] show ?thesis using s' by auto
qed
qed
qed
qed
lemma GTT_trancl_asound:
"agtt_lang (GTT_trancl G) \<subseteq> gtrancl_rel UNIV (agtt_lang G)"
proof (intro subrelI, goal_cases LR)
case (LR s t)
then obtain s' q t' where *: "(s, s') \<in> (gtt_lang G)\<^sup>+"
"q |\<in>| gta_der (fst G) s'" "q |\<in>| gta_der (snd G) t'" "(t', t) \<in> (gtt_lang G)\<^sup>+"
by (auto simp: agtt_lang_def GTT_trancl_def trancl_converse \<Delta>_trancl_inv
simp flip: gtt_lang_swap[of "fst G" "snd G", unfolded prod.collapse agtt_lang_def, simplified]
dest!: GTT_trancl_sound_aux)
then have "(s', t') \<in> agtt_lang G" using *(2,3)
by (auto simp: agtt_lang_def)
then show ?case using *(1,4) unfolding gtrancl_rel_def
by auto
qed
lemma GTT_trancl_sound:
"gtt_lang (GTT_trancl G) \<subseteq> (gtt_lang G)\<^sup>+"
proof -
note [dest] = GTT_trancl_sound_aux
have "gtt_accept (GTT_trancl G) s t \<Longrightarrow> (s, t) \<in> (gtt_lang G)\<^sup>+" for s t unfolding gtt_accept_def
proof (induct rule: gmctxt_cl.induct)
case (base s t)
from base obtain q where join: "q |\<in>| gta_der (fst (GTT_trancl G)) s" "q |\<in>| gta_der (snd (GTT_trancl G)) t"
by (auto simp: agtt_lang_def)
obtain s' where "(s, s') \<in> (gtt_lang G)\<^sup>+" "q |\<in>| gta_der (fst G) s'" using base join
by (auto simp: GTT_trancl_def \<Delta>_trancl_inv agtt_lang_def)
moreover obtain t' where "(t', t) \<in> (gtt_lang G)\<^sup>+" "q |\<in>| gta_der (snd G) t'" using join
by (auto simp: GTT_trancl_def gtt_lang_swap[of "fst G" "snd G", symmetric] trancl_converse \<Delta>_trancl_inv)
moreover have "(s', t') \<in> gtt_lang G" using calculation
by (auto simp: Restr_rtrancl_gtt_lang_eq_trancl_gtt_lang[symmetric] gtt_lang_join)
ultimately show "(s, t) \<in> (gtt_lang G)\<^sup>+" by (meson trancl.trancl_into_trancl trancl_trans)
qed (auto intro!: trancl_gtt_lang_arg_closed)
then show ?thesis by (auto simp: gtt_accept_def)
qed
lemma GTT_trancl_alang:
"agtt_lang (GTT_trancl G) = gtrancl_rel UNIV (agtt_lang G)"
using GTT_trancl_asound GTT_trancl_acomplete by blast
lemma GTT_trancl_lang:
"gtt_lang (GTT_trancl G) = (gtt_lang G)\<^sup>+"
using GTT_trancl_sound GTT_trancl_complete by blast
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Horn_Setup/Horn_Fset.thy b/thys/Regular_Tree_Relations/Horn_Setup/Horn_Fset.thy
--- a/thys/Regular_Tree_Relations/Horn_Setup/Horn_Fset.thy
+++ b/thys/Regular_Tree_Relations/Horn_Setup/Horn_Fset.thy
@@ -1,105 +1,105 @@
theory Horn_Fset
imports Horn_Inference FSet_Utils
begin
locale horn_fset_impl = horn +
fixes infer0_impl :: "'a list" and infer1_impl :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a list"
begin
lemma saturate_fold_simp [simp]:
"fold (\<lambda>xa. case_option None (f xa)) xs None = None"
by (induct xs) auto
lemma saturate_fold_mono [partial_function_mono]:
"option.mono_body (\<lambda>f. fold (\<lambda>x. case_option None (\<lambda>y. f (x, y))) xs b)"
unfolding monotone_def fun_ord_def flat_ord_def
proof (intro allI impI, induct xs arbitrary: b)
case (Cons a xs)
show ?case
using Cons(1)[OF Cons(2), of "x (a, the b)"] Cons(2)[rule_format, of "(a, the b)"]
by (cases b) auto
qed auto
partial_function (option) saturate_rec :: "'a \<Rightarrow> 'a fset \<Rightarrow> ('a fset) option" where
"saturate_rec x bs = (if x |\<in>| bs then Some bs else
fold (\<lambda>x. case_option None (saturate_rec x)) (infer1_impl x bs) (Some (finsert x bs)))"
definition saturate_impl where
"saturate_impl = fold (\<lambda>x. case_option None (saturate_rec x)) infer0_impl (Some {||})"
end
locale horn_fset = horn_fset_impl +
assumes infer0: "infer0 = set infer0_impl"
and infer1: "\<And>x bs. infer1 x (fset bs) = set (infer1_impl x bs)"
begin
lemma saturate_rec_sound:
"saturate_rec x bs = Some bs' \<Longrightarrow> ({x}, fset bs) \<turnstile> ({}, fset bs')"
proof (induct arbitrary: x bs bs' rule: saturate_rec.fixp_induct)
case 1 show ?case using option_admissible[of "\<lambda>(x, y) z. _ x y z"]
by fastforce
next
case (3 rec)
have [dest!]: "(set xs, fset ys) \<turnstile> ({}, fset bs')"
if "fold (\<lambda>x a. case a of None \<Rightarrow> None | Some a \<Rightarrow> rec x a) xs (Some ys) = Some bs'"
for xs ys using that
proof (induct xs arbitrary: ys)
case (Cons a xs)
show ?case using trans[OF step_mono[OF 3(1)], of a ys _ "set xs" "{}" "fset bs'"] Cons
by (cases "rec a ys") auto
qed (auto intro: refl)
show ?case using propagate[of x "{}" "fset bs", unfolded infer1 Un_empty_left] 3(2)
- by (auto simp: delete fmember_iff_member_fset split: if_splits intro: trans delete)
+ by (auto simp: delete split: if_splits intro: trans delete)
qed auto
lemma saturate_impl_sound:
assumes "saturate_impl = Some B'"
shows "fset B' = saturate"
proof -
have "(set xs, fset ys) \<turnstile> ({}, fset bs')"
if "fold (\<lambda>x a. case a of None \<Rightarrow> None | Some a \<Rightarrow> saturate_rec x a) xs (Some ys) = Some bs'"
for xs ys bs' using that
proof (induct xs arbitrary: ys)
case (Cons a xs)
show ?case
using trans[OF step_mono[OF saturate_rec_sound], of a ys _ "set xs" "{}" "fset bs'"] Cons
by (cases "saturate_rec a ys") auto
qed (auto intro: refl)
from this[of infer0_impl "{||}" B'] assms step_sound show ?thesis
by (auto simp: saturate_impl_def infer0)
qed
lemma saturate_impl_complete:
assumes "finite saturate"
shows "saturate_impl \<noteq> None"
proof -
have *: "fold (\<lambda>x. case_option None (saturate_rec x)) ds (Some bs) \<noteq> None"
if "fset bs \<subseteq> saturate" "set ds \<subseteq> saturate" for bs ds
using that
proof (induct "card (saturate - fset bs)" arbitrary: bs ds rule: less_induct)
case less
show ?case using less(3)
proof (induct ds)
case (Cons d ds)
have "infer1 d (fset bs) \<subseteq> saturate" using less(2) Cons(2)
unfolding infer1_def by (auto intro: saturate.infer)
moreover have "card (saturate - fset (finsert d bs)) < card (saturate - fset bs)" if "d \<notin> fset bs"
using Cons(2) assms that
by (metis DiffI Diff_insert card_Diff1_less finite_Diff finsert.rep_eq in_mono insertCI list.simps(15))
ultimately show ?case using less(1)[of "finsert d bs" "infer1_impl d bs @ ds"] less(2) Cons assms
unfolding fold.simps comp_def option.simps
apply (subst saturate_rec.simps)
apply (auto simp flip: saturate_rec.simps split!: if_splits simp: infer1)
- apply (simp add: notin_fset saturate_rec.simps)
+ apply (simp add: saturate_rec.simps)
done
qed simp
qed
show ?thesis using *[of "{||}" "infer0_impl"] inv_start by (simp add: saturate_impl_def infer0)
qed
end
lemmas [code] = horn_fset_impl.saturate_rec.simps horn_fset_impl.saturate_impl_def
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Pair_Automaton.thy b/thys/Regular_Tree_Relations/Pair_Automaton.thy
--- a/thys/Regular_Tree_Relations/Pair_Automaton.thy
+++ b/thys/Regular_Tree_Relations/Pair_Automaton.thy
@@ -1,370 +1,372 @@
theory Pair_Automaton
imports Tree_Automata_Complement GTT_Compose
begin
subsection \<open>Pair automaton and anchored GTTs\<close>
definition pair_at_lang :: "('q, 'f) gtt \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> 'f gterm rel" where
"pair_at_lang \<G> Q = {(s, t) | s t p q. q |\<in>| gta_der (fst \<G>) s \<and> p |\<in>| gta_der (snd \<G>) t \<and> (q, p) |\<in>| Q}"
lemma pair_at_lang_restr_states:
"pair_at_lang \<G> Q = pair_at_lang \<G> (Q |\<inter>| (\<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)))"
by (auto simp: pair_at_lang_def gta_der_def) (meson gterm_ta_der_states)
lemma pair_at_langE:
assumes "(s, t) \<in> pair_at_lang \<G> Q"
obtains q p where "(q, p) |\<in>| Q" and "q |\<in>| gta_der (fst \<G>) s" and "p |\<in>| gta_der (snd \<G>) t"
using assms by (auto simp: pair_at_lang_def)
lemma pair_at_langI:
assumes "q |\<in>| gta_der (fst \<G>) s" "p |\<in>| gta_der (snd \<G>) t" "(q, p) |\<in>| Q"
shows "(s, t) \<in> pair_at_lang \<G> Q"
using assms by (auto simp: pair_at_lang_def)
lemma pair_at_lang_fun_states:
assumes "finj_on f (\<Q> (fst \<G>))" and "finj_on g (\<Q> (snd \<G>))"
and "Q |\<subseteq>| \<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)"
shows "pair_at_lang \<G> Q = pair_at_lang (map_prod (fmap_states_ta f) (fmap_states_ta g) \<G>) (map_prod f g |`| Q)"
(is "?LS = ?RS")
proof
{fix s t assume "(s, t) \<in> ?LS"
then have "(s, t) \<in> ?RS" using ta_der_fmap_states_ta_mono[of f "fst \<G>" s]
using ta_der_fmap_states_ta_mono[of g "snd \<G>" t]
by (force simp: gta_der_def map_prod_def image_iff elim!: pair_at_langE split: prod.split intro!: pair_at_langI)}
then show "?LS \<subseteq> ?RS" by auto
next
{fix s t assume "(s, t) \<in> ?RS"
then obtain p q where rs: "p |\<in>| ta_der (fst \<G>) (term_of_gterm s)" "f p |\<in>| ta_der (fmap_states_ta f (fst \<G>)) (term_of_gterm s)" and
ts: "q |\<in>| ta_der (snd \<G>) (term_of_gterm t)" "g q |\<in>| ta_der (fmap_states_ta g (snd \<G>)) (term_of_gterm t)" and
st: "(f p, g q) |\<in>| (map_prod f g |`| Q)" using assms ta_der_fmap_states_inv[of f "fst \<G>" _ s]
using ta_der_fmap_states_inv[of g "snd \<G>" _ t]
by (auto simp: gta_der_def adapt_vars_term_of_gterm elim!: pair_at_langE)
- (metis (no_types, opaque_lifting) fimageE fmap_prod_fimageI ta_der_fmap_states_conv)
+ (metis (no_types, opaque_lifting) f_the_finv_into_f fimage.rep_eq fmap_prod_fimageI
+ fmap_states gterm_ta_der_states)
then have "p |\<in>| \<Q> (fst \<G>)" "q |\<in>| \<Q> (snd \<G>)" by auto
then have "(p, q) |\<in>| Q" using assms st unfolding fimage_iff fBex_def
by (auto dest!: fsubsetD simp: finj_on_eq_iff)
then have "(s, t) \<in> ?LS" using st rs(1) ts(1) by (auto simp: gta_der_def intro!: pair_at_langI)}
then show "?RS \<subseteq> ?LS" by auto
qed
lemma converse_pair_at_lang:
"(pair_at_lang \<G> Q)\<inverse> = pair_at_lang (prod.swap \<G>) (Q|\<inverse>|)"
by (auto simp: pair_at_lang_def)
lemma pair_at_agtt:
"agtt_lang \<G> = pair_at_lang \<G> (fId_on (gtt_interface \<G>))"
by (auto simp: agtt_lang_def gtt_interface_def pair_at_lang_def gtt_states_def gta_der_def fId_on_iff)
definition \<Delta>_eps_pair where
"\<Delta>_eps_pair \<G>\<^sub>1 Q\<^sub>1 \<G>\<^sub>2 Q\<^sub>2 \<equiv> Q\<^sub>1 |O| \<Delta>\<^sub>\<epsilon> (snd \<G>\<^sub>1) (fst \<G>\<^sub>2) |O| Q\<^sub>2"
lemma pair_comp_sound1:
assumes "(s, t) \<in> pair_at_lang \<G>\<^sub>1 Q\<^sub>1"
and "(t, u) \<in> pair_at_lang \<G>\<^sub>2 Q\<^sub>2"
shows "(s, u) \<in> pair_at_lang (fst \<G>\<^sub>1, snd \<G>\<^sub>2) (\<Delta>_eps_pair \<G>\<^sub>1 Q\<^sub>1 \<G>\<^sub>2 Q\<^sub>2)"
proof -
from pair_at_langE assms obtain p q q' r where
wit: "(p, q) |\<in>| Q\<^sub>1" "p |\<in>| gta_der (fst \<G>\<^sub>1) s" "q |\<in>| gta_der (snd \<G>\<^sub>1) t"
"(q', r) |\<in>| Q\<^sub>2" "q' |\<in>| gta_der (fst \<G>\<^sub>2) t" "r |\<in>| gta_der (snd \<G>\<^sub>2) u"
by metis
from wit(3, 5) have "(q, q') |\<in>| \<Delta>\<^sub>\<epsilon> (snd \<G>\<^sub>1) (fst \<G>\<^sub>2)"
by (auto simp: \<Delta>\<^sub>\<epsilon>_def' gta_der_def intro!: exI[of _ "term_of_gterm t"])
then have "(p, r) |\<in>| \<Delta>_eps_pair \<G>\<^sub>1 Q\<^sub>1 \<G>\<^sub>2 Q\<^sub>2" using wit(1, 4)
by (auto simp: \<Delta>_eps_pair_def)
then show ?thesis using wit(2, 6) unfolding pair_at_lang_def
by auto
qed
lemma pair_comp_sound2:
assumes "(s, u) \<in> pair_at_lang (fst \<G>\<^sub>1, snd \<G>\<^sub>2) (\<Delta>_eps_pair \<G>\<^sub>1 Q\<^sub>1 \<G>\<^sub>2 Q\<^sub>2)"
shows "\<exists> t. (s, t) \<in> pair_at_lang \<G>\<^sub>1 Q\<^sub>1 \<and> (t, u) \<in> pair_at_lang \<G>\<^sub>2 Q\<^sub>2"
using assms unfolding pair_at_lang_def \<Delta>_eps_pair_def
by (auto simp: \<Delta>\<^sub>\<epsilon>_def' gta_der_def) (metis gterm_of_term_inv)
lemma pair_comp_sound:
"pair_at_lang \<G>\<^sub>1 Q\<^sub>1 O pair_at_lang \<G>\<^sub>2 Q\<^sub>2 = pair_at_lang (fst \<G>\<^sub>1, snd \<G>\<^sub>2) (\<Delta>_eps_pair \<G>\<^sub>1 Q\<^sub>1 \<G>\<^sub>2 Q\<^sub>2)"
by (auto simp: pair_comp_sound1 pair_comp_sound2 relcomp.simps)
inductive_set \<Delta>_Atrans_set :: "('q \<times> 'q) fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) set" for Q \<A> \<B> where
base [simp]: "(p, q) |\<in>| Q \<Longrightarrow> (p, q) \<in> \<Delta>_Atrans_set Q \<A> \<B>"
| step [intro]: "(p, q) \<in> \<Delta>_Atrans_set Q \<A> \<B> \<Longrightarrow> (q, r) |\<in>| \<Delta>\<^sub>\<epsilon> \<B> \<A> \<Longrightarrow>
(r, v) \<in> \<Delta>_Atrans_set Q \<A> \<B> \<Longrightarrow> (p, v) \<in> \<Delta>_Atrans_set Q \<A> \<B>"
lemma \<Delta>_Atrans_set_states:
"(p, q) \<in> \<Delta>_Atrans_set Q \<A> \<B> \<Longrightarrow> (p, q) \<in> fset ((fst |`| Q |\<union>| \<Q> \<A>) |\<times>| (snd |`| Q |\<union>| \<Q> \<B>))"
- by (induct rule: \<Delta>_Atrans_set.induct) (auto simp: fimage_iff fBex_def simp flip: fmember_iff_member_fset)
+ by (induct rule: \<Delta>_Atrans_set.induct) (auto simp: image_iff intro!: bexI)
lemma finite_\<Delta>_Atrans_set: "finite (\<Delta>_Atrans_set Q \<A> \<B>)"
proof -
have "\<Delta>_Atrans_set Q \<A> \<B> \<subseteq> fset ((fst |`| Q |\<union>| \<Q> \<A>) |\<times>| (snd |`| Q |\<union>| \<Q> \<B>))"
- using \<Delta>_Atrans_set_states by auto
+ using \<Delta>_Atrans_set_states
+ by (metis subrelI)
from finite_subset[OF this] show ?thesis by simp
qed
context
includes fset.lifting
begin
lift_definition \<Delta>_Atrans :: "('q \<times> 'q) fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) fset" is \<Delta>_Atrans_set
by (simp add: finite_\<Delta>_Atrans_set)
lemmas \<Delta>_Atrans_base [simp] = \<Delta>_Atrans_set.base [Transfer.transferred]
lemmas \<Delta>_Atrans_step [intro] = \<Delta>_Atrans_set.step [Transfer.transferred]
lemmas \<Delta>_Atrans_cases = \<Delta>_Atrans_set.cases[Transfer.transferred]
lemmas \<Delta>_Atrans_induct [consumes 1, case_names base step] = \<Delta>_Atrans_set.induct[Transfer.transferred]
end
abbreviation "\<Delta>_Atrans_gtt \<G> Q \<equiv> \<Delta>_Atrans Q (fst \<G>) (snd \<G>)"
lemma pair_trancl_sound1:
assumes "(s, t) \<in> (pair_at_lang \<G> Q)\<^sup>+"
shows "\<exists> q p. p |\<in>| gta_der (fst \<G>) s \<and> q |\<in>| gta_der (snd \<G>) t \<and> (p, q) |\<in>| \<Delta>_Atrans_gtt \<G> Q"
using assms
proof (induct)
case (step t v)
obtain p q r r' where reach_t: "r |\<in>| gta_der (fst \<G>) t" "q |\<in>| gta_der (snd \<G>) t" and
reach: "p |\<in>| gta_der (fst \<G>) s" "r' |\<in>| gta_der (snd \<G>) v" and
st: "(p, q) |\<in>| \<Delta>_Atrans_gtt \<G> Q" "(r, r') |\<in>| Q" using step(2, 3)
by (auto simp: pair_at_lang_def)
from reach_t have "(q, r) |\<in>| \<Delta>\<^sub>\<epsilon> (snd \<G>) (fst \<G>)"
by (auto simp: \<Delta>\<^sub>\<epsilon>_def' gta_der_def intro: ground_term_of_gterm)
then have "(p, r') |\<in>| \<Delta>_Atrans_gtt \<G> Q" using st by auto
then show ?case using reach reach_t
by (auto simp: pair_at_lang_def gta_der_def \<Delta>\<^sub>\<epsilon>_def' intro: ground_term_of_gterm)
qed (auto simp: pair_at_lang_def intro: \<Delta>_Atrans_base)
lemma pair_trancl_sound2:
assumes "(p, q) |\<in>| \<Delta>_Atrans_gtt \<G> Q"
and "p |\<in>| gta_der (fst \<G>) s" "q |\<in>| gta_der (snd \<G>) t"
shows "(s, t) \<in> (pair_at_lang \<G> Q)\<^sup>+" using assms
proof (induct arbitrary: s t rule:\<Delta>_Atrans_induct)
case (step p q r v)
from step(2)[OF step(6)] step(5)[OF _ step(7)] step(3)
show ?case by (auto simp: gta_der_def \<Delta>\<^sub>\<epsilon>_def' intro!: ground_term_of_gterm)
(metis gterm_of_term_inv trancl_trans)
qed (auto simp: pair_at_lang_def)
lemma pair_trancl_sound:
"(pair_at_lang \<G> Q)\<^sup>+ = pair_at_lang \<G> (\<Delta>_Atrans_gtt \<G> Q)"
by (auto simp: pair_trancl_sound2 dest: pair_trancl_sound1 elim: pair_at_langE intro: pair_at_langI)
abbreviation "fst_pair_cl \<A> Q \<equiv> TA (rules \<A>) (eps \<A> |\<union>| (fId_on (\<Q> \<A>) |O| Q))"
definition pair_at_to_agtt :: "('q, 'f) gtt \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> ('q, 'f) gtt" where
"pair_at_to_agtt \<G> Q = (fst_pair_cl (fst \<G>) Q , TA (rules (snd \<G>)) (eps (snd \<G>)))"
lemma fst_pair_cl_eps:
assumes "(p, q) |\<in>| (eps (fst_pair_cl \<A> Q))|\<^sup>+|"
and "\<Q> \<A> |\<inter>| snd |`| Q = {||}"
shows "(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<or> (\<exists> r. (p = r \<or> (p, r) |\<in>| (eps \<A>)|\<^sup>+|) \<and> (r, q) |\<in>| Q)" using assms
proof (induct rule: ftrancl_induct)
case (Step p q r)
then have y: "q |\<in>| \<Q> \<A>" by (auto simp add: eps_trancl_statesD eps_statesD)
have [simp]: "(p, q) |\<in>| Q \<Longrightarrow> q |\<in>| snd |`| Q" for p q by (auto simp: fimage_iff) force
then show ?case using Step y
by auto (simp add: ftrancl_into_trancl)
qed auto
lemma fst_pair_cl_res_aux:
assumes "\<Q> \<A> |\<inter>| snd |`| Q = {||}"
and "q |\<in>| ta_der (fst_pair_cl \<A> Q) (term_of_gterm t)"
shows "\<exists> p. p |\<in>| ta_der \<A> (term_of_gterm t) \<and> (q |\<notin>| \<Q> \<A> \<longrightarrow> (p, q) |\<in>| Q) \<and> (q |\<in>| \<Q> \<A> \<longrightarrow> p = q)" using assms
proof (induct t arbitrary: q)
case (GFun f ts)
then obtain qs q' where rule: "TA_rule f qs q' |\<in>| rules \<A>" "length qs = length ts" and
eps: "q' = q \<or> (q', q) |\<in>| (eps (fst_pair_cl \<A> Q))|\<^sup>+|" and
reach: "\<forall> i < length ts. qs ! i |\<in>| ta_der (fst_pair_cl \<A> Q) (term_of_gterm (ts ! i))"
by auto
{fix i assume ass: "i < length ts" then have st: "qs ! i |\<in>| \<Q> \<A>" using rule
by (auto simp: rule_statesD)
then have "qs ! i |\<notin>| snd |`| Q" using GFun(2) by auto
then have "qs ! i |\<in>| ta_der \<A> (term_of_gterm (ts ! i))" using reach st ass
using fst_pair_cl_eps[OF _ GFun(2)] GFun(1)[OF nth_mem[OF ass] GFun(2), of "qs ! i"]
by blast} note IH = this
show ?case
proof (cases "q' = q")
case True
then show ?thesis using rule reach IH
by (auto dest: rule_statesD intro!: exI[of _ q'] exI[of _ qs])
next
case False note nt_eq = this
then have eps: "(q', q) |\<in>| (eps (fst_pair_cl \<A> Q))|\<^sup>+|" using eps by simp
from fst_pair_cl_eps[OF this assms(1)] show ?thesis
using False rule IH
proof (cases "q |\<notin>| \<Q> \<A>")
case True
from fst_pair_cl_eps[OF eps assms(1)] obtain r where
"q' = r \<or> (q', r) |\<in>| (eps \<A>)|\<^sup>+|" "(r, q) |\<in>| Q" using True
by (auto simp: eps_trancl_statesD)
then show ?thesis using nt_eq rule IH True
by (auto simp: fimage_iff eps_trancl_statesD)
next
case False
from fst_pair_cl_eps[OF eps assms(1)] False assms(1)
have "(q', q) |\<in>| (eps \<A>)|\<^sup>+|"
by (auto simp: fimage_iff) (metis fempty_iff fimage_eqI finterI snd_conv)+
then show ?thesis using IH rule
by (intro exI[of _ q]) (auto simp: eps_trancl_statesD)
qed
qed
qed
lemma restr_distjoing:
assumes "Q |\<subseteq>| \<Q> \<A> |\<times>| \<Q> \<BB>"
and "\<Q> \<A> |\<inter>| \<Q> \<BB> = {||}"
shows "\<Q> \<A> |\<inter>| snd |`| Q = {||}"
using assms by auto
lemma pair_at_agtt_conv:
assumes "Q |\<subseteq>| \<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)" and "\<Q> (fst \<G>) |\<inter>| \<Q> (snd \<G>) = {||}"
shows "pair_at_lang \<G> Q = agtt_lang (pair_at_to_agtt \<G> Q)" (is "?LS = ?RS")
proof
let ?TA = "fst_pair_cl (fst \<G>) Q"
{fix s t assume ls: "(s, t) \<in> ?LS"
then obtain q p where w: "(q, p) |\<in>| Q" "q |\<in>| gta_der (fst \<G>) s" "p |\<in>| gta_der (snd \<G>) t"
by (auto elim: pair_at_langE)
from w(2) have "q |\<in>| gta_der ?TA s" "q |\<in>| \<Q> (fst \<G>)"
using ta_der_mono'[of "fst \<G>" ?TA "term_of_gterm s"]
by (auto simp add: fin_mono ta_subset_def gta_der_def in_mono)
then have "(s, t) \<in> ?RS" using w(1, 3)
by (auto simp: pair_at_to_agtt_def agtt_lang_def gta_der_def ta_der_eps intro!: exI[of _ p])
(metis fId_onI frelcompI funionI2 ta.sel(2) ta_der_eps)}
then show "?LS \<subseteq> ?RS" by auto
next
{fix s t assume ls: "(s, t) \<in> ?RS"
then obtain q where w: "q |\<in>| ta_der (fst_pair_cl (fst \<G>) Q) (term_of_gterm s)"
"q |\<in>| ta_der (snd \<G>) (term_of_gterm t)"
by (auto simp: agtt_lang_def pair_at_to_agtt_def gta_der_def)
from w(2) have "q |\<in>| \<Q> (snd \<G>)" "q |\<notin>| \<Q> (fst \<G>)" using assms(2)
by auto
from fst_pair_cl_res_aux[OF restr_distjoing[OF assms] w(1)] this w(2)
have "(s, t) \<in> ?LS" by (auto simp: agtt_lang_def pair_at_to_agtt_def gta_der_def intro: pair_at_langI)}
then show "?RS \<subseteq> ?LS" by auto
qed
definition pair_at_to_agtt' where
"pair_at_to_agtt' \<G> Q = (let \<A> = fmap_states_ta Inl (fst \<G>) in
let \<B> = fmap_states_ta Inr (snd \<G>) in
let Q' = Q |\<inter>| (\<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)) in
pair_at_to_agtt (\<A>, \<B>) (map_prod Inl Inr |`| Q'))"
lemma pair_at_agtt_cost:
"pair_at_lang \<G> Q = agtt_lang (pair_at_to_agtt' \<G> Q)"
proof -
let ?G = "(fmap_states_ta CInl (fst \<G>), fmap_states_ta CInr (snd \<G>))"
let ?Q = "(Q |\<inter>| (\<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)))"
let ?Q' = "map_prod CInl CInr |`| ?Q"
have *: "pair_at_lang \<G> Q = pair_at_lang \<G> ?Q"
using pair_at_lang_restr_states by blast
have "pair_at_lang \<G> ?Q = pair_at_lang (map_prod (fmap_states_ta CInl) (fmap_states_ta CInr) \<G>) (map_prod CInl CInr |`| ?Q)"
by (intro pair_at_lang_fun_states[where ?\<G> = \<G> and ?Q = ?Q and ?f = CInl and ?g = CInr])
(auto simp: finj_CInl_CInr)
then have **:"pair_at_lang \<G> ?Q = pair_at_lang ?G ?Q'" by (simp add: map_prod_simp')
have "pair_at_lang ?G ?Q' = agtt_lang (pair_at_to_agtt ?G ?Q')"
by (intro pair_at_agtt_conv[where ?\<G> = ?G]) auto
then show ?thesis unfolding * ** pair_at_to_agtt'_def Let_def
by simp
qed
lemma \<Delta>_Atrans_states_stable:
assumes "Q |\<subseteq>| \<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)"
shows "\<Delta>_Atrans_gtt \<G> Q |\<subseteq>| \<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)"
proof
fix s assume ass: "s |\<in>| \<Delta>_Atrans_gtt \<G> Q"
then obtain t u where s: "s = (t, u)" by (cases s) blast
show "s |\<in>| \<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)" using ass assms unfolding s
by (induct rule: \<Delta>_Atrans_induct) auto
qed
lemma \<Delta>_Atrans_map_prod:
assumes "finj_on f (\<Q> (fst \<G>))" and "finj_on g (\<Q> (snd \<G>))"
and "Q |\<subseteq>| \<Q> (fst \<G>) |\<times>| \<Q> (snd \<G>)"
shows "map_prod f g |`| (\<Delta>_Atrans_gtt \<G> Q) = \<Delta>_Atrans_gtt (map_prod (fmap_states_ta f) (fmap_states_ta g) \<G>) (map_prod f g |`| Q)"
(is "?LS = ?RS")
proof -
{fix p q assume "(p, q) |\<in>| \<Delta>_Atrans_gtt \<G> Q"
then have "(f p, g q) |\<in>| ?RS" using assms
proof (induct rule: \<Delta>_Atrans_induct)
case (step p q r v)
from step(3, 6, 7) have "(g q, f r) |\<in>| \<Delta>\<^sub>\<epsilon> (fmap_states_ta g (snd \<G>)) (fmap_states_ta f (fst \<G>))"
by (auto simp: \<Delta>\<^sub>\<epsilon>_def' intro!: ground_term_of_gterm)
(metis ground_term_of_gterm ground_term_to_gtermD ta_der_to_fmap_states_der)
then show ?case using step by auto
- qed (auto simp add: fmap_prod_fimageI)}
+ qed (auto simp add: map_prod_imageI)}
moreover
{fix p q assume "(p, q) |\<in>| ?RS"
then have "(p, q) |\<in>| ?LS" using assms
proof (induct rule: \<Delta>_Atrans_induct)
case (step p q r v)
let ?f = "the_finv_into (\<Q> (fst \<G>)) f" let ?g = "the_finv_into (\<Q> (snd \<G>)) g"
have sub: "\<Delta>\<^sub>\<epsilon> (snd \<G>) (fst \<G>) |\<subseteq>| \<Q> (snd \<G>) |\<times>| \<Q> (fst \<G>)"
using \<Delta>\<^sub>\<epsilon>_statesD(1, 2) by fastforce
have s_e: "(?f p, ?g q) |\<in>| \<Delta>_Atrans_gtt \<G> Q" "(?f r, ?g v) |\<in>| \<Delta>_Atrans_gtt \<G> Q"
using step assms(1, 2) fsubsetD[OF \<Delta>_Atrans_states_stable[OF assms(3)]]
using finj_on_eq_iff[OF assms(1)] finj_on_eq_iff
using the_finv_into_f_f[OF assms(1)] the_finv_into_f_f[OF assms(2)]
by auto
from step(3) have "(?g q, ?f r) |\<in>| \<Delta>\<^sub>\<epsilon> (snd \<G>) (fst \<G>)"
using step(6-) sub
using ta_der_fmap_states_conv[OF assms(1)] ta_der_fmap_states_conv[OF assms(2)]
using the_finv_into_f_f[OF assms(1)] the_finv_into_f_f[OF assms(2)]
by (auto simp: \<Delta>\<^sub>\<epsilon>_fmember fimage_iff fBex_def)
(metis ground_term_of_gterm ground_term_to_gtermD ta_der_fmap_states_inv)
then have "(q, r) |\<in>| map_prod g f |`| \<Delta>\<^sub>\<epsilon> (snd \<G>) (fst \<G>)" using step
using the_finv_into_f_f[OF assms(1)] the_finv_into_f_f[OF assms(2)] sub
- by auto (smt \<Delta>\<^sub>\<epsilon>_statesD(1, 2) f_the_finv_into_f fmap_prod_fimageI fmap_states)
+ by (smt (verit, ccfv_threshold) \<Delta>\<^sub>\<epsilon>_statesD(1) \<Delta>\<^sub>\<epsilon>_statesD(2) f_the_finv_into_f fimage.rep_eq
+ fmap_states fst_map_prod map_prod_imageI snd_map_prod)
then show ?case using s_e assms(1, 2) s_e
using fsubsetD[OF sub]
using fsubsetD[OF \<Delta>_Atrans_states_stable[OF assms(3)]]
using \<Delta>_Atrans_step[of "?f p" "?g q" Q "fst \<G>" "snd \<G>" "?f r" "?g v"]
using the_finv_into_f_f[OF assms(1)] the_finv_into_f_f[OF assms(2)]
- by (auto simp: fimage_iff fBex_def)
- (smt Pair_inject prod_fun_fimageE step.hyps(2) step.hyps(5) step.prems(3))
+ using step.hyps(2) step.hyps(5) step.prems(3) by force
qed auto}
ultimately show ?thesis by auto
qed
\<comment> \<open>Section: Pair Automaton is closed under Determinization\<close>
definition Q_pow where
"Q_pow Q \<S>\<^sub>1 \<S>\<^sub>2 =
{|(Wrapp X, Wrapp Y) | X Y p q. X |\<in>| fPow \<S>\<^sub>1 \<and> Y |\<in>| fPow \<S>\<^sub>2 \<and> p |\<in>| X \<and> q |\<in>| Y \<and> (p, q) |\<in>| Q|}"
lemma Q_pow_fmember:
"(X, Y) |\<in>| Q_pow Q \<S>\<^sub>1 \<S>\<^sub>2 \<longleftrightarrow> (\<exists> p q. ex X |\<in>| fPow \<S>\<^sub>1 \<and> ex Y |\<in>| fPow \<S>\<^sub>2 \<and> p |\<in>| ex X \<and> q |\<in>| ex Y \<and> (p, q) |\<in>| Q)"
proof -
let ?S = "{(Wrapp X, Wrapp Y) | X Y p q. X |\<in>| fPow \<S>\<^sub>1 \<and> Y |\<in>| fPow \<S>\<^sub>2 \<and> p |\<in>| X \<and> q |\<in>| Y \<and> (p, q) |\<in>| Q}"
- have "?S \<subseteq> map_prod Wrapp Wrapp ` fset (fPow \<S>\<^sub>1 |\<times>| fPow \<S>\<^sub>2)" by (auto simp flip: fmember_iff_member_fset)
+ have "?S \<subseteq> map_prod Wrapp Wrapp ` fset (fPow \<S>\<^sub>1 |\<times>| fPow \<S>\<^sub>2)" by auto
from finite_subset[OF this] show ?thesis unfolding Q_pow_def
apply auto apply blast
by (meson FSet_Lex_Wrapper.exhaust_sel)
qed
lemma pair_automaton_det_lang_sound_complete:
"pair_at_lang \<G> Q = pair_at_lang (map_both ps_ta \<G>) (Q_pow Q (\<Q> (fst \<G>)) (\<Q> (snd \<G>)))" (is "?LS = ?RS")
proof -
{fix s t assume "(s, t) \<in> ?LS"
then obtain p q where
res : "p |\<in>| ta_der (fst \<G>) (term_of_gterm s)"
"q |\<in>| ta_der (snd \<G>) (term_of_gterm t)" "(p, q) |\<in>| Q"
by (auto simp: pair_at_lang_def gta_der_def)
from ps_rules_complete[OF this(1)] ps_rules_complete[OF this(2)] this(3)
have "(s, t) \<in> ?RS" using fPow_iff ps_ta_states'
- by (auto simp: pair_at_lang_def gta_der_def Q_pow_fmember)
- force}
+ apply (auto simp: pair_at_lang_def gta_der_def Q_pow_fmember)
+ by (smt (verit, best) dual_order.trans ground_ta_der_states ground_term_of_gterm ps_rules_sound)}
moreover
{fix s t assume "(s, t) \<in> ?RS" then have "(s, t) \<in> ?LS"
using ps_rules_sound
by (auto simp: pair_at_lang_def gta_der_def ps_ta_def Let_def Q_pow_fmember) blast}
ultimately show ?thesis by auto
qed
lemma pair_automaton_complement_sound_complete:
assumes "partially_completely_defined_on \<A> \<F>" and "partially_completely_defined_on \<B> \<F>"
and "ta_det \<A>" and "ta_det \<B>"
shows "pair_at_lang (\<A>, \<B>) (\<Q> \<A> |\<times>| \<Q> \<B> |-| Q) = gterms (fset \<F>) \<times> gterms (fset \<F>) - pair_at_lang (\<A>, \<B>) Q"
using assms unfolding partially_completely_defined_on_def pair_at_lang_def
apply (auto simp: gta_der_def)
apply (metis ta_detE)
apply fastforce
done
end
diff --git a/thys/Regular_Tree_Relations/RR2_Infinite.thy b/thys/Regular_Tree_Relations/RR2_Infinite.thy
--- a/thys/Regular_Tree_Relations/RR2_Infinite.thy
+++ b/thys/Regular_Tree_Relations/RR2_Infinite.thy
@@ -1,507 +1,515 @@
theory RR2_Infinite
imports RRn_Automata Tree_Automata_Pumping
begin
lemma map_ta_rule_id [simp]: "map_ta_rule f id r = (r_root r) (map f (r_lhs_states r)) \<rightarrow> (f (r_rhs r))" for f r
by (simp add: ta_rule.expand ta_rule.map_sel(1 - 3))
(* Finitness/Infinitness facts *)
lemma no_upper_bound_infinite:
assumes "\<forall>(n::nat). \<exists>t \<in> S. n < f t"
shows "infinite S"
proof (rule ccontr, simp)
assume "finite S"
then obtain n where "n = Max (f ` S)" "\<forall> t \<in> S. f t \<le> n" by auto
then show False using assms linorder_not_le by blast
qed
lemma set_constr_finite:
assumes "finite F"
shows "finite {h x | x. x \<in> F \<and> P x}" using assms
by (induct) auto
lemma bounded_depth_finite:
assumes fin_F: "finite \<F>" and "\<Union> (funas_term ` S) \<subseteq> \<F>"
and "\<forall>t \<in> S. depth t \<le> n" and "\<forall>t \<in> S. ground t"
shows "finite S" using assms(2-)
proof (induction n arbitrary: S)
case 0
{fix t assume elem: "t \<in> S"
from 0 have "depth t = 0" "ground t" "funas_term t \<subseteq> \<F>" using elem by auto
then have "\<exists> f. (f, 0) \<in> \<F> \<and> t = Fun f []" by (cases t rule: depth.cases) auto}
then have "S \<subseteq> {Fun f [] |f . (f, 0) \<in> \<F>}" by (auto simp add: image_iff)
from finite_subset[OF this] show ?case
using set_constr_finite[OF fin_F, of "\<lambda> (f, n). Fun f []" "\<lambda> x. snd x = 0"]
by auto
next
case (Suc n)
from Suc obtain S' where
S: "S' = {t :: ('a, 'b) term . ground t \<and> funas_term t \<subseteq> \<F> \<and> depth t \<le> n}" "finite S'"
by (auto simp add: SUP_le_iff)
then obtain L F where L: "set L = S'" "set F = \<F>" using fin_F by (meson finite_list)
let ?Sn = "{Fun f ts | f ts. (f, length ts) \<in> \<F> \<and> set ts \<subseteq> S'}"
let ?Ln = "concat (map (\<lambda> (f, n). map (\<lambda> ts. Fun f ts) (List.n_lists n L)) F)"
{fix t assume elem: "t \<in> S"
from Suc have "depth t \<le> Suc n" "ground t" "funas_term t \<subseteq> \<F>" using elem by auto
then have "funas_term t \<subseteq> \<F> \<and> (\<forall> x \<in> set (args t). depth x \<le> n) \<and> ground t"
by (cases t rule: depth.cases) auto
then have "t \<in> ?Sn \<union> S'"
using S by (cases t) auto} note sub = this
{fix t assume elem: "t \<in> ?Sn"
then obtain f ts where [simp]: "t = Fun f ts" and invar: "(f, length ts) \<in> \<F>" "set ts \<subseteq> S'"
by blast
then have "Fun f ts \<in> set (map (\<lambda> ts. Fun f ts) (List.n_lists (length ts) L))" using L(1)
by (auto simp: image_iff set_n_lists)
then have "t \<in> set ?Ln" using invar(1) L(2) by auto}
from this sub have sub: "?Sn \<subseteq> set ?Ln" "S \<subseteq> ?Sn \<union> S'" by blast+
from finite_subset[OF sub(1)] finite_subset[OF sub(2)] finite_UnI[of ?Sn, OF _ S(2)]
show ?case by blast
qed
lemma infinite_imageD:
"infinite (f ` S) \<Longrightarrow> inj_on f S \<Longrightarrow> infinite S"
by blast
lemma infinite_imageD2:
"infinite (f ` S) \<Longrightarrow> inj f \<Longrightarrow> infinite S"
by blast
lemma infinite_inj_image_infinite:
assumes "infinite S" and "inj_on f S"
shows "infinite (f ` S)"
using assms finite_image_iff by blast
(*The following lemma tells us that number of terms greater than a certain depth are infinite*)
lemma infinte_no_depth_limit:
assumes "infinite S" and "finite \<F>"
and "\<forall>t \<in> S. funas_term t \<subseteq> \<F>" and "\<forall>t \<in> S. ground t"
shows "\<forall>(n::nat). \<exists>t \<in> S. n < (depth t)"
proof(rule allI, rule ccontr)
fix n::nat
assume "\<not> (\<exists>t \<in> S. (depth t) > n)"
hence "\<forall>t \<in> S. depth t \<le> n" by auto
from bounded_depth_finite[OF assms(2) _ this] show False using assms
by auto
qed
lemma depth_gterm_conv:
"depth (term_of_gterm t) = depth (term_of_gterm t)"
by (metis leD nat_neq_iff poss_gposs_conv poss_length_bounded_by_depth poss_length_depth)
lemma funs_term_ctxt [simp]:
"funs_term C\<langle>s\<rangle> = funs_ctxt C \<union> funs_term s"
by (induct C) auto
lemma pigeonhole_ta_infinit_terms:
fixes t ::"'f gterm" and \<A> :: "('q, 'f) ta"
defines "t' \<equiv> term_of_gterm t :: ('f, 'q) term"
assumes "fcard (\<Q> \<A>) < depth t'" and "q |\<in>| gta_der \<A> t" and "P (funas_gterm t)"
shows "infinite {t . q |\<in>| gta_der \<A> t \<and> P (funas_gterm t)}"
proof -
from pigeonhole_tree_automata[OF _ assms(3)[unfolded gta_der_def]] assms(2,4)
obtain C C2 s v p where ctxt: "C2 \<noteq> \<box>" "C\<langle>s\<rangle> = t'" "C2\<langle>v\<rangle> = s" and
loop: "p |\<in>| ta_der \<A> v" "p |\<in>| ta_der \<A> C2\<langle>Var p\<rangle>" "q |\<in>| ta_der \<A> C\<langle>Var p\<rangle>"
unfolding assms(1) by auto
let ?terms = "\<lambda> n. C\<langle>(C2 ^n)\<langle>v\<rangle>\<rangle>" let ?inner = "\<lambda> n. (C2 ^n)\<langle>v\<rangle>"
have gr: "ground_ctxt C2" "ground_ctxt C" "ground v"
using arg_cong[OF ctxt(2), of ground] unfolding ctxt(3)[symmetric] assms(1)
by fastforce+
moreover have funas: "funas_term (?terms (Suc n)) = funas_term t'" for n
unfolding ctxt(2, 3)[symmetric] using ctxt_comp_n_pres_funas by auto
moreover have der: "q |\<in>| ta_der \<A> (?terms (Suc n))" for n using loop
by (meson ta_der_ctxt ta_der_ctxt_n_loop)
moreover have "n < depth (?terms (Suc n))" for n
by (meson ctxt(1) ctxt_comp_n_lower_bound depth_ctxt_less_eq less_le_trans)
ultimately have "q |\<in>| ta_der \<A> (?terms (Suc n)) \<and> ground (?terms (Suc n)) \<and>
P (funas_term (?terms (Suc n))) \<and> n < depth (?terms (Suc n))" for n using assms(4)
by (auto simp: assms(1) funas_term_of_gterm_conv)
then have inf: "infinite {t. q |\<in>| ta_der \<A> t \<and> ground t \<and> P (funas_term t)}"
by (intro no_upper_bound_infinite[of _ depth]) blast
have inj: "inj_on gterm_of_term {t. q |\<in>| ta_der \<A> t \<and> ground t \<and> P (funas_term t)}"
by (intro gterm_of_term_inj) simp
show ?thesis
by (intro infinite_super[OF _ infinite_inj_image_infinite[OF inf inj]])
(auto simp: image_def gta_der_def funas_gterm_gterm_of_term)
qed
lemma gterm_to_None_Some_funas [simp]:
"funas_gterm (gterm_to_None_Some t) \<subseteq> (\<lambda> (f, n). ((None, Some f), n)) ` \<F> \<longleftrightarrow> funas_gterm t \<subseteq> \<F>"
by (induct t) (auto simp: funas_gterm_def, blast)
lemma funas_gterm_bot_some_decomp:
assumes "funas_gterm s \<subseteq> (\<lambda> (f, n). ((None, Some f), n)) ` \<F>"
shows "\<exists> t. gterm_to_None_Some t = s \<and> funas_gterm t \<subseteq> \<F>" using assms
proof (induct s)
case (GFun f ts)
from GFun(1)[OF nth_mem] obtain ss where l: "length ss = length ts \<and> (\<forall>i<length ts. gterm_to_None_Some (ss ! i) = ts ! i)"
using Ex_list_of_length_P[of "length ts" "\<lambda> s i. gterm_to_None_Some s = ts ! i"] GFun(2-)
by (auto simp: funas_gterm_def) (meson UN_subset_iff nth_mem)
then have "i < length ss \<Longrightarrow> funas_gterm (ss ! i) \<subseteq> \<F>" for i using GFun(2)
by (auto simp: UN_subset_iff) (smt (z3) gterm_to_None_Some_funas nth_mem subsetD)
then show ?case using GFun(2-) l
by (cases f) (force simp: map_nth_eq_conv UN_subset_iff dest!: in_set_idx intro!: exI[of _ "GFun (the (snd f)) ss"])
qed
(* Definition INF, Q_infinity and automata construction *)
definition "Inf_branching_terms \<R> \<F> = {t . infinite {u. (t, u) \<in> \<R> \<and> funas_gterm u \<subseteq> fset \<F>} \<and> funas_gterm t \<subseteq> fset \<F>}"
definition "Q_infty \<A> \<F> = {|q | q. infinite {t | t. funas_gterm t \<subseteq> fset \<F> \<and> q |\<in>| ta_der \<A> (term_of_gterm (gterm_to_None_Some t))}|}"
lemma Q_infty_fmember:
"q |\<in>| Q_infty \<A> \<F> \<longleftrightarrow> infinite {t | t. funas_gterm t \<subseteq> fset \<F> \<and> q |\<in>| ta_der \<A> (term_of_gterm (gterm_to_None_Some t))}"
proof -
have "{q | q. infinite {t | t. funas_gterm t \<subseteq> fset \<F> \<and> q |\<in>| ta_der \<A> (term_of_gterm (gterm_to_None_Some t))}} \<subseteq> fset (\<Q> \<A>)"
- using not_finite_existsD notin_fset by fastforce
+ using not_finite_existsD by fastforce
from finite_subset[OF this] show ?thesis
by (auto simp: Q_infty_def)
qed
abbreviation q_inf_dash_intro_rules where
"q_inf_dash_intro_rules Q r \<equiv> if (r_rhs r) |\<in>| Q \<and> fst (r_root r) = None then {|(r_root r) (map CInl (r_lhs_states r)) \<rightarrow> CInr (r_rhs r)|} else {||}"
abbreviation args :: "'a list \<Rightarrow> nat \<Rightarrow> ('a + 'a) list" where
"args \<equiv> \<lambda> qs i. map CInl (take i qs) @ CInr (qs ! i) # map CInl (drop (Suc i) qs)"
abbreviation q_inf_dash_closure_rules :: "('q, 'f) ta_rule \<Rightarrow> ('q + 'q, 'f) ta_rule list" where
"q_inf_dash_closure_rules r \<equiv> (let (f, qs, q) = (r_root r, r_lhs_states r, r_rhs r) in
(map (\<lambda> i. f (args qs i) \<rightarrow> CInr q) [0 ..< length qs]))"
definition Inf_automata :: "('q, 'f option \<times> 'f option) ta \<Rightarrow> 'q fset \<Rightarrow> ('q + 'q, 'f option \<times> 'f option) ta" where
"Inf_automata \<A> Q = TA
(( |\<Union>| (q_inf_dash_intro_rules Q |`| rules \<A>)) |\<union>| ( |\<Union>| ((fset_of_list \<circ> q_inf_dash_closure_rules) |`| rules \<A>)) |\<union>|
map_ta_rule CInl id |`| rules \<A>) (map_both Inl |`| eps \<A> |\<union>| map_both CInr |`| eps \<A>)"
definition Inf_reg where
"Inf_reg \<A> Q = Reg (CInr |`| fin \<A>) (Inf_automata (ta \<A>) Q)"
lemma Inr_Inl_rel_comp:
"map_both CInr |`| S |O| map_both CInl |`| S = {||}" by auto
lemmas eps_split = ftrancl_Un2_separatorE[OF Inr_Inl_rel_comp]
lemma Inf_automata_eps_simp [simp]:
shows "(map_both Inl |`| eps \<A> |\<union>| map_both CInr |`| eps \<A>)|\<^sup>+| =
(map_both CInl |`| eps \<A>)|\<^sup>+| |\<union>| (map_both CInr |`| eps \<A>)|\<^sup>+|"
proof -
{fix x y z assume "(x, y) |\<in>| (map_both CInl |`| eps \<A>)|\<^sup>+|"
"(y, z) |\<in>| (map_both CInr |`| eps \<A>)|\<^sup>+|"
then have False
by (metis Inl_Inr_False eps_statesI(1, 2) eps_states_image fimageE ftranclD ftranclD2)}
then show ?thesis by (auto simp: Inf_automata_def eps_split)
qed
lemma map_both_CInl_ftrancl_conv:
"(map_both CInl |`| eps \<A>)|\<^sup>+| = map_both CInl |`| (eps \<A>)|\<^sup>+|"
by (intro ftrancl_map_both_fsubset) (auto simp: finj_CInl_CInr)
lemma map_both_CInr_ftrancl_conv:
"(map_both CInr |`| eps \<A>)|\<^sup>+| = map_both CInr |`| (eps \<A>)|\<^sup>+|"
by (intro ftrancl_map_both_fsubset) (auto simp: finj_CInl_CInr)
lemmas map_both_ftrancl_conv = map_both_CInl_ftrancl_conv map_both_CInr_ftrancl_conv
lemma Inf_automata_Inl_to_eps [simp]:
"(CInl p, CInl q) |\<in>| (map_both CInl |`| eps \<A>)|\<^sup>+| \<longleftrightarrow> (p, q) |\<in>| (eps \<A>)|\<^sup>+|"
"(CInr p, CInr q) |\<in>| (map_both CInr |`| eps \<A>)|\<^sup>+| \<longleftrightarrow> (p, q) |\<in>| (eps \<A>)|\<^sup>+|"
"(CInl q, CInl p) |\<in>| (map_both CInr |`| eps \<A>)|\<^sup>+| \<longleftrightarrow> False"
"(CInr q, CInr p) |\<in>| (map_both CInl |`| eps \<A>)|\<^sup>+| \<longleftrightarrow> False"
by (auto simp: map_both_ftrancl_conv dest: fmap_prod_fimageI)
lemma Inl_eps_Inr:
"(CInl q, CInl p) |\<in>| (eps (Inf_automata \<A> Q))|\<^sup>+| \<longleftrightarrow> (CInr q, CInr p) |\<in>| (eps (Inf_automata \<A> Q))|\<^sup>+|"
by (auto simp: Inf_automata_def)
lemma Inr_rhs_eps_Inr_lhs:
assumes "(q, CInr p) |\<in>| (eps (Inf_automata \<A> Q))|\<^sup>+|"
obtains q' where "q = CInr q'" using assms ftrancl_map_both_fsubset[OF finj_CInl_CInr(1)]
by (cases q) (auto simp: Inf_automata_def map_both_ftrancl_conv)
lemma Inl_rhs_eps_Inl_lhs:
assumes "(q, CInl p) |\<in>| (eps (Inf_automata \<A> Q))|\<^sup>+|"
obtains q' where "q = CInl q'" using assms
by (cases q) (auto simp: Inf_automata_def map_both_ftrancl_conv)
lemma Inf_automata_eps [simp]:
"(CInl q, CInr p) |\<in>| (eps (Inf_automata \<A> Q))|\<^sup>+| \<longleftrightarrow> False"
"(CInr q, CInl p) |\<in>| (eps (Inf_automata \<A> Q))|\<^sup>+| \<longleftrightarrow> False"
by (auto elim: Inr_rhs_eps_Inr_lhs Inl_rhs_eps_Inl_lhs)
lemma Inl_A_res_Inf_automata:
"ta_der (fmap_states_ta CInl \<A>) t |\<subseteq>| ta_der (Inf_automata \<A> Q) t"
- by (intro ta_der_mono) (auto simp: Inf_automata_def rev_fimage_eqI fmap_states_ta_def')
+proof (rule ta_der_mono)
+ show "rules (fmap_states_ta CInl \<A>) |\<subseteq>| rules (Inf_automata \<A> Q)"
+ apply (rule fsubsetI)
+ apply (simp add: Inf_automata_def fmap_states_ta_def')
+ by force
+next
+ show "eps (fmap_states_ta CInl \<A>) |\<subseteq>| eps (Inf_automata \<A> Q)"
+ by (rule fsubsetI) (simp add: Inf_automata_def fmap_states_ta_def')
+qed
lemma Inl_res_A_res_Inf_automata:
"CInl |`| ta_der \<A> (term_of_gterm t) |\<subseteq>| ta_der (Inf_automata \<A> Q) (term_of_gterm t)"
by (intro fsubset_trans[OF ta_der_fmap_states_ta_mono[of CInl \<A> t]]) (auto simp: Inl_A_res_Inf_automata)
lemma r_rhs_CInl_args_A_rule:
assumes "f qs \<rightarrow> CInl q |\<in>| rules (Inf_automata \<A> Q)"
obtains qs' where "qs = map CInl qs'" "f qs' \<rightarrow> q |\<in>| rules \<A>" using assms
by (auto simp: Inf_automata_def split!: if_splits)
lemma A_rule_to_dash_closure:
assumes "f qs \<rightarrow> q |\<in>| rules \<A>" and "i < length qs"
shows "f (args qs i) \<rightarrow> CInr q |\<in>| rules (Inf_automata \<A> Q)"
using assms by (auto simp add: Inf_automata_def fimage_iff fBall_def upt_fset intro!: fBexI[OF _ assms(1)])
lemma Inf_automata_reach_to_dash_reach:
assumes "CInl p |\<in>| ta_der (Inf_automata \<A> Q) C\<langle>Var (CInl q)\<rangle>"
shows "CInr p |\<in>| ta_der (Inf_automata \<A> Q) C\<langle>Var (CInr q)\<rangle>" (is "_ |\<in>| ta_der ?A _")
using assms
proof (induct C arbitrary: p)
case (More f ss C ts)
from More(2) obtain qs q' where
rule: "f qs \<rightarrow> q' |\<in>| rules ?A" "length qs = Suc (length ss + length ts)" and
eps: "q' = CInl p \<or> (q', CInl p) |\<in>| (eps ?A)|\<^sup>+|" and
reach: "\<forall> i < Suc (length ss + length ts). qs ! i |\<in>| ta_der ?A ((ss @ C\<langle>Var (CInl q)\<rangle> # ts) ! i)"
by auto
from eps obtain q'' where [simp]: "q' = CInl q''"
by (cases q') (auto simp add: Inf_automata_def eps_split elim: ftranclE converse_ftranclE)
from rule obtain qs' where args: "qs = map CInl qs'" "f qs' \<rightarrow> q'' |\<in>| rules \<A>"
using r_rhs_CInl_args_A_rule by (metis \<open>q' = CInl q''\<close>)
then have "CInl (qs' ! length ss) |\<in>| ta_der (Inf_automata \<A> Q) C\<langle>Var (CInl q)\<rangle>" using reach
by (auto simp: all_Suc_conv nth_append_Cons) (metis length_map less_add_Suc1 local.rule(2) nth_append_length nth_map reach)
from More(1)[OF this] More(2) show ?case
using rule args eps reach A_rule_to_dash_closure[OF args(2), of "length ss" Q]
by (auto simp: Inl_eps_Inr id_take_nth_drop all_Suc_conv
intro!: exI[of _ "CInr q''"] exI[of _ "map CInl (take (length ss) qs') @ CInr (qs' ! length ss) # map CInl (drop (Suc (length ss)) qs')"])
(auto simp: nth_append_Cons min_def)
qed (auto simp: Inf_automata_def)
lemma Inf_automata_dashI:
assumes "run \<A> r (gterm_to_None_Some t)" and "ex_rule_state r |\<in>| Q"
shows "CInr (ex_rule_state r) |\<in>| gta_der (Inf_automata \<A> Q) (gterm_to_None_Some t)"
proof (cases t)
case [simp]: (GFun f ts)
from run_root_rule[OF assms(1)] run_argsD[OF assms(1)] have
rule: "TA_rule (None, Some f) (map ex_comp_state (gargs r)) (ex_rule_state r) |\<in>| rules \<A>" "length (gargs r) = length ts" and
reach: "\<forall> i < length ts. ex_comp_state (gargs r ! i) |\<in>| ta_der \<A> (term_of_gterm (gterm_to_None_Some (ts ! i)))"
by (auto intro!: run_to_comp_st_gta_der[unfolded gta_der_def comp_def])
from rule assms(2) have "(None, Some f) (map (CInl \<circ> ex_comp_state) (gargs r)) \<rightarrow> CInr (ex_rule_state r) |\<in>| rules (Inf_automata \<A> Q)"
by (auto simp: Inf_automata_def) force
then show ?thesis using reach rule Inl_res_A_res_Inf_automata[of \<A> "gterm_to_None_Some (ts ! i)" Q for i]
by (auto simp: gta_der_def intro!: exI[of _ "CInr (ex_rule_state r)"] exI[of _ "map (CInl \<circ> ex_comp_state) (gargs r)"])
blast
qed
lemma Inf_automata_dash_reach_to_reach:
assumes "p |\<in>| ta_der (Inf_automata \<A> Q) t" (is "_ |\<in>| ta_der ?A _")
shows "remove_sum p |\<in>| ta_der \<A> (map_vars_term remove_sum t)" using assms
proof (induct t arbitrary: p)
case (Var x) then show ?case
by (cases p; cases x) (auto simp: Inf_automata_def ftrancl_map_both map_both_ftrancl_conv)
next
case (Fun f ss)
from Fun(2) obtain qs q' where
rule: "f qs \<rightarrow> q' |\<in>| rules ?A" "length qs = length ss" and
eps: "q' = p \<or> (q', p) |\<in>| (eps ?A)|\<^sup>+|" and
reach: "\<forall> i < length ss. qs ! i |\<in>| ta_der ?A (ss ! i)" by auto
from rule have r: "f (map (remove_sum) qs) \<rightarrow> (remove_sum q') |\<in>| rules \<A>"
by (auto simp: comp_def Inf_automata_def min_def id_take_nth_drop[symmetric] upt_fset
simp flip: drop_map take_map split!: if_splits)
moreover have "remove_sum q' = remove_sum p \<or> (remove_sum q', remove_sum p) |\<in>| (eps \<A>)|\<^sup>+|" using eps
by (cases "is_Inl q'"; cases "is_Inl p") (auto elim!: is_InlE is_InrE, auto simp: Inf_automata_def)
ultimately show ?case using reach rule(2) Fun(1)[OF nth_mem, of i "qs ! i" for i]
by auto (metis (mono_tags, lifting) length_map map_nth_eq_conv)+
qed
lemma depth_poss_split:
assumes "Suc (depth (term_of_gterm t) + n) < depth (term_of_gterm u)"
shows "\<exists> p q. p @ q \<in> gposs u \<and> n < length q \<and> p \<notin> gposs t"
proof -
from poss_length_depth obtain p m where p: "p \<in> gposs u" "length p = m" "depth (term_of_gterm u) = m"
using poss_gposs_conv by blast
then obtain m' where dt: "depth (term_of_gterm t) = m'" by blast
from assms dt p(2, 3) have "length (take (Suc m') p) = Suc m'"
by (metis Suc_leI depth_gterm_conv length_take less_add_Suc1 less_imp_le_nat less_le_trans min.absorb2)
then have nt: "take (Suc m') p \<notin> gposs t" using poss_length_bounded_by_depth dt depth_gterm_conv
by (metis Suc_n_not_le_n gposs_to_poss)
moreover have "n < length (drop (Suc m') p)" using assms depth_gterm_conv dt p(2-)
by (metis add_Suc diff_diff_left length_drop zero_less_diff)
ultimately show ?thesis by (metis append_take_drop_id p(1))
qed
lemma Inf_to_automata:
assumes "RR2_spec \<A> \<R>" and "t \<in> Inf_branching_terms \<R> \<F>"
shows "\<exists> u. gpair t u \<in> \<L> (Inf_reg \<A> (Q_infty (ta \<A>) \<F>))" (is "\<exists> u. gpair t u \<in> \<L> ?B")
proof -
let ?A = "Inf_automata (ta \<A>) (Q_infty (ta \<A>) \<F>)"
let ?t_of_g = "\<lambda> t. term_of_gterm t :: ('b, 'a) term"
obtain n where depth_card: "depth (?t_of_g t) + fcard (\<Q> (ta \<A>)) < n" by auto
from assms(1, 2) have fin: "infinite {u. gpair t u \<in> \<L> \<A> \<and> funas_gterm u \<subseteq> fset \<F>}"
by (auto simp: RR2_spec_def Inf_branching_terms_def)
from infinte_no_depth_limit[of "?t_of_g ` {u. gpair t u \<in> \<L> \<A> \<and> funas_gterm u \<subseteq> fset \<F>}" "fset \<F>"] this
have "\<forall> n. \<exists>t \<in> ?t_of_g ` {u. gpair t u \<in> \<L> \<A> \<and> funas_gterm u \<subseteq> fset \<F>}. n < depth t"
by (simp add: infinite_inj_image_infinite[OF fin] funas_term_of_gterm_conv inj_term_of_gterm)
from this depth_card obtain u where funas: "funas_gterm u \<subseteq> fset \<F>" and
depth: "Suc n < depth (?t_of_g u)" and lang: "gpair t u \<in> \<L> \<A>" by auto
have "Suc (depth (term_of_gterm t) + fcard (\<Q> (ta \<A>))) < depth (term_of_gterm u)"
using depth depth_card by (metis Suc_less_eq2 depth_gterm_conv less_trans)
from depth_poss_split[OF this] obtain p q where
pos: "p @ q \<in> gposs u" "p \<notin> gposs t" and card: "fcard (\<Q> (ta \<A>)) < length q" by auto
then have gp: "gsubt_at (gpair t u) p = gterm_to_None_Some (gsubt_at u p)"
using subst_at_gpair_nt_poss_None_Some[of p] by force
from lang obtain r where r: "run (ta \<A>) r (gpair t u)" "ex_comp_state r |\<in>| fin \<A>"
unfolding \<L>_def gta_lang_def by (fastforce dest: gta_der_to_run)
from pos have p_gtu: "p \<in> gposs (gpair t u)" and pu: "p \<in> gposs u"
using not_gposs_append by auto
have qinf: "ex_rule_state (gsubt_at r p) |\<in>| Q_infty (ta \<A>) \<F>"
using funas_gterm_gsubt_at_subseteq[OF pu] funas card
unfolding Q_infty_fmember gta_der_def[symmetric]
by (intro infinite_super[THEN infinite_imageD2[OF _ inj_gterm_to_None_Some],
OF _ pigeonhole_ta_infinit_terms[of "ta \<A>" "gsubt_at (gpair t u) p" _
"\<lambda> t. t \<subseteq> (\<lambda>(f, n). ((None, Some f), n)) ` fset \<F>",
OF _ run_to_gta_der_gsubt_at(1)[OF r(1) p_gtu]]])
(auto simp: poss_length_bounded_by_depth[of q] image_iff gp less_le_trans
pos(1) poss_gposs_conv pu dest!: funas_gterm_bot_some_decomp)
from Inf_automata_dashI[OF run_gsubt_cl[OF r(1) p_gtu, unfolded gp] qinf]
have dashI: "CInr (ex_rule_state (gsubt_at r p)) |\<in>| gta_der (Inf_automata (ta \<A>) (Q_infty (ta \<A>) \<F>)) (gsubt_at (gpair t u) p)"
unfolding gp[symmetric] .
have "CInl (ex_comp_state r) |\<in>| ta_der ?A (ctxt_at_pos (term_of_gterm (gpair t u)) p)\<langle>Var (CInl (ex_rule_state (gsubt_at r p)))\<rangle>"
using ta_der_fmap_states_ta[OF run_ta_der_ctxt_split2[OF r(1) p_gtu], of CInl, THEN fsubsetD[OF Inl_A_res_Inf_automata]]
unfolding replace_term_at_replace_at_conv[OF gposs_to_poss[OF p_gtu]]
by (auto simp: gterm.map_ident simp flip: map_term_replace_at_dist[OF gposs_to_poss[OF p_gtu]])
from ta_der_ctxt[OF dashI[unfolded gta_der_def] Inf_automata_reach_to_dash_reach[OF this]]
have "CInr (ex_comp_state r) |\<in>| gta_der (Inf_automata (ta \<A>) (Q_infty (ta \<A>) \<F>)) (gpair t u)"
unfolding replace_term_at_replace_at_conv[OF gposs_to_poss[OF p_gtu]]
unfolding replace_gterm_conv[OF p_gtu]
by (auto simp: gta_der_def)
moreover from r(2) have "CInr (ex_comp_state r) |\<in>| fin (Inf_reg \<A> (Q_infty (ta \<A>) \<F>))"
by (auto simp: Inf_reg_def)
ultimately show ?thesis using r(2)
by (auto simp: \<L>_def gta_der_def Inf_reg_def intro: exI[of _ u])
qed
lemma CInr_Inf_automata_to_q_state:
assumes "CInr p |\<in>| ta_der (Inf_automata \<A> Q) t" and "ground t"
shows "\<exists> C s q. C\<langle>s\<rangle> = t \<and> CInr q |\<in>| ta_der (Inf_automata \<A> Q) s \<and> q |\<in>| Q \<and>
CInr p |\<in>| ta_der (Inf_automata \<A> Q) C\<langle>Var (CInr q)\<rangle> \<and>
(fst \<circ> fst \<circ> the \<circ> root) s = None" using assms
proof (induct t arbitrary: p)
case (Fun f ts)
let ?A = "(Inf_automata \<A> Q)"
from Fun(2) obtain qs q' where
rule: "f qs \<rightarrow> CInr q' |\<in>| rules ?A" "length qs = length ts" and
eps: "q' = p \<or> (CInr q', CInr p) |\<in>| (eps ?A)|\<^sup>+|" and
reach: "\<forall> i < length ts. qs ! i |\<in>| ta_der ?A (ts ! i)"
by auto (metis Inr_rhs_eps_Inr_lhs)
consider (a) "\<And> i. i < length qs \<Longrightarrow> \<exists> q''. qs ! i = CInl q''" | (b) "\<exists> i < length qs. \<exists> q''. qs ! i = CInr q''"
by (meson remove_sum.cases)
then show ?case
proof cases
case a
then have "f qs \<rightarrow> CInr q' |\<in>| |\<Union>| (q_inf_dash_intro_rules Q |`| rules \<A>)" using rule
by (auto simp: Inf_automata_def min_def upt_fset split!: if_splits)
(metis (no_types, lifting) Inl_Inr_False Suc_pred append_eq_append_conv id_take_nth_drop
length_Cons length_drop length_greater_0_conv length_map
less_nat_zero_code list.size(3) nth_append_length rule(2))
then show ?thesis using reach eps rule
by (intro exI[of _ Hole] exI[of _ "Fun f ts"] exI[of _ q'])
(auto split!: if_splits)
next
case b
then obtain i q'' where b: "i < length ts" "qs ! i = CInr q''" using rule(2) by auto
then have "CInr q'' |\<in>| ta_der ?A (ts ! i)" using rule(2) reach by auto
from Fun(3) Fun(1)[OF nth_mem, OF b(1) this] b rule(2) obtain C s q''' where
ctxt: "C\<langle>s\<rangle> = ts ! i" and
qinf: "CInr q''' |\<in>| ta_der (Inf_automata \<A> Q) s \<and> q''' |\<in>| Q" and
reach2: "CInr q'' |\<in>| ta_der (Inf_automata \<A> Q) C\<langle>Var (CInr q''')\<rangle>" and
"(fst \<circ> fst \<circ> the \<circ> root) s = None"
by auto
then show ?thesis using rule eps reach ctxt qinf reach2 b(1) b(2)[symmetric]
by (auto simp: min_def nth_append_Cons simp flip: map_append id_take_nth_drop[OF b(1)]
intro!: exI[of _ "More f (take i ts) C (drop (Suc i) ts)"] exI[of _ "s"] exI[of _ "q'''"] exI[of _ "CInr q'"] exI[of _ "qs"])
qed
qed auto
lemma aux_lemma:
assumes "RR2_spec \<A> \<R>" and "\<R> \<subseteq> \<T>\<^sub>G (fset \<F>) \<times> \<T>\<^sub>G (fset \<F>)"
and "infinite {u | u. gpair t u \<in> \<L> \<A>}"
shows "t \<in> Inf_branching_terms \<R> \<F>"
proof -
from assms have [simp]: "gpair t u \<in> \<L> \<A> \<longleftrightarrow> (t, u) \<in> \<R> \<and> u \<in> \<T>\<^sub>G (fset \<F>)"
for u by (auto simp: RR2_spec_def)
from assms have "t \<in> \<T>\<^sub>G (fset \<F>)" unfolding RR2_spec_def
by (auto dest: not_finite_existsD)
then show ?thesis using assms unfolding Inf_branching_terms_def
by (auto simp: \<T>\<^sub>G_equivalent_def)
qed
lemma Inf_automata_to_Inf:
assumes "RR2_spec \<A> \<R>" and "\<R> \<subseteq> \<T>\<^sub>G (fset \<F>) \<times> \<T>\<^sub>G (fset \<F>)"
and "gpair t u \<in> \<L> (Inf_reg \<A> (Q_infty (ta \<A>) \<F>))"
shows "t \<in> Inf_branching_terms \<R> \<F>"
proof -
let ?con = "\<lambda> t. term_of_gterm (gterm_to_None_Some t)"
let ?A = "Inf_automata (ta \<A>) (Q_infty (ta \<A>) \<F>)"
from assms(3) obtain q where fin: "q |\<in>| fin \<A>" and
reach_fin: "CInr q |\<in>| ta_der ?A (term_of_gterm (gpair t u))"
by (auto simp: Inf_reg_def \<L>_def Inf_automata_def elim!: gta_langE)
from CInr_Inf_automata_to_q_state[OF reach_fin] obtain C s p where
ctxt: "C\<langle>s\<rangle> = term_of_gterm (gpair t u)" and
q_inf_st: "CInr p |\<in>| ta_der ?A s" "p |\<in>| Q_infty (ta \<A>) \<F>" and
reach: "CInr q |\<in>| ta_der ?A C\<langle>Var (CInr p)\<rangle>" and
none: "(fst \<circ> fst \<circ> the \<circ> root) s = None" by auto
have gr: "ground s" "ground_ctxt C" using arg_cong[OF ctxt, of ground]
by auto
have reach: "q |\<in>| ta_der (ta \<A>) (adapt_vars_ctxt C)\<langle>Var p\<rangle>"
using gr Inf_automata_dash_reach_to_reach[OF reach]
by (auto simp: map_vars_term_ctxt_apply)
from q_inf_st(2) have inf: "infinite {v. funas_gterm v \<subseteq> fset \<F> \<and> p |\<in>| ta_der (ta \<A>) (?con v)}"
by (simp add: Q_infty_fmember)
have inf: "infinite {v. funas_gterm v \<subseteq> fset \<F> \<and> q |\<in>| gta_der (ta \<A>) (gctxt_of_ctxt C)\<langle>gterm_to_None_Some v\<rangle>\<^sub>G}"
using reach ground_ctxt_adapt_ground[OF gr(2)] gr
by (intro infinite_super[OF _ inf], auto simp: gta_der_def)
(smt (z3) adapt_vars_ctxt adapt_vars_term_of_gterm ground_gctxt_of_ctxt_apply_gterm ta_der_ctxt)
have *: "gfun_at (gterm_of_term C\<langle>s\<rangle>) (hole_pos C) = gfun_at (gterm_of_term s) []"
by (induct C) (auto simp: nth_append_Cons)
from arg_cong[OF ctxt, of "\<lambda> t. gfun_at (gterm_of_term t) (hole_pos C)"] none
have hp_nt: "ghole_pos (gctxt_of_ctxt C) \<notin> gposs t" unfolding ground_hole_pos_to_ghole[OF gr(2)]
using gfun_at_gpair[of t u "hole_pos C"] gr *
by (cases s) (auto simp flip: poss_gposs_mem_conv split: if_splits elim: gfun_at_possE)
from gpair_ctxt_decomposition[OF hp_nt, of u "gsubt_at u (hole_pos C)"]
have to_gpair: "gpair t (gctxt_at_pos u (hole_pos C))\<langle>v\<rangle>\<^sub>G = (gctxt_of_ctxt C)\<langle>gterm_to_None_Some v\<rangle>\<^sub>G" for v
unfolding ground_hole_pos_to_ghole[OF gr(2)] using ctxt gr
using subst_at_gpair_nt_poss_None_Some[OF _ hp_nt, of u]
by (metis (no_types, lifting) UnE \<open>ghole_pos (gctxt_of_ctxt C) = hole_pos C\<close>
gposs_of_gpair gsubt_at_gctxt_apply_ghole hole_pos_in_ctxt_apply hp_nt poss_gposs_conv term_of_gterm_ctxt_apply)
have inf: "infinite {v. gpair t ((gctxt_at_pos u (hole_pos C))\<langle>v\<rangle>\<^sub>G) \<in> \<L> \<A>}" using fin
by (intro infinite_super[OF _ inf]) (auto simp: \<L>_def gta_der_def simp flip: to_gpair)
have "infinite {u |u. gpair t u \<in> \<L> \<A>}"
by (intro infinite_super[OF _ infinite_inj_image_infinite[OF inf gctxt_apply_inj_on_term[of "gctxt_at_pos u (hole_pos C)"]]])
(auto simp: image_def intro: infinite_super)
then show ?thesis using assms(1, 2)
by (intro aux_lemma[of \<A>]) simp
qed
lemma Inf_automata_subseteq:
"\<L> (Inf_reg \<A> (Q_infty (ta \<A>) \<F>)) \<subseteq> \<L> \<A>" (is "\<L> ?IA \<subseteq> _")
proof
fix s assume l: "s \<in> \<L> ?IA"
then obtain q where w: "q |\<in>| fin ?IA" "q |\<in>| ta_der (ta ?IA) (term_of_gterm s)"
by (auto simp: \<L>_def elim!: gta_langE)
from w(1) have "remove_sum q |\<in>| fin \<A>"
by (auto simp: Inf_reg_def Inf_automata_def)
then show "s \<in> \<L> \<A>" using Inf_automata_dash_reach_to_reach[of q "ta \<A>"] w(2)
by (auto simp: gterm.map_ident \<L>_def Inf_reg_def)
(metis gta_langI map_vars_term_term_of_gterm)
qed
lemma \<L>_Inf_reg:
assumes "RR2_spec \<A> \<R>" and "\<R> \<subseteq> \<T>\<^sub>G (fset \<F>) \<times> \<T>\<^sub>G (fset \<F>)"
shows "gfst ` \<L> (Inf_reg \<A> (Q_infty (ta \<A>) \<F>)) = Inf_branching_terms \<R> \<F>"
proof -
{fix s assume ass: "s \<in> \<L> (Inf_reg \<A> (Q_infty (ta \<A>) \<F>))"
then have "\<exists> t u. s = gpair t u" using Inf_automata_subseteq[of \<A> \<F>] assms(1)
by (auto simp: RR2_spec_def)
then have "gfst s \<in> Inf_branching_terms \<R> \<F>"
using ass Inf_automata_to_Inf[OF assms]
by (force simp: gfst_gpair)}
then show ?thesis using Inf_to_automata[OF assms(1), of _ \<F>]
by (auto simp: gfst_gpair) (metis gfst_gpair image_iff)
qed
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/RR2_Infinite_Q_infinity.thy b/thys/Regular_Tree_Relations/RR2_Infinite_Q_infinity.thy
--- a/thys/Regular_Tree_Relations/RR2_Infinite_Q_infinity.thy
+++ b/thys/Regular_Tree_Relations/RR2_Infinite_Q_infinity.thy
@@ -1,464 +1,464 @@
theory RR2_Infinite_Q_infinity
imports RR2_Infinite
begin
(* This section constructs an executable membership check for Q infinity,
since Inf_automata is already executable (for all sets Q where checking membership is executable)
*)
lemma if_cong':
"b = c \<Longrightarrow> x = u \<Longrightarrow> y = v \<Longrightarrow> (if b then x else y) = (if c then u else v)"
by auto
(* The reachable terms where eps transitions can only occur after the rule *)
fun ta_der_strict :: "('q,'f) ta \<Rightarrow> ('f,'q) term \<Rightarrow> 'q fset" where
"ta_der_strict \<A> (Var q) = {|q|}"
| "ta_der_strict \<A> (Fun f ts) = {| q' | q' q qs. TA_rule f qs q |\<in>| rules \<A> \<and> (q = q' \<or> (q, q') |\<in>| (eps \<A>)|\<^sup>+|) \<and>
length qs = length ts \<and> (\<forall> i < length ts. qs ! i |\<in>| ta_der_strict \<A> (ts ! i))|}"
lemma ta_der_strict_Var:
"q |\<in>| ta_der_strict \<A> (Var x) \<longleftrightarrow> x = q"
unfolding ta_der.simps by auto
lemma ta_der_strict_Fun:
"q |\<in>| ta_der_strict \<A> (Fun f ts) \<longleftrightarrow> (\<exists> ps p. TA_rule f ps p |\<in>| (rules \<A>) \<and>
(p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|) \<and> length ps = length ts \<and>
(\<forall> i < length ts. ps ! i |\<in>| ta_der_strict \<A> (ts ! i)))" (is "?Ls \<longleftrightarrow> ?Rs")
unfolding ta_der_strict.simps
by (intro iffI fCollect_memberI finite_Collect_less_eq[OF _ finite_eps[of \<A>]]) auto
declare ta_der_strict.simps[simp del]
lemmas ta_der_strict_simps [simp] = ta_der_strict_Var ta_der_strict_Fun
lemma ta_der_strict_sub_ta_der:
"ta_der_strict \<A> t |\<subseteq>| ta_der \<A> t"
proof (induct t)
case (Fun f ts)
then show ?case
by auto (metis fsubsetD nth_mem)+
qed auto
lemma ta_der_strict_ta_der_eq_on_ground:
assumes"ground t"
shows "ta_der \<A> t = ta_der_strict \<A> t"
proof
{fix q assume "q |\<in>| ta_der \<A> t" then have "q |\<in>| ta_der_strict \<A> t" using assms
proof (induct t arbitrary: q)
case (Fun f ts)
then show ?case apply auto
using nth_mem by blast+
qed auto}
then show "ta_der \<A> t |\<subseteq>| ta_der_strict \<A> t"
by auto
next
show "ta_der_strict \<A> t |\<subseteq>| ta_der \<A> t" using ta_der_strict_sub_ta_der .
qed
lemma ta_der_to_ta_strict:
assumes "q |\<in>| ta_der A C\<langle>Var p\<rangle>" and "ground_ctxt C"
shows "\<exists> q'. (p = q' \<or> (p, q') |\<in>| (eps A)|\<^sup>+|) \<and> q |\<in>| ta_der_strict A C\<langle>Var q'\<rangle>"
using assms
proof (induct C arbitrary: q p)
case (More f ss C ts)
from More(2) obtain qs q' where
r: "TA_rule f qs q' |\<in>| rules A" "length qs = Suc (length ss + length ts)" "q' = q \<or> (q', q) |\<in>| (eps A)|\<^sup>+|" and
rec: "\<forall> i < length qs. qs ! i |\<in>| ta_der A ((ss @ C\<langle>Var p\<rangle> # ts) ! i)"
by auto
from More(1)[of "qs ! length ss" p] More(3) rec r(2) obtain q'' where
mid: "(p = q'' \<or> (p, q'') |\<in>| (eps A)|\<^sup>+|) \<and> qs ! length ss |\<in>| ta_der_strict A C\<langle>Var q''\<rangle>"
by auto (metis length_map less_add_Suc1 nth_append_length)
then have "\<forall> i < length qs. qs ! i |\<in>| ta_der_strict A ((ss @ C\<langle>Var q''\<rangle> # ts) ! i)"
using rec r(2) More(3)
using ta_der_strict_ta_der_eq_on_ground[of _ A]
- by (auto simp: nth_append_Cons all_Suc_conv fmember_iff_member_fset split:if_splits cong: if_cong')
+ by (auto simp: nth_append_Cons all_Suc_conv split:if_splits cong: if_cong')
then show ?case using rec r conjunct1[OF mid]
by (rule_tac x = q'' in exI, auto intro!: exI[of _ q'] exI[of _ qs])
qed auto
fun root_ctxt where
"root_ctxt (More f ss C ts) = f"
| "root_ctxt \<box> = undefined"
lemma root_to_root_ctxt [simp]:
assumes "C \<noteq> \<box>"
shows "fst (the (root C\<langle>t\<rangle>)) \<longleftrightarrow> root_ctxt C"
using assms by (cases C) auto
(* Q_inf section *)
inductive_set Q_inf for \<A> where
trans: "(p, q) \<in> Q_inf \<A> \<Longrightarrow> (q, r) \<in> Q_inf \<A> \<Longrightarrow> (p, r) \<in> Q_inf \<A>"
| rule: "(None, Some f) qs \<rightarrow> q |\<in>| rules \<A> \<Longrightarrow> i < length qs \<Longrightarrow> (qs ! i, q) \<in> Q_inf \<A>"
| eps: "(p, q) \<in> Q_inf \<A> \<Longrightarrow> (q, r) |\<in>| eps \<A> \<Longrightarrow> (p, r) \<in> Q_inf \<A>"
abbreviation "Q_inf_e \<A> \<equiv> {q | p q. (p, p) \<in> Q_inf \<A> \<and> (p, q) \<in> Q_inf \<A>}"
lemma Q_inf_states_ta_states:
assumes "(p, q) \<in> Q_inf \<A>"
shows "p |\<in>| \<Q> \<A>" "q |\<in>| \<Q> \<A>"
using assms by (induct) (auto simp: rule_statesD eps_statesD)
lemma Q_inf_finite:
"finite (Q_inf \<A>)" "finite (Q_inf_e \<A>)"
proof -
have *: "Q_inf \<A> \<subseteq> fset (\<Q> \<A> |\<times>| \<Q> \<A>)" "Q_inf_e \<A> \<subseteq> fset (\<Q> \<A>)"
- by (auto simp add: Q_inf_states_ta_states(1, 2) subrelI simp flip: fmember_iff_member_fset)
+ by (auto simp add: Q_inf_states_ta_states(1, 2) subrelI)
show "finite (Q_inf \<A>)"
by (intro finite_subset[OF *(1)]) simp
show "finite (Q_inf_e \<A>)"
by (intro finite_subset[OF *(2)]) simp
qed
context
includes fset.lifting
begin
lift_definition fQ_inf :: "('a, 'b option \<times> 'c option) ta \<Rightarrow> ('a \<times> 'a) fset" is Q_inf
by (simp add: Q_inf_finite(1))
lift_definition fQ_inf_e :: "('a, 'b option \<times> 'c option) ta \<Rightarrow> 'a fset" is Q_inf_e
using Q_inf_finite(2) .
end
lemma Q_inf_ta_eps_Q_inf:
assumes "(p, q) \<in> Q_inf \<A>" and "(q, q') |\<in>| (eps \<A>)|\<^sup>+|"
shows "(p, q') \<in> Q_inf \<A>" using assms(2, 1)
by (induct rule: ftrancl_induct) (auto simp add: Q_inf.eps)
lemma lhs_state_rule:
assumes "(p, q) \<in> Q_inf \<A>"
shows "\<exists> f qs r. (None, Some f) qs \<rightarrow> r |\<in>| rules \<A> \<and> p |\<in>| fset_of_list qs"
using assms by induct (force intro: nth_mem)+
lemma Q_inf_reach_state_rule:
assumes "(p, q) \<in> Q_inf \<A>" and "\<Q> \<A> |\<subseteq>| ta_reachable \<A>"
shows "\<exists> ss ts f C. q |\<in>| ta_der \<A> (More (None, Some f) ss C ts)\<langle>Var p\<rangle> \<and> ground_ctxt (More (None, Some f) ss C ts)"
(is "\<exists> ss ts f C. ?P ss ts f C q p")
using assms
proof (induct)
case (trans p q r)
then obtain f1 f2 ss1 ts1 ss2 ts2 C1 C2 where
C: "?P ss1 ts1 f1 C1 q p" "?P ss2 ts2 f2 C2 r q" by blast
then show ?case
apply (rule_tac x = "ss2" in exI, rule_tac x = "ts2" in exI, rule_tac x = "f2" in exI,
rule_tac x = "C2 \<circ>\<^sub>c (More (None, Some f1) ss1 C1 ts1)" in exI)
apply (auto simp del: ctxt_apply_term.simps)
apply (metis Subterm_and_Context.ctxt_ctxt_compose ctxt_compose.simps(2) ta_der_ctxt)
done
next
case (rule f qs q i)
have "\<forall> i < length qs. \<exists> t. qs ! i |\<in>| ta_der \<A> t \<and> ground t"
using rule(1, 2) fset_mp[OF rule(3), of "qs ! i" for i]
by auto (meson fnth_mem rule_statesD(4) ta_reachableE)
then obtain ts where wit: "length ts = length qs"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> (ts ! i) \<and> ground (ts ! i)"
using Ex_list_of_length_P[of "length qs" "\<lambda> x i. qs ! i |\<in>| ta_der \<A> x \<and> ground x"] by blast
{fix j assume "j < length qs"
then have "qs ! j |\<in>| ta_der \<A> ((take i ts @ Var (qs ! i) # drop (Suc i) ts) ! j)"
using wit by (cases "j < i") (auto simp: min_def nth_append_Cons)}
then have "\<forall> i < length qs. qs ! i |\<in>| (map (ta_der \<A>) (take i ts @ Var (qs ! i) # drop (Suc i) ts)) ! i"
using wit rule(2) by (auto simp: nth_append_Cons)
then have res: "q |\<in>| ta_der \<A> (Fun (None, Some f) (take i ts @ Var (qs ! i) # drop (Suc i) ts))"
using rule(1, 2) wit by (auto simp: min_def nth_append_Cons intro!: exI[of _ q] exI[of _ qs])
then show ?case using rule(1, 2) wit
apply (rule_tac x = "take i ts" in exI, rule_tac x = "drop (Suc i) ts" in exI)
apply (auto simp: take_map drop_map dest!: in_set_takeD in_set_dropD simp del: ta_der_simps intro!: exI[of _ f] exI[of _ Hole])
apply (metis all_nth_imp_all_set)+
done
next
case (eps p q r)
then show ?case by (meson r_into_rtrancl ta_der_eps)
qed
lemma rule_target_Q_inf:
assumes "(None, Some f) qs \<rightarrow> q' |\<in>| rules \<A>" and "i < length qs"
shows "(qs ! i, q') \<in> Q_inf \<A>" using assms
by (intro rule) auto
lemma rule_target_eps_Q_inf:
assumes "(None, Some f) qs \<rightarrow> q' |\<in>| rules \<A>" "(q', q) |\<in>| (eps \<A>)|\<^sup>+|"
and "i < length qs"
shows "(qs ! i, q) \<in> Q_inf \<A>"
using assms(2, 1, 3) by (induct rule: ftrancl_induct) (auto intro: rule eps)
lemma step_in_Q_inf:
assumes "q |\<in>| ta_der_strict \<A> (map_funs_term (\<lambda>f. (None, Some f)) (Fun f (ss @ Var p # ts)))"
shows "(p, q) \<in> Q_inf \<A>"
using assms rule_target_eps_Q_inf[of f _ _ \<A> q] rule_target_Q_inf[of f _ q \<A>]
by (auto simp: comp_def nth_append_Cons split!: if_splits)
lemma ta_der_Q_inf:
assumes "q |\<in>| ta_der_strict \<A> (map_funs_term (\<lambda>f. (None, Some f)) (C\<langle>Var p\<rangle>))" and "C \<noteq> Hole"
shows "(p, q) \<in> Q_inf \<A>" using assms
proof (induct C arbitrary: q)
case (More f ss C ts)
then show ?case
proof (cases "C = Hole")
case True
then show ?thesis using More(2) by (auto simp: step_in_Q_inf)
next
case False
then obtain q' where q: "q' |\<in>| ta_der_strict \<A> (map_funs_term (\<lambda>f. (None, Some f)) C\<langle>Var p\<rangle>)"
"q |\<in>| ta_der_strict \<A> (map_funs_term (\<lambda>f. (None, Some f)) (Fun f (ss @ Var q' # ts)))"
using More(2) length_map
(* SLOW *)
by (auto simp: comp_def nth_append_Cons split: if_splits cong: if_cong')
(smt nat_neq_iff nth_map ta_der_strict_simps)+
have "(p, q') \<in> Q_inf \<A>" using More(1)[OF q(1) False] .
then show ?thesis using step_in_Q_inf[OF q(2)] by (auto intro: trans)
qed
qed auto
lemma Q_inf_e_infinite_terms_res:
assumes "q \<in> Q_inf_e \<A>" and "\<Q> \<A> |\<subseteq>| ta_reachable \<A>"
shows "infinite {t. q |\<in>| ta_der \<A> (term_of_gterm t) \<and> fst (groot_sym t) = None}"
proof -
let ?P ="\<lambda> u. ground u \<and> q |\<in>| ta_der \<A> u \<and> fst (fst (the (root u))) = None"
have groot[simp]: "fst (fst (the (root (term_of_gterm t)))) = fst (groot_sym t)" for t by (cases t) auto
have [simp]: "C \<noteq> \<box> \<Longrightarrow> fst (fst (the (root C\<langle>t\<rangle>))) = fst (root_ctxt C)" for C t by (cases C) auto
from assms(1) obtain p where cycle: "(p, p) \<in> Q_inf \<A>" "(p, q) \<in> Q_inf \<A>" by auto
from Q_inf_reach_state_rule[OF cycle(1) assms(2)] obtain C where
ctxt: "C \<noteq> \<box>" "ground_ctxt C" and reach: "p |\<in>| ta_der \<A> C\<langle>Var p\<rangle>"
by blast
obtain C2 where
closing_ctxt: "C2 \<noteq> \<box>" "ground_ctxt C2" "fst (root_ctxt C2) = None" and cl_reach: "q |\<in>| ta_der \<A> C2\<langle>Var p\<rangle>"
by (metis (full_types) Q_inf_reach_state_rule[OF cycle(2) assms(2)] ctxt.distinct(1) fst_conv root_ctxt.simps(1))
from assms(2) obtain t where t: "p |\<in>| ta_der \<A> t" and gr_t: "ground t"
by (meson Q_inf_states_ta_states(1) cycle(1) fsubsetD ta_reachableE)
let ?terms = "\<lambda> n. (C ^ Suc n)\<langle>t\<rangle>" let ?S = "{?terms n | n. p |\<in>| ta_der \<A> (?terms n) \<and> ground (?terms n)}"
have "ground (?terms n)" for n using ctxt(2) gr_t by auto
moreover have "p |\<in>| ta_der \<A> (?terms n)" for n using reach t(1)
by (auto simp: ta_der_ctxt) (meson ta_der_ctxt ta_der_ctxt_n_loop)
ultimately have inf: "infinite ?S" using ctxt_comp_n_lower_bound[OF ctxt(1)]
using no_upper_bound_infinite[of _ depth, of ?S] by blast
from infinite_inj_image_infinite[OF this] have inf:"infinite (ctxt_apply_term C2 ` ?S)"
by (smt ctxt_eq inj_on_def)
{fix u assume "u \<in> (ctxt_apply_term C2 ` ?S)"
then have "?P u" unfolding image_Collect using closing_ctxt cl_reach
by (auto simp: ta_der_ctxt)}
from this have inf: "infinite {u. ground u \<and> q |\<in>| ta_der \<A> u \<and> fst (fst (the (root u))) = None}"
by (intro infinite_super[OF _ inf] subsetI) fast
have inf: "infinite (gterm_of_term ` {u. ground u \<and> q |\<in>| ta_der \<A> u \<and> fst (fst (the (root u))) = None})"
by (intro infinite_inj_image_infinite[OF inf] gterm_of_term_inj) auto
show ?thesis
by (intro infinite_super[OF _ inf]) (auto dest: groot_sym_gterm_of_term)
qed
lemma gfun_at_after_hole_pos:
assumes "ghole_pos C \<le>\<^sub>p p"
shows "gfun_at C\<langle>t\<rangle>\<^sub>G p = gfun_at t (p -\<^sub>p ghole_pos C)" using assms
proof (induct C arbitrary: p)
case (GMore f ss C ts) then show ?case
by (cases p) auto
qed auto
lemma pos_diff_0 [simp]: "p -\<^sub>p p = []"
by (auto simp: pos_diff_def)
lemma Max_suffI: "finite A \<Longrightarrow> A = B \<Longrightarrow> Max A = Max B"
by (intro Max_eq_if) auto
lemma nth_args_depth_eqI:
assumes "length ss = length ts"
and "\<And> i. i < length ts \<Longrightarrow> depth (ss ! i) = depth (ts ! i)"
shows "depth (Fun f ss) = depth (Fun g ts)"
proof -
from assms(1, 2) have "depth ` set ss = depth ` set ts"
using nth_map_conv[OF assms(1), of depth depth]
by (simp flip: set_map)
from Max_suffI[OF _ this] show ?thesis using assms(1)
by (cases ss; cases ts) auto
qed
lemma subt_at_ctxt_apply_hole_pos [simp]: "C\<langle>s\<rangle> |_ hole_pos C = s"
by (induct C) auto
lemma ctxt_at_pos_ctxt_apply_hole_poss [simp]: "ctxt_at_pos C\<langle>s\<rangle> (hole_pos C) = C"
by (induct C) auto
abbreviation "map_funs_ctxt f \<equiv> map_ctxt f (\<lambda> x. x)"
lemma map_funs_term_ctxt_apply [simp]:
"map_funs_term f C\<langle>s\<rangle> = (map_funs_ctxt f C)\<langle>map_funs_term f s\<rangle>"
by (induct C) auto
lemma map_funs_term_ctxt_decomp:
assumes "map_funs_term fg t = C\<langle>s\<rangle>"
shows "\<exists> D u. C = map_funs_ctxt fg D \<and> s = map_funs_term fg u \<and> t = D\<langle>u\<rangle>"
using assms
proof (induct C arbitrary: t)
case Hole
show ?case
by (rule exI[of _ Hole], rule exI[of _ t], insert Hole, auto)
next
case (More g bef C aft)
from More(2) obtain f ts where t: "t = Fun f ts" by (cases t, auto)
from More(2)[unfolded t] have f: "fg f = g" and ts: "map (map_funs_term fg) ts = bef @ C\<langle>s\<rangle> # aft" (is "?ts = ?bca") by auto
from ts have "length ?ts = length ?bca" by auto
then have len: "length ts = length ?bca" by auto
note id = ts[unfolded map_nth_eq_conv[OF len], THEN spec, THEN mp]
let ?i = "length bef"
from len have i: "?i < length ts" by auto
from id[of ?i] have "map_funs_term fg (ts ! ?i) = C\<langle>s\<rangle>" by auto
from More(1)[OF this] obtain D u where D: "C = map_funs_ctxt fg D" and
u: "s = map_funs_term fg u" and id: "ts ! ?i = D\<langle>u\<rangle>" by auto
from ts have "take ?i ?ts = take ?i ?bca" by simp
also have "... = bef" by simp
finally have bef: "map (map_funs_term fg) (take ?i ts) = bef" by (simp add: take_map)
from ts have "drop (Suc ?i) ?ts = drop (Suc ?i) ?bca" by simp
also have "... = aft" by simp
finally have aft: "map (map_funs_term fg) (drop (Suc ?i) ts) = aft" by (simp add:drop_map)
let ?bda = "take ?i ts @ D\<langle>u\<rangle> # drop (Suc ?i) ts"
show ?case
proof (rule exI[of _ "More f (take ?i ts) D (drop (Suc ?i) ts)"],
rule exI[of _ u], simp add: u f D bef aft t)
have "ts = take ?i ts @ ts ! ?i # drop (Suc ?i) ts"
by (rule id_take_nth_drop[OF i])
also have "... = ?bda" by (simp add: id)
finally show "ts = ?bda" .
qed
qed
lemma prod_automata_from_none_root_dec:
assumes "gta_lang Q \<A> \<subseteq> {gpair s t| s t. funas_gterm s \<subseteq> \<F> \<and> funas_gterm t \<subseteq> \<F>}"
and "q |\<in>| ta_der \<A> (term_of_gterm t)" and "fst (groot_sym t) = None"
and "\<Q> \<A> |\<subseteq>| ta_reachable \<A>" and "q |\<in>| ta_productive Q \<A>"
shows "\<exists> u. t = gterm_to_None_Some u \<and> funas_gterm u \<subseteq> \<F>"
proof -
have *: "gfun_at t [] = Some (groot_sym t)" by (cases t) auto
from assms(4, 5) obtain C q\<^sub>f where ctxt: "ground_ctxt C" and
fin: "q\<^sub>f |\<in>| ta_der \<A> C\<langle>Var q\<rangle>" "q\<^sub>f |\<in>| Q"
by (auto simp: ta_productive_def'[OF assms(4)])
then obtain s v where gp: "gpair s v = (gctxt_of_ctxt C)\<langle>t\<rangle>\<^sub>G" and
funas: "funas_gterm v \<subseteq> \<F>"
using assms(1, 2) gta_langI[OF fin(2), of \<A> "(gctxt_of_ctxt C)\<langle>t\<rangle>\<^sub>G"]
by (auto simp: ta_der_ctxt ground_gctxt_of_ctxt_apply_gterm)
from gp have mem: "hole_pos C \<in> gposs s \<union> gposs v"
by auto (metis Un_iff ctxt ghole_pos_in_apply gposs_of_gpair ground_hole_pos_to_ghole)
from this have "hole_pos C \<notin> gposs s" using assms(3)
using arg_cong[OF gp, of "\<lambda> t. gfun_at t (hole_pos C)"]
using ground_hole_pos_to_ghole[OF ctxt]
using gfun_at_after_hole_pos[OF position_less_refl, of "gctxt_of_ctxt C"]
by (auto simp: gfun_at_gpair * split: if_splits)
(metis fstI gfun_at_None_ngposs_iff)+
from subst_at_gpair_nt_poss_None_Some[OF _ this, of v] this
have "t = gterm_to_None_Some (gsubt_at v (hole_pos C)) \<and> funas_gterm (gsubt_at v (hole_pos C)) \<subseteq> \<F>"
using funas mem funas_gterm_gsubt_at_subseteq
by (auto simp: gp intro!: exI[of _ "gsubt_at v (hole_pos C)"])
(metis ctxt ground_hole_pos_to_ghole gsubt_at_gctxt_apply_ghole)
then show ?thesis by blast
qed
lemma infinite_set_dec_infinite:
assumes "infinite S" and "\<And> s. s \<in> S \<Longrightarrow> \<exists> t. f t = s \<and> P t"
shows "infinite {t | t s. s \<in> S \<and> f t = s \<and> P t}" (is "infinite ?T")
proof (rule ccontr)
assume ass: "\<not> infinite ?T"
have "S \<subseteq> f ` {x . P x}" using assms(2) by auto
then show False using ass assms(1)
by (auto simp: subset_image_iff)
(smt Ball_Collect finite_imageI image_subset_iff infinite_iff_countable_subset subset_eq)
qed
lemma Q_inf_exec_impl_Q_inf:
assumes "gta_lang Q \<A> \<subseteq> {gpair s t| s t. funas_gterm s \<subseteq> fset \<F> \<and> funas_gterm t \<subseteq> fset \<F>}"
and "\<Q> \<A> |\<subseteq>| ta_reachable \<A>" and "\<Q> \<A> |\<subseteq>| ta_productive Q \<A>"
and "q \<in> Q_inf_e \<A>"
shows "q |\<in>| Q_infty \<A> \<F>"
proof -
let ?S = "{t. q |\<in>| ta_der \<A> (term_of_gterm t) \<and> fst (groot_sym t) = None}"
let ?P = "\<lambda> t. funas_gterm t \<subseteq> fset \<F> \<and> q |\<in>| ta_der \<A> (term_of_gterm (gterm_to_None_Some t))"
let ?F = "(\<lambda>(f, n). ((None, Some f), n)) |`| \<F>"
from Q_inf_e_infinite_terms_res[OF assms(4, 2)] have inf: "infinite ?S" by auto
{fix t assume "t \<in> ?S"
then have "\<exists> u. t = gterm_to_None_Some u \<and> funas_gterm u \<subseteq> fset \<F>"
using prod_automata_from_none_root_dec[OF assms(1)] assms(2, 3)
using fin_mono by fastforce}
then show ?thesis using infinite_set_dec_infinite[OF inf, of gterm_to_None_Some ?P]
by (auto simp: Q_infty_fmember) blast
qed
lemma Q_inf_impl_Q_inf_exec:
assumes "q |\<in>| Q_infty \<A> \<F>"
shows "q \<in> Q_inf_e \<A>"
proof -
let ?t_of_g = "\<lambda> t. term_of_gterm t :: ('b option \<times> 'b option, 'a) term"
let ?t_og_g2 = "\<lambda> t. term_of_gterm t :: ('b, 'a) term"
let ?inf = "(?t_og_g2 :: 'b gterm \<Rightarrow> ('b, 'a) term) ` {t |t. funas_gterm t \<subseteq> fset \<F> \<and> q |\<in>| ta_der \<A> (?t_of_g (gterm_to_None_Some t))}"
obtain n where card_st: "fcard (\<Q> \<A>) < n" by blast
from assms(1) have "infinite {t |t. funas_gterm t \<subseteq> fset \<F> \<and> q |\<in>| ta_der \<A> (?t_of_g (gterm_to_None_Some t))}"
unfolding Q_infty_def by blast
from infinite_inj_image_infinite[OF this, of "?t_og_g2"] have inf: "infinite ?inf" using inj_term_of_gterm by blast
{fix s assume "s \<in> ?inf" then have "ground s" "funas_term s \<subseteq> fset \<F>" by (auto simp: funas_term_of_gterm_conv subsetD)}
from infinte_no_depth_limit[OF inf, of "fset \<F>"] this obtain u where
funas: "funas_gterm u \<subseteq> fset \<F>" and card_d: "n < depth (?t_og_g2 u)" and reach: "q |\<in>| ta_der \<A> (?t_of_g (gterm_to_None_Some u))"
by auto blast
have "depth (?t_og_g2 u) = depth (?t_of_g (gterm_to_None_Some u))"
proof (induct u)
case (GFun f ts) then show ?case
by (auto simp: comp_def intro: nth_args_depth_eqI)
qed
from this pigeonhole_tree_automata[OF _ reach] card_st card_d obtain C2 C s v p where
ctxt: "C2 \<noteq> \<box>" "C\<langle>s\<rangle> = term_of_gterm (gterm_to_None_Some u)" "C2\<langle>v\<rangle> = s" and
loop: "p |\<in>| ta_der \<A> v \<and> p |\<in>| ta_der \<A> C2\<langle>Var p\<rangle> \<and> q |\<in>| ta_der \<A> C\<langle>Var p\<rangle>"
by auto
from ctxt have gr: "ground_ctxt C2" "ground_ctxt C" by auto (metis ground_ctxt_apply ground_term_of_gterm)+
from ta_der_to_ta_strict[OF _ gr(1)] loop obtain q' where
to_strict: "(p = q' \<or> (p, q') |\<in>| (eps \<A>)|\<^sup>+|) \<and> p |\<in>| ta_der_strict \<A> C2\<langle>Var q'\<rangle>" by fastforce
have *: "\<exists> C. C2 = map_funs_ctxt lift_None_Some C \<and> C \<noteq> \<box>" using ctxt(1, 2)
by (auto simp flip: ctxt(3)) (smt ctxt.simps(8) map_funs_term_ctxt_decomp map_term_of_gterm)
then have q_p: "(q', p) \<in> Q_inf \<A>" using to_strict ta_der_Q_inf[of p \<A> _ q'] ctxt
by auto
then have cycle: "(q', q') \<in> Q_inf \<A>" using to_strict by (auto intro: Q_inf_ta_eps_Q_inf)
show ?thesis
proof (cases "C = \<box>")
case True then show ?thesis
using cycle q_p loop by (auto intro: Q_inf_ta_eps_Q_inf)
next
case False
obtain q'' where r: "p = q'' \<or> (p, q'') |\<in>| (eps \<A>)|\<^sup>+|" "q |\<in>| ta_der_strict \<A> C\<langle>Var q''\<rangle>"
using ta_der_to_ta_strict[OF conjunct2[OF conjunct2[OF loop]] gr(2)] by auto
then have "(q'', q) \<in> Q_inf \<A>" using ta_der_Q_inf[of q \<A> _ q''] ctxt False
by auto (smt (z3) ctxt.simps(8) map_funs_term_ctxt_decomp map_term_of_gterm)+
then show ?thesis using r(1) cycle q_p
by (auto simp: Q_inf_ta_eps_Q_inf intro!: exI[of _ q'])
(meson Q_inf.trans Q_inf_ta_eps_Q_inf)+
qed
qed
lemma Q_infty_fQ_inf_e_conv:
assumes "gta_lang Q \<A> \<subseteq> {gpair s t| s t. funas_gterm s \<subseteq> fset \<F> \<and> funas_gterm t \<subseteq> fset \<F>}"
and "\<Q> \<A> |\<subseteq>| ta_reachable \<A>" and "\<Q> \<A> |\<subseteq>| ta_productive Q \<A>"
shows "Q_infty \<A> \<F> = fQ_inf_e \<A>"
using Q_inf_impl_Q_inf_exec Q_inf_exec_impl_Q_inf[OF assms]
- by (auto simp: fQ_inf_e.rep_eq fmember_iff_member_fset) fastforce
+ by (auto simp: fQ_inf_e.rep_eq) fastforce
definition Inf_reg_impl where
"Inf_reg_impl R = Inf_reg R (fQ_inf_e (ta R))"
lemma Inf_reg_impl_sound:
assumes "\<L> \<A> \<subseteq> {gpair s t| s t. funas_gterm s \<subseteq> fset \<F> \<and> funas_gterm t \<subseteq> fset \<F>}"
and "\<Q>\<^sub>r \<A> |\<subseteq>| ta_reachable (ta \<A>)" and "\<Q>\<^sub>r \<A> |\<subseteq>| ta_productive (fin \<A>) (ta \<A>)"
shows "\<L> (Inf_reg_impl \<A>) = \<L> (Inf_reg \<A> (Q_infty (ta \<A>) \<F>))"
using Q_infty_fQ_inf_e_conv[of "fin \<A>" "ta \<A>" \<F>] assms[unfolded \<L>_def]
by (simp add: Inf_reg_impl_def)
end
diff --git a/thys/Regular_Tree_Relations/RRn_Automata.thy b/thys/Regular_Tree_Relations/RRn_Automata.thy
--- a/thys/Regular_Tree_Relations/RRn_Automata.thy
+++ b/thys/Regular_Tree_Relations/RRn_Automata.thy
@@ -1,1536 +1,1568 @@
theory RRn_Automata
imports Tree_Automata_Complement Ground_Ctxt
begin
section \<open>Regular relations\<close>
subsection \<open>Encoding pairs of terms\<close>
text \<open>The encoding of two terms $s$ and $t$ is given by its tree domain, which is the union of the
domains of $s$ and $t$, and the labels, which arise from looking up each position in $s$ and $t$,
respectively.\<close>
definition gpair :: "'f gterm \<Rightarrow> 'g gterm \<Rightarrow> ('f option \<times> 'g option) gterm" where
"gpair s t = glabel (\<lambda>p. (gfun_at s p, gfun_at t p)) (gunion (gdomain s) (gdomain t))"
text \<open>We provide an efficient implementation of gpair.\<close>
definition zip_fill :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a option \<times> 'b option) list" where
"zip_fill xs ys = zip (map Some xs @ replicate (length ys - length xs) None)
(map Some ys @ replicate (length xs - length ys) None)"
lemma zip_fill_code [code]:
"zip_fill xs [] = map (\<lambda>x. (Some x, None)) xs"
"zip_fill [] ys = map (\<lambda>y. (None, Some y)) ys"
"zip_fill (x # xs) (y # ys) = (Some x, Some y) # zip_fill xs ys"
subgoal by (induct xs) (auto simp: zip_fill_def)
subgoal by (induct ys) (auto simp: zip_fill_def)
subgoal by (auto simp: zip_fill_def)
done
lemma length_zip_fill [simp]:
"length (zip_fill xs ys) = max (length xs) (length ys)"
by (auto simp: zip_fill_def)
lemma nth_zip_fill:
assumes "i < max (length xs) (length ys)"
shows "zip_fill xs ys ! i = (if i < length xs then Some (xs ! i) else None, if i < length ys then Some (ys ! i) else None)"
using assms by (auto simp: zip_fill_def nth_append)
fun gpair_impl :: "'f gterm option \<Rightarrow> 'g gterm option \<Rightarrow> ('f option \<times> 'g option) gterm" where
"gpair_impl (Some s) (Some t) = gpair s t"
| "gpair_impl (Some s) None = map_gterm (\<lambda>f. (Some f, None)) s"
| "gpair_impl None (Some t) = map_gterm (\<lambda>f. (None, Some f)) t"
| "gpair_impl None None = GFun (None, None) []"
declare gpair_impl.simps(2-4)[code]
lemma gpair_impl_code [simp, code]:
"gpair_impl (Some s) (Some t) =
(case s of GFun f ss \<Rightarrow> case t of GFun g ts \<Rightarrow>
GFun (Some f, Some g) (map (\<lambda>(s, t). gpair_impl s t) (zip_fill ss ts)))"
proof (induct "gdomain s" "gdomain t" arbitrary: s t rule: gunion.induct)
case (1 f ss g ts)
obtain f' ss' where [simp]: "s = GFun f' ss'" by (cases s)
obtain g' ts' where [simp]: "t = GFun g' ts'" by (cases t)
show ?case using 1(2,3) 1(1)[of i "ss' ! i" "ts' ! i" for i]
by (auto simp: gpair_def comp_def nth_zip_fill intro: glabel_map_gterm_conv[unfolded comp_def]
intro!: nth_equalityI)
qed
lemma gpair_code [code]:
"gpair s t = gpair_impl (Some s) (Some t)"
by simp
(* export_code gpair in Haskell *)
declare gpair_impl.simps(1)[simp del]
text \<open>We can easily prove some basic properties. I believe that proving them by induction with a
definition along the lines of @{const gpair_impl} would be very cumbersome.\<close>
lemma gpair_swap:
"map_gterm prod.swap (gpair s t) = gpair t s"
by (intro eq_gterm_by_gposs_gfun_at) (auto simp: gpair_def)
lemma gpair_assoc:
defines "f \<equiv> \<lambda>(f, gh). (f, gh \<bind> fst, gh \<bind> snd)"
defines "g \<equiv> \<lambda>(fg, h). (fg \<bind> fst, fg \<bind> snd, h)"
shows "map_gterm f (gpair s (gpair t u)) = map_gterm g (gpair (gpair s t) u)"
by (intro eq_gterm_by_gposs_gfun_at) (auto simp: gpair_def f_def g_def)
subsection \<open>Decoding of pairs\<close>
fun gcollapse :: "'f option gterm \<Rightarrow> 'f gterm option" where
"gcollapse (GFun None _) = None"
| "gcollapse (GFun (Some f) ts) = Some (GFun f (map the (filter (\<lambda>t. \<not> Option.is_none t) (map gcollapse ts))))"
lemma gcollapse_groot_None [simp]:
"groot_sym t = None \<Longrightarrow> gcollapse t = None"
"fst (groot t) = None \<Longrightarrow> gcollapse t = None"
by (cases t, simp)+
definition gfst :: "('f option \<times> 'g option) gterm \<Rightarrow> 'f gterm" where
"gfst = the \<circ> gcollapse \<circ> map_gterm fst"
definition gsnd :: "('f option \<times> 'g option) gterm \<Rightarrow> 'g gterm" where
"gsnd = the \<circ> gcollapse \<circ> map_gterm snd"
lemma filter_less_upt:
"[i\<leftarrow>[i..<m] . i < n] = [i..<min n m]"
proof (cases "i \<le> m")
case True then show ?thesis
proof (induct rule: inc_induct)
case (step n) then show ?case by (auto simp: upt_rec[of n])
qed simp
qed simp
lemma gcollapse_aux:
assumes "gposs s = {p. p \<in> gposs t \<and> gfun_at t p \<noteq> Some None}"
shows "gposs (the (gcollapse t)) = gposs s"
"\<And>p. p \<in> gposs s \<Longrightarrow> gfun_at (the (gcollapse t)) p = (gfun_at t p \<bind> id)"
proof (goal_cases)
define s' t' where "s' \<equiv> gdomain s" and "t' \<equiv> gdomain t"
have *: "gposs (the (gcollapse t)) = gposs s \<and>
(\<forall>p. p \<in> gposs s \<longrightarrow> gfun_at (the (gcollapse t)) p = (gfun_at t p \<bind> id))"
using assms s'_def t'_def
proof (induct s' t' arbitrary: s t rule: gunion.induct)
case (1 f' ss' g' ts')
obtain f ss where s [simp]: "s = GFun f ss" by (cases s)
obtain g ts where t [simp]: "t = GFun (Some g) ts"
using arg_cong[OF 1(2), of "\<lambda>P. [] \<in> P"] by (cases t) auto
have *: "i < length ts \<Longrightarrow> \<not> Option.is_none (gcollapse (ts ! i)) \<longleftrightarrow> i < length ss" for i
using arg_cong[OF 1(2), of "\<lambda>P. [i] \<in> P"] by (cases "ts ! i") auto
have l: "length ss \<le> length ts"
using arg_cong[OF 1(2), of "\<lambda>P. [length ss-1] \<in> P"] by auto
have [simp]: "[t\<leftarrow>map gcollapse ts . \<not> Option.is_none t] = take (length ss) (map gcollapse ts)"
by (subst (2) map_nth[symmetric]) (auto simp add: * filter_map comp_def filter_less_upt
cong: filter_cong[OF refl, of "[0..<length ts]", unfolded set_upt atLeastLessThan_iff]
intro!: nth_equalityI)
have [simp]: "i < length ss \<Longrightarrow> gposs (ss ! i) = gposs (the (gcollapse (ts ! i)))" for i
using conjunct1[OF 1(1), of i "ss ! i" "ts ! i"] l arg_cong[OF 1(2), of "\<lambda>P. {p. i # p \<in> P}"]
1(3,4) by simp
show ?case
proof (intro conjI allI, goal_cases A B)
case A show ?case using l by (auto simp: comp_def split: if_splits)
next
case (B p) show ?case
proof (cases p)
case (Cons i q) then show ?thesis using arg_cong[OF 1(2), of "\<lambda>P. {p. i # p \<in> P}"]
spec[OF conjunct2[OF 1(1)], of i "ss ! i" "ts ! i" q] 1(3,4) by auto
qed auto
qed
qed
{ case 1 then show ?case using * by blast
next
case 2 then show ?case using * by blast }
qed
lemma gfst_gpair:
"gfst (gpair s t) = s"
proof -
have *: "gposs s = {p \<in> gposs (map_gterm fst (gpair s t)). gfun_at (map_gterm fst (gpair s t)) p \<noteq> Some None}"
using gfun_at_nongposs
by (fastforce simp: gpair_def elim: gfun_at_possE)
show ?thesis unfolding gfst_def comp_def using gcollapse_aux[OF *]
by (auto intro!: eq_gterm_by_gposs_gfun_at simp: gpair_def)
qed
lemma gsnd_gpair:
"gsnd (gpair s t) = t"
using gfst_gpair[of t s] gpair_swap[of t s, symmetric]
by (simp add: gfst_def gsnd_def gpair_def gterm.map_comp comp_def)
lemma gpair_impl_None_Inv:
"map_gterm (the \<circ> snd) (gpair_impl None (Some t)) = t"
by (simp add: gterm.map_ident gterm.map_comp comp_def)
subsection \<open>Contexts to gpair\<close>
lemma gpair_context1:
assumes "length ts = length us"
shows "gpair (GFun f ts) (GFun f us) = GFun (Some f, Some f) (map (case_prod gpair) (zip ts us))"
using assms unfolding gpair_code by (auto intro!: nth_equalityI simp: zip_fill_def)
lemma gpair_context2:
assumes "\<And> i. i < length ts \<Longrightarrow> ts ! i = gpair (ss ! i) (us ! i)"
and "length ss = length ts" and "length us = length ts"
shows "GFun (Some f, Some h) ts = gpair (GFun f ss) (GFun h us)"
using assms unfolding gpair_code gpair_impl_code
by (auto simp: zip_fill_def intro!: nth_equalityI)
lemma map_funs_term_some_gpair:
shows "gpair t t = map_gterm (\<lambda>f. (Some f, Some f)) t"
proof (induct t)
case (GFun f ts)
then show ?case by (auto intro!: gpair_context2[symmetric])
qed
lemma gpair_inject [simp]:
"gpair s t = gpair s' t' \<longleftrightarrow> s = s' \<and> t = t'"
by (metis gfst_gpair gsnd_gpair)
abbreviation gterm_to_None_Some :: "'f gterm \<Rightarrow> ('f option \<times> 'f option) gterm" where
"gterm_to_None_Some t \<equiv> map_gterm (\<lambda>f. (None, Some f)) t"
abbreviation "gterm_to_Some_None t \<equiv> map_gterm (\<lambda>f. (Some f, None)) t"
lemma inj_gterm_to_None_Some: "inj gterm_to_None_Some"
by (meson Pair_inject gterm.inj_map inj_onI option.inject)
lemma zip_fill1:
assumes "length ss < length ts"
shows "zip_fill ss ts = zip (map Some ss) (map Some (take (length ss) ts)) @
map (\<lambda> x. (None, Some x)) (drop (length ss) ts)"
using assms by (auto simp: zip_fill_def list_eq_iff_nth_eq nth_append simp add: min.absorb2)
lemma zip_fill2:
assumes "length ts < length ss"
shows "zip_fill ss ts = zip (map Some (take (length ts) ss)) (map Some ts) @
map (\<lambda> x. (Some x, None)) (drop (length ts) ss)"
using assms by (auto simp: zip_fill_def list_eq_iff_nth_eq nth_append simp add: min.absorb2)
(* GPair position lemmas *)
(* MOVE me*)
lemma not_gposs_append [simp]:
assumes "p \<notin> gposs t"
shows "p @ q \<in> gposs t = False" using assms poss_gposs_conv
using poss_append_poss by blast
(*end Move *)
lemma gfun_at_gpair:
"gfun_at (gpair s t) p = (if p \<in> gposs s then (if p \<in> gposs t
then Some (gfun_at s p, gfun_at t p)
else Some (gfun_at s p, None)) else
(if p \<in> gposs t then Some (None, gfun_at t p) else None))"
using gfun_at_glabel by (auto simp: gpair_def)
lemma gposs_of_gpair [simp]:
shows "gposs (gpair s t) = gposs s \<union> gposs t"
by (auto simp: gpair_def)
lemma poss_to_gpair_poss:
"p \<in> gposs s \<Longrightarrow> p \<in> gposs (gpair s t)"
"p \<in> gposs t \<Longrightarrow> p \<in> gposs (gpair s t)"
by auto
lemma gsubt_at_gpair_poss:
assumes "p \<in> gposs s" and "p \<in> gposs t"
shows "gsubt_at (gpair s t) p = gpair (gsubt_at s p) (gsubt_at t p)" using assms
by (auto simp: gunion_gsubt_at_poss gfun_at_gpair intro!: eq_gterm_by_gposs_gfun_at)
lemma subst_at_gpair_nt_poss_Some_None:
assumes "p \<in> gposs s" and "p \<notin> gposs t"
shows "gsubt_at (gpair s t) p = gterm_to_Some_None (gsubt_at s p)" using assms gfun_at_poss
by (force simp: gunion_gsubt_at_poss gfun_at_gpair intro!: eq_gterm_by_gposs_gfun_at)
lemma subst_at_gpair_nt_poss_None_Some:
assumes "p \<in> gposs t" and "p \<notin> gposs s"
shows "gsubt_at (gpair s t) p = gterm_to_None_Some (gsubt_at t p)" using assms gfun_at_poss
by (force simp: gunion_gsubt_at_poss gfun_at_gpair intro!: eq_gterm_by_gposs_gfun_at)
lemma gpair_ctxt_decomposition:
fixes C defines "p \<equiv> ghole_pos C"
assumes "p \<notin> gposs s" and "gpair s t = C\<langle>gterm_to_None_Some u\<rangle>\<^sub>G"
shows "gpair s (gctxt_at_pos t p)\<langle>v\<rangle>\<^sub>G = C\<langle>gterm_to_None_Some v\<rangle>\<^sub>G"
using assms(2-)
proof -
note p[simp] = assms(1)
have pt: "p \<in> gposs t" and pc: "p \<in> gposs C\<langle>gterm_to_None_Some v\<rangle>\<^sub>G"
and pu: "p \<in> gposs C\<langle>gterm_to_None_Some u\<rangle>\<^sub>G"
using arg_cong[OF assms(3), of gposs] assms(2) ghole_pos_in_apply
by auto
have *: "gctxt_at_pos (gpair s (gctxt_at_pos t (ghole_pos C))\<langle>v\<rangle>\<^sub>G) (ghole_pos C) = gctxt_at_pos (gpair s t) (ghole_pos C)"
using assms(2) pt
by (intro eq_gctxt_at_pos)
(auto simp: gposs_gctxt_at_pos gunion_gsubt_at_poss gfun_at_gpair gfun_at_gctxt_at_pos_not_after)
have "gsubt_at (gpair s (gctxt_at_pos t p)\<langle>v\<rangle>\<^sub>G) p = gsubt_at C\<langle>gterm_to_None_Some v\<rangle>\<^sub>G p"
using pt assms(2) subst_at_gpair_nt_poss_None_Some[OF _ assms(2), of "(gctxt_at_pos t p)\<langle>v\<rangle>\<^sub>G"]
using ghole_pos_gctxt_at_pos
by (simp add: ghole_pos_in_apply)
then show ?thesis using assms(2) ghole_pos_gctxt_at_pos
using gsubst_at_gctxt_at_eq_gtermD[OF assms(3) pu]
by (intro gsubst_at_gctxt_at_eq_gtermI[OF _ pc])
(auto simp: ghole_pos_in_apply * gposs_gctxt_at_pos[OF pt, unfolded p])
qed
lemma groot_gpair [simp]:
"fst (groot (gpair s t)) = (Some (fst (groot s)), Some (fst (groot t)))"
by (cases s; cases t) (auto simp add: gpair_code)
lemma ground_ctxt_adapt_ground [intro]:
assumes "ground_ctxt C"
shows "ground_ctxt (adapt_vars_ctxt C)"
using assms by (induct C) auto
lemma adapt_vars_ctxt2 :
assumes "ground_ctxt C"
shows "adapt_vars_ctxt (adapt_vars_ctxt C) = adapt_vars_ctxt C" using assms
by (induct C) (auto simp: adapt_vars2)
subsection \<open>Encoding of lists of terms\<close>
definition gencode :: "'f gterm list \<Rightarrow> 'f option list gterm" where
"gencode ts = glabel (\<lambda>p. map (\<lambda>t. gfun_at t p) ts) (gunions (map gdomain ts))"
definition gdecode_nth :: "'f option list gterm \<Rightarrow> nat \<Rightarrow> 'f gterm" where
"gdecode_nth t i = the (gcollapse (map_gterm (\<lambda>f. f ! i) t))"
lemma gdecode_nth_gencode:
assumes "i < length ts"
shows "gdecode_nth (gencode ts) i = ts ! i"
proof -
have *: "gposs (ts ! i) = {p \<in> gposs (map_gterm (\<lambda>f. f ! i) (gencode ts)).
gfun_at (map_gterm (\<lambda>f. f ! i) (gencode ts)) p \<noteq> Some None}"
using assms
by (auto simp: gencode_def elim: gfun_at_possE dest: gfun_at_poss_gpossD) (force simp: fun_at_def' split: if_splits)
show ?thesis unfolding gdecode_nth_def comp_def using assms gcollapse_aux[OF *]
by (auto intro!: eq_gterm_by_gposs_gfun_at simp: gencode_def)
(metis (no_types) gposs_map_gterm length_map list.set_map map_nth_eq_conv nth_mem)
qed
definition gdecode :: "'f option list gterm \<Rightarrow> 'f gterm list" where
"gdecode t = (case t of GFun f ts \<Rightarrow> map (\<lambda>i. gdecode_nth t i) [0..<length f])"
lemma gdecode_gencode:
"gdecode (gencode ts) = ts"
proof (cases "gencode ts")
case (GFun f ts')
have "length f = length ts" using arg_cong[OF GFun, of "\<lambda>t. gfun_at t []"]
by (auto simp: gencode_def)
then show ?thesis using gdecode_nth_gencode[of _ ts]
by (auto intro!: nth_equalityI simp: gdecode_def GFun)
qed
definition gencode_impl :: "'f gterm option list \<Rightarrow> 'f option list gterm" where
"gencode_impl ts = glabel (\<lambda>p. map (\<lambda>t. t \<bind> (\<lambda>t. gfun_at t p)) ts) (gunions (map (case_option (GFun () []) gdomain) ts))"
lemma gencode_code [code]:
"gencode ts = gencode_impl (map Some ts)"
by (auto simp: gencode_def gencode_impl_def comp_def)
lemma gencode_singleton:
"gencode [t] = map_gterm (\<lambda>f. [Some f]) t"
using glabel_map_gterm_conv[unfolded comp_def, of "\<lambda>t. [t]" t]
by (simp add: gunions_def gencode_def)
lemma gencode_pair:
"gencode [t, u] = map_gterm (\<lambda>(f, g). [f, g]) (gpair t u)"
by (simp add: gunions_def gencode_def gpair_def map_gterm_glabel comp_def)
subsection \<open>RRn relations\<close>
definition RR1_spec where
"RR1_spec A T \<longleftrightarrow> \<L> A = T"
definition RR2_spec where
"RR2_spec A T \<longleftrightarrow> \<L> A = {gpair t u |t u. (t, u) \<in> T}"
definition RRn_spec where
"RRn_spec n A R \<longleftrightarrow> \<L> A = gencode ` R \<and> (\<forall>ts \<in> R. length ts = n)"
lemma RR1_to_RRn_spec:
assumes "RR1_spec A T"
shows "RRn_spec 1 (fmap_funs_reg (\<lambda>f. [Some f]) A) ((\<lambda>t. [t]) ` T)"
proof -
have [simp]: "inj_on (\<lambda>f. [Some f]) X" for X by (auto simp: inj_on_def)
show ?thesis using assms
by (auto simp: RR1_spec_def RRn_spec_def fmap_funs_\<L> image_comp comp_def gencode_singleton)
qed
lemma RR2_to_RRn_spec:
assumes "RR2_spec A T"
shows "RRn_spec 2 (fmap_funs_reg (\<lambda>(f, g). [f, g]) A) ((\<lambda>(t, u). [t, u]) ` T)"
proof -
have [simp]: "inj_on (\<lambda>(f, g). [f, g]) X" for X by (auto simp: inj_on_def)
show ?thesis using assms
by (auto simp: RR2_spec_def RRn_spec_def fmap_funs_\<L> image_comp comp_def prod.case_distrib gencode_pair)
qed
lemma RRn_to_RR2_spec:
assumes "RRn_spec 2 A T"
shows "RR2_spec (fmap_funs_reg (\<lambda> f. (f ! 0 , f ! 1)) A) ((\<lambda> f. (f ! 0, f ! 1)) ` T)" (is "RR2_spec ?A ?T")
proof -
{fix xs assume "xs \<in> T" then have "length xs = 2" using assms by (auto simp: RRn_spec_def)
then obtain t u where *: "xs = [t, u]"
by (metis (no_types, lifting) One_nat_def Suc_1 length_0_conv length_Suc_conv)
have **: "(\<lambda>f. (f ! 0, f ! Suc 0)) \<circ> (\<lambda>(f, g). [f, g]) = id" by auto
have "map_gterm (\<lambda>f. (f ! 0, f ! Suc 0)) (gencode xs) = gpair t u"
unfolding * gencode_pair gterm.map_comp ** gterm.map_id ..
then have "\<exists> t u. xs = [t, u] \<and> map_gterm (\<lambda>f. (f ! 0, f ! Suc 0)) (gencode xs) = gpair t u"
using * by blast}
then show ?thesis using assms
by (force simp: RR2_spec_def RRn_spec_def fmap_funs_\<L> image_comp comp_def prod.case_distrib gencode_pair image_iff Bex_def)
qed
lemma relabel_RR1_spec [simp]:
"RR1_spec (relabel_reg A) T \<longleftrightarrow> RR1_spec A T"
by (simp add: RR1_spec_def)
lemma relabel_RR2_spec [simp]:
"RR2_spec (relabel_reg A) T \<longleftrightarrow> RR2_spec A T"
by (simp add: RR2_spec_def)
lemma relabel_RRn_spec [simp]:
"RRn_spec n (relabel_reg A) T \<longleftrightarrow> RRn_spec n A T"
by (simp add: RRn_spec_def)
lemma trim_RR1_spec [simp]:
"RR1_spec (trim_reg A) T \<longleftrightarrow> RR1_spec A T"
by (simp add: RR1_spec_def \<L>_trim)
lemma trim_RR2_spec [simp]:
"RR2_spec (trim_reg A) T \<longleftrightarrow> RR2_spec A T"
by (simp add: RR2_spec_def \<L>_trim)
lemma trim_RRn_spec [simp]:
"RRn_spec n (trim_reg A) T \<longleftrightarrow> RRn_spec n A T"
by (simp add: RRn_spec_def \<L>_trim)
lemma swap_RR2_spec:
assumes "RR2_spec A R"
shows "RR2_spec (fmap_funs_reg prod.swap A) (prod.swap ` R)" using assms
by (force simp add: RR2_spec_def fmap_funs_\<L> gpair_swap image_iff)
subsection \<open>Nullary automata\<close>
lemma false_RRn_spec:
"RRn_spec n empty_reg {}"
by (auto simp: RRn_spec_def \<L>_epmty)
lemma true_RR0_spec:
"RRn_spec 0 (Reg {|q|} (TA {|[] [] \<rightarrow> q|} {||})) {[]}"
by (auto simp: RRn_spec_def \<L>_def const_ta_lang gencode_def gunions_def)
subsection \<open>Pairing RR1 languages\<close>
text \<open>cf. @{const "gpair"}.\<close>
abbreviation "lift_Some_None s \<equiv> (Some s, None)"
abbreviation "lift_None_Some s \<equiv> (None, Some s)"
abbreviation "pair_eps A B \<equiv> (\<lambda> (p, q). ((Some (fst p), q), (Some (snd p), q))) |`| (eps A |\<times>| finsert None (Some |`| \<Q> B))"
abbreviation "pair_rule \<equiv> (\<lambda> (ra, rb). TA_rule (Some (r_root ra), Some (r_root rb)) (zip_fill (r_lhs_states ra) (r_lhs_states rb)) (Some (r_rhs ra), Some (r_rhs rb)))"
lemma lift_Some_None_pord_swap [simp]:
"prod.swap \<circ> lift_Some_None = lift_None_Some"
"prod.swap \<circ> lift_None_Some = lift_Some_None"
by auto
lemma eps_to_pair_eps_Some_None:
"(p, q) |\<in>| eps \<A> \<Longrightarrow> (lift_Some_None p, lift_Some_None q) |\<in>| pair_eps \<A> \<B>"
by force
definition pair_automaton :: "('p, 'f) ta \<Rightarrow> ('q, 'g) ta \<Rightarrow> ('p option \<times> 'q option, 'f option \<times> 'g option) ta" where
"pair_automaton A B = TA
(map_ta_rule lift_Some_None lift_Some_None |`| rules A |\<union>|
map_ta_rule lift_None_Some lift_None_Some |`| rules B |\<union>|
pair_rule |`| (rules A |\<times>| rules B))
(pair_eps A B |\<union>| map_both prod.swap |`| (pair_eps B A))"
definition pair_automaton_reg where
"pair_automaton_reg R L = Reg (Some |`| fin R |\<times>| Some |`| fin L) (pair_automaton (ta R) (ta L))"
lemma pair_automaton_eps_simps:
"(lift_Some_None p, p') |\<in>| eps (pair_automaton A B) \<longleftrightarrow> (lift_Some_None p, p') |\<in>| pair_eps A B"
"(q , lift_Some_None q') |\<in>| eps (pair_automaton A B) \<longleftrightarrow> (q , lift_Some_None q') |\<in>| pair_eps A B"
by (auto simp: pair_automaton_def eps_to_pair_eps_Some_None)
lemma pair_automaton_eps_Some_SomeD:
"((Some p, Some p'), r) |\<in>| eps (pair_automaton A B) \<Longrightarrow> fst r \<noteq> None \<and> snd r \<noteq> None \<and> (Some p = fst r \<or> Some p' = snd r) \<and>
(Some p \<noteq> fst r \<longrightarrow> (p, the (fst r)) |\<in>| (eps A)) \<and> (Some p' \<noteq> snd r \<longrightarrow> (p', the (snd r)) |\<in>| (eps B))"
by (auto simp: pair_automaton_def)
lemma pair_automaton_eps_Some_SomeD2:
"(r, (Some p, Some p')) |\<in>| eps (pair_automaton A B) \<Longrightarrow> fst r \<noteq> None \<and> snd r \<noteq> None \<and> (fst r = Some p \<or> snd r = Some p') \<and>
(fst r \<noteq> Some p \<longrightarrow> (the (fst r), p) |\<in>| (eps A)) \<and> (snd r \<noteq> Some p' \<longrightarrow> (the (snd r), p') |\<in>| (eps B))"
by (auto simp: pair_automaton_def)
lemma pair_eps_Some_None:
fixes p q q'
defines "l \<equiv> (p, q)" and "r \<equiv> lift_Some_None q'"
assumes "(l, r) |\<in>| (eps (pair_automaton A B))|\<^sup>+|"
shows "q = None \<and> p \<noteq> None \<and> (the p, q') |\<in>| (eps A)|\<^sup>+|" using assms(3, 1, 2)
proof (induct arbitrary: q' q rule: ftrancl_induct)
case (Step b)
then show ?case unfolding pair_automaton_eps_simps
by (auto simp: pair_automaton_eps_simps)
(meson not_ftrancl_into)
qed (auto simp: pair_automaton_def)
lemma pair_eps_Some_Some:
fixes p q
defines "l \<equiv> (Some p, Some q)"
assumes "(l, r) |\<in>| (eps (pair_automaton A B))|\<^sup>+|"
shows "fst r \<noteq> None \<and> snd r \<noteq> None \<and>
(fst l \<noteq> fst r \<longrightarrow> (p, the (fst r)) |\<in>| (eps A)|\<^sup>+|) \<and>
(snd l \<noteq> snd r \<longrightarrow> (q, the (snd r)) |\<in>| (eps B)|\<^sup>+|)"
using assms(2, 1)
proof (induct arbitrary: p q rule: ftrancl_induct)
case (Step b c)
then obtain r r' where *: "b = (Some r, Some r')" by (cases b) auto
show ?case using Step(2)
using pair_automaton_eps_Some_SomeD[OF Step(3)[unfolded *]]
by (auto simp: *) (meson not_ftrancl_into)+
qed (auto simp: pair_automaton_def)
lemma pair_eps_Some_Some2:
fixes p q
defines "r \<equiv> (Some p, Some q)"
assumes "(l, r) |\<in>| (eps (pair_automaton A B))|\<^sup>+|"
shows "fst l \<noteq> None \<and> snd l \<noteq> None \<and>
(fst l \<noteq> fst r \<longrightarrow> (the (fst l), p) |\<in>| (eps A)|\<^sup>+|) \<and>
(snd l \<noteq> snd r \<longrightarrow> (the (snd l), q) |\<in>| (eps B)|\<^sup>+|)"
using assms(2, 1)
proof (induct arbitrary: p q rule: ftrancl_induct)
case (Step b c)
from pair_automaton_eps_Some_SomeD2[OF Step(3)]
obtain r r' where *: "c = (Some r, Some r')" by (cases c) auto
from Step(2)[OF this] show ?case
using pair_automaton_eps_Some_SomeD[OF Step(3)[unfolded *]]
by (auto simp: *) (meson not_ftrancl_into)+
qed (auto simp: pair_automaton_def)
lemma map_pair_automaton:
"pair_automaton (fmap_funs_ta f A) (fmap_funs_ta g B) =
fmap_funs_ta (\<lambda>(a, b). (map_option f a, map_option g b)) (pair_automaton A B)" (is "?Ls = ?Rs")
proof -
let ?ls = "pair_rule \<circ> map_prod (map_ta_rule id f) (map_ta_rule id g)"
let ?rs = "map_ta_rule id (\<lambda>(a, b). (map_option f a, map_option g b)) \<circ> pair_rule"
have *: "(\<lambda>(a, b). (map_option f a, map_option g b)) \<circ> lift_Some_None = lift_Some_None \<circ> f"
"(\<lambda>(a, b). (map_option f a, map_option g b)) \<circ> lift_None_Some = lift_None_Some \<circ> g"
by (auto simp: comp_def)
have "?ls x = ?rs x" for x
by (cases x) (auto simp: ta_rule.map_sel)
then have [simp]: "?ls = ?rs" by blast
then have "rules ?Ls = rules ?Rs"
unfolding pair_automaton_def fmap_funs_ta_def
by (simp add: fimage_funion map_ta_rule_comp * map_prod_ftimes)
moreover have "eps ?Ls = eps ?Rs"
unfolding pair_automaton_def fmap_funs_ta_def
by (simp add: fimage_funion \<Q>_def)
ultimately show ?thesis
by (intro TA_equalityI) simp
qed
lemmas map_pair_automaton_12 =
map_pair_automaton[of _ _ id, unfolded fmap_funs_ta_id option.map_id]
map_pair_automaton[of id _ _, unfolded fmap_funs_ta_id option.map_id]
lemma fmap_states_funs_ta_commute:
"fmap_states_ta f (fmap_funs_ta g A) = fmap_funs_ta g (fmap_states_ta f A)"
proof -
have [simp]: "map_ta_rule f id (map_ta_rule id g r) = map_ta_rule id g (map_ta_rule f id r)" for r
by (cases r) auto
show ?thesis
by (auto simp: ta_rule.case_distrib fmap_states_ta_def fmap_funs_ta_def fimage_iff fBex_def split: ta_rule.splits)
qed
lemma states_pair_automaton:
"\<Q> (pair_automaton A B) |\<subseteq>| (finsert None (Some |`| \<Q> A) |\<times>| (finsert None (Some |`| \<Q> B)))"
unfolding pair_automaton_def
apply (intro \<Q>_subseteq_I)
apply (auto simp: ta_rule.map_sel in_fset_conv_nth nth_zip_fill rule_statesD eps_statesD)
apply (simp add: rule_statesD)+
done
lemma swap_pair_automaton:
assumes "(p, q) |\<in>| ta_der (pair_automaton A B) (term_of_gterm t)"
shows "(q, p) |\<in>| ta_der (pair_automaton B A) (term_of_gterm (map_gterm prod.swap t))"
proof -
let ?m = "map_ta_rule prod.swap prod.swap"
have [simp]: "map prod.swap (zip_fill xs ys) = zip_fill ys xs" for xs ys
by (auto simp: zip_fill_def zip_nth_conv comp_def intro!: nth_equalityI)
let ?swp = "\<lambda>X. fmap_states_ta prod.swap (fmap_funs_ta prod.swap X)"
{ fix A B
let ?AB = "?swp (pair_automaton A B)" and ?BA = "pair_automaton B A"
have "eps ?AB |\<subseteq>| eps ?BA" "rules ?AB |\<subseteq>| rules ?BA"
- by (auto simp: fmap_states_ta_def fmap_funs_ta_def pair_automaton_def fimage_iff ta_rule.map_comp)
- force+
+ by (auto simp: fmap_states_ta_def fmap_funs_ta_def pair_automaton_def
+ fimage_funion comp_def prod.map_comp ta_rule.map_comp)
} note * = this
let ?BA = "?swp (?swp (pair_automaton B A))" and ?AB = "?swp (pair_automaton A B)"
have **: "r |\<in>| rules (pair_automaton B A) \<Longrightarrow> ?m r |\<in>| ?m |`| rules (pair_automaton B A)" for r
by blast
have "r |\<in>| rules ?BA \<Longrightarrow> r |\<in>| rules ?AB" "e |\<in>| eps ?BA \<Longrightarrow> e |\<in>| eps ?AB" for r e
using *[of B A] map_ta_rule_prod_swap_id
unfolding fmap_funs_ta_def fmap_states_ta_def
- by (auto simp: map_ta_rule_comp fimage_iff fBex_def ta_rule.map_id0 intro!: exI[of _ "?m r"])
+ by (auto simp: map_ta_rule_comp image_iff fBex_def ta_rule.map_id0 intro!: bexI[of _ "?m r"])
then have "eps ?BA |\<subseteq>| eps ?AB" "rules ?BA |\<subseteq>| rules ?AB"
by blast+
then have "fmap_states_ta prod.swap (fmap_funs_ta prod.swap (pair_automaton A B)) = pair_automaton B A"
using *[of A B] by (simp add: fmap_states_funs_ta_commute fmap_funs_ta_comp fmap_states_ta_comp TA_equalityI)
then show ?thesis
using ta_der_fmap_states_ta[OF ta_der_fmap_funs_ta[OF assms], of prod.swap prod.swap]
by (simp add: gterm.map_comp comp_def)
qed
lemma to_ta_der_pair_automaton:
"p |\<in>| ta_der A (term_of_gterm t) \<Longrightarrow>
(Some p, None) |\<in>| ta_der (pair_automaton A B) (term_of_gterm (map_gterm (\<lambda>f. (Some f, None)) t))"
"q |\<in>| ta_der B (term_of_gterm u) \<Longrightarrow>
(None, Some q) |\<in>| ta_der (pair_automaton A B) (term_of_gterm (map_gterm (\<lambda>f. (None, Some f)) u))"
"p |\<in>| ta_der A (term_of_gterm t) \<Longrightarrow> q |\<in>| ta_der B (term_of_gterm u) \<Longrightarrow>
(Some p, Some q) |\<in>| ta_der (pair_automaton A B) (term_of_gterm (gpair t u))"
proof (goal_cases sn ns ss)
let ?AB = "pair_automaton A B"
have 1: "(Some p, None) |\<in>| ta_der (pair_automaton A B) (term_of_gterm (map_gterm (\<lambda>f. (Some f, None)) s))"
if "p |\<in>| ta_der A (term_of_gterm s)" for A B p s
- by (intro fsubsetD[OF ta_der_mono, OF _ _ ta_der_fmap_states_ta[OF ta_der_fmap_funs_ta[OF that]],
- unfolded map_term_of_gterm id_def gterm.map_ident])
- (auto simp add: pair_automaton_def fmap_states_ta_def fmap_funs_ta_def ta_rule.map_comp image_iff eps_to_pair_eps_Some_None)
+ proof (rule fsubsetD[OF ta_der_mono])
+ show "rules (fmap_states_ta lift_Some_None (fmap_funs_ta lift_Some_None A)) |\<subseteq>|
+ rules (pair_automaton A B)"
+ by (auto simp: fmap_states_ta_def fmap_funs_ta_def comp_def ta_rule.map_comp
+ pair_automaton_def)
+ next
+ show "eps (fmap_states_ta lift_Some_None (fmap_funs_ta lift_Some_None A)) |\<subseteq>|
+ eps (pair_automaton A B)"
+ by (rule fsubsetI)
+ (auto simp: fmap_states_ta_def fmap_funs_ta_def pair_automaton_def comp_def fimage.rep_eq
+ dest: eps_to_pair_eps_Some_None)
+ next
+ show "lift_Some_None p |\<in>| ta_der
+ (fmap_states_ta lift_Some_None (fmap_funs_ta lift_Some_None A))
+ (term_of_gterm (gterm_to_Some_None s))"
+ using ta_der_fmap_states_ta[OF ta_der_fmap_funs_ta[OF that], of lift_Some_None]
+ using ta_der_fmap_funs_ta ta_der_to_fmap_states_der that by fastforce
+ qed
have 2: "q |\<in>| ta_der B (term_of_gterm t) \<Longrightarrow>
(None, Some q) |\<in>| ta_der ?AB (term_of_gterm (map_gterm (\<lambda>g. (None, Some g)) t))"
for q t using swap_pair_automaton[OF 1[of q B t A]] by (simp add: gterm.map_comp comp_def)
{
case sn then show ?case by (rule 1)
next
case ns then show ?case by (rule 2)
next
case ss then show ?case
proof (induct t arbitrary: p q u)
case (GFun f ts)
obtain g us where u [simp]: "u = GFun g us" by (cases u)
obtain p' ps where p': "f ps \<rightarrow> p' |\<in>| rules A" "p' = p \<or> (p', p) |\<in>| (eps A)|\<^sup>+|" "length ps = length ts"
"\<And>i. i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der A (term_of_gterm (ts ! i))" using GFun(2) by auto
obtain q' qs where q': "g qs \<rightarrow> q' |\<in>| rules B" "q' = q \<or> (q', q) |\<in>| (eps B)|\<^sup>+|" "length qs = length us"
"\<And>i. i < length us \<Longrightarrow> qs ! i |\<in>| ta_der B (term_of_gterm (us ! i))" using GFun(3) by auto
have *: "Some p |\<in>| Some |`| \<Q> A" "Some q' |\<in>| Some |`| \<Q> B"
using p'(2,1) q'(1)
by (auto simp: rule_statesD eps_trancl_statesD)
have eps: "p' = p \<and> q' = q \<or> ((Some p', Some q'), (Some p, Some q)) |\<in>| (eps ?AB)|\<^sup>+|"
proof (cases "p' = p")
case True note p = this then show ?thesis
proof (cases "q' = q")
case False
- then have "(q', q) |\<in>| (eps B)|\<^sup>+|" using q'(2) by auto
- then show ?thesis using p'(1)
- using ftrancl_map[of "eps B" "\<lambda>q. (Some p', Some q)" "eps ?AB" q' q]
- by (auto simp: p pair_automaton_def fimage_iff fBex_def rule_statesD)
+ then have "(q', q) |\<in>| (eps B)|\<^sup>+|" using q'(2) by auto
+ hence "((Some p', Some q'), Some p', Some q) |\<in>| (eps (pair_automaton A B))|\<^sup>+|"
+ proof (rule ftrancl_map[rotated])
+ fix x y
+ assume "(x, y) |\<in>| eps B"
+ thus "((Some p', Some x), Some p', Some y) |\<in>| eps (pair_automaton A B)"
+ using p'(1)[THEN rule_statesD(1), simplified]
+ apply (simp add: pair_automaton_def image_iff fSigma.rep_eq)
+ by fastforce
+ qed
+ thus ?thesis
+ by (simp add: p)
qed (simp add: p)
next
case False note p = this
then have *: "(p', p) |\<in>| (eps A)|\<^sup>+|" using p'(2) by auto
then have eps: "((Some p', Some q'), Some p, Some q') |\<in>| (eps (pair_automaton A B))|\<^sup>+|"
- using q'(1) ftrancl_map[of "eps A" "\<lambda>p. (Some p, Some q')" "eps ?AB" p' p]
- by (auto simp: p pair_automaton_def fimage_iff fBex_def rule_statesD)
- show ?thesis
+ proof (rule ftrancl_map[rotated])
+ fix x y
+ assume "(x, y) |\<in>| eps A"
+ then show "((Some x, Some q'), Some y, Some q') |\<in>| eps (pair_automaton A B)"
+ using q'(1)[THEN rule_statesD(1), simplified]
+ apply (simp add: pair_automaton_def image_iff fSigma.rep_eq)
+ by fastforce
+ qed
+ then show ?thesis
proof (cases "q' = q")
case True then show ?thesis using eps
by simp
next
case False
then have "(q', q) |\<in>| (eps B)|\<^sup>+|" using q'(2) by auto
then have "((Some p, Some q'), Some p, Some q) |\<in>| (eps (pair_automaton A B))|\<^sup>+|"
- using * ftrancl_map[of "eps B" "\<lambda>q. (Some p, Some q)" "eps ?AB" q' q]
- by (auto simp: p pair_automaton_def fimage_iff fBex_def eps_trancl_statesD)
+ apply (rule ftrancl_map[rotated])
+ using *[THEN eps_trancl_statesD]
+ apply (simp add: p pair_automaton_def image_iff fSigma.rep_eq)
+ by fastforce
+
then show ?thesis using eps
by (meson ftrancl_trans)
qed
qed
have "(Some f, Some g) zip_fill ps qs \<rightarrow> (Some p', Some q') |\<in>| rules ?AB"
using p'(1) q'(1) by (force simp: pair_automaton_def)
then show ?case using GFun(1)[OF nth_mem p'(4) q'(4)] p'(1 - 3) q'(1 - 3) eps
using 1[OF p'(4), of _ B] 2[OF q'(4)]
by (auto simp: gpair_code nth_zip_fill less_max_iff_disj
intro!: exI[of _ "zip_fill ps qs"] exI[of _ "Some p'"] exI[of _ "Some q'"])
qed
}
qed
lemma from_ta_der_pair_automaton:
"(None, None) |\<notin>| ta_der (pair_automaton A B) (term_of_gterm s)"
"(Some p, None) |\<in>| ta_der (pair_automaton A B) (term_of_gterm s) \<Longrightarrow>
\<exists>t. p |\<in>| ta_der A (term_of_gterm t) \<and> s = map_gterm (\<lambda>f. (Some f, None)) t"
"(None, Some q) |\<in>| ta_der (pair_automaton A B) (term_of_gterm s) \<Longrightarrow>
\<exists>u. q |\<in>| ta_der B (term_of_gterm u) \<and> s = map_gterm (\<lambda>f. (None, Some f)) u"
"(Some p, Some q) |\<in>| ta_der (pair_automaton A B) (term_of_gterm s) \<Longrightarrow>
\<exists>t u. p |\<in>| ta_der A (term_of_gterm t) \<and> q |\<in>| ta_der B (term_of_gterm u) \<and> s = gpair t u"
proof (goal_cases nn sn ns ss)
let ?AB = "pair_automaton A B"
{ fix p s A B assume "(Some p, None) |\<in>| ta_der (pair_automaton A B) (term_of_gterm s)"
then have "\<exists>t. p |\<in>| ta_der A (term_of_gterm t) \<and> s = map_gterm (\<lambda>f. (Some f, None)) t"
proof (induct s arbitrary: p)
case (GFun h ss)
obtain rs sp nq where r: "h rs \<rightarrow> (sp, nq) |\<in>| rules (pair_automaton A B)"
"sp = Some p \<and> nq = None \<or> ((sp, nq), (Some p, None)) |\<in>| (eps (pair_automaton A B))|\<^sup>+|" "length rs = length ss"
"\<And>i. i < length ss \<Longrightarrow> rs ! i |\<in>| ta_der (pair_automaton A B) (term_of_gterm (ss ! i))"
using GFun(2) by auto
obtain p' where pq: "sp = Some p'" "nq = None" "p' = p \<or> (p', p) |\<in>| (eps A)|\<^sup>+|"
using r(2) pair_eps_Some_None[of sp nq p A B]
by (cases sp) auto
obtain f ps where h: "h = lift_Some_None f" "rs = map lift_Some_None ps"
"f ps \<rightarrow> p' |\<in>| rules A"
using r(1) unfolding pq(1, 2) by (auto simp: pair_automaton_def map_ta_rule_cases)
obtain ts where "\<And>i. i < length ss \<Longrightarrow>
ps ! i |\<in>| ta_der A (term_of_gterm (ts i)) \<and> ss ! i = map_gterm (\<lambda>f. (Some f, None)) (ts i)"
using GFun(1)[OF nth_mem, of i "ps ! i" for i] r(4)[unfolded h(2)] r(3)[unfolded h(2) length_map]
by auto metis
then show ?case using h(3) pq(3) r(3)[unfolded h(2) length_map]
by (intro exI[of _ "GFun f (map ts [0..<length ss])"]) (auto simp: h(1) intro!: nth_equalityI)
qed
} note 1 = this
have 2: "\<exists>u. q |\<in>| ta_der B (term_of_gterm u) \<and> s = map_gterm (\<lambda>g. (None, Some g)) u"
if "(None, Some q) |\<in>| ta_der ?AB (term_of_gterm s)" for q s
using 1[OF swap_pair_automaton, OF that]
by (auto simp: gterm.map_comp comp_def gterm.map_ident
dest!: arg_cong[of "map_gterm prod.swap _" _ "map_gterm prod.swap", unfolded gterm.map_comp])
{
case nn
then show ?case
by (intro ta_der_not_reach) (auto simp: pair_automaton_def map_ta_rule_cases)
next
case sn then show ?case by (rule 1)
next
case ns then show ?case by (rule 2)
next
case ss then show ?case
proof (induct s arbitrary: p q)
case (GFun h ss)
obtain rs sp sq where r: "h rs \<rightarrow> (sp, sq) |\<in>| rules ?AB"
"sp = Some p \<and> sq = Some q \<or> ((sp, sq), (Some p, Some q)) |\<in>| (eps ?AB)|\<^sup>+|" "length rs = length ss"
"\<And>i. i < length ss \<Longrightarrow> rs ! i |\<in>| ta_der ?AB (term_of_gterm (ss ! i))"
using GFun(2) by auto
from r(2) have "sp \<noteq> None" "sq \<noteq> None" using pair_eps_Some_Some2[of "(sp, sq)" p q]
by auto
then obtain p' q' where pq: "sp = Some p'" "sq = Some q'"
"p' = p \<or> (p', p) |\<in>| (eps A)|\<^sup>+|" "q' = q \<or> (q', q) |\<in>| (eps B)|\<^sup>+|"
using r(2) pair_eps_Some_Some[where ?r = "(Some p, Some q)" and ?A = A and ?B = B]
using pair_eps_Some_Some2[of "(sp, sq)" p q]
by (cases sp; cases sq) auto
obtain f g ps qs where h: "h = (Some f, Some g)" "rs = zip_fill ps qs"
"f ps \<rightarrow> p' |\<in>| rules A" "g qs \<rightarrow> q' |\<in>| rules B"
using r(1) unfolding pq(1,2) by (auto simp: pair_automaton_def map_ta_rule_cases)
have *: "\<exists>t u. ps ! i |\<in>| ta_der A (term_of_gterm t) \<and> qs ! i |\<in>| ta_der B (term_of_gterm u) \<and> ss ! i = gpair t u"
if "i < length ps" "i < length qs" for i
using GFun(1)[OF nth_mem, of i "ps ! i" "qs ! i"] r(4)[unfolded pq(1,2) h(2), of i] that
r(3)[symmetric] by (auto simp: nth_zip_fill h(2))
{ fix i assume "i < length ss"
then have "\<exists>t u. (i < length ps \<longrightarrow> ps ! i |\<in>| ta_der A (term_of_gterm t)) \<and>
(i < length qs \<longrightarrow> qs ! i |\<in>| ta_der B (term_of_gterm u)) \<and>
ss ! i = gpair_impl (if i < length ps then Some t else None) (if i < length qs then Some u else None)"
using *[of i] 1[of "ps ! i" A B "ss ! i"] 2[of "qs ! i" "ss ! i"] r(4)[of i] r(3)[symmetric]
by (cases "i < length ps"; cases "i < length qs")
(auto simp add: h(2) nth_zip_fill less_max_iff_disj gpair_code split: gterm.splits) }
then obtain ts us where "\<And>i. i < length ss \<Longrightarrow>
(i < length ps \<longrightarrow> ps ! i |\<in>| ta_der A (term_of_gterm (ts i))) \<and>
(i < length qs \<longrightarrow> qs ! i |\<in>| ta_der B (term_of_gterm (us i))) \<and>
ss ! i = gpair_impl (if i < length ps then Some (ts i) else None) (if i < length qs then Some (us i) else None)"
by metis
moreover then have "\<And>i. i < length ps \<Longrightarrow> ps ! i |\<in>| ta_der A (term_of_gterm (ts i))"
"\<And>i. i < length qs \<Longrightarrow> qs ! i |\<in>| ta_der B (term_of_gterm (us i))"
using r(3)[unfolded h(2)] by auto
ultimately show ?case using h(3,4) pq(3,4) r(3)[symmetric]
by (intro exI[of _ "GFun f (map ts [0..<length ps])"] exI[of _ "GFun g (map us [0..<length qs])"])
(auto simp: gpair_code h(1,2) less_max_iff_disj nth_zip_fill intro!: nth_equalityI split: prod.splits gterm.splits)
qed
}
qed
lemma diagonal_automaton:
assumes "RR1_spec A R"
shows "RR2_spec (fmap_funs_reg (\<lambda>f. (Some f, Some f)) A) {(s, s) | s. s \<in> R}" using assms
by (auto simp: RR2_spec_def RR1_spec_def fmap_funs_reg_def fmap_funs_gta_lang map_funs_term_some_gpair \<L>_def)
lemma pair_automaton:
assumes "RR1_spec A T" "RR1_spec B U"
shows "RR2_spec (pair_automaton_reg A B) (T \<times> U)"
proof -
let ?AB = "pair_automaton (ta A) (ta B)"
{ fix t u assume t: "t \<in> T" and u: "u \<in> U"
obtain p where p: "p |\<in>| fin A" "p |\<in>| ta_der (ta A) (term_of_gterm t)" using t assms(1)
by (auto simp: RR1_spec_def gta_lang_def \<L>_def gta_der_def)
obtain q where q: "q |\<in>| fin B" "q |\<in>| ta_der (ta B) (term_of_gterm u)" using u assms(2)
by (auto simp: RR1_spec_def gta_lang_def \<L>_def gta_der_def)
have "gpair t u \<in> \<L> (pair_automaton_reg A B)" using p(1) q(1) to_ta_der_pair_automaton(3)[OF p(2) q(2)]
by (auto simp: pair_automaton_reg_def \<L>_def)
} moreover
{ fix s assume "s \<in> \<L> (pair_automaton_reg A B)"
from this[unfolded gta_lang_def \<L>_def]
obtain r where r: "r |\<in>| fin (pair_automaton_reg A B)" "r |\<in>| gta_der ?AB s"
by (auto simp: pair_automaton_reg_def)
obtain p q where pq: "r = (Some p, Some q)" "p |\<in>| fin A" "q |\<in>| fin B"
using r(1) by (auto simp: pair_automaton_reg_def)
from from_ta_der_pair_automaton(4)[OF r(2)[unfolded pq(1) gta_der_def]]
obtain t u where "p |\<in>| ta_der (ta A) (term_of_gterm t)" "q |\<in>| ta_der (ta B) (term_of_gterm u)" "s = gpair t u"
by (elim exE conjE) note * = this
then have "t \<in> \<L> A" "u \<in> \<L> B" using pq(2,3)
by (auto simp: \<L>_def)
then have "\<exists>t u. s = gpair t u \<and> t \<in> T \<and> u \<in> U" using assms
by (auto simp: RR1_spec_def *(3) intro!: exI[of _ t, OF exI[of _ u]])
} ultimately show ?thesis by (auto simp: RR2_spec_def)
qed
lemma pair_automaton':
shows "\<L> (pair_automaton_reg A B) = case_prod gpair ` (\<L> A \<times> \<L> B)"
using pair_automaton[of A _ B] by (auto simp: RR1_spec_def RR2_spec_def)
subsection \<open>Collapsing\<close>
text \<open>cf. @{const "gcollapse"}.\<close>
fun collapse_state_list where
"collapse_state_list Qn Qs [] = [[]]"
| "collapse_state_list Qn Qs (q # qs) = (let rec = collapse_state_list Qn Qs qs in
(if q |\<in>| Qn \<and> q |\<in>| Qs then map (Cons None) rec @ map (Cons (Some q)) rec
else if q |\<in>| Qn then map (Cons None) rec
else if q |\<in>| Qs then map (Cons (Some q)) rec
else [[]]))"
lemma collapse_state_list_inner_length:
assumes "qss = collapse_state_list Qn Qs qs"
and "\<forall> i < length qs. qs ! i |\<in>| Qn \<or> qs ! i |\<in>| Qs"
and "i < length qss"
shows "length (qss ! i) = length qs" using assms
proof (induct qs arbitrary: qss i)
case (Cons q qs)
have "\<forall>i<length qs. qs ! i |\<in>| Qn \<or> qs ! i |\<in>| Qs" using Cons(3) by auto
then show ?case using Cons(1)[of "collapse_state_list Qn Qs qs"] Cons(2-)
by (auto simp: Let_def nth_append)
qed auto
lemma collapse_fset_inv_constr:
assumes "\<forall> i < length qs'. qs ! i |\<in>| Qn \<and> qs' ! i = None \<or>
qs ! i |\<in>| Qs \<and> qs' ! i = Some (qs ! i)"
and "length qs = length qs'"
shows "qs' |\<in>| fset_of_list (collapse_state_list Qn Qs qs)" using assms
proof (induct qs arbitrary: qs')
case (Cons q qs)
have "(tl qs') |\<in>| fset_of_list (collapse_state_list Qn Qs qs)" using Cons(2-)
by (intro Cons(1)[of "tl qs'"]) (cases qs', auto)
then show ?case using Cons(2-)
by (cases qs') (auto simp: Let_def image_def)
qed auto
lemma collapse_fset_inv_constr2:
assumes "\<forall> i < length qs. qs ! i |\<in>| Qn \<or> qs ! i |\<in>| Qs"
and "qs' |\<in>| fset_of_list (collapse_state_list Qn Qs qs)" and "i < length qs'"
shows "qs ! i |\<in>| Qn \<and> qs' ! i = None \<or> qs ! i |\<in>| Qs \<and> qs' ! i = Some (qs ! i)"
using assms
proof (induct qs arbitrary: qs' i)
case (Cons a qs)
have "i \<noteq> 0 \<Longrightarrow> qs ! (i - 1) |\<in>| Qn \<and> tl qs' ! (i - 1) = None \<or> qs ! (i - 1) |\<in>| Qs \<and> tl qs' ! (i - 1) = Some (qs ! (i - 1))"
using Cons(2-)
by (intro Cons(1)[of "tl qs'" "i - 1"]) (auto simp: Let_def split: if_splits)
then show ?case using Cons(2-)
by (cases i) (auto simp: Let_def split: if_splits)
qed auto
definition collapse_rule where
"collapse_rule A Qn Qs =
|\<Union>| ((\<lambda> r. fset_of_list (map (\<lambda> qs. TA_rule (r_root r) qs (Some (r_rhs r))) (collapse_state_list Qn Qs (r_lhs_states r)))) |`|
ffilter (\<lambda> r. (\<forall> i < length (r_lhs_states r). r_lhs_states r ! i |\<in>| Qn \<or> r_lhs_states r ! i |\<in>| Qs))
(ffilter (\<lambda> r. r_root r \<noteq> None) (rules A)))"
definition collapse_rule_fset where
"collapse_rule_fset A Qn Qs = (\<lambda> r. TA_rule (the (r_root r)) (map the (filter (\<lambda>q. \<not> Option.is_none q) (r_lhs_states r))) (the (r_rhs r))) |`|
collapse_rule A Qn Qs"
lemma collapse_rule_set_conv:
"fset (collapse_rule_fset A Qn Qs) = {TA_rule f (map the (filter (\<lambda>q. \<not> Option.is_none q) qs')) q | f qs qs' q.
TA_rule (Some f) qs q |\<in>| rules A \<and> length qs = length qs' \<and>
(\<forall>i < length qs. qs ! i |\<in>| Qn \<and> qs' ! i = None \<or> qs ! i |\<in>| Qs \<and> (qs' ! i) = Some (qs ! i))} " (is "?Ls = ?Rs")
proof
{fix f qs' q qs assume ass: "TA_rule (Some f) qs q |\<in>| rules A"
"length qs = length qs'"
"\<forall>i < length qs. qs ! i |\<in>| Qn \<and> qs' ! i = None \<or> qs ! i |\<in>| Qs \<and> (qs' ! i) = Some (qs ! i)"
then have "TA_rule (Some f) qs' (Some q) |\<in>| collapse_rule A Qn Qs"
using collapse_fset_inv_constr[of qs' qs Qn Qs] unfolding collapse_rule_def
by (auto simp: fset_of_list.rep_eq fimage_iff intro!: fBexI[of _ "TA_rule (Some f) qs q"])
then have "TA_rule f (map the (filter (\<lambda>q. \<not> Option.is_none q) qs')) q \<in> ?Ls"
unfolding collapse_rule_fset_def
- by (auto simp: image_iff Bex_def fmember_iff_member_fset intro!: exI[of _"TA_rule (Some f) qs' (Some q)"])}
+ by (auto simp: image_iff Bex_def intro!: exI[of _"TA_rule (Some f) qs' (Some q)"])}
then show "?Rs \<subseteq> ?Ls" by blast
next
{fix f qs q assume ass: "TA_rule f qs q \<in> ?Ls"
then obtain ps qs' where w: "TA_rule (Some f) ps q |\<in>| rules A"
"\<forall> i < length ps. ps ! i |\<in>| Qn \<or> ps ! i |\<in>| Qs"
"qs' |\<in>| fset_of_list (collapse_state_list Qn Qs ps)"
"qs = map the (filter (\<lambda>q. \<not> Option.is_none q) qs')"
unfolding collapse_rule_fset_def collapse_rule_def
- by (auto simp: fmember_iff_member_fset ffUnion.rep_eq fset_of_list.rep_eq) (metis ta_rule.collapse)
+ by (auto simp: ffUnion.rep_eq fset_of_list.rep_eq) (metis ta_rule.collapse)
then have "\<forall> i < length qs'. ps ! i |\<in>| Qn \<and> qs' ! i = None \<or> ps ! i |\<in>| Qs \<and> qs' ! i = Some (ps ! i)"
using collapse_fset_inv_constr2
by metis
moreover have "length ps = length qs'"
using collapse_state_list_inner_length[of _ Qn Qs ps]
using w(2, 3) calculation(1)
by (auto simp: in_fset_conv_nth)
ultimately have "TA_rule f qs q \<in> ?Rs"
using w(1) unfolding w(4)
by auto fastforce}
then show "?Ls \<subseteq> ?Rs"
by (intro subsetI) (metis (no_types, lifting) ta_rule.collapse)
qed
lemma collapse_rule_fmember [simp]:
"TA_rule f qs q |\<in>| (collapse_rule_fset A Qn Qs) \<longleftrightarrow> (\<exists> qs' ps.
qs = map the (filter (\<lambda>q. \<not> Option.is_none q) qs') \<and> TA_rule (Some f) ps q |\<in>| rules A \<and> length ps = length qs' \<and>
(\<forall>i < length ps. ps ! i |\<in>| Qn \<and> qs' ! i = None \<or> ps ! i |\<in>| Qs \<and> (qs' ! i) = Some (ps ! i)))"
- unfolding fmember_iff_member_fset collapse_rule_set_conv
+ unfolding collapse_rule_set_conv
by auto
definition "Qn A \<equiv> (let S = (r_rhs |`| ffilter (\<lambda> r. r_root r = None) (rules A)) in (eps A)|\<^sup>+| |``| S |\<union>| S)"
definition "Qs A \<equiv> (let S = (r_rhs |`| ffilter (\<lambda> r. r_root r \<noteq> None) (rules A)) in (eps A)|\<^sup>+| |``| S |\<union>| S)"
lemma Qn_member_iff [simp]:
"q |\<in>| Qn A \<longleftrightarrow> (\<exists> ps p. TA_rule None ps p |\<in>| rules A \<and> (p = q \<or> (p, q) |\<in>| (eps A)|\<^sup>+|))" (is "?Ls \<longleftrightarrow> ?Rs")
proof -
{assume ass: "q |\<in>| Qn A" then obtain r where
"r_rhs r = q \<or> (r_rhs r, q) |\<in>| (eps A)|\<^sup>+|" "r |\<in>| rules A" "r_root r = None"
- by (force simp: Qn_def Image_def image_def Let_def fImage.rep_eq simp flip: fmember_iff_member_fset)
+ by (force simp: Qn_def Image_def image_def Let_def fImage.rep_eq)
then have "?Ls \<Longrightarrow> ?Rs" by (cases r) auto}
- moreover have "?Rs \<Longrightarrow> ?Ls" by (force simp: Qn_def Image_def image_def Let_def fImage.rep_eq fmember_iff_member_fset)
+ moreover have "?Rs \<Longrightarrow> ?Ls" by (force simp: Qn_def Image_def image_def Let_def fImage.rep_eq)
ultimately show ?thesis by blast
qed
lemma Qs_member_iff [simp]:
"q |\<in>| Qs A \<longleftrightarrow> (\<exists> f ps p. TA_rule (Some f) ps p |\<in>| rules A \<and> (p = q \<or> (p, q) |\<in>| (eps A)|\<^sup>+|))" (is "?Ls \<longleftrightarrow> ?Rs")
proof -
{assume ass: "q |\<in>| Qs A" then obtain f r where
"r_rhs r = q \<or> (r_rhs r, q) |\<in>| (eps A)|\<^sup>+|" "r |\<in>| rules A" "r_root r = Some f"
- by (force simp: Qs_def Image_def image_def Let_def fImage.rep_eq simp flip: fmember_iff_member_fset)
+ by (force simp: Qs_def Image_def image_def Let_def fImage.rep_eq)
then have "?Ls \<Longrightarrow> ?Rs" by (cases r) auto}
- moreover have "?Rs \<Longrightarrow> ?Ls" by (force simp: Qs_def Image_def image_def Let_def fImage.rep_eq fmember_iff_member_fset)
+ moreover have "?Rs \<Longrightarrow> ?Ls" by (force simp: Qs_def Image_def image_def Let_def fImage.rep_eq)
ultimately show ?thesis by blast
qed
lemma collapse_Qn_Qs_set_conv:
"fset (Qn A) = {q' |qs q q'. TA_rule None qs q |\<in>| rules A \<and> (q = q' \<or> (q, q') |\<in>| (eps A)|\<^sup>+|)}" (is "?Ls1 = ?Rs1")
"fset (Qs A) = {q' |f qs q q'. TA_rule (Some f) qs q |\<in>| rules A \<and> (q = q' \<or> (q, q') |\<in>| (eps A)|\<^sup>+|)}" (is "?Ls2 = ?Rs2")
- by (auto simp flip: fmember_iff_member_fset) force+
+ by auto force+
definition collapse_automaton :: "('q, 'f option) ta \<Rightarrow> ('q, 'f) ta" where
"collapse_automaton A = TA (collapse_rule_fset A (Qn A) (Qs A)) (eps A)"
definition collapse_automaton_reg where
"collapse_automaton_reg R = Reg (fin R) (collapse_automaton (ta R))"
lemma ta_states_collapse_automaton:
"\<Q> (collapse_automaton A) |\<subseteq>| \<Q> A"
apply (intro \<Q>_subseteq_I)
- apply (auto simp: collapse_automaton_def fmember_iff_member_fset collapse_Qn_Qs_set_conv collapse_rule_set_conv
- fset_of_list.rep_eq in_set_conv_nth rule_statesD[unfolded fmember_iff_member_fset] eps_statesD[unfolded fmember_iff_member_fset])
- apply (metis Option.is_none_def fnth_mem notin_fset option.sel rule_statesD(3) ta_rule.sel(2))
+ apply (auto simp: collapse_automaton_def collapse_Qn_Qs_set_conv collapse_rule_set_conv
+ fset_of_list.rep_eq in_set_conv_nth rule_statesD eps_statesD[unfolded])
+ apply (metis Option.is_none_def fnth_mem option.sel rule_statesD(3) ta_rule.sel(2))
done
lemma last_nthI:
assumes "i < length ts" "\<not> i < length ts - Suc 0"
shows "ts ! i = last ts" using assms
by (induct ts arbitrary: i)
(auto, metis last_conv_nth length_0_conv less_antisym nth_Cons')
lemma collapse_automaton':
assumes "\<Q> A |\<subseteq>| ta_reachable A" (* cf. ta_trim *)
shows "gta_lang Q (collapse_automaton A) = the ` (gcollapse ` gta_lang Q A - {None})"
proof -
define Qn' where "Qn' = Qn A"
define Qs' where "Qs' = Qs A"
{fix t assume t: "t \<in> gta_lang Q (collapse_automaton A)"
then obtain q where q: "q |\<in>| Q" "q |\<in>| ta_der (collapse_automaton A) (term_of_gterm t)" by auto
have "\<exists> t'. q |\<in>| ta_der A (term_of_gterm t') \<and> gcollapse t' \<noteq> None \<and> t = the (gcollapse t')" using q(2)
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
from GFun(1 - 3) obtain qs rs where ps: "TA_rule (Some f) qs p |\<in>| rules A" "length qs = length rs"
"\<And>i. i < length qs \<Longrightarrow> qs ! i |\<in>| Qn' \<and> rs ! i = None \<or> qs ! i |\<in>| Qs' \<and> rs ! i = Some (qs ! i)"
"ps = map the (filter (\<lambda>q. \<not> Option.is_none q) rs)"
by (auto simp: collapse_automaton_def Qn'_def Qs'_def)
obtain ts' where ts':
"ps ! i |\<in>| ta_der A (term_of_gterm (ts' i))" "gcollapse (ts' i) \<noteq> None" "ts ! i = the (gcollapse (ts' i))"
if "i < length ts" for i using GFun(5) by metis
from ps(2, 3, 4) have rs: "i < length qs \<Longrightarrow> rs ! i = None \<or> rs ! i = Some (qs ! i)" for i
by auto
{fix i assume "i < length qs" "rs ! i = None"
then have "\<exists> t'. groot_sym t' = None \<and> qs ! i |\<in>| ta_der A (term_of_gterm t')"
using ps(1, 2) ps(3)[of i]
by (auto simp: ta_der_trancl_eps Qn'_def groot_sym_groot_conv elim!: ta_reachable_rule_gtermE[OF assms])
(force dest: ta_der_trancl_eps)+}
note None = this
{fix i assume i: "i < length qs" "rs ! i = Some (qs ! i)"
have "map Some ps = filter (\<lambda>q. \<not> Option.is_none q) rs" using ps(4)
by (induct rs arbitrary: ps) (auto simp add: Option.is_none_def)
from filter_rev_nth_idx[OF _ _ this, of i]
have *: "rs ! i = map Some ps ! filter_rev_nth (\<lambda>q. \<not> Option.is_none q) rs i"
"filter_rev_nth (\<lambda>q. \<not> Option.is_none q) rs i < length ps"
using ps(2, 4) i by auto
then have "the (rs ! i) = ps ! filter_rev_nth (\<lambda>q. \<not> Option.is_none q) rs i"
"filter_rev_nth (\<lambda>q. \<not> Option.is_none q) rs i < length ps"
by auto} note Some = this
let ?P = "\<lambda> t i. qs ! i |\<in>| ta_der A (term_of_gterm t) \<and>
(rs ! i = None \<longrightarrow> groot_sym t = None) \<and>
(rs ! i = Some (qs ! i) \<longrightarrow> t = ts' (filter_rev_nth (\<lambda>q. \<not> Option.is_none q) rs i))"
{fix i assume i: "i < length qs"
then have "\<exists> t. ?P t i" using Some[OF i] None[OF i] ts' ps(2, 4) GFun(2) rs[OF i]
by (cases "rs ! i") auto}
then obtain ts'' where ts'': "length ts'' = length qs"
"i < length qs \<Longrightarrow> qs ! i |\<in>| ta_der A (term_of_gterm (ts'' ! i))"
"i < length qs \<Longrightarrow> rs ! i = None \<Longrightarrow> groot_sym (ts'' ! i) = None"
"i < length qs \<Longrightarrow> rs ! i = Some (qs ! i) \<Longrightarrow> ts'' ! i = ts' (filter_rev_nth (\<lambda>q. \<not> Option.is_none q) rs i)"
for i using that Ex_list_of_length_P[of "length qs" ?P] by auto
from ts''(1, 3, 4) Some ps(2, 4) GFun(2) rs ts'(2-)
have "map Some ts = (filter (\<lambda>q. \<not> Option.is_none q) (map gcollapse ts''))"
proof (induct ts'' arbitrary: qs rs ps ts rule: rev_induct)
case (snoc a us)
from snoc(2, 7) obtain r rs' where [simp]: "rs = rs' @ [r]"
by (metis append_butlast_last_id append_is_Nil_conv length_0_conv not_Cons_self2)
have l: "length us = length (butlast qs)" "length (butlast qs) = length (butlast rs)"
using snoc(2, 7) by auto
have *: "i < length (butlast qs) \<Longrightarrow> butlast rs ! i = None \<Longrightarrow> groot_sym (us ! i) = None" for i
using snoc(3)[of i] snoc(2, 7)
by (auto simp:nth_append l(1) nth_butlast split!: if_splits)
have **: "i < length (butlast qs) \<Longrightarrow> butlast rs ! i = None \<or> butlast rs ! i = Some (butlast qs ! i)" for i
using snoc(10)[of i] snoc(2, 7) l by (auto simp: nth_append nth_butlast)
have "i < length (butlast qs) \<Longrightarrow> butlast rs ! i = Some (butlast qs ! i) \<Longrightarrow>
us ! i = ts' (filter_rev_nth (\<lambda>q. \<not> Option.is_none q) (butlast rs) i)" for i
using snoc(4)[of i] snoc(2, 7) l
by (auto simp: nth_append nth_butlast filter_rev_nth_def take_butlast)
note IH = snoc(1)[OF l(1) * this _ _ l(2) _ _ **]
show ?case
proof (cases "r = None")
case True
then have "map Some ts = filter (\<lambda>q. \<not> Option.is_none q) (map gcollapse us)"
using snoc(2, 5-)
by (intro IH[of ps ts]) (auto simp: nth_append nth_butlast filter_rev_nth_butlast)
then show ?thesis using True snoc(2, 7) snoc(3)[of "length (butlast rs)"]
by (auto simp: nth_append l(1) last_nthI split!: if_splits)
next
case False
then obtain t' ss where *: "ts = ss @ [t']" using snoc(2, 7, 8, 9)
by (cases ts) (auto, metis append_butlast_last_id list.distinct(1))
let ?i = "filter_rev_nth (\<lambda>q. \<not> Option.is_none q) rs (length us)"
have "map Some (butlast ts) = filter (\<lambda>q. \<not> Option.is_none q) (map gcollapse us)"
using False snoc(2, 5-) l filter_rev_nth_idx
by (intro IH[of "butlast ps" "butlast ts"])
(fastforce simp: nth_butlast filter_rev_nth_butlast)+
moreover have "a = ts' ?i" "?i < length ps"
using False snoc(2, 9) l snoc(4, 6, 10)[of "length us"]
by (auto simp: nth_append)
moreover have "filter_rev_nth (\<lambda>q. \<not> Option.is_none q) (rs' @ [r]) (length rs') = length ss"
using l snoc(2, 7, 8, 9) False unfolding *
by (auto simp: filter_rev_nth_def)
ultimately show ?thesis using l snoc(2, 7, 9) snoc(11-)[of ?i]
by (auto simp: nth_append *)
qed
qed simp
then have "ts = map the (filter (\<lambda>t. \<not> Option.is_none t) (map gcollapse ts''))"
by (metis comp_the_Some list.map_id map_map)
then show ?case using ps(1, 2) ts''(1, 2) GFun(3)
by (auto simp: collapse_automaton_def intro!: exI[of _ "GFun (Some f) ts''"] exI[of _ qs] exI[of _ p])
qed
then have "t \<in> the ` (gcollapse ` gta_lang Q A - {None})"
by (meson Diff_iff gta_langI imageI q(1) singletonD)
} moreover
{ fix t assume t: "t \<in> gta_lang Q A" "gcollapse t \<noteq> None"
obtain q where q: "q |\<in>| Q" "q |\<in>| ta_der A (term_of_gterm t)" using t(1) by auto
have "q |\<in>| ta_der (collapse_automaton A) (term_of_gterm (the (gcollapse t)))" using q(2) t(2)
proof (induct t arbitrary: q)
case (GFun f ts)
obtain qs q' where q: "TA_rule f qs q' |\<in>| rules A" "q' = q \<or> (q', q) |\<in>| (eps (collapse_automaton A))|\<^sup>+|"
"length qs = length ts" "\<And>i. i < length ts \<Longrightarrow> qs ! i |\<in>| ta_der A (term_of_gterm (ts ! i))"
using GFun(2) by (auto simp: collapse_automaton_def)
obtain f' where f [simp]: "f = Some f'" using GFun(3) by (cases f) auto
define qs' where
"qs' = map (\<lambda>i. if Option.is_none (gcollapse (ts ! i)) then None else Some (qs ! i)) [0..<length qs]"
have "Option.is_none (gcollapse (ts ! i)) \<Longrightarrow> qs ! i |\<in>| Qn'" if "i < length qs" for i
using q(4)[of i] that
by (cases "ts ! i" rule: gcollapse.cases)
- (auto simp: q(3) Qn'_def fmember_iff_member_fset collapse_Qn_Qs_set_conv, meson notin_fset ta_der_Fun)
+ (auto simp: q(3) Qn'_def collapse_Qn_Qs_set_conv)
moreover have "\<not> Option.is_none (gcollapse (ts ! i)) \<Longrightarrow> qs ! i |\<in>| Qs'" if "i < length qs" for i
using q(4)[of i] that
by (cases "ts ! i" rule: gcollapse.cases)
- (auto simp: q(3) Qs'_def fmember_iff_member_fset collapse_Qn_Qs_set_conv, meson notin_fset ta_der_Fun)
+ (auto simp: q(3) Qs'_def collapse_Qn_Qs_set_conv)
ultimately have "f' (map the (filter (\<lambda>q. \<not> Option.is_none q) qs')) \<rightarrow> q' |\<in>| rules (collapse_automaton A)"
using q(1, 4) unfolding collapse_automaton_def Qn'_def[symmetric] Qs'_def[symmetric]
by (fastforce simp: qs'_def q(3) intro: exI[of _ qs] exI[of _ qs'] split: if_splits)
moreover have ***: "length (filter (\<lambda>i.\<not> Option.is_none (gcollapse (ts ! i))) [0..<length ts]) =
length (filter (\<lambda>t. \<not> Option.is_none (gcollapse t)) ts)" for ts
by (subst length_map[of "(!) ts", symmetric] filter_map[unfolded comp_def, symmetric] map_nth)+
(rule refl)
note this[of ts]
moreover have "the (filter (\<lambda>q. \<not> Option.is_none q) qs' ! i) |\<in>| ta_der (collapse_automaton A)
(term_of_gterm (the (filter (\<lambda>t. \<not> Option.is_none t) (map gcollapse ts) ! i)))"
if "i < length [x\<leftarrow>ts . \<not> Option.is_none (gcollapse x)]" for i
unfolding qs'_def using that q(3) GFun(1)[OF nth_mem q(4)]
proof (induct ts arbitrary: qs rule: List.rev_induct)
case (snoc t ts)
have x1 [simp]: "i < length xs \<Longrightarrow> (xs @ ys) ! i = xs ! i" for xs ys i by (auto simp: nth_append)
have x2: "i = length xs \<Longrightarrow> (xs @ ys) ! i = ys ! 0" for xs ys i by (auto simp: nth_append)
obtain q qs' where qs [simp]: "qs = qs' @ [q]" using snoc(3) by (cases "rev qs") auto
have [simp]:
"map (\<lambda>x. if Option.is_none (gcollapse ((ts @ [t]) ! x)) then None else Some ((qs' @ [q]) ! x)) [0..<length ts] =
map (\<lambda>x. if Option.is_none (gcollapse (ts ! x)) then None else Some (qs' ! x)) [0..<length ts]"
using snoc(3) by auto
show ?case
proof (cases "Option.is_none (gcollapse t)")
case True then show ?thesis using snoc(1)[of qs'] snoc(2,3)
snoc(4)[unfolded length_append list.size add_0 add_0_right add_Suc_right, OF less_SucI]
by (auto cong: if_cong)
next
case False note False' = this
show ?thesis
proof (cases "i = length [x\<leftarrow>ts . \<not> Option.is_none (gcollapse x)]")
case True
then show ?thesis using snoc(3) snoc(4)[of "length ts"]
unfolding qs length_append list.size add_0 add_0_right add_Suc_right
upt_Suc_append[OF zero_le] filter_append map_append
by (subst (5 6) x2) (auto simp: comp_def *** False' Option.is_none_def[symmetric])
next
case False
then show ?thesis using snoc(1)[of qs'] snoc(2,3)
snoc(4)[unfolded length_append list.size add_0 add_0_right add_Suc_right, OF less_SucI]
unfolding qs length_append list.size add_0 add_0_right add_Suc_right
upt_Suc_append[OF zero_le] filter_append map_append
by (subst (5 6) x1) (auto simp: comp_def *** False')
qed
qed
qed auto
ultimately show ?case using q(2) by (auto simp: qs'_def comp_def q(3)
intro!: exI[of _ q'] exI[of _ "map the (filter (\<lambda>q. \<not> Option.is_none q) qs')"])
qed
then have "the (gcollapse t) \<in> gta_lang Q (collapse_automaton A)"
by (metis q(1) gta_langI)
} ultimately show ?thesis by blast
qed
lemma \<L>_collapse_automaton':
assumes "\<Q>\<^sub>r A |\<subseteq>| ta_reachable (ta A)" (* cf. ta_trim *)
shows "\<L> (collapse_automaton_reg A) = the ` (gcollapse ` \<L> A - {None})"
using assms by (auto simp: collapse_automaton_reg_def \<L>_def collapse_automaton')
lemma collapse_automaton:
assumes "\<Q>\<^sub>r A |\<subseteq>| ta_reachable (ta A)" "RR1_spec A T"
shows "RR1_spec (collapse_automaton_reg A) (the ` (gcollapse ` \<L> A - {None}))"
using collapse_automaton'[OF assms(1)] assms(2)
by (simp add: collapse_automaton_reg_def \<L>_def RR1_spec_def)
subsection \<open>Cylindrification\<close>
(* cylindrification is a product ("pairing") of a RR1 language accepting all terms, and an RRn language,
modulo some fairly trivial renaming of symbols. *)
definition pad_with_Nones where
"pad_with_Nones n m = (\<lambda>(f, g). case_option (replicate n None) id f @ case_option (replicate m None) id g)"
lemma gencode_append:
"gencode (ss @ ts) = map_gterm (pad_with_Nones (length ss) (length ts)) (gpair (gencode ss) (gencode ts))"
proof -
have [simp]: "p \<notin> gposs (gunions (map gdomain ts)) \<Longrightarrow> map (\<lambda>t. gfun_at t p) ts = replicate (length ts) None"
for p ts by (intro nth_equalityI)
(fastforce simp: poss_gposs_mem_conv fun_at_def' image_def all_set_conv_all_nth)+
note [simp] = glabel_map_gterm_conv[of "\<lambda>(_ :: unit option). ()", unfolded comp_def gdomain_id]
show ?thesis by (auto intro!: arg_cong[of _ _ "\<lambda>x. glabel x _"] simp del: gposs_gunions
simp: pad_with_Nones_def gencode_def gunions_append gpair_def map_gterm_glabel comp_def)
qed
lemma append_automaton:
assumes "RRn_spec n A T" "RRn_spec m B U"
shows "RRn_spec (n + m) (fmap_funs_reg (pad_with_Nones n m) (pair_automaton_reg A B)) {ts @ us |ts us. ts \<in> T \<and> us \<in> U}"
using assms pair_automaton[of A "gencode ` T" B "gencode ` U"]
unfolding RRn_spec_def
proof (intro conjI set_eqI iffI, goal_cases)
case (1 s)
then obtain ts us where "ts \<in> T" "us \<in> U" "s = gencode (ts @ us)"
by (fastforce simp: \<L>_def fmap_funs_reg_def RR1_spec_def RR2_spec_def gencode_append fmap_funs_gta_lang)
then show ?case by blast
qed (fastforce simp: RR1_spec_def RR2_spec_def fmap_funs_reg_def \<L>_def gencode_append fmap_funs_gta_lang)+
lemma cons_automaton:
assumes "RR1_spec A T" "RRn_spec m B U"
shows "RRn_spec (Suc m) (fmap_funs_reg (\<lambda>(f, g). pad_with_Nones 1 m (map_option (\<lambda>f. [Some f]) f, g))
(pair_automaton_reg A B)) {t # us |t us. t \<in> T \<and> us \<in> U}"
proof -
have [simp]: "{ts @ us |ts us. ts \<in> (\<lambda>t. [t]) ` T \<and> us \<in> U} = {t # us |t us. t \<in> T \<and> us \<in> U}"
by (auto intro: exI[of _ "[_]", OF exI])
show ?thesis using append_automaton[OF RR1_to_RRn_spec, OF assms]
by (auto simp: \<L>_def fmap_funs_reg_def pair_automaton_reg_def comp_def
fmap_funs_gta_lang map_pair_automaton_12 fmap_funs_ta_comp prod.case_distrib
gencode_append[of "[_]", unfolded gencode_singleton List.append.simps])
qed
subsection \<open>Projection\<close>
(* projection is composed from fmap_funs_ta and collapse_automaton, corresponding to gsnd *)
abbreviation "drop_none_rule m fs \<equiv> if list_all (Option.is_none) (drop m fs) then None else Some (drop m fs)"
lemma drop_automaton_reg:
assumes "\<Q>\<^sub>r A |\<subseteq>| ta_reachable (ta A)" "m < n" "RRn_spec n A T"
defines "f \<equiv> \<lambda>fs. drop_none_rule m fs"
shows "RRn_spec (n - m) (collapse_automaton_reg (fmap_funs_reg f A)) (drop m ` T)"
proof -
have [simp]: "length ts = n - m ==> p \<in> gposs (gencode ts) \<Longrightarrow> \<exists>f. \<exists>t\<in>set ts. Some f = gfun_at t p" for p ts
proof (cases p, goal_cases Empty PCons)
case Empty
obtain t where "t \<in> set ts" using assms(2) Empty(1) by (cases ts) auto
moreover then obtain f where "Some f = gfun_at t p" using Empty(3) by (cases t rule: gterm.exhaust) auto
ultimately show ?thesis by auto
next
case (PCons i p')
then have "p \<noteq> []" by auto
then show ?thesis using PCons(2)
by (auto 0 3 simp: gencode_def eq_commute[of "gfun_at _ _" "Some _"] elim!: gfun_at_possE)
qed
{ fix p ts y assume that: "gfun_at (gencode ts) p = Some y"
have "p \<in> gposs (gencode ts) \<Longrightarrow> gfun_at (gencode ts) p = Some (map (\<lambda>t. gfun_at t p) ts)"
by (auto simp: gencode_def)
moreover have "gfun_at (gencode ts) p = Some y \<Longrightarrow> p \<in> gposs (gencode ts)"
using gfun_at_nongposs by force
ultimately have "y = map (\<lambda>t. gfun_at t p) ts \<and> p \<in> gposs (gencode ts)" by (simp add: that)
} note [dest!] = this
have [simp]: "list_all f (replicate n x) \<longleftrightarrow> n = 0 \<or> f x" for f n x by (induct n) auto
have [dest]: "x \<in> set xs \<Longrightarrow> list_all f xs \<Longrightarrow> f x" for f x xs by (induct xs) auto
have *: "f (pad_with_Nones m' n' z) = snd z"
if "fst z = None \<or> fst z \<noteq> None \<and> length (the (fst z)) = m"
"snd z = None \<or> snd z \<noteq> None \<and> length (the (snd z)) = n - m \<and> (\<exists>y. Some y \<in> set (the (snd z)))"
"m' = m" "n' = n - m" for z m' n'
using that by (auto simp: f_def pad_with_Nones_def split: option.splits prod.splits)
{ fix ts assume ts: "ts \<in> T" "length ts = n"
have "gencode (drop m ts) = the (gcollapse (map_gterm f (gencode ts)))"
"gcollapse (map_gterm f (gencode ts)) \<noteq> None"
proof (goal_cases)
case 1 show ?case
using ts assms(2)
apply (subst gsnd_gpair[of "gencode (take m ts)", symmetric])
apply (subst gencode_append[of "take m ts" "drop m ts", unfolded append_take_drop_id])
unfolding gsnd_def comp_def gterm.map_comp
apply (intro arg_cong[where f = "\<lambda>x. the (gcollapse x)"] gterm.map_cong refl)
by (subst *) (auto simp: gpair_def set_gterm_gposs_conv image_def)
next
case 2
have [simp]: "gcollapse t = None \<longleftrightarrow> gfun_at t [] = Some None" for t
by (cases t rule: gcollapse.cases) auto
obtain t where "t \<in> set (drop m ts)" using ts(2) assms(2) by (cases "drop m ts") auto
moreover then have "\<not> Option.is_none (gfun_at t [])" by (cases t rule: gterm.exhaust) auto
ultimately show ?case
by (auto simp: f_def gencode_def list_all_def drop_map)
qed
}
then show ?thesis using assms(3)
by (fastforce simp: \<L>_def collapse_automaton_reg_def fmap_funs_reg_def
RRn_spec_def fmap_funs_gta_lang gsnd_def gpair_def gterm.map_comp comp_def
glabel_map_gterm_conv[unfolded comp_def] assms(1) collapse_automaton')
qed
lemma gfst_collapse_simp:
"the (gcollapse (map_gterm fst t)) = gfst t"
by (simp add: gfst_def)
lemma gsnd_collapse_simp:
"the (gcollapse (map_gterm snd t)) = gsnd t"
by (simp add: gsnd_def)
definition proj_1_reg where
"proj_1_reg A = collapse_automaton_reg (fmap_funs_reg fst (trim_reg A))"
definition proj_2_reg where
"proj_2_reg A = collapse_automaton_reg (fmap_funs_reg snd (trim_reg A))"
lemmas proj_1_reg_simp = proj_1_reg_def collapse_automaton_reg_def fmap_funs_reg_def trim_reg_def
lemmas proj_2_reg_simp = proj_2_reg_def collapse_automaton_reg_def fmap_funs_reg_def trim_reg_def
lemma \<L>_proj_1_reg_collapse:
"\<L> (proj_1_reg \<A>) = the ` (gcollapse ` map_gterm fst ` (\<L> \<A>) - {None})"
proof -
have "\<Q>\<^sub>r (fmap_funs_reg fst (trim_reg \<A>)) |\<subseteq>| ta_reachable (ta (fmap_funs_reg fst (trim_reg \<A>)))"
by (auto simp: fmap_funs_reg_def)
note [simp] = \<L>_collapse_automaton'[OF this]
show ?thesis by (auto simp: proj_1_reg_def fmap_funs_\<L> \<L>_trim)
qed
lemma \<L>_proj_2_reg_collapse:
"\<L> (proj_2_reg \<A>) = the ` (gcollapse ` map_gterm snd ` (\<L> \<A>) - {None})"
proof -
have "\<Q>\<^sub>r (fmap_funs_reg snd (trim_reg \<A>)) |\<subseteq>| ta_reachable (ta (fmap_funs_reg snd (trim_reg \<A>)))"
by (auto simp: fmap_funs_reg_def)
note [simp] = \<L>_collapse_automaton'[OF this]
show ?thesis by (auto simp: proj_2_reg_def fmap_funs_\<L> \<L>_trim)
qed
lemma proj_1:
assumes "RR2_spec A R"
shows "RR1_spec (proj_1_reg A) (fst ` R)"
proof -
{fix s t assume ass: "(s, t) \<in> R"
from ass have s: "s = the (gcollapse (map_gterm fst (gpair s t)))"
by (auto simp: gfst_gpair gfst_collapse_simp)
then have "Some s = gcollapse (map_gterm fst (gpair s t))"
by (cases s; cases t) (auto simp: gpair_def)
then have "s \<in> \<L> (proj_1_reg A)" using assms ass s
by (auto simp: proj_1_reg_simp \<L>_def trim_ta_reach trim_gta_lang
image_def image_Collect RR2_spec_def fmap_funs_gta_lang
collapse_automaton'[of "fmap_funs_ta fst (trim_ta (fin A) (ta A))"])
force}
moreover
{fix s assume "s \<in> \<L> (proj_1_reg A)" then have "s \<in> fst ` R" using assms
by (auto simp: gfst_collapse_simp gfst_gpair rev_image_eqI RR2_spec_def trim_ta_reach trim_gta_lang
\<L>_def proj_1_reg_simp fmap_funs_gta_lang collapse_automaton'[of "fmap_funs_ta fst (trim_ta (fin A) (ta A))"])}
ultimately show ?thesis using assms unfolding RR2_spec_def RR1_spec_def \<L>_def proj_1_reg_simp
by auto
qed
lemma proj_2:
assumes "RR2_spec A R"
shows "RR1_spec (proj_2_reg A) (snd ` R)"
proof -
{fix s t assume ass: "(s, t) \<in> R"
then have s: "t = the (gcollapse (map_gterm snd (gpair s t)))"
by (auto simp: gsnd_gpair gsnd_collapse_simp)
then have "Some t = gcollapse (map_gterm snd (gpair s t))"
by (cases s; cases t) (auto simp: gpair_def)
then have "t \<in> \<L> (proj_2_reg A)" using assms ass s
by (auto simp: \<L>_def trim_ta_reach trim_gta_lang proj_2_reg_simp
image_def image_Collect RR2_spec_def fmap_funs_gta_lang
collapse_automaton'[of "fmap_funs_ta snd (trim_ta (fin A) (ta A))"])
fastforce}
moreover
{fix s assume "s \<in> \<L> (proj_2_reg A)" then have "s \<in> snd ` R" using assms
by (auto simp: \<L>_def gsnd_collapse_simp gsnd_gpair rev_image_eqI RR2_spec_def
trim_ta_reach trim_gta_lang proj_2_reg_simp
fmap_funs_gta_lang collapse_automaton'[of "fmap_funs_ta snd (trim_ta (fin A) (ta A))"])}
ultimately show ?thesis using assms unfolding RR2_spec_def RR1_spec_def
by auto
qed
lemma \<L>_proj:
assumes "RR2_spec A R"
shows "\<L> (proj_1_reg A) = gfst ` \<L> A" "\<L> (proj_2_reg A) = gsnd ` \<L> A"
proof -
have [simp]: "gfst ` {gpair t u |t u. (t, u) \<in> R} = fst ` R"
by (force simp: gfst_gpair image_def)
have [simp]: "gsnd ` {gpair t u |t u. (t, u) \<in> R} = snd ` R"
by (force simp: gsnd_gpair image_def)
show "\<L> (proj_1_reg A) = gfst ` \<L> A" "\<L> (proj_2_reg A) = gsnd ` \<L> A"
using proj_1[OF assms] proj_2[OF assms] assms gfst_gpair gsnd_gpair
by (auto simp: RR1_spec_def RR2_spec_def)
qed
lemmas proj_automaton_gta_lang = proj_1 proj_2
subsection \<open>Permutation\<close>
(* permutations are a direct application of fmap_funs_ta *)
lemma gencode_permute:
assumes "set ps = {0..<length ts}"
shows "gencode (map ((!) ts) ps) = map_gterm (\<lambda>xs. map ((!) xs) ps) (gencode ts)"
proof -
have *: "(!) ts ` set ps = set ts" using assms by (auto simp: image_def set_conv_nth)
show ?thesis using subsetD[OF equalityD1[OF assms]]
apply (intro eq_gterm_by_gposs_gfun_at)
unfolding gencode_def gposs_glabel gposs_map_gterm gposs_gunions gfun_at_map_gterm gfun_at_glabel
set_map * by auto
qed
lemma permute_automaton:
assumes "RRn_spec n A T" "set ps = {0..<n}"
shows "RRn_spec (length ps) (fmap_funs_reg (\<lambda>xs. map ((!) xs) ps) A) ((\<lambda>xs. map ((!) xs) ps) ` T)"
using assms by (auto simp: RRn_spec_def gencode_permute fmap_funs_reg_def \<L>_def fmap_funs_gta_lang image_def)
subsection \<open>Intersection\<close>
(* intersection is already defined in IsaFoR *)
lemma intersect_automaton:
assumes "RRn_spec n A T" "RRn_spec n B U"
shows "RRn_spec n (reg_intersect A B) (T \<inter> U)" using assms
by (simp add: RRn_spec_def \<L>_intersect)
(metis gdecode_gencode image_Int inj_on_def)
(*
lemma swap_union_automaton:
"fmap_states_ta (case_sum Inr Inl) (union_automaton B A) = union_automaton A B"
by (simp add: fmap_states_ta_def' union_automaton_def image_Un image_comp case_sum_o_inj
ta_rule.map_comp prod.map_comp comp_def id_def ac_simps)
*)
lemma union_automaton:
assumes "RRn_spec n A T" "RRn_spec n B U"
shows "RRn_spec n (reg_union A B) (T \<union> U)"
using assms by (auto simp: RRn_spec_def \<L>_union)
subsection \<open>Difference\<close>
lemma RR1_difference:
assumes "RR1_spec A T" "RR1_spec B U"
shows "RR1_spec (difference_reg A B) (T - U)"
using assms by (auto simp: RR1_spec_def \<L>_difference_reg)
lemma RR2_difference:
assumes "RR2_spec A T" "RR2_spec B U"
shows "RR2_spec (difference_reg A B) (T - U)"
using assms by (auto simp: RR2_spec_def \<L>_difference_reg)
lemma RRn_difference:
assumes "RRn_spec n A T" "RRn_spec n B U"
shows "RRn_spec n (difference_reg A B) (T - U)"
using assms by (auto simp: RRn_spec_def \<L>_difference_reg) (metis gdecode_gencode)+
subsection \<open>All terms over a signature\<close>
definition term_automaton :: "('f \<times> nat) fset \<Rightarrow> (unit, 'f) ta" where
"term_automaton \<F> = TA ((\<lambda> (f, n). TA_rule f (replicate n ()) ()) |`| \<F>) {||}"
definition term_reg where
"term_reg \<F> = Reg {|()|} (term_automaton \<F>)"
lemma term_automaton:
"RR1_spec (term_reg \<F>) (\<T>\<^sub>G (fset \<F>))"
unfolding RR1_spec_def gta_lang_def term_reg_def \<L>_def
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr t)
then have "() |\<in>| ta_der (term_automaton \<F>) (term_of_gterm t)"
by (auto simp: gta_der_def)
then show ?case
- by (induct t) (auto simp: term_automaton_def split: if_splits simp flip: fmember_iff_member_fset)
+ by (induct t) (auto simp: term_automaton_def split: if_splits)
next
case (rl t)
then have "() |\<in>| ta_der (term_automaton \<F>) (term_of_gterm t)"
proof (induct t rule: \<T>\<^sub>G.induct)
case (const a) then show ?case
- by (auto simp: term_automaton_def fimage_iff simp flip: fmember_iff_member_fset intro: fBexI[of _ "(a, 0)"])
+ by (auto simp: term_automaton_def image_iff intro: bexI[of _ "(a, 0)"])
next
case (ind f n ss) then show ?case
- by (auto simp: term_automaton_def fimage_iff simp flip: fmember_iff_member_fset intro: fBexI[of _ "(f, n)"])
+ by (auto simp: term_automaton_def image_iff intro: bexI[of _ "(f, n)"])
qed
then show ?case
by (auto simp: gta_der_def)
qed
fun true_RRn :: "('f \<times> nat) fset \<Rightarrow> nat \<Rightarrow> (nat, 'f option list) reg" where
"true_RRn \<F> 0 = Reg {|0|} (TA {|TA_rule [] [] 0|} {||})"
| "true_RRn \<F> (Suc 0) = relabel_reg (fmap_funs_reg (\<lambda>f. [Some f]) (term_reg \<F>))"
| "true_RRn \<F> (Suc n) = relabel_reg
(trim_reg (fmap_funs_reg (pad_with_Nones 1 n) (pair_automaton_reg (true_RRn \<F> 1) (true_RRn \<F> n))))"
lemma true_RRn_spec:
"RRn_spec n (true_RRn \<F> n) {ts. length ts = n \<and> set ts \<subseteq> \<T>\<^sub>G (fset \<F>)}"
proof (induct \<F> n rule: true_RRn.induct)
case (1 \<F>) show ?case
by (simp cong: conj_cong add: true_RR0_spec)
next
case (2 \<F>)
moreover have "{ts. length ts = 1 \<and> set ts \<subseteq> \<T>\<^sub>G (fset \<F>)} = (\<lambda>t. [t]) ` \<T>\<^sub>G (fset \<F>)"
apply (intro equalityI subsetI)
subgoal for ts by (cases ts) auto
by auto
ultimately show ?case
using RR1_to_RRn_spec[OF term_automaton, of \<F>] by auto
next
case (3 \<F> n)
have [simp]: "{ts @ us |ts us. length ts = n \<and> set ts \<subseteq> \<T>\<^sub>G (fset \<F>) \<and> length us = m \<and>
set us \<subseteq> \<T>\<^sub>G (fset \<F>)} = {ss. length ss = n + m \<and> set ss \<subseteq> \<T>\<^sub>G (fset \<F>)}" for n m
by (auto 0 4 intro!: exI[of _ "take n _", OF exI[of _ "drop n _"], of _ xs xs for xs]
dest!: subsetD[OF set_take_subset] subsetD[OF set_drop_subset])
show ?case using append_automaton[OF 3]
by simp
qed
subsection \<open>RR2 composition\<close>
abbreviation "RR2_to_RRn A \<equiv> fmap_funs_reg (\<lambda>(f, g). [f, g]) A"
abbreviation "RRn_to_RR2 A \<equiv> fmap_funs_reg (\<lambda>f. (f ! 0, f ! 1)) A"
definition rr2_compositon where
"rr2_compositon \<F> A B =
(let A' = RR2_to_RRn A in
let B' = RR2_to_RRn B in
let F = true_RRn \<F> 1 in
let CA = trim_reg (fmap_funs_reg (pad_with_Nones 2 1) (pair_automaton_reg A' F)) in
let CB = trim_reg (fmap_funs_reg (pad_with_Nones 1 2) (pair_automaton_reg F B')) in
let PI = trim_reg (fmap_funs_reg (\<lambda>xs. map ((!) xs) [1, 0, 2]) (reg_intersect CA CB)) in
RRn_to_RR2 (collapse_automaton_reg (fmap_funs_reg (drop_none_rule 1) PI))
)"
lemma list_length1E:
assumes "length xs = Suc 0" obtains x where "xs = [x]" using assms
by (cases xs) auto
lemma rr2_compositon:
assumes "\<R> \<subseteq> \<T>\<^sub>G (fset \<F>) \<times> \<T>\<^sub>G (fset \<F>)" "\<LL> \<subseteq> \<T>\<^sub>G (fset \<F>) \<times> \<T>\<^sub>G (fset \<F>)"
and "RR2_spec A \<R>" and "RR2_spec B \<LL>"
shows "RR2_spec (rr2_compositon \<F> A B) (\<R> O \<LL>)"
proof -
let ?R = "(\<lambda>(t, u). [t, u]) ` \<R>" let ?L = "(\<lambda>(t, u). [t, u]) ` \<LL>"
let ?A = "RR2_to_RRn A" let ?B = "RR2_to_RRn B" let ?F = "true_RRn \<F> 1"
let ?CA = "trim_reg (fmap_funs_reg (pad_with_Nones 2 1) (pair_automaton_reg ?A ?F))"
let ?CB = "trim_reg (fmap_funs_reg (pad_with_Nones 1 2) (pair_automaton_reg ?F ?B))"
let ?PI = "trim_reg (fmap_funs_reg (\<lambda>xs. map ((!) xs) [1, 0, 2]) (reg_intersect ?CA ?CB))"
let ?DR = "collapse_automaton_reg (fmap_funs_reg (drop_none_rule 1) ?PI)"
let ?Rs = "{ts @ us | ts us. ts \<in> ?R \<and> (\<exists>t. us = [t] \<and> t \<in> \<T>\<^sub>G (fset \<F>))}"
let ?Ls = "{us @ ts | ts us. ts \<in> ?L \<and> (\<exists>t. us = [t] \<and> t \<in> \<T>\<^sub>G (fset \<F>))}"
from RR2_to_RRn_spec assms(3, 4)
have rr2: "RRn_spec 2 ?A ?R" "RRn_spec 2 ?B ?L" by auto
have *: "{ts. length ts = 1 \<and> set ts \<subseteq> \<T>\<^sub>G (fset \<F>)} = {[t] | t. t \<in> \<T>\<^sub>G (fset \<F>)}"
by (auto elim!: list_length1E)
have F: "RRn_spec 1 ?F {[t] | t. t \<in> \<T>\<^sub>G (fset \<F>)}" using true_RRn_spec[of 1 \<F>] unfolding * .
have "RRn_spec 3 ?CA ?Rs" "RRn_spec 3 ?CB ?Ls"
using append_automaton[OF rr2(1) F] append_automaton[OF F rr2(2)]
by (auto simp: numeral_3_eq_3) (smt Collect_cong)
from permute_automaton[OF intersect_automaton[OF this], of "[1, 0, 2]"]
have "RRn_spec 3 ?PI ((\<lambda>xs. map ((!) xs) [1, 0, 2]) ` (?Rs \<inter> ?Ls))"
by (auto simp: atLeast0_lessThan_Suc insert_commute numeral_2_eq_2 numeral_3_eq_3)
from drop_automaton_reg[OF _ _ this, of 1]
have sp: "RRn_spec 2 ?DR (drop 1 ` (\<lambda>xs. map ((!) xs) [1, 0, 2]) ` (?Rs \<inter> ?Ls))"
by auto
{fix s assume "s \<in> (\<lambda>(t, u). [t, u]) ` (\<R> O \<LL>)"
then obtain t u v where comp: "s = [t, u]" "(t, v) \<in> \<R>" "(v, u) \<in> \<LL>"
by (auto simp: image_iff relcomp_unfold split!: prod.split)
then have "[t, v] \<in> ?R" "[v , u] \<in> ?L" "u \<in> \<T>\<^sub>G (fset \<F>)" "v \<in> \<T>\<^sub>G (fset \<F>)" "t \<in> \<T>\<^sub>G (fset \<F>)" using assms(1, 2)
by (auto simp: image_iff relcomp_unfold split!: prod.splits)
then have "[t, v, u] \<in> ?Rs" "[t, v, u] \<in> ?Ls"
apply (simp_all)
subgoal
apply (rule exI[of _ "[t, v]"], rule exI[of _ "[u]"])
apply simp
done
subgoal
apply (rule exI[of _ "[v, u]"], rule exI[of _ "[t]"])
apply simp
done
done
then have "s \<in> drop 1 ` (\<lambda>xs. map ((!) xs) [1, 0, 2]) ` (?Rs \<inter> ?Ls)" unfolding comp(1)
apply (simp add: image_def Bex_def)
apply (rule exI[of _ "[v, t, u]"]) apply simp
apply (rule exI[of _ "[t, v, u]"]) apply simp
done}
moreover have "drop 1 ` (\<lambda>xs. map ((!) xs) [1, 0, 2]) ` (?Rs \<inter> ?Ls) \<subseteq> (\<lambda>(t, u). [t, u]) ` (\<R> O \<LL>)"
by (auto simp: image_iff relcomp_unfold Bex_def split!: prod.splits)
ultimately have *: "drop 1 ` (\<lambda>xs. map ((!) xs) [1, 0, 2]) ` (?Rs \<inter> ?Ls) = (\<lambda>(t, u). [t, u]) ` (\<R> O \<LL>)"
by (simp add: subsetI subset_antisym)
have **: "(\<lambda>f. (f ! 0, f ! 1)) ` (\<lambda>(t, u). [t, u]) ` (\<R> O \<LL>) = \<R> O \<LL>"
by (force simp: image_def relcomp_unfold split!: prod.splits)
show ?thesis using sp unfolding *
using RRn_to_RR2_spec[where ?T = "(\<lambda>(t, u). [t, u]) ` (\<R> O \<LL>)" and ?A = ?DR]
unfolding ** by (auto simp: rr2_compositon_def Let_def image_iff)
qed
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Regular_Relation_Abstract_Impl.thy b/thys/Regular_Tree_Relations/Regular_Relation_Abstract_Impl.thy
--- a/thys/Regular_Tree_Relations/Regular_Relation_Abstract_Impl.thy
+++ b/thys/Regular_Tree_Relations/Regular_Relation_Abstract_Impl.thy
@@ -1,240 +1,240 @@
theory Regular_Relation_Abstract_Impl
imports Pair_Automaton
GTT_Transitive_Closure
RR2_Infinite_Q_infinity
Horn_Fset
begin
abbreviation TA_of_lists where
"TA_of_lists \<Delta> \<Delta>\<^sub>E \<equiv> TA (fset_of_list \<Delta>) (fset_of_list \<Delta>\<^sub>E)"
section \<open>Computing the epsilon transitions for the composition of GTT's\<close>
definition \<Delta>\<^sub>\<epsilon>_rules :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) horn set" where
"\<Delta>\<^sub>\<epsilon>_rules A B =
{zip ps qs \<rightarrow>\<^sub>h (p, q) |f ps p qs q. f ps \<rightarrow> p |\<in>| rules A \<and> f qs \<rightarrow> q |\<in>| rules B \<and> length ps = length qs} \<union>
{[(p, q)] \<rightarrow>\<^sub>h (p', q) |p p' q. (p, p') |\<in>| eps A} \<union>
{[(p, q)] \<rightarrow>\<^sub>h (p, q') |p q q'. (q, q') |\<in>| eps B}"
locale \<Delta>\<^sub>\<epsilon>_horn =
fixes A :: "('q, 'f) ta" and B :: "('q, 'f) ta"
begin
sublocale horn "\<Delta>\<^sub>\<epsilon>_rules A B" .
lemma \<Delta>\<^sub>\<epsilon>_infer0:
"infer0 = {(p, q) |f p q. f [] \<rightarrow> p |\<in>| rules A \<and> f [] \<rightarrow> q |\<in>| rules B}"
unfolding horn.infer0_def \<Delta>\<^sub>\<epsilon>_rules_def
using zip_Nil[of "[]"]
by auto (metis length_greater_0_conv zip_eq_Nil_iff)+
lemma \<Delta>\<^sub>\<epsilon>_infer1:
"infer1 pq X = {(p, q) |f ps p qs q. f ps \<rightarrow> p |\<in>| rules A \<and> f qs \<rightarrow> q |\<in>| rules B \<and> length ps = length qs \<and>
(fst pq, snd pq) \<in> set (zip ps qs) \<and> set (zip ps qs) \<subseteq> insert pq X} \<union>
{(p', snd pq) |p p'. (p, p') |\<in>| eps A \<and> p = fst pq} \<union>
{(fst pq, q') |q q'. (q, q') |\<in>| eps B \<and> q = snd pq}"
unfolding \<Delta>\<^sub>\<epsilon>_rules_def horn_infer1_union
apply (intro arg_cong2[of _ _ _ _ "(\<union>)"])
by (auto simp: horn.infer1_def simp flip: ex_simps(1)) force+
lemma \<Delta>\<^sub>\<epsilon>_sound:
"\<Delta>\<^sub>\<epsilon>_set A B = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using lr unfolding x
proof (induct)
case (\<Delta>\<^sub>\<epsilon>_set_cong f ps p qs q) show ?case
apply (intro infer[of "zip ps qs" "(p, q)"])
subgoal using \<Delta>\<^sub>\<epsilon>_set_cong(1-3) by (auto simp: \<Delta>\<^sub>\<epsilon>_rules_def)
subgoal using \<Delta>\<^sub>\<epsilon>_set_cong(3,5) by (auto simp: zip_nth_conv)
done
next
case (\<Delta>\<^sub>\<epsilon>_set_eps1 p p' q) then show ?case
by (intro infer[of "[(p, q)]" "(p', q)"]) (auto simp: \<Delta>\<^sub>\<epsilon>_rules_def)
next
case (\<Delta>\<^sub>\<epsilon>_set_eps2 q q' p) then show ?case
by (intro infer[of "[(p, q)]" "(p, q')"]) (auto simp: \<Delta>\<^sub>\<epsilon>_rules_def)
qed
next
case (rl x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using rl unfolding x
proof (induct)
case (infer as a) then show ?case
using \<Delta>\<^sub>\<epsilon>_set_cong[of _ "map fst as" "fst a" A "map snd as" "snd a" B]
\<Delta>\<^sub>\<epsilon>_set_eps1[of _ "fst a" A "snd a" B] \<Delta>\<^sub>\<epsilon>_set_eps2[of _ "snd a" B "fst a" A]
by (auto simp: \<Delta>\<^sub>\<epsilon>_rules_def)
qed
qed
end
section \<open>Computing the epsilon transitions for the transitive closure of GTT's\<close>
definition \<Delta>_trancl_rules :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) horn set" where
"\<Delta>_trancl_rules A B =
\<Delta>\<^sub>\<epsilon>_rules A B \<union> {[(p, q), (q, r)] \<rightarrow>\<^sub>h (p, r) |p q r. True}"
locale \<Delta>_trancl_horn =
fixes A :: "('q, 'f) ta" and B :: "('q, 'f) ta"
begin
sublocale horn "\<Delta>_trancl_rules A B" .
lemma \<Delta>_trancl_infer0:
"infer0 = horn.infer0 (\<Delta>\<^sub>\<epsilon>_rules A B)"
by (auto simp: \<Delta>\<^sub>\<epsilon>_rules_def \<Delta>_trancl_rules_def horn.infer0_def)
lemma \<Delta>_trancl_infer1:
"infer1 pq X = horn.infer1 (\<Delta>\<^sub>\<epsilon>_rules A B) pq X \<union>
{(r, snd pq) |r p'. (r, p') \<in> X \<and> p' = fst pq} \<union>
{(fst pq, r) |q' r. (q', r) \<in> (insert pq X) \<and> q' = snd pq}"
unfolding \<Delta>_trancl_rules_def horn_infer1_union Un_assoc
apply (intro arg_cong2[of _ _ _ _ "(\<union>)"] HOL.refl)
by (auto simp: horn.infer1_def simp flip: ex_simps(1)) force+
lemma \<Delta>_trancl_sound:
"\<Delta>_trancl_set A B = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using lr unfolding x
proof (induct)
case (\<Delta>_set_cong f ps p qs q) show ?case
apply (intro infer[of "zip ps qs" "(p, q)"])
subgoal using \<Delta>_set_cong(1-3) by (auto simp: \<Delta>_trancl_rules_def \<Delta>\<^sub>\<epsilon>_rules_def)
subgoal using \<Delta>_set_cong(3,5) by (auto simp: zip_nth_conv)
done
next
case (\<Delta>_set_eps1 p p' q) then show ?case
by (intro infer[of "[(p, q)]" "(p', q)"]) (auto simp: \<Delta>_trancl_rules_def \<Delta>\<^sub>\<epsilon>_rules_def)
next
case (\<Delta>_set_eps2 q q' p) then show ?case
by (intro infer[of "[(p, q)]" "(p, q')"]) (auto simp: \<Delta>_trancl_rules_def \<Delta>\<^sub>\<epsilon>_rules_def)
next
case (\<Delta>_set_trans p q r) then show ?case
by (intro infer[of "[(p,q), (q,r)]" "(p, r)"]) (auto simp: \<Delta>_trancl_rules_def \<Delta>\<^sub>\<epsilon>_rules_def)
qed
next
case (rl x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using rl unfolding x
proof (induct)
case (infer as a) then show ?case
using \<Delta>_set_cong[of _ "map fst as" "fst a" A "map snd as" "snd a" B]
\<Delta>_set_eps1[of _ "fst a" A "snd a" B] \<Delta>_set_eps2[of _ "snd a" B "fst a" A]
\<Delta>_set_trans[of "fst a" "snd (hd as)" A B "snd a"]
by (auto simp: \<Delta>_trancl_rules_def \<Delta>\<^sub>\<epsilon>_rules_def)
qed
qed
end
section \<open>Computing the epsilon transitions for the transitive closure of pair automata\<close>
definition \<Delta>_Atr_rules :: "('q \<times> 'q) fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q \<times> 'q) horn set" where
"\<Delta>_Atr_rules Q A B =
{[] \<rightarrow>\<^sub>h (p, q) | p q. (p , q) |\<in>| Q} \<union>
{[(p, q),(r, v)] \<rightarrow>\<^sub>h (p, v) |p q r v. (q, r) |\<in>| \<Delta>\<^sub>\<epsilon> B A}"
locale \<Delta>_Atr_horn =
fixes Q :: "('q \<times> 'q) fset" and A :: "('q, 'f) ta" and B :: "('q, 'f) ta"
begin
sublocale horn "\<Delta>_Atr_rules Q A B" .
lemma \<Delta>_Atr_infer0: "infer0 = fset Q"
- by (auto simp: horn.infer0_def \<Delta>_Atr_rules_def fmember_iff_member_fset)
+ by (auto simp: horn.infer0_def \<Delta>_Atr_rules_def)
lemma \<Delta>_Atr_infer1:
"infer1 pq X = {(p, snd pq) | p q. (p, q) \<in> X \<and> (q, fst pq) |\<in>| \<Delta>\<^sub>\<epsilon> B A} \<union>
{(fst pq, v) | q r v. (snd pq, r) |\<in>| \<Delta>\<^sub>\<epsilon> B A \<and> (r, v) \<in> X} \<union>
{(fst pq, snd pq) | q . (snd pq, fst pq) |\<in>| \<Delta>\<^sub>\<epsilon> B A}"
unfolding \<Delta>_Atr_rules_def horn_infer1_union
by (auto simp: horn.infer1_def simp flip: ex_simps(1)) force+
lemma \<Delta>_Atr_sound:
"\<Delta>_Atrans_set Q A B = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using lr unfolding x
proof (induct)
case (base p q)
then show ?case
by (intro infer[of "[]" "(p, q)"]) (auto simp: \<Delta>_Atr_rules_def)
next
case (step p q r v)
then show ?case
by (intro infer[of "[(p, q), (r, v)]" "(p, v)"]) (auto simp: \<Delta>_Atr_rules_def)
qed
next
case (rl x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using rl unfolding x
proof (induct)
case (infer as a) then show ?case
using base[of "fst a" "snd a" Q A B]
using \<Delta>_Atrans_set.step[of "fst a" _ Q A B "snd a"]
by (auto simp: \<Delta>_Atr_rules_def) blast
qed
qed
end
section \<open>Computing the Q infinity set for the infinity predicate automaton\<close>
definition Q_inf_rules :: "('q, 'f option \<times> 'g option) ta \<Rightarrow> ('q \<times> 'q) horn set" where
"Q_inf_rules A =
{[] \<rightarrow>\<^sub>h (ps ! i, p) |f ps p i. (None, Some f) ps \<rightarrow> p |\<in>| rules A \<and> i < length ps} \<union>
{[(p, q)] \<rightarrow>\<^sub>h (p, r) |p q r. (q, r) |\<in>| eps A} \<union>
{[(p, q), (q, r)] \<rightarrow>\<^sub>h (p, r) |p q r. True}"
locale Q_horn =
fixes A :: "('q, 'f option \<times> 'g option) ta"
begin
sublocale horn "Q_inf_rules A" .
lemma Q_infer0:
"infer0 = {(ps ! i, p) |f ps p i. (None, Some f) ps \<rightarrow> p |\<in>| rules A \<and> i < length ps}"
unfolding horn.infer0_def Q_inf_rules_def by auto
lemma Q_infer1:
"infer1 pq X = {(fst pq, r) | q r. (q, r) |\<in>| eps A \<and> q = snd pq} \<union>
{(r, snd pq) |r p'. (r, p') \<in> X \<and> p' = fst pq} \<union>
{(fst pq, r) |q' r. (q', r) \<in> (insert pq X) \<and> q' = snd pq}"
unfolding Q_inf_rules_def horn_infer1_union
by (auto simp: horn.infer1_def simp flip: ex_simps(1)) force+
lemma Q_sound:
"Q_inf A = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using lr unfolding x
proof (induct)
case (trans p q r)
then show ?case
by (intro infer[of "[(p,q), (q,r)]" "(p, r)"])
(auto simp: Q_inf_rules_def)
next
case (rule f qs q i)
then show ?case
by (intro infer[of "[]" "(qs ! i, q)"])
(auto simp: Q_inf_rules_def)
next
case (eps p q r)
then show ?case
by (intro infer[of "[(p, q)]" "(p, r)"])
(auto simp: Q_inf_rules_def)
qed
next
case (rl x) obtain p q where x: "x = (p, q)" by (cases x)
show ?case using rl unfolding x
proof (induct)
case (infer as a) then show ?case
using Q_inf.eps[of "fst a" _ A "snd a"]
using Q_inf.trans[of "fst a" "snd (hd as)" A "snd a"]
by (auto simp: Q_inf_rules_def intro: Q_inf.rule)
qed
qed
end
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Regular_Relation_Impl.thy b/thys/Regular_Tree_Relations/Regular_Relation_Impl.thy
--- a/thys/Regular_Tree_Relations/Regular_Relation_Impl.thy
+++ b/thys/Regular_Tree_Relations/Regular_Relation_Impl.thy
@@ -1,301 +1,301 @@
theory Regular_Relation_Impl
imports Tree_Automata_Impl
Regular_Relation_Abstract_Impl
Horn_Fset
begin
section \<open>Computing the epsilon transitions for the composition of GTT's\<close>
definition \<Delta>\<^sub>\<epsilon>_infer0_cont where
"\<Delta>\<^sub>\<epsilon>_infer0_cont \<Delta>\<^sub>A \<Delta>\<^sub>B =
(let arules = filter (\<lambda> r. r_lhs_states r = []) (sorted_list_of_fset \<Delta>\<^sub>A) in
let brules = filter (\<lambda> r. r_lhs_states r = []) (sorted_list_of_fset \<Delta>\<^sub>B) in
(map (map_prod r_rhs r_rhs) (filter (\<lambda>(ra, rb). r_root ra = r_root rb) (List.product arules brules))))"
definition \<Delta>\<^sub>\<epsilon>_infer1_cont where
"\<Delta>\<^sub>\<epsilon>_infer1_cont \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> =
(let (arules, aeps) = (sorted_list_of_fset \<Delta>\<^sub>A, sorted_list_of_fset \<Delta>\<^sub>A\<^sub>\<epsilon>) in
let (brules, beps) = (sorted_list_of_fset \<Delta>\<^sub>B, sorted_list_of_fset \<Delta>\<^sub>B\<^sub>\<epsilon>) in
let prules = List.product arules brules in
(\<lambda> pq bs.
map (map_prod r_rhs r_rhs) (filter (\<lambda>(ra, rb). case (ra, rb) of (TA_rule f ps p, TA_rule g qs q) \<Rightarrow>
f = g \<and> length ps = length qs \<and> (fst pq, snd pq) \<in> set (zip ps qs) \<and>
set (zip ps qs) \<subseteq> insert (fst pq, snd pq) (fset bs)) prules) @
map (\<lambda>(p, p'). (p', snd pq)) (filter (\<lambda>(p, p') \<Rightarrow> p = fst pq) aeps) @
map (\<lambda>(q, q'). (fst pq, q')) (filter (\<lambda>(q, q') \<Rightarrow> q = snd pq) beps)))"
locale \<Delta>\<^sub>\<epsilon>_fset =
fixes \<Delta>\<^sub>A :: "('q :: linorder, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>A\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
and \<Delta>\<^sub>B :: "('q, 'f) ta_rule fset" and \<Delta>\<^sub>B\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
begin
abbreviation A where "A \<equiv> TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>"
abbreviation B where "B \<equiv> TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>"
sublocale \<Delta>\<^sub>\<epsilon>_horn A B .
sublocale l: horn_fset "\<Delta>\<^sub>\<epsilon>_rules A B" "\<Delta>\<^sub>\<epsilon>_infer0_cont \<Delta>\<^sub>A \<Delta>\<^sub>B" "\<Delta>\<^sub>\<epsilon>_infer1_cont \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>"
apply (unfold_locales)
unfolding \<Delta>\<^sub>\<epsilon>_horn.\<Delta>\<^sub>\<epsilon>_infer0 \<Delta>\<^sub>\<epsilon>_horn.\<Delta>\<^sub>\<epsilon>_infer1 \<Delta>\<^sub>\<epsilon>_infer0_cont_def \<Delta>\<^sub>\<epsilon>_infer1_cont_def set_append Un_assoc[symmetric]
unfolding sorted_list_of_fset_simps union_fset
subgoal
apply (auto split!: prod.splits ta_rule.splits simp: comp_def fset_of_list_elem r_rhs_def
- map_prod_def fSigma.rep_eq image_def Bex_def simp flip: fmember_iff_member_fset)
+ map_prod_def fSigma.rep_eq image_def Bex_def)
apply (metis ta_rule.exhaust_sel)
done
unfolding Let_def prod.case set_append Un_assoc
apply (intro arg_cong2[of _ _ _ _ "(\<union>)"])
subgoal
- apply (auto split!: prod.splits ta_rule.splits simp flip: fmember_iff_member_fset )
+ apply (auto split!: prod.splits ta_rule.splits)
apply (smt (verit, del_insts) Pair_inject map_prod_imageI mem_Collect_eq ta_rule.inject ta_rule.sel(3))
done
-by (force simp add: image_def fmember_iff_member_fset split!: prod.splits)+
+by (force simp add: image_def split!: prod.splits)+
lemmas infer = l.infer0 l.infer1
lemmas saturate_impl_sound = l.saturate_impl_sound
lemmas saturate_impl_complete = l.saturate_impl_complete
end
definition \<Delta>\<^sub>\<epsilon>_impl where
"\<Delta>\<^sub>\<epsilon>_impl \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> = horn_fset_impl.saturate_impl (\<Delta>\<^sub>\<epsilon>_infer0_cont \<Delta>\<^sub>A \<Delta>\<^sub>B) (\<Delta>\<^sub>\<epsilon>_infer1_cont \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>)"
lemma \<Delta>\<^sub>\<epsilon>_impl_sound:
assumes "\<Delta>\<^sub>\<epsilon>_impl \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> = Some xs"
shows "xs = \<Delta>\<^sub>\<epsilon> (TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>) (TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>)"
using \<Delta>\<^sub>\<epsilon>_fset.saturate_impl_sound[OF assms[unfolded \<Delta>\<^sub>\<epsilon>_impl_def]]
unfolding \<Delta>\<^sub>\<epsilon>_horn.\<Delta>\<^sub>\<epsilon>_sound[symmetric]
- by (auto simp flip: \<Delta>\<^sub>\<epsilon>.rep_eq simp: fmember_iff_member_fset)
+ by (auto simp flip: \<Delta>\<^sub>\<epsilon>.rep_eq)
lemma \<Delta>\<^sub>\<epsilon>_impl_complete:
fixes \<Delta>\<^sub>A :: "('q :: linorder, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>B :: "('q, 'f) ta_rule fset"
and \<Delta>\<^sub>\<epsilon>\<^sub>A :: "('q \<times> 'q) fset" and \<Delta>\<^sub>\<epsilon>\<^sub>B :: "('q \<times> 'q) fset"
shows "\<Delta>\<^sub>\<epsilon>_impl \<Delta>\<^sub>A \<Delta>\<^sub>\<epsilon>\<^sub>A \<Delta>\<^sub>B \<Delta>\<^sub>\<epsilon>\<^sub>B \<noteq> None" unfolding \<Delta>\<^sub>\<epsilon>_impl_def
by (intro \<Delta>\<^sub>\<epsilon>_fset.saturate_impl_complete)
(auto simp flip: \<Delta>\<^sub>\<epsilon>_horn.\<Delta>\<^sub>\<epsilon>_sound)
lemma \<Delta>\<^sub>\<epsilon>_impl [code]:
"\<Delta>\<^sub>\<epsilon> (TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>) (TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>) = the (\<Delta>\<^sub>\<epsilon>_impl \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>)"
using \<Delta>\<^sub>\<epsilon>_impl_complete[of \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>] \<Delta>\<^sub>\<epsilon>_impl_sound[of \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>]
by force
section \<open>Computing the epsilon transitions for the transitive closure of GTT's\<close>
definition \<Delta>_trancl_infer0 where
"\<Delta>_trancl_infer0 \<Delta>\<^sub>A \<Delta>\<^sub>B = \<Delta>\<^sub>\<epsilon>_infer0_cont \<Delta>\<^sub>A \<Delta>\<^sub>B"
definition \<Delta>_trancl_infer1 :: "('q :: linorder , 'f :: linorder) ta_rule fset \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> ('q, 'f) ta_rule fset \<Rightarrow> ('q \<times> 'q) fset
\<Rightarrow> 'q \<times> 'q \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> ('q \<times> 'q) list" where
"\<Delta>_trancl_infer1 \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> pq bs =
\<Delta>\<^sub>\<epsilon>_infer1_cont \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> pq bs @
sorted_list_of_fset (
(\<lambda>(r, p'). (r, snd pq)) |`| (ffilter (\<lambda>(r, p') \<Rightarrow> p' = fst pq) bs) |\<union>|
(\<lambda>(q', r). (fst pq, r)) |`| (ffilter (\<lambda>(q', r) \<Rightarrow> q' = snd pq) (finsert pq bs)))"
locale \<Delta>_trancl_list =
fixes \<Delta>\<^sub>A :: "('q :: linorder, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>A\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
and \<Delta>\<^sub>B :: "('q, 'f) ta_rule fset" and \<Delta>\<^sub>B\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
begin
abbreviation A where "A \<equiv> TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>"
abbreviation B where "B \<equiv> TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>"
sublocale \<Delta>_trancl_horn A B .
sublocale l: horn_fset "\<Delta>_trancl_rules A B"
"\<Delta>_trancl_infer0 \<Delta>\<^sub>A \<Delta>\<^sub>B" "\<Delta>_trancl_infer1 \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>"
apply (unfold_locales)
unfolding \<Delta>_trancl_rules_def horn_infer0_union horn_infer1_union
\<Delta>_trancl_infer0_def \<Delta>_trancl_infer1_def \<Delta>\<^sub>\<epsilon>_fset.infer set_append
by (auto simp flip: ex_simps(1) simp: horn.infer0_def horn.infer1_def intro!: arg_cong2[of _ _ _ _ "(\<union>)"])
lemmas saturate_impl_sound = l.saturate_impl_sound
lemmas saturate_impl_complete = l.saturate_impl_complete
end
definition "\<Delta>_trancl_impl \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> =
horn_fset_impl.saturate_impl (\<Delta>_trancl_infer0 \<Delta>\<^sub>A \<Delta>\<^sub>B) (\<Delta>_trancl_infer1 \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>)"
lemma \<Delta>_trancl_impl_sound:
assumes "\<Delta>_trancl_impl \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> = Some xs"
shows "xs = \<Delta>_trancl (TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>) (TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>)"
using \<Delta>_trancl_list.saturate_impl_sound[OF assms[unfolded \<Delta>_trancl_impl_def]]
unfolding \<Delta>_trancl_horn.\<Delta>_trancl_sound[symmetric] \<Delta>_trancl.rep_eq[symmetric]
- by (auto simp: fmember_iff_member_fset)
+ by auto
lemma \<Delta>_trancl_impl_complete:
fixes \<Delta>\<^sub>A :: "('q :: linorder, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>B :: "('q, 'f) ta_rule fset"
and \<Delta>\<^sub>A\<^sub>\<epsilon> :: "('q \<times> 'q) fset" and \<Delta>\<^sub>B\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
shows "\<Delta>_trancl_impl \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> \<noteq> None"
unfolding \<Delta>_trancl_impl_def
by (intro \<Delta>_trancl_list.saturate_impl_complete)
(auto simp flip: \<Delta>_trancl_horn.\<Delta>_trancl_sound)
lemma \<Delta>_trancl_impl [code]:
"\<Delta>_trancl (TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>) (TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>) = (the (\<Delta>_trancl_impl \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>))"
using \<Delta>_trancl_impl_complete[of \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>]
using \<Delta>_trancl_impl_sound[of \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>]
by force
section \<open>Computing the epsilon transitions for the transitive closure of pair automata\<close>
definition \<Delta>_Atr_infer1_cont :: "('q :: linorder \<times> 'q) fset \<Rightarrow> ('q, 'f :: linorder) ta_rule fset \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow>
('q, 'f) ta_rule fset \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> 'q \<times> 'q \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> ('q \<times> 'q) list" where
"\<Delta>_Atr_infer1_cont Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> =
(let G = sorted_list_of_fset (the (\<Delta>\<^sub>\<epsilon>_impl \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>)) in
(\<lambda> pq bs.
(let bs_list = sorted_list_of_fset bs in
map (\<lambda> (p, q). (fst p, snd pq)) (filter (\<lambda> (p, q). snd p = fst q \<and> snd q = fst pq) (List.product bs_list G)) @
map (\<lambda> (p, q). (fst pq, snd q)) (filter (\<lambda> (p, q). snd p = fst q \<and> fst p = snd pq) (List.product G bs_list)) @
map (\<lambda> (p, q). (fst pq, snd pq)) (filter (\<lambda> (p, q). snd pq = p \<and> fst pq = q) G))))"
locale \<Delta>_Atr_fset =
fixes Q :: "('q :: linorder \<times> 'q) fset" and \<Delta>\<^sub>A :: "('q, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>A\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
and \<Delta>\<^sub>B :: "('q, 'f) ta_rule fset" and \<Delta>\<^sub>B\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
begin
abbreviation A where "A \<equiv> TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>"
abbreviation B where "B \<equiv> TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>"
sublocale \<Delta>_Atr_horn Q A B .
lemma infer1:
"infer1 pq (fset bs) = set (\<Delta>_Atr_infer1_cont Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> pq bs)"
proof -
have "{(p, snd pq) | p q. (p, q) \<in> (fset bs) \<and> (q, fst pq) |\<in>| \<Delta>\<^sub>\<epsilon> B A} \<union>
{(fst pq, v) | q r v. (snd pq, r) |\<in>| \<Delta>\<^sub>\<epsilon> B A \<and> (r, v) \<in> (fset bs)} \<union>
{(fst pq, snd pq) | q . (snd pq, fst pq) |\<in>| \<Delta>\<^sub>\<epsilon> B A} = set (\<Delta>_Atr_infer1_cont Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> pq bs)"
unfolding \<Delta>_Atr_infer1_cont_def set_append Un_assoc[symmetric] Let_def
unfolding sorted_list_of_fset_simps union_fset
apply (intro arg_cong2[of _ _ _ _ "(\<union>)"])
- apply (simp_all add: fSigma_repeq fmember_iff_member_fset flip: \<Delta>\<^sub>\<epsilon>_impl fset_of_list_elem)
+ apply (simp_all add: fSigma_repeq flip: \<Delta>\<^sub>\<epsilon>_impl fset_of_list_elem)
apply force+
done
then show ?thesis
using \<Delta>_Atr_horn.\<Delta>_Atr_infer1[of Q A B pq "fset bs"]
by simp
qed
sublocale l: horn_fset "\<Delta>_Atr_rules Q A B" "sorted_list_of_fset Q" "\<Delta>_Atr_infer1_cont Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>"
apply (unfold_locales)
unfolding \<Delta>_Atr_horn.\<Delta>_Atr_infer0 fset_of_list.rep_eq
using infer1
by auto
lemmas infer = l.infer0 l.infer1
lemmas saturate_impl_sound = l.saturate_impl_sound
lemmas saturate_impl_complete = l.saturate_impl_complete
end
definition "\<Delta>_Atr_impl Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> =
horn_fset_impl.saturate_impl (sorted_list_of_fset Q) (\<Delta>_Atr_infer1_cont Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>)"
lemma \<Delta>_Atr_impl_sound:
assumes "\<Delta>_Atr_impl Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> = Some xs"
shows "xs = \<Delta>_Atrans Q (TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>) (TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>)"
using \<Delta>_Atr_fset.saturate_impl_sound[OF assms[unfolded \<Delta>_Atr_impl_def]]
unfolding \<Delta>_Atr_horn.\<Delta>_Atr_sound[symmetric] \<Delta>_Atrans.rep_eq[symmetric]
by (simp add: fset_inject)
lemma \<Delta>_Atr_impl_complete:
shows "\<Delta>_Atr_impl Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon> \<noteq> None" unfolding \<Delta>_Atr_impl_def
by (intro \<Delta>_Atr_fset.saturate_impl_complete)
(auto simp: finite_\<Delta>_Atrans_set simp flip: \<Delta>_Atr_horn.\<Delta>_Atr_sound)
lemma \<Delta>_Atr_impl [code]:
"\<Delta>_Atrans Q (TA \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon>) (TA \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>) = (the (\<Delta>_Atr_impl Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>))"
using \<Delta>_Atr_impl_complete[of Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>] \<Delta>_Atr_impl_sound[of Q \<Delta>\<^sub>A \<Delta>\<^sub>A\<^sub>\<epsilon> \<Delta>\<^sub>B \<Delta>\<^sub>B\<^sub>\<epsilon>]
by force
section \<open>Computing the Q infinity set for the infinity predicate automaton\<close>
definition Q_infer0_cont :: "('q :: linorder, 'f :: linorder option \<times> 'g :: linorder option) ta_rule fset \<Rightarrow> ('q \<times> 'q) list" where
"Q_infer0_cont \<Delta> = concat (sorted_list_of_fset (
(\<lambda> r. case r of TA_rule f ps p \<Rightarrow> map (\<lambda> x. Pair x p) ps) |`|
(ffilter (\<lambda> r. case r of TA_rule f ps p \<Rightarrow> fst f = None \<and> snd f \<noteq> None \<and> ps \<noteq> []) \<Delta>)))"
definition Q_infer1_cont :: "('q ::linorder \<times> 'q) fset \<Rightarrow> 'q \<times> 'q \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> ('q \<times> 'q) list" where
"Q_infer1_cont \<Delta>\<epsilon> =
(let eps = sorted_list_of_fset \<Delta>\<epsilon> in
(\<lambda> pq bs.
let bs_list = sorted_list_of_fset bs in
map (\<lambda> (q, r). (fst pq, r)) (filter (\<lambda> (q, r) \<Rightarrow> q = snd pq) eps) @
map (\<lambda>(r, p'). (r, snd pq)) (filter (\<lambda>(r, p') \<Rightarrow> p' = fst pq) bs_list) @
map (\<lambda>(q', r). (fst pq, r)) (filter (\<lambda>(q', r) \<Rightarrow> q' = snd pq) (pq # bs_list))))"
locale Q_fset =
fixes \<Delta> :: "('q :: linorder, 'f :: linorder option \<times> 'g :: linorder option) ta_rule fset" and \<Delta>\<epsilon> :: "('q \<times> 'q) fset"
begin
abbreviation A where "A \<equiv> TA \<Delta> \<Delta>\<epsilon>"
sublocale Q_horn A .
sublocale l: horn_fset "Q_inf_rules A" "Q_infer0_cont \<Delta>" "Q_infer1_cont \<Delta>\<epsilon>"
apply (unfold_locales)
unfolding Q_horn.Q_infer0 Q_horn.Q_infer1 Q_infer0_cont_def Q_infer1_cont_def set_append Un_assoc[symmetric]
unfolding sorted_list_of_fset_simps union_fset
subgoal
- apply (auto simp add: Bex_def fmember_iff_member_fset split!: ta_rule.splits)
+ apply (auto simp add: Bex_def split!: ta_rule.splits)
apply (rule_tac x = "TA_rule (lift_None_Some f) ps p" in exI)
apply (force dest: in_set_idx)+
done
unfolding Let_def set_append Un_assoc
- by (intro arg_cong2[of _ _ _ _ "(\<union>)"]) (auto simp add: fmember_iff_member_fset)
+ by (intro arg_cong2[of _ _ _ _ "(\<union>)"]) auto
lemmas saturate_impl_sound = l.saturate_impl_sound
lemmas saturate_impl_complete = l.saturate_impl_complete
end
definition Q_impl where
"Q_impl \<Delta> \<Delta>\<epsilon> = horn_fset_impl.saturate_impl (Q_infer0_cont \<Delta>) (Q_infer1_cont \<Delta>\<epsilon>)"
lemma Q_impl_sound:
"Q_impl \<Delta> \<Delta>\<epsilon> = Some xs \<Longrightarrow> fset xs = Q_inf (TA \<Delta> \<Delta>\<epsilon>)"
using Q_fset.saturate_impl_sound unfolding Q_impl_def Q_horn.Q_sound .
lemma Q_impl_complete:
"Q_impl \<Delta> \<Delta>\<epsilon> \<noteq> None"
proof -
let ?A = "TA \<Delta> \<Delta>\<epsilon>"
have *: "Q_inf ?A \<subseteq> fset (\<Q> ?A |\<times>| \<Q> ?A)"
- by (auto simp add: Q_inf_states_ta_states(1, 2) subrelI simp flip: fmember_iff_member_fset)
+ by (auto simp add: Q_inf_states_ta_states(1, 2) subrelI)
have "finite (Q_inf ?A)"
by (intro finite_subset[OF *]) simp
then show ?thesis unfolding Q_impl_def
by (intro Q_fset.saturate_impl_complete) (auto simp: Q_horn.Q_sound)
qed
definition "Q_infinity_impl \<Delta> \<Delta>\<epsilon> = (let Q = the (Q_impl \<Delta> \<Delta>\<epsilon>) in
snd |`| ((ffilter (\<lambda> (p, q). p = q) Q) |O| Q))"
lemma Q_infinity_impl_fmember:
"q |\<in>| Q_infinity_impl \<Delta> \<Delta>\<epsilon> \<longleftrightarrow> (\<exists> p. (p, p) |\<in>| the (Q_impl \<Delta> \<Delta>\<epsilon>) \<and>
(p, q) |\<in>| the (Q_impl \<Delta> \<Delta>\<epsilon>))"
unfolding Q_infinity_impl_def
- by (auto simp: Let_def fimage_iff fBex_def) fastforce
+ by (auto simp: Let_def image_iff Bex_def) fastforce
lemma loop_sound_correct [simp]:
"fset (Q_infinity_impl \<Delta> \<Delta>\<epsilon>) = Q_inf_e (TA \<Delta> \<Delta>\<epsilon>)"
proof -
obtain Q where [simp]: "Q_impl \<Delta> \<Delta>\<epsilon> = Some Q" using Q_impl_complete[of \<Delta> \<Delta>\<epsilon>]
by blast
have "fset Q = (Q_inf (TA \<Delta> \<Delta>\<epsilon>))"
using Q_impl_sound[of \<Delta> \<Delta>\<epsilon>]
by (auto simp: fset_of_list.rep_eq)
then show ?thesis
by (force simp: Q_infinity_impl_fmember Let_def fset_of_list_elem
- fset_of_list.rep_eq simp flip: fmember_iff_member_fset)
+ fset_of_list.rep_eq)
qed
lemma fQ_inf_e_code [code]:
"fQ_inf_e (TA \<Delta> \<Delta>\<epsilon>) = Q_infinity_impl \<Delta> \<Delta>\<epsilon>"
using loop_sound_correct
- by (auto simp add: fQ_inf_e.rep_eq fmember_iff_member_fset)
+ by (auto simp add: fQ_inf_e.rep_eq)
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Tree_Automata/Myhill_Nerode.thy b/thys/Regular_Tree_Relations/Tree_Automata/Myhill_Nerode.thy
--- a/thys/Regular_Tree_Relations/Tree_Automata/Myhill_Nerode.thy
+++ b/thys/Regular_Tree_Relations/Tree_Automata/Myhill_Nerode.thy
@@ -1,80 +1,80 @@
theory Myhill_Nerode
imports Tree_Automata Ground_Ctxt
begin
subsection \<open>Myhill Nerode characterization for regular tree languages\<close>
lemma ground_ctxt_apply_pres_der:
assumes "ta_der \<A> (term_of_gterm s) = ta_der \<A> (term_of_gterm t)"
shows "ta_der \<A> (term_of_gterm C\<langle>s\<rangle>\<^sub>G) = ta_der \<A> (term_of_gterm C\<langle>t\<rangle>\<^sub>G)" using assms
by (induct C) (auto, (metis append_Cons_nth_not_middle nth_append_length)+)
locale myhill_nerode =
fixes \<F> \<L> assumes term_subset: "\<L> \<subseteq> \<T>\<^sub>G \<F>"
begin
definition myhill ("_ \<equiv>\<^sub>\<L> _") where
"myhill s t \<equiv> s \<in> \<T>\<^sub>G \<F> \<and> t \<in> \<T>\<^sub>G \<F> \<and> (\<forall> C. C\<langle>s\<rangle>\<^sub>G \<in> \<L> \<and> C\<langle>t\<rangle>\<^sub>G \<in> \<L> \<or> C\<langle>s\<rangle>\<^sub>G \<notin> \<L> \<and> C\<langle>t\<rangle>\<^sub>G \<notin> \<L>)"
lemma myhill_sound: "s \<equiv>\<^sub>\<L> t \<Longrightarrow> s \<in> \<T>\<^sub>G \<F>" "s \<equiv>\<^sub>\<L> t \<Longrightarrow> t \<in> \<T>\<^sub>G \<F>"
unfolding myhill_def by auto
lemma myhill_refl [simp]: "s \<in> \<T>\<^sub>G \<F> \<Longrightarrow> s \<equiv>\<^sub>\<L> s"
unfolding myhill_def by auto
lemma myhill_symmetric: "s \<equiv>\<^sub>\<L> t \<Longrightarrow> t \<equiv>\<^sub>\<L> s"
unfolding myhill_def by auto
lemma myhill_trans [trans]:
"s \<equiv>\<^sub>\<L> t \<Longrightarrow> t \<equiv>\<^sub>\<L> u \<Longrightarrow> s \<equiv>\<^sub>\<L> u"
unfolding myhill_def by auto
abbreviation myhill_r ("MN\<^sub>\<L>") where
"myhill_r \<equiv> {(s, t) | s t. s \<equiv>\<^sub>\<L> t}"
lemma myhill_equiv:
"equiv (\<T>\<^sub>G \<F>) MN\<^sub>\<L>"
apply (intro equivI) apply (auto simp: myhill_sound myhill_symmetric sym_def trans_def refl_on_def)
using myhill_trans by blast
lemma rtl_der_image_on_myhill_inj:
assumes "gta_lang Q\<^sub>f \<A> = \<L>"
shows "inj_on (\<lambda> X. gta_der \<A> ` X) (\<T>\<^sub>G \<F> // MN\<^sub>\<L>)" (is "inj_on ?D ?R")
proof -
{fix S T assume eq_rel: "S \<in> ?R" "T \<in> ?R" "?D S = ?D T"
have "\<And> s t. s \<in> S \<Longrightarrow> t \<in> T \<Longrightarrow> s \<equiv>\<^sub>\<L> t"
proof -
fix s t assume mem: "s \<in> S" "t \<in> T"
then obtain t' where res: "t' \<in> T" "gta_der \<A> s = gta_der \<A> t'" using eq_rel(3)
by (metis image_iff)
from res(1) mem have "s \<in> \<T>\<^sub>G \<F>" "t \<in> \<T>\<^sub>G \<F>" "t' \<in> \<T>\<^sub>G \<F>" using eq_rel(1, 2)
using in_quotient_imp_subset myhill_equiv by blast+
then have "s \<equiv>\<^sub>\<L> t'" using assms res ground_ctxt_apply_pres_der[of \<A> s]
by (auto simp: myhill_def gta_der_def simp flip: ctxt_of_gctxt_apply
elim!: gta_langE intro: gta_langI)
moreover have "t' \<equiv>\<^sub>\<L> t" using quotient_eq_iff[OF myhill_equiv eq_rel(2) eq_rel(2) res(1) mem(2)]
by simp
ultimately show "s \<equiv>\<^sub>\<L> t" using myhill_trans by blast
qed
then have "\<And> s t. s \<in> S \<Longrightarrow> t \<in> T \<Longrightarrow> (s, t) \<in> MN\<^sub>\<L>" by blast
then have "S = T" using quotient_eq_iff[OF myhill_equiv eq_rel(1, 2)]
using eq_rel(3) by fastforce}
then show inj: "inj_on ?D ?R" by (meson inj_onI)
qed
lemma rtl_implies_finite_indexed_myhill_relation:
assumes "gta_lang Q\<^sub>f \<A> = \<L>"
shows "finite (\<T>\<^sub>G \<F> // MN\<^sub>\<L>)" (is "finite ?R")
proof -
let ?D = "\<lambda> X. gta_der \<A> ` X"
have image: "?D ` ?R \<subseteq> Pow (fset (fPow (\<Q> \<A>)))" unfolding gta_der_def
- by (meson PowI fPowI ground_ta_der_states ground_term_of_gterm image_subsetI notin_fset)
+ by (meson PowI fPowI ground_ta_der_states ground_term_of_gterm image_subsetI)
then have "finite (Pow (fset (fPow (\<Q> \<A>))))" by simp
then have "finite (?D ` ?R)" using finite_subset[OF image] by fastforce
then show ?thesis using finite_image_iff[OF rtl_der_image_on_myhill_inj[OF assms]]
by blast
qed
end
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy
--- a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy
+++ b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy
@@ -1,2184 +1,2191 @@
section \<open>Tree automaton\<close>
theory Tree_Automata
imports FSet_Utils
"HOL-Library.Product_Lexorder"
"HOL-Library.Option_ord"
begin
subsection \<open>Tree automaton definition and functionality\<close>
datatype ('q, 'f) ta_rule = TA_rule (r_root: 'f) (r_lhs_states: "'q list") (r_rhs: 'q) ("_ _ \<rightarrow> _" [51, 51, 51] 52)
datatype ('q, 'f) ta = TA (rules: "('q, 'f) ta_rule fset") (eps: "('q \<times> 'q) fset")
text \<open>In many application we are interested in specific subset of all terms. If these
can be captured by a tree automaton (identified by a state) then we say the set is regular.
This gives the motivation for the following definition\<close>
datatype ('q, 'f) reg = Reg (fin: "'q fset") (ta: "('q, 'f) ta")
text \<open>The state set induced by a tree automaton is implicit in our representation.
We compute it based on the rules and epsilon transitions of a given tree automaton\<close>
abbreviation rule_arg_states where "rule_arg_states \<Delta> \<equiv> |\<Union>| ((fset_of_list \<circ> r_lhs_states) |`| \<Delta>)"
abbreviation rule_target_states where "rule_target_states \<Delta> \<equiv> (r_rhs |`| \<Delta>)"
definition rule_states where "rule_states \<Delta> \<equiv> rule_arg_states \<Delta> |\<union>| rule_target_states \<Delta>"
definition eps_states where "eps_states \<Delta>\<^sub>\<epsilon> \<equiv> (fst |`| \<Delta>\<^sub>\<epsilon>) |\<union>| (snd |`| \<Delta>\<^sub>\<epsilon>)"
definition "\<Q> \<A> = rule_states (rules \<A>) |\<union>| eps_states (eps \<A>)"
abbreviation "\<Q>\<^sub>r \<A> \<equiv> \<Q> (ta \<A>)"
definition ta_rhs_states :: "('q, 'f) ta \<Rightarrow> 'q fset" where
"ta_rhs_states \<A> \<equiv> {| q | p q. (p |\<in>| rule_target_states (rules \<A>)) \<and> (p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|)|}"
definition "ta_sig \<A> = (\<lambda> r. (r_root r, length (r_lhs_states r))) |`| (rules \<A>)"
subsubsection \<open>Rechability of a term induced by a tree automaton\<close>
(* The reachable states of some term. *)
fun ta_der :: "('q, 'f) ta \<Rightarrow> ('f, 'q) term \<Rightarrow> 'q fset" where
"ta_der \<A> (Var q) = {|q' | q'. q = q' \<or> (q, q') |\<in>| (eps \<A>)|\<^sup>+| |}"
| "ta_der \<A> (Fun f ts) = {|q' | q' q qs.
TA_rule f qs q |\<in>| (rules \<A>) \<and> (q = q' \<or> (q, q') |\<in>| (eps \<A>)|\<^sup>+|) \<and> length qs = length ts \<and>
(\<forall> i < length ts. qs ! i |\<in>| ta_der \<A> (ts ! i))|}"
(* The reachable mixed terms of some term. *)
fun ta_der' :: "('q,'f) ta \<Rightarrow> ('f,'q) term \<Rightarrow> ('f,'q) term fset" where
"ta_der' \<A> (Var p) = {|Var q | q. p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+| |}"
| "ta_der' \<A> (Fun f ts) = {|Var q | q. q |\<in>| ta_der \<A> (Fun f ts)|} |\<union>|
{|Fun f ss | ss. length ss = length ts \<and>
(\<forall>i < length ts. ss ! i |\<in>| ta_der' \<A> (ts ! i))|}"
text \<open>Sometimes it is useful to analyse a concrete computation done by a tree automaton.
To do this we introduce the notion of run which keeps track which states are computed in each
subterm to reach a certain state.\<close>
abbreviation "ex_rule_state \<equiv> fst \<circ> groot_sym"
abbreviation "ex_comp_state \<equiv> snd \<circ> groot_sym"
inductive run for \<A> where
step: "length qs = length ts \<Longrightarrow> (\<forall> i < length ts. run \<A> (qs ! i) (ts ! i)) \<Longrightarrow>
TA_rule f (map ex_comp_state qs) q |\<in>| (rules \<A>) \<Longrightarrow> (q = q' \<or> (q, q') |\<in>| (eps \<A>)|\<^sup>+|) \<Longrightarrow>
run \<A> (GFun (q, q') qs) (GFun f ts)"
subsubsection \<open>Language acceptance\<close>
definition ta_lang :: "'q fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('f, 'v) terms" where
[code del]: "ta_lang Q \<A> = {adapt_vars t | t. ground t \<and> Q |\<inter>| ta_der \<A> t \<noteq> {||}}"
definition gta_der where
"gta_der \<A> t = ta_der \<A> (term_of_gterm t)"
definition gta_lang where
"gta_lang Q \<A> = {t. Q |\<inter>| gta_der \<A> t \<noteq> {||}}"
definition \<L> where
"\<L> \<A> = gta_lang (fin \<A>) (ta \<A>)"
definition reg_Restr_Q\<^sub>f where
"reg_Restr_Q\<^sub>f R = Reg (fin R |\<inter>| \<Q>\<^sub>r R) (ta R)"
subsubsection \<open>Trimming\<close>
definition ta_restrict where
"ta_restrict \<A> Q = TA {| TA_rule f qs q| f qs q. TA_rule f qs q |\<in>| rules \<A> \<and> fset_of_list qs |\<subseteq>| Q \<and> q |\<in>| Q |} (fRestr (eps \<A>) Q)"
definition ta_reachable :: "('q, 'f) ta \<Rightarrow> 'q fset" where
"ta_reachable \<A> = {|q| q. \<exists> t. ground t \<and> q |\<in>| ta_der \<A> t |}"
definition ta_productive :: "'q fset \<Rightarrow> ('q,'f) ta \<Rightarrow> 'q fset" where
"ta_productive P \<A> \<equiv> {|q| q q' C. q' |\<in>| ta_der \<A> (C\<langle>Var q\<rangle>) \<and> q' |\<in>| P |}"
text \<open>An automaton is trim if all its states are reachable and productive.\<close>
definition ta_is_trim :: "'q fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> bool" where
"ta_is_trim P \<A> \<equiv> \<forall> q. q |\<in>| \<Q> \<A> \<longrightarrow> q |\<in>| ta_reachable \<A> \<and> q |\<in>| ta_productive P \<A>"
definition reg_is_trim :: "('q, 'f) reg \<Rightarrow> bool" where
"reg_is_trim R \<equiv> ta_is_trim (fin R) (ta R)"
text \<open>We obtain a trim automaton by restriction it to reachable and productive states.\<close>
abbreviation ta_only_reach :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta" where
"ta_only_reach \<A> \<equiv> ta_restrict \<A> (ta_reachable \<A>)"
abbreviation ta_only_prod :: "'q fset \<Rightarrow> ('q,'f) ta \<Rightarrow> ('q,'f) ta" where
"ta_only_prod P \<A> \<equiv> ta_restrict \<A> (ta_productive P \<A>)"
definition reg_reach where
"reg_reach R = Reg (fin R) (ta_only_reach (ta R))"
definition reg_prod where
"reg_prod R = Reg (fin R) (ta_only_prod (fin R) (ta R))"
definition trim_ta :: "'q fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> ('q, 'f) ta" where
"trim_ta P \<A> = ta_only_prod P (ta_only_reach \<A>)"
definition trim_reg where
"trim_reg R = Reg (fin R) (trim_ta (fin R) (ta R))"
subsubsection \<open>Mapping over tree automata\<close>
definition fmap_states_ta :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'f) ta \<Rightarrow> ('b, 'f) ta" where
"fmap_states_ta f \<A> = TA (map_ta_rule f id |`| rules \<A>) (map_both f |`| eps \<A>)"
definition fmap_funs_ta :: "('f \<Rightarrow> 'g) \<Rightarrow> ('a, 'f) ta \<Rightarrow> ('a, 'g) ta" where
"fmap_funs_ta f \<A> = TA (map_ta_rule id f |`| rules \<A>) (eps \<A>)"
definition fmap_states_reg :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'f) reg \<Rightarrow> ('b, 'f) reg" where
"fmap_states_reg f R = Reg (f |`| fin R) (fmap_states_ta f (ta R))"
definition fmap_funs_reg :: "('f \<Rightarrow> 'g) \<Rightarrow> ('a, 'f) reg \<Rightarrow> ('a, 'g) reg" where
"fmap_funs_reg f R = Reg (fin R) (fmap_funs_ta f (ta R))"
subsubsection \<open>Product construction (language intersection)\<close>
definition prod_ta_rules :: "('q1,'f) ta \<Rightarrow> ('q2,'f) ta \<Rightarrow> ('q1 \<times> 'q2, 'f) ta_rule fset" where
"prod_ta_rules \<A> \<B> = {| TA_rule f qs q | f qs q. TA_rule f (map fst qs) (fst q) |\<in>| rules \<A> \<and>
TA_rule f (map snd qs) (snd q) |\<in>| rules \<B>|}"
declare prod_ta_rules_def [simp]
definition prod_epsLp where
"prod_epsLp \<A> \<B> = (\<lambda> (p, q). (fst p, fst q) |\<in>| eps \<A> \<and> snd p = snd q \<and> snd q |\<in>| \<Q> \<B>)"
definition prod_epsRp where
"prod_epsRp \<A> \<B> = (\<lambda> (p, q). (snd p, snd q) |\<in>| eps \<B> \<and> fst p = fst q \<and> fst q |\<in>| \<Q> \<A>)"
definition prod_ta :: "('q1,'f) ta \<Rightarrow> ('q2,'f) ta \<Rightarrow> ('q1 \<times> 'q2, 'f) ta" where
"prod_ta \<A> \<B> = TA (prod_ta_rules \<A> \<B>)
(fCollect (prod_epsLp \<A> \<B>) |\<union>| fCollect (prod_epsRp \<A> \<B>))"
definition reg_intersect where
"reg_intersect R L = Reg (fin R |\<times>| fin L) (prod_ta (ta R) (ta L))"
subsubsection \<open>Union construction (language union)\<close>
definition ta_union where
"ta_union \<A> \<B> = TA (rules \<A> |\<union>| rules \<B>) (eps \<A> |\<union>| eps \<B>)"
definition reg_union where
"reg_union R L = Reg (Inl |`| (fin R |\<inter>| \<Q>\<^sub>r R) |\<union>| Inr |`| (fin L |\<inter>| \<Q>\<^sub>r L))
(ta_union (fmap_states_ta Inl (ta R)) (fmap_states_ta Inr (ta L)))"
subsubsection \<open>Epsilon free and tree automaton accepting empty language\<close>
definition eps_free_rulep where
"eps_free_rulep \<A> = (\<lambda> r. \<exists> f qs q q'. r = TA_rule f qs q' \<and> TA_rule f qs q |\<in>| rules \<A> \<and> (q = q' \<or> (q, q') |\<in>| (eps \<A>)|\<^sup>+|))"
definition eps_free :: "('q, 'f) ta \<Rightarrow> ('q, 'f) ta" where
"eps_free \<A> = TA (fCollect (eps_free_rulep \<A>)) {||}"
definition is_ta_eps_free :: "('q, 'f) ta \<Rightarrow> bool" where
"is_ta_eps_free \<A> \<longleftrightarrow> eps \<A> = {||}"
definition ta_empty :: "'q fset \<Rightarrow> ('q,'f) ta \<Rightarrow> bool" where
"ta_empty Q \<A> \<longleftrightarrow> ta_reachable \<A> |\<inter>| Q |\<subseteq>| {||}"
definition eps_free_reg where
"eps_free_reg R = Reg (fin R) (eps_free (ta R))"
definition reg_empty where
"reg_empty R = ta_empty (fin R) (ta R)"
subsubsection \<open>Relabeling tree automaton states to natural numbers\<close>
definition map_fset_to_nat :: "('a :: linorder) fset \<Rightarrow> 'a \<Rightarrow> nat" where
"map_fset_to_nat X = (\<lambda>x. the (mem_idx x (sorted_list_of_fset X)))"
definition map_fset_fset_to_nat :: "('a :: linorder) fset fset \<Rightarrow> 'a fset \<Rightarrow> nat" where
"map_fset_fset_to_nat X = (\<lambda>x. the (mem_idx (sorted_list_of_fset x) (sorted_list_of_fset (sorted_list_of_fset |`| X))))"
definition relabel_ta :: "('q :: linorder, 'f) ta \<Rightarrow> (nat, 'f) ta" where
"relabel_ta \<A> = fmap_states_ta (map_fset_to_nat (\<Q> \<A>)) \<A>"
definition relabel_Q\<^sub>f :: "('q :: linorder) fset \<Rightarrow> ('q :: linorder, 'f) ta \<Rightarrow> nat fset" where
"relabel_Q\<^sub>f Q \<A> = map_fset_to_nat (\<Q> \<A>) |`| (Q |\<inter>| \<Q> \<A>)"
definition relabel_reg :: "('q :: linorder, 'f) reg \<Rightarrow> (nat, 'f) reg" where
"relabel_reg R = Reg (relabel_Q\<^sub>f (fin R) (ta R)) (relabel_ta (ta R))"
\<comment> \<open>The instantiation of $<$ and $\leq$ for finite sets are $\mid \subset \mid$ and $\mid \subseteq \mid$
which don't give rise to a total order and therefore it cannot be an instance of the type class linorder.
However taking the lexographic order of the sorted list of each finite set gives rise
to a total order. Therefore we provide a relabeling for tree automata where the states
are finite sets. This allows us to relabel the well known power set construction.\<close>
definition relabel_fset_ta :: "(('q :: linorder) fset, 'f) ta \<Rightarrow> (nat, 'f) ta" where
"relabel_fset_ta \<A> = fmap_states_ta (map_fset_fset_to_nat (\<Q> \<A>)) \<A>"
definition relabel_fset_Q\<^sub>f :: "('q :: linorder) fset fset \<Rightarrow> (('q :: linorder) fset, 'f) ta \<Rightarrow> nat fset" where
"relabel_fset_Q\<^sub>f Q \<A> = map_fset_fset_to_nat (\<Q> \<A>) |`| (Q |\<inter>| \<Q> \<A>)"
definition relable_fset_reg :: "(('q :: linorder) fset, 'f) reg \<Rightarrow> (nat, 'f) reg" where
"relable_fset_reg R = Reg (relabel_fset_Q\<^sub>f (fin R) (ta R)) (relabel_fset_ta (ta R))"
definition "srules \<A> = fset (rules \<A>)"
definition "seps \<A> = fset (eps \<A>)"
lemma rules_transfer [transfer_rule]:
"rel_fun (=) (pcr_fset (=)) srules rules" unfolding rel_fun_def
by (auto simp add: cr_fset_def fset.pcr_cr_eq srules_def)
lemma eps_transfer [transfer_rule]:
"rel_fun (=) (pcr_fset (=)) seps eps" unfolding rel_fun_def
by (auto simp add: cr_fset_def fset.pcr_cr_eq seps_def)
lemma TA_equalityI:
"rules \<A> = rules \<B> \<Longrightarrow> eps \<A> = eps \<B> \<Longrightarrow> \<A> = \<B>"
using ta.expand by blast
lemma rule_states_code [code]:
"rule_states \<Delta> = |\<Union>| ((\<lambda> r. finsert (r_rhs r) (fset_of_list (r_lhs_states r))) |`| \<Delta>)"
unfolding rule_states_def
by fastforce
lemma eps_states_code [code]:
"eps_states \<Delta>\<^sub>\<epsilon> = |\<Union>| ((\<lambda> (q,q'). {|q,q'|}) |`| \<Delta>\<^sub>\<epsilon>)" (is "?Ls = ?Rs")
unfolding eps_states_def
by (force simp add: case_prod_beta')
lemma rule_states_empty [simp]:
"rule_states {||} = {||}"
by (auto simp: rule_states_def)
lemma eps_states_empty [simp]:
"eps_states {||} = {||}"
by (auto simp: eps_states_def)
lemma rule_states_union [simp]:
"rule_states (\<Delta> |\<union>| \<Gamma>) = rule_states \<Delta> |\<union>| rule_states \<Gamma>"
unfolding rule_states_def
by fastforce
lemma rule_states_mono:
"\<Delta> |\<subseteq>| \<Gamma> \<Longrightarrow> rule_states \<Delta> |\<subseteq>| rule_states \<Gamma>"
unfolding rule_states_def
by force
lemma eps_states_union [simp]:
"eps_states (\<Delta> |\<union>| \<Gamma>) = eps_states \<Delta> |\<union>| eps_states \<Gamma>"
unfolding eps_states_def
by auto
lemma eps_states_image [simp]:
"eps_states (map_both f |`| \<Delta>\<^sub>\<epsilon>) = f |`| eps_states \<Delta>\<^sub>\<epsilon>"
unfolding eps_states_def map_prod_def
by (force simp: fimage_iff)
lemma eps_states_mono:
"\<Delta> |\<subseteq>| \<Gamma> \<Longrightarrow> eps_states \<Delta> |\<subseteq>| eps_states \<Gamma>"
unfolding eps_states_def
by transfer auto
lemma eps_statesI [intro]:
"(p, q) |\<in>| \<Delta> \<Longrightarrow> p |\<in>| eps_states \<Delta>"
"(p, q) |\<in>| \<Delta> \<Longrightarrow> q |\<in>| eps_states \<Delta>"
unfolding eps_states_def
- by (auto simp add: rev_fimage_eqI)
+ by (auto simp add: rev_image_eqI)
lemma eps_statesE [elim]:
assumes "p |\<in>| eps_states \<Delta>"
obtains q where "(p, q) |\<in>| \<Delta> \<or> (q, p) |\<in>| \<Delta>" using assms
unfolding eps_states_def
by (transfer, auto)+
lemma rule_statesE [elim]:
assumes "q |\<in>| rule_states \<Delta>"
obtains f ps p where "TA_rule f ps p |\<in>| \<Delta>" "q |\<in>| (fset_of_list ps) \<or> q = p" using assms
proof -
assume ass: "(\<And>f ps p. f ps \<rightarrow> p |\<in>| \<Delta> \<Longrightarrow> q |\<in>| fset_of_list ps \<or> q = p \<Longrightarrow> thesis)"
from assms obtain r where "r |\<in>| \<Delta>" "q |\<in>| fset_of_list (r_lhs_states r) \<or> q = r_rhs r"
by (auto simp: rule_states_def)
then show thesis using ass
by (cases r) auto
qed
lemma rule_statesI [intro]:
assumes "r |\<in>| \<Delta>" "q |\<in>| finsert (r_rhs r) (fset_of_list (r_lhs_states r))"
shows "q |\<in>| rule_states \<Delta>" using assms
by (auto simp: rule_states_def)
text \<open>Destruction rule for states\<close>
lemma rule_statesD:
"r |\<in>| (rules \<A>) \<Longrightarrow> r_rhs r |\<in>| \<Q> \<A>" "f qs \<rightarrow> q |\<in>| (rules \<A>) \<Longrightarrow> q |\<in>| \<Q> \<A>"
"r |\<in>| (rules \<A>) \<Longrightarrow> p |\<in>| fset_of_list (r_lhs_states r) \<Longrightarrow> p |\<in>| \<Q> \<A>"
"f qs \<rightarrow> q |\<in>| (rules \<A>) \<Longrightarrow> p |\<in>| fset_of_list qs \<Longrightarrow> p |\<in>| \<Q> \<A>"
by (force simp: \<Q>_def rule_states_def fimage_iff)+
lemma eps_states [simp]: "(eps \<A>) |\<subseteq>| \<Q> \<A> |\<times>| \<Q> \<A>"
unfolding \<Q>_def eps_states_def rule_states_def
- by (auto simp add: rev_fimage_eqI)
+ by (auto simp add: rev_image_eqI)
lemma eps_statesD: "(p, q) |\<in>| (eps \<A>) \<Longrightarrow> p |\<in>| \<Q> \<A> \<and> q |\<in>| \<Q> \<A>"
using eps_states by (auto simp add: \<Q>_def)
lemma eps_trancl_statesD:
"(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> p |\<in>| \<Q> \<A> \<and> q |\<in>| \<Q> \<A>"
by (induct rule: ftrancl_induct) (auto dest: eps_statesD)
lemmas eps_dest_all = eps_statesD eps_trancl_statesD
text \<open>Mapping over function symbols/states\<close>
lemma finite_Collect_ta_rule:
"finite {TA_rule f qs q | f qs q. TA_rule f qs q |\<in>| rules \<A>}" (is "finite ?S")
proof -
have "{f qs \<rightarrow> q |f qs q. f qs \<rightarrow> q |\<in>| rules \<A>} \<subseteq> fset (rules \<A>)"
- by (auto simp flip: fmember_iff_member_fset)
+ by auto
from finite_subset[OF this] show ?thesis by simp
qed
lemma map_ta_rule_finite:
"finite \<Delta> \<Longrightarrow> finite {TA_rule (g h) (map f qs) (f q) | h qs q. TA_rule h qs q \<in> \<Delta>}"
proof (induct rule: finite.induct)
case (insertI A a)
have union: "{TA_rule (g h) (map f qs) (f q) |h qs q. TA_rule h qs q \<in> insert a A} =
{TA_rule (g h) (map f qs) (f q) | h qs q. TA_rule h qs q = a} \<union> {TA_rule (g h) (map f qs) (f q) |h qs q. TA_rule h qs q \<in> A}"
by auto
have "finite {g h map f qs \<rightarrow> f q |h qs q. h qs \<rightarrow> q = a}"
by (cases a) auto
from finite_UnI[OF this insertI(2)] show ?case unfolding union .
qed auto
-lemmas map_ta_rule_fset_finite [simp] = map_ta_rule_finite[of "fset \<Delta>" for \<Delta>, simplified, unfolded fmember_iff_member_fset[symmetric]]
-lemmas map_ta_rule_states_finite [simp] = map_ta_rule_finite[of "fset \<Delta>" id for \<Delta>, simplified, unfolded fmember_iff_member_fset[symmetric]]
-lemmas map_ta_rule_funsym_finite [simp] = map_ta_rule_finite[of "fset \<Delta>" _ id for \<Delta>, simplified, unfolded fmember_iff_member_fset[symmetric]]
+lemmas map_ta_rule_fset_finite [simp] = map_ta_rule_finite[of "fset \<Delta>" for \<Delta>, simplified]
+lemmas map_ta_rule_states_finite [simp] = map_ta_rule_finite[of "fset \<Delta>" id for \<Delta>, simplified]
+lemmas map_ta_rule_funsym_finite [simp] = map_ta_rule_finite[of "fset \<Delta>" _ id for \<Delta>, simplified]
lemma map_ta_rule_comp:
"map_ta_rule f g \<circ> map_ta_rule f' g' = map_ta_rule (f \<circ> f') (g \<circ> g')"
using ta_rule.map_comp[of f g]
by (auto simp: comp_def)
lemma map_ta_rule_cases:
"map_ta_rule f g r = TA_rule (g (r_root r)) (map f (r_lhs_states r)) (f (r_rhs r))"
by (cases r) auto
lemma map_ta_rule_prod_swap_id [simp]:
"map_ta_rule prod.swap prod.swap (map_ta_rule prod.swap prod.swap r) = r"
by (auto simp: map_ta_rule_cases)
lemma rule_states_image [simp]:
"rule_states (map_ta_rule f g |`| \<Delta>) = f |`| rule_states \<Delta>" (is "?Ls = ?Rs")
proof -
{fix q assume "q |\<in>| ?Ls"
then obtain r where "r |\<in>| \<Delta>"
"q |\<in>| finsert (r_rhs (map_ta_rule f g r)) (fset_of_list (r_lhs_states (map_ta_rule f g r)))"
by (auto simp: rule_states_def)
then have "q |\<in>| ?Rs" by (cases r) (force simp: fimage_iff)}
moreover
{fix q assume "q |\<in>| ?Rs"
then obtain r p where "r |\<in>| \<Delta>" "f p = q"
"p |\<in>| finsert (r_rhs r) (fset_of_list (r_lhs_states r))"
by (auto simp: rule_states_def)
then have "q |\<in>| ?Ls" by (cases r) (force simp: fimage_iff)}
ultimately show ?thesis by blast
qed
lemma \<Q>_mono:
"(rules \<A>) |\<subseteq>| (rules \<B>) \<Longrightarrow> (eps \<A>) |\<subseteq>| (eps \<B>) \<Longrightarrow> \<Q> \<A> |\<subseteq>| \<Q> \<B>"
using rule_states_mono eps_states_mono unfolding \<Q>_def
by blast
lemma \<Q>_subseteq_I:
assumes "\<And> r. r |\<in>| rules \<A> \<Longrightarrow> r_rhs r |\<in>| S"
and "\<And> r. r |\<in>| rules \<A> \<Longrightarrow> fset_of_list (r_lhs_states r) |\<subseteq>| S"
and "\<And> e. e |\<in>| eps \<A> \<Longrightarrow> fst e |\<in>| S \<and> snd e |\<in>| S"
shows "\<Q> \<A> |\<subseteq>| S" using assms unfolding \<Q>_def
by (auto simp: rule_states_def) blast
lemma finite_states:
"finite {q. \<exists> f p ps. f ps \<rightarrow> p |\<in>| rules \<A> \<and> (p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|)}" (is "finite ?set")
proof -
have "?set \<subseteq> fset (\<Q> \<A>)"
by (intro subsetI, drule CollectD)
- (metis eps_trancl_statesD notin_fset rule_statesD(2))
+ (metis eps_trancl_statesD rule_statesD(2))
from finite_subset[OF this] show ?thesis by auto
qed
text \<open>Collecting all states reachable from target of rules\<close>
lemma finite_ta_rhs_states [simp]:
"finite {q. \<exists>p. p |\<in>| rule_target_states (rules \<A>) \<and> (p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|)}" (is "finite ?Set")
proof -
have "?Set \<subseteq> fset (\<Q> \<A>)"
by (auto dest: rule_statesD)
- (metis eps_trancl_statesD notin_fset rule_statesD(1))+
+ (metis eps_trancl_statesD rule_statesD(1))+
from finite_subset[OF this] show ?thesis
by auto
qed
text \<open>Computing the signature induced by the rule set of given tree automaton\<close>
lemma ta_sigI [intro]:
"TA_rule f qs q |\<in>| (rules \<A>) \<Longrightarrow> length qs = n \<Longrightarrow> (f, n) |\<in>| ta_sig \<A>" unfolding ta_sig_def
using mk_disjoint_finsert by fastforce
lemma ta_sig_mono:
"(rules \<A>) |\<subseteq>| (rules \<B>) \<Longrightarrow> ta_sig \<A> |\<subseteq>| ta_sig \<B>"
by (auto simp: ta_sig_def)
lemma finite_eps:
"finite {q. \<exists> f ps p. f ps \<rightarrow> p |\<in>| rules \<A> \<and> (p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|)}" (is "finite ?S")
- by (intro finite_subset[OF _ finite_ta_rhs_states[of \<A>]]) auto
+ by (intro finite_subset[OF _ finite_ta_rhs_states[of \<A>]]) (auto intro!: bexI)
lemma collect_snd_trancl_fset:
"{p. (q, p) |\<in>| (eps \<A>)|\<^sup>+|} = fset (snd |`| (ffilter (\<lambda> x. fst x = q) ((eps \<A>)|\<^sup>+|)))"
- by (auto simp: image_iff fmember_iff_member_fset) force
+ by (auto simp: image_iff) force
lemma ta_der_Var:
"q |\<in>| ta_der \<A> (Var x) \<longleftrightarrow> x = q \<or> (x, q) |\<in>| (eps \<A>)|\<^sup>+|"
by (auto simp: collect_snd_trancl_fset)
lemma ta_der_Fun:
"q |\<in>| ta_der \<A> (Fun f ts) \<longleftrightarrow> (\<exists> ps p. TA_rule f ps p |\<in>| (rules \<A>) \<and>
(p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|) \<and> length ps = length ts \<and>
(\<forall> i < length ts. ps ! i |\<in>| ta_der \<A> (ts ! i)))" (is "?Ls \<longleftrightarrow> ?Rs")
unfolding ta_der.simps
by (intro iffI fCollect_memberI finite_Collect_less_eq[OF _ finite_eps[of \<A>]]) auto
declare ta_der.simps[simp del]
declare ta_der.simps[code del]
lemmas ta_der_simps [simp] = ta_der_Var ta_der_Fun
lemma ta_der'_Var:
"Var q |\<in>| ta_der' \<A> (Var x) \<longleftrightarrow> x = q \<or> (x, q) |\<in>| (eps \<A>)|\<^sup>+|"
by (auto simp: collect_snd_trancl_fset)
lemma ta_der'_Fun:
"Var q |\<in>| ta_der' \<A> (Fun f ts) \<longleftrightarrow> q |\<in>| ta_der \<A> (Fun f ts)"
unfolding ta_der'.simps
by (intro iffI funionI1 fCollect_memberI)
(auto simp del: ta_der_Fun ta_der_Var simp: fset_image_conv)
lemma ta_der'_Fun2:
"Fun f ps |\<in>| ta_der' \<A> (Fun g ts) \<longleftrightarrow> f = g \<and> length ps = length ts \<and> (\<forall>i<length ts. ps ! i |\<in>| ta_der' \<A> (ts ! i))"
proof -
have f: "finite {ss. set ss \<subseteq> fset ( |\<Union>| (fset_of_list (map (ta_der' \<A>) ts))) \<and> length ss = length ts}"
by (intro finite_lists_length_eq) auto
have "finite {ss. length ss = length ts \<and> (\<forall>i<length ts. ss ! i |\<in>| ta_der' \<A> (ts ! i))}"
by (intro finite_subset[OF _ f])
- (force simp: in_fset_conv_nth simp flip: fset_of_list_elem fmember_iff_member_fset)
+ (force simp: in_fset_conv_nth simp flip: fset_of_list_elem)
then show ?thesis unfolding ta_der'.simps
by (intro iffI funionI2 fCollect_memberI)
(auto simp del: ta_der_Fun ta_der_Var)
qed
declare ta_der'.simps[simp del]
declare ta_der'.simps[code del]
lemmas ta_der'_simps [simp] = ta_der'_Var ta_der'_Fun ta_der'_Fun2
text \<open>Induction schemes for the most used cases\<close>
lemma ta_der_induct[consumes 1, case_names Var Fun]:
assumes reach: "q |\<in>| ta_der \<A> t"
and VarI: "\<And> q v. v = q \<or> (v, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> P (Var v) q"
and FunI: "\<And>f ts ps p q. f ps \<rightarrow> p |\<in>| rules \<A> \<Longrightarrow> length ts = length ps \<Longrightarrow> p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow>
(\<And>i. i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der \<A> (ts ! i)) \<Longrightarrow>
(\<And>i. i < length ts \<Longrightarrow> P (ts ! i) (ps ! i)) \<Longrightarrow> P (Fun f ts) q"
shows "P t q" using assms(1)
by (induct t arbitrary: q) (auto simp: VarI FunI)
lemma ta_der_gterm_induct[consumes 1, case_names GFun]:
assumes reach: "q |\<in>| ta_der \<A> (term_of_gterm t)"
and Fun: "\<And>f ts ps p q. TA_rule f ps p |\<in>| rules \<A> \<Longrightarrow> length ts = length ps \<Longrightarrow> p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow>
(\<And>i. i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der \<A> (term_of_gterm (ts ! i))) \<Longrightarrow>
(\<And>i. i < length ts \<Longrightarrow> P (ts ! i) (ps ! i)) \<Longrightarrow> P (GFun f ts) q"
shows "P t q" using assms(1)
by (induct t arbitrary: q) (auto simp: Fun)
lemma ta_der_rule_empty:
assumes "q |\<in>| ta_der (TA {||} \<Delta>\<^sub>\<epsilon>) t"
obtains p where "t = Var p" "p = q \<or> (p, q) |\<in>| \<Delta>\<^sub>\<epsilon>|\<^sup>+|"
using assms by (cases t) auto
lemma ta_der_eps:
assumes "(p, q) |\<in>| (eps \<A>)" and "p |\<in>| ta_der \<A> t"
shows "q |\<in>| ta_der \<A> t" using assms
by (cases t) (auto intro: ftrancl_into_trancl)
lemma ta_der_trancl_eps:
assumes "(p, q) |\<in>| (eps \<A>)|\<^sup>+|" and "p |\<in>| ta_der \<A> t"
shows "q |\<in>| ta_der \<A> t" using assms
by (induct rule: ftrancl_induct) (auto intro: ftrancl_into_trancl ta_der_eps)
lemma ta_der_mono:
"(rules \<A>) |\<subseteq>| (rules \<B>) \<Longrightarrow> (eps \<A>) |\<subseteq>| (eps \<B>) \<Longrightarrow> ta_der \<A> t |\<subseteq>| ta_der \<B> t"
proof (induct t)
case (Var x) then show ?case
by (auto dest: ftrancl_mono[of _ "eps \<A>" "eps \<B>"])
next
case (Fun f ts)
show ?case using Fun(1)[OF nth_mem Fun(2, 3)]
by (auto dest!: fsubsetD[OF Fun(2)] ftrancl_mono[OF _ Fun(3)]) blast+
qed
lemma ta_der_el_mono:
"(rules \<A>) |\<subseteq>| (rules \<B>) \<Longrightarrow> (eps \<A>) |\<subseteq>| (eps \<B>) \<Longrightarrow> q |\<in>| ta_der \<A> t \<Longrightarrow> q |\<in>| ta_der \<B> t"
using ta_der_mono by blast
lemma ta_der'_ta_der:
assumes "t |\<in>| ta_der' \<A> s" "p |\<in>| ta_der \<A> t"
shows "p |\<in>| ta_der \<A> s" using assms
proof (induction arbitrary: p t rule: ta_der'.induct)
case (2 \<A> f ts) show ?case using 2(2-)
proof (induction t)
case (Var x) then show ?case
by auto (meson ftrancl_trans)
next
case (Fun g ss)
have ss_props: "g = f" "length ss = length ts" "\<forall>i < length ts. ss ! i |\<in>| ta_der' \<A> (ts ! i)"
using Fun(2) by auto
then show ?thesis using Fun(1)[OF nth_mem] Fun(2-)
by (auto simp: ss_props)
(metis (no_types, lifting) "2.IH" ss_props(3))+
qed
qed (auto dest: ftrancl_trans simp: ta_der'.simps)
lemma ta_der'_empty:
assumes "t |\<in>| ta_der' (TA {||} {||}) s"
shows "t = s" using assms
by (induct s arbitrary: t) (auto simp add: ta_der'.simps nth_equalityI)
lemma ta_der'_to_ta_der:
"Var q |\<in>| ta_der' \<A> s \<Longrightarrow> q |\<in>| ta_der \<A> s"
using ta_der'_ta_der by fastforce
lemma ta_der_to_ta_der':
"q |\<in>| ta_der \<A> s \<longleftrightarrow> Var q |\<in>| ta_der' \<A> s "
by (induct s arbitrary: q) auto
lemma ta_der'_poss:
assumes "t |\<in>| ta_der' \<A> s"
shows "poss t \<subseteq> poss s" using assms
proof (induct s arbitrary: t)
case (Fun f ts)
show ?case using Fun(2) Fun(1)[OF nth_mem, of i "args t ! i" for i]
by (cases t) auto
qed (auto simp: ta_der'.simps)
lemma ta_der'_refl[simp]: "t |\<in>| ta_der' \<A> t"
by (induction t) fastforce+
lemma ta_der'_eps:
assumes "Var p |\<in>| ta_der' \<A> s" and "(p, q) |\<in>| (eps \<A>)|\<^sup>+|"
shows "Var q |\<in>| ta_der' \<A> s" using assms
by (cases s, auto dest: ftrancl_trans) (meson ftrancl_trans)
lemma ta_der'_trans:
assumes "t |\<in>| ta_der' \<A> s" and "u |\<in>| ta_der' \<A> t"
shows "u |\<in>| ta_der' \<A> s" using assms
proof (induct t arbitrary: u s)
case (Fun f ts) note IS = Fun(2-) note IH = Fun(1)[OF nth_mem, of i "args s ! i" for i]
show ?case
proof (cases s)
case (Var x1)
then show ?thesis using IS by (auto simp: ta_der'.simps)
next
case [simp]: (Fun g ss)
show ?thesis using IS IH
by (cases u, auto) (metis ta_der_to_ta_der')+
qed
qed (auto simp: ta_der'.simps ta_der'_eps)
text \<open>Connecting contexts to derivation definition\<close>
lemma ta_der_ctxt:
assumes p: "p |\<in>| ta_der \<A> t" "q |\<in>| ta_der \<A> C\<langle>Var p\<rangle>"
shows "q |\<in>| ta_der \<A> C\<langle>t\<rangle>" using assms(2)
proof (induct C arbitrary: q)
case Hole then show ?case using assms
by (auto simp: ta_der_trancl_eps)
next
case (More f ss C ts)
from More(2) obtain qs r where
rule: "f qs \<rightarrow> r |\<in>| rules \<A>" "length qs = Suc (length ss + length ts)" and
reach: "\<forall> i < Suc (length ss + length ts). qs ! i |\<in>| ta_der \<A> ((ss @ C\<langle>Var p\<rangle> # ts) ! i)" "r = q \<or> (r, q) |\<in>| (eps \<A>)|\<^sup>+|"
by auto
have "i < Suc (length ss + length ts) \<Longrightarrow> qs ! i |\<in>| ta_der \<A> ((ss @ C\<langle>t\<rangle> # ts) ! i)" for i
using More(1)[of "qs ! length ss"] assms rule(2) reach(1)
unfolding nth_append_Cons by presburger
then show ?case using rule reach(2) by auto
qed
lemma ta_der_eps_ctxt:
assumes "p |\<in>| ta_der A C\<langle>Var q'\<rangle>" and "(q, q') |\<in>| (eps A)|\<^sup>+|"
shows "p |\<in>| ta_der A C\<langle>Var q\<rangle>"
using assms by (meson ta_der_Var ta_der_ctxt)
lemma rule_reachable_ctxt_exist:
assumes rule: "f qs \<rightarrow> q |\<in>| rules \<A>" and "i < length qs"
shows "\<exists> C. q |\<in>| ta_der \<A> (C \<langle>Var (qs ! i)\<rangle>)" using assms
by (intro exI[of _ "More f (map Var (take i qs)) \<box> (map Var (drop (Suc i) qs))"])
(auto simp: min_def nth_append_Cons intro!: exI[of _ q] exI[of _ qs])
lemma ta_der_ctxt_decompose:
assumes "q |\<in>| ta_der \<A> C\<langle>t\<rangle>"
shows "\<exists> p . p |\<in>| ta_der \<A> t \<and> q |\<in>| ta_der \<A> C\<langle>Var p\<rangle>" using assms
proof (induct C arbitrary: q)
case (More f ss C ts)
from More(2) obtain qs r where
rule: "f qs \<rightarrow> r |\<in>| rules \<A>" "length qs = Suc (length ss + length ts)" and
reach: "\<forall> i < Suc (length ss + length ts). qs ! i |\<in>| ta_der \<A> ((ss @ C\<langle>t\<rangle> # ts) ! i)"
"r = q \<or> (r, q) |\<in>| (eps \<A>)|\<^sup>+|"
by auto
obtain p where p: "p |\<in>| ta_der \<A> t" "qs ! length ss |\<in>| ta_der \<A> C\<langle>Var p\<rangle>"
using More(1)[of "qs ! length ss"] reach(1) rule(2)
by (metis less_add_Suc1 nth_append_length)
have "i < Suc (length ss + length ts) \<Longrightarrow> qs ! i |\<in>| ta_der \<A> ((ss @ C\<langle>Var p\<rangle> # ts) ! i)" for i
using reach rule(2) p by (auto simp: p(2) nth_append_Cons)
then have "q |\<in>| ta_der \<A> (More f ss C ts)\<langle>Var p\<rangle>" using rule reach
by auto
then show ?case using p(1) by (intro exI[of _ p]) blast
qed auto
\<comment> \<open>Relation between reachable states and states of a tree automaton\<close>
lemma ta_der_states:
"ta_der \<A> t |\<subseteq>| \<Q> \<A> |\<union>| fvars_term t"
proof (induct t)
case (Var x) then show ?case
- by (auto simp: eq_onp_same_args fmember.abs_eq)
+ by (auto simp: eq_onp_same_args)
(metis eps_trancl_statesD)
case (Fun f ts) then show ?case
by (auto simp: rule_statesD(2) eps_trancl_statesD)
qed
lemma ground_ta_der_states:
"ground t \<Longrightarrow> ta_der \<A> t |\<subseteq>| \<Q> \<A>"
using ta_der_states[of \<A> t] by auto
lemmas ground_ta_der_statesD = fsubsetD[OF ground_ta_der_states]
lemma gterm_ta_der_states [simp]:
"q |\<in>| ta_der \<A> (term_of_gterm t) \<Longrightarrow> q |\<in>| \<Q> \<A>"
by (intro ground_ta_der_states[THEN fsubsetD, of "term_of_gterm t"]) simp
lemma ta_der_states':
"q |\<in>| ta_der \<A> t \<Longrightarrow> q |\<in>| \<Q> \<A> \<Longrightarrow> fvars_term t |\<subseteq>| \<Q> \<A>"
proof (induct rule: ta_der_induct)
case (Fun f ts ps p r)
then have "i < length ts \<Longrightarrow> fvars_term (ts ! i) |\<subseteq>| \<Q> \<A>" for i
by (auto simp: in_fset_conv_nth dest!: rule_statesD(3))
then show ?case by (force simp: in_fset_conv_nth)
qed (auto simp: eps_trancl_statesD)
lemma ta_der_not_stateD:
"q |\<in>| ta_der \<A> t \<Longrightarrow> q |\<notin>| \<Q> \<A> \<Longrightarrow> t = Var q"
using fsubsetD[OF ta_der_states, of q \<A> t]
by (cases t) (auto dest: rule_statesD eps_trancl_statesD)
lemma ta_der_is_fun_stateD:
"is_Fun t \<Longrightarrow> q |\<in>| ta_der \<A> t \<Longrightarrow> q |\<in>| \<Q> \<A>"
using ta_der_not_stateD[of q \<A> t]
by (cases t) auto
lemma ta_der_is_fun_fvars_stateD:
"is_Fun t \<Longrightarrow> q |\<in>| ta_der \<A> t \<Longrightarrow> fvars_term t |\<subseteq>| \<Q> \<A>"
using ta_der_is_fun_stateD[of t q \<A>]
using ta_der_states'[of q \<A> t]
by (cases t) auto
lemma ta_der_not_reach:
assumes "\<And> r. r |\<in>| rules \<A> \<Longrightarrow> r_rhs r \<noteq> q"
and "\<And> e. e |\<in>| eps \<A> \<Longrightarrow> snd e \<noteq> q"
shows "q |\<notin>| ta_der \<A> (term_of_gterm t)" using assms
by (cases t) (fastforce dest!: assms(1) ftranclD2[of _ q])
lemma ta_rhs_states_subset_states: "ta_rhs_states \<A> |\<subseteq>| \<Q> \<A>"
by (auto simp: ta_rhs_states_def dest: rtranclD rule_statesD eps_trancl_statesD)
(* a resulting state is always some rhs of a rule (or epsilon transition) *)
lemma ta_rhs_states_res: assumes "is_Fun t"
shows "ta_der \<A> t |\<subseteq>| ta_rhs_states \<A>"
proof
fix q assume q: "q |\<in>| ta_der \<A> t"
from \<open>is_Fun t\<close> obtain f ts where t: "t = Fun f ts" by (cases t, auto)
from q[unfolded t] obtain q' qs where "TA_rule f qs q' |\<in>| rules \<A>"
and q: "q' = q \<or> (q', q) |\<in>| (eps \<A>)|\<^sup>+|" by auto
then show "q |\<in>| ta_rhs_states \<A>" unfolding ta_rhs_states_def
- by auto
+ by (auto intro!: bexI)
qed
text \<open>Reachable states of ground terms are preserved over the @{const adapt_vars} function\<close>
lemma ta_der_adapt_vars_ground [simp]:
"ground t \<Longrightarrow> ta_der A (adapt_vars t) = ta_der A t"
by (induct t) auto
lemma gterm_of_term_inv':
"ground t \<Longrightarrow> term_of_gterm (gterm_of_term t) = adapt_vars t"
by (induct t) (auto 0 0 intro!: nth_equalityI)
lemma map_vars_term_term_of_gterm:
"map_vars_term f (term_of_gterm t) = term_of_gterm t"
by (induct t) auto
lemma adapt_vars_term_of_gterm:
"adapt_vars (term_of_gterm t) = term_of_gterm t"
by (induct t) auto
(* a term can be reduced to a state, only if all symbols appear in the automaton *)
lemma ta_der_term_sig:
"q |\<in>| ta_der \<A> t \<Longrightarrow> ffunas_term t |\<subseteq>| ta_sig \<A>"
proof (induct rule: ta_der_induct)
case (Fun f ts ps p q)
show ?case using Fun(1 - 4) Fun(5)[THEN fsubsetD]
by (auto simp: in_fset_conv_nth)
qed auto
lemma ta_der_gterm_sig:
"q |\<in>| ta_der \<A> (term_of_gterm t) \<Longrightarrow> ffunas_gterm t |\<subseteq>| ta_sig \<A>"
using ta_der_term_sig ffunas_term_of_gterm_conv
by fastforce
text \<open>@{const ta_lang} for terms with arbitrary variable type\<close>
lemma ta_langE: assumes "t \<in> ta_lang Q \<A>"
obtains t' q where "ground t'" "q |\<in>| Q" "q |\<in>| ta_der \<A> t'" "t = adapt_vars t'"
using assms unfolding ta_lang_def by blast
lemma ta_langI: assumes "ground t'" "q |\<in>| Q" "q |\<in>| ta_der \<A> t'" "t = adapt_vars t'"
shows "t \<in> ta_lang Q \<A>"
using assms unfolding ta_lang_def by blast
lemma ta_lang_def2: "(ta_lang Q (\<A> :: ('q,'f)ta) :: ('f,'v)terms) = {t. ground t \<and> Q |\<inter>| ta_der \<A> (adapt_vars t) \<noteq> {||}}"
by (auto elim!: ta_langE) (metis adapt_vars_adapt_vars ground_adapt_vars ta_langI)
text \<open>@{const ta_lang} for @{const gterms}\<close>
lemma ta_lang_to_gta_lang [simp]:
"ta_lang Q \<A> = term_of_gterm ` gta_lang Q \<A>" (is "?Ls = ?Rs")
proof -
{fix t assume "t \<in> ?Ls"
from ta_langE[OF this] obtain q t' where "ground t'" "q |\<in>| Q" "q |\<in>| ta_der \<A> t'" "t = adapt_vars t'"
by blast
then have "t \<in> ?Rs" unfolding gta_lang_def gta_der_def
by (auto simp: image_iff gterm_of_term_inv' intro!: exI[of _ "gterm_of_term t'"])}
moreover
{fix t assume "t \<in> ?Rs" then have "t \<in> ?Ls"
using ta_langI[OF ground_term_of_gterm _ _ gterm_of_term_inv'[OF ground_term_of_gterm]]
by (force simp: gta_lang_def gta_der_def)}
ultimately show ?thesis by blast
qed
lemma term_of_gterm_in_ta_lang_conv:
"term_of_gterm t \<in> ta_lang Q \<A> \<longleftrightarrow> t \<in> gta_lang Q \<A>"
by (metis (mono_tags, lifting) image_iff ta_lang_to_gta_lang term_of_gterm_inv)
lemma gta_lang_def_sym:
"gterm_of_term ` ta_lang Q \<A> = gta_lang Q \<A>"
(* this is nontrivial because the lhs has a more general type than the rhs of gta_lang_def *)
unfolding gta_lang_def image_def
by (intro Collect_cong) (simp add: gta_lang_def)
lemma gta_langI [intro]:
assumes "q |\<in>| Q" and "q |\<in>| ta_der \<A> (term_of_gterm t)"
shows "t \<in> gta_lang Q \<A>" using assms
by (metis adapt_vars_term_of_gterm ground_term_of_gterm ta_langI term_of_gterm_in_ta_lang_conv)
lemma gta_langE [elim]:
assumes "t \<in> gta_lang Q \<A>"
obtains q where "q |\<in>| Q" and "q |\<in>| ta_der \<A> (term_of_gterm t)" using assms
by (metis adapt_vars_adapt_vars adapt_vars_term_of_gterm ta_langE term_of_gterm_in_ta_lang_conv)
lemma gta_lang_mono:
assumes "\<And> t. ta_der \<A> t |\<subseteq>| ta_der \<BB> t" and "Q\<^sub>\<A> |\<subseteq>| Q\<^sub>\<BB>"
shows "gta_lang Q\<^sub>\<A> \<A> \<subseteq> gta_lang Q\<^sub>\<BB> \<BB>"
using assms by (auto elim!: gta_langE intro!: gta_langI)
lemma gta_lang_term_of_gterm [simp]:
"term_of_gterm t \<in> term_of_gterm ` gta_lang Q \<A> \<longleftrightarrow> t \<in> gta_lang Q \<A>"
by (auto elim!: gta_langE intro!: gta_langI) (metis term_of_gterm_inv)
(* terms can be accepted, only if all their symbols appear in the automaton *)
lemma gta_lang_subset_rules_funas:
"gta_lang Q \<A> \<subseteq> \<T>\<^sub>G (fset (ta_sig \<A>))"
using ta_der_gterm_sig[THEN fsubsetD]
- by (force simp: \<T>\<^sub>G_equivalent_def simp flip: fmember_iff_member_fset ffunas_gterm.rep_eq)
+ by (force simp: \<T>\<^sub>G_equivalent_def ffunas_gterm.rep_eq)
lemma reg_funas:
"\<L> \<A> \<subseteq> \<T>\<^sub>G (fset (ta_sig (ta \<A>)))" using gta_lang_subset_rules_funas
by (auto simp: \<L>_def)
lemma ta_syms_lang: "t \<in> ta_lang Q \<A> \<Longrightarrow> ffunas_term t |\<subseteq>| ta_sig \<A>"
using gta_lang_subset_rules_funas ffunas_gterm_gterm_of_term ta_der_gterm_sig ta_lang_def2
by fastforce
lemma gta_lang_Rest_states_conv:
"gta_lang Q \<A> = gta_lang (Q |\<inter>| \<Q> \<A>) \<A>"
by (auto elim!: gta_langE)
lemma reg_Rest_fin_states [simp]:
"\<L> (reg_Restr_Q\<^sub>f \<A>) = \<L> \<A>"
using gta_lang_Rest_states_conv
by (auto simp: \<L>_def reg_Restr_Q\<^sub>f_def)
text \<open>Deterministic tree automatons\<close>
definition ta_det :: "('q,'f) ta \<Rightarrow> bool" where
"ta_det \<A> \<longleftrightarrow> eps \<A> = {||} \<and>
(\<forall> f qs q q'. TA_rule f qs q |\<in>| rules \<A> \<longrightarrow> TA_rule f qs q' |\<in>| rules \<A> \<longrightarrow> q = q')"
definition "ta_subset \<A> \<B> \<longleftrightarrow> rules \<A> |\<subseteq>| rules \<B> \<and> eps \<A> |\<subseteq>| eps \<B>"
(* determinism implies unique results *)
lemma ta_detE[elim, consumes 1]: assumes det: "ta_det \<A>"
shows "q |\<in>| ta_der \<A> t \<Longrightarrow> q' |\<in>| ta_der \<A> t \<Longrightarrow> q = q'" using assms
by (induct t arbitrary: q q') (auto simp: ta_det_def, metis nth_equalityI nth_mem)
lemma ta_subset_states: "ta_subset \<A> \<B> \<Longrightarrow> \<Q> \<A> |\<subseteq>| \<Q> \<B>"
using \<Q>_mono by (auto simp: ta_subset_def)
lemma ta_subset_refl[simp]: "ta_subset \<A> \<A>"
unfolding ta_subset_def by auto
lemma ta_subset_trans: "ta_subset \<A> \<B> \<Longrightarrow> ta_subset \<B> \<CC> \<Longrightarrow> ta_subset \<A> \<CC>"
unfolding ta_subset_def by auto
lemma ta_subset_det: "ta_subset \<A> \<B> \<Longrightarrow> ta_det \<B> \<Longrightarrow> ta_det \<A>"
unfolding ta_det_def ta_subset_def by blast
lemma ta_der_mono': "ta_subset \<A> \<B> \<Longrightarrow> ta_der \<A> t |\<subseteq>| ta_der \<B> t"
using ta_der_mono unfolding ta_subset_def by auto
lemma ta_lang_mono': "ta_subset \<A> \<B> \<Longrightarrow> Q\<^sub>\<A> |\<subseteq>| Q\<^sub>\<B> \<Longrightarrow> ta_lang Q\<^sub>\<A> \<A> \<subseteq> ta_lang Q\<^sub>\<B> \<B>"
using gta_lang_mono[of \<A> \<B>] ta_der_mono'[of \<A> \<B>]
by auto blast
(* the restriction of an automaton to a given set of states *)
lemma ta_restrict_subset: "ta_subset (ta_restrict \<A> Q) \<A>"
unfolding ta_subset_def ta_restrict_def
by auto
lemma ta_restrict_states_Q: "\<Q> (ta_restrict \<A> Q) |\<subseteq>| Q"
by (auto simp: \<Q>_def ta_restrict_def rule_states_def eps_states_def dest!: fsubsetD)
lemma ta_restrict_states: "\<Q> (ta_restrict \<A> Q) |\<subseteq>| \<Q> \<A>"
using ta_subset_states[OF ta_restrict_subset] by fastforce
lemma ta_restrict_states_eq_imp_eq [simp]:
assumes eq: "\<Q> (ta_restrict \<A> Q) = \<Q> \<A>"
shows "ta_restrict \<A> Q = \<A>" using assms
apply (auto simp: ta_restrict_def
intro!: ta.expand finite_subset[OF _ finite_Collect_ta_rule, of _ \<A>])
apply (metis (no_types, lifting) eq fsubsetD fsubsetI rule_statesD(1) rule_statesD(4) ta_restrict_states_Q ta_rule.collapse)
apply (metis eps_statesD eq fin_mono ta_restrict_states_Q)
by (metis eps_statesD eq fsubsetD ta_restrict_states_Q)
lemma ta_der_ta_derict_states:
"fvars_term t |\<subseteq>| Q \<Longrightarrow> q |\<in>| ta_der (ta_restrict \<A> Q) t \<Longrightarrow> q |\<in>| Q"
by (induct t arbitrary: q) (auto simp: ta_restrict_def elim: ftranclE)
lemma ta_derict_ruleI [intro]:
"TA_rule f qs q |\<in>| rules \<A> \<Longrightarrow> fset_of_list qs |\<subseteq>| Q \<Longrightarrow> q |\<in>| Q \<Longrightarrow> TA_rule f qs q |\<in>| rules (ta_restrict \<A> Q)"
by (auto simp: ta_restrict_def intro!: ta.expand finite_subset[OF _ finite_Collect_ta_rule, of _ \<A>])
text \<open>Reachable and productive states: There always is a trim automaton\<close>
lemma finite_ta_reachable [simp]:
"finite {q. \<exists>t. ground t \<and> q |\<in>| ta_der \<A> t}"
proof -
have "{q. \<exists>t. ground t \<and> q |\<in>| ta_der \<A> t} \<subseteq> fset (\<Q> \<A>)"
using ground_ta_der_states[of _ \<A>]
- by auto (metis fsubsetD notin_fset)
+ by auto
from finite_subset[OF this] show ?thesis by auto
qed
lemma ta_reachable_states:
"ta_reachable \<A> |\<subseteq>| \<Q> \<A>"
unfolding ta_reachable_def using ground_ta_der_states
by force
lemma ta_reachableE:
assumes "q |\<in>| ta_reachable \<A>"
obtains t where "ground t" "q |\<in>| ta_der \<A> t"
using assms[unfolded ta_reachable_def] by auto
lemma ta_reachable_gtermE [elim]:
assumes "q |\<in>| ta_reachable \<A>"
obtains t where "q |\<in>| ta_der \<A> (term_of_gterm t)"
using ta_reachableE[OF assms]
by (metis ground_term_to_gtermD)
lemma ta_reachableI [intro]:
assumes "ground t" and "q |\<in>| ta_der \<A> t"
shows "q |\<in>| ta_reachable \<A>"
using assms finite_ta_reachable
by (auto simp: ta_reachable_def)
lemma ta_reachable_gtermI [intro]:
"q |\<in>| ta_der \<A> (term_of_gterm t) \<Longrightarrow> q |\<in>| ta_reachable \<A>"
by (intro ta_reachableI[of "term_of_gterm t"]) simp
lemma ta_reachableI_rule:
assumes sub: "fset_of_list qs |\<subseteq>| ta_reachable \<A>"
and rule: "TA_rule f qs q |\<in>| rules \<A>"
shows "q |\<in>| ta_reachable \<A>"
"\<exists> ts. length qs = length ts \<and> (\<forall> i < length ts. ground (ts ! i)) \<and>
(\<forall> i < length ts. qs ! i |\<in>| ta_der \<A> (ts ! i))" (is "?G")
proof -
{
fix i
assume i: "i < length qs"
then have "qs ! i |\<in>| fset_of_list qs" by auto
with sub have "qs ! i |\<in>| ta_reachable \<A>" by auto
from ta_reachableE[OF this] have "\<exists> t. ground t \<and> qs ! i |\<in>| ta_der \<A> t" by auto
}
then have "\<forall> i. \<exists> t. i < length qs \<longrightarrow> ground t \<and> qs ! i |\<in>| ta_der \<A> t" by auto
from choice[OF this] obtain ts where ts: "\<And> i. i < length qs \<Longrightarrow> ground (ts i) \<and> qs ! i |\<in>| ta_der \<A> (ts i)" by blast
let ?t = "Fun f (map ts [0 ..< length qs])"
have gt: "ground ?t" using ts by auto
have r: "q |\<in>| ta_der \<A> ?t" unfolding ta_der_Fun using rule ts
by (intro exI[of _ qs] exI[of _ q]) simp
with gt show "q |\<in>| ta_reachable \<A>" by blast
from gt ts show ?G by (intro exI[of _ "map ts [0..<length qs]"]) simp
qed
lemma ta_reachable_rule_gtermE:
assumes "\<Q> \<A> |\<subseteq>| ta_reachable \<A>"
and "TA_rule f qs q |\<in>| rules \<A>"
obtains t where "groot t = (f, length qs)" "q |\<in>| ta_der \<A> (term_of_gterm t)"
proof -
assume *: "\<And>t. groot t = (f, length qs) \<Longrightarrow> q |\<in>| ta_der \<A> (term_of_gterm t) \<Longrightarrow> thesis"
from assms have "fset_of_list qs |\<subseteq>| ta_reachable \<A>"
by (auto dest: rule_statesD(3))
from ta_reachableI_rule[OF this assms(2)] obtain ts where args: "length qs = length ts"
"\<forall> i < length ts. ground (ts ! i)" "\<forall> i < length ts. qs ! i |\<in>| ta_der \<A> (ts ! i)"
using assms by force
then show ?thesis using assms(2)
by (intro *[of "GFun f (map gterm_of_term ts)"]) auto
qed
lemma ta_reachableI_eps':
assumes reach: "q |\<in>| ta_reachable \<A>"
and eps: "(q, q') |\<in>| (eps \<A>)|\<^sup>+|"
shows "q' |\<in>| ta_reachable \<A>"
proof -
from ta_reachableE[OF reach] obtain t where g: "ground t" and res: "q |\<in>| ta_der \<A> t" by auto
from ta_der_trancl_eps[OF eps res] g show ?thesis by blast
qed
lemma ta_reachableI_eps:
assumes reach: "q |\<in>| ta_reachable \<A>"
and eps: "(q, q') |\<in>| eps \<A>"
shows "q' |\<in>| ta_reachable \<A>"
by (rule ta_reachableI_eps'[OF reach], insert eps, auto)
\<comment> \<open>Automata are productive on a set P if all states can reach a state in P\<close>
lemma finite_ta_productive:
"finite {p. \<exists>q q' C. p = q \<and> q' |\<in>| ta_der \<A> C\<langle>Var q\<rangle> \<and> q' |\<in>| P}"
proof -
{fix x q C assume ass: "x \<notin> fset P" "q |\<in>| P" "q |\<in>| ta_der \<A> C\<langle>Var x\<rangle>"
then have "x \<in> fset (\<Q> \<A>)"
proof (cases "is_Fun C\<langle>Var x\<rangle>")
case True
then show ?thesis using ta_der_is_fun_fvars_stateD[OF _ ass(3)]
- by auto (metis notin_fset)
+ by auto
next
case False
then show ?thesis using ass
- by (cases C, auto, (metis eps_trancl_statesD notin_fset)+)
+ by (cases C, auto, (metis eps_trancl_statesD)+)
qed}
then have "{q | q q' C. q' |\<in>| ta_der \<A> (C\<langle>Var q\<rangle>) \<and> q' |\<in>| P} \<subseteq> fset (\<Q> \<A>) \<union> fset P" by auto
from finite_subset[OF this] show ?thesis by auto
qed
lemma ta_productiveE: assumes "q |\<in>| ta_productive P \<A>"
obtains q' C where "q' |\<in>| ta_der \<A> (C\<langle>Var q\<rangle>)" "q' |\<in>| P"
using assms[unfolded ta_productive_def] by auto
lemma ta_productiveI:
assumes "q' |\<in>| ta_der \<A> (C\<langle>Var q\<rangle>)" "q' |\<in>| P"
shows "q |\<in>| ta_productive P \<A>"
using assms unfolding ta_productive_def
using finite_ta_productive
by auto
lemma ta_productiveI':
assumes "q |\<in>| ta_der \<A> (C\<langle>Var p\<rangle>)" "q |\<in>| ta_productive P \<A>"
shows "p |\<in>| ta_productive P \<A>"
using assms unfolding ta_productive_def
by auto (metis (mono_tags, lifting) ctxt_ctxt_compose ta_der_ctxt)
lemma ta_productive_setI:
"q |\<in>| P \<Longrightarrow> q |\<in>| ta_productive P \<A>"
using ta_productiveI[of q \<A> \<box> q]
by simp
lemma ta_reachable_empty_rules [simp]:
"rules \<A> = {||} \<Longrightarrow> ta_reachable \<A> = {||}"
by (auto simp: ta_reachable_def)
(metis ground.simps(1) ta.exhaust_sel ta_der_rule_empty)
lemma ta_reachable_mono:
"ta_subset \<A> \<B> \<Longrightarrow> ta_reachable \<A> |\<subseteq>| ta_reachable \<B>" using ta_der_mono'
by (auto simp: ta_reachable_def) blast
lemma ta_reachabe_rhs_states:
"ta_reachable \<A> |\<subseteq>| ta_rhs_states \<A>"
proof -
{fix q assume "q |\<in>| ta_reachable \<A>"
then obtain t where "ground t" "q |\<in>| ta_der \<A> t"
by (auto simp: ta_reachable_def)
then have "q |\<in>| ta_rhs_states \<A>"
- by (cases t) (auto simp: ta_rhs_states_def)}
+ by (cases t) (auto simp: ta_rhs_states_def intro!: bexI)}
then show ?thesis by blast
qed
lemma ta_reachable_eps:
"(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> p |\<in>| ta_reachable \<A> \<Longrightarrow> (p, q) |\<in>| (fRestr (eps \<A>) (ta_reachable \<A>))|\<^sup>+|"
proof (induct rule: ftrancl_induct)
case (Base a b)
then show ?case
by (metis fSigmaI finterI fr_into_trancl ta_reachableI_eps)
next
case (Step p q r)
then have "q |\<in>| ta_reachable \<A>" "r |\<in>| ta_reachable \<A>"
by (metis ta_reachableI_eps ta_reachableI_eps')+
then show ?case using Step
by (metis fSigmaI finterI ftrancl_into_trancl)
qed
(* major lemma to show that one can restrict to reachable states *)
lemma ta_der_only_reach:
assumes "fvars_term t |\<subseteq>| ta_reachable \<A>"
shows "ta_der \<A> t = ta_der (ta_only_reach \<A>) t" (is "?LS = ?RS")
proof -
have "?RS |\<subseteq>| ?LS" using ta_der_mono'[OF ta_restrict_subset]
by fastforce
moreover
{fix q assume "q |\<in>| ?LS"
then have "q |\<in>| ?RS" using assms
proof (induct rule: ta_der_induct)
case (Fun f ts ps p q)
from Fun(2, 6) have ta_reach [simp]: "i < length ps \<Longrightarrow> fvars_term (ts ! i) |\<subseteq>| ta_reachable \<A>" for i
by auto (metis ffUnionI fimage_fset fnth_mem funionI2 length_map nth_map sup.orderE)
from Fun have r: "i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der (ta_only_reach \<A>) (ts ! i)"
"i < length ts \<Longrightarrow> ps ! i |\<in>| ta_reachable \<A>" for i
by (auto) (metis ta_reach ta_der_ta_derict_states)+
then have "f ps \<rightarrow> p |\<in>| rules (ta_only_reach \<A>)"
using Fun(1, 2)
by (intro ta_derict_ruleI)
(fastforce simp: in_fset_conv_nth intro!: ta_reachableI_rule[OF _ Fun(1)])+
then show ?case using ta_reachable_eps[of p q] ta_reachableI_rule[OF _ Fun(1)] r Fun(2, 3)
by (auto simp: ta_restrict_def intro!: exI[of _ p] exI[of _ ps])
qed (auto simp: ta_restrict_def intro: ta_reachable_eps)}
ultimately show ?thesis by blast
qed
lemma ta_der_gterm_only_reach:
"ta_der \<A> (term_of_gterm t) = ta_der (ta_only_reach \<A>) (term_of_gterm t)"
using ta_der_only_reach[of "term_of_gterm t" \<A>]
by simp
lemma ta_reachable_ta_only_reach [simp]:
"ta_reachable (ta_only_reach \<A>) = ta_reachable \<A>" (is "?LS = ?RS")
proof -
have "?LS |\<subseteq>| ?RS" using ta_der_mono'[OF ta_restrict_subset]
by (auto simp: ta_reachable_def) fastforce
moreover
{fix t assume "ground (t :: ('b, 'a) term)"
then have "ta_der \<A> t = ta_der (ta_only_reach \<A>) t" using ta_der_only_reach[of t \<A>]
by simp}
ultimately show ?thesis unfolding ta_reachable_def
by auto
qed
lemma ta_only_reach_reachable:
"\<Q> (ta_only_reach \<A>) |\<subseteq>| ta_reachable (ta_only_reach \<A>)"
using ta_restrict_states_Q[of \<A> "ta_reachable \<A>"]
by auto
(* It is sound to restrict to reachable states. *)
lemma gta_only_reach_lang:
"gta_lang Q (ta_only_reach \<A>) = gta_lang Q \<A>"
using ta_der_gterm_only_reach
by (auto elim!: gta_langE intro!: gta_langI) force+
lemma \<L>_only_reach: "\<L> (reg_reach R) = \<L> R"
using gta_only_reach_lang
by (auto simp: \<L>_def reg_reach_def)
lemma ta_only_reach_lang:
"ta_lang Q (ta_only_reach \<A>) = ta_lang Q \<A>"
using gta_only_reach_lang
by (metis ta_lang_to_gta_lang)
lemma ta_prod_epsD:
"(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> q |\<in>| ta_productive P \<A> \<Longrightarrow> p |\<in>| ta_productive P \<A>"
using ta_der_ctxt[of q \<A> "\<box>\<langle>Var p\<rangle>"]
by (auto simp: ta_productive_def ta_der_trancl_eps)
lemma ta_only_prod_eps:
"(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> q |\<in>| ta_productive P \<A> \<Longrightarrow> (p, q) |\<in>| (eps (ta_only_prod P \<A>))|\<^sup>+|"
proof (induct rule: ftrancl_induct)
case (Base p q)
then show ?case
by (metis (no_types, lifting) fSigmaI finterI fr_into_trancl ta.sel(2) ta_prod_epsD ta_restrict_def)
next
case (Step p q r) note IS = this
show ?case using IS(2 - 4) ta_prod_epsD[OF fr_into_trancl[OF IS(3)] IS(4)]
by (auto simp: ta_restrict_def) (simp add: ftrancl_into_trancl)
qed
(* Major lemma to show that it is sound to restrict to productive states. *)
lemma ta_der_only_prod:
"q |\<in>| ta_der \<A> t \<Longrightarrow> q |\<in>| ta_productive P \<A> \<Longrightarrow> q |\<in>| ta_der (ta_only_prod P \<A>) t"
proof (induct rule: ta_der_induct)
case (Fun f ts ps p q)
let ?\<A> = "ta_only_prod P \<A>"
have pr: "p |\<in>| ta_productive P \<A>" "i < length ts \<Longrightarrow> ps ! i |\<in>| ta_productive P \<A>" for i
using Fun(2) ta_prod_epsD[of p q] Fun(3, 6) rule_reachable_ctxt_exist[OF Fun(1)]
using ta_productiveI'[of p \<A> _ "ps ! i" P]
by auto
then have "f ps \<rightarrow> p |\<in>| rules ?\<A>" using Fun(1, 2) unfolding ta_restrict_def
by (auto simp: in_fset_conv_nth intro: finite_subset[OF _ finite_Collect_ta_rule, of _ \<A>])
then show ?case using pr Fun ta_only_prod_eps[of p q \<A> P] Fun(3, 6)
by auto
qed (auto intro: ta_only_prod_eps)
lemma ta_der_ta_only_prod_ta_der:
"q |\<in>| ta_der (ta_only_prod P \<A>) t \<Longrightarrow> q |\<in>| ta_der \<A> t"
by (meson ta_der_el_mono ta_restrict_subset ta_subset_def)
(* It is sound to restrict to productive states. *)
lemma gta_only_prod_lang:
"gta_lang Q (ta_only_prod Q \<A>) = gta_lang Q \<A>" (is "gta_lang Q ?\<A> = _")
proof
show "gta_lang Q ?\<A> \<subseteq> gta_lang Q \<A>"
using gta_lang_mono[OF ta_der_mono'[OF ta_restrict_subset]]
by blast
next
{fix t assume "t \<in> gta_lang Q \<A>"
from gta_langE[OF this] obtain q where
reach: "q |\<in>| ta_der \<A> (term_of_gterm t)" "q |\<in>| Q" .
from ta_der_only_prod[OF reach(1) ta_productive_setI[OF reach(2)]] reach(2)
have "t \<in> gta_lang Q ?\<A>" by (auto intro: gta_langI)}
then show "gta_lang Q \<A> \<subseteq> gta_lang Q ?\<A>" by blast
qed
lemma \<L>_only_prod: "\<L> (reg_prod R) = \<L> R"
using gta_only_prod_lang
by (auto simp: \<L>_def reg_prod_def)
lemma ta_only_prod_lang:
"ta_lang Q (ta_only_prod Q \<A>) = ta_lang Q \<A>"
using gta_only_prod_lang
by (metis ta_lang_to_gta_lang)
(* the productive states are also productive w.r.t. the new automaton *)
lemma ta_prodictive_ta_only_prod [simp]:
"ta_productive P (ta_only_prod P \<A>) = ta_productive P \<A>" (is "?LS = ?RS")
proof -
have "?LS |\<subseteq>| ?RS" using ta_der_mono'[OF ta_restrict_subset]
using finite_ta_productive[of \<A> P]
by (auto simp: ta_productive_def) fastforce
moreover have "?RS |\<subseteq>| ?LS" using ta_der_only_prod
by (auto elim!: ta_productiveE)
(smt (z3) ta_der_only_prod ta_productiveI ta_productive_setI)
ultimately show ?thesis by blast
qed
lemma ta_only_prod_productive:
"\<Q> (ta_only_prod P \<A>) |\<subseteq>| ta_productive P (ta_only_prod P \<A>)"
using ta_restrict_states_Q by force
lemma ta_only_prod_reachable:
assumes all_reach: "\<Q> \<A> |\<subseteq>| ta_reachable \<A>"
shows "\<Q> (ta_only_prod P \<A>) |\<subseteq>| ta_reachable (ta_only_prod P \<A>)" (is "?Ls |\<subseteq>| ?Rs")
proof -
{fix q assume "q |\<in>| ?Ls"
then obtain t where "ground t" "q |\<in>| ta_der \<A> t" "q |\<in>| ta_productive P \<A>"
using fsubsetD[OF ta_only_prod_productive[of \<A> P]]
using fsubsetD[OF fsubset_trans[OF ta_restrict_states all_reach, of "ta_productive P \<A>"]]
by (auto elim!: ta_reachableE)
then have "q |\<in>| ?Rs"
by (intro ta_reachableI[where ?\<A> = "ta_only_prod P \<A>" and ?t = t]) (auto simp: ta_der_only_prod)}
then show ?thesis by blast
qed
lemma ta_prod_reach_subset:
"ta_subset (ta_only_prod P (ta_only_reach \<A>)) \<A>"
by (rule ta_subset_trans, (rule ta_restrict_subset)+)
lemma ta_prod_reach_states:
"\<Q> (ta_only_prod P (ta_only_reach \<A>)) |\<subseteq>| \<Q> \<A>"
by (rule ta_subset_states[OF ta_prod_reach_subset])
(* If all states are reachable then there exists a ground context for all productive states *)
lemma ta_productive_aux:
assumes "\<Q> \<A> |\<subseteq>| ta_reachable \<A>" "q |\<in>| ta_der \<A> (C\<langle>t\<rangle>)"
shows "\<exists>C'. ground_ctxt C' \<and> q |\<in>| ta_der \<A> (C'\<langle>t\<rangle>)" using assms(2)
proof (induct C arbitrary: q)
case Hole then show ?case by (intro exI[of _ "\<box>"]) auto
next
case (More f ts1 C ts2)
from More(2) obtain qs q' where q': "f qs \<rightarrow> q' |\<in>| rules \<A>" "q' = q \<or> (q', q) |\<in>| (eps \<A>)|\<^sup>+|"
"qs ! length ts1 |\<in>| ta_der \<A> (C\<langle>t\<rangle>)" "length qs = Suc (length ts1 + length ts2)"
by simp (metis less_add_Suc1 nth_append_length)
{ fix i assume "i < length qs"
then have "qs ! i |\<in>| \<Q> \<A>" using q'(1)
by (auto dest!: rule_statesD(4))
then have "\<exists>t. ground t \<and> qs ! i |\<in>| ta_der \<A> t" using assms(1)
by (simp add: ta_reachable_def) force}
then obtain ts where ts: "i < length qs \<Longrightarrow> ground (ts i) \<and> qs ! i |\<in>| ta_der \<A> (ts i)" for i by metis
obtain C' where C: "ground_ctxt C'" "qs ! length ts1 |\<in>| ta_der \<A> C'\<langle>t\<rangle>" using More(1)[OF q'(3)] by blast
define D where "D \<equiv> More f (map ts [0..<length ts1]) C' (map ts [Suc (length ts1)..<Suc (length ts1 + length ts2)])"
have "ground_ctxt D" unfolding D_def using ts C(1) q'(4) by auto
moreover have "q |\<in>| ta_der \<A> D\<langle>t\<rangle>" using ts C(2) q' unfolding D_def
by (auto simp: nth_append_Cons not_le not_less le_less_Suc_eq Suc_le_eq intro!: exI[of _ qs] exI[of _ q'])
ultimately show ?case by blast
qed
lemma ta_productive_def':
assumes "\<Q> \<A> |\<subseteq>| ta_reachable \<A>"
shows "ta_productive Q \<A> = {| q| q q' C. ground_ctxt C \<and> q' |\<in>| ta_der \<A> (C\<langle>Var q\<rangle>) \<and> q' |\<in>| Q |}"
using ta_productive_aux[OF assms]
by (auto simp: ta_productive_def intro!: finite_subset[OF _ finite_ta_productive, of _ \<A> Q]) force+
(* turn a finite automaton into a trim one, by removing
first all unreachable and then all non-productive states *)
lemma trim_gta_lang: "gta_lang Q (trim_ta Q \<A>) = gta_lang Q \<A>"
unfolding trim_ta_def gta_only_reach_lang gta_only_prod_lang ..
lemma trim_ta_subset: "ta_subset (trim_ta Q \<A>) \<A>"
unfolding trim_ta_def by (rule ta_prod_reach_subset)
theorem trim_ta: "ta_is_trim Q (trim_ta Q \<A>)" unfolding ta_is_trim_def
by (metis fin_mono ta_only_prod_reachable ta_only_reach_reachable
ta_prodictive_ta_only_prod ta_restrict_states_Q trim_ta_def)
lemma reg_is_trim_trim_reg [simp]: "reg_is_trim (trim_reg R)"
unfolding reg_is_trim_def trim_reg_def
by (simp add: trim_ta)
lemma trim_reg_reach [simp]:
"\<Q>\<^sub>r (trim_reg A) |\<subseteq>| ta_reachable (ta (trim_reg A))"
by (auto simp: trim_reg_def) (meson ta_is_trim_def trim_ta)
lemma trim_reg_prod [simp]:
"\<Q>\<^sub>r (trim_reg A) |\<subseteq>| ta_productive (fin (trim_reg A)) (ta (trim_reg A))"
by (auto simp: trim_reg_def) (meson ta_is_trim_def trim_ta)
(* Proposition 7: every tree automaton can be turned into an equivalent trim one *)
lemmas obtain_trimmed_ta = trim_ta trim_gta_lang ta_subset_det[OF trim_ta_subset]
(* Trim tree automaton signature *)
lemma \<L>_trim_ta_sig:
assumes "reg_is_trim R" "\<L> R \<subseteq> \<T>\<^sub>G (fset \<F>)"
shows "ta_sig (ta R) |\<subseteq>| \<F>"
proof -
{fix r assume r: "r |\<in>| rules (ta R)"
then obtain f ps p where [simp]: "r = f ps \<rightarrow> p" by (cases r) auto
from r assms(1) have "fset_of_list ps |\<subseteq>| ta_reachable (ta R)"
by (auto simp add: rule_statesD(4) reg_is_trim_def ta_is_trim_def)
from ta_reachableI_rule[OF this, of f p] r
obtain ts where ts: "length ts = length ps" "\<forall> i < length ps. ground (ts ! i)"
"\<forall> i < length ps. ps ! i |\<in>| ta_der (ta R) (ts ! i)"
by auto
obtain C q where ctxt: "ground_ctxt C" "q |\<in>| ta_der (ta R) (C\<langle>Var p\<rangle>)" "q |\<in>| fin R"
using assms(1) unfolding reg_is_trim_def
by (metis \<open>r = f ps \<rightarrow> p\<close> fsubsetI r rule_statesD(2) ta_productiveE ta_productive_aux ta_is_trim_def)
from ts ctxt r have reach: "q |\<in>| ta_der (ta R) C\<langle>Fun f ts\<rangle>"
by auto (metis ta_der_Fun ta_der_ctxt)
have gr: "ground C\<langle>Fun f ts\<rangle>" using ts(1, 2) ctxt(1)
by (auto simp: in_set_conv_nth)
then have "C\<langle>Fun f ts\<rangle> \<in> ta_lang (fin R) (ta R)" using ctxt(1, 3) ts(1, 2)
apply (intro ta_langI[OF _ _ reach, of "fin R" "C\<langle>Fun f ts\<rangle>"])
apply (auto simp del: adapt_vars_ctxt)
by (metis gr adapt_vars2 adapt_vars_adapt_vars)
then have *: "gterm_of_term C\<langle>Fun f ts\<rangle> \<in> \<L> R" using gr
by (auto simp: \<L>_def)
then have "funas_gterm (gterm_of_term C\<langle>Fun f ts\<rangle>) \<subseteq> fset \<F>" using assms(2) gr
by (auto simp: \<T>\<^sub>G_equivalent_def)
moreover have "(f, length ps) \<in> funas_gterm (gterm_of_term C\<langle>Fun f ts\<rangle>)"
using ts(1) by (auto simp: funas_gterm_gterm_of_term[OF gr])
ultimately have "(r_root r, length (r_lhs_states r)) |\<in>| \<F>"
- by (auto simp: fmember_iff_member_fset)}
+ by auto}
then show ?thesis
by (auto simp: ta_sig_def)
qed
text \<open>Map function over TA rules which change states/signature\<close>
lemma map_ta_rule_iff:
"map_ta_rule f g |`| \<Delta> = {|TA_rule (g h) (map f qs) (f q) | h qs q. TA_rule h qs q |\<in>| \<Delta>|}"
apply (intro fequalityI fsubsetI)
- apply (auto simp add: rev_fimage_eqI)
+ apply (auto simp add: rev_image_eqI)
apply (metis map_ta_rule_cases ta_rule.collapse)
done
lemma \<L>_trim: "\<L> (trim_reg R) = \<L> R"
by (auto simp: trim_gta_lang \<L>_def trim_reg_def)
lemma fmap_funs_ta_def':
"fmap_funs_ta h \<A> = TA {|(h f) qs \<rightarrow> q |f qs q. f qs \<rightarrow> q |\<in>| rules \<A>|} (eps \<A>)"
unfolding fmap_funs_ta_def map_ta_rule_iff by auto
lemma fmap_states_ta_def':
"fmap_states_ta h \<A> = TA {|f (map h qs) \<rightarrow> h q |f qs q. f qs \<rightarrow> q |\<in>| rules \<A>|} (map_both h |`| eps \<A>)"
unfolding fmap_states_ta_def map_ta_rule_iff by auto
lemma fmap_states [simp]:
"\<Q> (fmap_states_ta h \<A>) = h |`| \<Q> \<A>"
unfolding fmap_states_ta_def \<Q>_def
by auto
lemma fmap_states_ta_sig [simp]:
"ta_sig (fmap_states_ta f \<A>) = ta_sig \<A>"
- by (auto simp: fBex_def fmap_states_ta_def ta_sig_def fimage_iff)
- (metis id_def length_map ta_rule.map_sel(1, 2))+
+ by (auto simp: fmap_states_ta_def ta_sig_def ta_rule.map_sel intro: fset.map_cong0)
lemma fmap_states_ta_eps_wit:
assumes "(h p, q) |\<in>| (map_both h |`| eps \<A>)|\<^sup>+|" "finj_on h (\<Q> \<A>)" "p |\<in>| \<Q> \<A>"
obtains q' where "q = h q'" "(p, q') |\<in>| (eps \<A>)|\<^sup>+|" "q' |\<in>| \<Q> \<A>" using assms
by (auto simp: fimage_iff finj_on_def' ftrancl_map_both_fsubset[OF assms(2), of "eps \<A>"])
(metis (mono_tags, lifting) assms(2) eps_trancl_statesD finj_on_eq_iff)
lemma ta_der_fmap_states_inv_superset:
assumes "\<Q> \<A> |\<subseteq>| \<B>" "finj_on h \<B>"
and "q |\<in>| ta_der (fmap_states_ta h \<A>) (term_of_gterm t)"
shows "the_finv_into \<B> h q |\<in>| ta_der \<A> (term_of_gterm t)" using assms(3)
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
from assms(1, 2) have inj: "finj_on h (\<Q> \<A>)" using fsubset_finj_on by blast
have "x |\<in>| \<Q> \<A> \<Longrightarrow> the_finv_into (\<Q> \<A>) h (h x) = the_finv_into \<B> h (h x)" for x
using assms(1, 2) by (metis fsubsetD inj the_finv_into_f_f)
then show ?case using GFun the_finv_into_f_f[OF inj] assms(1)
by (auto simp: fmap_states_ta_def' finj_on_def' rule_statesD eps_statesD
elim!: fmap_states_ta_eps_wit[OF _ inj]
intro!: exI[of _ "the_finv_into \<B> h p"])
qed
lemma ta_der_fmap_states_inv:
assumes "finj_on h (\<Q> \<A>)" "q |\<in>| ta_der (fmap_states_ta h \<A>) (term_of_gterm t)"
shows "the_finv_into (\<Q> \<A>) h q |\<in>| ta_der \<A> (term_of_gterm t)"
by (simp add: ta_der_fmap_states_inv_superset assms)
lemma ta_der_to_fmap_states_der:
assumes "q |\<in>| ta_der \<A> (term_of_gterm t)"
shows "h q |\<in>| ta_der (fmap_states_ta h \<A>) (term_of_gterm t)" using assms
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
then show ?case
using ftrancl_map_prod_mono[of h "eps \<A>"]
by (auto simp: fmap_states_ta_def' intro!: exI[of _ "h p"] exI[of _ "map h ps"])
qed
lemma ta_der_fmap_states_conv:
assumes "finj_on h (\<Q> \<A>)"
shows "ta_der (fmap_states_ta h \<A>) (term_of_gterm t) = h |`| ta_der \<A> (term_of_gterm t)"
using ta_der_to_fmap_states_der[of _ \<A> t] ta_der_fmap_states_inv[OF assms]
using f_the_finv_into_f[OF assms] finj_on_the_finv_into[OF assms]
using gterm_ta_der_states
by (auto intro!: rev_fimage_eqI) fastforce
lemma fmap_states_ta_det:
assumes "finj_on f (\<Q> \<A>)"
shows "ta_det (fmap_states_ta f \<A>) = ta_det \<A>" (is "?Ls = ?Rs")
proof
{fix g ps p q assume ass: "?Ls" "TA_rule g ps p |\<in>| rules \<A>" "TA_rule g ps q |\<in>| rules \<A>"
then have "TA_rule g (map f ps) (f p) |\<in>| rules (fmap_states_ta f \<A>)"
"TA_rule g (map f ps) (f q) |\<in>| rules (fmap_states_ta f \<A>)"
by (force simp: fmap_states_ta_def)+
then have "p = q" using ass finj_on_eq_iff[OF assms]
by (auto simp: ta_det_def) (meson rule_statesD(2))}
then show "?Ls \<Longrightarrow> ?Rs"
by (auto simp: ta_det_def fmap_states_ta_def')
next
{fix g ps qs p q assume ass: "?Rs" "TA_rule g ps p |\<in>| rules \<A>" "TA_rule g qs q |\<in>| rules \<A>"
then have "map f ps = map f qs \<Longrightarrow> ps = qs" using finj_on_eq_iff[OF assms]
by (auto simp: map_eq_nth_conv in_fset_conv_nth dest!: rule_statesD(4) intro!: nth_equalityI)}
then show "?Rs \<Longrightarrow> ?Ls" using finj_on_eq_iff[OF assms]
by (auto simp: ta_det_def fmap_states_ta_def') blast
qed
lemma fmap_states_ta_lang:
"finj_on f (\<Q> \<A>) \<Longrightarrow> Q |\<subseteq>| \<Q> \<A> \<Longrightarrow> gta_lang (f |`| Q) (fmap_states_ta f \<A>) = gta_lang Q \<A>"
using ta_der_fmap_states_conv[of f \<A>]
by (auto simp: finj_on_def' finj_on_eq_iff fsubsetD elim!: gta_langE intro!: gta_langI)
lemma fmap_states_ta_lang2:
"finj_on f (\<Q> \<A> |\<union>| Q) \<Longrightarrow> gta_lang (f |`| Q) (fmap_states_ta f \<A>) = gta_lang Q \<A>"
using ta_der_fmap_states_conv[OF fsubset_finj_on[of f "\<Q> \<A> |\<union>| Q" "\<Q> \<A>"]]
by (auto simp: finj_on_def' elim!: gta_langE intro!: gta_langI) fastforce
definition funs_ta :: "('q, 'f) ta \<Rightarrow> 'f fset" where
"funs_ta \<A> = {|f |f qs q. TA_rule f qs q |\<in>| rules \<A>|}"
lemma funs_ta[code]:
"funs_ta \<A> = (\<lambda> r. case r of TA_rule f ps p \<Rightarrow> f) |`| (rules \<A>)" (is "?Ls = ?Rs")
- by (force simp: funs_ta_def rev_fimage_eqI simp flip: fset.set_map fmember_iff_member_fset
+ by (force simp: funs_ta_def rev_fimage_eqI simp flip: fset.set_map
split!: ta_rule.splits intro!: finite_subset[of "{f. \<exists>qs q. TA_rule f qs q |\<in>| rules \<A>}" "fset ?Rs"])
lemma finite_funs_ta [simp]:
"finite {f. \<exists>qs q. TA_rule f qs q |\<in>| rules \<A>}"
by (intro finite_subset[of "{f. \<exists>qs q. TA_rule f qs q |\<in>| rules \<A>}" "fset (funs_ta \<A>)"])
- (auto simp: funs_ta rev_fimage_eqI simp flip: fset.set_map fmember_iff_member_fset split!: ta_rule.splits)
+ (auto simp: funs_ta rev_fimage_eqI simp flip: fset.set_map split!: ta_rule.splits)
lemma funs_taE [elim]:
assumes "f |\<in>| funs_ta \<A>"
obtains ps p where "TA_rule f ps p |\<in>| rules \<A>" using assms
by (auto simp: funs_ta_def)
lemma funs_taI [intro]:
"TA_rule f ps p |\<in>| rules \<A> \<Longrightarrow> f |\<in>| funs_ta \<A>"
by (auto simp: funs_ta_def)
lemma fmap_funs_ta_cong:
"(\<And>x. x |\<in>| funs_ta \<A> \<Longrightarrow> h x = k x) \<Longrightarrow> \<A> = \<B> \<Longrightarrow> fmap_funs_ta h \<A> = fmap_funs_ta k \<B>"
by (force simp: fmap_funs_ta_def')
lemma [simp]: "{|TA_rule f qs q |f qs q. TA_rule f qs q |\<in>| X|} = X"
by (intro fset_eqI; case_tac x) auto
lemma fmap_funs_ta_id [simp]:
"fmap_funs_ta id \<A> = \<A>" by (simp add: fmap_funs_ta_def')
lemma fmap_states_ta_id [simp]:
"fmap_states_ta id \<A> = \<A>"
by (auto simp: fmap_states_ta_def map_ta_rule_iff prod.map_id0)
lemmas fmap_funs_ta_id' [simp] = fmap_funs_ta_id[unfolded id_def]
lemma fmap_funs_ta_comp:
"fmap_funs_ta h (fmap_funs_ta k A) = fmap_funs_ta (h \<circ> k) A"
proof -
have "r |\<in>| rules A \<Longrightarrow> map_ta_rule id h (map_ta_rule id k r) = map_ta_rule id (\<lambda>x. h (k x)) r" for r
by (cases r) (auto)
then show ?thesis
by (force simp: fmap_funs_ta_def fimage_iff cong: fmap_funs_ta_cong)
qed
lemma fmap_funs_reg_comp:
"fmap_funs_reg h (fmap_funs_reg k A) = fmap_funs_reg (h \<circ> k) A"
using fmap_funs_ta_comp unfolding fmap_funs_reg_def
by auto
lemma fmap_states_ta_comp:
"fmap_states_ta h (fmap_states_ta k A) = fmap_states_ta (h \<circ> k) A"
by (auto simp: fmap_states_ta_def ta_rule.map_comp comp_def id_def prod.map_comp)
lemma funs_ta_fmap_funs_ta [simp]:
"funs_ta (fmap_funs_ta f A) = f |`| funs_ta A"
by (auto simp: funs_ta fmap_funs_ta_def' comp_def fimage_iff
split!: ta_rule.splits) force+
lemma ta_der_funs_ta:
"q |\<in>| ta_der A t \<Longrightarrow> ffuns_term t |\<subseteq>| funs_ta A"
proof (induct t arbitrary: q)
case (Fun f ts)
then have "f |\<in>| funs_ta A" by (auto simp: funs_ta_def)
then show ?case using Fun(1)[OF nth_mem, THEN fsubsetD] Fun(2)
by (auto simp: in_fset_conv_nth) blast+
qed auto
lemma ta_der_fmap_funs_ta:
"q |\<in>| ta_der A t \<Longrightarrow> q |\<in>| ta_der (fmap_funs_ta f A) (map_funs_term f t)"
by (induct t arbitrary: q) (auto 0 4 simp: fmap_funs_ta_def')
lemma ta_der_fmap_states_ta:
assumes "q |\<in>| ta_der A t"
shows "h q |\<in>| ta_der (fmap_states_ta h A) (map_vars_term h t)"
proof -
have [intro]: "(q, q') |\<in>| (eps A)|\<^sup>+| \<Longrightarrow> (h q, h q') |\<in>| (eps (fmap_states_ta h A))|\<^sup>+|" for q q'
by (force intro!: ftrancl_map[of "eps A"] simp: fmap_states_ta_def)
show ?thesis using assms
proof (induct rule: ta_der_induct)
case (Fun f ts ps p q)
have "f (map h ps) \<rightarrow> h p |\<in>| rules (fmap_states_ta h A)"
using Fun(1) by (force simp: fmap_states_ta_def')
then show ?case using Fun by (auto 0 4)
qed auto
qed
lemma ta_der_fmap_states_ta_mono:
shows "f |`| ta_der A (term_of_gterm s) |\<subseteq>| ta_der (fmap_states_ta f A) (term_of_gterm s)"
using ta_der_fmap_states_ta[of _ A "term_of_gterm s" f]
by (simp add: fimage_fsubsetI ta_der_to_fmap_states_der)
lemma ta_der_fmap_states_ta_mono2:
assumes "finj_on f (\<Q> A)"
shows "ta_der (fmap_states_ta f A) (term_of_gterm s) |\<subseteq>| f |`| ta_der A (term_of_gterm s)"
using ta_der_fmap_states_conv[OF assms] by auto
lemma fmap_funs_ta_der':
"q |\<in>| ta_der (fmap_funs_ta h A) t \<Longrightarrow> \<exists>t'. q |\<in>| ta_der A t' \<and> map_funs_term h t' = t"
proof (induct rule: ta_der_induct)
case (Var q v)
then show ?case by (auto simp: fmap_funs_ta_def intro!: exI[of _ "Var v"])
next
case (Fun f ts ps p q)
obtain f' ts' where root: "f = h f'" "f' ps \<rightarrow> p |\<in>| rules A" and
"\<And>i. i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der A (ts' i) \<and> map_funs_term h (ts' i) = ts ! i"
using Fun(1, 5) unfolding fmap_funs_ta_def'
by auto metis
note [simp] = conjunct1[OF this(3)] conjunct2[OF this(3), unfolded id_def]
have [simp]: "p = q \<Longrightarrow> f' ps \<rightarrow> q |\<in>| rules A" using root(2) by auto
show ?case using Fun(3)
by (auto simp: comp_def Fun root fmap_funs_ta_def'
intro!: exI[of _ "Fun f' (map ts' [0..<length ts])"] exI[of _ ps] exI[of _ p] nth_equalityI)
qed
lemma fmap_funs_gta_lang:
"gta_lang Q (fmap_funs_ta h \<A>) = map_gterm h ` gta_lang Q \<A>" (is "?Ls = ?Rs")
proof -
{fix s assume "s \<in> ?Ls" then obtain q where
lang: "q |\<in>| Q" "q |\<in>| ta_der (fmap_funs_ta h \<A>) (term_of_gterm s)"
by auto
from fmap_funs_ta_der'[OF this(2)] obtain t where
t: "q |\<in>| ta_der \<A> t" "map_funs_term h t = term_of_gterm s" "ground t"
by (metis ground_map_term ground_term_of_gterm)
then have "s \<in> ?Rs" using map_gterm_of_term[OF t(3), of h id] lang
by (auto simp: gta_lang_def gta_der_def image_iff)
(metis fempty_iff finterI ground_term_to_gtermD map_term_of_gterm term_of_gterm_inv)}
moreover have "?Rs \<subseteq> ?Ls" using ta_der_fmap_funs_ta[of _ \<A> _ h]
by (auto elim!: gta_langE intro!: gta_langI) fastforce
ultimately show ?thesis by blast
qed
lemma fmap_funs_\<L>:
"\<L> (fmap_funs_reg h R) = map_gterm h ` \<L> R"
using fmap_funs_gta_lang[of "fin R" h]
by (auto simp: fmap_funs_reg_def \<L>_def)
lemma ta_states_fmap_funs_ta [simp]: "\<Q> (fmap_funs_ta f A) = \<Q> A"
by (auto simp: fmap_funs_ta_def \<Q>_def)
lemma ta_reachable_fmap_funs_ta [simp]:
"ta_reachable (fmap_funs_ta f A) = ta_reachable A" unfolding ta_reachable_def
by (metis (mono_tags, lifting) fmap_funs_ta_der' ta_der_fmap_funs_ta ground_map_term)
lemma fin_in_states:
"fin (reg_Restr_Q\<^sub>f R) |\<subseteq>| \<Q>\<^sub>r (reg_Restr_Q\<^sub>f R)"
by (auto simp: reg_Restr_Q\<^sub>f_def)
lemma fmap_states_reg_Restr_Q\<^sub>f_fin:
"finj_on f (\<Q> \<A>) \<Longrightarrow> fin (fmap_states_reg f (reg_Restr_Q\<^sub>f R)) |\<subseteq>| \<Q>\<^sub>r (fmap_states_reg f (reg_Restr_Q\<^sub>f R))"
by (auto simp: fmap_states_reg_def reg_Restr_Q\<^sub>f_def)
lemma \<L>_fmap_states_reg_Inl_Inr [simp]:
"\<L> (fmap_states_reg Inl R) = \<L> R"
"\<L> (fmap_states_reg Inr R) = \<L> R"
unfolding \<L>_def fmap_states_reg_def
by (auto simp: finj_Inl_Inr intro!: fmap_states_ta_lang2)
lemma finite_Collect_prod_ta_rules:
"finite {f qs \<rightarrow> (a, b) |f qs a b. f map fst qs \<rightarrow> a |\<in>| rules \<A> \<and> f map snd qs \<rightarrow> b |\<in>| rules \<BB>}" (is "finite ?set")
proof -
have "?set \<subseteq> (\<lambda> (ra, rb). case ra of f ps \<rightarrow> p \<Rightarrow> case rb of g qs \<rightarrow> q \<Rightarrow> f (zip ps qs) \<rightarrow> (p, q)) ` (srules \<A> \<times> srules \<BB>)"
- by (auto simp: srules_def image_iff fmember_iff_member_fset split!: ta_rule.splits)
+ by (auto simp: srules_def image_iff split!: ta_rule.splits)
(metis ta_rule.inject zip_map_fst_snd)
from finite_imageI[of "srules \<A> \<times> srules \<BB>", THEN finite_subset[OF this]]
show ?thesis by (auto simp: srules_def)
qed
\<comment> \<open>The product automaton of the automata A and B is constructed
by applying the rules on pairs of states\<close>
lemmas prod_eps_def = prod_epsLp_def prod_epsRp_def
lemma finite_prod_epsLp:
"finite (Collect (prod_epsLp \<A> \<B>))"
by (intro finite_subset[of "Collect (prod_epsLp \<A> \<B>)" "fset ((\<Q> \<A> |\<times>| \<Q> \<B>) |\<times>| \<Q> \<A> |\<times>| \<Q> \<B>)"])
- (auto simp: prod_epsLp_def simp flip: fmember_iff_member_fset dest: eps_statesD)
+ (auto simp: prod_epsLp_def dest: eps_statesD)
lemma finite_prod_epsRp:
"finite (Collect (prod_epsRp \<A> \<B>))"
by (intro finite_subset[of "Collect (prod_epsRp \<A> \<B>)" "fset ((\<Q> \<A> |\<times>| \<Q> \<B>) |\<times>| \<Q> \<A> |\<times>| \<Q> \<B>)"])
- (auto simp: prod_epsRp_def simp flip: fmember_iff_member_fset dest: eps_statesD)
+ (auto simp: prod_epsRp_def dest: eps_statesD)
lemmas finite_prod_eps [simp] = finite_prod_epsLp[unfolded prod_epsLp_def] finite_prod_epsRp[unfolded prod_epsRp_def]
lemma [simp]: "f qs \<rightarrow> q |\<in>| rules (prod_ta \<A> \<B>) \<longleftrightarrow> f qs \<rightarrow> q |\<in>| prod_ta_rules \<A> \<B>"
"r |\<in>| rules (prod_ta \<A> \<B>) \<longleftrightarrow> r |\<in>| prod_ta_rules \<A> \<B>"
by (auto simp: prod_ta_def)
lemma prod_ta_states:
"\<Q> (prod_ta \<A> \<B>) |\<subseteq>| \<Q> \<A> |\<times>| \<Q> \<B>"
proof -
{fix q assume "q |\<in>| rule_states (rules (prod_ta \<A> \<B>))"
then obtain f ps p where "f ps \<rightarrow> p |\<in>| rules (prod_ta \<A> \<B>)" and "q |\<in>| fset_of_list ps \<or> p = q"
by (metis rule_statesE)
then have "fst q |\<in>| \<Q> \<A> \<and> snd q |\<in>| \<Q> \<B>"
using rule_statesD(2, 4)[of f "map fst ps" "fst p" \<A>]
using rule_statesD(2, 4)[of f "map snd ps" "snd p" \<B>]
by auto}
moreover
{fix q assume "q |\<in>| eps_states (eps (prod_ta \<A> \<B>))" then have "fst q |\<in>| \<Q> \<A> \<and> snd q |\<in>| \<Q> \<B>"
by (auto simp: eps_states_def prod_ta_def prod_eps_def dest: eps_statesD)}
ultimately show ?thesis
by (auto simp: \<Q>_def) blast+
qed
lemma prod_ta_det:
assumes "ta_det \<A>" and "ta_det \<B>"
shows "ta_det (prod_ta \<A> \<B>)"
using assms unfolding ta_det_def prod_ta_def prod_eps_def
by auto
lemma prod_ta_sig:
"ta_sig (prod_ta \<A> \<B>) |\<subseteq>| ta_sig \<A> |\<union>| ta_sig \<B>"
- by (auto simp add: ta_sig_def fimage_iff fBall_def)+
+proof (rule fsubsetI)
+ fix x
+ assume "x |\<in>| ta_sig (prod_ta \<A> \<B>)"
+ hence "x |\<in>| ta_sig \<A> \<or> x |\<in>| ta_sig \<B>"
+ unfolding ta_sig_def prod_ta_def
+ using image_iff by fastforce
+ thus "x |\<in>| ta_sig (prod_ta \<A> \<B>) \<Longrightarrow> x |\<in>| ta_sig \<A> |\<union>| ta_sig \<B>"
+ by simp
+qed
lemma from_prod_eps:
"(p, q) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+| \<Longrightarrow> (snd p, snd q) |\<notin>| (eps \<B>)|\<^sup>+| \<Longrightarrow> snd p = snd q \<and> (fst p, fst q) |\<in>| (eps \<A>)|\<^sup>+|"
"(p, q) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+| \<Longrightarrow> (fst p, fst q) |\<notin>| (eps \<A>)|\<^sup>+| \<Longrightarrow> fst p = fst q \<and> (snd p, snd q) |\<in>| (eps \<B>)|\<^sup>+|"
apply (induct rule: ftrancl_induct)
apply (auto simp: prod_ta_def prod_eps_def intro: ftrancl_into_trancl )
apply (simp add: fr_into_trancl not_ftrancl_into)+
done
lemma to_prod_eps\<A>:
"(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> r |\<in>| \<Q> \<B> \<Longrightarrow> ((p, r), (q, r)) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|"
by (induct rule: ftrancl_induct)
(auto simp: prod_ta_def prod_eps_def intro: fr_into_trancl ftrancl_into_trancl)
lemma to_prod_eps\<B>:
"(p, q) |\<in>| (eps \<B>)|\<^sup>+| \<Longrightarrow> r |\<in>| \<Q> \<A> \<Longrightarrow> ((r, p), (r, q)) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|"
by (induct rule: ftrancl_induct)
(auto simp: prod_ta_def prod_eps_def intro: fr_into_trancl ftrancl_into_trancl)
lemma to_prod_eps:
"(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> (p', q') |\<in>| (eps \<B>)|\<^sup>+| \<Longrightarrow> ((p, p'), (q, q')) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|"
proof (induct rule: ftrancl_induct)
case (Base a b)
show ?case using Base(2, 1)
proof (induct rule: ftrancl_induct)
case (Base c d)
then have "((a, c), b, c) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|" using finite_prod_eps
by (auto simp: prod_ta_def prod_eps_def dest: eps_statesD intro!: fr_into_trancl ftrancl_into_trancl)
moreover have "((b, c), b, d) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|" using finite_prod_eps Base
by (auto simp: prod_ta_def prod_eps_def dest: eps_statesD intro!: fr_into_trancl ftrancl_into_trancl)
ultimately show ?case
by (auto intro: ftrancl_trans)
next
case (Step p q r)
then have "((b, q), b, r) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|" using finite_prod_eps
by (auto simp: prod_ta_def prod_eps_def dest: eps_statesD intro!: fr_into_trancl)
then show ?case using Step
by (auto intro: ftrancl_trans)
qed
next
case (Step a b c)
from Step have "q' |\<in>| \<Q> \<B>"
by (auto dest: eps_trancl_statesD)
then have "((b, q'), (c,q')) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|"
using Step(3) finite_prod_eps
by (auto simp: prod_ta_def prod_eps_def intro!: fr_into_trancl)
then show ?case using ftrancl_trans Step
by auto
qed
lemma prod_ta_der_to_\<A>_\<B>_der1:
assumes "q |\<in>| ta_der (prod_ta \<A> \<B>) (term_of_gterm t)"
shows "fst q |\<in>| ta_der \<A> (term_of_gterm t)" using assms
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
then show ?case
by (auto dest: from_prod_eps intro!: exI[of _ "map fst ps"] exI[of _ "fst p"])
qed
lemma prod_ta_der_to_\<A>_\<B>_der2:
assumes "q |\<in>| ta_der (prod_ta \<A> \<B>) (term_of_gterm t)"
shows "snd q |\<in>| ta_der \<B> (term_of_gterm t)" using assms
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
then show ?case
by (auto dest: from_prod_eps intro!: exI[of _ "map snd ps"] exI[of _ "snd p"])
qed
lemma \<A>_\<B>_der_to_prod_ta:
assumes "fst q |\<in>| ta_der \<A> (term_of_gterm t)" "snd q |\<in>| ta_der \<B> (term_of_gterm t)"
shows "q |\<in>| ta_der (prod_ta \<A> \<B>) (term_of_gterm t)" using assms
proof (induct t arbitrary: q)
case (GFun f ts)
from GFun(2, 3) obtain ps qs p q' where
rules: "f ps \<rightarrow> p |\<in>| rules \<A>" "f qs \<rightarrow> q' |\<in>| rules \<B>" "length ps = length ts" "length ps = length qs" and
eps: "p = fst q \<or> (p, fst q) |\<in>| (eps \<A>)|\<^sup>+|" "q' = snd q \<or> (q', snd q) |\<in>| (eps \<B>)|\<^sup>+|" and
steps: "\<forall> i < length qs. ps ! i |\<in>| ta_der \<A> (term_of_gterm (ts ! i))"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<B> (term_of_gterm (ts ! i))"
by auto
from rules have st: "p |\<in>| \<Q> \<A>" "q' |\<in>| \<Q> \<B>" by (auto dest: rule_statesD)
have "(p, snd q) = q \<or> ((p, q'), q) |\<in>| (eps (prod_ta \<A> \<B>))|\<^sup>+|" using eps st
using to_prod_eps\<B>[of q' "snd q" \<B> "fst q" \<A>]
using to_prod_eps\<A>[of p "fst q" \<A> "snd q" \<B>]
using to_prod_eps[of p "fst q" \<A> q' "snd q" \<B>]
by (cases "p = fst q"; cases "q' = snd q") (auto simp: prod_ta_def)
then show ?case using rules eps steps GFun(1) st
by (cases "(p, snd q) = q")
(auto simp: finite_Collect_prod_ta_rules dest: to_prod_eps\<B> intro!: exI[of _ p] exI[of _ q'] exI[of _ "zip ps qs"])
qed
lemma prod_ta_der:
"q |\<in>| ta_der (prod_ta \<A> \<B>) (term_of_gterm t) \<longleftrightarrow>
fst q |\<in>| ta_der \<A> (term_of_gterm t) \<and> snd q |\<in>| ta_der \<B> (term_of_gterm t)"
using prod_ta_der_to_\<A>_\<B>_der1 prod_ta_der_to_\<A>_\<B>_der2 \<A>_\<B>_der_to_prod_ta
by blast
lemma intersect_ta_gta_lang:
"gta_lang (Q\<^sub>\<A> |\<times>| Q\<^sub>\<B>) (prod_ta \<A> \<B>) = gta_lang Q\<^sub>\<A> \<A> \<inter> gta_lang Q\<^sub>\<B> \<B>"
by (auto simp: prod_ta_der elim!: gta_langE intro: gta_langI)
lemma \<L>_intersect: "\<L> (reg_intersect R L) = \<L> R \<inter> \<L> L"
by (auto simp: intersect_ta_gta_lang \<L>_def reg_intersect_def)
lemma intersect_ta_ta_lang:
"ta_lang (Q\<^sub>\<A> |\<times>| Q\<^sub>\<B>) (prod_ta \<A> \<B>) = ta_lang Q\<^sub>\<A> \<A> \<inter> ta_lang Q\<^sub>\<B> \<B>"
using intersect_ta_gta_lang[of Q\<^sub>\<A> Q\<^sub>\<B> \<A> \<B>]
by auto (metis IntI imageI term_of_gterm_inv)
\<comment> \<open>Union of tree automata\<close>
lemma ta_union_ta_subset:
"ta_subset \<A> (ta_union \<A> \<B>)" "ta_subset \<B> (ta_union \<A> \<B>)"
unfolding ta_subset_def ta_union_def
by auto
lemma ta_union_states [simp]:
"\<Q> (ta_union \<A> \<B>) = \<Q> \<A> |\<union>| \<Q> \<B>"
by (auto simp: ta_union_def \<Q>_def)
lemma ta_union_sig [simp]:
"ta_sig (ta_union \<A> \<B>) = ta_sig \<A> |\<union>| ta_sig \<B>"
by (auto simp: ta_union_def ta_sig_def)
lemma ta_union_eps_disj_states:
assumes "\<Q> \<A> |\<inter>| \<Q> \<B> = {||}" and "(p, q) |\<in>| (eps (ta_union \<A> \<B>))|\<^sup>+|"
shows "(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<or> (p, q) |\<in>| (eps \<B>)|\<^sup>+|" using assms(2, 1)
by (induct rule: ftrancl_induct)
(auto simp: ta_union_def ftrancl_into_trancl dest: eps_statesD eps_trancl_statesD)
lemma eps_ta_union_eps [simp]:
"(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> (p, q) |\<in>| (eps (ta_union \<A> \<B>))|\<^sup>+|"
"(p, q) |\<in>| (eps \<B>)|\<^sup>+| \<Longrightarrow> (p, q) |\<in>| (eps (ta_union \<A> \<B>))|\<^sup>+|"
by (auto simp add: in_ftrancl_UnI ta_union_def)
lemma disj_states_eps [simp]:
"\<Q> \<A> |\<inter>| \<Q> \<B> = {||} \<Longrightarrow> f ps \<rightarrow> p |\<in>| rules \<A> \<Longrightarrow> (p, q) |\<in>| (eps \<B>)|\<^sup>+| \<longleftrightarrow> False"
"\<Q> \<A> |\<inter>| \<Q> \<B> = {||} \<Longrightarrow> f ps \<rightarrow> p |\<in>| rules \<B> \<Longrightarrow> (p, q) |\<in>| (eps \<A>)|\<^sup>+| \<longleftrightarrow> False"
by (auto simp: rtrancl_eq_or_trancl dest: rule_statesD eps_trancl_statesD)
lemma ta_union_der_disj_states:
assumes "\<Q> \<A> |\<inter>| \<Q> \<B> = {||}" and "q |\<in>| ta_der (ta_union \<A> \<B>) t"
shows "q |\<in>| ta_der \<A> t \<or> q |\<in>| ta_der \<B> t" using assms(2)
proof (induct rule: ta_der_induct)
case (Var q v)
then show ?case using ta_union_eps_disj_states[OF assms(1)]
by auto
next
case (Fun f ts ps p q)
have dist: "fset_of_list ps |\<subseteq>| \<Q> \<A> \<Longrightarrow> i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der \<A> (ts ! i)"
"fset_of_list ps |\<subseteq>| \<Q> \<B> \<Longrightarrow> i < length ts \<Longrightarrow> ps ! i |\<in>| ta_der \<B> (ts ! i)" for i
using Fun(2) Fun(5)[of i] assms(1)
by (auto dest!: ta_der_not_stateD fsubsetD)
from Fun(1) consider (a) "fset_of_list ps |\<subseteq>| \<Q> \<A>" | (b) "fset_of_list ps |\<subseteq>| \<Q> \<B>"
by (auto simp: ta_union_def dest: rule_statesD)
then show ?case using dist Fun(1, 2) assms(1) ta_union_eps_disj_states[OF assms(1), of p q] Fun(3)
by (cases) (auto simp: fsubsetI rule_statesD ta_union_def intro!: exI[of _ p] exI[of _ ps])
qed
lemma ta_union_der_disj_states':
assumes "\<Q> \<A> |\<inter>| \<Q> \<B> = {||}"
shows "ta_der (ta_union \<A> \<B>) t = ta_der \<A> t |\<union>| ta_der \<B> t"
using ta_union_der_disj_states[OF assms] ta_der_mono' ta_union_ta_subset
by (auto, fastforce) blast
lemma ta_union_gta_lang:
assumes "\<Q> \<A> |\<inter>| \<Q> \<B> = {||}" and "Q\<^sub>\<A> |\<subseteq>| \<Q> \<A>" and "Q\<^sub>\<B> |\<subseteq>| \<Q> \<B>"
shows"gta_lang (Q\<^sub>\<A> |\<union>| Q\<^sub>\<B>) (ta_union \<A> \<B>) = gta_lang Q\<^sub>\<A> \<A> \<union> gta_lang Q\<^sub>\<B> \<B>" (is "?Ls = ?Rs")
proof -
{fix s assume "s \<in> ?Ls" then obtain q
where w: "q |\<in>| Q\<^sub>\<A> |\<union>| Q\<^sub>\<B>" "q |\<in>| ta_der (ta_union \<A> \<B>) (term_of_gterm s)"
by (auto elim: gta_langE)
from ta_union_der_disj_states[OF assms(1) w(2)] consider
(a) "q |\<in>| ta_der \<A> (term_of_gterm s)" | "q |\<in>| ta_der \<B> (term_of_gterm s)" by blast
then have "s \<in> ?Rs" using w(1) assms
by (cases, auto simp: gta_langI)
(metis fempty_iff finterI funion_iff gterm_ta_der_states sup.orderE)}
moreover have "?Rs \<subseteq> ?Ls" using ta_union_der_disj_states'[OF assms(1)]
by (auto elim!: gta_langE intro!: gta_langI)
ultimately show ?thesis by blast
qed
lemma \<L>_union: "\<L> (reg_union R L) = \<L> R \<union> \<L> L"
proof -
let ?inl = "Inl :: 'b \<Rightarrow> 'b + 'c" let ?inr = "Inr :: 'c \<Rightarrow> 'b + 'c"
let ?fr = "?inl |`| (fin R |\<inter>| \<Q>\<^sub>r R)" let ?fl = "?inr |`| (fin L |\<inter>| \<Q>\<^sub>r L)"
have [simp]:"gta_lang (?fr |\<union>| ?fl) (ta_union (fmap_states_ta ?inl (ta R)) (fmap_states_ta ?inr (ta L))) =
gta_lang ?fr (fmap_states_ta ?inl (ta R)) \<union> gta_lang ?fl (fmap_states_ta ?inr (ta L))"
by (intro ta_union_gta_lang) (auto simp: fimage_iff)
have [simp]: "gta_lang ?fr (fmap_states_ta ?inl (ta R)) = gta_lang (fin R |\<inter>| \<Q>\<^sub>r R) (ta R)"
by (intro fmap_states_ta_lang) (auto simp: finj_Inl_Inr)
have [simp]: "gta_lang ?fl (fmap_states_ta ?inr (ta L)) = gta_lang (fin L |\<inter>| \<Q>\<^sub>r L) (ta L)"
by (intro fmap_states_ta_lang) (auto simp: finj_Inl_Inr)
show ?thesis
using gta_lang_Rest_states_conv
by (auto simp: \<L>_def reg_union_def ta_union_gta_lang) fastforce
qed
lemma reg_union_states:
"\<Q>\<^sub>r (reg_union A B) = (Inl |`| \<Q>\<^sub>r A) |\<union>| (Inr |`| \<Q>\<^sub>r B)"
by (auto simp: reg_union_def)
\<comment> \<open>Deciding emptiness\<close>
lemma ta_empty [simp]:
"ta_empty Q \<A> = (gta_lang Q \<A> = {})"
by (auto simp: ta_empty_def elim!: gta_langE ta_reachable_gtermE
intro: ta_reachable_gtermI gta_langI)
lemma reg_empty [simp]:
"reg_empty R = (\<L> R = {})"
by (simp add: \<L>_def reg_empty_def)
text \<open>Epsilon free automaton\<close>
lemma finite_eps_free_rulep [simp]:
"finite (Collect (eps_free_rulep \<A>))"
proof -
let ?par = "(\<lambda> r. case r of f qs \<rightarrow> q \<Rightarrow> (f, qs)) |`| (rules \<A>)"
let ?st = "(\<lambda> (r, q). case r of (f, qs) \<Rightarrow> TA_rule f qs q) |`| (?par |\<times>| \<Q> \<A>)"
show ?thesis using rule_statesD eps_trancl_statesD
by (intro finite_subset[of "Collect (eps_free_rulep \<A>)" "fset ?st"])
(auto simp: eps_free_rulep_def fimage_iff
- simp flip: fset.set_map fmember_iff_member_fset
+ simp flip: fset.set_map
split!: ta_rule.splits, fastforce+)
qed
lemmas finite_eps_free_rule [simp] = finite_eps_free_rulep[unfolded eps_free_rulep_def]
lemma ta_res_eps_free:
"ta_der (eps_free \<A>) (term_of_gterm t) = ta_der \<A> (term_of_gterm t)" (is "?Ls = ?Rs")
proof -
{fix q assume "q |\<in>| ?Ls" then have "q |\<in>| ?Rs"
by (induct rule: ta_der_gterm_induct)
(auto simp: eps_free_def eps_free_rulep_def)}
moreover
{fix q assume "q |\<in>| ?Rs" then have "q |\<in>| ?Ls"
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
then show ?case
by (auto simp: eps_free_def eps_free_rulep_def intro!: exI[of _ ps])
qed}
ultimately show ?thesis by blast
qed
lemma ta_lang_eps_free [simp]:
"gta_lang Q (eps_free \<A>) = gta_lang Q \<A>"
by (auto simp add: ta_res_eps_free elim!: gta_langE intro: gta_langI)
lemma \<L>_eps_free: "\<L> (eps_free_reg R) = \<L> R"
by (auto simp: \<L>_def eps_free_reg_def)
text \<open>Sufficient criterion for containment\<close>
(* sufficient criterion to check whether automaton accepts at least T_g(F) where F is a subset of
the signature *)
definition ta_contains_aux :: "('f \<times> nat) set \<Rightarrow> 'q fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> 'q fset \<Rightarrow> bool" where
"ta_contains_aux \<F> Q\<^sub>1 \<A> Q\<^sub>2 \<equiv> (\<forall> f qs. (f, length qs) \<in> \<F> \<and> fset_of_list qs |\<subseteq>| Q\<^sub>1 \<longrightarrow>
(\<exists> q q'. TA_rule f qs q |\<in>| rules \<A> \<and> q' |\<in>| Q\<^sub>2 \<and> (q = q' \<or> (q, q') |\<in>| (eps \<A>)|\<^sup>+|)))"
lemma ta_contains_aux_state_set:
assumes "ta_contains_aux \<F> Q \<A> Q" "t \<in> \<T>\<^sub>G \<F>"
shows "\<exists> q. q |\<in>| Q \<and> q |\<in>| ta_der \<A> (term_of_gterm t)" using assms(2)
proof (induct rule: \<T>\<^sub>G.induct)
case (const a)
then show ?case using assms(1)
by (force simp: ta_contains_aux_def)
next
case (ind f n ss)
obtain qs where "fset_of_list qs |\<subseteq>| Q" "length ss = length qs"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> (term_of_gterm (ss ! i))"
using ind(4) Ex_list_of_length_P[of "length ss" "\<lambda> q i. q |\<in>| Q \<and> q |\<in>| ta_der \<A> (term_of_gterm (ss ! i))"]
by (auto simp: fset_list_fsubset_eq_nth_conv) metis
then show ?case using ind(1 - 3) assms(1)
by (auto simp: ta_contains_aux_def) blast
qed
lemma ta_contains_aux_mono:
assumes "ta_subset \<A> \<B>" and "Q\<^sub>2 |\<subseteq>| Q\<^sub>2'"
shows "ta_contains_aux \<F> Q\<^sub>1 \<A> Q\<^sub>2 \<Longrightarrow> ta_contains_aux \<F> Q\<^sub>1 \<B> Q\<^sub>2'"
using assms unfolding ta_contains_aux_def ta_subset_def
by (meson fin_mono ftrancl_mono)
definition ta_contains :: "('f \<times> nat) set \<Rightarrow> ('f \<times> nat) set \<Rightarrow> ('q, 'f) ta \<Rightarrow> 'q fset \<Rightarrow> 'q fset \<Rightarrow> bool"
where "ta_contains \<F> \<G> \<A> Q Q\<^sub>f \<equiv> ta_contains_aux \<F> Q \<A> Q \<and> ta_contains_aux \<G> Q \<A> Q\<^sub>f"
lemma ta_contains_mono:
assumes "ta_subset \<A> \<B>" and "Q\<^sub>f |\<subseteq>| Q\<^sub>f'"
shows "ta_contains \<F> \<G> \<A> Q Q\<^sub>f \<Longrightarrow> ta_contains \<F> \<G> \<B> Q Q\<^sub>f'"
unfolding ta_contains_def
using ta_contains_aux_mono[OF assms(1) fsubset_refl]
using ta_contains_aux_mono[OF assms]
by blast
lemma ta_contains_both:
assumes contain: "ta_contains \<F> \<G> \<A> Q Q\<^sub>f"
shows "\<And> t. groot t \<in> \<G> \<Longrightarrow> \<Union> (funas_gterm ` set (gargs t)) \<subseteq> \<F> \<Longrightarrow> t \<in> gta_lang Q\<^sub>f \<A>"
proof -
fix t :: "'a gterm"
assume F: "\<Union> (funas_gterm ` set (gargs t)) \<subseteq> \<F>" and G: "groot t \<in> \<G>"
obtain g ss where t[simp]: "t = GFun g ss" by (cases t, auto)
then have "\<exists> qs. length qs = length ss \<and> (\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> (term_of_gterm (ss ! i)) \<and> qs ! i |\<in>| Q)"
using contain ta_contains_aux_state_set[of \<F> Q \<A> "ss ! i" for i] F
unfolding ta_contains_def \<T>\<^sub>G_funas_gterm_conv
using Ex_list_of_length_P[of "length ss" "\<lambda> q i. q |\<in>| Q \<and> q |\<in>| ta_der \<A> (term_of_gterm (ss ! i))"]
by auto (metis SUP_le_iff nth_mem)
then obtain qs where " length qs = length ss"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> (term_of_gterm (ss ! i))"
"\<forall> i < length qs. qs ! i |\<in>| Q"
by blast
then obtain q where "q |\<in>| Q\<^sub>f" "q |\<in>| ta_der \<A> (term_of_gterm t)"
using conjunct2[OF contain[unfolded ta_contains_def]] G
by (auto simp: ta_contains_def ta_contains_aux_def fset_list_fsubset_eq_nth_conv) metis
then show "t \<in> gta_lang Q\<^sub>f \<A>"
by (intro gta_langI) simp
qed
lemma ta_contains:
assumes contain: "ta_contains \<F> \<F> \<A> Q Q\<^sub>f"
shows "\<T>\<^sub>G \<F> \<subseteq> gta_lang Q\<^sub>f \<A>" (is "?A \<subseteq> _")
proof -
have [simp]: "funas_gterm t \<subseteq> \<F> \<Longrightarrow> groot t \<in> \<F>" for t by (cases t) auto
have [simp]: "funas_gterm t \<subseteq> \<F> \<Longrightarrow> \<Union> (funas_gterm ` set (gargs t)) \<subseteq> \<F>" for t
by (cases t) auto
show ?thesis using ta_contains_both[OF contain]
by (auto simp: \<T>\<^sub>G_equivalent_def)
qed
text \<open>Relabeling, map finite set to natural numbers\<close>
lemma map_fset_to_nat_inj:
assumes "Y |\<subseteq>| X"
shows "finj_on (map_fset_to_nat X) Y"
proof -
{ fix x y assume "x |\<in>| X" "y |\<in>| X"
then have "x |\<in>| fset_of_list (sorted_list_of_fset X)" "y |\<in>| fset_of_list (sorted_list_of_fset X)"
by simp_all
note this[unfolded mem_idx_fset_sound]
then have "x = y" if "map_fset_to_nat X x = map_fset_to_nat X y"
using that nth_eq_iff_index_eq[OF distinct_sorted_list_of_fset[of X]]
by (force dest: mem_idx_sound_output simp: map_fset_to_nat_def) }
then show ?thesis using assms
by (auto simp add: finj_on_def' fBall_def)
qed
lemma map_fset_fset_to_nat_inj:
assumes "Y |\<subseteq>| X"
shows "finj_on (map_fset_fset_to_nat X) Y" using assms
proof -
let ?f = "map_fset_fset_to_nat X"
{ fix x y assume "x |\<in>| X" "y |\<in>| X"
then have "sorted_list_of_fset x |\<in>| fset_of_list (sorted_list_of_fset (sorted_list_of_fset |`| X))"
"sorted_list_of_fset y |\<in>| fset_of_list (sorted_list_of_fset (sorted_list_of_fset |`| X))"
unfolding map_fset_fset_to_nat_def by auto
note this[unfolded mem_idx_fset_sound]
then have "x = y" if "?f x = ?f y"
using that nth_eq_iff_index_eq[OF distinct_sorted_list_of_fset[of "sorted_list_of_fset |`| X"]]
by (auto simp: map_fset_fset_to_nat_def)
- (metis mem_idx_sound_output notin_fset sorted_list_of_fset_simps(1))+}
+ (metis mem_idx_sound_output sorted_list_of_fset_simps(1))+}
then show ?thesis using assms
by (auto simp add: finj_on_def' fBall_def)
qed
lemma relabel_gta_lang [simp]:
"gta_lang (relabel_Q\<^sub>f Q \<A>) (relabel_ta \<A>) = gta_lang Q \<A>"
proof -
have "gta_lang (relabel_Q\<^sub>f Q \<A>) (relabel_ta \<A>) = gta_lang (Q |\<inter>| \<Q> \<A>) \<A>"
unfolding relabel_ta_def relabel_Q\<^sub>f_def
by (intro fmap_states_ta_lang2 map_fset_to_nat_inj) simp
then show ?thesis by fastforce
qed
lemma \<L>_relable [simp]: "\<L> (relabel_reg R) = \<L> R"
by (auto simp: \<L>_def relabel_reg_def)
lemma relabel_ta_lang [simp]:
"ta_lang (relabel_Q\<^sub>f Q \<A>) (relabel_ta \<A>) = ta_lang Q \<A>"
unfolding ta_lang_to_gta_lang
using relabel_gta_lang
by simp
lemma relabel_fset_gta_lang [simp]:
"gta_lang (relabel_fset_Q\<^sub>f Q \<A>) (relabel_fset_ta \<A>) = gta_lang Q \<A>"
proof -
have "gta_lang (relabel_fset_Q\<^sub>f Q \<A>) (relabel_fset_ta \<A>) = gta_lang (Q |\<inter>| \<Q> \<A>) \<A>"
unfolding relabel_fset_Q\<^sub>f_def relabel_fset_ta_def
by (intro fmap_states_ta_lang2 map_fset_fset_to_nat_inj) simp
then show ?thesis by fastforce
qed
lemma \<L>_relable_fset [simp]: "\<L> (relable_fset_reg R) = \<L> R"
by (auto simp: \<L>_def relable_fset_reg_def)
lemma ta_states_trim_ta:
"\<Q> (trim_ta Q \<A>) |\<subseteq>| \<Q> \<A>"
unfolding trim_ta_def using ta_prod_reach_states .
lemma trim_ta_reach: "\<Q> (trim_ta Q \<A>) |\<subseteq>| ta_reachable (trim_ta Q \<A>)"
unfolding trim_ta_def using ta_only_prod_reachable ta_only_reach_reachable
by metis
lemma trim_ta_prod: "\<Q> (trim_ta Q A) |\<subseteq>| ta_productive Q (trim_ta Q A)"
unfolding trim_ta_def using ta_only_prod_productive
by metis
lemma empty_gta_lang:
"gta_lang Q (TA {||} {||}) = {}"
using ta_reachable_gtermI
by (force simp: gta_lang_def gta_der_def elim!: ta_langE)
abbreviation empty_reg where
"empty_reg \<equiv> Reg {||} (TA {||} {||})"
lemma \<L>_epmty:
"\<L> empty_reg = {}"
by (auto simp: \<L>_def empty_gta_lang)
lemma const_ta_lang:
"gta_lang {|q|} (TA {| TA_rule f [] q |} {||}) = {GFun f []}"
proof -
have [dest!]: "q' |\<in>| ta_der (TA {| TA_rule f [] q |} {||}) t' \<Longrightarrow> ground t' \<Longrightarrow> t' = Fun f []" for t' q'
by (induct t') auto
show ?thesis
by (auto simp: gta_lang_def gta_der_def elim!: gta_langE)
(metis gterm_of_term.simps list.simps(8) term_of_gterm_inv)
qed
lemma run_argsD:
"run \<A> s t \<Longrightarrow> length (gargs s) = length (gargs t) \<and> (\<forall> i < length (gargs t). run \<A> (gargs s ! i) (gargs t ! i))"
using run.cases by fastforce
lemma run_root_rule:
"run \<A> s t \<Longrightarrow> TA_rule (groot_sym t) (map ex_comp_state (gargs s)) (ex_rule_state s) |\<in>| (rules \<A>) \<and>
(ex_rule_state s = ex_comp_state s \<or> (ex_rule_state s, ex_comp_state s) |\<in>| (eps \<A>)|\<^sup>+|)"
by (cases s; cases t) (auto elim: run.cases)
lemma run_poss_eq: "run \<A> s t \<Longrightarrow> gposs s = gposs t"
by (induct rule: run.induct) auto
lemma run_gsubt_cl:
assumes "run \<A> s t" and "p \<in> gposs t"
shows "run \<A> (gsubt_at s p) (gsubt_at t p)" using assms
proof (induct p arbitrary: s t)
case (Cons i p) show ?case using Cons(1) Cons(2-)
by (cases s; cases t) (auto dest: run_argsD)
qed auto
lemma run_replace_at:
assumes "run \<A> s t" and "run \<A> u v" and "p \<in> gposs s"
and "ex_comp_state (gsubt_at s p) = ex_comp_state u"
shows "run \<A> s[p \<leftarrow> u]\<^sub>G t[p \<leftarrow> v]\<^sub>G" using assms
proof (induct p arbitrary: s t)
case (Cons i p)
obtain r q qs f ts where [simp]: "s = GFun (r, q) qs" "t = GFun f ts" by (cases s, cases t) auto
have *: "j < length qs \<Longrightarrow> ex_comp_state (qs[i := (qs ! i)[p \<leftarrow> u]\<^sub>G] ! j) = ex_comp_state (qs ! j)" for j
using Cons(5) by (cases "i = j", cases p) auto
have [simp]: "map ex_comp_state (qs[i := (qs ! i)[p \<leftarrow> u]\<^sub>G]) = map ex_comp_state qs" using Cons(5)
by (auto simp: *[unfolded comp_def] intro!: nth_equalityI)
have "run \<A> (qs ! i)[p \<leftarrow> u]\<^sub>G (ts ! i)[p \<leftarrow> v]\<^sub>G" using Cons(2-)
by (intro Cons(1)) (auto dest: run_argsD)
moreover have "i < length qs" "i < length ts" using Cons(4) run_poss_eq[OF Cons(2)]
by force+
ultimately show ?case using Cons(2) run_root_rule[OF Cons(2)]
by (fastforce simp: nth_list_update dest: run_argsD intro!: run.intros)
qed simp
text \<open>relating runs to derivation definition\<close>
lemma run_to_comp_st_gta_der:
"run \<A> s t \<Longrightarrow> ex_comp_state s |\<in>| gta_der \<A> t"
proof (induct s arbitrary: t)
case (GFun q qs)
show ?case using GFun(1)[OF nth_mem, of i "gargs t ! i" for i]
using run_argsD[OF GFun(2)] run_root_rule[OF GFun(2-)]
by (cases t) (auto simp: gta_der_def intro!: exI[of _ "map ex_comp_state qs"] exI[of _ "fst q"])
qed
lemma run_to_rule_st_gta_der:
assumes "run \<A> s t" shows "ex_rule_state s |\<in>| gta_der \<A> t"
proof (cases s)
case [simp]: (GFun q qs)
have "i < length qs \<Longrightarrow> ex_comp_state (qs ! i) |\<in>| gta_der \<A> (gargs t ! i)" for i
using run_to_comp_st_gta_der[of \<A>] run_argsD[OF assms] by force
then show ?thesis using conjunct1[OF run_argsD[OF assms]] run_root_rule[OF assms]
by (cases t) (auto simp: gta_der_def intro!: exI[of _ "map ex_comp_state qs"] exI[of _ "fst q"])
qed
lemma run_to_gta_der_gsubt_at:
assumes "run \<A> s t" and "p \<in> gposs t"
shows "ex_rule_state (gsubt_at s p) |\<in>| gta_der \<A> (gsubt_at t p)"
"ex_comp_state (gsubt_at s p) |\<in>| gta_der \<A> (gsubt_at t p)"
using assms run_gsubt_cl[THEN run_to_comp_st_gta_der] run_gsubt_cl[THEN run_to_rule_st_gta_der]
by blast+
lemma gta_der_to_run:
"q |\<in>| gta_der \<A> t \<Longrightarrow> (\<exists> p qs. run \<A> (GFun (p, q) qs) t)" unfolding gta_der_def
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
from GFun(5) Ex_list_of_length_P[of "length ts" "\<lambda> qs i. run \<A> (GFun (fst qs, ps ! i) (snd qs)) (ts ! i)"]
obtain qss where mid: "length qss = length ts" "\<forall> i < length ts .run \<A> (GFun (fst (qss ! i), ps ! i) (snd (qss ! i))) (ts ! i)"
by auto
have [simp]: "map (ex_comp_state \<circ> (\<lambda>(qs, y). GFun (fst y, qs) (snd y))) (zip ps qss) = ps" using GFun(2) mid(1)
by (intro nth_equalityI) auto
show ?case using mid GFun(1 - 4)
by (intro exI[of _ p] exI[of _ "map2 (\<lambda> f args. GFun (fst args, f) (snd args)) ps qss"])
(auto intro: run.intros)
qed
lemma run_ta_der_ctxt_split1:
assumes "run \<A> s t" "p \<in> gposs t"
shows "ex_comp_state s |\<in>| ta_der \<A> (ctxt_at_pos (term_of_gterm t) p)\<langle>Var (ex_comp_state (gsubt_at s p))\<rangle>"
using assms
proof (induct p arbitrary: s t)
case (Cons i p)
obtain q f qs ts where [simp]: "s = GFun q qs" "t = GFun f ts" and l: "length qs = length ts"
using run_argsD[OF Cons(2)] by (cases s, cases t) auto
from Cons(2, 3) l have "ex_comp_state (qs ! i) |\<in>| ta_der \<A> (ctxt_at_pos (term_of_gterm (ts ! i)) p)\<langle>Var (ex_comp_state (gsubt_at (qs ! i) p))\<rangle>"
by (intro Cons(1)) (auto dest: run_argsD)
then show ?case using Cons(2-) l
by (fastforce simp: nth_append_Cons min_def dest: run_root_rule run_argsD
intro!: exI[of _ "map ex_comp_state (gargs s)"] exI[of _ "ex_rule_state s"]
run_to_comp_st_gta_der[of \<A> "qs ! i" "ts ! i" for i, unfolded comp_def gta_der_def])
qed auto
lemma run_ta_der_ctxt_split2:
assumes "run \<A> s t" "p \<in> gposs t"
shows "ex_comp_state s |\<in>| ta_der \<A> (ctxt_at_pos (term_of_gterm t) p)\<langle>Var (ex_rule_state (gsubt_at s p))\<rangle>"
proof (cases "ex_rule_state (gsubt_at s p) = ex_comp_state (gsubt_at s p)")
case False then show ?thesis
using run_root_rule[OF run_gsubt_cl[OF assms]]
by (intro ta_der_eps_ctxt[OF run_ta_der_ctxt_split1[OF assms]]) auto
qed (auto simp: run_ta_der_ctxt_split1[OF assms, unfolded comp_def])
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Abstract_Impl.thy b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Abstract_Impl.thy
--- a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Abstract_Impl.thy
+++ b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Abstract_Impl.thy
@@ -1,333 +1,334 @@
theory Tree_Automata_Abstract_Impl
imports Tree_Automata_Det Horn_Fset
begin
section \<open>Computing state derivation\<close>
lemma ta_der_Var_code [code]:
"ta_der \<A> (Var q) = finsert q ((eps \<A>)|\<^sup>+| |``| {|q|})"
by (auto)
lemma ta_der_Fun_code [code]:
"ta_der \<A> (Fun f ts) =
(let args = map (ta_der \<A>) ts in
let P = (\<lambda> r. case r of TA_rule g ps p \<Rightarrow> f = g \<and> list_all2 fmember ps args) in
let S = r_rhs |`| ffilter P (rules \<A>) in
S |\<union>| (eps \<A>)|\<^sup>+| |``| S)" (is "?Ls = ?Rs")
proof
{fix q assume "q |\<in>| ?Ls" then have "q |\<in>| ?Rs"
- by (auto simp: Let_def ffmember_filter fimage_iff fBex_def list_all2_conv_all_nth fImage_iff
- split!: ta_rule.splits) force}
+ apply (simp add: Let_def fImage_iff fBex_def image_iff)
+ by (smt (verit, ccfv_threshold) IntI length_map list_all2_conv_all_nth mem_Collect_eq nth_map
+ ta_rule.case ta_rule.sel(3))
+ }
then show "?Ls |\<subseteq>| ?Rs" by blast
next
{fix q assume "q |\<in>| ?Rs" then have "q |\<in>| ?Ls"
apply (auto simp: Let_def ffmember_filter fimage_iff fBex_def list_all2_conv_all_nth fImage_iff
split!: ta_rule.splits)
apply (metis ta_rule.collapse)
apply blast
done}
then show "?Rs |\<subseteq>| ?Ls" by blast
qed
definition eps_free_automata where
"eps_free_automata epscl \<A> =
(let ruleps = (\<lambda> r. finsert (r_rhs r) (epscl |``| {|r_rhs r|})) in
let rules = (\<lambda> r. (\<lambda> q. TA_rule (r_root r) (r_lhs_states r) q) |`| (ruleps r)) |`| (rules \<A>) in
TA ( |\<Union>| rules) {||})"
lemma eps_free [code]:
"eps_free \<A> = eps_free_automata ((eps \<A>)|\<^sup>+|) \<A>"
apply (intro TA_equalityI)
apply (auto simp: eps_free_def eps_free_rulep_def eps_free_automata_def)
using fBex_def apply fastforce
apply (metis ta_rule.exhaust_sel)+
done
lemma eps_of_eps_free_automata [simp]:
"eps (eps_free_automata S \<A>) = {||}"
by (auto simp add: eps_free_automata_def)
lemma eps_free_automata_empty [simp]:
"eps \<A> = {||} \<Longrightarrow> eps_free_automata {||} \<A> = \<A>"
by (auto simp add: eps_free_automata_def intro!: TA_equalityI)
section \<open>Computing the restriction of tree automata to state set\<close>
lemma ta_restrict [code]:
"ta_restrict \<A> Q =
(let rules = ffilter (\<lambda> r. case r of TA_rule f ps p \<Rightarrow> fset_of_list ps |\<subseteq>| Q \<and> p |\<in>| Q) (rules \<A>) in
let eps = ffilter (\<lambda> r. case r of (p, q) \<Rightarrow> p |\<in>| Q \<and> q |\<in>| Q) (eps \<A>) in
TA rules eps)"
by (auto simp: Let_def ta_restrict_def split!: ta_rule.splits intro: finite_subset[OF _ finite_Collect_ta_rule])
section \<open>Computing the epsilon transition for the product automaton\<close>
lemma prod_eps[code_unfold]:
"fCollect (prod_epsLp \<A> \<B>) = (\<lambda> ((p, q), r). ((p, r), (q, r))) |`| (eps \<A> |\<times>| \<Q> \<B>)"
"fCollect (prod_epsRp \<A> \<B>) = (\<lambda> ((p, q), r). ((r, p), (r, q))) |`| (eps \<B> |\<times>| \<Q> \<A>)"
- by (auto simp: finite_prod_epsLp prod_epsLp_def finite_prod_epsRp prod_epsRp_def fimage_iff fBex_def)
+ by (auto simp: finite_prod_epsLp prod_epsLp_def finite_prod_epsRp prod_epsRp_def image_iff
+ fSigma.rep_eq) force+
section \<open>Computing reachability\<close>
inductive_set ta_reach for \<A> where
rule [intro]: "f qs \<rightarrow> q |\<in>| rules \<A> \<Longrightarrow> \<forall> i < length qs. qs ! i \<in> ta_reach \<A> \<Longrightarrow> q \<in> ta_reach \<A>"
| eps [intro]: "q \<in> ta_reach \<A> \<Longrightarrow> (q, r) |\<in>| eps \<A> \<Longrightarrow> r \<in> ta_reach \<A>"
lemma ta_reach_eps_transI:
assumes "(p, q) |\<in>| (eps \<A>)|\<^sup>+|" "p \<in> ta_reach \<A>"
shows "q \<in> ta_reach \<A>" using assms
by (induct rule: ftrancl_induct) auto
lemma ta_reach_ground_term_der:
assumes "q \<in> ta_reach \<A>"
shows "\<exists> t. ground t \<and> q |\<in>| ta_der \<A> t" using assms
proof (induct)
case (rule f qs q)
then obtain ts where "length ts = length qs"
"\<forall> i < length qs. ground (ts ! i)"
"\<forall> i < length qs. qs ! i |\<in>| ta_der \<A> (ts ! i)"
using Ex_list_of_length_P[of "length qs" "\<lambda> t i. ground t \<and> qs ! i |\<in>| ta_der \<A> t"]
by auto
then show ?case using rule(1)
by (auto dest!: in_set_idx intro!: exI[of _ "Fun f ts"]) blast
qed (auto, meson ta_der_eps)
lemma ground_term_der_ta_reach:
assumes "ground t" "q |\<in>| ta_der \<A> t"
shows "q \<in> ta_reach \<A>" using assms(2, 1)
by (induct rule: ta_der_induct) (auto simp add: rule ta_reach_eps_transI)
lemma ta_reach_reachable:
"ta_reach \<A> = fset (ta_reachable \<A>)"
using ta_reach_ground_term_der[of _ \<A>]
using ground_term_der_ta_reach[of _ _ \<A>]
unfolding ta_reachable_def
- by (auto simp flip: fmember_iff_member_fset)
+ by auto
subsection \<open>Horn setup for reachable states\<close>
definition "reach_rules \<A> =
{qs \<rightarrow>\<^sub>h q | f qs q. TA_rule f qs q |\<in>| rules \<A>} \<union>
{[q] \<rightarrow>\<^sub>h r | q r. (q, r) |\<in>| eps \<A>}"
locale reach_horn =
fixes \<A> :: "('q, 'f) ta"
begin
sublocale horn "reach_rules \<A>" .
lemma reach_infer0: "infer0 = {q | f q. TA_rule f [] q |\<in>| rules \<A>}"
by (auto simp: horn.infer0_def reach_rules_def)
lemma reach_infer1:
"infer1 p X = {r | f qs r. TA_rule f qs r |\<in>| rules \<A> \<and> p \<in> set qs \<and> set qs \<subseteq> insert p X} \<union>
{r | r. (p, r) |\<in>| eps \<A>}"
unfolding reach_rules_def
by (auto simp: horn.infer1_def simp flip: ex_simps(1))
lemma reach_sound:
"ta_reach \<A> = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr x) obtain p where x: "p = ta_reach \<A>" by auto
show ?case using lr unfolding x
proof (induct)
case (rule f qs q)
then show ?case
by (intro infer[of qs q]) (auto simp: reach_rules_def dest: in_set_idx)
next
case (eps q r)
then show ?case
by (intro infer[of "[q]" r]) (auto simp: reach_rules_def)
qed
next
case (rl x)
then show ?case
by (induct) (auto simp: reach_rules_def)
qed
end
subsection \<open>Computing productivity\<close>
text \<open>First, use an alternative definition of productivity\<close>
inductive_set ta_productive_ind :: "'q fset \<Rightarrow> ('q,'f) ta \<Rightarrow> 'q set" for P and \<A> :: "('q,'f) ta" where
basic [intro]: "q |\<in>| P \<Longrightarrow> q \<in> ta_productive_ind P \<A>"
| eps [intro]: "(p, q) |\<in>| (eps \<A>)|\<^sup>+| \<Longrightarrow> q \<in> ta_productive_ind P \<A> \<Longrightarrow> p \<in> ta_productive_ind P \<A>"
| rule: "TA_rule f qs q |\<in>| rules \<A> \<Longrightarrow> q \<in> ta_productive_ind P \<A> \<Longrightarrow> q' \<in> set qs \<Longrightarrow> q' \<in> ta_productive_ind P \<A>"
lemma ta_productive_ind:
"ta_productive_ind P \<A> = fset (ta_productive P \<A>)" (is "?LS = ?RS")
proof -
{fix q assume "q \<in> ?LS" then have "q \<in> ?RS"
- by (induct) (auto dest: ta_prod_epsD simp flip: fmember_iff_member_fset intro: ta_productive_setI,
+ by (induct) (auto dest: ta_prod_epsD intro: ta_productive_setI,
metis (full_types) in_set_conv_nth rule_reachable_ctxt_exist ta_productiveI')}
moreover
- {fix q assume "q \<in> ?RS" note * = this[unfolded fmember_iff_member_fset[symmetric]]
+ {fix q assume "q \<in> ?RS" note * = this
from ta_productiveE[OF *] obtain r C where
reach : "r |\<in>| ta_der \<A> C\<langle>Var q\<rangle>" and f: "r |\<in>| P" by auto
from f have "r \<in> ta_productive_ind P \<A>" "r |\<in>| ta_productive P \<A>"
by (auto intro: ta_productive_setI)
then have "q \<in> ?LS" using reach
proof (induct C arbitrary: q r)
case (More f ss C ts)
from iffD1 ta_der_Fun[THEN iffD1, OF More(4)[unfolded ctxt_apply_term.simps]] obtain ps p where
inv: "f ps \<rightarrow> p |\<in>| rules \<A>" "p = r \<or> (p, r) |\<in>| (eps \<A>)|\<^sup>+|" "length ps = length (ss @ C\<langle>Var q\<rangle> # ts)"
"ps ! length ss |\<in>| ta_der \<A> C\<langle>Var q\<rangle>"
by (auto simp: nth_append_Cons split: if_splits)
then have "p \<in> ta_productive_ind P \<A> \<Longrightarrow> p |\<in>| ta_der \<A> C\<langle>Var q\<rangle> \<Longrightarrow> q \<in> ta_productive_ind P \<A>" for p
- using More(1) calculation by (auto simp flip: fmember_iff_member_fset)
+ using More(1) calculation by auto
note [intro!] = this[of "ps ! length ss"]
show ?case using More(2) inv
- by (auto simp flip: fmember_iff_member_fset simp: nth_append_Cons ta_productive_ind.rule)
+ by (auto simp: nth_append_Cons ta_productive_ind.rule)
(metis less_add_Suc1 nth_mem ta_productive_ind.simps)
qed (auto intro: ta_productive_setI)
}
ultimately show ?thesis by auto
qed
subsubsection \<open>Horn setup for productive states\<close>
definition "productive_rules P \<A> = {[] \<rightarrow>\<^sub>h q | q. q |\<in>| P} \<union>
{[r] \<rightarrow>\<^sub>h q | q r. (q, r) |\<in>| eps \<A>} \<union>
{[q] \<rightarrow>\<^sub>h r | f qs q r. TA_rule f qs q |\<in>| rules \<A> \<and> r \<in> set qs}"
locale productive_horn =
fixes \<A> :: "('q, 'f) ta" and P :: "'q fset"
begin
sublocale horn "productive_rules P \<A>" .
lemma productive_infer0: "infer0 = fset P"
- by (auto simp: productive_rules_def horn.infer0_def simp flip: fmember_iff_member_fset)
+ by (auto simp: productive_rules_def horn.infer0_def)
lemma productive_infer1:
"infer1 p X = {r | r. (r, p) |\<in>| eps \<A>} \<union>
{r | f qs r. TA_rule f qs p |\<in>| rules \<A> \<and> r \<in> set qs}"
unfolding productive_rules_def horn_infer1_union
by (auto simp add: horn.infer1_def)
(metis insertCI list.set(1) list.simps(15) singletonD subsetI)
lemma productive_sound:
"ta_productive_ind P \<A> = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr p) then show ?case using lr
proof (induct)
case (basic q)
then show ?case
by (intro infer[of "[]" q]) (auto simp: productive_rules_def)
next
case (eps p q) then show ?case
proof (induct rule: ftrancl_induct)
case (Base p q)
then show ?case using infer[of "[q]" p]
by (auto simp: productive_rules_def)
next
case (Step p q r)
then show ?case using infer[of "[r]" q]
by (auto simp: productive_rules_def)
qed
next
case (rule f qs q p)
then show ?case
by (intro infer[of "[q]" p]) (auto simp: productive_rules_def)
qed
next
case (rl p)
then show ?case
by (induct) (auto simp: productive_rules_def ta_productive_ind.rule)
qed
end
subsection \<open>Horn setup for power set construction states\<close>
lemma prod_list_exists:
assumes "fst p \<in> set qs" "set qs \<subseteq> insert (fst p) (fst ` X)"
obtains as where "p \<in> set as" "map fst as = qs" "set as \<subseteq> insert p X"
proof -
from assms have "qs \<in> lists (fst ` (insert p X))" by blast
then obtain ts where ts: "map fst ts = qs" "ts \<in> lists (insert p X)"
by (metis image_iff lists_image)
then obtain i where mem: "i < length qs" "qs ! i = fst p" using assms(1)
by (metis in_set_idx)
from ts have p: "ts[i := p] \<in> lists (insert p X)"
using set_update_subset_insert by fastforce
then have "p \<in> set (ts[i := p])" "map fst (ts[i := p]) = qs" "set (ts[i := p]) \<subseteq> insert p X"
using mem ts(1)
by (auto simp add: nth_list_update set_update_memI intro!: nth_equalityI)
then show ?thesis using that
by blast
qed
definition "ps_states_rules \<A> = {rs \<rightarrow>\<^sub>h (Wrapp q) | rs f q.
q = ps_reachable_states \<A> f (map ex rs) \<and> q \<noteq> {||}}"
locale ps_states_horn =
fixes \<A> :: "('q, 'f) ta"
begin
sublocale horn "ps_states_rules \<A>" .
lemma ps_construction_infer0: "infer0 =
{Wrapp q | f q. q = ps_reachable_states \<A> f [] \<and> q \<noteq> {||}}"
- by (auto simp: ps_states_rules_def horn.infer0_def simp flip: fmember_iff_member_fset)
+ by (auto simp: ps_states_rules_def horn.infer0_def)
lemma ps_construction_infer1:
"infer1 p X = {Wrapp q | f qs q. q = ps_reachable_states \<A> f (map ex qs) \<and> q \<noteq> {||} \<and>
p \<in> set qs \<and> set qs \<subseteq> insert p X}"
unfolding ps_states_rules_def horn_infer1_union
by (auto simp add: horn.infer1_def ps_reachable_states_def comp_def elim!: prod_list_exists)
lemma ps_states_sound:
"ps_states_set \<A> = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
case (lr p) then show ?case using lr
proof (induct)
case (1 ps f)
then have "ps \<rightarrow>\<^sub>h (Wrapp (ps_reachable_states \<A> f (map ex ps))) \<in> ps_states_rules \<A>"
by (auto simp: ps_states_rules_def)
then show ?case using horn.saturate.simps 1
by fastforce
qed
next
case (rl p)
then obtain q where "q \<in> saturate" "q = p" by blast
then show ?case
by (induct arbitrary: p)
(auto simp: ps_states_rules_def intro!: ps_states_set.intros)
qed
end
definition ps_reachable_states_cont where
"ps_reachable_states_cont \<Delta> \<Delta>\<^sub>\<epsilon> f ps =
(let R = ffilter (\<lambda> r. case r of TA_rule g qs q \<Rightarrow> f = g \<and> list_all2 (|\<in>|) qs ps) \<Delta> in
let S = r_rhs |`| R in
S |\<union>| \<Delta>\<^sub>\<epsilon>|\<^sup>+| |``| S)"
lemma ps_reachable_states [code]:
"ps_reachable_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>) f ps = ps_reachable_states_cont \<Delta> \<Delta>\<^sub>\<epsilon> f ps"
by (auto simp: ps_reachable_states_fmember ps_reachable_states_cont_def Let_def fimage_iff fBex_def
split!: ta_rule.splits) force+
definition ps_rules_cont where
"ps_rules_cont \<A> Q =
(let sig = ta_sig \<A> in
let qss = (\<lambda> (f, n). (f, n, fset_of_list (List.n_lists n (sorted_list_of_fset Q)))) |`| sig in
let res = (\<lambda> (f, n, Qs). (\<lambda> qs. TA_rule f qs (Wrapp (ps_reachable_states \<A> f (map ex qs)))) |`| Qs) |`| qss in
ffilter (\<lambda> r. ex (r_rhs r) \<noteq> {||}) ( |\<Union>| res))"
lemma ps_rules [code]:
"ps_rules \<A> Q = ps_rules_cont \<A> Q"
using ps_reachable_states_sig finite_ps_rulesp_unfolded[of Q \<A>]
unfolding ps_rules_cont_def
apply (auto simp: fset_of_list_elem ps_rules_def fin_mono ps_rulesp_def
- fimage_iff set_n_lists simp flip: fmember_iff_member_fset split!: prod.splits dest!: in_set_idx)
- apply fastforce
- apply (meson fmember_iff_member_fset nth_mem subsetD)
- done
+ image_iff set_n_lists split!: prod.splits dest!: in_set_idx)
+ by fastforce+
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Complement.thy b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Complement.thy
--- a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Complement.thy
+++ b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Complement.thy
@@ -1,193 +1,192 @@
theory Tree_Automata_Complement
imports Tree_Automata_Det
begin
subsection \<open>Complement closure of regular languages\<close>
definition partially_completely_defined_on where
"partially_completely_defined_on \<A> \<F> \<longleftrightarrow>
(\<forall> t. funas_gterm t \<subseteq> fset \<F> \<longleftrightarrow> (\<exists> q. q |\<in>| ta_der \<A> (term_of_gterm t)))"
definition sig_ta where
"sig_ta \<F> = TA ((\<lambda> (f, n). TA_rule f (replicate n ()) ()) |`| \<F>) {||}"
lemma sig_ta_rules_fmember:
"TA_rule f qs q |\<in>| rules (sig_ta \<F>) \<longleftrightarrow> (\<exists> n. (f, n) |\<in>| \<F> \<and> qs = replicate n () \<and> q = ())"
by (auto simp: sig_ta_def fimage_iff fBex_def)
lemma sig_ta_completely_defined:
"partially_completely_defined_on (sig_ta \<F>) \<F>"
proof -
{fix t assume "funas_gterm t \<subseteq> fset \<F>"
then have "() |\<in>| ta_der (sig_ta \<F>) (term_of_gterm t)"
proof (induct t)
case (GFun f ts)
then show ?case
by (auto simp: sig_ta_rules_fmember SUP_le_iff
- simp flip: fmember_iff_member_fset intro!: exI[of _ "replicate (length ts) ()"])
+ intro!: exI[of _ "replicate (length ts) ()"])
qed}
moreover
{fix t q assume "q |\<in>| ta_der (sig_ta \<F>) (term_of_gterm t)"
then have "funas_gterm t \<subseteq> fset \<F>"
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
from GFun(1 - 4) GFun(5)[THEN subsetD] show ?case
- by (auto simp: sig_ta_rules_fmember simp flip: fmember_iff_member_fset dest!: in_set_idx)
+ by (auto simp: sig_ta_rules_fmember dest!: in_set_idx)
qed}
ultimately show ?thesis
unfolding partially_completely_defined_on_def
by blast
qed
lemma ta_der_fsubset_sig_ta_completely:
assumes "ta_subset (sig_ta \<F>) \<A>" "ta_sig \<A> |\<subseteq>| \<F>"
shows "partially_completely_defined_on \<A> \<F>"
proof -
have "ta_der (sig_ta \<F>) t |\<subseteq>| ta_der \<A> t" for t
using assms by (simp add: ta_der_mono')
then show ?thesis using sig_ta_completely_defined assms(2)
by (auto simp: partially_completely_defined_on_def)
- (metis ffunas_gterm.rep_eq fin_mono notin_fset ta_der_gterm_sig)
+ (metis ffunas_gterm.rep_eq fin_mono ta_der_gterm_sig)
qed
lemma completely_definied_ps_taI:
"partially_completely_defined_on \<A> \<F> \<Longrightarrow> partially_completely_defined_on (ps_ta \<A>) \<F>"
unfolding partially_completely_defined_on_def
using ps_rules_not_empty_reach[of \<A>]
using fsubsetD[OF ps_rules_sound[of _ \<A>]] ps_rules_complete[of _ \<A>]
by (metis FSet_Lex_Wrapper.collapse fsubsetI fsubset_fempty)
lemma completely_definied_ta_union1I:
"partially_completely_defined_on \<A> \<F> \<Longrightarrow> ta_sig \<B> |\<subseteq>| \<F> \<Longrightarrow> \<Q> \<A> |\<inter>| \<Q> \<B> = {||} \<Longrightarrow>
partially_completely_defined_on (ta_union \<A> \<B>) \<F>"
unfolding partially_completely_defined_on_def
using ta_union_der_disj_states'[of \<A> \<B>]
by (auto simp: ta_union_der_disj_states)
(metis ffunas_gterm.rep_eq fsubset_trans less_eq_fset.rep_eq ta_der_gterm_sig)
lemma completely_definied_fmaps_statesI:
"partially_completely_defined_on \<A> \<F> \<Longrightarrow> finj_on f (\<Q> \<A>) \<Longrightarrow> partially_completely_defined_on (fmap_states_ta f \<A>) \<F>"
unfolding partially_completely_defined_on_def
using fsubsetD[OF ta_der_fmap_states_ta_mono2, of f \<A>]
using ta_der_to_fmap_states_der[of _ \<A> _ f]
by (auto simp: fimage_iff fBex_def) fastforce+
lemma det_completely_defined_complement:
assumes "partially_completely_defined_on \<A> \<F>" "ta_det \<A>"
shows "gta_lang (\<Q> \<A> |-| Q) \<A> = \<T>\<^sub>G (fset \<F>) - gta_lang Q \<A>" (is "?Ls = ?Rs")
proof -
{fix t assume "t \<in> ?Ls"
then obtain p where p: "p |\<in>| \<Q> \<A>" "p |\<notin>| Q" "p |\<in>| ta_der \<A> (term_of_gterm t)"
by auto
from ta_detE[OF assms(2) p(3)] have "\<forall> q. q |\<in>| ta_der \<A> (term_of_gterm t) \<longrightarrow> q = p"
by blast
moreover have "funas_gterm t \<subseteq> fset \<F>"
using p(3) assms(1) unfolding partially_completely_defined_on_def
by (auto simp: less_eq_fset.rep_eq ffunas_gterm.rep_eq)
ultimately have "t \<in> ?Rs" using p(2)
by (auto simp: \<T>\<^sub>G_equivalent_def)}
moreover
{fix t assume "t \<in> ?Rs"
then have f: "funas_gterm t \<subseteq> fset \<F>" "\<forall> q. q |\<in>| ta_der \<A> (term_of_gterm t) \<longrightarrow> q |\<notin>| Q"
by (auto simp: \<T>\<^sub>G_equivalent_def)
from f(1) obtain p where "p |\<in>| ta_der \<A> (term_of_gterm t)" using assms(1)
by (force simp: partially_completely_defined_on_def)
then have "t \<in> ?Ls" using f(2)
by (auto simp: gterm_ta_der_states intro: gta_langI[of p])}
ultimately show ?thesis by blast
qed
lemma ta_der_gterm_sig_fset:
"q |\<in>| ta_der \<A> (term_of_gterm t) \<Longrightarrow> funas_gterm t \<subseteq> fset (ta_sig \<A>)"
using ta_der_gterm_sig
by (metis ffunas_gterm.rep_eq less_eq_fset.rep_eq)
definition filter_ta_sig where
"filter_ta_sig \<F> \<A> = TA (ffilter (\<lambda> r. (r_root r, length (r_lhs_states r)) |\<in>| \<F>) (rules \<A>)) (eps \<A>)"
definition filter_ta_reg where
"filter_ta_reg \<F> R = Reg (fin R) (filter_ta_sig \<F> (ta R))"
lemma filter_ta_sig:
"ta_sig (filter_ta_sig \<F> \<A>) |\<subseteq>| \<F>"
by (auto simp: ta_sig_def filter_ta_sig_def)
lemma filter_ta_sig_lang:
"gta_lang Q (filter_ta_sig \<F> \<A>) = gta_lang Q \<A> \<inter> \<T>\<^sub>G (fset \<F>)" (is "?Ls = ?Rs")
proof -
let ?A = "filter_ta_sig \<F> \<A>"
{fix t assume "t \<in> ?Ls"
then obtain q where q: "q |\<in>| Q" "q |\<in>| ta_der ?A (term_of_gterm t)"
by auto
then have "funas_gterm t \<subseteq> fset \<F>"
using subset_trans[OF ta_der_gterm_sig_fset[OF q(2)] filter_ta_sig[unfolded less_eq_fset.rep_eq]]
by blast
then have "t \<in> ?Rs" using q
by (auto simp: \<T>\<^sub>G_equivalent_def filter_ta_sig_def
intro!: gta_langI[of q] ta_der_el_mono[where ?q = q and \<B> = \<A> and \<A> = ?A])}
moreover
{fix t assume ass: "t \<in> ?Rs"
then have funas: "funas_gterm t \<subseteq> fset \<F>"
by (auto simp: \<T>\<^sub>G_equivalent_def)
from ass obtain p where p: "p |\<in>| Q" "p |\<in>| ta_der \<A> (term_of_gterm t)"
by auto
from this(2) funas have "p |\<in>| ta_der ?A (term_of_gterm t)"
proof (induct rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
then show ?case
- by (auto simp: filter_ta_sig_def SUP_le_iff simp flip: fmember_iff_member_fset
- intro!: exI[of _ ps] exI[of _ p])
+ by (auto simp: filter_ta_sig_def SUP_le_iff intro!: exI[of _ ps] exI[of _ p])
qed
then have "t \<in> ?Ls" using p(1) by auto}
ultimately show ?thesis by blast
qed
lemma \<L>_filter_ta_reg:
"\<L> (filter_ta_reg \<F> \<A>) = \<L> \<A> \<inter> \<T>\<^sub>G (fset \<F>)"
using filter_ta_sig_lang
by (auto simp: \<L>_def filter_ta_reg_def)
definition sig_ta_reg where
"sig_ta_reg \<F> = Reg {||} (sig_ta \<F>)"
lemma \<L>_sig_ta_reg:
"\<L> (sig_ta_reg \<F>) = {}"
by (auto simp: \<L>_def sig_ta_reg_def)
definition complement_reg where
"complement_reg R \<F> = (let \<A> = ps_reg (reg_union (sig_ta_reg \<F>) R) in
Reg (\<Q>\<^sub>r \<A> |-| fin \<A>) (ta \<A>))"
lemma \<L>_complement_reg:
assumes "ta_sig (ta \<A>) |\<subseteq>| \<F>"
shows "\<L> (complement_reg \<A> \<F>) = \<T>\<^sub>G (fset \<F>) - \<L> \<A>"
proof -
have "\<L> (complement_reg \<A> \<F>) = \<T>\<^sub>G (fset \<F>) - \<L> (ps_reg (reg_union (sig_ta_reg \<F>) \<A>))"
unfolding \<L>_def complement_reg_def using assms
by (auto simp: complement_reg_def Let_def ps_reg_def reg_union_def sig_ta_reg_def
sig_ta_completely_defined finj_Inl_Inr
intro!: det_completely_defined_complement completely_definied_ps_taI
completely_definied_ta_union1I completely_definied_fmaps_statesI)
then show ?thesis
by (auto simp: \<L>_ps_reg \<L>_union \<L>_sig_ta_reg)
qed
lemma \<L>_complement_filter_reg:
"\<L> (complement_reg (filter_ta_reg \<F> \<A>) \<F>) = \<T>\<^sub>G (fset \<F>) - \<L> \<A>"
proof -
have *: "ta_sig (ta (filter_ta_reg \<F> \<A>)) |\<subseteq>| \<F>"
by (auto simp: filter_ta_reg_def filter_ta_sig)
show ?thesis unfolding \<L>_complement_reg[OF *] \<L>_filter_ta_reg
by blast
qed
definition difference_reg where
"difference_reg R L = (let F = ta_sig (ta R) in
reg_intersect R (trim_reg (complement_reg (filter_ta_reg F L) F)))"
lemma \<L>_difference_reg:
"\<L> (difference_reg R L) = \<L> R - \<L> L" (is "?Ls = ?Rs")
unfolding difference_reg_def Let_def \<L>_trim \<L>_intersect \<L>_complement_filter_reg
using reg_funas by blast
end
\ No newline at end of file
diff --git a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Det.thy b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Det.thy
--- a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Det.thy
+++ b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Det.thy
@@ -1,268 +1,268 @@
theory Tree_Automata_Det
imports
Tree_Automata
begin
subsection \<open>Powerset Construction for Tree Automata\<close>
text \<open>
The idea to treat states and transitions separately is from arXiv:1511.03595. Some parts of
the implementation are also based on that paper. (The Algorithm corresponds roughly to the one
in "Step 5")
\<close>
text \<open>Abstract Definitions and Correctness Proof\<close>
definition ps_reachable_statesp where
"ps_reachable_statesp \<A> f ps = (\<lambda> q'. \<exists> qs q. TA_rule f qs q |\<in>| rules \<A> \<and> list_all2 (|\<in>|) qs ps \<and>
(q = q' \<or> (q,q') |\<in>| (eps \<A>)|\<^sup>+|))"
lemma ps_reachable_statespE:
assumes "ps_reachable_statesp \<A> f qs q"
obtains ps p where "TA_rule f ps p |\<in>| rules \<A>" "list_all2 (|\<in>|) ps qs" "(p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|)"
using assms unfolding ps_reachable_statesp_def
by auto
lemma ps_reachable_statesp_\<Q>:
"ps_reachable_statesp \<A> f ps q \<Longrightarrow> q |\<in>| \<Q> \<A>"
- by (auto simp: ps_reachable_statesp_def simp flip: fmember_iff_member_fset dest: rule_statesD eps_trancl_statesD)
+ by (auto simp: ps_reachable_statesp_def simp flip: dest: rule_statesD eps_trancl_statesD)
lemma finite_Collect_ps_statep [simp]:
"finite (Collect (ps_reachable_statesp \<A> f ps))" (is "finite ?S")
by (intro finite_subset[of ?S "fset (\<Q> \<A>)"])
- (auto simp: ps_reachable_statesp_\<Q> simp flip: fmember_iff_member_fset)
+ (auto simp: ps_reachable_statesp_\<Q>)
lemmas finite_Collect_ps_statep_unfolded [simp] = finite_Collect_ps_statep[unfolded ps_reachable_statesp_def, simplified]
definition "ps_reachable_states \<A> f ps \<equiv> fCollect (ps_reachable_statesp \<A> f ps)"
lemmas ps_reachable_states_simp = ps_reachable_statesp_def ps_reachable_states_def
lemma ps_reachable_states_fmember:
"q' |\<in>| ps_reachable_states \<A> f ps \<longleftrightarrow> (\<exists> qs q. TA_rule f qs q |\<in>| rules \<A> \<and>
list_all2 (|\<in>|) qs ps \<and> (q = q' \<or> (q, q') |\<in>| (eps \<A>)|\<^sup>+|))"
by (auto simp: ps_reachable_states_simp)
lemma ps_reachable_statesI:
assumes "TA_rule f ps p |\<in>| rules \<A>" "list_all2 (|\<in>|) ps qs" "(p = q \<or> (p, q) |\<in>| (eps \<A>)|\<^sup>+|)"
shows "p |\<in>| ps_reachable_states \<A> f qs"
using assms unfolding ps_reachable_states_simp
by auto
lemma ps_reachable_states_sig:
"ps_reachable_states \<A> f ps \<noteq> {||} \<Longrightarrow> (f, length ps) |\<in>| ta_sig \<A>"
- by (auto simp: ps_reachable_states_simp ta_sig_def fimage_iff fBex_def dest!: list_all2_lengthD)
+ by (auto simp: ps_reachable_states_simp ta_sig_def image_iff intro!: bexI dest!: list_all2_lengthD)
text \<open>
A set of "powerset states" is complete if it is sufficient to capture all (non)deterministic
derivations.
\<close>
definition ps_states_complete_it :: "('a, 'b) ta \<Rightarrow> 'a FSet_Lex_Wrapper fset \<Rightarrow> 'a FSet_Lex_Wrapper fset \<Rightarrow> bool"
where "ps_states_complete_it \<A> Q Qnext \<equiv>
\<forall>f ps. fset_of_list ps |\<subseteq>| Q \<and> ps_reachable_states \<A> f (map ex ps) \<noteq> {||} \<longrightarrow> Wrapp (ps_reachable_states \<A> f (map ex ps)) |\<in>| Qnext"
lemma ps_states_complete_itD:
"ps_states_complete_it \<A> Q Qnext \<Longrightarrow> fset_of_list ps |\<subseteq>| Q \<Longrightarrow>
ps_reachable_states \<A> f (map ex ps) \<noteq> {||} \<Longrightarrow> Wrapp (ps_reachable_states \<A> f (map ex ps)) |\<in>| Qnext"
unfolding ps_states_complete_it_def by blast
abbreviation "ps_states_complete \<A> Q \<equiv> ps_states_complete_it \<A> Q Q"
text \<open>The least complete set of states\<close>
inductive_set ps_states_set for \<A> where
"\<forall> p \<in> set ps. p \<in> ps_states_set \<A> \<Longrightarrow> ps_reachable_states \<A> f (map ex ps) \<noteq> {||} \<Longrightarrow>
Wrapp (ps_reachable_states \<A> f (map ex ps)) \<in> ps_states_set \<A>"
lemma ps_states_Pow:
"ps_states_set \<A> \<subseteq> fset (Wrapp |`| fPow (\<Q> \<A>))"
proof -
{fix q assume "q \<in> ps_states_set \<A>" then have "q \<in> fset (Wrapp |`| fPow (\<Q> \<A>))"
- by induct (auto simp: ps_reachable_statesp_\<Q> ps_reachable_states_def image_iff simp flip: fmember_iff_member_fset)}
+ by induct (auto simp: ps_reachable_statesp_\<Q> ps_reachable_states_def image_iff)}
then show ?thesis by blast
qed
context
includes fset.lifting
begin
lift_definition ps_states :: "('a, 'b) ta \<Rightarrow> 'a FSet_Lex_Wrapper fset" is ps_states_set
by (auto intro: finite_subset[OF ps_states_Pow])
lemma ps_states: "ps_states \<A> |\<subseteq>| Wrapp |`| fPow (\<Q> \<A>)" using ps_states_Pow
by (simp add: ps_states_Pow less_eq_fset.rep_eq ps_states.rep_eq)
lemmas ps_states_cases = ps_states_set.cases[Transfer.transferred]
lemmas ps_states_induct = ps_states_set.induct[Transfer.transferred]
lemmas ps_states_simps = ps_states_set.simps[Transfer.transferred]
lemmas ps_states_intros= ps_states_set.intros[Transfer.transferred]
end
lemma ps_states_complete:
"ps_states_complete \<A> (ps_states \<A>)"
unfolding ps_states_complete_it_def
by (auto intro: ps_states_intros)
lemma ps_states_least_complete:
assumes "ps_states_complete_it \<A> Q Qnext" "Qnext |\<subseteq>| Q"
shows "ps_states \<A> |\<subseteq>| Q"
proof standard
fix q assume ass: "q |\<in>| ps_states \<A>" then show "q |\<in>| Q"
using ps_states_complete_itD[OF assms(1)] fsubsetD[OF assms(2)]
by (induct rule: ps_states_induct[of _ \<A>]) (fastforce intro: ass)+
qed
definition ps_rulesp :: "('a, 'b) ta \<Rightarrow> 'a FSet_Lex_Wrapper fset \<Rightarrow> ('a FSet_Lex_Wrapper, 'b) ta_rule \<Rightarrow> bool" where
"ps_rulesp \<A> Q = (\<lambda> r. \<exists> f ps p. r = TA_rule f ps (Wrapp p) \<and> fset_of_list ps |\<subseteq>| Q \<and>
p = ps_reachable_states \<A> f (map ex ps) \<and> p \<noteq> {||})"
definition "ps_rules" where
"ps_rules \<A> Q \<equiv> fCollect (ps_rulesp \<A> Q)"
lemma finite_ps_rulesp [simp]:
"finite (Collect (ps_rulesp \<A> Q))" (is "finite ?S")
proof -
let ?Q = "fset (Wrapp |`| fPow (\<Q> \<A>) |\<union>| Q)" let ?sig = "fset (ta_sig \<A>)"
define args where "args \<equiv> \<Union> (f,n) \<in> ?sig. {qs| qs. set qs \<subseteq> ?Q \<and> length qs = n}"
define bound where "bound \<equiv> \<Union>(f,_) \<in> ?sig. \<Union>q \<in> ?Q. \<Union>qs \<in> args. {TA_rule f qs q}"
have finite: "finite ?Q" "finite ?sig" by (auto intro: finite_subset)
then have "finite args" using finite_lists_length_eq[OF \<open>finite ?Q\<close>]
by (force simp: args_def)
with finite have "finite bound" unfolding bound_def by (auto simp only: finite_UN)
moreover have "Collect (ps_rulesp \<A> Q) \<subseteq> bound"
proof standard
fix r assume *: "r \<in> Collect (ps_rulesp \<A> Q)"
obtain f ps p where r[simp]: "r = TA_rule f ps p" by (cases r)
from * obtain qs q where "TA_rule f qs q |\<in>| rules \<A>" and len: "length ps = length qs"
unfolding ps_rulesp_def ps_reachable_states_simp
using list_all2_lengthD by fastforce
from this have sym: "(f, length qs) \<in> ?sig"
- by (auto simp flip: fmember_iff_member_fset)
+ by auto
moreover from * have "set ps \<subseteq> ?Q" unfolding ps_rulesp_def
- by (auto simp flip: fset_of_list_elem fmember_iff_member_fset simp: ps_reachable_statesp_def)
+ by (auto simp flip: fset_of_list_elem simp: ps_reachable_statesp_def)
ultimately have ps: "ps \<in> args"
by (auto simp only: args_def UN_iff intro!: bexI[of _ "(f, length qs)"] len)
from * have "p \<in> ?Q" unfolding ps_rulesp_def ps_reachable_states_def
- using fmember_iff_member_fset ps_reachable_statesp_\<Q>
+ using ps_reachable_statesp_\<Q>
by (fastforce simp add: image_iff)
with ps sym show "r \<in> bound"
by (auto simp only: r bound_def UN_iff intro!: bexI[of _ "(f, length qs)"] bexI[of _ "p"] bexI[of _ "ps"])
qed
ultimately show ?thesis by (blast intro: finite_subset)
qed
lemmas finite_ps_rulesp_unfolded = finite_ps_rulesp[unfolded ps_rulesp_def, simplified]
lemmas ps_rulesI [intro!] = fCollect_memberI[OF finite_ps_rulesp]
lemma ps_rules_states:
"rule_states (fCollect (ps_rulesp \<A> Q)) |\<subseteq>| (Wrapp |`| fPow (\<Q> \<A>) |\<union>| Q)"
by (auto simp: ps_rulesp_def rule_states_def ps_reachable_states_def ps_reachable_statesp_\<Q>) blast
definition ps_ta :: "('q, 'f) ta \<Rightarrow> ('q FSet_Lex_Wrapper, 'f) ta" where
"ps_ta \<A> = (let Q = ps_states \<A> in
TA (ps_rules \<A> Q) {||})"
definition ps_ta_Q\<^sub>f :: "'q fset \<Rightarrow> ('q, 'f) ta \<Rightarrow> 'q FSet_Lex_Wrapper fset" where
"ps_ta_Q\<^sub>f Q \<A> = (let Q' = ps_states \<A> in
ffilter (\<lambda> S. Q |\<inter>| (ex S) \<noteq> {||}) Q')"
lemma ps_rules_sound:
assumes "p |\<in>| ta_der (ps_ta \<A>) (term_of_gterm t)"
shows "ex p |\<subseteq>| ta_der \<A> (term_of_gterm t)" using assms
proof (induction rule: ta_der_gterm_induct)
case (GFun f ts ps p q)
then have IH: "\<forall>i < length ts. ex (ps ! i) |\<subseteq>| gta_der \<A> (ts ! i)" unfolding gta_der_def by auto
show ?case
proof standard
fix r assume "r |\<in>| ex q"
with GFun(1 - 3) obtain qs q' where "TA_rule f qs q' |\<in>| rules \<A>"
"q' = r \<or> (q', r) |\<in>| (eps \<A>)|\<^sup>+|" "list_all2 (|\<in>|) qs (map ex ps)"
by (auto simp: Let_def ps_ta_def ps_rulesp_def ps_reachable_states_simp ps_rules_def)
then show "r |\<in>| ta_der \<A> (term_of_gterm (GFun f ts))"
using GFun(2) IH unfolding gta_der_def
by (force dest!: fsubsetD intro!: exI[of _ q'] exI[of _ qs] simp: list_all2_conv_all_nth)
qed
qed
lemma ps_ta_nt_empty_set:
"TA_rule f qs (Wrapp {||}) |\<in>| rules (ps_ta \<A>) \<Longrightarrow> False"
by (auto simp: ps_ta_def ps_rulesp_def ps_rules_def)
lemma ps_rules_not_empty_reach:
assumes "Wrapp {||} |\<in>| ta_der (ps_ta \<A>) (term_of_gterm t)"
shows False using assms
proof (induction t)
case (GFun f ts)
then show ?case using ps_ta_nt_empty_set[of f _ \<A>]
by (auto simp: ps_ta_def)
qed
lemma ps_rules_complete:
assumes "q |\<in>| ta_der \<A> (term_of_gterm t)"
shows "\<exists>p. q |\<in>| ex p \<and> p |\<in>| ta_der (ps_ta \<A>) (term_of_gterm t) \<and> p |\<in>| ps_states \<A>" using assms
proof (induction rule: ta_der_gterm_induct)
let ?P = "\<lambda>t q p. q |\<in>| ex p \<and> p |\<in>| ta_der (ps_ta \<A>) (term_of_gterm t) \<and> p |\<in>| ps_states \<A>"
case (GFun f ts ps p q)
then have "\<forall>i. \<exists>p. i < length ts \<longrightarrow> ?P (ts ! i) (ps ! i) p" by auto
with choice[OF this] obtain psf where ps: "\<forall>i < length ts. ?P (ts ! i) (ps ! i) (psf i)" by auto
define qs where "qs = map psf [0 ..< length ts]"
let ?p = "ps_reachable_states \<A> f (map ex qs)"
from ps have in_Q: "fset_of_list qs |\<subseteq>| ps_states \<A>"
by (auto simp: qs_def fset_of_list_elem)
from ps GFun(2) have all: "list_all2 (|\<in>|) ps (map ex qs)"
by (auto simp: list_all2_conv_all_nth qs_def)
then have in_p: "q |\<in>| ?p" using GFun(1, 3)
unfolding ps_reachable_statesp_def ps_reachable_states_def by auto
then have rule: "TA_rule f qs (Wrapp ?p) |\<in>| ps_rules \<A> (ps_states \<A>)" using in_Q unfolding ps_rules_def
by (intro ps_rulesI) (auto simp: ps_rulesp_def)
from in_Q in_p have "Wrapp ?p |\<in>| (ps_states \<A>)"
by (auto intro!: ps_states_complete[unfolded ps_states_complete_it_def, rule_format])
with in_p ps rule show ?case
by (auto intro!: exI[of _ "Wrapp ?p"] exI[of _ qs] simp: ps_ta_def qs_def)
qed
lemma ps_ta_eps[simp]: "eps (ps_ta \<A>) = {||}" by (auto simp: Let_def ps_ta_def)
lemma ps_ta_det[iff]: "ta_det (ps_ta \<A>)" by (auto simp: Let_def ps_ta_def ta_det_def ps_rulesp_def ps_rules_def)
lemma ps_gta_lang:
"gta_lang (ps_ta_Q\<^sub>f Q \<A>) (ps_ta \<A>) = gta_lang Q \<A>" (is "?R = ?L")
proof standard
show "?L \<subseteq> ?R" proof standard
fix t assume "t \<in> ?L"
then obtain q where q_res: "q |\<in>| ta_der \<A> (term_of_gterm t)" and q_final: "q |\<in>| Q"
by auto
from ps_rules_complete[OF q_res] obtain p where
"p |\<in>| ps_states \<A>" "q |\<in>| ex p" "p |\<in>| ta_der (ps_ta \<A>) (term_of_gterm t)"
by auto
moreover with q_final have "p |\<in>| ps_ta_Q\<^sub>f Q \<A>"
by (auto simp: ps_ta_Q\<^sub>f_def)
ultimately show "t \<in> ?R" by auto
qed
show "?R \<subseteq> ?L" proof standard
fix t assume "t \<in> ?R"
then obtain p where
p_res: "p |\<in>| ta_der (ps_ta \<A>) (term_of_gterm t)" and p_final: "p |\<in>| ps_ta_Q\<^sub>f Q \<A>"
by (auto simp: ta_lang_def)
from ps_rules_sound[OF p_res] have "ex p |\<subseteq>| ta_der \<A> (term_of_gterm t)"
by auto
moreover from p_final obtain q where "q |\<in>| ex p" "q |\<in>| Q" by (auto simp: ps_ta_Q\<^sub>f_def)
ultimately show "t \<in> ?L" by auto
qed
qed
definition ps_reg where
"ps_reg R = Reg (ps_ta_Q\<^sub>f (fin R) (ta R)) (ps_ta (ta R))"
lemma \<L>_ps_reg:
"\<L> (ps_reg R) = \<L> R"
by (auto simp: \<L>_def ps_gta_lang ps_reg_def)
lemma ps_ta_states: "\<Q> (ps_ta \<A>) |\<subseteq>| Wrapp |`| fPow (\<Q> \<A>)"
using ps_rules_states ps_states unfolding ps_ta_def \<Q>_def
- by (auto simp: Let_def ps_rules_def) blast
+ by (auto simp: Let_def ps_rules_def) force
lemma ps_ta_states': "ex |`| \<Q> (ps_ta \<A>) |\<subseteq>| fPow (\<Q> \<A>)"
using ps_ta_states[of \<A>]
by fastforce
end
diff --git a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Impl.thy b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Impl.thy
--- a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Impl.thy
+++ b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata_Impl.thy
@@ -1,406 +1,402 @@
theory Tree_Automata_Impl
imports Tree_Automata_Abstract_Impl
"HOL-Library.List_Lexorder"
"HOL-Library.AList_Mapping"
Tree_Automata_Class_Instances_Impl
Containers.Containers
begin
definition map_val_of_list :: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'c list) \<Rightarrow> 'b list \<Rightarrow> ('a, 'c list) mapping" where
"map_val_of_list ek ev xs = foldr (\<lambda>x m. Mapping.update (ek x) (ev x @ case_option Nil id (Mapping.lookup m (ek x))) m) xs Mapping.empty"
abbreviation "map_of_list ek ev xs \<equiv> map_val_of_list ek (\<lambda> x. [ev x]) xs"
lemma map_val_of_list_tabulate_conv:
"map_val_of_list ek ev xs = Mapping.tabulate (sort (remdups (map ek xs))) (\<lambda> k. concat (map ev (filter (\<lambda> x. k = ek x) xs)))"
unfolding map_val_of_list_def
proof (induct xs)
case (Cons x xs) then show ?case
by (intro mapping_eqI) (auto simp: lookup_combine lookup_update' lookup_empty lookup_tabulate image_iff)
qed (simp add: empty_Mapping tabulate_Mapping)
lemmas map_val_of_list_simp = map_val_of_list_tabulate_conv lookup_tabulate
subsection \<open>Setup for the list implementation of reachable states\<close>
definition reach_infer0_cont where
"reach_infer0_cont \<Delta> =
map r_rhs (filter (\<lambda> r. case r of TA_rule f ps p \<Rightarrow> ps = []) (sorted_list_of_fset \<Delta>))"
definition reach_infer1_cont :: "('q :: linorder, 'f :: linorder) ta_rule fset \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> 'q \<Rightarrow> 'q fset \<Rightarrow> 'q list" where
"reach_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon> =
(let rules = sorted_list_of_fset \<Delta> in
let eps = sorted_list_of_fset \<Delta>\<^sub>\<epsilon> in
let mapp_r = map_val_of_list fst snd (concat (map (\<lambda> r. map (\<lambda> q. (q, [r])) (r_lhs_states r)) rules)) in
let mapp_e = map_of_list fst snd eps in
(\<lambda> p bs.
(map r_rhs (filter (\<lambda> r. case r of TA_rule f qs q \<Rightarrow>
fset_of_list qs |\<subseteq>| finsert p bs) (case_option Nil id (Mapping.lookup mapp_r p)))) @
case_option Nil id (Mapping.lookup mapp_e p)))"
locale reach_rules_fset =
fixes \<Delta> :: "('q :: linorder, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
begin
sublocale reach_horn "TA \<Delta> \<Delta>\<^sub>\<epsilon>" .
lemma infer1:
"infer1 p (fset bs) = set (reach_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon> p bs)"
unfolding reach_infer1 reach_infer1_cont_def set_append Un_assoc[symmetric] Let_def
unfolding sorted_list_of_fset_simps union_fset
apply (intro arg_cong2[of _ _ _ _ "(\<union>)"])
subgoal
apply (auto simp: fset_of_list_elem less_eq_fset.rep_eq fset_of_list.rep_eq image_iff
- map_val_of_list_simp fmember_iff_member_fset split!: ta_rule.splits)
+ map_val_of_list_simp split!: ta_rule.splits)
apply (metis list.set_intros(1) ta_rule.sel(2, 3))
apply (metis in_set_simps(2) ta_rule.exhaust_sel)
done
subgoal
- apply (simp add: image_def Bex_def fmember_iff_member_fset map_val_of_list_simp)
+ apply (simp add: image_def Bex_def map_val_of_list_simp)
done
done
sublocale l: horn_fset "reach_rules (TA \<Delta> \<Delta>\<^sub>\<epsilon>)" "reach_infer0_cont \<Delta>" "reach_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon>"
apply (unfold_locales)
unfolding reach_infer0 reach_infer0_cont_def
subgoal
- apply (auto simp: image_iff ta_rule.case_eq_if Bex_def fset_of_list_elem fmember_iff_member_fset)
+ apply (auto simp: image_iff ta_rule.case_eq_if Bex_def fset_of_list_elem)
apply force
apply (metis ta_rule.collapse)+
done
subgoal using infer1
apply blast
done
done
lemmas infer = l.infer0 l.infer1
lemmas saturate_impl_sound = l.saturate_impl_sound
lemmas saturate_impl_complete = l.saturate_impl_complete
end
definition "reach_cont_impl \<Delta> \<Delta>\<^sub>\<epsilon> =
horn_fset_impl.saturate_impl (reach_infer0_cont \<Delta>) (reach_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon>)"
lemma reach_fset_impl_sound:
"reach_cont_impl \<Delta> \<Delta>\<^sub>\<epsilon> = Some xs \<Longrightarrow> fset xs = ta_reach (TA \<Delta> \<Delta>\<^sub>\<epsilon>)"
using reach_rules_fset.saturate_impl_sound unfolding reach_cont_impl_def
unfolding reach_horn.reach_sound .
lemma reach_fset_impl_complete:
"reach_cont_impl \<Delta> \<Delta>\<^sub>\<epsilon> \<noteq> None"
proof -
have "finite (ta_reach (TA \<Delta> \<Delta>\<^sub>\<epsilon>))"
unfolding ta_reach_reachable by simp
then show ?thesis unfolding reach_cont_impl_def
by (intro reach_rules_fset.saturate_impl_complete)
(auto simp: reach_horn.reach_sound)
qed
lemma reach_impl [code]:
"ta_reachable (TA \<Delta> \<Delta>\<^sub>\<epsilon>) = the (reach_cont_impl \<Delta> \<Delta>\<^sub>\<epsilon>)"
using reach_fset_impl_sound[of \<Delta> \<Delta>\<^sub>\<epsilon>]
- apply (auto simp add: ta_reach_reachable reach_fset_impl_complete fset_of_list_elem)
- apply (metis fset_inject option.exhaust_sel reach_fset_impl_complete)+
- done
+ by (auto simp add: ta_reach_reachable reach_fset_impl_complete fset_of_list_elem)
subsection \<open>Setup for list implementation of productive states\<close>
definition productive_infer1_cont :: "('q :: linorder, 'f :: linorder) ta_rule fset \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow> 'q \<Rightarrow> 'q fset \<Rightarrow> 'q list" where
"productive_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon> =
(let rules = sorted_list_of_fset \<Delta> in
let eps = sorted_list_of_fset \<Delta>\<^sub>\<epsilon> in
let mapp_r = map_of_list (\<lambda> r. r_rhs r) r_lhs_states rules in
let mapp_e = map_of_list snd fst eps in
(\<lambda> p bs.
(case_option Nil id (Mapping.lookup mapp_e p)) @
concat (case_option Nil id (Mapping.lookup mapp_r p))))"
locale productive_rules_fset =
fixes \<Delta> :: "('q :: linorder, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>\<epsilon> :: "('q \<times> 'q) fset" and P :: "'q fset"
begin
sublocale productive_horn "TA \<Delta> \<Delta>\<^sub>\<epsilon>" P .
lemma infer1:
"infer1 p (fset bs) = set (productive_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon> p bs)"
unfolding productive_infer1 productive_infer1_cont_def set_append Un_assoc[symmetric]
unfolding union_fset sorted_list_of_fset_simps Let_def set_append
apply (intro arg_cong2[of _ _ _ _ "(\<union>)"])
subgoal
- apply (simp add: image_def Bex_def fmember_iff_member_fset map_val_of_list_simp)
- done
+ by (simp add: image_def Bex_def map_val_of_list_simp)
subgoal
- apply (auto simp flip: fmember_iff_member_fset simp: map_val_of_list_simp image_iff)
+ apply (auto simp: map_val_of_list_simp image_iff)
apply (metis ta_rule.sel(2, 3))
apply (metis ta_rule.collapse)
- apply (metis notin_fset ta_rule.sel(3))
done
done
sublocale l: horn_fset "productive_rules P (TA \<Delta> \<Delta>\<^sub>\<epsilon>)" "sorted_list_of_fset P" "productive_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon>"
apply (unfold_locales)
using infer1 productive_infer0 fset_of_list.rep_eq
by fastforce+
lemmas infer = l.infer0 l.infer1
lemmas saturate_impl_sound = l.saturate_impl_sound
lemmas saturate_impl_complete = l.saturate_impl_complete
end
definition "productive_cont_impl P \<Delta> \<Delta>\<^sub>\<epsilon> =
horn_fset_impl.saturate_impl (sorted_list_of_fset P) (productive_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon>)"
lemma productive_cont_impl_sound:
"productive_cont_impl P \<Delta> \<Delta>\<^sub>\<epsilon> = Some xs \<Longrightarrow> fset xs = ta_productive_ind P (TA \<Delta> \<Delta>\<^sub>\<epsilon>)"
using productive_rules_fset.saturate_impl_sound unfolding productive_cont_impl_def
unfolding productive_horn.productive_sound .
lemma productive_cont_impl_complete:
"productive_cont_impl P \<Delta> \<Delta>\<^sub>\<epsilon> \<noteq> None"
proof -
have "finite (ta_productive_ind P (TA \<Delta> \<Delta>\<^sub>\<epsilon>))"
unfolding ta_productive_ind by simp
then show ?thesis unfolding productive_cont_impl_def
by (intro productive_rules_fset.saturate_impl_complete)
(auto simp: productive_horn.productive_sound)
qed
lemma productive_impl [code]:
"ta_productive P (TA \<Delta> \<Delta>\<^sub>\<epsilon>) = the (productive_cont_impl P \<Delta> \<Delta>\<^sub>\<epsilon>)"
using productive_cont_impl_complete[of P \<Delta>] productive_cont_impl_sound[of P \<Delta>]
- by (auto simp add: ta_productive_ind fset_of_list_elem fmember_iff_member_fset)
+ by (auto simp add: ta_productive_ind fset_of_list_elem)
subsection \<open>Setup for the implementation of power set construction states\<close>
abbreviation "r_statesl r \<equiv> length (r_lhs_states r)"
definition ps_reachable_states_list where
"ps_reachable_states_list mapp_r mapp_e f ps =
(let R = filter (\<lambda> r. list_all2 (|\<in>|) (r_lhs_states r) ps)
(case_option Nil id (Mapping.lookup mapp_r (f, length ps))) in
let S = map r_rhs R in
S @ concat (map (case_option Nil id \<circ> Mapping.lookup mapp_e) S))"
lemma ps_reachable_states_list_sound:
assumes "length ps = n"
and mapp_r: "case_option Nil id (Mapping.lookup mapp_r (f, n)) =
filter (\<lambda>r. r_root r = f \<and> r_statesl r = n) (sorted_list_of_fset \<Delta>)"
and mapp_e: "\<And>p. case_option Nil id (Mapping.lookup mapp_e p) =
map snd (filter (\<lambda> q. fst q = p) (sorted_list_of_fset (\<Delta>\<^sub>\<epsilon>|\<^sup>+|)))"
shows "fset_of_list (ps_reachable_states_list mapp_r mapp_e f (map ex ps)) =
ps_reachable_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>) f (map ex ps)" (is "?Ls = ?Rs")
proof -
have *: "length ps = n" "length (map ex ps) = n" using assms by auto
{fix q assume "q |\<in>| ?Ls"
then obtain qs p where "TA_rule f qs p |\<in>| \<Delta>" "length ps = length qs"
"list_all2 (|\<in>|) qs (map ex ps)" "p = q \<or> (p, q) |\<in>| \<Delta>\<^sub>\<epsilon>|\<^sup>+|"
unfolding ps_reachable_states_list_def Let_def comp_def assms(1, 2, 3) *
- by (force simp add: fset_of_list_elem image_iff fBex_def simp flip: fmember_iff_member_fset)
+ by (force simp add: fset_of_list_elem image_iff fBex_def)
then have "q |\<in>| ?Rs"
by (force simp add: ps_reachable_states_fmember image_iff)}
moreover
{fix q assume "q |\<in>| ?Rs"
then obtain qs p where "TA_rule f qs p |\<in>| \<Delta>" "length ps = length qs"
"list_all2 (|\<in>|) qs (map ex ps)" "p = q \<or> (p, q) |\<in>| \<Delta>\<^sub>\<epsilon>|\<^sup>+|"
by (auto simp add: ps_reachable_states_fmember list_all2_iff)
then have "q |\<in>| ?Ls"
unfolding ps_reachable_states_list_def Let_def * comp_def assms(2, 3)
- by (force simp add: fset_of_list_elem image_iff simp flip: fmember_iff_member_fset)}
+ by (force simp add: fset_of_list_elem image_iff)}
ultimately show ?thesis by blast
qed
lemma rule_target_statesI:
"\<exists> r |\<in>| \<Delta>. r_rhs r = q \<Longrightarrow> q |\<in>| rule_target_states \<Delta>"
by auto
definition ps_states_infer0_cont :: "('q :: linorder, 'f :: linorder) ta_rule fset \<Rightarrow>
('q \<times> 'q) fset \<Rightarrow> 'q FSet_Lex_Wrapper list" where
"ps_states_infer0_cont \<Delta> \<Delta>\<^sub>\<epsilon> =
(let sig = filter (\<lambda> r. r_lhs_states r = []) (sorted_list_of_fset \<Delta>) in
filter (\<lambda> p. ex p \<noteq> {||}) (map (\<lambda> r. Wrapp (ps_reachable_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>) (r_root r) [])) sig))"
definition ps_states_infer1_cont :: "('q :: linorder , 'f :: linorder) ta_rule fset \<Rightarrow> ('q \<times> 'q) fset \<Rightarrow>
'q FSet_Lex_Wrapper \<Rightarrow> 'q FSet_Lex_Wrapper fset \<Rightarrow> 'q FSet_Lex_Wrapper list" where
"ps_states_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon> =
(let sig = remdups (map (\<lambda> r. (r_root r, r_statesl r)) (filter (\<lambda> r. r_lhs_states r \<noteq> []) (sorted_list_of_fset \<Delta>))) in
let arities = remdups (map snd sig) in
let etr = sorted_list_of_fset (\<Delta>\<^sub>\<epsilon>|\<^sup>+|) in
let mapp_r = map_of_list (\<lambda> r. (r_root r, r_statesl r)) id (sorted_list_of_fset \<Delta>) in
let mapp_e = map_of_list fst snd etr in
(\<lambda> p bs.
(let states = sorted_list_of_fset (finsert p bs) in
let arity_to_states_map = Mapping.tabulate arities (\<lambda> n. list_of_permutation_element_n p n states) in
let res = map (\<lambda> (f, n).
map (\<lambda> s. let rules = the (Mapping.lookup mapp_r (f, n)) in
Wrapp (fset_of_list (ps_reachable_states_list mapp_r mapp_e f (map ex s))))
(the (Mapping.lookup arity_to_states_map n)))
sig in
filter (\<lambda> p. ex p \<noteq> {||}) (concat res))))"
locale ps_states_fset =
fixes \<Delta> :: "('q :: linorder, 'f :: linorder) ta_rule fset" and \<Delta>\<^sub>\<epsilon> :: "('q \<times> 'q) fset"
begin
sublocale ps_states_horn "TA \<Delta> \<Delta>\<^sub>\<epsilon>" .
lemma infer0: "infer0 = set (ps_states_infer0_cont \<Delta> \<Delta>\<^sub>\<epsilon>)"
unfolding ps_states_horn.ps_construction_infer0
unfolding ps_states_infer0_cont_def Let_def
using ps_reachable_states_fmember
by (auto simp add: image_def Ball_def Bex_def)
- (metis fmember_iff_member_fset list_all2_Nil2 ps_reachable_states_fmember ta.sel(1) ta_rule.sel(1, 2))
+ (metis list_all2_Nil2 ps_reachable_states_fmember ta.sel(1) ta_rule.sel(1, 2))
lemma r_lhs_states_nConst:
"r_lhs_states r \<noteq> [] \<Longrightarrow> r_statesl r \<noteq> 0" for r by auto
lemma filter_empty_conv':
"[] = filter P xs \<longleftrightarrow> (\<forall>x\<in>set xs. \<not> P x)"
by (metis filter_empty_conv)
lemma infer1:
"infer1 p (fset bs) = set (ps_states_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon> p bs)" (is "?Ls = ?Rs")
proof -
let ?mapp_r = "map_of_list (\<lambda>r. (r_root r, r_statesl r)) (\<lambda>x. x) (sorted_list_of_fset \<Delta>)"
let ?mapp_e = "map_of_list fst snd (sorted_list_of_fset (\<Delta>\<^sub>\<epsilon>|\<^sup>+|))"
have mapr: "case_option Nil id (Mapping.lookup ?mapp_r (f, n)) =
filter (\<lambda>r. r_root r = f \<and> r_statesl r = n) (sorted_list_of_fset \<Delta>)" for f n
by (auto simp: map_val_of_list_simp image_iff filter_empty_conv' intro: filter_cong)
have epsr: "\<And>p. case_option Nil id (Mapping.lookup ?mapp_e p) =
map snd (filter (\<lambda> q. fst q = p) (sorted_list_of_fset (\<Delta>\<^sub>\<epsilon>|\<^sup>+|)))"
by (auto simp: map_val_of_list_simp image_iff filter_empty_conv) metis
have *: "p \<in> set qs \<Longrightarrow> x |\<in>| ps_reachable_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>) f (map ex qs) \<Longrightarrow>
(\<exists> ps q. TA_rule f ps q |\<in>| \<Delta> \<and> length ps = length qs)" for x f qs
by (auto simp: ps_reachable_states_fmember list_all2_conv_all_nth)
{fix q assume "q \<in> ?Ls"
then obtain f qss where sp: "q = Wrapp (ps_reachable_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>) f (map ex qss))"
"ps_reachable_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>) f (map ex qss) \<noteq> {||}" "p \<in> set qss" "set qss \<subseteq> insert p (fset bs)"
by (auto simp add: ps_construction_infer1 ps_reachable_states_fmember)
from sp(2, 3) obtain ps p' where r: "TA_rule f ps p' |\<in>| \<Delta>" "length ps = length qss" using *
by blast
then have mem: "qss \<in> set (list_of_permutation_element_n p (length ps) (sorted_list_of_fset (finsert p bs)))" using sp(2-)
by (auto simp: list_of_permutation_element_n_iff)
(meson in_set_idx insertE set_list_subset_eq_nth_conv)
then have "q \<in> ?Rs" using sp r
unfolding ps_construction_infer1 ps_states_infer1_cont_def Let_def
- apply (simp add: lookup_tabulate ps_reachable_states_fmember image_iff flip: fmember_iff_member_fset)
+ apply (simp add: lookup_tabulate ps_reachable_states_fmember image_iff)
apply (rule_tac x = "f ps \<rightarrow> p'" in exI)
apply (auto simp: Bex_def ps_reachable_states_list_sound[OF _ mapr epsr] intro: exI[of _ qss])
done}
moreover
{fix q assume ass: "q \<in> ?Rs"
then obtain r qss where "r |\<in>| \<Delta>" "r_lhs_states r \<noteq> []" "qss \<in> set (list_of_permutation_element_n p (r_statesl r) (sorted_list_of_fset (finsert p bs)))"
"q = Wrapp (ps_reachable_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>) (r_root r) (map ex qss))"
unfolding ps_states_infer1_cont_def Let_def
by (auto simp add: lookup_tabulate ps_reachable_states_fmember image_iff
- ps_reachable_states_list_sound[OF _ mapr epsr] split: if_splits simp flip: fmember_iff_member_fset)
+ ps_reachable_states_list_sound[OF _ mapr epsr] split: if_splits)
moreover have "q \<noteq> Wrapp {||}" using ass
by (auto simp: ps_states_infer1_cont_def Let_def)
ultimately have "q \<in> ?Ls" unfolding ps_construction_infer1
apply (auto simp: list_of_permutation_element_n_iff intro!: exI[of _ "r_root r"] exI[of _ qss])
apply (metis in_set_idx)
done}
ultimately show ?thesis by blast
qed
sublocale l: horn_fset "ps_states_rules (TA \<Delta> \<Delta>\<^sub>\<epsilon>)" "ps_states_infer0_cont \<Delta> \<Delta>\<^sub>\<epsilon>" "ps_states_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon>"
apply (unfold_locales)
using infer0 infer1
by fastforce+
lemmas infer = l.infer0 l.infer1
lemmas saturate_impl_sound = l.saturate_impl_sound
lemmas saturate_impl_complete = l.saturate_impl_complete
end
definition "ps_states_fset_impl \<Delta> \<Delta>\<^sub>\<epsilon> =
horn_fset_impl.saturate_impl (ps_states_infer0_cont \<Delta> \<Delta>\<^sub>\<epsilon>) (ps_states_infer1_cont \<Delta> \<Delta>\<^sub>\<epsilon>)"
lemma ps_states_fset_impl_sound:
assumes "ps_states_fset_impl \<Delta> \<Delta>\<^sub>\<epsilon> = Some xs"
shows "xs = ps_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>)"
using ps_states_fset.saturate_impl_sound[OF assms[unfolded ps_states_fset_impl_def]]
using ps_states_horn.ps_states_sound[of "TA \<Delta> \<Delta>\<^sub>\<epsilon>"]
- by (auto simp: fset_of_list_elem fmember_iff_member_fset ps_states.rep_eq fset_of_list.rep_eq)
+ by (auto simp: fset_of_list_elem ps_states.rep_eq fset_of_list.rep_eq)
lemma ps_states_fset_impl_complete:
"ps_states_fset_impl \<Delta> \<Delta>\<^sub>\<epsilon> \<noteq> None"
proof -
let ?R = "ps_states (TA \<Delta> \<Delta>\<^sub>\<epsilon>)"
let ?S = "horn.saturate (ps_states_rules (TA \<Delta> \<Delta>\<^sub>\<epsilon>))"
have "?S \<subseteq> fset ?R"
using ps_states_horn.ps_states_sound
by (simp add: ps_states_horn.ps_states_sound ps_states.rep_eq)
from finite_subset[OF this] show ?thesis
unfolding ps_states_fset_impl_def
by (intro ps_states_fset.saturate_impl_complete) simp
qed
lemma ps_ta_impl [code]:
"ps_ta (TA \<Delta> \<Delta>\<^sub>\<epsilon>) =
(let xs = the (ps_states_fset_impl \<Delta> \<Delta>\<^sub>\<epsilon>) in
TA (ps_rules (TA \<Delta> \<Delta>\<^sub>\<epsilon>) xs) {||})"
using ps_states_fset_impl_complete
using ps_states_fset_impl_sound
unfolding ps_ta_def Let_def
by (metis option.exhaust_sel)
lemma ps_reg_impl [code]:
"ps_reg (Reg Q (TA \<Delta> \<Delta>\<^sub>\<epsilon>)) =
(let xs = the (ps_states_fset_impl \<Delta> \<Delta>\<^sub>\<epsilon>) in
Reg (ffilter (\<lambda> S. Q |\<inter>| ex S \<noteq> {||}) xs)
(TA (ps_rules (TA \<Delta> \<Delta>\<^sub>\<epsilon>) xs) {||}))"
using ps_states_fset_impl_complete[of \<Delta> \<Delta>\<^sub>\<epsilon>]
using ps_states_fset_impl_sound[of \<Delta> \<Delta>\<^sub>\<epsilon>]
unfolding ps_reg_def ps_ta_Q\<^sub>f_def Let_def
unfolding ps_ta_def Let_def
using eq_ffilter by auto
lemma prod_ta_zip [code]:
"prod_ta_rules (\<A> :: ('q1 :: linorder, 'f :: linorder) ta) (\<B> :: ('q2 :: linorder, 'f :: linorder) ta) =
(let sig = sorted_list_of_fset (ta_sig \<A> |\<inter>| ta_sig \<B>) in
let mapA = map_of_list (\<lambda>r. (r_root r, r_statesl r)) id (sorted_list_of_fset (rules \<A>)) in
let mapB = map_of_list (\<lambda>r. (r_root r, r_statesl r)) id (sorted_list_of_fset (rules \<B>)) in
let merge = (\<lambda> (ra, rb). TA_rule (r_root ra) (zip (r_lhs_states ra) (r_lhs_states rb)) (r_rhs ra, r_rhs rb)) in
fset_of_list (
concat (map (\<lambda> (f, n). map merge
(List.product (the (Mapping.lookup mapA (f, n))) (the (Mapping.lookup mapB (f, n))))) sig)))"
(is "?Ls = ?Rs")
proof -
have [simp]: "distinct (sorted_list_of_fset (ta_sig \<A>))" "distinct (sorted_list_of_fset (ta_sig \<B>))"
by (simp_all add: distinct_sorted_list_of_fset)
have *: "sort (remdups (map (\<lambda>r. (r_root r, r_statesl r)) (sorted_list_of_fset (rules \<A>)))) = sorted_list_of_fset (ta_sig \<A>)"
"sort (remdups (map (\<lambda>r. (r_root r, r_statesl r)) (sorted_list_of_fset (rules \<B>)))) = sorted_list_of_fset (ta_sig \<B>)"
by (auto simp: ta_sig_def sorted_list_of_fset_fimage_dist)
{fix r assume ass: "r |\<in>| ?Ls"
then obtain f qs q where [simp]: "r = f qs \<rightarrow> q" by auto
then have "(f, length qs) |\<in>| ta_sig \<A> |\<inter>| ta_sig \<B>" using ass by auto
then have "r |\<in>| ?Rs" using ass unfolding map_val_of_list_tabulate_conv *
- by (auto simp: Let_def fset_of_list_elem image_iff case_prod_beta lookup_tabulate simp flip: fmember_iff_member_fset intro!: bexI[of _ "(f, length qs)"])
+ by (auto simp: Let_def fset_of_list_elem image_iff case_prod_beta lookup_tabulate intro!: bexI[of _ "(f, length qs)"])
(metis (no_types, lifting) length_map ta_rule.sel(1 - 3) zip_map_fst_snd)}
moreover
{fix r assume ass: "r |\<in>| ?Rs" then have "r |\<in>| ?Ls" unfolding map_val_of_list_tabulate_conv *
- by (auto simp: fset_of_list_elem finite_Collect_prod_ta_rules lookup_tabulate simp flip: fmember_iff_member_fset)
+ by (auto simp: fset_of_list_elem finite_Collect_prod_ta_rules lookup_tabulate)
(metis ta_rule.collapse)}
ultimately show ?thesis by blast
qed
(*
export_code ta_der in Haskell
export_code ta_reachable in Haskell
export_code ta_productive in Haskell
export_code trim_ta in Haskell
export_code ta_restrict in Haskell
export_code ps_reachable_states in Haskell
export_code prod_ta_rules in Haskell
export_code ps_ta in Haskell
export_code ps_reg in Haskell
export_code reg_intersect in Haskell
*)
end
\ No newline at end of file
diff --git a/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy b/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy
--- a/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy
+++ b/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy
@@ -1,1127 +1,1125 @@
section \<open>Soundness for graphical diagrams\<close>
theory Ribbons_Graphical_Soundness imports
Ribbons_Graphical
More_Finite_Map
begin
text \<open>We prove that the proof rules for graphical ribbon proofs are sound
with respect to the rules of separation logic.
We impose an additional assumption to achieve soundness: that the
Frame rule has no side-condition. This assumption is reasonable because there
are several separation logics that lack such a side-condition, such as
``variables-as-resource''.
We first describe how to extract proofchains from a diagram. This process is
similar to the process of extracting commands from a diagram, which was
described in @{theory Ribbon_Proofs.Ribbons_Graphical}. When we extract a proofchain, we
don't just include the commands, but the assertions in between them. Our
main lemma for proving soundness says that each of these proofchains
corresponds to a valid separation logic proof.
\<close>
subsection \<open>Proofstate chains\<close>
text \<open>When extracting a proofchain from a diagram, we need to keep track
of which nodes we have processed and which ones we haven't. A
proofstate, defined below, maps a node to ``Top'' if it hasn't been
processed and ``Bot'' if it has.\<close>
datatype topbot = Top | Bot
type_synonym proofstate = "node \<rightharpoonup>\<^sub>f topbot"
text \<open>A proofstate chain contains all the nodes and edges of a graphical
diagram, interspersed with proofstates that track which nodes have been
processed at each point.\<close>
type_synonym ps_chain = "(proofstate, node + edge) chain"
text \<open>The @{term "next_ps \<sigma>"} function processes one node or one edge in a
diagram, given the current proofstate @{term \<sigma>}. It processes a node
@{term v} by replacing the mapping from @{term v} to @{term Top} with a
mapping from @{term v} to @{term Bot}. It processes an edge @{term e}
(whose source and target nodes are @{term vs} and @{term ws} respectively)
by removing all the mappings from @{term vs} to @{term Bot}, and adding
mappings from @{term ws} to @{term Top}.\<close>
fun next_ps :: "proofstate \<Rightarrow> node + edge \<Rightarrow> proofstate"
where
"next_ps \<sigma> (Inl v) = \<sigma> \<ominus> {|v|} ++\<^sub>f [{|v|} |=> Bot]"
| "next_ps \<sigma> (Inr e) = \<sigma> \<ominus> fst3 e ++\<^sub>f [thd3 e |=> Top]"
text \<open>The function @{term "mk_ps_chain \<Pi> \<pi>"} generates from @{term \<pi>}, which
is a list of nodes and edges, a proofstate chain, by interspersing the
elements of @{term \<pi>} with the appropriate proofstates. The first argument
@{term \<Pi>} is the part of the chain that has already been converted.\<close>
definition
mk_ps_chain :: "[ps_chain, (node + edge) list] \<Rightarrow> ps_chain"
where
"mk_ps_chain \<equiv> foldl (\<lambda>\<Pi> x. cSnoc \<Pi> x (next_ps (post \<Pi>) x))"
lemma mk_ps_chain_preserves_length:
fixes \<pi> \<Pi>
shows "chainlen (mk_ps_chain \<Pi> \<pi>) = chainlen \<Pi> + length \<pi>"
proof (induct \<pi> arbitrary: \<Pi>)
case Nil
show ?case by (unfold mk_ps_chain_def, auto)
next
case (Cons x \<pi>)
show ?case
apply (unfold mk_ps_chain_def list.size foldl.simps)
apply (fold mk_ps_chain_def)
apply (auto simp add: Cons len_snoc)
done
qed
text \<open>Distributing @{term mk_ps_chain} over @{term Cons}.\<close>
lemma mk_ps_chain_cons:
"mk_ps_chain \<Pi> (x # \<pi>) = mk_ps_chain (cSnoc \<Pi> x (next_ps (post \<Pi>) x)) \<pi>"
by (auto simp add: mk_ps_chain_def)
text \<open>Distributing @{term mk_ps_chain} over @{term snoc}.\<close>
lemma mk_ps_chain_snoc:
"mk_ps_chain \<Pi> (\<pi> @ [x])
= cSnoc (mk_ps_chain \<Pi> \<pi>) x (next_ps (post (mk_ps_chain \<Pi> \<pi>)) x)"
by (unfold mk_ps_chain_def, auto)
text \<open>Distributing @{term mk_ps_chain} over @{term cCons}.\<close>
lemma mk_ps_chain_ccons:
fixes \<pi> \<Pi>
shows "mk_ps_chain (\<lbrace> \<sigma> \<rbrace> \<cdot> x \<cdot> \<Pi>) \<pi> = \<lbrace> \<sigma> \<rbrace> \<cdot> x \<cdot> mk_ps_chain \<Pi> \<pi> "
by (induct \<pi> arbitrary: \<Pi>, auto simp add: mk_ps_chain_cons mk_ps_chain_def)
lemma pre_mk_ps_chain:
fixes \<Pi> \<pi>
shows "pre (mk_ps_chain \<Pi> \<pi>) = pre \<Pi>"
apply (induct \<pi> arbitrary: \<Pi>)
apply (auto simp add: mk_ps_chain_def mk_ps_chain_cons pre_snoc)
done
text \<open>A chain which is obtained from the list @{term \<pi>}, has @{term \<pi>}
as its list of commands. The following lemma states this in a slightly
more general form, that allows for part of the chain to have already
been processed.\<close>
lemma comlist_mk_ps_chain:
"comlist (mk_ps_chain \<Pi> \<pi>) = comlist \<Pi> @ \<pi>"
proof (induct \<pi> arbitrary: \<Pi>)
case Nil
thus ?case by (auto simp add: mk_ps_chain_def)
next
case (Cons x \<pi>')
show ?case
apply (unfold mk_ps_chain_def foldl.simps, fold mk_ps_chain_def)
apply (auto simp add: Cons comlist_snoc)
done
qed
text \<open>In order to perform induction over our diagrams, we shall wish
to obtain ``smaller'' diagrams, by removing nodes or edges. However, the
syntax and well-formedness constraints for diagrams are such that although
we can always remove an edge from a diagram, we cannot (in general) remove
a node -- the resultant diagram would not be a well-formed if an edge
connected to that node.
Hence, we consider ``partially-processed diagrams'' @{term "(G,S)"}, which
comprise a diagram @{term G} and a set @{term S} of nodes. @{term S} denotes
the subset of @{term G}'s initial nodes that have already been processed,
and can be thought of as having been removed from @{term G}.
We now give an updated version of the @{term "lins G"} function. This was
originally defined in @{theory Ribbon_Proofs.Ribbons_Graphical}. We provide an extra
parameter, @{term S}, which denotes the subset of @{term G}'s initial nodes
that shouldn't be included in the linear extensions.\<close>
definition lins2 :: "[node fset, diagram] \<Rightarrow> lin set"
where
"lins2 S G \<equiv> {\<pi> :: lin .
(distinct \<pi>)
\<and> (set \<pi> = (fset G^V - fset S) <+> set G^E)
\<and> (\<forall>i j v e. i < length \<pi> \<and> j < length \<pi>
\<and> \<pi>!i = Inl v \<and> \<pi>!j = Inr e \<and> v |\<in>| fst3 e \<longrightarrow> i<j)
\<and> (\<forall>j k w e. j < length \<pi> \<and> k < length \<pi>
\<and> \<pi>!j = Inr e \<and> \<pi>!k = Inl w \<and> w |\<in>| thd3 e \<longrightarrow> j<k) }"
lemma lins2D:
assumes "\<pi> \<in> lins2 S G"
shows "distinct \<pi>"
and "set \<pi> = (fset G^V - fset S) <+> set G^E"
and "\<And>i j v e. \<lbrakk> i < length \<pi> ; j < length \<pi> ;
\<pi>!i = Inl v ; \<pi>!j = Inr e ; v |\<in>| fst3 e \<rbrakk> \<Longrightarrow> i<j"
and "\<And>i k w e. \<lbrakk> j < length \<pi> ; k < length \<pi> ;
\<pi>!j = Inr e ; \<pi>!k = Inl w ; w |\<in>| thd3 e \<rbrakk> \<Longrightarrow> j<k"
using assms
apply (unfold lins2_def Collect_iff)
apply (elim conjE, assumption)+
apply blast+
done
lemma lins2I:
assumes "distinct \<pi>"
and "set \<pi> = (fset G^V - fset S) <+> set G^E"
and "\<And>i j v e. \<lbrakk> i < length \<pi> ; j < length \<pi> ;
\<pi>!i = Inl v ; \<pi>!j = Inr e ; v |\<in>| fst3 e \<rbrakk> \<Longrightarrow> i<j"
and "\<And>j k w e. \<lbrakk> j < length \<pi> ; k < length \<pi> ;
\<pi>!j = Inr e ; \<pi>!k = Inl w ; w |\<in>| thd3 e \<rbrakk> \<Longrightarrow> j<k"
shows "\<pi> \<in> lins2 S G"
using assms
apply (unfold lins2_def Collect_iff, intro conjI)
apply assumption+
apply blast+
done
text \<open>When @{term S} is empty, the two definitions coincide.\<close>
lemma lins_is_lins2_with_empty_S:
"lins G = lins2 {||} G"
by (unfold lins_def lins2_def, auto)
text \<open>The first proofstate for a diagram @{term G} is obtained by
mapping each of its initial nodes to @{term Top}.\<close>
definition
initial_ps :: "diagram \<Rightarrow> proofstate"
where
"initial_ps G \<equiv> [ initials G |=> Top ]"
text \<open>The first proofstate for the partially-processed diagram @{term G} is
obtained by mapping each of its initial nodes to @{term Top}, except those
in @{term S}, which are mapped to @{term Bot}.\<close>
definition
initial_ps2 :: "[node fset, diagram] \<Rightarrow> proofstate"
where
"initial_ps2 S G \<equiv> [ initials G - S |=> Top ] ++\<^sub>f [ S |=> Bot ]"
text \<open>When @{term S} is empty, the above two definitions coincide.\<close>
lemma initial_ps_is_initial_ps2_with_empty_S:
"initial_ps = initial_ps2 {||}"
apply (unfold fun_eq_iff, intro allI)
apply (unfold initial_ps_def initial_ps2_def)
apply simp
done
text \<open>The following function extracts the set of proofstate chains from
a diagram.\<close>
definition
ps_chains :: "diagram \<Rightarrow> ps_chain set"
where
"ps_chains G \<equiv> mk_ps_chain (cNil (initial_ps G)) ` lins G"
text \<open>The following function extracts the set of proofstate chains from
a partially-processed diagram. Nodes in @{term S} are excluded from
the resulting chains.\<close>
definition
ps_chains2 :: "[node fset, diagram] \<Rightarrow> ps_chain set"
where
"ps_chains2 S G \<equiv> mk_ps_chain (cNil (initial_ps2 S G)) ` lins2 S G"
text \<open>When @{term S} is empty, the above two definitions coincide.\<close>
lemma ps_chains_is_ps_chains2_with_empty_S:
"ps_chains = ps_chains2 {||}"
apply (unfold fun_eq_iff, intro allI)
apply (unfold ps_chains_def ps_chains2_def)
apply (fold initial_ps_is_initial_ps2_with_empty_S)
apply (fold lins_is_lins2_with_empty_S)
apply auto
done
text \<open>We now wish to describe proofstates chain that are well-formed. First,
let us say that @{term "f ++\<^sub>fdisjoint g"} is defined, when @{term f} and
@{term g} have disjoint domains, as @{term "f ++\<^sub>f g"}. Then, a well-formed
proofstate chain consists of triples of the form @{term "(\<sigma> ++\<^sub>fdisjoint
[{| v |} |=> Top], Inl v, \<sigma> ++\<^sub>fdisjoint [{| v |} |=> Bot])"}, where @{term v}
is a node, or of the form @{term "(\<sigma> ++\<^sub>fdisjoint [{| vs |} |=> Bot], Inr e,
\<sigma> ++\<^sub>fdisjoint [{| ws |} |=> Top])"}, where @{term e} is an edge with source
and target nodes @{term vs} and @{term ws} respectively.
The definition below describes a well-formed triple; we then lift this
to complete chains shortly.\<close>
definition
wf_ps_triple :: "proofstate \<times> (node + edge) \<times> proofstate \<Rightarrow> bool"
where
"wf_ps_triple T = (case snd3 T of
Inl v \<Rightarrow> (\<exists>\<sigma>. v |\<notin>| fmdom \<sigma>
\<and> fst3 T = [ {|v|} |=> Top ] ++\<^sub>f \<sigma>
\<and> thd3 T = [ {|v|} |=> Bot ] ++\<^sub>f \<sigma>)
| Inr e \<Rightarrow> (\<exists>\<sigma>. (fst3 e |\<union>| thd3 e) |\<inter>| fmdom \<sigma> = {||}
\<and> fst3 T = [ fst3 e |=> Bot ] ++\<^sub>f \<sigma>
\<and> thd3 T = [ thd3 e |=> Top ] ++\<^sub>f \<sigma>))"
lemma wf_ps_triple_nodeI:
assumes "\<exists>\<sigma>. v |\<notin>| fmdom \<sigma> \<and>
\<sigma>1 = [ {|v|} |=> Top ] ++\<^sub>f \<sigma> \<and>
\<sigma>2 = [ {|v|} |=> Bot ] ++\<^sub>f \<sigma>"
shows "wf_ps_triple (\<sigma>1, Inl v, \<sigma>2)"
using assms unfolding wf_ps_triple_def
by (auto simp add: fst3_simp snd3_simp thd3_simp)
lemma wf_ps_triple_edgeI:
assumes "\<exists>\<sigma>. (fst3 e |\<union>| thd3 e) |\<inter>| fmdom \<sigma> = {||}
\<and> \<sigma>1 = [ fst3 e |=> Bot ] ++\<^sub>f \<sigma>
\<and> \<sigma>2 = [ thd3 e |=> Top ] ++\<^sub>f \<sigma>"
shows "wf_ps_triple (\<sigma>1, Inr e, \<sigma>2)"
using assms unfolding wf_ps_triple_def
by (auto simp add: fst3_simp snd3_simp thd3_simp)
definition
wf_ps_chain :: "ps_chain \<Rightarrow> bool"
where
"wf_ps_chain \<equiv> chain_all wf_ps_triple"
lemma next_initial_ps2_vertex:
"initial_ps2 ({|v|} |\<union>| S) G
= initial_ps2 S G \<ominus> {|v|} ++\<^sub>f [ {|v|} |=> Bot ]"
apply (unfold initial_ps2_def)
apply transfer
apply (auto simp add: make_map_def map_diff_def map_add_def restrict_map_def)
done
lemma next_initial_ps2_edge:
assumes "G = Graph V \<Lambda> E" and "G' = Graph V' \<Lambda> E'" and
"V' = V - fst3 e" and "E' = removeAll e E" and "e \<in> set E" and
"fst3 e |\<subseteq>| S" and "S |\<subseteq>| initials G" and "wf_dia G"
shows "initial_ps2 (S - fst3 e) G' =
initial_ps2 S G \<ominus> fst3 e ++\<^sub>f [ thd3 e |=> Top ]"
proof (insert assms, unfold initial_ps2_def, transfer)
fix G V \<Lambda> E G' V' E' e S
assume G_def: "G = Graph V \<Lambda> E" and G'_def: "G' = Graph V' \<Lambda> E'" and
V'_def: "V' = V - fst3 e" and E'_def: "E' = removeAll e E" and
e_in_E: "e \<in> set E" and fst_e_in_S: "fst3 e |\<subseteq>| S" and
S_initials: "S |\<subseteq>| initials G" and wf_G: "wf_dia G"
have "thd3 e |\<inter>| initials G = {||}"
by (auto simp add: initials_def G_def e_in_E)
show "make_map (initials G' - (S - fst3 e)) Top ++ make_map (S - fst3 e) Bot
= map_diff (make_map (initials G - S) Top ++ make_map S Bot) (fst3 e)
++ make_map (thd3 e) Top"
apply (unfold make_map_def map_diff_def)
apply (unfold map_add_def restrict_map_def)
apply (unfold minus_fset)
apply (unfold fun_eq_iff initials_def)
apply (unfold G_def G'_def V'_def E'_def)
apply (unfold edges.simps vertices.simps)
- apply (simp add: less_eq_fset.rep_eq fmember_iff_member_fset e_in_E)
+ apply (simp add: less_eq_fset.rep_eq e_in_E)
apply safe
apply (insert \<open>thd3 e |\<inter>| initials G = {||}\<close>)[1]
apply (insert S_initials, fold fset_cong)[2]
apply (unfold less_eq_fset.rep_eq initials_def filter_fset)
- apply (auto simp add: fmember_iff_member_fset G_def e_in_E)[1]
- apply (auto simp add: fmember_iff_member_fset G_def e_in_E)[1]
- apply (auto simp add: fmember_iff_member_fset G_def e_in_E)[1]
+ apply (auto simp add: G_def e_in_E)[1]
+ apply (auto simp add: G_def e_in_E)[1]
+ apply (auto simp add: G_def e_in_E)[1]
apply (insert wf_G)[1]
apply (unfold G_def vertices.simps edges.simps)
apply (drule wf_dia_inv(3))
apply (unfold acyclicity_def)
apply (metis fst_e_in_S inter_fset le_iff_inf subsetD)
apply (insert wf_G)[1]
apply (unfold G_def vertices.simps edges.simps)
apply (drule wf_dia_inv(4))
apply (drule linearityD2)
apply (fold fset_cong, unfold inter_fset fset_simps)
apply (insert e_in_E, blast)[1]
apply (insert wf_G)[1]
apply (unfold G_def vertices.simps edges.simps)
apply (drule wf_dia_inv(3))
apply (metis (lifting) e_in_E G_def empty_iff fset_simps(1)
- finter_iff linearityD(2) notin_fset wf_G wf_dia_inv(4))
+ finter_iff linearityD(2) wf_G wf_dia_inv(4))
apply (insert wf_G)[1]
apply (unfold G_def vertices.simps edges.simps)
apply (drule wf_dia_inv(4))
apply (drule linearityD2)
apply (fold fset_cong, unfold inter_fset fset_simps)
apply (insert e_in_E, blast)[1]
apply (insert wf_G)[1]
apply (unfold G_def vertices.simps edges.simps)
apply (drule wf_dia_inv(3))
apply (metis (lifting) e_in_E G_def empty_iff fset_simps(1)
- finter_iff linearityD(2) notin_fset wf_G wf_dia_inv(4))
+ finter_iff linearityD(2) wf_G wf_dia_inv(4))
apply (insert wf_G)
apply (unfold G_def vertices.simps edges.simps)
apply (drule wf_dia_inv(5))
apply (unfold less_eq_fset.rep_eq union_fset)
apply auto[1]
apply (drule wf_dia_inv(5))
apply (unfold less_eq_fset.rep_eq union_fset)
apply auto[1]
apply (drule wf_dia_inv(5))
apply (unfold less_eq_fset.rep_eq union_fset)
apply (auto simp add: e_in_E)[1]
apply (drule wf_dia_inv(5))
apply (unfold less_eq_fset.rep_eq union_fset)
apply (auto simp add: e_in_E)[1]
done
qed
lemma next_lins2_vertex:
assumes "Inl v # \<pi> \<in> lins2 S G"
assumes "v |\<notin>| S"
shows "\<pi> \<in> lins2 ({|v|} |\<union>| S) G"
proof -
note lins2D = lins2D[OF assms(1)]
show ?thesis
proof (intro lins2I)
show "distinct \<pi>" using lins2D(1) by auto
next
have "set \<pi> = set (Inl v # \<pi>) - {Inl v}" using lins2D(1) by auto
also have "... = (fset G^V - fset ({|v|} |\<union>| S)) <+> set G^E"
using lins2D(2) by auto
finally show "set \<pi> = (fset G^V - fset ({|v|} |\<union>| S)) <+> set G^E"
by auto
next
fix i j v e
assume "i < length \<pi>" "j < length \<pi>" "\<pi> ! i = Inl v"
"\<pi> ! j = Inr e" "v |\<in>| fst3 e"
thus "i < j" using lins2D(3)[of "i+1" "j+1"] by auto
next
fix j k w e
assume "j < length \<pi>" "k < length \<pi>" "\<pi> ! j = Inr e"
"\<pi> ! k = Inl w" "w |\<in>| thd3 e"
thus "j < k" using lins2D(4)[of "j+1" "k+1"] by auto
qed
qed
lemma next_lins2_edge:
assumes "Inr e # \<pi> \<in> lins2 S (Graph V \<Lambda> E)"
and "vs |\<subseteq>| S"
and "e = (vs,c,ws)"
shows "\<pi> \<in> lins2 (S - vs) (Graph (V - vs) \<Lambda> (removeAll e E))"
proof -
note lins2D = lins2D[OF assms(1)]
show ?thesis
proof (intro lins2I, unfold vertices.simps edges.simps)
show "distinct \<pi>"
using lins2D(1) by auto
next
show "set \<pi> = (fset (V - vs) - fset (S - vs))
<+> set (removeAll e E)"
apply (insert lins2D(1) lins2D(2) assms(2))
apply (unfold assms(3) vertices.simps edges.simps less_eq_fset.rep_eq, simp)
apply (unfold diff_diff_eq)
proof -
have "\<forall>a aa b.
insert (Inr (vs, c, ws)) (set \<pi>) = (fset V - fset S) <+> set E \<longrightarrow>
fset vs \<subseteq> fset S \<longrightarrow>
Inr (vs, c, ws) \<notin> set \<pi> \<longrightarrow>
distinct \<pi> \<longrightarrow> (a, aa, b) \<in> set E \<longrightarrow> Inr (a, aa, b) \<notin> set \<pi> \<longrightarrow> b = ws"
by (metis (lifting) InrI List.set_simps(2)
prod.inject set_ConsD sum.simps(2))
moreover have "\<forall>a aa b.
insert (Inr (vs, c, ws)) (set \<pi>) = (fset V - fset S) <+> set E \<longrightarrow>
fset vs \<subseteq> fset S \<longrightarrow>
Inr (vs, c, ws) \<notin> set \<pi> \<longrightarrow>
distinct \<pi> \<longrightarrow> (a, aa, b) \<in> set E \<longrightarrow> Inr (a, aa, b) \<notin> set \<pi> \<longrightarrow> aa = c"
by (metis (lifting) InrI List.set_simps(2)
prod.inject set_ConsD sum.simps(2))
moreover have "\<forall>x. insert (Inr (vs, c, ws)) (set \<pi>) = (fset V - fset S) <+> set E \<longrightarrow>
fset vs \<subseteq> fset S \<longrightarrow>
Inr (vs, c, ws) \<notin> set \<pi> \<longrightarrow>
distinct \<pi> \<longrightarrow> x \<in> set \<pi> \<longrightarrow> x \<in> (fset V - fset S) <+> set E - {(vs, c, ws)}"
apply (unfold insert_is_Un[of _ "set \<pi>"])
apply (fold assms(3))
apply clarify
apply (subgoal_tac "set \<pi> = ((fset V - fset S) <+> set E) - {Inr e}")
by auto
ultimately show "Inr (vs, c, ws) \<notin> set \<pi> \<and> distinct \<pi> \<Longrightarrow>
insert (Inr (vs, c, ws)) (set \<pi>) = (fset V - fset S) <+> set E \<Longrightarrow>
fset vs \<subseteq> fset S \<Longrightarrow> set \<pi> = (fset V - fset S) <+> set E - {(vs, c, ws)}"
by blast
qed
next
fix i j v e
assume "i < length \<pi>" "j < length \<pi>" "\<pi> ! i = Inl v"
"\<pi> ! j = Inr e" "v |\<in>| fst3 e"
thus "i < j" using lins2D(3)[of "i+1" "j+1"] by auto
next
fix j k w e
assume "j < length \<pi>" "k < length \<pi>" "\<pi> ! j = Inr e"
"\<pi> ! k = Inl w" "w |\<in>| thd3 e"
thus "j < k" using lins2D(4)[of "j+1" "k+1"] by auto
qed
qed
text \<open>We wish to prove that every proofstate chain that can be obtained from
a linear extension of @{term G} is well-formed and has as its final
proofstate that state in which every terminal node in @{term G} is mapped
to @{term Bot}.
We first prove this for partially-processed diagrams, for
then the result for ordinary diagrams follows as an easy corollary.
We use induction on the size of the partially-processed diagram. The size of
a partially-processed diagram @{term "(G,S)"} is defined as the number of
nodes in @{term G}, plus the number of edges, minus the number of nodes in
@{term S}.\<close>
-lemmas [simp] = fmember_iff_member_fset
-
lemma wf_chains2:
fixes k
assumes "S |\<subseteq>| initials G"
and "wf_dia G"
and "\<Pi> \<in> ps_chains2 S G"
and "fcard G^V + length G^E = k + fcard S"
shows "wf_ps_chain \<Pi> \<and> (post \<Pi> = [ terminals G |=> Bot ])"
using assms
proof (induct k arbitrary: S G \<Pi>)
case 0
obtain V \<Lambda> E where G_def: "G = Graph V \<Lambda> E" by (metis diagram.exhaust)
have "S |\<subseteq>| V"
using "0.prems"(1) initials_in_vertices[of "G"]
by (auto simp add: G_def)
have "fcard V \<le> fcard S"
using "0.prems"(4)
by (unfold G_def, auto)
from fcard_seteq[OF \<open>S |\<subseteq>| V\<close> this] have "S = V" by auto
hence "E = []" using "0.prems"(4) by (unfold G_def, auto)
have "initials G = V"
by (unfold G_def \<open>E=[]\<close>, rule no_edges_imp_all_nodes_initial)
have "terminals G = V"
by (unfold G_def \<open>E=[]\<close>, rule no_edges_imp_all_nodes_terminal)
have "{} <+> {} = {}" by auto
have "lins2 S G = { [] }"
apply (unfold G_def \<open>S=V\<close> \<open>E=[]\<close>)
apply (unfold lins2_def, auto simp add: \<open>{} <+> {} = {}\<close>)
done
hence \<Pi>_def: "\<Pi> = \<lbrace> initial_ps2 S G \<rbrace>"
using "0.prems"(3)
by (auto simp add: ps_chains2_def mk_ps_chain_def)
show ?case
apply (intro conjI)
apply (unfold \<Pi>_def wf_ps_chain_def, auto)
apply (unfold post.simps initial_ps2_def \<open>initials G = V\<close> \<open>terminals G = V\<close>)
apply (unfold \<open>S=V\<close>)
apply (subgoal_tac "V - V = {||}", simp_all)
done
next
case (Suc k)
obtain V \<Lambda> E where G_def: "G = Graph V \<Lambda> E" by (metis diagram.exhaust)
from Suc.prems(3) obtain \<pi> where
\<Pi>_def: "\<Pi> = mk_ps_chain \<lbrace> initial_ps2 S G \<rbrace> \<pi>" and
\<pi>_in: "\<pi> \<in> lins2 S G"
by (auto simp add: ps_chains2_def)
note lins2 = lins2D[OF \<pi>_in]
have "S |\<subseteq>| V"
using Suc.prems(1) initials_in_vertices[of "G"]
by (auto simp add: G_def)
show ?case
proof (cases \<pi>)
case Nil
from \<pi>_in have "V = S" "E = []"
apply (-, unfold \<open>\<pi> = []\<close> lins2_def, simp_all)
apply (unfold empty_eq_Plus_conv)
apply (unfold G_def vertices.simps edges.simps, auto)
by (metis \<open>S |\<subseteq>| V\<close> less_eq_fset.rep_eq subset_antisym)
with Suc.prems(4) have False by (simp add: G_def)
thus ?thesis by auto
next
case (Cons x \<pi>')
note \<pi>_def = this
show ?thesis
proof (cases x)
case (Inl v)
note x_def = this
have "v |\<notin>| S \<and> v |\<in>| V"
apply (subgoal_tac "v \<in> fset V - fset S")
apply (simp)
apply (subgoal_tac "Inl v \<in> (fset V - fset S) <+> set E")
apply (metis Inl_inject Inr_not_Inl PlusE)
apply (metis lins2(1) lins2(2) Cons G_def Inl distinct.simps(2)
distinct_length_2_or_more edges.simps vertices.simps)
done
hence v_notin_S: "v |\<notin>| S" and v_in_V: "v |\<in>| V" by auto
have v_initial_not_S: "v |\<in>| initials G - S"
apply (simp only: G_def initials_def vertices.simps edges.simps)
apply (simp only: fminus_iff)
apply (simp only: conj_commute, intro conjI, rule v_notin_S)
apply (subgoal_tac
"v \<in> fset (ffilter (\<lambda>v. \<forall>e\<in>set E. v |\<notin>| thd3 e) V)")
apply simp
apply (simp only: filter_fset, simp, simp only: conj_commute)
apply (intro conjI ballI notI)
apply (insert v_in_V, simp)
proof -
fix e :: edge
assume "v \<in> fset (thd3 e)"
then have "v |\<in>| (thd3 e)" by auto
assume "e \<in> set E"
hence "Inr e \<in> set \<pi>" using lins2(2) by (auto simp add: G_def)
then obtain j where
"j < length \<pi>" "0 < length \<pi>" "\<pi>!j = Inr e" "\<pi>!0 = Inl v"
by (metis \<pi>_def x_def in_set_conv_nth length_pos_if_in_set nth_Cons_0)
with lins2(4)[OF this \<open>v |\<in>| (thd3 e)\<close>] show False by auto
qed
define S' where "S' = {|v|} |\<union>| S"
define \<Pi>' where "\<Pi>' = mk_ps_chain \<lbrace> initial_ps2 S' G \<rbrace> \<pi>'"
hence pre_\<Pi>': "pre \<Pi>' = initial_ps2 S' G"
by (metis pre.simps(1) pre_mk_ps_chain)
define \<sigma> where "\<sigma> = [ initials G - ({|v|} |\<union>| S) |=> Top ] ++\<^sub>f [ S |=> Bot ]"
have "wf_ps_chain \<Pi>' \<and> (post \<Pi>' = [terminals G |=> Bot])"
proof (intro Suc.hyps[of "S'"])
show "S' |\<subseteq>| initials G"
apply (unfold S'_def, auto)
- apply (metis fmember_iff_member_fset fminus_iff v_initial_not_S)
- by (metis Suc.prems(1) fmember_iff_member_fset fset_rev_mp)
+ apply (metis fminus_iff v_initial_not_S)
+ by (metis Suc.prems(1) fset_rev_mp)
next
show "wf_dia G" by (rule Suc.prems(2))
next
show "\<Pi>' \<in> ps_chains2 S' G"
apply (unfold ps_chains2_def \<Pi>'_def)
apply (intro imageI)
apply (unfold S'_def)
apply (intro next_lins2_vertex)
apply (fold x_def, fold \<pi>_def)
apply (rule \<pi>_in)
by (metis v_notin_S)
next
show "fcard G^V + length G^E = k + fcard S'"
apply (unfold S'_def)
by (auto simp add: Suc.prems(4) fcard_finsert_disjoint[OF v_notin_S])
qed
hence
wf_\<Pi>': "wf_ps_chain \<Pi>'" and
post_\<Pi>': "post \<Pi>' = [terminals G |=> Bot]"
by auto
show ?thesis
proof (intro conjI)
have 1: "fmdom [ {|v|} |=> Bot ]
|\<inter>| fmdom ([ initials G - ({|v|} |\<union>| S) |=> Top ] ++\<^sub>f
[ S |=> Bot ]) = {||}"
by (metis (no_types) fdom_make_fmap fmdom_add
bot_least funion_iff finter_finsert_left le_iff_inf
fminus_iff finsert_fsubset sup_ge1 v_initial_not_S)
show "wf_ps_chain \<Pi>"
using [[unfold_abs_def = false]]
apply (simp only: \<Pi>_def \<pi>_def x_def mk_ps_chain_cons)
apply simp
apply (unfold mk_ps_chain_ccons)
apply (fold next_initial_ps2_vertex S'_def)
apply (fold \<Pi>'_def)
apply (unfold wf_ps_chain_def chain_all.simps conj_commute)
apply (intro conjI)
apply (fold wf_ps_chain_def, rule wf_\<Pi>')
apply (intro wf_ps_triple_nodeI exI[of _ "\<sigma>"] conjI)
apply (unfold \<sigma>_def fmdom_add fdom_make_fmap)
apply (metis finsertI1 fminus_iff funion_iff v_notin_S)
apply (unfold pre_\<Pi>' initial_ps2_def S'_def)
apply (unfold fmap_add_commute[OF 1])
apply (unfold fmadd_assoc)
apply (fold fmadd_assoc[of _ "[ S |=> Bot ]"])
apply (unfold make_fmap_union sup.commute[of "{|v|}"])
apply (unfold fminus_funion)
using v_initial_not_S apply auto
by (metis (opaque_lifting, no_types) finsert_absorb finsert_fminus_single finter_fminus
inf_commute inf_idem v_initial_not_S)
next
show "post \<Pi> = [ terminals G |=> Bot ]"
apply (unfold \<Pi>_def \<pi>_def x_def mk_ps_chain_cons, simp)
apply (unfold mk_ps_chain_ccons post.simps)
apply (fold next_initial_ps2_vertex S'_def)
apply (fold \<Pi>'_def, rule post_\<Pi>')
done
qed
next
case (Inr e)
note x_def = this
define vs where "vs = fst3 e"
define ws where "ws = thd3 e"
obtain c where e_def: "e = (vs, c, ws)"
by (metis vs_def ws_def fst3_simp thd3_simp prod_cases3)
have "linearity E" and "acyclicity E" and
e_in_V: "\<And>e. e \<in> set E \<Longrightarrow> fst3 e |\<union>| thd3 e |\<subseteq>| V"
by (insert Suc.prems(2) wf_dia_inv, unfold G_def, blast)+
note lin = linearityD[OF this(1)]
have acy: "\<And>e. e \<in> set E \<Longrightarrow> fst3 e |\<inter>| thd3 e = {||}"
apply (fold fset_cong, insert \<open>acyclicity E\<close>)
apply (unfold acyclicity_def acyclic_def, auto)
done
note lins = lins2D[OF \<pi>_in]
have e_in_E: "e \<in> set E"
apply (subgoal_tac "set \<pi> = (fset G^V - fset S) <+> set G^E")
apply (unfold \<pi>_def x_def G_def edges.simps, auto)[1]
apply (simp add: lins(2))
done
have vs_in_S: "vs |\<subseteq>| S"
apply (insert e_in_V[OF e_in_E])
apply (unfold less_eq_fset.rep_eq)
apply (intro subsetI)
apply (unfold vs_def)
apply (rule ccontr)
apply (subgoal_tac "x \<in> fset V")
prefer 2
apply (auto)
proof -
fix v
assume a: "v \<in> fset (fst3 e)"
assume "v \<notin> fset S" and "v \<in> fset V"
hence "Inl v \<in> set \<pi>"
by (metis (lifting) DiffI G_def InlI lins(2) vertices.simps)
then obtain i where
"i < length \<pi>" "0 < length \<pi>" "\<pi>!i = Inl v" "\<pi>!0 = Inr e"
by (metis Cons Inr in_set_conv_nth length_pos_if_in_set nth_Cons_0)
from lins(3)[OF this] show "False" by (auto simp add: a)
qed
have "ws |\<inter>| (initials G) = {||}"
apply (insert e_in_V[OF e_in_E])
- apply (unfold initials_def less_eq_fset.rep_eq fmember_iff_member_fset, fold fset_cong)
+ apply (unfold initials_def less_eq_fset.rep_eq, fold fset_cong)
apply (unfold ws_def G_def, auto simp add: e_in_E)
done
define S' where "S' = S - vs"
define V' where "V' = V - vs"
define E' where "E' = removeAll e E"
define G' where "G' = Graph V' \<Lambda> E'"
define \<Pi>' where "\<Pi>' = mk_ps_chain \<lbrace> initial_ps2 S' G' \<rbrace> \<pi>'"
hence pre_\<Pi>': "pre \<Pi>' = initial_ps2 S' G'"
by (metis pre.simps(1) pre_mk_ps_chain)
define \<sigma> where "\<sigma> = [ initials G - S |=> Top ] ++\<^sub>f [ S - vs |=> Bot ]"
have next_initial_ps2: "initial_ps2 S' G'
= initial_ps2 S G \<ominus> vs ++\<^sub>f [ws |=> Top]"
using next_initial_ps2_edge[OF G_def _ _ _ e_in_E _ Suc.prems(1)
Suc.prems(2)] G'_def E'_def vs_def ws_def V'_def vs_in_S S'_def
by auto
have "wf_ps_chain \<Pi>' \<and> post \<Pi>' = [ terminals G' |=> Bot ]"
proof (intro Suc.hyps[of "S'"])
show "S' |\<subseteq>| initials G'"
apply (insert Suc.prems(1))
apply (unfold G'_def G_def initials_def)
apply (unfold less_eq_fset.rep_eq S'_def E'_def V'_def)
apply auto
done
next
from Suc.prems(2) have "wf_dia (Graph V \<Lambda> E)"
by (unfold G_def)
note wf_G = wf_dia_inv[OF this]
show "wf_dia G'"
apply (unfold G'_def V'_def E'_def)
apply (insert wf_G e_in_E vs_in_S Suc.prems(1))
apply (unfold vs_def)
apply (intro wf_dia)
apply (unfold linearity_def initials_def G_def)
- apply (fold fset_cong, unfold less_eq_fset.rep_eq fmember_iff_member_fset)
+ apply (fold fset_cong, unfold less_eq_fset.rep_eq)
apply (simp, simp)
apply (unfold acyclicity_def, rule acyclic_subset)
apply (auto simp add: distinct_removeAll)
apply (metis (lifting) IntI empty_iff)
done
next
show "\<Pi>' \<in> ps_chains2 S' G'"
apply (unfold \<Pi>_def \<Pi>'_def ps_chains2_def)
apply (intro imageI)
apply (unfold S'_def G'_def V'_def E'_def)
apply (intro next_lins2_edge)
apply (metis \<pi>_def G_def x_def \<pi>_in)
by (simp only: vs_in_S e_def)+
next
have "vs |\<subseteq>| V" by (metis (lifting) \<open>S |\<subseteq>| V\<close> order_trans vs_in_S)
have "distinct E" using \<open>linearity E\<close> linearity_def by auto
show "fcard G'^V + length G'^E = k + fcard S'"
apply (insert Suc.prems(4))
apply (unfold G_def G'_def vertices.simps edges.simps)
apply (unfold V'_def E'_def S'_def)
apply (unfold fcard_funion_fsubset[OF \<open>vs |\<subseteq>| V\<close>])
apply (unfold fcard_funion_fsubset[OF \<open>vs |\<subseteq>| S\<close>])
apply (fold distinct_remove1_removeAll[OF \<open>distinct E\<close>])
apply (unfold length_remove1)
apply (simp add: e_in_E)
apply (drule arg_cong[of _ _ "\<lambda>x. x - fcard vs - 1"])
apply (subst (asm) add_diff_assoc2[symmetric])
apply (simp add: fcard_mono[OF \<open>vs |\<subseteq>| V\<close>])
apply (subst add_diff_assoc, insert length_pos_if_in_set[OF e_in_E], arith, auto)
apply (subst add_diff_assoc, auto simp add: fcard_mono[OF \<open>vs |\<subseteq>| S\<close>])
done
qed
hence
wf_\<Pi>': "wf_ps_chain \<Pi>'" and
post_\<Pi>': "post \<Pi>' = [ terminals G' |=> Bot ]"
by auto
have terms_same: "terminals G = terminals G'"
apply (unfold G'_def G_def terminals_def edges.simps vertices.simps)
apply (unfold E'_def V'_def)
apply (fold fset_cong, auto simp add: e_in_E vs_def)
done
have 1: "fmdom [ fst3 e |=> Bot ] |\<inter>|
fmdom([ ffilter (\<lambda>v. \<forall>e\<in>set E. v |\<notin>| thd3 e) V - S |=> Top ]
++\<^sub>f [ S - fst3 e |=> Bot ]) = {||}"
apply (unfold fmdom_add fdom_make_fmap)
apply (fold fset_cong)
apply auto
apply (metis in_mono less_eq_fset.rep_eq vs_def vs_in_S)
done
show ?thesis
proof (intro conjI)
show "wf_ps_chain \<Pi>"
using [[unfold_abs_def = false]]
apply (unfold \<Pi>_def \<pi>_def x_def mk_ps_chain_cons)
apply simp
apply (unfold mk_ps_chain_ccons)
apply (fold vs_def ws_def)
apply (fold next_initial_ps2)
apply (fold \<Pi>'_def)
apply (unfold wf_ps_chain_def chain_all.simps conj_commute)
apply (intro conjI)
apply (fold wf_ps_chain_def)
apply (rule wf_\<Pi>')
apply (intro wf_ps_triple_edgeI exI[of _ "\<sigma>"])
apply (unfold e_def fst3_simp thd3_simp \<sigma>_def, intro conjI)
apply (insert Suc.prems(1))
apply (unfold pre_\<Pi>' initial_ps2_def initials_def)
apply (insert vs_in_S acy[OF e_in_E])
apply (fold fset_cong)
apply (unfold less_eq_fset.rep_eq)[1]
apply (unfold G_def G'_def vs_def ws_def V'_def E'_def S'_def)
apply (unfold vertices.simps edges.simps)
apply (unfold fmap_add_commute[OF 1])
apply (fold fmadd_assoc)
apply (unfold make_fmap_union)
apply (auto simp add: fdom_make_fmap e_in_E)[1]
apply simp
apply (unfold fmadd_assoc)
apply (unfold make_fmap_union)
apply (metis (lifting) funion_absorb2 vs_def vs_in_S)
apply (intro arg_cong2[of _ _ "[ S - fst3 e |=> Bot ]"
"[ S - fst3 e |=> Bot ]" "(++\<^sub>f)"])
apply (intro arg_cong2[of _ _ "Top" "Top" "make_fmap"])
defer 1
apply (simp, simp)
apply (fold fset_cong)
- apply (unfold less_eq_fset.rep_eq fmember_iff_member_fset, simp)
+ apply (unfold less_eq_fset.rep_eq, simp)
apply (elim conjE)
apply (intro set_eqI iffI, simp_all)
apply (elim conjE, intro disjI conjI ballI, simp)
apply (case_tac "ea=e", simp_all)
apply (elim disjE conjE, intro conjI ballI impI, simp_all)
apply (insert e_in_E lin(2))[1]
apply (subst (asm) (2) fset_cong[symmetric])
apply (elim conjE)
apply (subst (asm) inter_fset)
apply (subst (asm) fset_simps)
apply (insert disjoint_iff_not_equal)[1]
apply blast
apply (metis G_def Suc(3) e_in_E subsetD less_eq_fset.rep_eq wf_dia_inv')
prefer 2
apply (metis (lifting) IntI Suc(2) \<open>ws |\<inter>| initials G = {||}\<close>
empty_iff fset_simps(1) in_mono inter_fset less_eq_fset.rep_eq ws_def)
apply auto
done
next
show "post \<Pi> = [terminals G |=> Bot]"
apply (unfold \<Pi>_def \<pi>_def x_def mk_ps_chain_cons)
apply simp
apply (unfold mk_ps_chain_ccons post.simps)
apply (fold vs_def ws_def)
apply (fold next_initial_ps2)
apply (fold \<Pi>'_def)
apply (unfold terms_same)
apply (rule post_\<Pi>')
done
qed
qed
qed
qed
corollary wf_chains:
assumes "wf_dia G"
assumes "\<Pi> \<in> ps_chains G"
shows "wf_ps_chain \<Pi> \<and> post \<Pi> = [ terminals G |=> Bot ]"
apply (intro wf_chains2[of "{||}"], insert assms(2))
by (auto simp add: assms(1) ps_chains_is_ps_chains2_with_empty_S fcard_fempty)
subsection \<open>Interface chains\<close>
type_synonym int_chain = "(interface, assertion_gadget + command_gadget) chain"
text \<open>An interface chain is similar to a proofstate chain. However, where a
proofstate chain talks about nodes and edges, an interface chain talks about
the assertion-gadgets and command-gadgets that label those nodes and edges
in a diagram. And where a proofstate chain talks about proofstates, an
interface chain talks about the interfaces obtained from those proofstates.
The following functions convert a proofstate chain into an
interface chain.\<close>
definition
ps_to_int :: "[diagram, proofstate] \<Rightarrow> interface"
where
"ps_to_int G \<sigma> \<equiv>
\<Otimes>v |\<in>| fmdom \<sigma>. case_topbot top_ass bot_ass (lookup \<sigma> v) (G^\<Lambda> v)"
definition
ps_chain_to_int_chain :: "[diagram, ps_chain] \<Rightarrow> int_chain"
where
"ps_chain_to_int_chain G \<Pi> \<equiv>
chainmap (ps_to_int G) ((case_sum (Inl \<circ> G^\<Lambda>) (Inr \<circ> snd3))) \<Pi>"
lemma ps_chain_to_int_chain_simp:
"ps_chain_to_int_chain (Graph V \<Lambda> E) \<Pi> =
chainmap (ps_to_int (Graph V \<Lambda> E)) ((case_sum (Inl \<circ> \<Lambda>) (Inr \<circ> snd3))) \<Pi>"
by (simp add: ps_chain_to_int_chain_def)
subsection \<open>Soundness proof\<close>
text \<open>We assume that @{term wr_com} always returns @{term "{}"}. This is
equivalent to changing our axiomatization of separation logic such that the
frame rule has no side-condition. One way to obtain a separation logic
lacking a side-condition on its frame rule is to use variables-as-
resource.
We proceed by induction on the proof rules for graphical diagrams. We
show that: (1) if a diagram @{term G} is provable w.r.t. interfaces
@{term P} and @{term Q}, then @{term P} and @{term Q} are the top and bottom
interfaces of @{term G}, and that the Hoare triple @{term "(asn P,
c, asn Q)"} is provable for each command @{term c} that can be extracted
from @{term G}; (2) if a command-gadget @{term C} is provable w.r.t.
interfaces @{term P} and @{term Q}, then the Hoare triple @{term "(asn P,
c, asn Q)"} is provable for each command @{term c} that can be extracted
from @{term C}; and (3) if an assertion-gadget @{term A} is provable, and if
the top and bottom interfaces of @{term A} are @{term P} and @{term Q}
respectively, then the Hoare triple @{term "(asn P, c, asn Q)"} is provable
for each command @{term c} that can be extracted from @{term A}.\<close>
lemma soundness_graphical_helper:
assumes no_var_interference: "\<And>c. wr_com c = {}"
shows
"(prov_dia G P Q \<longrightarrow>
(P = top_dia G \<and> Q = bot_dia G \<and>
(\<forall>c. coms_dia G c \<longrightarrow> prov_triple (asn P, c, asn Q))))
\<and> (prov_com C P Q \<longrightarrow>
(\<forall>c. coms_com C c \<longrightarrow> prov_triple (asn P, c, asn Q)))
\<and> (prov_ass A \<longrightarrow>
(\<forall>c. coms_ass A c \<longrightarrow> prov_triple (asn (top_ass A), c, asn (bot_ass A))))"
proof (induct rule: prov_dia_prov_com_prov_ass.induct)
case (Skip p)
thus ?case
apply (intro allI impI, elim conjE coms_skip_inv)
apply (auto simp add: prov_triple.skip)
done
next
case (Exists G P Q x)
thus ?case
apply (intro allI impI, elim conjE coms_exists_inv)
apply (auto simp add: prov_triple.exists)
done
next
case (Basic P c Q)
thus ?case
by (intro allI impI, elim conjE coms_basic_inv, auto)
next
case (Choice G P Q H)
thus ?case
apply (intro allI impI, elim conjE coms_choice_inv)
apply (auto simp add: prov_triple.choose)
done
next
case (Loop G P)
thus ?case
apply (intro allI impI, elim conjE coms_loop_inv)
apply (auto simp add: prov_triple.loop)
done
next
case (Main G)
thus ?case
apply (intro conjI)
apply (simp, simp)
apply (intro allI impI)
apply (elim coms_main_inv, simp)
proof -
fix c V \<Lambda> E
fix \<pi>::"lin"
fix cs::"command list"
assume wf_G: "wf_dia (Graph V \<Lambda> E)"
assume "\<And>v. v \<in> fset V \<Longrightarrow> \<forall>c. coms_ass (\<Lambda> v) c \<longrightarrow>
prov_triple (asn (top_ass (\<Lambda> v)), c, asn (bot_ass (\<Lambda> v)))"
hence prov_vertex: "\<And>v c P Q F. \<lbrakk> coms_ass (\<Lambda> v) c; v \<in> fset V;
P = (top_ass (\<Lambda> v) \<otimes> F) ; Q = (bot_ass (\<Lambda> v) \<otimes> F) \<rbrakk>
\<Longrightarrow> prov_triple (asn P, c, asn Q)"
by (auto simp add: prov_triple.frame no_var_interference)
assume "\<And>e. e \<in> set E \<Longrightarrow> \<forall>c. coms_com (snd3 e) c \<longrightarrow> prov_triple
(asn (\<Otimes>v|\<in>|fst3 e. bot_ass (\<Lambda> v)),c,asn (\<Otimes>v|\<in>|thd3 e. top_ass (\<Lambda> v)))"
hence prov_edge: "\<And>e c P Q F. \<lbrakk> e \<in> set E ; coms_com (snd3 e) c ;
P = ((\<Otimes>v|\<in>|fst3 e. bot_ass (\<Lambda> v)) \<otimes> F) ;
Q = ((\<Otimes>v|\<in>|thd3 e. top_ass (\<Lambda> v)) \<otimes> F) \<rbrakk>
\<Longrightarrow> prov_triple (asn P, c, asn Q)"
by (auto simp add: prov_triple.frame no_var_interference)
assume len_cs: "length cs = length \<pi>"
assume "\<forall>i<length \<pi>.
case_sum (coms_ass \<circ> \<Lambda>) (coms_com \<circ> snd3) (\<pi> ! i) (cs ! i)"
hence \<pi>_cs: "\<And>i. i < length \<pi> \<Longrightarrow>
case_sum (coms_ass \<circ> \<Lambda>) (coms_com \<circ> snd3) (\<pi> ! i) (cs ! i)" by auto
assume G_def: "G = Graph V \<Lambda> E"
assume c_def: "c = foldr (;;) cs Skip"
assume \<pi>_lin: "\<pi> \<in> lins (Graph V \<Lambda> E)"
note lins = linsD[OF \<pi>_lin]
define \<Pi> where "\<Pi> = mk_ps_chain \<lbrace> initial_ps G \<rbrace> \<pi>"
have "\<Pi> \<in> ps_chains G" by (simp add: \<pi>_lin \<Pi>_def ps_chains_def G_def)
hence 1: "post \<Pi> = [ terminals G |=> Bot ]"
and 2: "chain_all wf_ps_triple \<Pi>"
by (insert wf_chains G_def wf_G, auto simp add: wf_ps_chain_def)
show "prov_triple (asn (\<Otimes>v|\<in>|initials (Graph V \<Lambda> E). top_ass (\<Lambda> v)),
foldr (;;) cs Skip, asn (\<Otimes>v|\<in>|terminals (Graph V \<Lambda> E). bot_ass (\<Lambda> v)))"
using [[unfold_abs_def = false]]
apply (intro seq_fold[of _ "ps_chain_to_int_chain G \<Pi>"])
apply (unfold len_cs)
apply (unfold ps_chain_to_int_chain_def chainmap_preserves_length \<Pi>_def)
apply (unfold mk_ps_chain_preserves_length, simp)
apply (unfold pre_chainmap post_chainmap)
apply (unfold pre_mk_ps_chain pre.simps)
apply (fold \<Pi>_def, unfold 1)
apply (unfold initial_ps_def)
apply (unfold ps_to_int_def)
apply (unfold fdom_make_fmap)
apply (unfold G_def labelling.simps, fold G_def)
apply (subgoal_tac "\<forall>v \<in> fset (initials G). top_ass (\<Lambda> v) =
case_topbot top_ass bot_ass (lookup [ initials G |=> Top ] v) (\<Lambda> v)")
apply (unfold iter_hcomp_cong, simp)
apply (metis lookup_make_fmap topbot.simps(3))
apply (subgoal_tac "\<forall>v \<in> fset (terminals G). bot_ass (\<Lambda> v) =
case_topbot top_ass bot_ass (lookup [ terminals G |=> Bot ] v) (\<Lambda> v)")
apply (unfold iter_hcomp_cong, simp)
apply (metis lookup_make_fmap topbot.simps(4), simp)
apply (unfold G_def, fold ps_chain_to_int_chain_simp G_def)
proof -
fix i
assume "i < length \<pi>"
hence "i < chainlen \<Pi>"
by (metis \<Pi>_def add_0_left chainlen.simps(1)
mk_ps_chain_preserves_length)
hence wf_\<Pi>i: "wf_ps_triple (nthtriple \<Pi> i)"
by (insert 2, simp add: chain_all_nthtriple)
show "prov_triple (asn (fst3 (nthtriple (ps_chain_to_int_chain G \<Pi>) i)),
cs ! i, asn (thd3 (nthtriple (ps_chain_to_int_chain G \<Pi>) i)))"
apply (unfold ps_chain_to_int_chain_def)
apply (unfold nthtriple_chainmap[OF \<open>i < chainlen \<Pi>\<close>])
apply (unfold fst3_simp thd3_simp)
proof (cases "\<pi>!i")
case (Inl v)
have "snd3 (nthtriple \<Pi> i) = Inl v"
apply (unfold snds_of_triples_form_comlist[OF \<open>i < chainlen \<Pi>\<close>])
apply (auto simp add: \<Pi>_def comlist_mk_ps_chain Inl)
done
with wf_\<Pi>i wf_ps_triple_def obtain \<sigma> where
v_notin_\<sigma>: "v |\<notin>| fmdom \<sigma>" and
fst_\<Pi>i: "fst3 (nthtriple \<Pi> i) = [ {|v|} |=> Top ] ++\<^sub>f \<sigma>" and
thd_\<Pi>i: "thd3 (nthtriple \<Pi> i) = [ {|v|} |=> Bot ] ++\<^sub>f \<sigma>" by auto
show "prov_triple (asn (ps_to_int G (fst3 (nthtriple \<Pi> i))),
cs ! i, asn (ps_to_int G (thd3 (nthtriple \<Pi> i))))"
apply (intro prov_vertex[where v=v])
apply (metis (no_types) Inl \<open>i < length \<pi>\<close> \<pi>_cs o_def sum.simps(5))
apply (metis (lifting) Inl lins(2) Inl_not_Inr PlusE \<open>i < length \<pi>\<close>
nth_mem sum.simps(1) vertices.simps)
apply (unfold fst_\<Pi>i thd_\<Pi>i)
apply (unfold ps_to_int_def)
apply (unfold fmdom_add fdom_make_fmap)
apply (unfold finsert_is_funion[symmetric])
apply (insert v_notin_\<sigma>)
apply (unfold iter_hcomp_insert)
apply (unfold lookup_union2 lookup_make_fmap1)
apply (unfold G_def labelling.simps)
apply (subgoal_tac "\<forall>va \<in> fset (fmdom \<sigma>). case_topbot top_ass bot_ass
(lookup ([ {|v|} |=> Top ] ++\<^sub>f \<sigma>) va) (\<Lambda> va) =
case_topbot top_ass bot_ass (lookup ([{|v|} |=> Bot] ++\<^sub>f \<sigma>) va)(\<Lambda> va)")
apply (unfold iter_hcomp_cong, simp)
- apply (metis fmember_iff_member_fset lookup_union1, simp)
+ apply (metis lookup_union1, simp)
done
next
case (Inr e)
have "snd3 (nthtriple \<Pi> i) = Inr e"
apply (unfold snds_of_triples_form_comlist[OF \<open>i < chainlen \<Pi>\<close>])
apply (auto simp add: \<Pi>_def comlist_mk_ps_chain Inr)
done
with wf_\<Pi>i wf_ps_triple_def obtain \<sigma> where
fst_e_disjoint_\<sigma>: "fst3 e |\<inter>| fmdom \<sigma> = {||}" and
thd_e_disjoint_\<sigma>: "thd3 e |\<inter>| fmdom \<sigma> = {||}" and
fst_\<Pi>i: "fst3 (nthtriple \<Pi> i) = [ fst3 e |=> Bot ] ++\<^sub>f \<sigma>" and
thd_\<Pi>i: "thd3 (nthtriple \<Pi> i) = [ thd3 e |=> Top ] ++\<^sub>f \<sigma>"
by (auto simp add: inf_sup_distrib2)
show "prov_triple (asn (ps_to_int G (fst3 (nthtriple \<Pi> i))),
cs ! i, asn (ps_to_int G (thd3 (nthtriple \<Pi> i))))"
apply (intro prov_edge[where e=e])
apply (subgoal_tac "Inr e \<in> set \<pi>")
apply (metis Inr_not_Inl PlusE edges.simps lins(2) sum.simps(2))
apply (metis Inr \<open>i < length \<pi>\<close> nth_mem)
apply (metis (no_types) Inr \<open>i < length \<pi>\<close> \<pi>_cs o_def sum.simps(6))
apply (unfold fst_\<Pi>i thd_\<Pi>i)
apply (unfold ps_to_int_def)
apply (unfold G_def labelling.simps)
apply (unfold fmdom_add fdom_make_fmap)
apply (insert fst_e_disjoint_\<sigma>)
apply (unfold iter_hcomp_union)
apply (subgoal_tac "\<forall>v \<in> fset (fst3 e). case_topbot top_ass bot_ass
(lookup ([ fst3 e |=> Bot ] ++\<^sub>f \<sigma>) v) (\<Lambda> v) = bot_ass (\<Lambda> v)")
apply (unfold iter_hcomp_cong)
apply (simp)
apply (intro ballI)
apply (subgoal_tac "v |\<notin>| fmdom \<sigma>")
apply (unfold lookup_union2)
apply (metis lookup_make_fmap topbot.simps(4))
- apply (metis fempty_iff finterI fmember_iff_member_fset)
+ apply (metis fempty_iff finterI)
apply (insert thd_e_disjoint_\<sigma>)
apply (unfold iter_hcomp_union)
apply (subgoal_tac "\<forall>v \<in> fset (thd3 e). case_topbot top_ass bot_ass
(lookup ([ thd3 e |=> Top ] ++\<^sub>f \<sigma>) v) (\<Lambda> v) = top_ass (\<Lambda> v)")
apply (unfold iter_hcomp_cong)
apply (subgoal_tac "\<forall>v \<in> fset (fmdom \<sigma>). case_topbot top_ass bot_ass
(lookup ([ thd3 e |=> Top ] ++\<^sub>f \<sigma>) v) (\<Lambda> v) =
case_topbot top_ass bot_ass (lookup ([fst3 e |=> Bot] ++\<^sub>f \<sigma>) v) (\<Lambda> v)")
apply (unfold iter_hcomp_cong)
apply simp
apply (intro ballI)
apply (subgoal_tac "v |\<in>| fmdom \<sigma>")
apply (unfold lookup_union1, auto)
apply (subgoal_tac "v |\<notin>| fmdom \<sigma>")
apply (unfold lookup_union2)
apply (metis lookup_make_fmap topbot.simps(3))
- by (metis fempty_iff finterI fmember_iff_member_fset)
+ by (metis fempty_iff finterI)
qed
qed
qed
qed
text \<open>The soundness theorem states that any diagram provable using the
proof rules for ribbons can be recreated as a valid proof in separation
logic.\<close>
corollary soundness_graphical:
assumes "\<And>c. wr_com c = {}"
assumes "prov_dia G P Q"
shows "\<forall>c. coms_dia G c \<longrightarrow> prov_triple (asn P, c, asn Q)"
using soundness_graphical_helper[OF assms(1)] and assms(2) by auto
end
diff --git a/thys/Ribbon_Proofs/Ribbons_Interfaces.thy b/thys/Ribbon_Proofs/Ribbons_Interfaces.thy
--- a/thys/Ribbon_Proofs/Ribbons_Interfaces.thy
+++ b/thys/Ribbon_Proofs/Ribbons_Interfaces.thy
@@ -1,210 +1,211 @@
section \<open>Ribbon proof interfaces\<close>
theory Ribbons_Interfaces imports
Ribbons_Basic
Proofchain
"HOL-Library.FSet"
begin
text \<open>Interfaces are the top and bottom boundaries through which diagrams
can be connected into a surrounding context. For instance, when composing two
diagrams vertically, the bottom interface of the upper diagram must match the
top interface of the lower diagram.
We define a datatype of concrete interfaces. We then quotient by the
associativity, commutativity and unity properties of our
horizontal-composition operator.\<close>
subsection \<open>Syntax of interfaces\<close>
datatype conc_interface =
Ribbon_conc "assertion"
| HComp_int_conc "conc_interface" "conc_interface" (infix "\<otimes>\<^sub>c" 50)
| Emp_int_conc ("\<epsilon>\<^sub>c")
| Exists_int_conc "string" "conc_interface"
text \<open>We define an equivalence on interfaces. The first three rules make this
an equivalence relation. The next three make it a congruence. The next two
identify interfaces up to associativity and commutativity of @{term "(\<otimes>\<^sub>c)"}
The final two make @{term "\<epsilon>\<^sub>c"} the left and right unit of @{term "(\<otimes>\<^sub>c)"}.
\<close>
inductive
equiv_int :: "conc_interface \<Rightarrow> conc_interface \<Rightarrow> bool" (infix "\<simeq>" 45)
where
refl: "P \<simeq> P"
| sym: "P \<simeq> Q \<Longrightarrow> Q \<simeq> P"
| trans: "\<lbrakk>P \<simeq> Q; Q \<simeq> R\<rbrakk> \<Longrightarrow> P \<simeq> R"
| cong_hcomp1: "P \<simeq> Q \<Longrightarrow> P' \<otimes>\<^sub>c P \<simeq> P' \<otimes>\<^sub>c Q"
| cong_hcomp2: "P \<simeq> Q \<Longrightarrow> P \<otimes>\<^sub>c P' \<simeq> Q \<otimes>\<^sub>c P'"
| cong_exists: "P \<simeq> Q \<Longrightarrow> Exists_int_conc x P \<simeq> Exists_int_conc x Q"
| hcomp_conc_assoc: "P \<otimes>\<^sub>c (Q \<otimes>\<^sub>c R) \<simeq> (P \<otimes>\<^sub>c Q) \<otimes>\<^sub>c R"
| hcomp_conc_comm: "P \<otimes>\<^sub>c Q \<simeq> Q \<otimes>\<^sub>c P"
| hcomp_conc_unit1: "\<epsilon>\<^sub>c \<otimes>\<^sub>c P \<simeq> P"
| hcomp_conc_unit2: "P \<otimes>\<^sub>c \<epsilon>\<^sub>c \<simeq> P"
lemma equiv_int_cong_hcomp:
"\<lbrakk> P \<simeq> Q ; P' \<simeq> Q' \<rbrakk> \<Longrightarrow> P \<otimes>\<^sub>c P' \<simeq> Q \<otimes>\<^sub>c Q'"
by (metis cong_hcomp1 cong_hcomp2 equiv_int.trans)
quotient_type interface = "conc_interface" / "equiv_int"
apply (intro equivpI)
apply (intro reflpI, simp add: equiv_int.refl)
apply (intro sympI, simp add: equiv_int.sym)
apply (intro transpI, elim equiv_int.trans, simp add: equiv_int.refl)
done
lift_definition
Ribbon :: "assertion \<Rightarrow> interface"
is "Ribbon_conc" .
lift_definition
Emp_int :: "interface" ("\<epsilon>")
is "\<epsilon>\<^sub>c" .
lift_definition
Exists_int :: "string \<Rightarrow> interface \<Rightarrow> interface"
is "Exists_int_conc"
by (rule equiv_int.cong_exists)
lift_definition
HComp_int :: "interface \<Rightarrow> interface \<Rightarrow> interface" (infix "\<otimes>" 50)
is "HComp_int_conc" by (rule equiv_int_cong_hcomp)
lemma hcomp_comm:
"(P \<otimes> Q) = (Q \<otimes> P)"
by (rule hcomp_conc_comm[Transfer.transferred])
lemma hcomp_assoc:
"(P \<otimes> (Q \<otimes> R)) = ((P \<otimes> Q) \<otimes> R)"
by (rule hcomp_conc_assoc[Transfer.transferred])
lemma emp_hcomp:
"\<epsilon> \<otimes> P = P"
by (rule hcomp_conc_unit1[Transfer.transferred])
lemma hcomp_emp:
"P \<otimes> \<epsilon> = P"
by (rule hcomp_conc_unit2[Transfer.transferred])
lemma comp_fun_commute_hcomp:
"comp_fun_commute (\<otimes>)"
by standard (simp add: hcomp_assoc fun_eq_iff, metis hcomp_comm)
subsection \<open>An iterated horizontal-composition operator\<close>
definition iter_hcomp :: "('a fset) \<Rightarrow> ('a \<Rightarrow> interface) \<Rightarrow> interface"
where
"iter_hcomp X f \<equiv> ffold ((\<otimes>) \<circ> f) \<epsilon> X"
syntax "iter_hcomp_syntax" ::
"'a \<Rightarrow> ('a fset) \<Rightarrow> ('a \<Rightarrow> interface) \<Rightarrow> interface"
("(\<Otimes>_|\<in>|_. _)" [0,0,10] 10)
translations "\<Otimes>x|\<in>|M. e" == "CONST iter_hcomp M (\<lambda>x. e)"
term "\<Otimes>P|\<in>|Ps. f P" \<comment> \<open>this is eta-expanded, so prints in expanded form\<close>
term "\<Otimes>P|\<in>|Ps. f" \<comment> \<open>this isn't eta-expanded, so prints as written\<close>
lemma iter_hcomp_cong:
assumes "\<forall>v \<in> fset vs. \<phi> v = \<phi>' v"
shows "(\<Otimes>v|\<in>|vs. \<phi> v) = (\<Otimes>v|\<in>|vs. \<phi>' v)"
using assms unfolding iter_hcomp_def
-by (auto simp add: fmember_iff_member_fset comp_fun_commute.comp_comp_fun_commute comp_fun_commute_hcomp
+by (auto simp add: comp_fun_commute.comp_comp_fun_commute comp_fun_commute_hcomp
intro: ffold_cong)
lemma iter_hcomp_empty:
shows "(\<Otimes>x |\<in>| {||}. p x) = \<epsilon>"
-by (metis comp_fun_commute.ffold_empty comp_fun_commute_hcomp iter_hcomp_def)
+ by (simp add: comp_fun_commute.comp_comp_fun_commute comp_fun_commute.ffold_empty
+ comp_fun_commute_hcomp iter_hcomp_def)
lemma iter_hcomp_insert:
assumes "v |\<notin>| ws"
shows "(\<Otimes>x |\<in>| finsert v ws. p x) = (p v \<otimes> (\<Otimes>x |\<in>| ws. p x))"
proof -
interpret comp_fun_commute "((\<otimes>) \<circ> p)"
by (metis comp_fun_commute.comp_comp_fun_commute comp_fun_commute_hcomp)
from assms show ?thesis unfolding iter_hcomp_def by auto
qed
lemma iter_hcomp_union:
assumes "vs |\<inter>| ws = {||}"
shows "(\<Otimes>x |\<in>| vs |\<union>| ws. p x) = ((\<Otimes>x |\<in>| vs. p x) \<otimes> (\<Otimes>x |\<in>| ws. p x))"
using assms
by (induct vs) (auto simp add: emp_hcomp iter_hcomp_empty iter_hcomp_insert hcomp_assoc)
subsection \<open>Semantics of interfaces\<close>
text \<open>The semantics of an interface is an assertion.\<close>
fun
conc_asn :: "conc_interface \<Rightarrow> assertion"
where
"conc_asn (Ribbon_conc p) = p"
| "conc_asn (P \<otimes>\<^sub>c Q) = (conc_asn P) \<star> (conc_asn Q)"
| "conc_asn (\<epsilon>\<^sub>c) = Emp"
| "conc_asn (Exists_int_conc x P) = Exists x (conc_asn P)"
lift_definition
asn :: "interface \<Rightarrow> assertion"
is "conc_asn"
by (induct_tac rule:equiv_int.induct) (auto simp add: star_assoc star_comm star_rot emp_star)
lemma asn_simps [simp]:
"asn (Ribbon p) = p"
"asn (P \<otimes> Q) = (asn P) \<star> (asn Q)"
"asn \<epsilon> = Emp"
"asn (Exists_int x P) = Exists x (asn P)"
by (transfer, simp)+
subsection \<open>Program variables mentioned in an interface.\<close>
fun
rd_conc_int :: "conc_interface \<Rightarrow> string set"
where
"rd_conc_int (Ribbon_conc p) = rd_ass p"
| "rd_conc_int (P \<otimes>\<^sub>c Q) = rd_conc_int P \<union> rd_conc_int Q"
| "rd_conc_int (\<epsilon>\<^sub>c) = {}"
| "rd_conc_int (Exists_int_conc x P) = rd_conc_int P"
lift_definition
rd_int :: "interface \<Rightarrow> string set"
is "rd_conc_int"
by (induct_tac rule: equiv_int.induct) auto
text \<open>The program variables read by an interface are the same as those read
by its corresponding assertion.\<close>
lemma rd_int_is_rd_ass:
"rd_ass (asn P) = rd_int P"
by (transfer, induct_tac P, auto simp add: rd_star rd_exists rd_emp)
text \<open>Here is an iterated version of the Hoare logic sequencing rule.\<close>
lemma seq_fold:
"\<And>\<Pi>. \<lbrakk> length cs = chainlen \<Pi> ; p1 = asn (pre \<Pi>) ; p2 = asn (post \<Pi>) ;
\<And>i. i < chainlen \<Pi> \<Longrightarrow> prov_triple
(asn (fst3 (nthtriple \<Pi> i)), cs ! i, asn (thd3 (nthtriple \<Pi> i))) \<rbrakk>
\<Longrightarrow> prov_triple (p1, foldr (;;) cs Skip, p2)"
proof (induct cs arbitrary: p1 p2)
case Nil
thus ?case
by (cases \<Pi>, auto simp add: prov_triple.skip)
next
case (Cons c cs)
obtain p x \<Pi>' where \<Pi>_def: "\<Pi> = \<lbrace> p \<rbrace> \<cdot> x \<cdot> \<Pi>'"
by (metis Cons.prems(1) chain.exhaust chainlen.simps(1) impossible_Cons le0)
show ?case
apply (unfold foldr_Cons o_def)
apply (rule prov_triple.seq[where q = "asn (pre \<Pi>')"])
apply (unfold Cons.prems(2) Cons.prems(3) \<Pi>_def pre.simps post.simps)
apply (subst nth_Cons_0[of c cs, symmetric])
apply (subst fst3_simp[of p x "pre \<Pi>'", symmetric])
apply (subst(2) thd3_simp[of p x "pre \<Pi>'", symmetric])
apply (subst(1 2) nthtriple.simps(1)[of p x "\<Pi>'", symmetric])
apply (fold \<Pi>_def, intro Cons.prems(4), simp add: \<Pi>_def)
apply (intro Cons.hyps, insert \<Pi>_def Cons.prems(1), auto)
apply (fold nth_Cons_Suc[of c cs] nthtriple.simps(2)[of p x "\<Pi>'"])
apply (fold \<Pi>_def, intro Cons.prems(4), simp add: \<Pi>_def)
done
qed
end
diff --git a/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy b/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy
--- a/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy
+++ b/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy
@@ -1,2344 +1,2344 @@
(***********************************************************************************
* 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 \<open>Core SC DOM Components\<close>
theory Core_DOM_DOM_Components
imports Core_SC_DOM.Core_DOM
begin
subsection \<open>Components\<close>
locale l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs =
l_get_root_node_defs get_root_node get_root_node_locs +
l_to_tree_order_defs to_tree_order
for get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
begin
definition a_get_dom_component :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_get_dom_component ptr = do {
root \<leftarrow> get_root_node ptr;
to_tree_order root
}"
definition a_is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"a_is_strongly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = (
let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
let arg_components =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_dom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h). set |h \<turnstile> a_get_dom_component ptr|\<^sub>r) in
let arg_components' =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_dom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h'). set |h' \<turnstile> a_get_dom_component ptr|\<^sub>r) in
removed_pointers \<subseteq> arg_components \<and>
added_pointers \<subseteq> arg_components' \<and>
S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \<subseteq> arg_components' \<and>
(\<forall>outside_ptr \<in> fset (object_ptr_kinds h) \<inter> fset (object_ptr_kinds h') -
(\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_dom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))"
definition a_is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"a_is_weakly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = (
let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
let arg_components =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_dom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h). set |h \<turnstile> a_get_dom_component ptr|\<^sub>r) in
let arg_components' =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_dom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h'). set |h' \<turnstile> a_get_dom_component ptr|\<^sub>r) in
removed_pointers \<subseteq> arg_components \<and>
S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \<subseteq> arg_components' \<union> added_pointers \<and>
(\<forall>outside_ptr \<in> fset (object_ptr_kinds h) \<inter> fset (object_ptr_kinds h') -
(\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_dom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))"
lemma "a_is_strongly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' \<Longrightarrow> a_is_weakly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h'"
by(auto simp add: a_is_strongly_dom_component_safe_def a_is_weakly_dom_component_safe_def Let_def)
definition is_document_component :: "(_) object_ptr list \<Rightarrow> bool"
where
"is_document_component c = is_document_ptr_kind (hd c)"
definition is_disconnected_component :: "(_) object_ptr list \<Rightarrow> bool"
where
"is_disconnected_component c = is_node_ptr_kind (hd c)"
end
global_interpretation l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs to_tree_order
defines get_dom_component = a_get_dom_component
and is_strongly_dom_component_safe = a_is_strongly_dom_component_safe
and is_weakly_dom_component_safe = a_is_weakly_dom_component_safe
.
locale l_get_dom_component_defs =
fixes get_dom_component :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
fixes is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
fixes is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
locale l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_to_tree_order_wf +
l_get_dom_component_defs +
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +
l_get_ancestors +
l_get_ancestors_wf +
l_get_root_node +
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent +
l_get_parent_wf +
l_get_element_by +
l_to_tree_order_wf_get_root_node_wf +
(* l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ get_child_nodes +
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ get_child_nodes+
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ "l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.a_to_tree_order get_child_nodes"
for get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog" *)
assumes get_dom_component_impl: "get_dom_component = a_get_dom_component"
assumes is_strongly_dom_component_safe_impl:
"is_strongly_dom_component_safe = a_is_strongly_dom_component_safe"
assumes is_weakly_dom_component_safe_impl:
"is_weakly_dom_component_safe = a_is_weakly_dom_component_safe"
begin
lemmas get_dom_component_def = a_get_dom_component_def[folded get_dom_component_impl]
lemmas is_strongly_dom_component_safe_def =
a_is_strongly_dom_component_safe_def[folded is_strongly_dom_component_safe_impl]
lemmas is_weakly_dom_component_safe_def =
a_is_weakly_dom_component_safe_def[folded is_weakly_dom_component_safe_impl]
lemma get_dom_component_ptr_in_heap:
assumes "h \<turnstile> ok (get_dom_component ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms get_root_node_ptr_in_heap
by(auto simp add: get_dom_component_def)
lemma get_dom_component_ok:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_dom_component ptr)"
using assms
apply(auto simp add: get_dom_component_def a_get_root_node_def intro!: bind_is_OK_pure_I)[1]
using get_root_node_ok to_tree_order_ok get_root_node_ptr_in_heap
apply blast
by (simp add: local.get_root_node_root_in_heap local.to_tree_order_ok)
lemma get_dom_component_ptr:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
shows "ptr \<in> set c"
proof(insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev )
case (step child)
then show ?case
proof (cases "is_node_ptr_kind child")
case True
obtain node_ptr where
node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child"
using \<open>is_node_ptr_kind child\<close> node_ptr_casts_commute3 by blast
have "child |\<in>| object_ptr_kinds h"
using \<open>h \<turnstile> get_dom_component child \<rightarrow>\<^sub>r c\<close> get_dom_component_ptr_in_heap by fast
with node_ptr have "node_ptr |\<in>| node_ptr_kinds h"
by auto
then obtain parent_opt where
parent: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt"
using get_parent_ok \<open>type_wf h\<close> \<open>known_ptrs h\<close>
by fast
then show ?thesis
proof (induct parent_opt)
case None
then have "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
by (simp add: local.get_root_node_no_parent)
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> node_ptr step(2)
apply(auto simp add: get_dom_component_def a_get_root_node_def elim!: bind_returns_result_E2)[1]
using to_tree_order_ptr_in_result returns_result_eq by fastforce
next
case (Some parent_ptr)
then have "h \<turnstile> get_dom_component parent_ptr \<rightarrow>\<^sub>r c"
using step(2) node_ptr \<open>type_wf h\<close> \<open>known_ptrs h\<close> \<open>heap_is_wellformed h\<close>
get_root_node_parent_same
by(auto simp add: get_dom_component_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I)
then have "parent_ptr \<in> set c"
using step node_ptr Some by blast
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> \<open>heap_is_wellformed h\<close> step(2) node_ptr Some
apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
using to_tree_order_parent by blast
qed
next
case False
then show ?thesis
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> step(2)
apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
by (metis is_OK_returns_result_I local.get_root_node_not_node_same
local.get_root_node_ptr_in_heap local.to_tree_order_ptr_in_result returns_result_eq)
qed
qed
lemma get_dom_component_parent_inside:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "cast node_ptr \<in> set c"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some parent"
shows "parent \<in> set c"
proof -
have "parent |\<in>| object_ptr_kinds h"
using assms(6) get_parent_parent_in_heap by blast
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr" and c: "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using assms(4)
by (metis (no_types, opaque_lifting) bind_returns_result_E2 get_dom_component_def get_root_node_pure)
then have "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r root_ptr"
using assms(1) assms(2) assms(3) assms(5) to_tree_order_same_root by blast
then have "h \<turnstile> get_root_node parent \<rightarrow>\<^sub>r root_ptr"
using assms(6) get_root_node_parent_same by blast
then have "h \<turnstile> get_dom_component parent \<rightarrow>\<^sub>r c"
using c get_dom_component_def by auto
then show ?thesis
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
qed
lemma get_dom_component_subset:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "ptr' \<in> set c"
shows "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
proof(insert assms(1) assms(5), induct ptr' rule: heap_wellformed_induct_rev )
case (step child)
then show ?case
proof (cases "is_node_ptr_kind child")
case True
obtain node_ptr where
node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child"
using \<open>is_node_ptr_kind child\<close> node_ptr_casts_commute3 by blast
have "child |\<in>| object_ptr_kinds h"
using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2)
unfolding get_dom_component_def
by (meson bind_returns_result_E2 get_root_node_pure)
with node_ptr have "node_ptr |\<in>| node_ptr_kinds h"
by auto
then obtain parent_opt where
parent: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt"
using get_parent_ok \<open>type_wf h\<close> \<open>known_ptrs h\<close>
by fast
then show ?thesis
proof (induct parent_opt)
case None
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) get_root_node_no_parent node_ptr by blast
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> node_ptr step(2) assms(4) assms(1)
by (metis (no_types) bind_pure_returns_result_I2 bind_returns_result_E2
get_dom_component_def get_root_node_pure is_OK_returns_result_I returns_result_eq
to_tree_order_same_root)
next
case (Some parent_ptr)
then have "h \<turnstile> get_dom_component parent_ptr \<rightarrow>\<^sub>r c"
using step get_dom_component_parent_inside assms node_ptr by blast
then show ?case
using Some node_ptr
apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2
del: bind_pure_returns_result_I intro!: bind_pure_returns_result_I)[1]
using get_root_node_parent_same by blast
qed
next
case False
then have "child |\<in>| object_ptr_kinds h"
using assms(1) assms(4) step(2)
by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure
get_dom_component_def to_tree_order_ptrs_in_heap)
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) False get_root_node_not_node_same by blast
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) step.prems
by (metis (no_types) False \<open>child |\<in>| object_ptr_kinds h\<close> bind_pure_returns_result_I2
bind_returns_result_E2 get_dom_component_def get_root_node_ok get_root_node_pure returns_result_eq
to_tree_order_node_ptrs)
qed
qed
lemma get_dom_component_to_tree_order_subset:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
shows "set nodes \<subseteq> set c"
using assms
apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
by (meson to_tree_order_subset assms(5) contra_subsetD get_dom_component_ptr)
lemma get_dom_component_to_tree_order:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to"
assumes "ptr \<in> set to"
shows "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
get_dom_component_ok get_dom_component_subset get_dom_component_to_tree_order_subset
is_OK_returns_result_E local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap
select_result_I2 subsetCE)
lemma get_dom_component_root_node_same:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
assumes "x \<in> set c"
shows "h \<turnstile> get_root_node x \<rightarrow>\<^sub>r root_ptr"
proof(insert assms(1) assms(6), induct x rule: heap_wellformed_induct_rev )
case (step child)
then show ?case
proof (cases "is_node_ptr_kind child")
case True
obtain node_ptr where
node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child"
using \<open>is_node_ptr_kind child\<close> node_ptr_casts_commute3 by blast
have "child |\<in>| object_ptr_kinds h"
using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2)
unfolding get_dom_component_def
by (meson bind_returns_result_E2 get_root_node_pure)
with node_ptr have "node_ptr |\<in>| node_ptr_kinds h"
by auto
then obtain parent_opt where
parent: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt"
using get_parent_ok \<open>type_wf h\<close> \<open>known_ptrs h\<close>
by fast
then show ?thesis
proof (induct parent_opt)
case None
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) get_root_node_no_parent node_ptr by blast
then show ?case
using \<open>type_wf h\<close> \<open>known_ptrs h\<close> node_ptr step(2) assms(4) assms(1) assms(5)
by (metis (no_types) \<open>child |\<in>| object_ptr_kinds h\<close> bind_pure_returns_result_I
get_dom_component_def get_dom_component_ptr get_dom_component_subset get_root_node_pure
is_OK_returns_result_E returns_result_eq to_tree_order_ok to_tree_order_same_root)
next
case (Some parent_ptr)
then have "h \<turnstile> get_dom_component parent_ptr \<rightarrow>\<^sub>r c"
using step get_dom_component_parent_inside assms node_ptr
by (meson get_dom_component_subset)
then show ?case
using Some node_ptr
apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
using get_root_node_parent_same
using \<open>h \<turnstile> get_dom_component parent_ptr \<rightarrow>\<^sub>r c\<close> assms(1) assms(2) assms(3)
get_dom_component_ptr step.hyps by blast
qed
next
case False
then have "child |\<in>| object_ptr_kinds h"
using assms(1) assms(4) step(2)
by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure
get_dom_component_def to_tree_order_ptrs_in_heap)
then have "h \<turnstile> get_root_node child \<rightarrow>\<^sub>r child"
using assms(1) False get_root_node_not_node_same by auto
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) step.prems assms(5)
by (metis (no_types, opaque_lifting) bind_returns_result_E2 get_dom_component_def
get_root_node_pure returns_result_eq to_tree_order_same_root)
qed
qed
lemma get_dom_component_no_overlap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c'"
shows "set c \<inter> set c' = {} \<or> c = c'"
proof (rule ccontr, auto)
fix x
assume 1: "c \<noteq> c'" and 2: "x \<in> set c" and 3: "x \<in> set c'"
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
using assms(4) unfolding get_dom_component_def
by (meson bind_is_OK_E is_OK_returns_result_I)
moreover obtain root_ptr' where root_ptr': "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr'"
using assms(5) unfolding get_dom_component_def
by (meson bind_is_OK_E is_OK_returns_result_I)
ultimately have "root_ptr \<noteq> root_ptr'"
using 1 assms
unfolding get_dom_component_def
by (meson bind_returns_result_E3 get_root_node_pure returns_result_eq)
moreover have "h \<turnstile> get_root_node x \<rightarrow>\<^sub>r root_ptr"
using 2 root_ptr get_dom_component_root_node_same assms by blast
moreover have "h \<turnstile> get_root_node x \<rightarrow>\<^sub>r root_ptr'"
using 3 root_ptr' get_dom_component_root_node_same assms by blast
ultimately show False
using select_result_I2 by force
qed
lemma get_dom_component_separates_tree_order:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c'"
assumes "ptr' \<notin> set c"
shows "set to \<inter> set c' = {}"
proof -
have "c \<noteq> c'"
using assms get_dom_component_ptr by blast
then have "set c \<inter> set c' = {}"
using assms get_dom_component_no_overlap by blast
moreover have "set to \<subseteq> set c"
using assms get_dom_component_to_tree_order_subset by blast
ultimately show ?thesis
by blast
qed
lemma get_dom_component_separates_tree_order_general:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> to_tree_order ptr'' \<rightarrow>\<^sub>r to''"
assumes "ptr'' \<in> set c"
assumes "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c'"
assumes "ptr' \<notin> set c"
shows "set to'' \<inter> set c' = {}"
proof -
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
using assms(4)
by (metis bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
then have c: "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using assms(4) unfolding get_dom_component_def
by (simp add: bind_returns_result_E3)
with root_ptr show ?thesis
using assms get_dom_component_separates_tree_order get_dom_component_subset
by meson
qed
end
interpretation i_get_dom_component?: l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name
by(auto simp add: l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def
get_dom_component_def is_strongly_dom_component_safe_def is_weakly_dom_component_safe_def instances)
declare l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_child\_nodes\<close>
locale l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_child_nodes_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "cast child \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast child \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual
local.get_root_node_parent_same by blast
then have "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual
local.get_dom_component_parent_inside local.get_dom_component_subset by blast
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume 1: "ptr' \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root_ptr"
using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual
local.get_root_node_parent_same by blast
then have "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual
local.get_dom_component_parent_inside local.get_dom_component_subset by blast
then show "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \<in> set c"
by (smt (verit) \<open>h \<turnstile> get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r root_ptr\<close> assms(1) assms(2) assms(3)
assms(5) assms(6) disjoint_iff_not_equal is_OK_returns_result_E is_OK_returns_result_I
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_no_overlap local.child_parent_dual local.get_dom_component_ok
local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_root_node_ptr_in_heap
local.l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms)
qed
lemma get_child_nodes_get_dom_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
shows "cast ` set children \<subseteq> set c"
using assms get_child_nodes_is_strongly_dom_component_safe
using local.get_dom_component_ptr by auto
lemma
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} (cast ` set children) h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_child_nodes_pure pure_returns_heap_eq)
then show ?thesis
using assms
- apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
- by (smt (verit) IntI fmember_iff_member_fset get_child_nodes_is_strongly_dom_component_safe
+ apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def Bex_def)[1]
+ by (smt (verit) IntI get_child_nodes_is_strongly_dom_component_safe
is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_dom_component_impl
local.get_dom_component_ok local.get_dom_component_ptr returns_result_select_result)
qed
end
interpretation i_get_dom_component_get_child_nodes?: l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_parent\<close>
locale l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_parent_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_parent ptr' \<rightarrow>\<^sub>r Some parent"
shows "parent \<in> set c \<longleftrightarrow> cast ptr' \<in> set c"
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) is_OK_returns_result_E
l_to_tree_order_wf.to_tree_order_parent local.get_dom_component_parent_inside
local.get_dom_component_subset local.get_dom_component_to_tree_order_subset
local.get_parent_parent_in_heap local.l_to_tree_order_wf_axioms local.to_tree_order_ok
local.to_tree_order_ptr_in_result subsetCE)
lemma get_parent_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some parent"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {cast node_ptr} {parent} h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_parent_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
- by (metis IntI finite_set_in local.get_dom_component_impl local.get_dom_component_ok
+ by (metis IntI local.get_dom_component_impl local.get_dom_component_ok
local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_parent_parent_in_heap
local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap local.parent_child_rel_parent
returns_result_select_result)
qed
end
interpretation i_get_dom_component_get_parent?: l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_root\_node\<close>
locale l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_root_node_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root"
shows "root \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "root \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node root \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
by blast
moreover have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
by (metis (no_types, lifting) calculation assms(1) assms(2) assms(3) assms(5)
is_OK_returns_result_E local.get_root_node_root_in_heap local.to_tree_order_ok
local.to_tree_order_ptr_in_result local.to_tree_order_same_root select_result_I2)
ultimately have "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
apply(auto simp add: get_dom_component_def)[1]
using assms(4) bind_returns_result_E3 local.get_dom_component_def root_ptr by fastforce
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume 1: "ptr' \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node root \<rightarrow>\<^sub>r root_ptr"
by (metis assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E
local.get_root_node_root_in_heap local.to_tree_order_ok local.to_tree_order_ptr_in_result
local.to_tree_order_same_root returns_result_eq)
then have "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
using "1" assms(1) assms(2) assms(3) assms(4) local.get_dom_component_subset by blast
then show "root \<in> set c"
using assms(5) bind_returns_result_E3 local.get_dom_component_def local.to_tree_order_ptr_in_result
by fastforce
qed
lemma get_root_node_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} {root} h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_root_node_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
- by (metis (no_types, lifting) IntI finite_set_in get_root_node_is_strongly_dom_component_safe_step
+ by (metis (no_types, opaque_lifting) IntI get_root_node_is_strongly_dom_component_safe_step
is_OK_returns_result_I local.get_dom_component_impl local.get_dom_component_ok
local.get_dom_component_ptr local.get_root_node_ptr_in_heap returns_result_select_result)
qed
end
interpretation i_get_dom_component_get_root_node?: l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_element\_by\_id\<close>
locale l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_element_by_id_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_element_by_id ptr' idd \<rightarrow>\<^sub>r Some result"
shows "cast result \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast result \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
then have "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c\<close>
by (simp add: get_dom_component_def bind_returns_result_E3)
obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
using \<open>h \<turnstile> get_element_by_id ptr' idd \<rightarrow>\<^sub>r Some result\<close>
apply(simp add: get_element_by_id_def first_in_tree_order_def)
by (meson bind_is_OK_E is_OK_returns_result_I)
then have "cast result \<in> set to'"
using \<open>h \<turnstile> get_element_by_id ptr' idd \<rightarrow>\<^sub>r Some result\<close> get_element_by_id_result_in_tree_order
by blast
have "h \<turnstile> get_root_node (cast result) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using \<open>cast result \<in> set to'\<close> \<open>h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'\<close>
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_dom_component_root_node_same
get_dom_component_subset get_dom_component_to_tree_order
by blast
then have "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c\<close>
using get_dom_component_def by auto
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume "ptr' \<in> set c"
moreover obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
ultimately have "set to' \<subseteq> set c"
using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset
get_dom_component_to_tree_order_subset
by blast
moreover have "cast result \<in> set to'"
using assms(5) get_element_by_id_result_in_tree_order to' by blast
ultimately show "cast result \<in> set c"
by blast
qed
lemma get_element_by_id_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_element_by_id ptr idd \<rightarrow>\<^sub>r Some result"
assumes "h \<turnstile> get_element_by_id ptr idd \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} {cast result} h h'"
proof -
have "h = h'"
using assms(5)
by(auto simp add: preserved_def get_element_by_id_def first_in_tree_order_def
elim!: bind_returns_heap_E2 intro!: map_filter_M_pure bind_pure_I
split: option.splits list.splits)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_element_by_id_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast result \<in> set to"
using assms(4) local.get_element_by_id_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_dom_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_dom_component_impl
local.get_dom_component_ok by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def get_element_by_id_def
first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I
split: option.splits list.splits)[1]
- by (metis (no_types, lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4) finite_set_in
+ by (metis (no_types, opaque_lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4)
get_element_by_id_is_strongly_dom_component_safe_step local.get_dom_component_impl
local.get_dom_component_ptr select_result_I2)
qed
end
interpretation i_get_dom_component_get_element_by_id?: l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_elements\_by\_class\_name\<close>
locale l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_elements_by_class_name_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_elements_by_class_name ptr' idd \<rightarrow>\<^sub>r results"
assumes "result \<in> set results"
shows "cast result \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast result \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
then have "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c\<close>
by (simp add: get_dom_component_def bind_returns_result_E3)
obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
using assms(5)
apply(simp add: get_elements_by_class_name_def first_in_tree_order_def)
by (meson bind_is_OK_E is_OK_returns_result_I)
then have "cast result \<in> set to'"
using assms get_elements_by_class_name_result_in_tree_order by blast
have "h \<turnstile> get_root_node (cast result) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using \<open>cast result \<in> set to'\<close> \<open>h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'\<close>
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_dom_component_root_node_same
get_dom_component_subset get_dom_component_to_tree_order
by blast
then have "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c\<close>
using get_dom_component_def by auto
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume "ptr' \<in> set c"
moreover obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
ultimately have "set to' \<subseteq> set c"
using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset
get_dom_component_to_tree_order_subset
by blast
moreover have "cast result \<in> set to'"
using assms get_elements_by_class_name_result_in_tree_order to' by blast
ultimately show "cast result \<in> set c"
by blast
qed
lemma get_elements_by_class_name_pure [simp]:
"pure (get_elements_by_class_name ptr name) h"
by(auto simp add: get_elements_by_class_name_def intro!: bind_pure_I map_filter_M_pure
split: option.splits)
lemma get_elements_by_class_name_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_elements_by_class_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> get_elements_by_class_name ptr name \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'"
proof -
have "h = h'"
using assms(5)
by (meson get_elements_by_class_name_pure pure_returns_heap_eq)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_elements_by_class_name_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast ` set results \<subseteq> set to"
using assms(4) local.get_elements_by_class_name_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_dom_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_dom_component_impl
local.get_dom_component_ok by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def
get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2
intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
- by (metis (no_types, lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4) finite_set_in
+ by (metis (no_types, opaque_lifting) Int_iff \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4)
get_elements_by_class_name_is_strongly_dom_component_safe_step local.get_dom_component_impl
local.get_dom_component_ptr select_result_I2)
qed
end
interpretation i_get_dom_component_get_elements_by_class_name?:
l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_elements\_by\_tag\_name\<close>
locale l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_elements_by_tag_name_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
assumes "h \<turnstile> get_elements_by_tag_name ptr' idd \<rightarrow>\<^sub>r results"
assumes "result \<in> set results"
shows "cast result \<in> set c \<longleftrightarrow> ptr' \<in> set c"
proof
assume 1: "cast result \<in> set c"
obtain root_ptr where
root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
then have "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c\<close>
by (simp add: get_dom_component_def bind_returns_result_E3)
obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
using assms(5)
apply(simp add: get_elements_by_tag_name_def first_in_tree_order_def)
by (meson bind_is_OK_E is_OK_returns_result_I)
then have "cast result \<in> set to'"
using assms get_elements_by_tag_name_result_in_tree_order by blast
have "h \<turnstile> get_root_node (cast result) \<rightarrow>\<^sub>r root_ptr"
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
by blast
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root_ptr"
using \<open>cast result \<in> set to'\<close> \<open>h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'\<close>
using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr
get_dom_component_root_node_same get_dom_component_subset get_dom_component_to_tree_order
by blast
then have "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c"
using \<open>h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c\<close>
using get_dom_component_def by auto
then show "ptr' \<in> set c"
using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
assume "ptr' \<in> set c"
moreover obtain to' where to': "h \<turnstile> to_tree_order ptr' \<rightarrow>\<^sub>r to'"
by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
ultimately have "set to' \<subseteq> set c"
using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset
get_dom_component_to_tree_order_subset
by blast
moreover have "cast result \<in> set to'"
using assms get_elements_by_tag_name_result_in_tree_order to' by blast
ultimately show "cast result \<in> set c"
by blast
qed
lemma get_elements_by_tag_name_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_elements_by_tag_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> get_elements_by_tag_name ptr name \<rightarrow>\<^sub>h h'"
shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'"
proof -
have "h = h'"
using assms(5)
by (meson get_elements_by_tag_name_pure pure_returns_heap_eq)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_elements_by_tag_name_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast ` set results \<subseteq> set to"
using assms(4) local.get_elements_by_tag_name_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_dom_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_dom_component_impl
local.get_dom_component_ok by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def
get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2
intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
- by (metis (no_types, lifting) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close> finite_set_in
+ by (metis (no_types, opaque_lifting) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close>
get_elements_by_tag_name_is_strongly_dom_component_safe_step local.get_dom_component_impl
local.get_dom_component_ptr select_result_I2)
qed
end
interpretation i_get_dom_component_get_elements_by_tag_name?:
l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>remove\_child\<close>
lemma remove_child_unsafe: "\<not>(\<forall>(h
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) h' ptr child. heap_is_wellformed h \<longrightarrow> type_wf h \<longrightarrow> known_ptrs h \<longrightarrow>
h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h' \<longrightarrow> is_weakly_dom_component_safe {ptr, cast child} {} h h')"
proof -
obtain h document_ptr e1 e2 where h: "Inr ((document_ptr, e1, e2), h) = (Heap (fmempty)
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) \<turnstile> (do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
e2 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
return (document_ptr, e1, e2)
})"
by(code_simp, auto simp add: equal_eq List.member_def)+
then obtain h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
h': "h \<turnstile> remove_child (cast e1) (cast e2) \<rightarrow>\<^sub>h h'" and
"\<not>(is_weakly_dom_component_safe {cast e1, cast e2} {} h h')"
apply(code_simp)
apply(clarify)
by(code_simp, auto simp add: equal_eq List.member_def)+
then show ?thesis
by auto
qed
subsubsection \<open>adopt\_node\<close>
lemma adopt_node_unsafe: "\<not>(\<forall>(h
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) h' document_ptr child. heap_is_wellformed h \<longrightarrow> type_wf h \<longrightarrow> known_ptrs h \<longrightarrow> h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h' \<longrightarrow> is_weakly_dom_component_safe {cast document_ptr, cast child} {} h h')"
proof -
obtain h document_ptr document_ptr2 e1 where h: "Inr ((document_ptr, document_ptr2, e1), h) = (Heap (fmempty)
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) \<turnstile> (do {
document_ptr \<leftarrow> create_document;
document_ptr2 \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
adopt_node document_ptr2 (cast e1);
return (document_ptr, document_ptr2, e1)
})"
by(code_simp, auto simp add: equal_eq List.member_def)+
then obtain h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
h': "h \<turnstile> adopt_node document_ptr (cast e1) \<rightarrow>\<^sub>h h'" and
"\<not>(is_weakly_dom_component_safe {cast document_ptr, cast e1} {} h h')"
apply(code_simp)
apply(clarify)
by(code_simp, auto simp add: equal_eq List.member_def)+
then show ?thesis
by auto
qed
subsubsection \<open>create\_element\<close>
lemma create_element_not_strongly_dom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and document_ptr and new_element_ptr and tag where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {cast document_ptr} {cast new_element_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "create_document"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and document_ptr="?document_ptr"])
by code_simp+
qed
locale l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name +
l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr +
l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs +
l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_tag_name set_tag_name_locs +
l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs +
l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs +
l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr
get_child_nodes get_child_nodes_locs +
l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name
set_tag_name_locs
set_disconnected_nodes set_disconnected_nodes_locs create_element
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_dom_component :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and is_strongly_dom_component_safe :: "(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and is_weakly_dom_component_safe :: "(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_element_by_id :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_elements_by_class_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_elements_by_tag_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
begin
lemma create_element_is_weakly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<noteq> cast |h \<turnstile> create_element document_ptr tag|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain new_element_ptr h2 h3 disc_nodes where
new_element_ptr: "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile>set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: create_element_def
elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
using new_element_ptr h2 h3 disc_nodes h'
apply(auto simp add: create_element_def intro!: bind_returns_result_I
bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "preserved (get_M ptr getter) h h2"
using h2 new_element_ptr
apply(rule new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
using new_element_ptr assms(6) \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close>
by simp
have "preserved (get_M ptr getter) h2 h3"
using set_tag_name_writes h3
apply(rule reads_writes_preserved2)
apply(auto simp add: set_tag_name_locs_impl a_set_tag_name_locs_def all_args_def)[1]
by (metis (no_types, lifting) \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> assms(6)
get_M_Element_preserved8 select_result_I2)
have "document_ptr |\<in>| document_ptr_kinds h"
using create_element_document_in_heap
using assms(4)
by blast
then
have "ptr \<noteq> (cast document_ptr)"
using assms(5) assms(1) assms(2) assms(3) local.get_dom_component_ok local.get_dom_component_ptr
by auto
have "preserved (get_M ptr getter) h3 h'"
using set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved2)
apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1]
by (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
show ?thesis
using \<open>preserved (get_M ptr getter) h h2\<close> \<open>preserved (get_M ptr getter) h2 h3\<close>
\<open>preserved (get_M ptr getter) h3 h'\<close>
by(auto simp add: preserved_def)
qed
lemma create_element_is_weakly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'"
shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'"
proof -
obtain new_element_ptr h2 h3 disc_nodes_h3 where
new_element_ptr: "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: create_element_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_element_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_element_ptr)"
using \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr
by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2
by blast
then
have "known_ptrs h3"
using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast
then
have "known_ptrs h'"
using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []"
using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr]
new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes
by blast
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_child_nodes)
then have children_eq2_h2: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_element_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_tag_name_writes h3]
using set_tag_name_types_preserved
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3: "\<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_element_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close>
using \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def acyclic_heap_def)
also have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting)
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally have "a_acyclic_heap h'"
by (simp add: acyclic_heap_def)
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h2"
apply(auto simp add: a_all_ptrs_in_heap_def)[1]
apply (metis \<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> assms
funion_iff local.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap
local.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h
returns_result_select_result)
by (metis assms disconnected_nodes_eq2_h document_ptr_kinds_eq_h funion_iff
local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h
returns_result_select_result)
then have "a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "a_all_ptrs_in_heap h'"
using assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1) local.heap_is_wellformed_def
by blast
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_element_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
heap_is_wellformed_children_in_heap
by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms fset_mp
fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result)
then have "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_element_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_element_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_element_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_element_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms empty_iff local.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
by (metis \<open>local.a_distinct_lists h\<close> \<open>type_wf h2\<close> disconnected_nodes_eq_h document_ptr_kinds_eq_h
local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
children_eq2_h2 object_ptr_kinds_eq_h2)
then have "a_distinct_lists h'"
proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3
object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3
intro!: distinct_concat_map_I)[1]
fix x
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> 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) \<open>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 \<notin> set disc_nodes_h3\<close>
\<open>a_distinct_lists h3\<close> \<open>type_wf h'\<close> disc_nodes_h3 distinct.simps(2)
distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq
returns_result_select_result)
next
fix x y xa
assume "distinct (concat (map (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 (metis (mono_tags, lifting) Int_Collect Int_iff \<open>type_wf h'\<close> assms(1) assms(2) assms(3)
assms(5) bot_set_def document_ptr_kinds_eq_h3 empty_Collect_eq
l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_disconnected_nodes_ok
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_disc_parent
local.create_element_preserves_wellformedness(1)
local.l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms
local.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms returns_result_select_result)
by (metis (no_types, opaque_lifting) \<open>type_wf h'\<close> assms(1) assms(2) assms(3) assms(5)
disjoint_iff document_ptr_kinds_eq_h3 local.create_element_preserves_wellformedness(1)
local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent
returns_result_select_result)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply -
apply(cases "xb = document_ptr")
apply (metis (no_types, opaque_lifting) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3
\<Longrightarrow> 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 \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
\<open>a_distinct_lists h3\<close> children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h'
select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes)
by (metis "3" "4" "5" "6" \<open>a_distinct_lists h3\<close> \<open>type_wf h3\<close> children_eq2_h3
distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result)
qed
have "a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(auto simp add: a_owner_document_valid_def)[1]
apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1]
apply(auto simp add: object_ptr_kinds_eq_h2)[1]
apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1]
apply(auto simp add: document_ptr_kinds_eq_h2)[1]
apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1]
apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1]
apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric]
disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3)[1]
apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1)
local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
apply(simp add: object_ptr_kinds_eq_h)
- by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M
+ by(metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> 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'
+ disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
node_ptr_kinds_commutes select_result_I2)
have "parent_child_rel h = parent_child_rel h'"
proof -
have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally show ?thesis
by simp
qed
have root: "h \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> local.get_root_node_not_node_same)
then
have root': "h' \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h
local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3)
have "heap_is_wellformed h'" and "known_ptrs h'"
using create_element_preserves_wellformedness assms
by blast+
have "cast result |\<notin>| object_ptr_kinds h"
using \<open>cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> assms(4) returns_result_eq)
obtain to where to: "h \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to"
by (meson \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(1) assms(2) assms(3)
document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok)
then
have "h \<turnstile> a_get_dom_component (cast document_ptr) \<rightarrow>\<^sub>r to"
using root
by(auto simp add: a_get_dom_component_def)
moreover
obtain to' where to': "h' \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to'"
by (meson \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> is_OK_returns_result_E
local.get_root_node_root_in_heap local.to_tree_order_ok root')
then
have "h' \<turnstile> a_get_dom_component (cast document_ptr) \<rightarrow>\<^sub>r to'"
using root'
by(auto simp add: a_get_dom_component_def)
moreover
have "\<And>child. child \<in> set to \<longleftrightarrow> child \<in> set to'"
by (metis \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>parent_child_rel h = parent_child_rel h'\<close>
\<open>type_wf h'\<close> assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to')
ultimately
have "set |h \<turnstile> local.a_get_dom_component (cast document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_dom_component (cast document_ptr)|\<^sub>r"
by(auto simp add: a_get_dom_component_def)
show ?thesis
apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1]
using \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> assms(2) assms(3)
children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result
apply (metis is_OK_returns_result_I)
apply (metis \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> assms(4)
element_ptr_kinds_commutes h2 new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2
node_ptr_kinds_eq_h3 returns_result_eq returns_result_heap_def)
using \<open>cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r result |\<notin>| object_ptr_kinds h\<close> element_ptr_kinds_commutes
node_ptr_kinds_commutes apply blast
using assms(1) assms(2) assms(3) \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'\<close>
apply(rule create_element_is_weakly_dom_component_safe_step)
apply (simp add: local.get_dom_component_impl)
using \<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close>
by auto
qed
end
interpretation i_get_dom_component_create_element?: l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name
set_tag_name_locs create_element
by(auto simp add: l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>create\_character\_data\<close>
lemma create_character_data_not_strongly_dom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder},
'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and document_ptr and create_character_datanew_character_data_ptr and tag where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> create_character_data document_ptr tag \<rightarrow>\<^sub>r create_character_datanew_character_data_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_dom_component_safe {cast document_ptr} {cast create_character_datanew_character_data_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "create_document"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and document_ptr="?document_ptr"])
by code_simp+
qed
locale l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name +
l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr +
l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs +
l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_val set_val_locs +
l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val
set_val_locs set_disconnected_nodes set_disconnected_nodes_locs
create_character_data known_ptrs
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_dom_component :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and is_strongly_dom_component_safe :: "(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and is_weakly_dom_component_safe :: "(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_element_by_id :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_elements_by_class_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_elements_by_tag_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) character_data_ptr) prog"
begin
lemma create_character_data_is_weakly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<noteq> cast |h \<turnstile> create_character_data document_ptr text|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain new_character_data_ptr h2 h3 disc_nodes where
new_character_data_ptr: "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: create_character_data_def elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
using new_character_data_ptr h2 h3 disc_nodes h'
apply(auto simp add: create_character_data_def intro!: bind_returns_result_I
bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "preserved (get_M ptr getter) h h2"
using h2 new_character_data_ptr
apply(rule new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
using new_character_data_ptr assms(6)
\<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
by simp
have "preserved (get_M ptr getter) h2 h3"
using set_val_writes h3
apply(rule reads_writes_preserved2)
apply(auto simp add: set_val_locs_impl a_set_val_locs_def all_args_def)[1]
by (metis (mono_tags) CharacterData_simp11
\<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> assms(4) assms(6)
is_OK_returns_heap_I is_OK_returns_result_E returns_result_eq select_result_I2)
have "document_ptr |\<in>| document_ptr_kinds h"
using create_character_data_document_in_heap
using assms(4)
by blast
then
have "ptr \<noteq> (cast document_ptr)"
using assms(5) assms(1) assms(2) assms(3) local.get_dom_component_ok local.get_dom_component_ptr
by auto
have "preserved (get_M ptr getter) h3 h'"
using set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved2)
apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1]
by (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
show ?thesis
using \<open>preserved (get_M ptr getter) h h2\<close> \<open>preserved (get_M ptr getter) h2 h3\<close>
\<open>preserved (get_M ptr getter) h3 h'\<close>
by(auto simp add: preserved_def)
qed
lemma create_character_data_is_weakly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>h h'"
shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'"
proof -
obtain new_character_data_ptr h2 h3 disc_nodes_h3 where
new_character_data_ptr: "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: create_character_data_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then
have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_character_data_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h:
"\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2
get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_child_nodes)
then have children_eq2_h2:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_character_data_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_val_writes h3]
using set_val_types_preserved
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_character_data_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close> using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
a_all_ptrs_in_heap_def heap_is_wellformed_def
using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast
have "parent_child_rel h = parent_child_rel h'"
proof -
have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally show ?thesis
by simp
qed
have root: "h \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> local.get_root_node_not_node_same)
then
have root': "h' \<turnstile> get_root_node (cast document_ptr) \<rightarrow>\<^sub>r cast document_ptr"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h
local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3)
have "heap_is_wellformed h'" and "known_ptrs h'"
using create_character_data_preserves_wellformedness assms
by blast+
have "cast result |\<notin>| object_ptr_kinds h"
using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> assms(4) returns_result_eq)
obtain to where to: "h \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to"
by (meson \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(1) assms(2) assms(3)
document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok)
then
have "h \<turnstile> a_get_dom_component (cast document_ptr) \<rightarrow>\<^sub>r to"
using root
by(auto simp add: a_get_dom_component_def)
moreover
obtain to' where to': "h' \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to'"
by (meson \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> is_OK_returns_result_E
local.get_root_node_root_in_heap local.to_tree_order_ok root')
then
have "h' \<turnstile> a_get_dom_component (cast document_ptr) \<rightarrow>\<^sub>r to'"
using root'
by(auto simp add: a_get_dom_component_def)
moreover
have "\<And>child. child \<in> set to \<longleftrightarrow> child \<in> set to'"
by (metis \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>parent_child_rel h = parent_child_rel h'\<close>
\<open>type_wf h'\<close> assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to')
ultimately
have "set |h \<turnstile> local.a_get_dom_component (cast document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_dom_component (cast document_ptr)|\<^sub>r"
by(auto simp add: a_get_dom_component_def)
show ?thesis
apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1]
using assms(2) assms(3) children_eq_h local.get_child_nodes_ok
local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr object_ptr_kinds_eq_h2
object_ptr_kinds_eq_h3 returns_result_select_result
apply (metis \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
is_OK_returns_result_I)
apply (metis \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> assms(4)
character_data_ptr_kinds_commutes h2 new_character_data_ptr new_character_data_ptr_in_heap
node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 returns_result_eq)
using \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
\<open>new_character_data_ptr \<notin> set |h \<turnstile> character_data_ptr_kinds_M|\<^sub>r\<close> assms(4) returns_result_eq
apply fastforce
using assms(2) assms(3) children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap
local.known_ptrs_known_ptr object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result
apply (smt (verit, best) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h
\<turnstile> object_ptr_kinds_M|\<^sub>r\<close> \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r
new_character_data_ptr\<close> assms(1) assms(5)
create_character_data_is_weakly_dom_component_safe_step local.a_get_dom_component_def
local.get_dom_component_def select_result_I2)
done
qed
end
interpretation i_get_dom_component_create_character_data?: l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_val set_val_locs
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs
create_character_data
by(auto simp add: l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>create\_document\<close>
lemma create_document_unsafe: "\<not>(\<forall>(h
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) h' new_document_ptr. heap_is_wellformed h \<longrightarrow> type_wf h \<longrightarrow> known_ptrs h \<longrightarrow>
h \<turnstile> create_document \<rightarrow>\<^sub>r new_document_ptr \<longrightarrow> h \<turnstile> create_document \<rightarrow>\<^sub>h h' \<longrightarrow>
is_strongly_dom_component_safe {} {cast new_document_ptr} h h')"
proof -
obtain h document_ptr where h: "Inr (document_ptr, h) = (Heap (fmempty)
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) \<turnstile> (do {
document_ptr \<leftarrow> create_document;
return (document_ptr)
})"
by(code_simp, auto simp add: equal_eq List.member_def)+
then obtain h' new_document_ptr where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
h': "h \<turnstile> create_document \<rightarrow>\<^sub>r new_document_ptr" and
h': "h \<turnstile> create_document \<rightarrow>\<^sub>h h'" and
"\<not>(is_strongly_dom_component_safe {} {cast new_document_ptr} h h')"
by(code_simp, auto simp add: equal_eq List.member_def)+
then show ?thesis
by blast
qed
locale l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name +
l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_dom_component :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and is_strongly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and is_weakly_dom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_ancestors :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_element_by_id ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_elements_by_class_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_elements_by_tag_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and create_document :: "((_) heap, exception, (_) document_ptr) prog"
begin
lemma create_document_is_weakly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>h h'"
assumes "ptr \<noteq> cast |h \<turnstile> create_document|\<^sub>r"
shows "preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
using assms
apply(auto simp add: create_document_def)[1]
by (metis assms(4) assms(5) is_OK_returns_heap_I local.create_document_def new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
select_result_I)
lemma create_document_is_weakly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>h h'"
shows "is_weakly_dom_component_safe {} {cast result} h h'"
proof -
have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast result|}"
using assms(4) assms(5) local.create_document_def new_document_new_ptr by auto
moreover have "result |\<notin>| document_ptr_kinds h"
using assms(4) assms(5) local.create_document_def new_document_ptr_not_in_heap by auto
ultimately show ?thesis
using assms
apply(auto simp add: is_weakly_dom_component_safe_def Let_def local.create_document_def
new_document_ptr_not_in_heap)[1]
by (metis \<open>result |\<notin>| document_ptr_kinds h\<close> document_ptr_kinds_commutes new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t)
qed
end
interpretation i_get_dom_component_create_document?: l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name create_document
by(auto simp add: l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>insert\_before\<close>
lemma insert_before_unsafe: "\<not>(\<forall>(h
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) h' ptr child. heap_is_wellformed h \<longrightarrow> type_wf h \<longrightarrow> known_ptrs h \<longrightarrow>
h \<turnstile> insert_before ptr child None \<rightarrow>\<^sub>h h' \<longrightarrow> is_weakly_dom_component_safe {ptr, cast child} {} h h')"
proof -
obtain h document_ptr e1 e2 where h: "Inr ((document_ptr, e1, e2), h) = (Heap (fmempty)
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) \<turnstile> (do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
e2 \<leftarrow> create_element document_ptr ''div'';
return (document_ptr, e1, e2)
})"
by(code_simp, auto simp add: equal_eq List.member_def)+
then obtain h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
h': "h \<turnstile> insert_before (cast e1) (cast e2) None \<rightarrow>\<^sub>h h'" and
"\<not>(is_weakly_dom_component_safe {cast e1, cast e2} {} h h')"
by(code_simp, auto simp add: equal_eq List.member_def)+
then show ?thesis
by auto
qed
lemma insert_before_unsafe2: "\<not>(\<forall>(h
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) h' ptr child ref. heap_is_wellformed h \<longrightarrow> type_wf h \<longrightarrow> known_ptrs h \<longrightarrow>
h \<turnstile> insert_before ptr child (Some ref) \<rightarrow>\<^sub>h h' \<longrightarrow>
is_weakly_dom_component_safe {ptr, cast child, cast ref} {} h h')"
proof -
obtain h document_ptr e1 e2 e3 where h: "Inr ((document_ptr, e1, e2, e3), h) = (Heap (fmempty)
:: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
) \<turnstile> (do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
e2 \<leftarrow> create_element document_ptr ''div'';
e3 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
return (document_ptr, e1, e2, e3)
})"
by(code_simp, auto simp add: equal_eq List.member_def)+
then obtain h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
h': "h \<turnstile> insert_before (cast e1) (cast e3) (Some (cast e2)) \<rightarrow>\<^sub>h h'" and
"\<not>(is_weakly_dom_component_safe {cast e1, cast e3, cast e2} {} h h')"
apply(code_simp)
apply(clarify)
by(code_simp, auto simp add: equal_eq List.member_def)+
then show ?thesis
by fast
qed
lemma append_child_unsafe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and ptr and child where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'" and
"\<not> is_weakly_dom_component_safe {ptr, cast child} {} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
e2 \<leftarrow> create_element document_ptr ''div'';
return (e1, e2)
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e1 = "fst |?h0 \<turnstile> ?P|\<^sub>r"
let ?e2 = "snd |?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1" and child="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e2"])
by code_simp+
qed
subsubsection \<open>get\_owner\_document\<close>
lemma get_owner_document_unsafe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and ptr and owner_document where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<rightarrow>\<^sub>h h'" and
"\<not>is_weakly_dom_component_safe {ptr} {cast owner_document} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
e1 \<leftarrow> create_element document_ptr ''div'';
return (document_ptr, e1)
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "fst |?h0 \<turnstile> ?P|\<^sub>r"
let ?e1 = "snd |?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1" and owner_document="?document_ptr"])
by code_simp+
qed
end
diff --git a/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy b/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy
--- a/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy
+++ b/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy
@@ -1,3761 +1,3760 @@
(***********************************************************************************
* 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 \<open>Core SC DOM Components II\<close>
theory Core_DOM_SC_DOM_Components
imports
Core_DOM_DOM_Components
begin
declare [[smt_timeout=2400]]
section \<open>Scope Components\<close>
subsection \<open>Definition\<close>
locale l_get_scdom_component\<^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 +
l_get_owner_document_defs get_owner_document +
l_to_tree_order_defs to_tree_order
for get_owner_document :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
begin
definition a_get_scdom_component :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_get_scdom_component ptr = do {
document \<leftarrow> get_owner_document ptr;
disc_nodes \<leftarrow> get_disconnected_nodes document;
tree_order \<leftarrow> to_tree_order (cast document);
disconnected_tree_orders \<leftarrow> map_M (to_tree_order \<circ> cast) disc_nodes;
return (tree_order @ (concat disconnected_tree_orders))
}"
definition a_is_strongly_scdom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"a_is_strongly_scdom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = (
let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
let arg_components =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h). set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r) in
let arg_components' =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h'). set |h' \<turnstile> a_get_scdom_component ptr|\<^sub>r) in
removed_pointers \<subseteq> arg_components \<and>
added_pointers \<subseteq> arg_components' \<and>
S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \<subseteq> arg_components' \<and>
(\<forall>outside_ptr \<in> fset (object_ptr_kinds h) \<inter> fset (object_ptr_kinds h') -
(\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))"
definition a_is_weakly_scdom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"a_is_weakly_scdom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = (
let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
let arg_components =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h). set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r) in
let arg_components' =
(\<Union>ptr \<in> (\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r) \<inter>
fset (object_ptr_kinds h'). set |h' \<turnstile> a_get_scdom_component ptr|\<^sub>r) in
removed_pointers \<subseteq> arg_components \<and>
S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \<subseteq> arg_components' \<union> added_pointers \<and>
(\<forall>outside_ptr \<in> fset (object_ptr_kinds h) \<inter> fset (object_ptr_kinds h') -
(\<Union>ptr \<in> S\<^sub>a\<^sub>r\<^sub>g. set |h \<turnstile> a_get_scdom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))"
end
global_interpretation l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_disconnected_nodes
get_disconnected_nodes_locs to_tree_order
defines get_scdom_component = "l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_scdom_component
get_owner_document get_disconnected_nodes to_tree_order"
and is_strongly_scdom_component_safe = a_is_strongly_scdom_component_safe
and is_weakly_scdom_component_safe = a_is_weakly_scdom_component_safe
.
locale l_get_scdom_component_defs =
fixes get_scdom_component :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
fixes is_strongly_scdom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
fixes is_weakly_scdom_component_safe ::
"(_) object_ptr set \<Rightarrow> (_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
locale l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_scdom_component_defs +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +
assumes get_scdom_component_impl: "get_scdom_component = a_get_scdom_component"
assumes is_strongly_scdom_component_safe_impl:
"is_strongly_scdom_component_safe = a_is_strongly_scdom_component_safe"
assumes is_weakly_scdom_component_safe_impl:
"is_weakly_scdom_component_safe = a_is_weakly_scdom_component_safe"
begin
lemmas get_scdom_component_def = a_get_scdom_component_def[folded get_scdom_component_impl]
lemmas is_strongly_scdom_component_safe_def =
a_is_strongly_scdom_component_safe_def[folded is_strongly_scdom_component_safe_impl]
lemmas is_weakly_scdom_component_safe_def =
a_is_weakly_scdom_component_safe_def[folded is_weakly_scdom_component_safe_impl]
end
interpretation i_get_scdom_component?: l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order
by(auto simp add: l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def get_scdom_component_def
is_strongly_scdom_component_safe_def is_weakly_scdom_component_safe_def instances)
declare l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed +
l_get_owner_document +
l_get_owner_document_wf +
l_get_disconnected_nodes +
l_to_tree_order +
l_known_ptr +
l_known_ptrs +
l_get_owner_document_wf_get_root_node_wf +
assumes known_ptr_impl: "known_ptr = DocumentClass.known_ptr"
begin
lemma known_ptr_node_or_document: "known_ptr ptr \<Longrightarrow> is_node_ptr_kind ptr \<or> is_document_ptr_kind ptr"
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
lemma get_scdom_component_ptr_in_heap2:
assumes "h \<turnstile> ok (get_scdom_component ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms get_root_node_ptr_in_heap
apply(auto simp add: get_scdom_component_def elim!: bind_is_OK_E3 intro!: map_M_pure_I)[1]
by (simp add: is_OK_returns_result_I local.get_owner_document_ptr_in_heap)
lemma get_scdom_component_subset_get_dom_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
shows "set c \<subseteq> set sc"
proof -
obtain document disc_nodes tree_order disconnected_tree_orders where
document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r document"
and disc_nodes: "h \<turnstile> get_disconnected_nodes document \<rightarrow>\<^sub>r disc_nodes"
and tree_order: "h \<turnstile> to_tree_order (cast document) \<rightarrow>\<^sub>r tree_order"
and disconnected_tree_orders: "h \<turnstile> map_M (to_tree_order \<circ> cast) disc_nodes \<rightarrow>\<^sub>r disconnected_tree_orders"
and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
using assms(4)
by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
)
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
and c: "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using assms(5)
by(auto simp add: get_dom_component_def elim!: bind_returns_result_E2[rotated, OF get_root_node_pure, rotated])
show ?thesis
proof (cases "is_document_ptr_kind root_ptr")
case True
then have "cast document = root_ptr"
using get_root_node_document assms(1) assms(2) assms(3) root_ptr document
by (metis document_ptr_casts_commute3 returns_result_eq)
then have "c = tree_order"
using tree_order c
by auto
then show ?thesis
by(simp add: sc)
next
case False
moreover have "root_ptr |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap root_ptr by blast
ultimately have "is_node_ptr_kind root_ptr"
using assms(3) known_ptrs_known_ptr known_ptr_node_or_document
by auto
then obtain root_node_ptr where root_node_ptr: "root_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 root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> get_owner_document root_ptr \<rightarrow>\<^sub>r document"
using get_root_node_same_owner_document
using assms(1) assms(2) assms(3) document root_ptr by blast
then have "root_node_ptr \<in> set disc_nodes"
using assms(1) assms(2) assms(3) disc_nodes in_disconnected_nodes_no_parent root_node_ptr
using local.get_root_node_same_no_parent root_ptr by blast
then have "c \<in> set disconnected_tree_orders"
using c root_node_ptr
using map_M_pure_E[OF disconnected_tree_orders]
by (metis (mono_tags, lifting) comp_apply local.to_tree_order_pure select_result_I2)
then show ?thesis
by(auto simp add: sc)
qed
qed
lemma get_scdom_component_ptrs_same_owner_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
shows "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document"
proof -
obtain document disc_nodes tree_order disconnected_tree_orders where
document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r document"
and disc_nodes: "h \<turnstile> get_disconnected_nodes document \<rightarrow>\<^sub>r disc_nodes"
and tree_order: "h \<turnstile> to_tree_order (cast document) \<rightarrow>\<^sub>r tree_order"
and disconnected_tree_orders: "h \<turnstile> map_M (to_tree_order \<circ> cast) disc_nodes \<rightarrow>\<^sub>r disconnected_tree_orders"
and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
using assms(4)
by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
)
show ?thesis
proof (cases "ptr' \<in> set tree_order")
case True
have "owner_document = document"
using assms(6) document by fastforce
then show ?thesis
by (metis (no_types) True assms(1) assms(2) assms(3) 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_inject document
document_ptr_casts_commute3 document_ptr_document_ptr_cast document_ptr_kinds_commutes
local.get_owner_document_owner_document_in_heap local.get_root_node_document
local.get_root_node_not_node_same local.to_tree_order_same_root node_ptr_no_document_ptr_cast tree_order)
next
case False
then obtain disconnected_tree_order where disconnected_tree_order:
"ptr' \<in> set disconnected_tree_order" and "disconnected_tree_order \<in> set disconnected_tree_orders"
using sc \<open>ptr' \<in> set sc\<close>
by auto
obtain root_ptr' where
root_ptr': "root_ptr' \<in> set disc_nodes" and
"h \<turnstile> to_tree_order (cast root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order"
using map_M_pure_E2[OF disconnected_tree_orders \<open>disconnected_tree_order \<in> set disconnected_tree_orders\<close>]
by (metis comp_apply local.to_tree_order_pure)
have "\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). root_ptr' \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)"
using disc_nodes
by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok
- local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset
+ local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr
returns_result_select_result root_ptr')
then
have "h \<turnstile> get_parent root_ptr' \<rightarrow>\<^sub>r None"
using disc_nodes
- by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember_iff_member_fset local.get_parent_child_dual
+ by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3) local.get_parent_child_dual
local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_disc_nodes_in_heap
returns_result_select_result root_ptr' select_result_I2 split_option_ex)
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast root_ptr'"
using \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order\<close> assms(1)
assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent
local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result
by blast
then have "h \<turnstile> get_owner_document (cast root_ptr') \<rightarrow>\<^sub>r document"
using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr'
by blast
then have "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r document"
using \<open>h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr'\<close> assms(1) assms(2) assms(3)
local.get_root_node_same_owner_document
by blast
then show ?thesis
using assms(6) document returns_result_eq by force
qed
qed
lemma get_scdom_component_ptrs_same_scope_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
shows "h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc"
proof -
obtain document disc_nodes tree_order disconnected_tree_orders where
document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r document"
and disc_nodes: "h \<turnstile> get_disconnected_nodes document \<rightarrow>\<^sub>r disc_nodes"
and tree_order: "h \<turnstile> to_tree_order (cast document) \<rightarrow>\<^sub>r tree_order"
and disconnected_tree_orders: "h \<turnstile> map_M (to_tree_order \<circ> cast) disc_nodes \<rightarrow>\<^sub>r disconnected_tree_orders"
and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
using assms(4)
by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
)
show ?thesis
proof (cases "ptr' \<in> set tree_order")
case True
then have "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r document"
by (metis assms(1) assms(2) assms(3) 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_inject document
document_ptr_casts_commute3 document_ptr_kinds_commutes known_ptr_node_or_document
local.get_owner_document_owner_document_in_heap local.get_root_node_document
local.get_root_node_not_node_same local.known_ptrs_known_ptr local.to_tree_order_get_root_node
local.to_tree_order_ptr_in_result node_ptr_no_document_ptr_cast tree_order)
then show ?thesis
using disc_nodes tree_order disconnected_tree_orders sc
by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
next
case False
then obtain disconnected_tree_order where disconnected_tree_order:
"ptr' \<in> set disconnected_tree_order" and "disconnected_tree_order \<in> set disconnected_tree_orders"
using sc \<open>ptr' \<in> set sc\<close>
by auto
obtain root_ptr' where
root_ptr': "root_ptr' \<in> set disc_nodes" and
"h \<turnstile> to_tree_order (cast root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order"
using map_M_pure_E2[OF disconnected_tree_orders \<open>disconnected_tree_order \<in> set disconnected_tree_orders\<close>]
by (metis comp_apply local.to_tree_order_pure)
have "\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). root_ptr' \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)"
using disc_nodes
by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok
- local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset
+ local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr
returns_result_select_result root_ptr')
then
have "h \<turnstile> get_parent root_ptr' \<rightarrow>\<^sub>r None"
using disc_nodes
- by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember_iff_member_fset
+ by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3)
local.get_parent_child_dual local.get_parent_ok local.get_parent_parent_in_heap
local.heap_is_wellformed_disc_nodes_in_heap returns_result_select_result root_ptr'
select_result_I2 split_option_ex)
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast root_ptr'"
using \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order\<close> assms(1)
assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent
local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result
by blast
then have "h \<turnstile> get_owner_document (cast root_ptr') \<rightarrow>\<^sub>r document"
using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr'
by blast
then have "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r document"
using \<open>h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr'\<close> assms(1) assms(2) assms(3)
local.get_root_node_same_owner_document
by blast
then show ?thesis
using disc_nodes tree_order disconnected_tree_orders sc
by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
qed
qed
lemma get_scdom_component_ok:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_scdom_component ptr)"
using assms
apply(auto simp add: get_scdom_component_def intro!: bind_is_OK_pure_I map_M_pure_I map_M_ok_I)[1]
using get_owner_document_ok
apply blast
apply (simp add: local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap)
apply (simp add: local.get_owner_document_owner_document_in_heap local.to_tree_order_ok)
using local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok node_ptr_kinds_commutes
by blast
lemma get_scdom_component_ptr_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
shows "ptr' |\<in>| object_ptr_kinds h"
apply(insert assms )
apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
using local.to_tree_order_ptrs_in_heap apply blast
by (metis (no_types, lifting) assms(4) assms(5) bind_returns_result_E
get_scdom_component_ptrs_same_scope_component is_OK_returns_result_I get_scdom_component_def
local.get_owner_document_ptr_in_heap)
lemma get_scdom_component_contains_get_dom_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
obtains c where "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c" and "set c \<subseteq> set sc"
proof -
have "h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc"
using assms(1) assms(2) assms(3) assms(4) assms(5) get_scdom_component_ptrs_same_scope_component
by blast
then show ?thesis
by (meson assms(1) assms(2) assms(3) assms(5) get_scdom_component_ptr_in_heap
get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok that)
qed
lemma get_scdom_component_owner_document_same:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
obtains owner_document where "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document" and "cast owner_document \<in> set sc"
using assms
apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
apply (metis (no_types, lifting) assms(4) assms(5) document_ptr_casts_commute3
document_ptr_document_ptr_cast get_scdom_component_contains_get_dom_component
local.get_dom_component_ptr local.get_dom_component_root_node_same local.get_dom_component_to_tree_order
local.get_root_node_document local.get_root_node_not_node_same local.to_tree_order_ptr_in_result
local.to_tree_order_ptrs_in_heap node_ptr_no_document_ptr_cast)
apply(rule map_M_pure_E2)
apply(simp)
apply(simp)
apply(simp)
using assms(4) assms(5) get_scdom_component_ptrs_same_owner_document local.to_tree_order_ptr_in_result
by blast
lemma get_scdom_component_different_owner_documents:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
assumes "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document'"
assumes "owner_document \<noteq> owner_document'"
shows "set |h \<turnstile> get_scdom_component ptr|\<^sub>r \<inter> set |h \<turnstile> get_scdom_component ptr'|\<^sub>r = {}"
using assms get_scdom_component_ptrs_same_owner_document
by (smt (verit) disjoint_iff_not_equal get_scdom_component_ok is_OK_returns_result_I
local.get_owner_document_ptr_in_heap returns_result_eq returns_result_select_result)
lemma get_scdom_component_ptr:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r c"
shows "ptr \<in> set c"
using assms
by (meson get_scdom_component_ptr_in_heap2 get_scdom_component_subset_get_dom_component
is_OK_returns_result_E is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_ptr
subsetD)
end
locale l_get_dom_component_get_scdom_component = l_get_owner_document_defs + l_heap_is_wellformed_defs +
l_type_wf + l_known_ptrs + l_get_scdom_component_defs + l_get_dom_component_defs +
assumes get_scdom_component_subset_get_dom_component:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc \<Longrightarrow>
h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c \<Longrightarrow> set c \<subseteq> set sc"
assumes get_scdom_component_ptrs_same_scope_component:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc \<Longrightarrow>
ptr' \<in> set sc \<Longrightarrow> h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc"
assumes get_scdom_component_ptrs_same_owner_document:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc \<Longrightarrow>
ptr' \<in> set sc \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<Longrightarrow> h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document"
assumes get_scdom_component_ok:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow>
h \<turnstile> ok (get_scdom_component ptr)"
assumes get_scdom_component_ptr_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc \<Longrightarrow>
ptr' \<in> set sc \<Longrightarrow> ptr' |\<in>| object_ptr_kinds h"
assumes get_scdom_component_contains_get_dom_component:
"(\<And>c. h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c \<Longrightarrow> set c \<subseteq> set sc \<Longrightarrow> thesis) \<Longrightarrow> heap_is_wellformed h \<Longrightarrow>
type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc \<Longrightarrow> ptr' \<in> set sc \<Longrightarrow> thesis"
assumes get_scdom_component_owner_document_same:
"(\<And>owner_document. h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document \<Longrightarrow> cast owner_document \<in> set sc \<Longrightarrow> thesis) \<Longrightarrow>
heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc \<Longrightarrow>
ptr' \<in> set sc \<Longrightarrow> thesis"
assumes get_scdom_component_different_owner_documents:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<Longrightarrow>
h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document' \<Longrightarrow> owner_document \<noteq> owner_document' \<Longrightarrow>
set |h \<turnstile> get_scdom_component ptr|\<^sub>r \<inter> set |h \<turnstile> get_scdom_component ptr'|\<^sub>r = {}"
interpretation i_get_dom_component_get_scdom_component?: l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_owner_document
get_disconnected_nodes get_disconnected_nodes_locs to_tree_order heap_is_wellformed parent_child_rel
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors
get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
by(auto simp add: l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)
declare l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_dom_component_get_scdom_component_is_l_get_dom_component_get_scdom_component [instances]:
"l_get_dom_component_get_scdom_component get_owner_document heap_is_wellformed type_wf known_ptr
known_ptrs get_scdom_component get_dom_component"
apply(auto simp add: l_get_dom_component_get_scdom_component_def l_get_dom_component_get_scdom_component_axioms_def instances)[1]
using get_scdom_component_subset_get_dom_component apply fast
using get_scdom_component_ptrs_same_scope_component apply fast
using get_scdom_component_ptrs_same_owner_document apply fast
using get_scdom_component_ok apply fast
using get_scdom_component_ptr_in_heap apply fast
using get_scdom_component_contains_get_dom_component apply blast
using get_scdom_component_owner_document_same apply blast
using get_scdom_component_different_owner_documents apply fast
done
subsubsection \<open>get\_child\_nodes\<close>
locale l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_child_nodes_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "cast child \<in> set sc \<longleftrightarrow> ptr' \<in> set sc"
apply(auto)[1]
apply (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) contra_subsetD
get_scdom_component_ptrs_same_scope_component get_scdom_component_subset_get_dom_component
is_OK_returns_result_E local.get_child_nodes_is_strongly_dom_component_safe local.get_dom_component_ok
local.get_dom_component_ptr local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes)
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
get_scdom_component_contains_get_dom_component is_OK_returns_result_E is_OK_returns_result_I
get_child_nodes_is_strongly_dom_component_safe local.get_child_nodes_ptr_in_heap
local.get_dom_component_ok local.get_dom_component_ptr set_rev_mp)
lemma get_child_nodes_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr} (cast ` set children) h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_child_nodes_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1]
- by (smt (verit, del_insts) IntI finite_set_in
+ by (smt (verit, del_insts) IntI
get_child_nodes_is_strongly_scdom_component_safe_step is_OK_returns_result_I
local.get_child_nodes_ptr_in_heap local.get_dom_component_ok local.get_dom_component_ptr
local.get_scdom_component_impl local.get_scdom_component_ok
local.get_scdom_component_subset_get_dom_component returns_result_select_result subsetD)
qed
end
interpretation i_get_scdom_component_get_child_nodes?: l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes
get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_parent\<close>
locale l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_parent_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_parent ptr' \<rightarrow>\<^sub>r Some parent"
shows "parent \<in> set sc \<longleftrightarrow> cast ptr' \<in> set sc"
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
get_scdom_component_contains_get_dom_component local.get_dom_component_ptr
local.get_parent_is_strongly_dom_component_safe_step)
lemma get_parent_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some parent"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {cast node_ptr} {parent} h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_parent_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1]
by (smt (verit) Int_iff get_parent_is_strongly_scdom_component_safe_step in_mono
- l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_ptr local.get_dom_component_ok
+ get_dom_component_ptr local.get_dom_component_ok
local.get_parent_parent_in_heap local.get_scdom_component_impl local.get_scdom_component_ok
local.get_scdom_component_ptr_in_heap local.get_scdom_component_ptrs_same_scope_component
local.get_scdom_component_subset_get_dom_component
- local.l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms notin_fset returns_result_eq
- returns_result_select_result)
+ returns_result_eq returns_result_select_result)
qed
end
interpretation i_get_scdom_component_get_parent?: l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_root\_node\<close>
locale l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_root_node_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r root"
shows "root \<in> set sc \<longleftrightarrow> ptr' \<in> set sc"
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
get_scdom_component_contains_get_dom_component local.get_dom_component_ptr
local.get_root_node_is_strongly_dom_component_safe_step)
lemma get_root_node_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr} {root} h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_root_node_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1]
- by (smt (verit) Int_iff finite_set_in is_OK_returns_result_I local.get_dom_component_ok
+ by (smt (verit) Int_iff is_OK_returns_result_I local.get_dom_component_ok
local.get_dom_component_ptr local.get_root_node_is_strongly_dom_component_safe_step
local.get_root_node_ptr_in_heap local.get_scdom_component_impl local.get_scdom_component_ok
local.get_scdom_component_subset_get_dom_component returns_result_select_result subset_eq)
qed
end
interpretation i_get_scdom_component_get_root_node?: l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes
get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order
get_attribute get_attribute_locs
by(auto simp add: l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_element\_by\_id\<close>
locale l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_element_by_id_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_element_by_id ptr' idd \<rightarrow>\<^sub>r Some result"
shows "cast result \<in> set sc \<longleftrightarrow> ptr' \<in> set sc"
by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
get_element_by_id_is_strongly_dom_component_safe_step get_scdom_component_contains_get_dom_component
local.get_dom_component_ptr)
lemma get_element_by_id_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_element_by_id ptr idd \<rightarrow>\<^sub>r Some result"
assumes "h \<turnstile> get_element_by_id ptr idd \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr} {cast result} h h'"
proof -
have "h = h'"
using assms(5)
by(auto simp add: preserved_def get_element_by_id_def first_in_tree_order_def
elim!: bind_returns_heap_E2 intro!: map_filter_M_pure bind_pure_I
split: option.splits list.splits)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_element_by_id_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast result \<in> set to"
using assms(4) local.get_element_by_id_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_scdom_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_scdom_component_impl
local.get_scdom_component_ok
by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def
get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2
intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
- by (smt (verit) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4) finite_set_in
+ by (smt (verit) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(4)
get_element_by_id_is_strongly_scdom_component_safe_step local.get_dom_component_ok
local.get_dom_component_ptr local.get_scdom_component_impl
local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2
subsetD)
qed
end
interpretation i_get_scdom_component_get_element_by_id?: l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes
get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order
get_attribute get_attribute_locs
by(auto simp add: l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_elements\_by\_class\_name\<close>
locale l_get_scdom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_elements_by_class_name_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_elements_by_class_name ptr' idd \<rightarrow>\<^sub>r results"
assumes "result \<in> set results"
shows "cast result \<in> set sc \<longleftrightarrow> ptr' \<in> set sc"
by (meson assms local.get_dom_component_ptr
local.get_elements_by_class_name_is_strongly_dom_component_safe_step
local.get_scdom_component_contains_get_dom_component subsetD)
lemma get_elements_by_class_name_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_elements_by_class_name ptr idd \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> get_elements_by_class_name ptr idd \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr} (cast ` set results) h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_elements_by_class_name_pure pure_returns_heap_eq)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_elements_by_class_name_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast ` set results \<subseteq> set to"
using assms(4) local.get_elements_by_class_name_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_scdom_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_scdom_component_impl
local.get_scdom_component_ok by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def
get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure
bind_pure_I split: option.splits list.splits)[1]
- by (smt (verit) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close> finite_set_in
+ by (smt (verit) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close>
get_elements_by_class_name_is_strongly_scdom_component_safe_step local.get_dom_component_ok
local.get_dom_component_ptr local.get_scdom_component_impl
local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2 subsetD)
qed
end
interpretation i_get_scdom_component_get_elements_by_class_name?: l_get_scdom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_scdom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_elements\_by\_tag\_name\<close>
locale l_get_scdom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_elements_by_tag_name_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_elements_by_tag_name ptr' idd \<rightarrow>\<^sub>r results"
assumes "result \<in> set results"
shows "cast result \<in> set sc \<longleftrightarrow> ptr' \<in> set sc"
by (meson assms local.get_dom_component_ptr
local.get_elements_by_tag_name_is_strongly_dom_component_safe_step
local.get_scdom_component_contains_get_dom_component subsetD)
lemma get_elements_by_tag_name_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_elements_by_tag_name ptr idd \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> get_elements_by_tag_name ptr idd \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr} (cast ` set results) h h'"
proof -
have "h = h'"
using assms(5)
by (meson local.get_elements_by_tag_name_pure pure_returns_heap_eq)
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
apply(auto simp add: get_elements_by_tag_name_def)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
obtain to where to: "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.to_tree_order_ok)
then have "cast ` set results \<subseteq> set to"
using assms(4) local.get_elements_by_tag_name_result_in_tree_order by auto
obtain c where c: "h \<turnstile> a_get_scdom_component ptr \<rightarrow>\<^sub>r c"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) local.get_scdom_component_impl
local.get_scdom_component_ok by blast
then show ?thesis
using assms \<open>h = h'\<close>
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def
get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!:
map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
- by (smt (verit) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close> finite_set_in
+ by (smt (verit) IntI \<open>ptr |\<in>| object_ptr_kinds h\<close>
get_elements_by_tag_name_is_strongly_scdom_component_safe_step local.get_dom_component_ok
local.get_dom_component_ptr local.get_scdom_component_impl
local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2
subsetD)
qed
end
interpretation i_get_scdom_component_get_elements_by_tag_name?:
l_get_scdom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe
get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs
get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors
get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
first_in_tree_order get_attribute get_attribute_locs
by(auto simp add: l_get_scdom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>remove\_child\<close>
locale l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^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 +
l_get_owner_document_wf +
l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs
get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf
known_ptr known_ptrs heap_is_wellformed parent_child_rel
begin
lemma remove_child_is_component_unsafe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component ptr'|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r"
(* assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast child)|\<^sub>r" *)
shows "preserved (get_M ptr getter) h h'"
proof -
have "ptr \<noteq> ptr'"
using assms(5)
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I
is_OK_returns_result_E local.get_dom_component_ok local.get_dom_component_ptr
local.remove_child_ptr_in_heap select_result_I2)
obtain owner_document where owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document"
by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_result_E local.get_owner_document_ok
local.remove_child_child_in_heap node_ptr_kinds_commutes)
then
obtain c where "h \<turnstile> get_dom_component (cast owner_document) \<rightarrow>\<^sub>r c"
using get_dom_component_ok owner_document assms(1) assms(2) assms(3)
by (meson document_ptr_kinds_commutes get_owner_document_owner_document_in_heap select_result_I)
then
have "ptr \<noteq> cast owner_document"
using assms(6) assms(1) assms(2) assms(3) local.get_dom_component_ptr owner_document
by auto
show ?thesis
using remove_child_writes assms(4)
apply(rule reads_writes_preserved2)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: option.splits)[1]
apply (metis \<open>ptr \<noteq> ptr'\<close> document_ptr_casts_commute3 get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis \<open>ptr \<noteq> ptr'\<close> element_ptr_casts_commute3 get_M_Element_preserved8)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
done
qed
lemma remove_child_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component ptr'|\<^sub>r"
(* assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r" *)
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast child)|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain sc where sc: "h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc"
using get_scdom_component_ok
by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I local.remove_child_ptr_in_heap
returns_result_select_result)
have "child |\<in>| node_ptr_kinds h"
using assms(4) remove_child_child_in_heap by blast
then
obtain child_sc where child_sc: "h \<turnstile> get_scdom_component (cast child) \<rightarrow>\<^sub>r child_sc"
using get_scdom_component_ok
by (meson assms(1) assms(2) assms(3) node_ptr_kinds_commutes select_result_I)
then obtain owner_document where owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document"
by (meson \<open>child |\<in>| node_ptr_kinds h\<close> assms(1) assms(2) assms(3) contra_subsetD
get_scdom_component_owner_document_same is_OK_returns_result_E
get_scdom_component_subset_get_dom_component local.get_dom_component_ok local.get_dom_component_ptr
node_ptr_kinds_commutes)
then have "h \<turnstile> get_scdom_component (cast owner_document) \<rightarrow>\<^sub>r child_sc"
using child_sc
by (smt (verit) \<open>child |\<in>| node_ptr_kinds h\<close> assms(1) assms(2) assms(3) contra_subsetD
get_scdom_component_subset_get_dom_component get_scdom_component_owner_document_same
get_scdom_component_ptrs_same_scope_component local.get_dom_component_ok local.get_dom_component_ptr
node_ptr_kinds_commutes returns_result_select_result select_result_I2)
have "ptr \<notin> set |h \<turnstile> get_dom_component ptr'|\<^sub>r"
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
get_scdom_component_subset_get_dom_component is_OK_returns_heap_I local.get_dom_component_ok
local.remove_child_ptr_in_heap returns_result_select_result sc select_result_I2)
moreover have "ptr \<notin> set |h \<turnstile> get_scdom_component (cast |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r"
using get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component
by (metis (no_types, lifting)
\<open>h \<turnstile> get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \<rightarrow>\<^sub>r child_sc\<close> assms(6) child_sc
owner_document select_result_I2)
have "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r"
using get_scdom_component_owner_document_same
by (metis (no_types, lifting)
\<open>h \<turnstile> get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \<rightarrow>\<^sub>r child_sc\<close>
\<open>ptr \<notin> set |h \<turnstile> get_scdom_component (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 |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r\<close>
assms(1) assms(2) assms(3) contra_subsetD document_ptr_kinds_commutes get_scdom_component_subset_get_dom_component
is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap owner_document
select_result_I2)
ultimately show ?thesis
using assms(1) assms(2) assms(3) assms(4) remove_child_is_component_unsafe by blast
qed
lemma remove_child_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr, cast child} {} h h'"
proof -
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(4)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1]
using pure_returns_heap_eq by fastforce
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>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)
then have object_ptr_kinds_eq: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq2: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
using select_result_eq by force
then have node_ptr_kinds_eq2: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by auto
then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
using node_ptr_kinds_M_eq by auto
have document_ptr_kinds_eq2: "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
using document_ptr_kinds_M_eq by auto
have children_eq:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(4)])
unfolding remove_child_locs_def
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
then have children_eq2:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq: "\<And>document_ptr disconnected_nodes. document_ptr \<noteq> owner_document
\<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes"
apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(4)])
unfolding remove_child_locs_def
using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers
by (metis (no_types, lifting) Un_iff owner_document select_result_I2)
then have disconnected_nodes_eq2:
"\<And>document_ptr. document_ptr \<noteq> owner_document
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] )
by (simp add: set_disconnected_nodes_get_child_nodes)
have "known_ptrs h'"
using object_ptr_kinds_eq3 known_ptrs_preserved \<open>known_ptrs h\<close> by blast
have "known_ptr ptr"
using assms(3)
using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h2]
using set_disconnected_nodes_types_preserved assms(2)
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_child_nodes_writes h']
using set_child_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r remove1 child children_h"
using assms(4) owner_document h2 disconnected_nodes_h children_h
apply(auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto split: if_splits)[1]
apply(simp)
apply(auto split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E4)
apply(auto)[1]
apply simp
using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h'
by blast
have disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
using owner_document assms(4) h2 disconnected_nodes_h
apply (auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E2)
apply(auto split: if_splits)[1]
apply(simp)
by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits)
then have disconnected_nodes_h': "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h'])
by (simp add: set_child_nodes_get_disconnected_nodes)
moreover have "a_acyclic_heap h"
using assms(1) by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h"
proof (standard, safe)
fix parent child
assume a1: "(parent, child) \<in> parent_child_rel h'"
then show "(parent, child) \<in> parent_child_rel h"
proof (cases "parent = ptr")
case True
then show ?thesis
using a1 remove_child_removes_parent[OF assms(1) assms(4)] children_h children_h'
get_child_nodes_ptr_in_heap
apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1]
by (metis imageI notin_set_remove1)
next
case False
then show ?thesis
using a1
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2)
qed
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h"
using assms(1) by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1]
apply (metis (no_types, lifting) \<open>type_wf h'\<close> assms local.get_child_nodes_ok local.known_ptrs_known_ptr
- local.remove_child_children_subset notin_fset object_ptr_kinds_eq3 returns_result_select_result subset_code(1))
+ local.remove_child_children_subset object_ptr_kinds_eq3 returns_result_select_result subset_code(1))
apply (metis (no_types, lifting) assms(4) disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h'
- document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2
+ document_ptr_kinds_eq3 local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2
set_ConsD subset_code(1))
done
moreover have "a_owner_document_valid h"
using assms(1) by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3
node_ptr_kinds_eq3)[1]
proof -
fix node_ptr
assume 0: "\<forall>node_ptr\<in>fset (node_ptr_kinds h'). (\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or> (\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<and>
node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
and 1: "node_ptr |\<in>| node_ptr_kinds h'"
and 2: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<longrightarrow> node_ptr \<notin> set |h' \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
then show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h'
\<and> node_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
proof (cases "node_ptr = child")
case True
show ?thesis
apply(rule exI[where x=owner_document])
using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True
by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I
list.set_intros(1) select_result_I2)
next
case False
then show ?thesis
using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h
disconnected_nodes_h'
apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1]
- by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2))
+ by (metis children_eq2 disconnected_nodes_eq2 in_set_remove1 list.set_intros(2))
qed
qed
moreover
{
have h0: "a_distinct_lists h"
using assms(1) by (simp add: heap_is_wellformed_def)
moreover have ha1: "(\<Union>x\<in>set |h \<turnstile> object_ptr_kinds_M|\<^sub>r. set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r. set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
using \<open>a_distinct_lists h\<close>
unfolding a_distinct_lists_def
by(auto)
have ha2: "ptr |\<in>| object_ptr_kinds h"
using children_h get_child_nodes_ptr_in_heap by blast
have ha3: "child \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r"
using child_in_children_h children_h
by(simp)
have child_not_in: "\<And>document_ptr. document_ptr |\<in>| document_ptr_kinds h
\<Longrightarrow> child \<notin> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using ha1 ha2 ha3
apply(simp)
using IntI by fastforce
moreover have "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
apply(rule select_result_I)
by(auto simp add: object_ptr_kinds_M_defs)
moreover have "distinct |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
apply(rule select_result_I)
by(auto simp add: document_ptr_kinds_M_defs)
ultimately have "a_distinct_lists h'"
proof(simp (no_asm) add: a_distinct_lists_def, safe)
assume 1: "a_distinct_lists h"
and 3: "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
assume 1: "a_distinct_lists h"
and 3: "distinct |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
have 4: "distinct (concat ((map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r)))"
using 1 by(auto simp add: a_distinct_lists_def)
show "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
proof(rule distinct_concat_map_I[OF 3[unfolded object_ptr_kinds_eq2], simplified])
fix x
assume 5: "x |\<in>| object_ptr_kinds h'"
then have 6: "distinct |h \<turnstile> get_child_nodes x|\<^sub>r"
using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce
obtain children where children: "h \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children"
and distinct_children: "distinct children"
by (metis "5" "6" assms get_child_nodes_ok local.known_ptrs_known_ptr
object_ptr_kinds_eq3 select_result_I)
obtain children' where children': "h' \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children'"
using children children_eq children_h' by fastforce
then have "distinct children'"
proof (cases "ptr = x")
case True
then show ?thesis
using children distinct_children children_h children_h'
by (metis children' distinct_remove1 returns_result_eq)
next
case False
then show ?thesis
using children distinct_children children_eq[OF False]
using children' distinct_lists_children h0
using select_result_I2 by fastforce
qed
then show "distinct |h' \<turnstile> get_child_nodes x|\<^sub>r"
using children' by(auto simp add: )
next
fix x y
assume 5: "x |\<in>| object_ptr_kinds h'" and 6: "y |\<in>| object_ptr_kinds h'" and 7: "x \<noteq> y"
obtain children_x where children_x: "h \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children_x"
by (metis "5" assms get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children_y where children_y: "h \<turnstile> get_child_nodes y \<rightarrow>\<^sub>r children_y"
by (metis "6" assms get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children_x' where children_x': "h' \<turnstile> get_child_nodes x \<rightarrow>\<^sub>r children_x'"
using children_eq children_h' children_x by fastforce
obtain children_y' where children_y': "h' \<turnstile> get_child_nodes y \<rightarrow>\<^sub>r children_y'"
using children_eq children_h' children_y by fastforce
have "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r) |h \<turnstile> object_ptr_kinds_M|\<^sub>r))"
using h0 by(auto simp add: a_distinct_lists_def)
then have 8: "set children_x \<inter> set children_y = {}"
using "7" assms(1) children_x children_y local.heap_is_wellformed_one_parent by blast
have "set children_x' \<inter> set children_y' = {}"
proof (cases "ptr = x")
case True
then have "ptr \<noteq> y"
by(simp add: 7)
have "children_x' = remove1 child children_x"
using children_h children_h' children_x children_x' True returns_result_eq by fastforce
moreover have "children_y' = children_y"
using children_y children_y' children_eq[OF \<open>ptr \<noteq> y\<close>] by auto
ultimately show ?thesis
using 8 set_remove1_subset by fastforce
next
case False
then show ?thesis
proof (cases "ptr = y")
case True
have "children_y' = remove1 child children_y"
using children_h children_h' children_y children_y' True returns_result_eq by fastforce
moreover have "children_x' = children_x"
using children_x children_x' children_eq[OF \<open>ptr \<noteq> x\<close>] by auto
ultimately show ?thesis
using 8 set_remove1_subset by fastforce
next
case False
have "children_x' = children_x"
using children_x children_x' children_eq[OF \<open>ptr \<noteq> x\<close>] by auto
moreover have "children_y' = children_y"
using children_y children_y' children_eq[OF \<open>ptr \<noteq> y\<close>] by auto
ultimately show ?thesis
using 8 by simp
qed
qed
then show "set |h' \<turnstile> get_child_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_child_nodes y|\<^sub>r = {}"
using children_x' children_y'
by (metis (no_types, lifting) select_result_I2)
qed
next
assume 2: "distinct |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
then have 4: "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))"
by simp
have 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))))"
using h0
by(simp add: a_distinct_lists_def document_ptr_kinds_eq3)
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h')))))"
proof(rule distinct_concat_map_I[OF 4[unfolded document_ptr_kinds_eq3]])
fix x
assume 4: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 5: "distinct |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using distinct_lists_disconnected_nodes[OF h0] 4 get_disconnected_nodes_ok
by (simp add: assms document_ptr_kinds_eq3 select_result_I)
show "distinct |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "x = owner_document")
case True
have "child \<notin> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using child_not_in document_ptr_kinds_eq2 "4" by fastforce
moreover have "|h' \<turnstile> get_disconnected_nodes x|\<^sub>r = child # |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using disconnected_nodes_h' disconnected_nodes_h unfolding True
by(simp)
ultimately show ?thesis
using 5 unfolding True
by simp
next
case False
show ?thesis
using "5" False disconnected_nodes_eq2 by auto
qed
next
fix x y
assume 4: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and 5: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))" and "x \<noteq> y"
obtain disc_nodes_x where disc_nodes_x: "h \<turnstile> get_disconnected_nodes x \<rightarrow>\<^sub>r disc_nodes_x"
using 4 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of x] document_ptr_kinds_eq2
by auto
obtain disc_nodes_y where disc_nodes_y: "h \<turnstile> get_disconnected_nodes y \<rightarrow>\<^sub>r disc_nodes_y"
using 5 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of y] document_ptr_kinds_eq2
by auto
obtain disc_nodes_x' where disc_nodes_x': "h' \<turnstile> get_disconnected_nodes x \<rightarrow>\<^sub>r disc_nodes_x'"
using 4 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of x] document_ptr_kinds_eq2
by auto
obtain disc_nodes_y' where disc_nodes_y': "h' \<turnstile> get_disconnected_nodes y \<rightarrow>\<^sub>r disc_nodes_y'"
using 5 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of y] document_ptr_kinds_eq2
by auto
have "distinct
(concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) |h \<turnstile> document_ptr_kinds_M|\<^sub>r))"
using h0 by (simp add: a_distinct_lists_def)
then have 6: "set disc_nodes_x \<inter> set disc_nodes_y = {}"
using \<open>x \<noteq> y\<close> assms(1) disc_nodes_x disc_nodes_y local.heap_is_wellformed_one_disc_parent
by blast
have "set disc_nodes_x' \<inter> set disc_nodes_y' = {}"
proof (cases "x = owner_document")
case True
then have "y \<noteq> owner_document"
using \<open>x \<noteq> y\<close> by simp
then have "disc_nodes_y' = disc_nodes_y"
using disconnected_nodes_eq[OF \<open>y \<noteq> owner_document\<close>] disc_nodes_y disc_nodes_y'
by auto
have "disc_nodes_x' = child # disc_nodes_x"
using disconnected_nodes_h' disc_nodes_x disc_nodes_x' True disconnected_nodes_h returns_result_eq
by fastforce
have "child \<notin> set disc_nodes_y"
using child_not_in disc_nodes_y 5
using document_ptr_kinds_eq2 by fastforce
then show ?thesis
apply(unfold \<open>disc_nodes_x' = child # disc_nodes_x\<close> \<open>disc_nodes_y' = disc_nodes_y\<close>)
using 6 by auto
next
case False
then show ?thesis
proof (cases "y = owner_document")
case True
then have "disc_nodes_x' = disc_nodes_x"
using disconnected_nodes_eq[OF \<open>x \<noteq> owner_document\<close>] disc_nodes_x disc_nodes_x' by auto
have "disc_nodes_y' = child # disc_nodes_y"
using disconnected_nodes_h' disc_nodes_y disc_nodes_y' True disconnected_nodes_h returns_result_eq
by fastforce
have "child \<notin> set disc_nodes_x"
using child_not_in disc_nodes_x 4
using document_ptr_kinds_eq2 by fastforce
then show ?thesis
apply(unfold \<open>disc_nodes_y' = child # disc_nodes_y\<close> \<open>disc_nodes_x' = disc_nodes_x\<close>)
using 6 by auto
next
case False
have "disc_nodes_x' = disc_nodes_x"
using disconnected_nodes_eq[OF \<open>x \<noteq> owner_document\<close>] disc_nodes_x disc_nodes_x' by auto
have "disc_nodes_y' = disc_nodes_y"
using disconnected_nodes_eq[OF \<open>y \<noteq> owner_document\<close>] disc_nodes_y disc_nodes_y' by auto
then show ?thesis
apply(unfold \<open>disc_nodes_y' = disc_nodes_y\<close> \<open>disc_nodes_x' = disc_nodes_x\<close>)
using 6 by auto
qed
qed
then show "set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using disc_nodes_x' disc_nodes_y' by auto
qed
next
fix x xa xb
assume 1: "xa \<in> fset (object_ptr_kinds h')"
and 2: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 3: "xb \<in> fset (document_ptr_kinds h')"
and 4: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
obtain disc_nodes where disc_nodes: "h \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r disc_nodes"
using 3 get_disconnected_nodes_ok[OF \<open>type_wf h\<close>, of xb] document_ptr_kinds_eq2 by auto
obtain disc_nodes' where disc_nodes': "h' \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r disc_nodes'"
using 3 get_disconnected_nodes_ok[OF \<open>type_wf h'\<close>, of xb] document_ptr_kinds_eq2 by auto
obtain children where children: "h \<turnstile> get_child_nodes xa \<rightarrow>\<^sub>r children"
- by (metis "1" assms finite_set_in get_child_nodes_ok is_OK_returns_result_E
+ by (metis "1" assms get_child_nodes_ok is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
obtain children' where children': "h' \<turnstile> get_child_nodes xa \<rightarrow>\<^sub>r children'"
using children children_eq children_h' by fastforce
have "\<And>x. x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r \<Longrightarrow> x \<in> set |h \<turnstile> get_disconnected_nodes xb|\<^sub>r \<Longrightarrow> False"
using 1 3
apply(fold \<open> object_ptr_kinds h = object_ptr_kinds h'\<close>)
apply(fold \<open> document_ptr_kinds h = document_ptr_kinds h'\<close>)
using children disc_nodes h0 apply(auto simp add: a_distinct_lists_def)[1]
by (metis (no_types, lifting) h0 local.distinct_lists_no_parent select_result_I2)
then have 5: "\<And>x. x \<in> set children \<Longrightarrow> x \<in> set disc_nodes \<Longrightarrow> False"
using children disc_nodes by fastforce
have 6: "|h' \<turnstile> get_child_nodes xa|\<^sub>r = children'"
using children' by simp
have 7: "|h' \<turnstile> get_disconnected_nodes xb|\<^sub>r = disc_nodes'"
using disc_nodes' by simp
have "False"
proof (cases "xa = ptr")
case True
have "distinct children_h"
using children_h distinct_lists_children h0 \<open>known_ptr ptr\<close> by blast
have "|h' \<turnstile> get_child_nodes ptr|\<^sub>r = remove1 child children_h"
using children_h'
by simp
have "children = children_h"
using True children children_h by auto
show ?thesis
using disc_nodes' children' 5 2 4 children_h \<open>distinct children_h\<close> disconnected_nodes_h'
apply(auto simp add: 6 7
\<open>xa = ptr\<close> \<open>|h' \<turnstile> get_child_nodes ptr|\<^sub>r = remove1 child children_h\<close> \<open>children = children_h\<close>)[1]
by (metis (no_types, lifting) disc_nodes disconnected_nodes_eq2 disconnected_nodes_h
select_result_I2 set_ConsD)
next
case False
have "children' = children"
using children' children children_eq[OF False[symmetric]]
by auto
then show ?thesis
proof (cases "xb = owner_document")
case True
then show ?thesis
using disc_nodes disconnected_nodes_h disconnected_nodes_h'
using "2" "4" "5" "6" "7" False \<open>children' = children\<close> assms(1) child_in_children_h
child_parent_dual children children_h disc_nodes' get_child_nodes_ptr_in_heap
list.set_cases list.simps(3) option.simps(1) returns_result_eq set_ConsD
by (metis (no_types, opaque_lifting) assms)
next
case False
then show ?thesis
using "2" "4" "5" "6" "7" \<open>children' = children\<close> disc_nodes disc_nodes'
disconnected_nodes_eq returns_result_eq
by metis
qed
qed
then show "x \<in> {}"
by simp
qed
}
ultimately have "heap_is_wellformed h'"
using heap_is_wellformed_def by blast
show ?thesis
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def object_ptr_kinds_eq3)[1]
using assms(1) assms(2) assms(3) assms(4) local.get_scdom_component_impl
remove_child_is_strongly_dom_component_safe_step
by blast
qed
end
interpretation i_get_scdom_component_remove_child?: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors
get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id
get_elements_by_class_name get_elements_by_tag_name set_child_nodes set_child_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove
by(auto simp add: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>adopt\_node\<close>
locale l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^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 +
l_adopt_node_wf +
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document_wf +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma adopt_node_is_component_unsafe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> adopt_node document_ptr node_ptr \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast node_ptr)|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain owner_document where owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr) \<rightarrow>\<^sub>r owner_document"
using assms(4) local.adopt_node_def by auto
then
obtain c where "h \<turnstile> get_dom_component (cast owner_document) \<rightarrow>\<^sub>r c"
using get_dom_component_ok assms(1) assms(2) assms(3) get_owner_document_owner_document_in_heap
by (meson document_ptr_kinds_commutes select_result_I)
then
have "ptr \<noteq> cast owner_document"
using assms(6) assms(1) assms(2) assms(3) local.get_dom_component_ptr owner_document
by (metis (no_types, lifting) assms(7) select_result_I2)
have "document_ptr |\<in>| document_ptr_kinds h"
using adopt_node_document_in_heap assms(1) assms(2) assms(3) assms(4) by auto
then
have "ptr \<noteq> cast document_ptr"
using assms(5)
using assms(1) assms(2) assms(3) local.get_dom_component_ptr get_dom_component_ok
by (meson document_ptr_kinds_commutes returns_result_select_result)
have "\<And>parent. |h \<turnstile> get_parent node_ptr|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr"
by (metis assms(1) assms(2) assms(3) assms(6) is_OK_returns_result_I local.get_dom_component_ok
local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_owner_document_ptr_in_heap
local.get_parent_ok node_ptr_kinds_commutes owner_document returns_result_select_result)
show ?thesis
using adopt_node_writes assms(4)
apply(rule reads_writes_preserved2)
apply(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def)[1]
apply (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply(drule \<open>\<And>parent. |h \<turnstile> get_parent node_ptr|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>)[1] apply (metis element_ptr_casts_commute3 get_M_Element_preserved8 is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply(drule \<open>\<And>parent. |h \<turnstile> get_parent node_ptr|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>)[1] apply (metis document_ptr_casts_commute3 get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply(drule \<open>\<And>parent. |h \<turnstile> get_parent node_ptr|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>)[1]
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\<close> get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> get_M_Mdocument_preserved3 owner_document select_result_I2)
done
qed
lemma adopt_node_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> adopt_node document_ptr node_ptr \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast node_ptr)|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
have "document_ptr |\<in>| document_ptr_kinds h"
by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I local.adopt_node_document_in_heap)
then
obtain sc where sc: "h \<turnstile> get_scdom_component (cast document_ptr) \<rightarrow>\<^sub>r sc"
using get_scdom_component_ok
by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes returns_result_select_result)
have "node_ptr |\<in>| node_ptr_kinds h"
using assms(4)
by (meson is_OK_returns_heap_I local.adopt_node_child_in_heap)
then
obtain child_sc where child_sc: "h \<turnstile> get_scdom_component (cast node_ptr) \<rightarrow>\<^sub>r child_sc"
using get_scdom_component_ok
by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E node_ptr_kinds_commutes)
then obtain owner_document where owner_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document"
by (meson \<open>node_ptr |\<in>| node_ptr_kinds h\<close> assms(1) assms(2) assms(3) contra_subsetD
get_scdom_component_owner_document_same is_OK_returns_result_E
get_scdom_component_subset_get_dom_component local.get_dom_component_ok local.get_dom_component_ptr
node_ptr_kinds_commutes)
then have "h \<turnstile> get_scdom_component (cast owner_document) \<rightarrow>\<^sub>r child_sc"
using child_sc
by (metis (no_types, lifting) \<open>node_ptr |\<in>| node_ptr_kinds h\<close> assms(1) assms(2) assms(3)
get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component
get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok
local.get_dom_component_ptr node_ptr_kinds_commutes select_result_I2 subset_code(1))
have "ptr \<notin> set |h \<turnstile> get_dom_component (cast document_ptr)|\<^sub>r"
by (metis (no_types, lifting) \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(1) assms(2) assms(3)
assms(5) contra_subsetD document_ptr_kinds_commutes get_scdom_component_subset_get_dom_component
local.get_dom_component_ok returns_result_select_result sc select_result_I2)
moreover have "ptr \<notin> set |h \<turnstile> get_dom_component (cast node_ptr)|\<^sub>r"
by (metis (no_types, lifting) \<open>node_ptr |\<in>| node_ptr_kinds h\<close> assms(1) assms(2) assms(3) assms(6)
child_sc contra_subsetD get_scdom_component_subset_get_dom_component local.get_dom_component_ok
node_ptr_kinds_commutes returns_result_select_result select_result_I2)
moreover have "ptr \<notin> set |h \<turnstile> get_scdom_component (cast |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r"
using get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component
by (metis (no_types, lifting)
\<open>h \<turnstile> get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \<rightarrow>\<^sub>r child_sc\<close> assms(6) child_sc
owner_document select_result_I2)
have "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r"
using get_scdom_component_owner_document_same
by (metis (no_types, opaque_lifting)
\<open>\<And>thesis. (\<And>owner_document. h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr) \<rightarrow>\<^sub>r owner_document \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
\<open>h \<turnstile> get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \<rightarrow>\<^sub>r child_sc\<close>
\<open>ptr \<notin> set |h \<turnstile> get_scdom_component (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 |h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r\<close>
assms(1) assms(2) assms(3) contra_subsetD document_ptr_kinds_commutes get_scdom_component_subset_get_dom_component
is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap owner_document
returns_result_eq select_result_I2)
ultimately show ?thesis
using assms(1) assms(2) assms(3) assms(4) adopt_node_is_component_unsafe
by blast
qed
lemma adopt_node_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h"
assumes "h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {cast document_ptr, cast child} {} h h'"
proof -
obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r old_document"
and
parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt"
and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent child | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2"
and
h': "h2 \<turnstile> (if document_ptr \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 child old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes document_ptr;
set_disconnected_nodes document_ptr (child # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: adopt_node_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_parent_pure])
have 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="\<lambda>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:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have wellformed_h2: "heap_is_wellformed h2"
using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "type_wf h2"
using h2 remove_child_preserves_type_wf known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "known_ptrs h2"
using h2 remove_child_preserves_known_ptrs known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "heap_is_wellformed h' \<and> known_ptrs h' \<and> type_wf h'"
proof(cases "document_ptr = old_document")
case True
then show ?thesis
using h' wellformed_h2 \<open>type_wf h2\<close> \<open>known_ptrs h2\<close> by auto
next
case False
then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where
docs_neq: "document_ptr \<noteq> old_document" and
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 child old_disc_nodes) \<rightarrow>\<^sub>h h3" and
disc_nodes_document_ptr_h3:
"h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \<rightarrow>\<^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="\<lambda>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:
"\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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:
"\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r = |h3 \<turnstile> 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="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h3: "|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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:
"\<And>ptr children. h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r = |h' \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. old_document \<noteq> doc_ptr
\<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. old_document \<noteq> doc_ptr
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
using old_disc_nodes by blast
then have disc_nodes_old_document_h3:
"h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h']
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
have "known_ptrs h3"
using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 by blast
then have "known_ptrs h'"
using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2:
"h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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 |\<in>| 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 |\<in>| 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 \<in> set disc_nodes_old_document_h2"
proof (insert parent_opt h2, induct parent_opt)
case None
then have "h = h2"
by(auto)
moreover have "a_owner_document_valid h"
using assms(1) heap_is_wellformed_def by(simp add: heap_is_wellformed_def)
ultimately show ?case
using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)]
in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast
next
case (Some option)
then show ?case
apply(simp split: option.splits)
using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs
by blast
qed
have "child \<notin> set (remove1 child disc_nodes_old_document_h2)"
using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \<open>distinct disc_nodes_old_document_h2\<close>
by auto
have "child \<notin> set disc_nodes_document_ptr_h3"
proof -
have "a_distinct_lists h2"
using heap_is_wellformed_def wellformed_h2 by blast
then have 0: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r))"
by(simp add: a_distinct_lists_def)
show ?thesis
using distinct_concat_map_E(1)[OF 0] \<open>child \<in> set disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h2 disc_nodes_document_ptr_h2
by (meson \<open>type_wf h2\<close> 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 |\<in>| node_ptr_kinds h"
using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]]
node_ptr_kinds_commutes by blast
have "a_acyclic_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h2"
proof
fix x
assume "x \<in> parent_child_rel h'"
then show "x \<in> parent_child_rel h2"
using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3
mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong
unfolding parent_child_rel_def
by(simp)
qed
then have "a_acyclic_heap h'"
using \<open>a_acyclic_heap h2\<close> acyclic_heap_def acyclic_subset by blast
moreover have "a_all_ptrs_in_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h3"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1]
apply (simp add: children_eq2_h2 object_ptr_kinds_h2_eq3 subset_code(1))
by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> \<open>type_wf h2\<close>
disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2
in_set_remove1 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2
returns_result_select_result select_result_I2 wellformed_h2)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1]
apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1))
by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> 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
+ local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
select_result_I2 set_ConsD subset_code(1) wellformed_h2)
moreover have "a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_owner_document_valid h'"
apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2
document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 )
by (smt (verit) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2
disc_nodes_old_document_h2 disc_nodes_old_document_h3
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap
document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1
list.set_intros(1) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2
set_subset_Cons subset_code(1))
have a_distinct_lists_h2: "a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h'"
apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2
children_eq2_h2 children_eq2_h3)[1]
proof -
assume 1: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 3: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> 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 \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 4: "distinct |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "old_document \<noteq> x")
case True
then show ?thesis
proof (cases "document_ptr \<noteq> x")
case True
then show ?thesis
using disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>]
disconnected_nodes_eq2_h3[OF \<open>document_ptr \<noteq> x\<close>] 4
by(auto)
next
case False
then show ?thesis
using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4
\<open>child \<notin> set disc_nodes_document_ptr_h3\<close>
by(auto simp add: disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>] )
qed
next
case False
then show ?thesis
by (metis (no_types, opaque_lifting) \<open>distinct disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h3 disconnected_nodes_eq2_h3
distinct_remove1 docs_neq select_result_I2)
qed
next
fix x y
assume a0: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a1: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a2: "x \<noteq> y"
moreover have 5: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
proof(cases "old_document = x")
case True
have "old_document \<noteq> y"
using \<open>x \<noteq> y\<close> \<open>old_document = x\<close> by simp
have "document_ptr \<noteq> x"
using docs_neq \<open>old_document = x\<close> 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] \<open>old_document = x\<close>
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
\<open>document_ptr \<noteq> x\<close> 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 \<open>old_document = x\<close>
docs_neq \<open>old_document \<noteq> y\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr = x\<close>
apply(simp)
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr \<noteq> x\<close>
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 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False
\<open>type_wf h2\<close> 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 \<noteq> y"
using \<open>x \<noteq> y\<close> by auto
have "set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
using \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document \<noteq> y\<close> \<open>document_ptr = x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3
\<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
by(auto)
next
case False
then show ?thesis
proof(cases "document_ptr = y")
case True
have f1: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set disc_nodes_document_ptr_h3 = {}"
using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>document_ptr \<noteq> x\<close> 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 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = {}"
using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>old_document \<noteq> x\<close>
by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2
- document_ptr_kinds_eq3_h3 finite_fset fmember_iff_member_fset set_sorted_list_of_set)
+ document_ptr_kinds_eq3_h3 finite_fset set_sorted_list_of_set)
ultimately show ?thesis
using 5 select_result_I2[OF disc_nodes_document_ptr_h']
select_result_I2[OF disc_nodes_old_document_h2] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr = y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> 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] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close>
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
by (metis \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
empty_iff inf.idem)
qed
qed
qed
qed
qed
next
fix x xa xb
assume 0: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 2: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h'"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h'"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
then show False
using \<open>child \<in> set disc_nodes_old_document_h2\<close> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a3: "h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r remove1 child disc_nodes_old_document_h2"
assume a4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a5: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f6: "old_document |\<in>| document_ptr_kinds h'"
using a1 \<open>xb |\<in>| document_ptr_kinds h'\<close> by blast
have f7: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a2 by simp
have "x \<in> set disc_nodes_old_document_h2"
using f6 a3 a1 by (metis (no_types) \<open>type_wf h'\<close> \<open>x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r\<close>
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' \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using f7 f6 a5 a4 \<open>xa |\<in>| object_ptr_kinds h'\<close>
by fastforce
then show ?thesis
using \<open>x \<in> set disc_nodes_old_document_h2\<close> a1 a4 f7 by blast
next
assume a1: "xb \<noteq> old_document"
assume a2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3"
assume a3: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a4: "xa |\<in>| object_ptr_kinds h'"
assume a5: "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r child # disc_nodes_document_ptr_h3"
assume a6: "old_document |\<in>| document_ptr_kinds h'"
assume a7: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
assume a8: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a10: "\<And>doc_ptr. old_document \<noteq> doc_ptr
\<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a11: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a12: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f13: "\<And>d. d \<notin> set |h' \<turnstile> document_ptr_kinds_M|\<^sub>r \<or> h2 \<turnstile> ok get_disconnected_nodes d"
using a9 \<open>type_wf h2\<close> get_disconnected_nodes_ok
by simp
then have f14: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a6 a3 by simp
have "x \<notin> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
using a12 a8 a4 \<open>xb |\<in>| document_ptr_kinds h'\<close>
- by (meson UN_I disjoint_iff_not_equal fmember_iff_member_fset)
+ by (meson UN_I disjoint_iff_not_equal)
then have "x = child"
using f13 a11 a10 a7 a5 a2 a1
by (metis (no_types, lifting) select_result_I2 set_ConsD)
then have "child \<notin> set disc_nodes_old_document_h2"
using f14 a12 a8 a6 a4
by (metis \<open>type_wf h'\<close> adopt_node_removes_child assms 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 \<open>child \<in> set disc_nodes_old_document_h2\<close> by fastforce
qed
qed
ultimately show ?thesis
using \<open>type_wf h'\<close> \<open>known_ptrs h'\<close> \<open>a_owner_document_valid h'\<close> heap_is_wellformed_def by blast
qed
then have "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'"
by auto
have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF adopt_node_writes assms(4)])
unfolding adopt_node_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def split: if_splits)
show ?thesis
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def object_ptr_kinds_eq3 )[1]
using adopt_node_is_strongly_dom_component_safe_step get_scdom_component_impl assms by blast
qed
end
interpretation i_get_scdom_component_adopt_node?: l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe 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 get_child_nodes get_child_nodes_locs
set_child_nodes set_child_nodes_locs remove to_tree_order get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
by(auto simp add: l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>create\_element\<close>
locale l_get_scdom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma create_element_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<noteq> cast |h \<turnstile> create_element document_ptr tag|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain new_element_ptr h2 h3 disc_nodes where
new_element_ptr: "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile>set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes) \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated,
OF get_disconnected_nodes_pure, rotated])
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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 "heap_is_wellformed h'"
using assms(4)
using assms(1) assms(2) assms(3) local.create_element_preserves_wellformedness(1) by blast
have "type_wf h'"
using assms(1) assms(2) assms(3) assms(4) local.create_element_preserves_wellformedness(2) by blast
have "known_ptrs h'"
using assms(1) assms(2) assms(3) assms(4) local.create_element_preserves_wellformedness(3) by blast
have "document_ptr |\<in>| document_ptr_kinds h"
by (meson assms(4) is_OK_returns_heap_I local.create_element_document_in_heap)
then
obtain sc where sc: "h \<turnstile> get_scdom_component (cast document_ptr) \<rightarrow>\<^sub>r sc"
using get_scdom_component_ok
by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes returns_result_select_result)
have "document_ptr |\<in>| document_ptr_kinds h'"
using \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h
using document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3 by blast
then
obtain sc' where sc': "h' \<turnstile> get_scdom_component (cast document_ptr) \<rightarrow>\<^sub>r sc'"
using get_scdom_component_ok
by (meson \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> document_ptr_kinds_commutes
returns_result_select_result)
obtain c where c: "h \<turnstile> get_dom_component (cast document_ptr) \<rightarrow>\<^sub>r c"
by (meson \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(1) assms(2) assms(3)
document_ptr_kinds_commutes is_OK_returns_result_E local.get_dom_component_ok)
have "set c \<subseteq> set sc"
using assms(1) assms(2) assms(3) c get_scdom_component_subset_get_dom_component sc by blast
have "ptr \<notin> set c"
using \<open>set c \<subseteq> set sc\<close> assms(5) sc
by auto
then
show ?thesis
using create_element_is_weakly_dom_component_safe
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(6) c
local.create_element_is_weakly_dom_component_safe_step select_result_I2)
qed
lemma create_element_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {cast document_ptr} {cast result} h h'"
proof -
obtain new_element_ptr h2 h3 disc_nodes_h3 where
new_element_ptr: "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: create_element_def returns_result_heap_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
then have "result = new_element_ptr"
using assms(4) by auto
have "new_element_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_element_ptr)"
using \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2
by blast
then
have "known_ptrs h3"
using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast
then
have "known_ptrs h'"
using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>ptr' children. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []"
using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr]
new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes
by blast
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_child_nodes)
then have children_eq2_h2: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_tag_name_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_tag_name_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_element_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_tag_name_writes h3]
using set_tag_name_types_preserved
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3: "\<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_element_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close>
using \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
a_all_ptrs_in_heap_def heap_is_wellformed_def
using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast
have "parent_child_rel h = parent_child_rel h'"
proof -
have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting)
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally show ?thesis
by simp
qed
have "document_ptr |\<in>| document_ptr_kinds h'"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h
document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3)
have "known_ptr (cast document_ptr)"
using \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(3) document_ptr_kinds_commutes
local.known_ptrs_known_ptr by blast
have "h \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr"
using \<open>known_ptr (cast document_ptr)\<close> \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, rule conjI)+
by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits)
have "h' \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr"
using \<open>known_ptr (cast document_ptr)\<close> \<open>document_ptr |\<in>| document_ptr_kinds h'\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, rule conjI)+
by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits)
obtain to where to: "h \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to"
by (meson \<open>h \<turnstile> get_owner_document (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r document_ptr\<close> assms(1)
assms(2) assms(3) is_OK_returns_result_E is_OK_returns_result_I local.get_owner_document_ptr_in_heap
local.to_tree_order_ok)
obtain to' where to': "h' \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to'"
by (metis \<open>document_ptr |\<in>| document_ptr_kinds h\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> assms(1) assms(2)
assms(3) assms(5) document_ptr_kinds_commutes document_ptr_kinds_eq_h document_ptr_kinds_eq_h2
document_ptr_kinds_eq_h3 is_OK_returns_result_E local.create_element_preserves_wellformedness(1)
local.to_tree_order_ok)
have "set to = set to'"
proof safe
fix x
assume "x \<in> set to"
show "x \<in> set to'"
using to to'
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to\<close> assms(1) assms(2) assms(3) assms(5)
local.create_element_preserves_wellformedness(1))
next
fix x
assume "x \<in> set to'"
show "x \<in> set to"
using to to'
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to'\<close> assms(1) assms(2) assms(3) assms(5)
local.create_element_preserves_wellformedness(1))
qed
have "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_element_ptr # disc_nodes_h3"
using h' local.set_disconnected_nodes_get_disconnected_nodes by auto
obtain disc_nodes_h' where disc_nodes_h': "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h'"
and "cast new_element_ptr \<in> set disc_nodes_h'"
and "disc_nodes_h' = cast new_element_ptr # disc_nodes_h3"
by (simp add: \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\<close>)
have "\<And>disc_ptr to to'. disc_ptr \<in> set disc_nodes_h3 \<Longrightarrow> h \<turnstile> to_tree_order (cast disc_ptr) \<rightarrow>\<^sub>r to \<Longrightarrow>
h' \<turnstile> to_tree_order (cast disc_ptr) \<rightarrow>\<^sub>r to' \<Longrightarrow> set to = set to'"
proof safe
fix disc_ptr to to' x
assume "disc_ptr \<in> set disc_nodes_h3"
assume "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to"
assume "h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'"
assume "x \<in> set to"
show "x \<in> set to'"
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to\<close>
\<open>h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to\<close>
assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1))
next
fix disc_ptr to to' x
assume "disc_ptr \<in> set disc_nodes_h3"
assume "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to"
assume "h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'"
assume "x \<in> set to'"
show "x \<in> set to"
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to\<close>
\<open>h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to'\<close>
assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1))
qed
have "heap_is_wellformed h'"
using assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1)
by blast
have "cast new_element_ptr |\<in>| object_ptr_kinds h'"
using \<open>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 \<in> set disc_nodes_h'\<close> \<open>heap_is_wellformed h'\<close> disc_nodes_h'
local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes by blast
then
have "new_element_ptr |\<in>| element_ptr_kinds h'"
by simp
have "\<And>node_ptr. node_ptr \<in> set disc_nodes_h3 \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h'"
by (meson \<open>heap_is_wellformed h'\<close> h' local.heap_is_wellformed_disc_nodes_in_heap
local.set_disconnected_nodes_get_disconnected_nodes set_subset_Cons subset_code(1))
have "h \<turnstile> ok (map_M (to_tree_order \<circ> 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) disc_nodes_h3)"
using assms(1) assms(2) assms(3) to_tree_order_ok
apply(auto intro!: map_M_ok_I)[1]
using disc_nodes_document_ptr_h local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes
by blast
then
obtain disc_tree_orders where disc_tree_orders:
"h \<turnstile> map_M (to_tree_order \<circ> 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) disc_nodes_h3 \<rightarrow>\<^sub>r disc_tree_orders"
by auto
have "h' \<turnstile> ok (map_M (to_tree_order \<circ> 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) disc_nodes_h')"
apply(auto intro!: map_M_ok_I)[1]
apply(simp add: \<open>disc_nodes_h' = cast new_element_ptr # disc_nodes_h3\<close>)
using \<open>\<And>node_ptr. node_ptr \<in> set disc_nodes_h3 \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h'\<close>
\<open>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 \<in> set disc_nodes_h'\<close> \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close>
\<open>type_wf h'\<close> disc_nodes_h' local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok
node_ptr_kinds_commutes by blast
then
obtain disc_tree_orders' where disc_tree_orders':
"h' \<turnstile> map_M (to_tree_order \<circ> 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) disc_nodes_h' \<rightarrow>\<^sub>r disc_tree_orders'"
by auto
have "h' \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []"
using \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> children_eq_h2
children_eq_h3 by auto
obtain new_tree_order where new_tree_order:
"h' \<turnstile> to_tree_order (cast new_element_ptr) \<rightarrow>\<^sub>r new_tree_order" and
"new_tree_order \<in> set disc_tree_orders'"
using map_M_pure_E[OF disc_tree_orders' \<open>cast new_element_ptr \<in> set disc_nodes_h'\<close>]
by auto
then have "new_tree_order = [cast new_element_ptr]"
using \<open>h' \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto simp add: to_tree_order_def
dest!: bind_returns_result_E3[rotated, OF \<open>h' \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>, rotated])
obtain foo where foo: "h' \<turnstile> map_M (to_tree_order \<circ> 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)
(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>r [cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr] # foo"
apply(auto intro!: bind_pure_returns_result_I map_M_pure_I)[1]
using \<open>new_tree_order = [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]\<close> new_tree_order apply auto[1]
by (smt (verit) \<open>disc_nodes_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 # disc_nodes_h3\<close>
bind_pure_returns_result_I bind_returns_result_E2 comp_apply disc_tree_orders'
local.to_tree_order_pure map_M.simps(2) map_M_pure_I return_returns_result returns_result_eq)
then have "set (concat foo) = set (concat disc_tree_orders)"
apply(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
apply (smt (verit) \<open>\<And>to' toa disc_ptr. \<lbrakk>disc_ptr \<in> set disc_nodes_h3; h \<turnstile> to_tree_order
(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r toa; h' \<turnstile> to_tree_order
(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<rbrakk> \<Longrightarrow> set toa = set to'\<close>
comp_eq_dest_lhs disc_tree_orders local.to_tree_order_pure map_M_pure_E map_M_pure_E2)
by (smt (verit) \<open>\<And>to' toa disc_ptr. \<lbrakk>disc_ptr \<in> set disc_nodes_h3; h \<turnstile> to_tree_order
(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r toa; h' \<turnstile> to_tree_order
(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<rbrakk> \<Longrightarrow> set toa = set to'\<close>
comp_eq_dest_lhs disc_tree_orders local.to_tree_order_pure map_M_pure_E map_M_pure_E2)
have "disc_tree_orders' = [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] # foo"
using foo disc_tree_orders'
by (simp add: \<open>disc_nodes_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 # disc_nodes_h3\<close> returns_result_eq)
have "set (concat disc_tree_orders') = {cast new_element_ptr} \<union> set (concat disc_tree_orders)"
apply(auto simp add: \<open>disc_tree_orders' = [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] # foo\<close>)[1]
using \<open>set (concat foo) = set (concat disc_tree_orders)\<close> by auto
have "h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r to' @ concat disc_tree_orders'"
using \<open>h' \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr\<close> disc_nodes_h' to' disc_tree_orders'
by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
then
have "set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to' \<union> set (concat disc_tree_orders')"
by auto
have "h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r to @ concat disc_tree_orders"
using \<open>h \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr\<close> disc_nodes_document_ptr_h
to disc_tree_orders
by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
then
have "set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \<union> set (concat disc_tree_orders)"
by auto
have "{cast new_element_ptr} \<union> set |h \<turnstile> local.a_get_scdom_component (cast document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_scdom_component (cast document_ptr)|\<^sub>r"
proof(safe)
show "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
\<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
using \<open>h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r to' @ concat disc_tree_orders'\<close>
apply(auto simp add: a_get_scdom_component_def)[1]
by (meson \<open>\<And>thesis. (\<And>new_tree_order. \<lbrakk>h' \<turnstile> to_tree_order (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) \<rightarrow>\<^sub>r new_tree_order;
new_tree_order \<in> set disc_tree_orders'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> local.to_tree_order_ptr_in_result)
next
fix x
assume " x \<in> set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
then
show "x \<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
using \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union>set (concat disc_tree_orders)\<close>
using \<open>set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to' \<union> set (concat disc_tree_orders')\<close>
using \<open>set to = set to'\<close>
using \<open>set (concat disc_tree_orders') = {cast new_element_ptr} \<union> set (concat disc_tree_orders)\<close>
by(auto)
next
fix x
assume " x \<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
assume "x \<notin> set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
show "x = 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"
using \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union> set (concat disc_tree_orders)\<close>
using \<open>set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to' \<union> set (concat disc_tree_orders')\<close>
using \<open>set to = set to'\<close>
using \<open>set (concat disc_tree_orders') = {cast new_element_ptr} \<union> set (concat disc_tree_orders)\<close>
using \<open>x \<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close>
\<open>x \<notin> set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close>
by auto
qed
have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using object_ptr_kinds_eq_h object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 by auto
then
show ?thesis
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def)[1]
apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"])
using \<open>{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} \<union>
set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close>
apply auto[2]
using \<open>set to = set to'\<close> \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union> set (concat disc_tree_orders)\<close> local.to_tree_order_ptr_in_result to'
apply auto[1]
using \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply blast
apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"])
using \<open>result = new_element_ptr\<close>
\<open>{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} \<union> set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close> apply auto[1]
apply(auto)[1]
using \<open>set to = set to'\<close> \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union> set (concat disc_tree_orders)\<close> local.to_tree_order_ptr_in_result to' apply auto[1]
apply (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close>)
using \<open>\<And>thesis. (\<And>new_element_ptr h2 h3 disc_nodes_h3. \<lbrakk>h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr;
h \<turnstile> new_element \<rightarrow>\<^sub>h h2; h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3;
h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3;
h3 \<turnstile> set_disconnected_nodes document_ptr (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
new_element_ptr new_element_ptr_not_in_heap
apply auto[1]
using create_element_is_strongly_scdom_component_safe_step
by (smt (verit, best) ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile>
object_ptr_kinds_M|\<^sub>r\<close> \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> assms(1)
assms(2) assms(3) assms(5) local.get_scdom_component_impl select_result_I2)
qed
end
interpretation i_get_scdom_component_remove_child?: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component
is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_child_nodes set_child_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove
by(auto simp add: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>create\_character\_data\<close>
locale l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma create_character_data_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast document_ptr)|\<^sub>r"
assumes "ptr \<noteq> cast |h \<turnstile> create_character_data document_ptr text|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
have "document_ptr |\<in>| document_ptr_kinds h"
by (meson assms(4) is_OK_returns_heap_I local.create_character_data_document_in_heap)
then
obtain sc where sc: "h \<turnstile> get_scdom_component (cast document_ptr) \<rightarrow>\<^sub>r sc"
using get_scdom_component_ok
by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes returns_result_select_result)
obtain c where c: "h \<turnstile> get_dom_component (cast document_ptr) \<rightarrow>\<^sub>r c"
by (meson \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(1) assms(2) assms(3)
document_ptr_kinds_commutes is_OK_returns_result_E local.get_dom_component_ok)
have "set c \<subseteq> set sc"
using assms(1) assms(2) assms(3) c get_scdom_component_subset_get_dom_component sc by blast
have "ptr \<notin> set c"
using \<open>set c \<subseteq> set sc\<close> assms(5) sc
by auto
then
show ?thesis
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(6) c
local.create_character_data_is_weakly_dom_component_safe_step select_result_I2)
qed
lemma create_character_data_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {cast document_ptr} {cast result} h h'"
proof -
obtain new_character_data_ptr h2 h3 disc_nodes_h3 where
new_character_data_ptr: "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: create_character_data_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
then have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1]
apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
pure_returns_heap_eq)
by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
then have "result = new_character_data_ptr"
using assms(4) by auto
have "new_character_data_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>ptr' children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h:
"\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_disconnected_nodes_writes h'])
using set_disconnected_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "document_ptr |\<in>| document_ptr_kinds h"
using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
get_disconnected_nodes_ptr_in_heap \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>ptr' children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have children_eq2_h: "\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2
get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_child_nodes)
then have children_eq2_h2:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_val_writes h3
apply(rule reads_writes_preserved)
by(auto simp add: set_val_get_disconnected_nodes)
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_character_data_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_val_writes h3]
using set_val_types_preserved
by(auto simp add: reflp_def transp_def)
then have "type_wf h'"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h3:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then have disconnected_nodes_eq2_h3: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_character_data_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close> using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
a_all_ptrs_in_heap_def heap_is_wellformed_def
using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast
have "parent_child_rel h = parent_child_rel h'"
proof -
have "parent_child_rel h = parent_child_rel h2"
proof(auto simp add: parent_child_rel_def)[1]
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = parent_child_rel h'"
by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
finally show ?thesis
by simp
qed
have "known_ptr (cast new_character_data_ptr)"
using \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
create_character_data_known_ptr by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2
by blast
then
have "known_ptrs h3"
using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast
then
have "known_ptrs h'"
using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast
have "document_ptr |\<in>| document_ptr_kinds h'"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h
document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3)
have "known_ptr (cast document_ptr)"
using \<open>document_ptr |\<in>| document_ptr_kinds h\<close> assms(3) document_ptr_kinds_commutes
local.known_ptrs_known_ptr by blast
have "h \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr"
using \<open>known_ptr (cast document_ptr)\<close> \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, rule conjI)+
by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_pure_returns_result_I split: option.splits)
have "h' \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr"
using \<open>known_ptr (cast document_ptr)\<close> \<open>document_ptr |\<in>| document_ptr_kinds h'\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, rule conjI)+
by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_pure_returns_result_I split: option.splits)
obtain to where to: "h \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to"
by (meson \<open>h \<turnstile> get_owner_document (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r document_ptr\<close>
assms(1) assms(2) assms(3) is_OK_returns_result_E is_OK_returns_result_I
local.get_owner_document_ptr_in_heap local.to_tree_order_ok)
obtain to' where to': "h' \<turnstile> to_tree_order (cast document_ptr) \<rightarrow>\<^sub>r to'"
by (metis \<open>document_ptr |\<in>| document_ptr_kinds h\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> assms(1) assms(2)
assms(3) assms(5) document_ptr_kinds_commutes document_ptr_kinds_eq_h document_ptr_kinds_eq_h2
document_ptr_kinds_eq_h3 is_OK_returns_result_E local.create_character_data_preserves_wellformedness(1)
local.to_tree_order_ok)
have "set to = set to'"
proof safe
fix x
assume "x \<in> set to"
show "x \<in> set to'"
using to to'
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to\<close> assms(1) assms(2) assms(3) assms(5)
local.create_character_data_preserves_wellformedness(1))
next
fix x
assume "x \<in> set to'"
show "x \<in> set to"
using to to'
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to'\<close> assms(1) assms(2) assms(3) assms(5)
local.create_character_data_preserves_wellformedness(1))
qed
have "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_character_data_ptr # disc_nodes_h3"
using h' local.set_disconnected_nodes_get_disconnected_nodes by auto
obtain disc_nodes_h' where disc_nodes_h': "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h'"
and "cast new_character_data_ptr \<in> set disc_nodes_h'"
and "disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3"
by (simp add: \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_character_data_ptr # disc_nodes_h3\<close>)
have "\<And>disc_ptr to to'. disc_ptr \<in> set disc_nodes_h3 \<Longrightarrow> h \<turnstile> to_tree_order (cast disc_ptr) \<rightarrow>\<^sub>r to \<Longrightarrow>
h' \<turnstile> to_tree_order (cast disc_ptr) \<rightarrow>\<^sub>r to' \<Longrightarrow> set to = set to'"
proof safe
fix disc_ptr to to' x
assume "disc_ptr \<in> set disc_nodes_h3"
assume "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to"
assume "h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'"
assume "x \<in> set to"
show "x \<in> set to'"
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to\<close>
\<open>h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to\<close>
assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1))
next
fix disc_ptr to to' x
assume "disc_ptr \<in> set disc_nodes_h3"
assume "h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to"
assume "h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'"
assume "x \<in> set to'"
show "x \<in> set to"
using to_tree_order_parent_child_rel \<open>parent_child_rel h = parent_child_rel h'\<close>
by (metis \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to\<close>
\<open>h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> \<open>x \<in> set to'\<close>
assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1))
qed
have "heap_is_wellformed h'"
using assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1)
by blast
have "cast new_character_data_ptr |\<in>| object_ptr_kinds h'"
using \<open>cast new_character_data_ptr \<in> set disc_nodes_h'\<close> \<open>heap_is_wellformed h'\<close> disc_nodes_h'
local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes by blast
then
have "new_character_data_ptr |\<in>| character_data_ptr_kinds h'"
by simp
have "\<And>node_ptr. node_ptr \<in> set disc_nodes_h3 \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h'"
by (meson \<open>heap_is_wellformed h'\<close> h' local.heap_is_wellformed_disc_nodes_in_heap
local.set_disconnected_nodes_get_disconnected_nodes set_subset_Cons subset_code(1))
have "h \<turnstile> ok (map_M (to_tree_order \<circ> 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) disc_nodes_h3)"
using assms(1) assms(2) assms(3) to_tree_order_ok
apply(auto intro!: map_M_ok_I)[1]
using disc_nodes_document_ptr_h local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes
by blast
then
obtain disc_tree_orders where disc_tree_orders:
"h \<turnstile> map_M (to_tree_order \<circ> 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) disc_nodes_h3 \<rightarrow>\<^sub>r disc_tree_orders"
by auto
have "h' \<turnstile> ok (map_M (to_tree_order \<circ> 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) disc_nodes_h')"
apply(auto intro!: map_M_ok_I)[1]
apply(simp add: \<open>disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\<close>)
using \<open>\<And>node_ptr. node_ptr \<in> set disc_nodes_h3 \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h'\<close>
\<open>cast new_character_data_ptr \<in> set disc_nodes_h'\<close> \<open>heap_is_wellformed h'\<close> \<open>known_ptrs h'\<close>
\<open>type_wf h'\<close> disc_nodes_h' local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok
node_ptr_kinds_commutes by blast
then
obtain disc_tree_orders' where disc_tree_orders':
"h' \<turnstile> map_M (to_tree_order \<circ> 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) disc_nodes_h' \<rightarrow>\<^sub>r disc_tree_orders'"
by auto
have "h' \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []"
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close> children_eq_h2 children_eq_h3 by auto
obtain new_tree_order where new_tree_order:
"h' \<turnstile> to_tree_order (cast new_character_data_ptr) \<rightarrow>\<^sub>r new_tree_order" and
"new_tree_order \<in> set disc_tree_orders'"
using map_M_pure_E[OF disc_tree_orders' \<open>cast new_character_data_ptr \<in> set disc_nodes_h'\<close>]
by auto
then have "new_tree_order = [cast new_character_data_ptr]"
using \<open>h' \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto simp add: to_tree_order_def
dest!: bind_returns_result_E3[rotated, OF \<open>h' \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>, rotated])
obtain foo where foo: "h' \<turnstile> map_M (to_tree_order \<circ> 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)
(cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^sub>r [cast new_character_data_ptr] # foo"
apply(auto intro!: bind_pure_returns_result_I map_M_pure_I)[1]
using \<open>new_tree_order = [cast new_character_data_ptr]\<close> new_tree_order apply auto[1]
using \<open>disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\<close> bind_pure_returns_result_I
bind_returns_result_E2 comp_apply disc_tree_orders' local.to_tree_order_pure map_M.simps(2)
map_M_pure_I return_returns_result returns_result_eq
apply simp
by (smt (verit) \<open>disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\<close> bind_pure_returns_result_I
bind_returns_result_E2 comp_apply disc_tree_orders' local.to_tree_order_pure map_M.simps(2) map_M_pure_I
return_returns_result returns_result_eq)
then have "set (concat foo) = set (concat disc_tree_orders)"
apply(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
apply (smt (verit) \<open>\<And>to' toa disc_ptr. \<lbrakk>disc_ptr \<in> set disc_nodes_h3;
h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r toa; h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<rbrakk> \<Longrightarrow>
set toa = set to'\<close> comp_apply disc_tree_orders local.to_tree_order_pure map_M_pure_E map_M_pure_E2)
by (smt (verit) \<open>\<And>to' toa disc_ptr. \<lbrakk>disc_ptr \<in> set disc_nodes_h3; h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r toa;
h' \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \<rightarrow>\<^sub>r to'\<rbrakk> \<Longrightarrow> set toa = set to'\<close> comp_apply disc_tree_orders
local.to_tree_order_pure map_M_pure_E map_M_pure_E2)
have "disc_tree_orders' = [cast new_character_data_ptr] # foo"
using foo disc_tree_orders'
by (simp add: \<open>disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\<close> returns_result_eq)
have "set (concat disc_tree_orders') = {cast new_character_data_ptr} \<union> set (concat disc_tree_orders)"
apply(auto simp add: \<open>disc_tree_orders' = [cast new_character_data_ptr] # foo\<close>)[1]
using \<open>set (concat foo) = set (concat disc_tree_orders)\<close> by auto
have "h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r to' @ concat disc_tree_orders'"
using \<open>h' \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr\<close> disc_nodes_h' to' disc_tree_orders'
by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
then
have "set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to' \<union> set (concat disc_tree_orders')"
by auto
have "h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r to @ concat disc_tree_orders"
using \<open>h \<turnstile> get_owner_document (cast document_ptr) \<rightarrow>\<^sub>r document_ptr\<close> disc_nodes_document_ptr_h to disc_tree_orders
by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
then
have "set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \<union> set (concat disc_tree_orders)"
by auto
have "{cast new_character_data_ptr} \<union> set |h \<turnstile> local.a_get_scdom_component (cast document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_scdom_component (cast document_ptr)|\<^sub>r"
proof(safe)
show "cast new_character_data_ptr
\<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
using \<open>h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \<rightarrow>\<^sub>r to' @ concat disc_tree_orders'\<close>
apply(auto simp add: a_get_scdom_component_def)[1]
by (meson \<open>\<And>thesis. (\<And>new_tree_order. \<lbrakk>h' \<turnstile> to_tree_order (cast new_character_data_ptr) \<rightarrow>\<^sub>r new_tree_order;
new_tree_order \<in> set disc_tree_orders'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> local.to_tree_order_ptr_in_result)
next
fix x
assume " x \<in> set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
then
show "x \<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
using \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union> set (concat disc_tree_orders)\<close>
using \<open>set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to' \<union> set (concat disc_tree_orders')\<close>
using \<open>set to = set to'\<close>
using \<open>set (concat disc_tree_orders') = {cast new_character_data_ptr} \<union> set (concat disc_tree_orders)\<close>
by(auto)
next
fix x
assume " x \<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
assume "x \<notin> set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r"
show "x = cast new_character_data_ptr"
using \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union> set (concat disc_tree_orders)\<close>
using \<open>set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to' \<union> set (concat disc_tree_orders')\<close>
using \<open>set to = set to'\<close>
using \<open>set (concat disc_tree_orders') = {cast new_character_data_ptr} \<union> set (concat disc_tree_orders)\<close>
using \<open>x \<in> set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close>
\<open>x \<notin> set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close>
by auto
qed
have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using object_ptr_kinds_eq_h object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 by auto
then
show ?thesis
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def)[1]
apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"])
using \<open>{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} \<union> set |h \<turnstile> local.a_get_scdom_component
(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close>
apply auto[2]
using \<open>set to = set to'\<close> \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union> set (concat disc_tree_orders)\<close> local.to_tree_order_ptr_in_result to'
apply auto[1]
using \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply blast
apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"])
using \<open>result = new_character_data_ptr\<close> \<open>{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} \<union>
set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set |h' \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\<close>
apply auto[1]
apply(auto)[1]
using \<open>set to = set to'\<close> \<open>set |h \<turnstile> local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r =
set to \<union> set (concat disc_tree_orders)\<close> local.to_tree_order_ptr_in_result to' apply auto[1]
apply (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close>)
using \<open>\<And>thesis. (\<And>new_character_data_ptr h2 h3 disc_nodes_h3. \<lbrakk>h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr;
h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2; h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3;
h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3;
h3 \<turnstile> set_disconnected_nodes document_ptr (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 # disc_nodes_h3) \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
new_character_data_ptr new_character_data_ptr_not_in_heap
apply auto[1]
using create_character_data_is_strongly_dom_component_safe_step
by (smt (verit) ObjectMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>result = new_character_data_ptr\<close> assms(1) assms(2) assms(3) assms(4) assms(5) local.get_scdom_component_impl select_result_I2)
qed
end
interpretation i_get_scdom_component_create_character_data?: l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component
is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs create_character_data
by(auto simp add: l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>create\_document\<close>
lemma create_document_not_strongly_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
h' and new_document_ptr where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> create_document \<rightarrow>\<^sub>r new_document_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_scdom_component_safe {} {cast new_document_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
let ?P = "create_document"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?document_ptr = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1"])
by code_simp+
qed
locale l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma create_document_is_weakly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> create_document \<rightarrow>\<^sub>h h'"
shows "is_weakly_scdom_component_safe {} {cast result} h h'"
proof -
have "object_ptr_kinds h' = {|cast result|} |\<union>| object_ptr_kinds h"
using assms(4) assms(5) local.create_document_def new_document_new_ptr by blast
have "result |\<notin>| document_ptr_kinds h"
using assms(4) assms(5) local.create_document_def new_document_ptr_not_in_heap by auto
show ?thesis
using assms
apply(auto simp add: is_weakly_scdom_component_safe_def Let_def)[1]
using \<open>object_ptr_kinds h' = {|cast result|} |\<union>| object_ptr_kinds h\<close> apply(auto)[1]
apply (simp add: local.create_document_def new_document_ptr_in_heap)
using \<open>result |\<notin>| document_ptr_kinds h\<close> apply auto[1]
apply (metis (no_types, lifting) \<open>result |\<notin>| document_ptr_kinds h\<close> document_ptr_kinds_commutes
local.create_document_is_weakly_dom_component_safe_step select_result_I2)
done
qed
end
interpretation i_get_scdom_component_create_document?: l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe to_tree_order get_parent get_parent_locs
get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_element_by_id get_elements_by_class_name get_elements_by_tag_name create_document
get_disconnected_nodes get_disconnected_nodes_locs
by(auto simp add: l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>insert\_before\<close>
locale l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^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 +
l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document_wf +
l_get_dom_component_get_scdom_component +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_child_nodes_get_disconnected_nodes +
l_remove_child +
l_get_root_node_wf +
l_set_disconnected_nodes_get_disconnected_nodes_wf +
l_set_disconnected_nodes_get_ancestors +
l_get_ancestors_wf +
l_get_owner_document +
l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma insert_before_is_component_unsafe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> insert_before ptr' child ref \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component ptr'|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast child)|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document ptr'|\<^sub>r)|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document (cast child)|\<^sub>r)|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
obtain owner_document where owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document"
using assms(4)
by(auto simp add: local.adopt_node_def insert_before_def elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF ensure_pre_insertion_validity_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] split: if_splits)
then
obtain c where "h \<turnstile> get_dom_component (cast owner_document) \<rightarrow>\<^sub>r c"
using get_dom_component_ok assms(1) assms(2) assms(3) get_owner_document_owner_document_in_heap
by (meson document_ptr_kinds_commutes select_result_I)
then
have "ptr \<noteq> cast owner_document"
using assms(6) assms(1) assms(2) assms(3) local.get_dom_component_ptr owner_document
by (metis (no_types, lifting) assms(8) select_result_I2)
obtain owner_document' where owner_document': "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document'"
using assms(4)
by(auto simp add: local.adopt_node_def insert_before_def elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF ensure_pre_insertion_validity_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] split: if_splits)
then
obtain c where "h \<turnstile> get_dom_component (cast owner_document') \<rightarrow>\<^sub>r c"
using get_dom_component_ok assms(1) assms(2) assms(3) get_owner_document_owner_document_in_heap
by (meson document_ptr_kinds_commutes select_result_I)
then
have "ptr \<noteq> cast owner_document'"
using assms(1) assms(2) assms(3) assms(7) local.get_dom_component_ptr owner_document' by auto
then
have "ptr \<noteq> cast |h \<turnstile> get_owner_document ptr'|\<^sub>r"
using owner_document' by auto
have "ptr \<noteq> ptr'"
by (metis (mono_tags, opaque_lifting) assms(1) assms(2) assms(3) assms(5) assms(7) is_OK_returns_result_I
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_ok l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_ptr
l_get_owner_document.get_owner_document_ptr_in_heap local.l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms
local.l_get_owner_document_axioms owner_document' return_returns_result returns_result_select_result)
have "\<And>parent. h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent \<Longrightarrow> parent \<noteq> ptr"
by (meson assms(1) assms(2) assms(3) assms(6) l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_ptr
local.get_dom_component_ok local.get_dom_component_to_tree_order local.get_parent_parent_in_heap
local.l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.to_tree_order_ok local.to_tree_order_parent
local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap returns_result_select_result)
then
have "\<And>parent. |h \<turnstile> get_parent child|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr"
by (metis assms(2) assms(3) assms(4) is_OK_returns_heap_I local.get_parent_ok
local.insert_before_child_in_heap select_result_I)
show ?thesis
using insert_before_writes assms(4)
apply(rule reads_writes_preserved2)
apply(auto simp add: insert_before_locs_def adopt_node_locs_def all_args_def)[1]
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
apply (metis \<open>ptr \<noteq> 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 |h \<turnstile> get_owner_document ptr'|\<^sub>r\<close>
get_M_Mdocument_preserved3)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
apply (metis \<open>ptr \<noteq> ptr'\<close> document_ptr_casts_commute3 get_M_Mdocument_preserved3)
apply(auto split: option.splits)[1]
apply (metis \<open>ptr \<noteq> ptr'\<close> element_ptr_casts_commute3 get_M_Element_preserved8)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def
all_args_def split: if_splits)[1]
apply (metis \<open>ptr \<noteq> 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 |h \<turnstile> get_owner_document ptr'|\<^sub>r\<close> get_M_Mdocument_preserved3)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
apply (metis (no_types, lifting) \<open>\<And>parent. |h \<turnstile> get_parent child|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>
element_ptr_casts_commute3 get_M_Element_preserved8 node_ptr_casts_commute option.case_eq_if option.collapse)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis \<open>\<And>parent. |h \<turnstile> get_parent child|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>
document_ptr_casts_commute3 get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
apply (metis \<open>ptr \<noteq> 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 |h \<turnstile> get_owner_document ptr'|\<^sub>r\<close>
get_M_Mdocument_preserved3)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
apply (metis (no_types, lifting) \<open>ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close>
get_M_Mdocument_preserved3 owner_document select_result_I2)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
apply (metis \<open>ptr \<noteq> ptr'\<close> document_ptr_casts_commute3 get_M_Mdocument_preserved3)
apply (metis (no_types, lifting) \<open>ptr \<noteq> ptr'\<close> element_ptr_casts_commute3
get_M_Element_preserved8 node_ptr_casts_commute option.case_eq_if option.collapse)
apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
set_disconnected_nodes_locs_def all_args_def split: if_splits)[1]
by (metis \<open>ptr \<noteq> 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 |h \<turnstile> get_owner_document ptr'|\<^sub>r\<close> get_M_Mdocument_preserved3)
qed
lemma insert_before_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> insert_before ptr' child ref \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component ptr'|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast child)|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
proof -
have "ptr' |\<in>| object_ptr_kinds h"
by (meson assms(4) is_OK_returns_heap_I local.insert_before_ptr_in_heap)
then
obtain sc' where sc': "h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc'"
by (meson assms(1) assms(2) assms(3) get_scdom_component_ok is_OK_returns_result_E)
moreover
obtain c' where c': "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c'"
by (meson \<open>ptr' |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.get_dom_component_ok)
ultimately have "set c' \<subseteq> set sc'"
using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast
have "child |\<in>| node_ptr_kinds h"
by (meson assms(4) is_OK_returns_heap_I local.insert_before_child_in_heap)
then
obtain child_sc where child_sc: "h \<turnstile> get_scdom_component (cast child) \<rightarrow>\<^sub>r child_sc"
by (meson assms(1) assms(2) assms(3) get_scdom_component_ok is_OK_returns_result_E
node_ptr_kinds_commutes)
moreover
obtain child_c where child_c: "h \<turnstile> get_dom_component (cast child) \<rightarrow>\<^sub>r child_c"
by (meson \<open>child |\<in>| node_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.get_dom_component_ok node_ptr_kinds_commutes)
ultimately have "set child_c \<subseteq> set child_sc"
using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast
obtain ptr'_owner_document where ptr'_owner_document: "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r ptr'_owner_document"
by (meson \<open>set c' \<subseteq> set sc'\<close> assms(1) assms(2) assms(3) c' get_scdom_component_owner_document_same
local.get_dom_component_ptr sc' subset_code(1))
then
have "h \<turnstile> get_scdom_component (cast ptr'_owner_document) \<rightarrow>\<^sub>r sc'"
by (metis (no_types, lifting) \<open>set c' \<subseteq> set sc'\<close> assms(1) assms(2) assms(3) c'
get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component
local.get_dom_component_ptr sc' select_result_I2 subset_code(1))
moreover
obtain ptr'_owner_document_c where ptr'_owner_document_c:
"h \<turnstile> get_dom_component (cast ptr'_owner_document) \<rightarrow>\<^sub>r ptr'_owner_document_c"
by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E
local.get_dom_component_ok local.get_owner_document_owner_document_in_heap ptr'_owner_document)
ultimately have "set ptr'_owner_document_c \<subseteq> set sc'"
using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast
obtain child_owner_document where child_owner_document: "h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r child_owner_document"
by (meson \<open>set child_c \<subseteq> set child_sc\<close> assms(1) assms(2) assms(3) child_c child_sc
get_scdom_component_owner_document_same local.get_dom_component_ptr subset_code(1))
have "child_owner_document |\<in>| document_ptr_kinds h"
using assms(1) assms(2) assms(3) child_owner_document local.get_owner_document_owner_document_in_heap
by blast
then
have "h \<turnstile> get_scdom_component (cast child_owner_document) \<rightarrow>\<^sub>r child_sc"
using get_scdom_component_ok assms(1) assms(2) assms(3) child_sc
by (metis (no_types, lifting) \<open>set child_c \<subseteq> set child_sc\<close> child_c child_owner_document
get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component
local.get_dom_component_ptr returns_result_eq set_mp)
moreover
obtain child_owner_document_c where child_owner_document_c:
"h \<turnstile> get_dom_component (cast child_owner_document) \<rightarrow>\<^sub>r child_owner_document_c"
by (meson assms(1) assms(2) assms(3) child_owner_document document_ptr_kinds_commutes
is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap)
ultimately have "set child_owner_document_c \<subseteq> set child_sc"
using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast
have "ptr \<notin> set |h \<turnstile> get_dom_component ptr'|\<^sub>r"
using \<open>set c' \<subseteq> set sc'\<close> assms(5) c' sc' by auto
moreover have "ptr \<notin> set |h \<turnstile> get_dom_component (cast child)|\<^sub>r"
using \<open>set child_c \<subseteq> set child_sc\<close> assms(6) child_c child_sc by auto
moreover have "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document ptr'|\<^sub>r)|\<^sub>r"
using \<open>set ptr'_owner_document_c \<subseteq> set sc'\<close> assms(5) ptr'_owner_document ptr'_owner_document_c sc'
by auto
moreover have "ptr \<notin> set |h \<turnstile> get_dom_component (cast |h \<turnstile> get_owner_document (cast child)|\<^sub>r)|\<^sub>r"
using \<open>set child_owner_document_c \<subseteq> set child_sc\<close> assms(6) child_owner_document child_owner_document_c
child_sc by auto
ultimately show ?thesis
using assms insert_before_is_component_unsafe
by blast
qed
lemma insert_before_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> insert_before ptr node child \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe ({ptr, cast node} \<union> (case child of Some ref \<Rightarrow> {cast ref} | None \<Rightarrow> {} )) {} h h'"
proof -
obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where
ancestors: "h \<turnstile> get_ancestors ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
reference_child:
"h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^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 object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs )
then have object_ptr_kinds_M_eq2_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have "known_ptrs h2"
using assms(3) object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast
have wellformed_h2: "heap_is_wellformed h2"
using adopt_node_preserves_wellformedness[OF assms(1) h2] assms(3) assms(2) .
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h2: "|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto
have "known_ptrs h3"
using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \<open>known_ptrs h2\<close> by blast
have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h3:
"|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h3: "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto
have "object_ptr_kinds h = object_ptr_kinds h'"
by (simp add: object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2)
then
show ?thesis
using assms
apply(auto simp add: is_strongly_scdom_component_safe_def)[1]
using insert_before_is_strongly_dom_component_safe_step local.get_scdom_component_impl by blast
qed
lemma append_child_is_strongly_dom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> append_child ptr' child \<rightarrow>\<^sub>h h'"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component ptr'|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast child)|\<^sub>r"
shows "preserved (get_M ptr getter) h h'"
by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
insert_before_is_strongly_dom_component_safe_step local.append_child_def)
lemma append_child_is_strongly_dom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> append_child ptr child \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr, cast child} {} h h'"
using assms unfolding append_child_def
using insert_before_is_strongly_dom_component_safe
by fastforce
end
interpretation i_get_dom_component_insert_before?: l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name set_child_nodes set_child_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs
get_owner_document remove_child remove_child_locs remove adopt_node adopt_node_locs insert_before
insert_before_locs append_child get_scdom_component is_strongly_scdom_component_safe
is_weakly_scdom_component_safe
by(auto simp add: l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_owner\_document\<close>
locale l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_scdom_component +
l_get_owner_document_wf_get_root_node_wf
begin
lemma get_owner_document_is_strongly_scdom_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document"
shows "cast owner_document \<in> set sc \<longleftrightarrow> ptr' \<in> set sc"
proof -
have "h \<turnstile> get_owner_document (cast owner_document) \<rightarrow>\<^sub>r owner_document"
by (metis assms(1) assms(2) assms(3) assms(5) 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_inject
document_ptr_casts_commute3 document_ptr_document_ptr_cast document_ptr_kinds_commutes
local.get_owner_document_owner_document_in_heap local.get_root_node_document
local.get_root_node_not_node_same node_ptr_no_document_ptr_cast)
then show ?thesis
using assms
using bind_returns_result_E contra_subsetD get_scdom_component_ok
get_scdom_component_ptrs_same_scope_component get_scdom_component_subset_get_dom_component
is_OK_returns_result_E is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_ptr
local.get_owner_document_ptr_in_heap local.get_owner_document_pure local.get_scdom_component_def
pure_returns_heap_eq returns_result_eq
by (smt (verit) local.get_scdom_component_ptrs_same_owner_document subsetD)
qed
lemma get_owner_document_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>h h'"
shows "is_strongly_scdom_component_safe {ptr} {cast owner_document} h h'"
proof -
have "h = h'"
by (meson assms(5) local.get_owner_document_pure pure_returns_heap_eq)
then show ?thesis
using assms
apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1]
by (smt (verit) get_owner_document_is_strongly_scdom_component_safe_step inf.orderE is_OK_returns_result_I
local.get_dom_component_ok local.get_dom_component_to_tree_order_subset local.get_owner_document_ptr_in_heap
local.get_scdom_component_impl local.get_scdom_component_ok local.get_scdom_component_ptr_in_heap
local.get_scdom_component_subset_get_dom_component local.to_tree_order_ok
- local.to_tree_order_ptr_in_result notin_fset returns_result_select_result subset_eq)
+ local.to_tree_order_ptr_in_result returns_result_select_result subset_eq)
qed
end
interpretation i_get_owner_document_scope_component?: l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order known_ptr
known_ptrs type_wf heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs
get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_element_by_id
get_elements_by_class_name get_elements_by_tag_name
by(auto simp add: l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
end
diff --git a/thys/SC_DOM_Components/Shadow_DOM_SC_DOM_Components.thy b/thys/SC_DOM_Components/Shadow_DOM_SC_DOM_Components.thy
--- a/thys/SC_DOM_Components/Shadow_DOM_SC_DOM_Components.thy
+++ b/thys/SC_DOM_Components/Shadow_DOM_SC_DOM_Components.thy
@@ -1,1005 +1,1005 @@
(***********************************************************************************
* 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 \<open>Shadow SC DOM Components II\<close>
theory Shadow_DOM_SC_DOM_Components
imports
Core_DOM_SC_DOM_Components
Shadow_DOM_DOM_Components
begin
section \<open>Shadow root scope components\<close>
subsection \<open>get\_scope\_component\<close>
global_interpretation l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_disconnected_nodes
get_disconnected_nodes_locs to_tree_order
defines get_scdom_component = a_get_scdom_component
and is_strongly_scdom_component_safe = a_is_strongly_scdom_component_safe
and is_weakly_scdom_component_safe = a_is_weakly_scdom_component_safe
.
interpretation i_get_scdom_component?: l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order
by(auto simp add: l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def get_scdom_component_def
is_strongly_scdom_component_safe_def is_weakly_scdom_component_safe_def instances)
declare l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_component\<close>
locale l_get_dom_component_get_scdom_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed +
l_get_owner_document +
l_get_owner_document_wf +
l_get_disconnected_nodes +
l_to_tree_order +
l_known_ptr +
l_known_ptrs +
l_get_owner_document_wf_get_root_node_wf +
assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr"
begin
lemma known_ptr_node_or_document: "known_ptr ptr \<Longrightarrow> is_node_ptr_kind ptr \<or> is_document_ptr_kind ptr"
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
lemma get_scdom_component_subset_get_dom_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_dom_component ptr \<rightarrow>\<^sub>r c"
shows "set c \<subseteq> set sc"
proof -
obtain document disc_nodes tree_order disconnected_tree_orders where document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r document"
and disc_nodes: "h \<turnstile> get_disconnected_nodes document \<rightarrow>\<^sub>r disc_nodes"
and tree_order: "h \<turnstile> to_tree_order (cast document) \<rightarrow>\<^sub>r tree_order"
and disconnected_tree_orders: "h \<turnstile> map_M (to_tree_order \<circ> cast) disc_nodes \<rightarrow>\<^sub>r disconnected_tree_orders"
and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
using assms(4)
by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
)
obtain root_ptr where root_ptr: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root_ptr"
and c: "h \<turnstile> to_tree_order root_ptr \<rightarrow>\<^sub>r c"
using assms(5)
by(auto simp add: get_dom_component_def elim!: bind_returns_result_E2[rotated, OF get_root_node_pure, rotated])
show ?thesis
proof (cases "is_document_ptr_kind root_ptr")
case True
then have "cast document = root_ptr"
using get_root_node_document assms(1) assms(2) assms(3) root_ptr document
by (metis document_ptr_casts_commute3 returns_result_eq)
then have "c = tree_order"
using tree_order c
by auto
then show ?thesis
by(simp add: sc)
next
case False
moreover have "root_ptr |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap root_ptr by blast
ultimately have "is_node_ptr_kind root_ptr"
using assms(3) known_ptrs_known_ptr known_ptr_node_or_document
by auto
then obtain root_node_ptr where root_node_ptr: "root_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 root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> get_owner_document root_ptr \<rightarrow>\<^sub>r document"
using get_root_node_same_owner_document
using assms(1) assms(2) assms(3) document root_ptr by blast
then have "root_node_ptr \<in> set disc_nodes"
using assms(1) assms(2) assms(3) disc_nodes in_disconnected_nodes_no_parent root_node_ptr
using local.get_root_node_same_no_parent root_ptr by blast
then have "c \<in> set disconnected_tree_orders"
using c root_node_ptr
using map_M_pure_E[OF disconnected_tree_orders]
by (metis (mono_tags, lifting) comp_apply local.to_tree_order_pure select_result_I2)
then show ?thesis
by(auto simp add: sc)
qed
qed
lemma get_scdom_component_ptrs_same_owner_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
shows "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document"
proof -
obtain document disc_nodes tree_order disconnected_tree_orders where document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r document"
and disc_nodes: "h \<turnstile> get_disconnected_nodes document \<rightarrow>\<^sub>r disc_nodes"
and tree_order: "h \<turnstile> to_tree_order (cast document) \<rightarrow>\<^sub>r tree_order"
and disconnected_tree_orders: "h \<turnstile> map_M (to_tree_order \<circ> cast) disc_nodes \<rightarrow>\<^sub>r disconnected_tree_orders"
and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
using assms(4)
by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
)
show ?thesis
proof (cases "ptr' \<in> set tree_order")
case True
have "owner_document = document"
using assms(6) document by fastforce
then show ?thesis
by (metis (no_types) True assms(1) assms(2) assms(3) 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_inject document
document_ptr_casts_commute3 document_ptr_document_ptr_cast document_ptr_kinds_commutes
local.get_owner_document_owner_document_in_heap local.get_root_node_document
local.get_root_node_not_node_same local.to_tree_order_same_root node_ptr_no_document_ptr_cast
tree_order)
next
case False
then obtain disconnected_tree_order where disconnected_tree_order:
"ptr' \<in> set disconnected_tree_order" and "disconnected_tree_order \<in> set disconnected_tree_orders"
using sc \<open>ptr' \<in> set sc\<close>
by auto
obtain root_ptr' where
root_ptr': "root_ptr' \<in> set disc_nodes" and
"h \<turnstile> to_tree_order (cast root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order"
using map_M_pure_E2[OF disconnected_tree_orders \<open>disconnected_tree_order \<in> set disconnected_tree_orders\<close>]
by (metis comp_apply local.to_tree_order_pure)
have "\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). root_ptr' \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)"
using disc_nodes
by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok
- local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset
+ local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr
returns_result_select_result root_ptr')
then
have "h \<turnstile> get_parent root_ptr' \<rightarrow>\<^sub>r None"
using disc_nodes
- by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember_iff_member_fset local.get_parent_child_dual
+ by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3) local.get_parent_child_dual
local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_disc_nodes_in_heap
returns_result_select_result root_ptr' select_result_I2 split_option_ex)
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast root_ptr'"
using \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order\<close> assms(1)
assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent local.to_tree_order_get_root_node
local.to_tree_order_ptr_in_result
by blast
then have "h \<turnstile> get_owner_document (cast root_ptr') \<rightarrow>\<^sub>r document"
using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr'
by blast
then have "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r document"
using \<open>h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr'\<close> assms(1) assms(2) assms(3)
local.get_root_node_same_owner_document
by blast
then show ?thesis
using assms(6) document returns_result_eq by force
qed
qed
lemma get_scdom_component_ptrs_same_scope_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
shows "h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc"
proof -
obtain document disc_nodes tree_order disconnected_tree_orders where document:
"h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r document"
and disc_nodes: "h \<turnstile> get_disconnected_nodes document \<rightarrow>\<^sub>r disc_nodes"
and tree_order: "h \<turnstile> to_tree_order (cast document) \<rightarrow>\<^sub>r tree_order"
and disconnected_tree_orders: "h \<turnstile> map_M (to_tree_order \<circ> cast) disc_nodes \<rightarrow>\<^sub>r disconnected_tree_orders"
and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
using assms(4)
by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
)
show ?thesis
proof (cases "ptr' \<in> set tree_order")
case True
then have "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r document"
by (metis assms(1) assms(2) assms(3) 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_inject document document_ptr_casts_commute3
document_ptr_kinds_commutes known_ptr_node_or_document local.get_owner_document_owner_document_in_heap
local.get_root_node_document local.get_root_node_not_node_same local.known_ptrs_known_ptr
local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result node_ptr_no_document_ptr_cast tree_order)
then show ?thesis
using disc_nodes tree_order disconnected_tree_orders sc
by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
next
case False
then obtain disconnected_tree_order where disconnected_tree_order:
"ptr' \<in> set disconnected_tree_order" and "disconnected_tree_order \<in> set disconnected_tree_orders"
using sc \<open>ptr' \<in> set sc\<close>
by auto
obtain root_ptr' where
root_ptr': "root_ptr' \<in> set disc_nodes" and
"h \<turnstile> to_tree_order (cast root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order"
using map_M_pure_E2[OF disconnected_tree_orders \<open>disconnected_tree_order \<in> set disconnected_tree_orders\<close>]
by (metis comp_apply local.to_tree_order_pure)
have "\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). root_ptr' \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)"
using disc_nodes
by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok
- local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset
+ local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr
returns_result_select_result root_ptr')
then
have "h \<turnstile> get_parent root_ptr' \<rightarrow>\<^sub>r None"
using disc_nodes
- by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember_iff_member_fset local.get_parent_child_dual
+ by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3) local.get_parent_child_dual
local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_disc_nodes_in_heap
returns_result_select_result root_ptr' select_result_I2 split_option_ex)
then have "h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast root_ptr'"
using \<open>h \<turnstile> to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr') \<rightarrow>\<^sub>r disconnected_tree_order\<close> assms(1)
assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent local.to_tree_order_get_root_node
local.to_tree_order_ptr_in_result
by blast
then have "h \<turnstile> get_owner_document (cast root_ptr') \<rightarrow>\<^sub>r document"
using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr'
by blast
then have "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r document"
using \<open>h \<turnstile> get_root_node ptr' \<rightarrow>\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr'\<close> assms(1) assms(2) assms(3)
local.get_root_node_same_owner_document
by blast
then show ?thesis
using disc_nodes tree_order disconnected_tree_orders sc
by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
qed
qed
lemma get_scdom_component_ok:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_scdom_component ptr)"
using assms
apply(auto simp add: get_scdom_component_def intro!: bind_is_OK_pure_I map_M_pure_I map_M_ok_I)[1]
using get_owner_document_ok
apply blast
apply (simp add: local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap)
apply (simp add: local.get_owner_document_owner_document_in_heap local.to_tree_order_ok)
using local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok node_ptr_kinds_commutes
by blast
lemma get_scdom_component_ptr_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
shows "ptr' |\<in>| object_ptr_kinds h"
using assms
apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
using local.to_tree_order_ptrs_in_heap apply blast
by (metis (no_types, lifting) assms(4) assms(5) bind_returns_result_E get_scdom_component_impl
get_scdom_component_ptrs_same_scope_component is_OK_returns_result_I
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_scdom_component_def local.get_owner_document_ptr_in_heap)
lemma get_scdom_component_contains_get_dom_component:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
obtains c where "h \<turnstile> get_dom_component ptr' \<rightarrow>\<^sub>r c" and "set c \<subseteq> set sc"
proof -
have "h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc"
using assms(1) assms(2) assms(3) assms(4) assms(5) get_scdom_component_ptrs_same_scope_component
by blast
then show ?thesis
by (meson assms(1) assms(2) assms(3) assms(5) get_scdom_component_ptr_in_heap
get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok that)
qed
lemma get_scdom_component_owner_document_same:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "ptr' \<in> set sc"
obtains owner_document where "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document" and "cast owner_document \<in> set sc"
using assms
apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
apply (metis (no_types, lifting) assms(4) assms(5) document_ptr_casts_commute3
document_ptr_document_ptr_cast get_scdom_component_contains_get_dom_component local.get_dom_component_ptr
local.get_dom_component_root_node_same local.get_dom_component_to_tree_order local.get_root_node_document
local.get_root_node_not_node_same local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap
node_ptr_no_document_ptr_cast)
apply(rule map_M_pure_E2)
apply(simp)
apply(simp)
apply(simp)
by (smt (verit) assms(4) assms(5) comp_apply get_scdom_component_ptr_in_heap is_OK_returns_result_E
local.get_owner_document_disconnected_nodes local.get_root_node_ok local.get_root_node_same_owner_document
local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result)
lemma get_scdom_component_different_owner_documents:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
assumes "h \<turnstile> get_owner_document ptr' \<rightarrow>\<^sub>r owner_document'"
assumes "owner_document \<noteq> owner_document'"
shows "set |h \<turnstile> get_scdom_component ptr|\<^sub>r \<inter> set |h \<turnstile> get_scdom_component ptr'|\<^sub>r = {}"
using assms get_scdom_component_ptrs_same_owner_document
by (smt (verit) disjoint_iff_not_equal get_scdom_component_ok is_OK_returns_result_I
local.get_owner_document_ptr_in_heap returns_result_eq returns_result_select_result)
end
interpretation i_get_dom_component_get_scdom_component?: l_get_dom_component_get_scdom_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_owner_document
get_disconnected_nodes get_disconnected_nodes_locs to_tree_order heap_is_wellformed parent_child_rel
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node
get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name
get_elements_by_tag_name
by(auto simp add: l_get_dom_component_get_scdom_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_dom_component_get_scdom_component\<^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_dom_component_get_scdom_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_dom_component_get_scdom_component_is_l_get_dom_component_get_scdom_component [instances]:
"l_get_dom_component_get_scdom_component get_owner_document heap_is_wellformed type_wf known_ptr known_ptrs get_scdom_component get_dom_component"
apply(auto simp add: l_get_dom_component_get_scdom_component_def
l_get_dom_component_get_scdom_component_axioms_def instances)[1]
using get_scdom_component_subset_get_dom_component apply fast
using get_scdom_component_ptrs_same_scope_component apply fast
using get_scdom_component_ptrs_same_owner_document apply fast
using get_scdom_component_ok apply fast
using get_scdom_component_ptr_in_heap apply fast
using get_scdom_component_contains_get_dom_component apply blast
using get_scdom_component_owner_document_same apply blast
using get_scdom_component_different_owner_documents apply fast
done
subsection \<open>attach\_shadow\_root\<close>
lemma attach_shadow_root_not_strongly_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}, 'ShadowRoot::{equal,linorder}) heap" and
h' and host and new_shadow_root_ptr where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> attach_shadow_root host m \<rightarrow>\<^sub>r new_shadow_root_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_strongly_scdom_component_safe {cast host} {cast new_shadow_root_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder},
'ShadowRoot::{equal,linorder}) heap"
let ?P = "do {
doc \<leftarrow> create_document;
create_element doc ''div''
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e1 = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and host="?e1"])
by code_simp+
qed
locale l_get_scdom_component_attach_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_dom_component_attach_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_dom_component_get_scdom_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma attach_shadow_root_is_weakly_component_safe_step:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^sub>h h'"
assumes "ptr \<noteq> cast |h \<turnstile> attach_shadow_root element_ptr shadow_root_mode|\<^sub>r"
assumes "ptr \<notin> set |h \<turnstile> get_scdom_component (cast element_ptr)|\<^sub>r"
shows "preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
proof -
have "element_ptr |\<in>| element_ptr_kinds h"
by (meson assms(4) is_OK_returns_heap_I local.attach_shadow_root_element_ptr_in_heap)
obtain sc where sc: "h \<turnstile> get_scdom_component (cast element_ptr) \<rightarrow>\<^sub>r sc"
using get_scdom_component_ok
by (meson assms(1) assms(2) assms(3) assms(4) element_ptr_kinds_commutes is_OK_returns_heap_I
local.attach_shadow_root_element_ptr_in_heap node_ptr_kinds_commutes select_result_I)
then have "ptr \<notin> set |h \<turnstile> get_dom_component (cast element_ptr)|\<^sub>r"
by (metis (no_types, lifting) \<open>element_ptr |\<in>| element_ptr_kinds h\<close> assms(1) assms(2) assms(3)
assms(6) element_ptr_kinds_commutes local.get_dom_component_ok
local.get_scdom_component_subset_get_dom_component node_ptr_kinds_commutes
returns_result_select_result select_result_I2 set_rev_mp)
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) assms(5) local.attach_shadow_root_is_weakly_dom_component_safe
by blast
qed
lemma attach_shadow_root_is_weakly_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^sub>r result"
assumes "h \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^sub>h h'"
shows "is_weakly_scdom_component_safe {cast element_ptr} {cast result} h h'"
proof -
obtain h2 h3 new_shadow_root_ptr where
h2: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h2" and
new_shadow_root_ptr: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr" and
h3: "h2 \<turnstile> set_mode new_shadow_root_ptr shadow_root_mode \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> set_shadow_root element_ptr (Some new_shadow_root_ptr) \<rightarrow>\<^sub>h h'"
using assms(5)
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 \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^sub>r new_shadow_root_ptr"
using new_shadow_root_ptr h2 h3 h'
using assms(5)
by(auto simp add: attach_shadow_root_def intro!: bind_returns_result_I
bind_pure_returns_result_I[OF get_tag_name_pure] bind_pure_returns_result_I[OF get_shadow_root_pure]
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 "object_ptr_kinds h2 = {|cast result|} |\<union>| object_ptr_kinds h"
using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr
using new_shadow_root_ptr
using assms(4) by auto
moreover
have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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)
moreover
have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_shadow_root_writes h'])
using set_shadow_root_pointers_preserved
apply blast
by (auto simp add: reflp_def transp_def)
ultimately have "object_ptr_kinds h' = {|cast result|} |\<union>| object_ptr_kinds h"
by simp
moreover
have "result |\<notin>| shadow_root_ptr_kinds h"
using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr
using \<open>h \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^sub>r new_shadow_root_ptr\<close> assms(4)
returns_result_eq by metis
ultimately
show ?thesis
using assms
apply(auto simp add: is_weakly_scdom_component_safe_def Let_def)[1]
using attach_shadow_root_is_weakly_component_safe_step
by (smt (verit) document_ptr_kinds_commutes local.get_scdom_component_impl select_result_I2
shadow_root_ptr_kinds_commutes)
qed
end
interpretation i_get_scdom_component_attach_shadow_root?: l_get_scdom_component_attach_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe to_tree_order
get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root get_disconnected_nodes
get_disconnected_nodes_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs
by(auto simp add: l_get_scdom_component_attach_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_attach_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_shadow\_root\<close>
lemma get_shadow_root_not_weakly_scdom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder},
'Document::{equal,linorder}, 'Shadowroot::{equal,linorder}) heap" and
element_ptr and shadow_root_ptr_opt and h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> get_shadow_root_safe element_ptr \<rightarrow>\<^sub>r shadow_root_ptr_opt \<rightarrow>\<^sub>h h'" and
"\<not> is_weakly_scdom_component_safe {cast element_ptr} (cast ` set_option shadow_root_ptr_opt) h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder},
'Document::{equal,linorder}, 'Shadowroot::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
html \<leftarrow> create_element document_ptr ''html'';
append_child (cast document_ptr) (cast html);
head \<leftarrow> create_element document_ptr ''head'';
append_child (cast html) (cast head);
body \<leftarrow> create_element document_ptr ''body'';
append_child (cast html) (cast body);
e1 \<leftarrow> create_element document_ptr ''div'';
append_child (cast body) (cast e1);
e2 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
s1 \<leftarrow> attach_shadow_root e1 Open;
e3 \<leftarrow> create_element document_ptr ''slot'';
append_child (cast s1) (cast e3);
return e1
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e1 = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and element_ptr="?e1"])
by code_simp+
qed
locale l_get_shadow_root_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_dom_component_get_scdom_component +
l_get_shadow_root_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_create_document +
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_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_mode +
l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^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 +
assumes known_ptrs_impl: "known_ptrs = a_known_ptrs"
begin
lemma get_shadow_root_components_disjunct:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr"
shows "set |h \<turnstile> get_scdom_component (cast host)|\<^sub>r \<inter> set | h \<turnstile> get_scdom_component (cast shadow_root_ptr)|\<^sub>r = {}"
proof -
obtain owner_document where owner_document: "h \<turnstile> get_owner_document (cast host) \<rightarrow>\<^sub>r owner_document"
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(5) element_ptr_kinds_commutes
is_OK_returns_result_E is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_ptr
local.get_scdom_component_ok local.get_scdom_component_owner_document_same
local.get_scdom_component_subset_get_dom_component local.get_shadow_root_ptr_in_heap node_ptr_kinds_commutes
subset_code(1))
have "owner_document \<noteq> cast shadow_root_ptr"
proof
assume "owner_document = cast shadow_root_ptr"
then have "(cast owner_document, cast host) \<in>
(parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using get_owner_document_rel owner_document
by (metis (no_types, lifting) assms(1) assms(2) assms(3) cast_document_ptr_not_node_ptr(2)
in_rtrancl_UnI inf_sup_aci(6) inf_sup_aci(7))
then have "(cast shadow_root_ptr, cast host) \<in>
(parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
by (simp add: \<open>owner_document = cast shadow_root_ptr\<close>)
moreover have "(cast host, cast shadow_root_ptr) \<in> a_host_shadow_root_rel h"
by (metis (mono_tags, lifting) assms(5) is_OK_returns_result_I local.get_shadow_root_ptr_in_heap
local.a_host_shadow_root_rel_def mem_Collect_eq pair_imageI prod.simps(2) select_result_I2)
then have "(cast host, cast shadow_root_ptr) \<in>
(parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
by (simp)
ultimately show False
using assms(1)
unfolding heap_is_wellformed_def \<open>owner_document = cast shadow_root_ptr\<close> acyclic_def
by (meson rtrancl_into_trancl1)
qed
moreover
have "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
using assms(1) assms(5) local.get_shadow_root_shadow_root_ptr_in_heap by blast
then
have "cast shadow_root_ptr \<in> fset (object_ptr_kinds h)"
by auto
have "is_shadow_root_ptr shadow_root_ptr"
using assms(3)[unfolded known_ptrs_impl ShadowRootClass.known_ptrs_defs
ShadowRootClass.known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
ElementClass.known_ptr_defs NodeClass.known_ptr_defs, simplified, rule_format, OF
\<open>cast shadow_root_ptr \<in> fset (object_ptr_kinds h)\<close>]
by(auto simp add: is_document_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: option.splits document_ptr.splits)
then
have "h \<turnstile> get_owner_document (cast shadow_root_ptr) \<rightarrow>\<^sub>r cast shadow_root_ptr"
using \<open>shadow_root_ptr |\<in>| shadow_root_ptr_kinds h\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
by(auto simp add: is_node_ptr_kind_none a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
ultimately show ?thesis
using assms(1) assms(2) assms(3) get_scdom_component_different_owner_documents owner_document
by blast
qed
lemma get_shadow_root_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_shadow_root_safe element_ptr \<rightarrow>\<^sub>r shadow_root_ptr_opt \<rightarrow>\<^sub>h h'"
assumes "\<forall>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h). h \<turnstile> get_mode shadow_root_ptr \<rightarrow>\<^sub>r Closed"
shows "is_strongly_scdom_component_safe {cast element_ptr} (cast ` set_option shadow_root_ptr_opt) h h'"
proof -
have "h = h'"
using assms(4)
by(auto simp add: returns_result_heap_def pure_returns_heap_eq)
moreover have "shadow_root_ptr_opt = None"
using assms(4)
apply(auto simp add: returns_result_heap_def get_shadow_root_safe_def elim!: bind_returns_result_E2
split: option.splits if_splits)[1]
using get_shadow_root_shadow_root_ptr_in_heap
- by (meson assms(5) is_OK_returns_result_I local.get_mode_ptr_in_heap notin_fset returns_result_eq
+ by (meson assms(5) is_OK_returns_result_I local.get_mode_ptr_in_heap returns_result_eq
shadow_root_mode.distinct(1))
ultimately show ?thesis
by(simp add: is_strongly_scdom_component_safe_def preserved_def)
qed
end
interpretation i_get_shadow_root_scope_component?: l_get_shadow_root_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
is_strongly_dom_component_safe is_weakly_dom_component_safe 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\<^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 to_tree_order get_parent get_parent_locs get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
remove_shadow_root remove_shadow_root_locs create_document DocumentClass.known_ptr DocumentClass.type_wf
CD.a_get_owner_document get_mode get_mode_locs get_shadow_root_safe get_shadow_root_safe_locs
by(auto simp add: l_get_shadow_root_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_shadow_root_scope_component\<^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_shadow_root_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_host\<close>
lemma get_host_not_weakly_scdom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}, 'Shadowroot::{equal,linorder}) heap" and
shadow_root_ptr and element_ptr and h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r element_ptr \<rightarrow>\<^sub>h h'" and
"\<not> is_weakly_scdom_component_safe {cast shadow_root_ptr} {cast element_ptr} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder},
'Shadowroot::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
html \<leftarrow> create_element document_ptr ''html'';
append_child (cast document_ptr) (cast html);
head \<leftarrow> create_element document_ptr ''head'';
append_child (cast html) (cast head);
body \<leftarrow> create_element document_ptr ''body'';
append_child (cast html) (cast body);
e1 \<leftarrow> create_element document_ptr ''div'';
append_child (cast body) (cast e1);
e2 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
s1 \<leftarrow> attach_shadow_root e1 Open;
e3 \<leftarrow> create_element document_ptr ''slot'';
append_child (cast s1) (cast e3);
return s1
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?s1 = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and shadow_root_ptr="?s1"])
by code_simp+
qed
locale l_get_host_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_shadow_root_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_host_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_host_components_disjunct:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_scdom_component ptr \<rightarrow>\<^sub>r sc"
assumes "h \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r host"
shows "set |h \<turnstile> get_scdom_component (cast host)|\<^sub>r \<inter> set | h \<turnstile> get_scdom_component (cast shadow_root_ptr)|\<^sub>r = {}"
using assms get_shadow_root_components_disjunct local.shadow_root_host_dual
by blast
end
interpretation i_get_host_scope_component?: l_get_host_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr
known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe 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\<^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 to_tree_order get_parent get_parent_locs get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
remove_shadow_root remove_shadow_root_locs create_document DocumentClass.known_ptr DocumentClass.type_wf
CD.a_get_owner_document get_mode get_mode_locs get_shadow_root_safe get_shadow_root_safe_locs
by(auto simp add: l_get_host_scope_component\<^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_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsection \<open>get\_root\_node\_si\<close>
lemma get_composed_root_node_not_weakly_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}, 'Shadowroot::{equal,linorder}) heap" and
ptr and root and h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r root \<rightarrow>\<^sub>h h'" and
"\<not> is_weakly_scdom_component_safe {ptr} {root} h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder},
'Shadowroot::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
html \<leftarrow> create_element document_ptr ''html'';
append_child (cast document_ptr) (cast html);
head \<leftarrow> create_element document_ptr ''head'';
append_child (cast html) (cast head);
body \<leftarrow> create_element document_ptr ''body'';
append_child (cast html) (cast body);
e1 \<leftarrow> create_element document_ptr ''div'';
append_child (cast body) (cast e1);
e2 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
s1 \<leftarrow> attach_shadow_root e1 Closed;
e3 \<leftarrow> create_element document_ptr ''slot'';
append_child (cast s1) (cast e3);
return e3
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e3 = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e3"])
by code_simp+
qed
locale l_get_scdom_component_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_dom_component_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_dom_component_get_scdom_component
begin
lemma get_root_node_si_is_component_unsafe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node_si ptr' \<rightarrow>\<^sub>r root"
shows "set |h \<turnstile> get_scdom_component ptr'|\<^sub>r = set |h \<turnstile> get_scdom_component root|\<^sub>r \<or>
set |h \<turnstile> get_scdom_component ptr'|\<^sub>r \<inter> set |h \<turnstile> get_scdom_component root|\<^sub>r = {}"
proof -
have "ptr' |\<in>| object_ptr_kinds h"
using get_ancestors_si_ptr_in_heap assms(4)
by(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2)
then
obtain sc where "h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc"
by (meson assms(1) assms(2) assms(3) local.get_scdom_component_ok select_result_I)
moreover
have "root |\<in>| object_ptr_kinds h"
using get_ancestors_si_ptr assms(4)
apply(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2)[1]
by (metis (no_types, lifting) assms(1) assms(2) assms(3) empty_iff empty_set
get_ancestors_si_ptrs_in_heap last_in_set)
then
obtain sc' where "h \<turnstile> get_scdom_component root \<rightarrow>\<^sub>r sc'"
by (meson assms(1) assms(2) assms(3) local.get_scdom_component_ok select_result_I)
ultimately show ?thesis
by (metis (no_types, opaque_lifting) IntE \<open>\<And>thesis. (\<And>sc'. h \<turnstile> get_scdom_component root \<rightarrow>\<^sub>r sc' \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
\<open>\<And>thesis. (\<And>sc. h \<turnstile> get_scdom_component ptr' \<rightarrow>\<^sub>r sc \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> assms(1) assms(2) assms(3) empty_subsetI
local.get_scdom_component_ptrs_same_scope_component returns_result_eq select_result_I2 subsetI subset_antisym)
qed
end
interpretation i_get_scdom_component_get_root_node_si?: l_get_scdom_component_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 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 to_tree_order
get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
get_owner_document get_scdom_component is_strongly_scdom_component_safe
by(auto simp add: l_get_scdom_component_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_scdom_component_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]
subsection \<open>get\_assigned\_nodes\<close>
lemma assigned_nodes_not_weakly_scdom_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}, 'Shadowroot::{equal,linorder}) heap" and
node_ptr and nodes and h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> assigned_nodes node_ptr \<rightarrow>\<^sub>r nodes \<rightarrow>\<^sub>h h'" and
"\<not> is_weakly_scdom_component_safe {cast node_ptr} (cast ` set nodes) h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}, 'Shadowroot::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
html \<leftarrow> create_element document_ptr ''html'';
append_child (cast document_ptr) (cast html);
head \<leftarrow> create_element document_ptr ''head'';
append_child (cast html) (cast head);
body \<leftarrow> create_element document_ptr ''body'';
append_child (cast html) (cast body);
e1 \<leftarrow> create_element document_ptr ''div'';
append_child (cast body) (cast e1);
e2 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
s1 \<leftarrow> attach_shadow_root e1 Closed;
e3 \<leftarrow> create_element document_ptr ''slot'';
append_child (cast s1) (cast e3);
return e3
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e3 = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and node_ptr="?e3"])
by code_simp+
qed
lemma assigned_slot_not_weakly_component_safe:
obtains
h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}, 'Shadowroot::{equal,linorder}) heap" and
node_ptr and slot_opt and h' where
"heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
"h \<turnstile> assigned_slot node_ptr \<rightarrow>\<^sub>r slot_opt \<rightarrow>\<^sub>h h'" and
"\<not> is_weakly_scdom_component_safe {cast node_ptr} (cast ` set_option slot_opt) h h'"
proof -
let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder},
'Shadowroot::{equal,linorder}) heap"
let ?P = "do {
document_ptr \<leftarrow> create_document;
html \<leftarrow> create_element document_ptr ''html'';
append_child (cast document_ptr) (cast html);
head \<leftarrow> create_element document_ptr ''head'';
append_child (cast html) (cast head);
body \<leftarrow> create_element document_ptr ''body'';
append_child (cast html) (cast body);
e1 \<leftarrow> create_element document_ptr ''div'';
append_child (cast body) (cast e1);
e2 \<leftarrow> create_element document_ptr ''div'';
append_child (cast e1) (cast e2);
s1 \<leftarrow> attach_shadow_root e1 Open;
e3 \<leftarrow> create_element document_ptr ''slot'';
append_child (cast s1) (cast e3);
return e2
}"
let ?h1 = "|?h0 \<turnstile> ?P|\<^sub>h"
let ?e2 = "|?h0 \<turnstile> ?P|\<^sub>r"
show thesis
apply(rule that[where h="?h1" and node_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>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e2"])
by code_simp+
qed
locale l_assigned_nodes_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_assigned_nodes_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_shadow_root_scope_component\<^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
begin
lemma find_slot_is_component_unsafe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> find_slot open_flag node_ptr \<rightarrow>\<^sub>r Some slot"
shows "set |h \<turnstile> get_scdom_component (cast node_ptr)|\<^sub>r \<inter> set |h \<turnstile> get_scdom_component (cast slot)|\<^sub>r = {}"
proof -
obtain host shadow_root_ptr to where
"h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some (cast host)" and
"h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr" and
"h \<turnstile> to_tree_order (cast shadow_root_ptr) \<rightarrow>\<^sub>r to" and
"cast slot \<in> set to"
using assms(4)
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]
by (metis element_ptr_casts_commute3)+
have "node_ptr |\<in>| node_ptr_kinds h"
using assms(4) find_slot_ptr_in_heap by blast
then obtain node_ptr_c where node_ptr_c: "h \<turnstile> get_scdom_component (cast node_ptr) \<rightarrow>\<^sub>r node_ptr_c"
using assms(1) assms(2) assms(3) get_scdom_component_ok is_OK_returns_result_E
node_ptr_kinds_commutes[symmetric]
by metis
then have "cast host \<in> set node_ptr_c"
using \<open>h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some (cast host)\<close> assms(1) assms(2) assms(3)
by (meson assms(4) is_OK_returns_result_E local.find_slot_ptr_in_heap local.get_dom_component_ok
local.get_dom_component_parent_inside local.get_dom_component_ptr
local.get_scdom_component_subset_get_dom_component node_ptr_kinds_commutes subsetCE)
then have "h \<turnstile> get_scdom_component (cast host) \<rightarrow>\<^sub>r node_ptr_c"
using \<open>h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r Some (cast host)\<close> a_heap_is_wellformed_def assms(1) assms(2)
assms(3) node_ptr_c
using local.get_scdom_component_ptrs_same_scope_component by blast
moreover have "slot |\<in>| element_ptr_kinds h"
using assms(4) find_slot_slot_in_heap by blast
then obtain slot_c where slot_c: "h \<turnstile> get_scdom_component (cast slot) \<rightarrow>\<^sub>r slot_c"
using a_heap_is_wellformed_def assms(1) assms(2) assms(3) get_scdom_component_ok
is_OK_returns_result_E node_ptr_kinds_commutes[symmetric] element_ptr_kinds_commutes[symmetric]
by (smt (verit, best))
then have "cast shadow_root_ptr \<in> set slot_c"
using \<open>h \<turnstile> to_tree_order (cast shadow_root_ptr) \<rightarrow>\<^sub>r to\<close> \<open>cast slot \<in> set to\<close> assms(1) assms(2) assms(3)
by (meson is_OK_returns_result_E local.get_dom_component_ok local.get_dom_component_ptr
local.get_dom_component_to_tree_order local.get_scdom_component_subset_get_dom_component local.to_tree_order_ptrs_in_heap subsetCE)
then have "h \<turnstile> get_scdom_component (cast shadow_root_ptr) \<rightarrow>\<^sub>r slot_c"
using \<open>h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> assms(1) assms(2) assms(3) slot_c
using local.get_scdom_component_ptrs_same_scope_component by blast
ultimately show ?thesis
using get_shadow_root_components_disjunct assms \<open>h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr\<close>
node_ptr_c slot_c
by (metis (no_types, lifting) select_result_I2)
qed
lemma assigned_nodes_is_component_unsafe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> assigned_nodes element_ptr \<rightarrow>\<^sub>r nodes"
assumes "node_ptr \<in> set nodes"
shows "set |h \<turnstile> get_scdom_component (cast element_ptr)|\<^sub>r \<inter> set |h \<turnstile> get_scdom_component (cast node_ptr)|\<^sub>r = {}"
using assms
proof -
have "h \<turnstile> find_slot False node_ptr \<rightarrow>\<^sub>r Some element_ptr"
using assms(4) assms(5)
by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2
dest!: filter_M_holds_for_result[where x=node_ptr] intro!: bind_pure_I split: if_splits)
then show ?thesis
using assms find_slot_is_component_unsafe
by blast
qed
lemma assigned_slot_is_strongly_scdom_component_safe:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> assigned_slot element_ptr \<rightarrow>\<^sub>r slot_opt \<rightarrow>\<^sub>h h'"
assumes "\<forall>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h). h \<turnstile> get_mode shadow_root_ptr \<rightarrow>\<^sub>r Closed"
shows "is_strongly_scdom_component_safe {cast element_ptr} (cast ` set_option slot_opt) h h'"
proof -
have "h = h'"
using assms(4) find_slot_pure
by(auto simp add: assigned_slot_def returns_result_heap_def pure_returns_heap_eq find_slot_impl)
moreover have "slot_opt = None"
using assms(4) assms(5)
apply(auto simp add: returns_result_heap_def assigned_slot_def a_find_slot_def
elim!: bind_returns_result_E2 split: option.splits if_splits
dest!: get_shadow_root_shadow_root_ptr_in_heap[OF assms(1)])[1]
- apply (meson finite_set_in returns_result_eq shadow_root_mode.distinct(1))
- apply (meson finite_set_in returns_result_eq shadow_root_mode.distinct(1))
+ apply (meson returns_result_eq shadow_root_mode.distinct(1))
+ apply (meson returns_result_eq shadow_root_mode.distinct(1))
done
ultimately show ?thesis
by(auto simp add: is_strongly_scdom_component_safe_def preserved_def)
qed
end
interpretation i_assigned_nodes_scope_component?: l_assigned_nodes_scope_component\<^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 known_ptr get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_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 get_parent get_parent_locs
get_mode get_mode_locs get_attribute get_attribute_locs first_in_tree_order find_slot
assigned_slot known_ptrs to_tree_order assigned_nodes assigned_nodes_flatten flatten_dom
get_root_node get_root_node_locs remove insert_before insert_before_locs append_child
remove_shadow_root remove_shadow_root_locs set_shadow_root set_shadow_root_locs remove_child
remove_child_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe
get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
get_owner_document set_disconnected_nodes set_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs
adopt_node adopt_node_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs
get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe create_document
DocumentClass.known_ptr DocumentClass.type_wf CD.a_get_owner_document get_shadow_root_safe
get_shadow_root_safe_locs
by(auto simp add: l_assigned_nodes_scope_component\<^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_scope_component\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
end
diff --git a/thys/Safe_OCL/Finite_Map_Ext.thy b/thys/Safe_OCL/Finite_Map_Ext.thy
--- a/thys/Safe_OCL/Finite_Map_Ext.thy
+++ b/thys/Safe_OCL/Finite_Map_Ext.thy
@@ -1,445 +1,445 @@
(* Title: Safe OCL
Author: Denis Nikiforov, March 2019
Maintainer: Denis Nikiforov <denis.nikif at gmail.com>
License: LGPL
*)
section \<open>Finite Maps\<close>
theory Finite_Map_Ext
imports "HOL-Library.Finite_Map"
begin
type_notation fmap ("(_ \<rightharpoonup>\<^sub>f /_)" [22, 21] 21)
nonterminal fmaplets and fmaplet
syntax
"_fmaplet" :: "['a, 'a] \<Rightarrow> fmaplet" ("_ /\<mapsto>\<^sub>f/ _")
"_fmaplets" :: "['a, 'a] \<Rightarrow> fmaplet" ("_ /[\<mapsto>\<^sub>f]/ _")
"" :: "fmaplet \<Rightarrow> fmaplets" ("_")
"_FMaplets" :: "[fmaplet, fmaplets] \<Rightarrow> fmaplets" ("_,/ _")
"_FMapUpd" :: "['a \<rightharpoonup> 'b, fmaplets] \<Rightarrow> 'a \<rightharpoonup> 'b" ("_/'(_')" [900, 0] 900)
"_FMap" :: "fmaplets \<Rightarrow> 'a \<rightharpoonup> 'b" ("(1[_])")
syntax (ASCII)
"_fmaplet" :: "['a, 'a] \<Rightarrow> fmaplet" ("_ /|->f/ _")
"_fmaplets" :: "['a, 'a] \<Rightarrow> fmaplet" ("_ /[|->f]/ _")
translations
"_FMapUpd m (_FMaplets xy ms)" \<rightleftharpoons> "_FMapUpd (_FMapUpd m xy) ms"
"_FMapUpd m (_fmaplet x y)" \<rightleftharpoons> "CONST fmupd x y m"
"_FMap ms" \<rightleftharpoons> "_FMapUpd (CONST fmempty) ms"
"_FMap (_FMaplets ms1 ms2)" \<leftharpoondown> "_FMapUpd (_FMap ms1) ms2"
"_FMaplets ms1 (_FMaplets ms2 ms3)" \<leftharpoondown> "_FMaplets (_FMaplets ms1 ms2) ms3"
(*** Helper Lemmas **********************************************************)
subsection \<open>Helper Lemmas\<close>
lemma fmrel_on_fset_fmdom:
"fmrel_on_fset (fmdom ym) f xm ym \<Longrightarrow>
k |\<in>| fmdom ym \<Longrightarrow>
k |\<in>| fmdom xm"
by (metis fmdom_notD fmdom_notI fmrel_on_fsetD option.rel_sel)
(*** Finite Map Merge *******************************************************)
subsection \<open>Merge Operation\<close>
definition "fmmerge f xm ym \<equiv>
fmap_of_list (map
(\<lambda>k. (k, f (the (fmlookup xm k)) (the (fmlookup ym k))))
(sorted_list_of_fset (fmdom xm |\<inter>| fmdom ym)))"
lemma fmdom_fmmerge [simp]:
"fmdom (fmmerge g xm ym) = fmdom xm |\<inter>| fmdom ym"
by (auto simp add: fmmerge_def fmdom_of_list)
lemma fmmerge_commut:
assumes "\<And>x y. x \<in> fmran' xm \<Longrightarrow> f x y = f y x"
shows "fmmerge f xm ym = fmmerge f ym xm"
proof -
obtain zm where zm: "zm = sorted_list_of_fset (fmdom xm |\<inter>| fmdom ym)"
by auto
with assms have
"map (\<lambda>k. (k, f (the (fmlookup xm k)) (the (fmlookup ym k)))) zm =
map (\<lambda>k. (k, f (the (fmlookup ym k)) (the (fmlookup xm k)))) zm"
- by (auto) (metis fmdom_notI fmran'I notin_fset option.collapse)
+ by (auto) (metis fmdom_notI fmran'I option.collapse)
thus ?thesis
unfolding fmmerge_def zm
by (metis (no_types, lifting) inf_commute)
qed
lemma fmrel_on_fset_fmmerge1 [intro]:
assumes "\<And>x y z. z \<in> fmran' zm \<Longrightarrow> f x z \<Longrightarrow> f y z \<Longrightarrow> f (g x y) z"
assumes "fmrel_on_fset (fmdom zm) f xm zm"
assumes "fmrel_on_fset (fmdom zm) f ym zm"
shows "fmrel_on_fset (fmdom zm) f (fmmerge g xm ym) zm"
proof -
{
fix x a b c
assume "x |\<in>| fmdom zm"
moreover hence "x |\<in>| fmdom xm |\<inter>| fmdom ym"
by (meson assms(2) assms(3) finterI fmrel_on_fset_fmdom)
moreover assume "fmlookup xm x = Some a"
and "fmlookup ym x = Some b"
and "fmlookup zm x = Some c"
moreover from assms calculation have "f (g a b) c"
by (metis fmran'I fmrel_on_fsetD option.rel_inject(2))
ultimately have
"rel_option f (fmlookup (fmmerge g xm ym) x) (fmlookup zm x)"
unfolding fmmerge_def fmlookup_of_list apply auto
unfolding option_rel_Some2 apply (rule_tac ?x="g a b" in exI)
unfolding map_of_map_restrict restrict_map_def
- by (auto simp: fmember_iff_member_fset)
+ by auto
}
with assms(2) assms(3) show ?thesis
by (meson fmdomE fmrel_on_fsetI fmrel_on_fset_fmdom)
qed
lemma fmrel_on_fset_fmmerge2 [intro]:
assumes "\<And>x y. x \<in> fmran' xm \<Longrightarrow> f x (g x y)"
shows "fmrel_on_fset (fmdom ym) f xm (fmmerge g xm ym)"
proof -
{
fix x a b
assume "x |\<in>| fmdom xm |\<inter>| fmdom ym"
and "fmlookup xm x = Some a"
and "fmlookup ym x = Some b"
hence "rel_option f (fmlookup xm x) (fmlookup (fmmerge g xm ym) x)"
unfolding fmmerge_def fmlookup_of_list apply auto
unfolding option_rel_Some1 apply (rule_tac ?x="g a b" in exI)
- by (auto simp add: map_of_map_restrict fmember_iff_member_fset assms fmran'I)
+ by (auto simp add: map_of_map_restrict assms fmran'I)
}
with assms show ?thesis
apply auto
apply (rule fmrel_on_fsetI)
by (metis (full_types) finterD1 fmdomE fmdom_fmmerge fmdom_notD rel_option_None2)
qed
(*** Acyclicity *************************************************************)
subsection \<open>Acyclicity\<close>
abbreviation "acyclic_on xs r \<equiv> (\<forall>x. x \<in> xs \<longrightarrow> (x, x) \<notin> r\<^sup>+)"
abbreviation "acyclicP_on xs r \<equiv> acyclic_on xs {(x, y). r x y}"
lemma fmrel_acyclic:
"acyclicP_on (fmran' xm) R \<Longrightarrow>
fmrel R\<^sup>+\<^sup>+ xm ym \<Longrightarrow>
fmrel R ym xm \<Longrightarrow>
xm = ym"
by (metis (full_types) fmap_ext fmran'I fmrel_cases option.sel
tranclp.trancl_into_trancl tranclp_unfold)
lemma fmrel_acyclic':
assumes "acyclicP_on (fmran' ym) R"
assumes "fmrel R\<^sup>+\<^sup>+ xm ym"
assumes "fmrel R ym xm"
shows "xm = ym"
proof -
{
fix x
from assms(1) have
"rel_option R\<^sup>+\<^sup>+ (fmlookup xm x) (fmlookup ym x) \<Longrightarrow>
rel_option R (fmlookup ym x) (fmlookup xm x) \<Longrightarrow>
rel_option R (fmlookup xm x) (fmlookup ym x)"
by (metis (full_types) fmdom'_notD fmlookup_dom'_iff
fmran'I option.rel_sel option.sel
tranclp_into_tranclp2 tranclp_unfold)
}
with assms show ?thesis
unfolding fmrel_iff
by (metis fmap.rel_mono_strong fmrelI fmrel_acyclic tranclp.simps)
qed
lemma fmrel_on_fset_acyclic:
"acyclicP_on (fmran' xm) R \<Longrightarrow>
fmrel_on_fset (fmdom ym) R\<^sup>+\<^sup>+ xm ym \<Longrightarrow>
fmrel_on_fset (fmdom xm) R ym xm \<Longrightarrow>
xm = ym"
unfolding fmrel_on_fset_fmrel_restrict
by (metis (no_types, lifting) fmdom_filter fmfilter_alt_defs(5)
fmfilter_cong fmlookup_filter fmrel_acyclic fmrel_fmdom_eq
fmrestrict_fset_dom option.simps(3))
lemma fmrel_on_fset_acyclic':
"acyclicP_on (fmran' ym) R \<Longrightarrow>
fmrel_on_fset (fmdom ym) R\<^sup>+\<^sup>+ xm ym \<Longrightarrow>
fmrel_on_fset (fmdom xm) R ym xm \<Longrightarrow>
xm = ym"
unfolding fmrel_on_fset_fmrel_restrict
by (metis (no_types, lifting) ffmember_filter fmdom_filter
fmfilter_alt_defs(5) fmfilter_cong fmrel_acyclic'
fmrel_fmdom_eq fmrestrict_fset_dom)
(*** Transitive Closures ****************************************************)
subsection \<open>Transitive Closures\<close>
lemma fmrel_trans:
"(\<And>x y z. x \<in> fmran' xm \<Longrightarrow> P x y \<Longrightarrow> Q y z \<Longrightarrow> R x z) \<Longrightarrow>
fmrel P xm ym \<Longrightarrow> fmrel Q ym zm \<Longrightarrow> fmrel R xm zm"
unfolding fmrel_iff
by (metis fmdomE fmdom_notD fmran'I option.rel_inject(2) option.rel_sel)
lemma fmrel_on_fset_trans:
"(\<And>x y z. x \<in> fmran' xm \<Longrightarrow> P x y \<Longrightarrow> Q y z \<Longrightarrow> R x z) \<Longrightarrow>
fmrel_on_fset (fmdom ym) P xm ym \<Longrightarrow>
fmrel_on_fset (fmdom zm) Q ym zm \<Longrightarrow>
fmrel_on_fset (fmdom zm) R xm zm"
apply (rule fmrel_on_fsetI)
unfolding option.rel_sel apply auto
apply (meson fmdom_notI fmrel_on_fset_fmdom)
by (metis fmdom_notI fmran'I fmrel_on_fsetD fmrel_on_fset_fmdom
option.rel_sel option.sel)
lemma trancl_to_fmrel:
"(fmrel f)\<^sup>+\<^sup>+ xm ym \<Longrightarrow> fmrel f\<^sup>+\<^sup>+ xm ym"
apply (induct rule: tranclp_induct)
apply (simp add: fmap.rel_mono_strong)
by (rule fmrel_trans; auto)
lemma fmrel_trancl_fmdom_eq:
"(fmrel f)\<^sup>+\<^sup>+ xm ym \<Longrightarrow> fmdom xm = fmdom ym"
by (induct rule: tranclp_induct; simp add: fmrel_fmdom_eq)
text \<open>
The proof was derived from the accepted answer on the website
Stack Overflow that is available at
@{url "https://stackoverflow.com/a/53585232/632199"}
and provided with the permission of the author of the answer.\<close>
lemma fmupd_fmdrop:
"fmlookup xm k = Some x \<Longrightarrow>
xm = fmupd k x (fmdrop k xm)"
apply (rule fmap_ext)
unfolding fmlookup_drop fmupd_lookup
by auto
lemma fmap_eqdom_Cons1:
assumes "fmlookup xm i = None"
and "fmdom (fmupd i x xm) = fmdom ym"
and "fmrel R (fmupd i x xm) ym"
shows "(\<exists>z zm. fmlookup zm i = None \<and> ym = (fmupd i z zm) \<and>
R x z \<and> fmrel R xm zm)"
proof -
from assms(2) obtain y where "fmlookup ym i = Some y" by force
then obtain z zm where z_zm: "ym = fmupd i z zm \<and> fmlookup zm i = None"
using fmupd_fmdrop by force
{
assume "\<not> R x z"
with z_zm have "\<not> fmrel R (fmupd i x xm) ym"
by (metis fmrel_iff fmupd_lookup option.simps(11))
}
with assms(3) moreover have "R x z" by auto
{
assume "\<not> fmrel R xm zm"
with assms(1) have "\<not> fmrel R (fmupd i x xm) ym"
by (metis fmrel_iff fmupd_lookup option.rel_sel z_zm)
}
with assms(3) moreover have "fmrel R xm zm" by auto
ultimately show ?thesis using z_zm by blast
qed
text \<open>
The proof was derived from the accepted answer on the website
Stack Overflow that is available at
@{url "https://stackoverflow.com/a/53585232/632199"}
and provided with the permission of the author of the answer.\<close>
lemma fmap_eqdom_induct [consumes 2, case_names nil step]:
assumes R: "fmrel R xm ym"
and dom_eq: "fmdom xm = fmdom ym"
and nil: "P (fmap_of_list []) (fmap_of_list [])"
and step:
"\<And>x xm y ym i.
\<lbrakk>R x y; fmrel R xm ym; fmdom xm = fmdom ym; P xm ym\<rbrakk> \<Longrightarrow>
P (fmupd i x xm) (fmupd i y ym)"
shows "P xm ym"
using R dom_eq
proof (induct xm arbitrary: ym)
case fmempty thus ?case
by (metis fempty_iff fmdom_empty fmempty_of_list fmfilter_alt_defs(5)
fmfilter_false fmrestrict_fset_dom local.nil)
next
case (fmupd i x xm) show ?case
proof -
obtain y where "fmlookup ym i = Some y"
by (metis fmupd.prems(1) fmrel_cases fmupd_lookup option.discI)
from fmupd.hyps(2) fmupd.prems(1) fmupd.prems(2) obtain z zm where
"fmlookup zm i = None" and
ym_eq_z_zm: "ym = (fmupd i z zm)" and
R_x_z: "R x z" and
R_xm_zm: "fmrel R xm zm"
using fmap_eqdom_Cons1 by metis
hence dom_xm_eq_dom_zm: "fmdom xm = fmdom zm"
using fmrel_fmdom_eq by blast
with R_xm_zm fmupd.hyps(1) have "P xm zm" by blast
with R_x_z R_xm_zm dom_xm_eq_dom_zm have
"P (fmupd i x xm) (fmupd i z zm)"
by (rule step)
thus ?thesis by (simp add: ym_eq_z_zm)
qed
qed
text \<open>
The proof was derived from the accepted answer on the website
Stack Overflow that is available at
@{url "https://stackoverflow.com/a/53585232/632199"}
and provided with the permission of the author of the answer.\<close>
lemma fmrel_to_rtrancl:
assumes as_r: "reflp r"
and rel_rpp_xm_ym: "fmrel r\<^sup>*\<^sup>* xm ym"
shows "(fmrel r)\<^sup>*\<^sup>* xm ym"
proof -
from rel_rpp_xm_ym have "fmdom xm = fmdom ym"
using fmrel_fmdom_eq by blast
with rel_rpp_xm_ym show "(fmrel r)\<^sup>*\<^sup>* xm ym"
proof (induct rule: fmap_eqdom_induct)
case nil show ?case by auto
next
case (step x xm y ym i) show ?case
proof -
from step.hyps(1) have "(fmrel r)\<^sup>*\<^sup>* (fmupd i x xm) (fmupd i y xm)"
proof (induct rule: rtranclp_induct)
case base show ?case by simp
next
case (step y z) show ?case
proof -
from as_r have "fmrel r xm xm"
by (simp add: fmap.rel_reflp reflpD)
with step.hyps(2) have "(fmrel r)\<^sup>*\<^sup>* (fmupd i y xm) (fmupd i z xm)"
by (simp add: fmrel_upd r_into_rtranclp)
with step.hyps(3) show ?thesis by simp
qed
qed
also from step.hyps(4) have "(fmrel r)\<^sup>*\<^sup>* (fmupd i y xm) (fmupd i y ym)"
proof (induct rule: rtranclp_induct)
case base show ?case by simp
next
case (step ya za) show ?case
proof -
from step.hyps(2) as_r have "(fmrel r)\<^sup>*\<^sup>* (fmupd i y ya) (fmupd i y za)"
by (simp add: fmrel_upd r_into_rtranclp reflp_def)
with step.hyps(3) show ?thesis by simp
qed
qed
finally show ?thesis by simp
qed
qed
qed
text \<open>
The proof was derived from the accepted answer on the website
Stack Overflow that is available at
@{url "https://stackoverflow.com/a/53585232/632199"}
and provided with the permission of the author of the answer.\<close>
lemma fmrel_to_trancl:
assumes "reflp r"
and "fmrel r\<^sup>+\<^sup>+ xm ym"
shows "(fmrel r)\<^sup>+\<^sup>+ xm ym"
proof -
from assms(2) have "fmrel r\<^sup>*\<^sup>* xm ym"
by (drule_tac ?Ra="r\<^sup>*\<^sup>*" in fmap.rel_mono_strong; auto)
with assms(1) have "(fmrel r)\<^sup>*\<^sup>* xm ym"
by (simp add: fmrel_to_rtrancl)
with assms(1) show ?thesis
by (metis fmap.rel_reflp reflpD rtranclpD tranclp.r_into_trancl)
qed
lemma fmrel_tranclp_induct:
"fmrel r\<^sup>+\<^sup>+ a b \<Longrightarrow>
reflp r \<Longrightarrow>
(\<And>y. fmrel r a y \<Longrightarrow> P y) \<Longrightarrow>
(\<And>y z. (fmrel r)\<^sup>+\<^sup>+ a y \<Longrightarrow> fmrel r y z \<Longrightarrow> P y \<Longrightarrow> P z) \<Longrightarrow> P b"
apply (drule fmrel_to_trancl, simp)
by (erule tranclp_induct; simp)
lemma fmrel_converse_tranclp_induct:
"fmrel r\<^sup>+\<^sup>+ a b \<Longrightarrow>
reflp r \<Longrightarrow>
(\<And>y. fmrel r y b \<Longrightarrow> P y) \<Longrightarrow>
(\<And>y z. fmrel r y z \<Longrightarrow> fmrel r\<^sup>+\<^sup>+ z b \<Longrightarrow> P z \<Longrightarrow> P y) \<Longrightarrow> P a"
apply (drule fmrel_to_trancl, simp)
by (erule converse_tranclp_induct; simp add: trancl_to_fmrel)
lemma fmrel_tranclp_trans_induct:
"fmrel r\<^sup>+\<^sup>+ a b \<Longrightarrow>
reflp r \<Longrightarrow>
(\<And>x y. fmrel r x y \<Longrightarrow> P x y) \<Longrightarrow>
(\<And>x y z. fmrel r\<^sup>+\<^sup>+ x y \<Longrightarrow> P x y \<Longrightarrow> fmrel r\<^sup>+\<^sup>+ y z \<Longrightarrow> P y z \<Longrightarrow> P x z) \<Longrightarrow>
P a b"
apply (drule fmrel_to_trancl, simp)
apply (erule tranclp_trans_induct, simp)
using trancl_to_fmrel by blast
(*** Finite Map Size Calculation ********************************************)
subsection \<open>Size Calculation\<close>
text \<open>
The contents of the subsection was derived from the accepted answer
on the website Stack Overflow that is available at
@{url "https://stackoverflow.com/a/53244203/632199"}
and provided with the permission of the author of the answer.\<close>
abbreviation "tcf \<equiv> (\<lambda> v::('a \<times> nat). (\<lambda> r::nat. snd v + r))"
interpretation tcf: comp_fun_commute tcf
proof
fix x y :: "'a \<times> nat"
show "tcf y \<circ> tcf x = tcf x \<circ> tcf y"
proof -
fix z
have "(tcf y \<circ> tcf x) z = snd y + snd x + z" by auto
also have "(tcf x \<circ> tcf y) z = snd y + snd x + z" by auto
finally have "(tcf y \<circ> tcf x) z = (tcf x \<circ> tcf y) z" by auto
thus "(tcf y \<circ> tcf x) = (tcf x \<circ> tcf y)" by auto
qed
qed
lemma ffold_rec_exp:
assumes "k |\<in>| fmdom x"
and "ky = (k, the (fmlookup (fmmap f x) k))"
shows "ffold tcf 0 (fset_of_fmap (fmmap f x)) =
tcf ky (ffold tcf 0 ((fset_of_fmap (fmmap f x)) |-| {|ky|}))"
proof -
have "ky |\<in>| (fset_of_fmap (fmmap f x))"
using assms by auto
thus ?thesis
by (simp add: tcf.ffold_rec)
qed
lemma elem_le_ffold [intro]:
"k |\<in>| fmdom x \<Longrightarrow>
f (the (fmlookup x k)) < Suc (ffold tcf 0 (fset_of_fmap (fmmap f x)))"
by (subst ffold_rec_exp, auto)
lemma elem_le_ffold' [intro]:
"z \<in> fmran' x \<Longrightarrow>
f z < Suc (ffold tcf 0 (fset_of_fmap (fmmap f x)))"
apply (erule fmran'E)
apply (frule fmdomI)
by (subst ffold_rec_exp, auto)
(*** Code Setup *************************************************************)
subsection \<open>Code Setup\<close>
abbreviation "fmmerge_fun f xm ym \<equiv>
fmap_of_list (map
(\<lambda>k. if k |\<in>| fmdom xm \<and> k |\<in>| fmdom ym
then (k, f (the (fmlookup xm k)) (the (fmlookup ym k)))
else (k, undefined))
(sorted_list_of_fset (fmdom xm |\<inter>| fmdom ym)))"
lemma fmmerge_fun_simp [code_abbrev, simp]:
"fmmerge_fun f xm ym = fmmerge f xm ym"
unfolding fmmerge_def
apply (rule_tac ?f="fmap_of_list" in HOL.arg_cong)
- by (simp add: notin_fset)
+ by simp
end
diff --git a/thys/Safe_OCL/OCL_Typing.thy b/thys/Safe_OCL/OCL_Typing.thy
--- a/thys/Safe_OCL/OCL_Typing.thy
+++ b/thys/Safe_OCL/OCL_Typing.thy
@@ -1,1095 +1,1096 @@
(* Title: Safe OCL
Author: Denis Nikiforov, March 2019
Maintainer: Denis Nikiforov <denis.nikif at gmail.com>
License: LGPL
*)
chapter \<open>Typing\<close>
theory OCL_Typing
imports OCL_Object_Model "HOL-Library.Transitive_Closure_Table"
begin
text \<open>
The following rules are more restrictive than rules given in
the OCL specification. This allows one to identify more errors
in expressions. However, these restrictions may be revised if necessary.
Perhaps some of them could be separated and should cause warnings
instead of errors.\<close>
(*** Operations Typing ******************************************************)
section \<open>Operations Typing\<close>
subsection \<open>Metaclass Operations\<close>
text \<open>
All basic types in the theory are either nullable or non-nullable.
For example, instead of @{text Boolean} type we have two types:
@{text "Boolean[1]"} and @{text "Boolean[?]"}.
The @{text "allInstances()"} operation is extended accordingly:\<close>
text \<open>
\<^verbatim>\<open>Boolean[1].allInstances() = Set{true, false}
Boolean[?].allInstances() = Set{true, false, null}\<close>\<close>
inductive mataop_type where
"mataop_type \<tau> AllInstancesOp (Set \<tau>)"
subsection \<open>Type Operations\<close>
text \<open>
At first we decided to allow casting only to subtypes.
However sometimes it is necessary to cast expressions to supertypes,
for example, to access overridden attributes of a supertype.
So we allow casting to subtypes and supertypes.
Casting to other types is meaningless.\<close>
text \<open>
According to the Section 7.4.7 of the OCL specification
@{text "oclAsType()"} can be applied to collections as well as
to single values. I guess we can allow @{text "oclIsTypeOf()"}
and @{text "oclIsKindOf()"} for collections too.\<close>
text \<open>
Please take a note that the following expressions are prohibited,
because they always return true or false:\<close>
text \<open>
\<^verbatim>\<open>1.oclIsKindOf(OclAny[?])
1.oclIsKindOf(String[1])\<close>\<close>
text \<open>
Please take a note that:\<close>
text \<open>
\<^verbatim>\<open>Set{1,2,null,'abc'}->selectByKind(Integer[1]) = Set{1,2}
Set{1,2,null,'abc'}->selectByKind(Integer[?]) = Set{1,2,null}\<close>\<close>
text \<open>
The following expressions are prohibited, because they always
returns either the same or empty collections:\<close>
text \<open>
\<^verbatim>\<open>Set{1,2,null,'abc'}->selectByKind(OclAny[?])
Set{1,2,null,'abc'}->selectByKind(Collection(Boolean[1]))\<close>\<close>
inductive typeop_type where
"\<sigma> < \<tau> \<or> \<tau> < \<sigma> \<Longrightarrow>
typeop_type DotCall OclAsTypeOp \<tau> \<sigma> \<sigma>"
| "\<sigma> < \<tau> \<Longrightarrow>
typeop_type DotCall OclIsTypeOfOp \<tau> \<sigma> Boolean[1]"
| "\<sigma> < \<tau> \<Longrightarrow>
typeop_type DotCall OclIsKindOfOp \<tau> \<sigma> Boolean[1]"
| "element_type \<tau> \<rho> \<Longrightarrow> \<sigma> < \<rho> \<Longrightarrow>
update_element_type \<tau> \<sigma> \<upsilon> \<Longrightarrow>
typeop_type ArrowCall SelectByKindOp \<tau> \<sigma> \<upsilon>"
| "element_type \<tau> \<rho> \<Longrightarrow> \<sigma> < \<rho> \<Longrightarrow>
update_element_type \<tau> \<sigma> \<upsilon> \<Longrightarrow>
typeop_type ArrowCall SelectByTypeOp \<tau> \<sigma> \<upsilon>"
subsection \<open>OclSuper Operations\<close>
text \<open>
It makes sense to compare values only with compatible types.\<close>
(* We have to specify the predicate type explicitly to let
a generated code work *)
inductive super_binop_type
:: "super_binop \<Rightarrow> ('a :: order) type \<Rightarrow> 'a type \<Rightarrow> 'a type \<Rightarrow> bool" where
"\<tau> \<le> \<sigma> \<or> \<sigma> < \<tau> \<Longrightarrow>
super_binop_type EqualOp \<tau> \<sigma> Boolean[1]"
| "\<tau> \<le> \<sigma> \<or> \<sigma> < \<tau> \<Longrightarrow>
super_binop_type NotEqualOp \<tau> \<sigma> Boolean[1]"
subsection \<open>OclAny Operations\<close>
text \<open>
The OCL specification defines @{text "toString()"} operation
only for boolean and numeric types. However, I guess it is a good
idea to define it once for all basic types. Maybe it should be defined
for collections as well.\<close>
inductive any_unop_type where
"\<tau> \<le> OclAny[?] \<Longrightarrow>
any_unop_type OclAsSetOp \<tau> (Set (to_required_type \<tau>))"
| "\<tau> \<le> OclAny[?] \<Longrightarrow>
any_unop_type OclIsNewOp \<tau> Boolean[1]"
| "\<tau> \<le> OclAny[?] \<Longrightarrow>
any_unop_type OclIsUndefinedOp \<tau> Boolean[1]"
| "\<tau> \<le> OclAny[?] \<Longrightarrow>
any_unop_type OclIsInvalidOp \<tau> Boolean[1]"
| "\<tau> \<le> OclAny[?] \<Longrightarrow>
any_unop_type OclLocaleOp \<tau> String[1]"
| "\<tau> \<le> OclAny[?] \<Longrightarrow>
any_unop_type ToStringOp \<tau> String[1]"
subsection \<open>Boolean Operations\<close>
text \<open>
Please take a note that:\<close>
text \<open>
\<^verbatim>\<open>true or false : Boolean[1]
true and null : Boolean[?]
null and null : OclVoid[?]\<close>\<close>
inductive boolean_unop_type where
"\<tau> \<le> Boolean[?] \<Longrightarrow>
boolean_unop_type NotOp \<tau> \<tau>"
inductive boolean_binop_type where
"\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> \<le> Boolean[?] \<Longrightarrow>
boolean_binop_type AndOp \<tau> \<sigma> \<rho>"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> \<le> Boolean[?] \<Longrightarrow>
boolean_binop_type OrOp \<tau> \<sigma> \<rho>"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> \<le> Boolean[?] \<Longrightarrow>
boolean_binop_type XorOp \<tau> \<sigma> \<rho>"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> \<le> Boolean[?] \<Longrightarrow>
boolean_binop_type ImpliesOp \<tau> \<sigma> \<rho>"
subsection \<open>Numeric Operations\<close>
text \<open>
The expression @{text "1 + null"} is not well-typed.
Nullable numeric values should be converted to non-nullable ones.
This is a significant difference from the OCL specification.\<close>
text \<open>
Please take a note that many operations automatically casts
unlimited naturals to integers.\<close>
text \<open>
The difference between @{text "oclAsType(Integer)"} and
@{text "toInteger()"} for unlimited naturals is unclear.\<close>
inductive numeric_unop_type where
"\<tau> = Real[1] \<Longrightarrow>
numeric_unop_type UMinusOp \<tau> Real[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
numeric_unop_type UMinusOp \<tau> Integer[1]"
| "\<tau> = Real[1] \<Longrightarrow>
numeric_unop_type AbsOp \<tau> Real[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
numeric_unop_type AbsOp \<tau> Integer[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_unop_type FloorOp \<tau> Integer[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_unop_type RoundOp \<tau> Integer[1]"
| "\<tau> = UnlimitedNatural[1] \<Longrightarrow>
numeric_unop_type numeric_unop.ToIntegerOp \<tau> Integer[1]"
inductive numeric_binop_type where
"\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type PlusOp \<tau> \<sigma> \<rho>"
| "\<tau> \<squnion> \<sigma> = Real[1] \<Longrightarrow>
numeric_binop_type MinusOp \<tau> \<sigma> Real[1]"
| "\<tau> \<squnion> \<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
numeric_binop_type MinusOp \<tau> \<sigma> Integer[1]"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type MultOp \<tau> \<sigma> \<rho>"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type DivideOp \<tau> \<sigma> Real[1]"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
numeric_binop_type DivOp \<tau> \<sigma> \<rho>"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
numeric_binop_type ModOp \<tau> \<sigma> \<rho>"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type MaxOp \<tau> \<sigma> \<rho>"
| "\<tau> \<squnion> \<sigma> = \<rho> \<Longrightarrow> \<rho> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type MinOp \<tau> \<sigma> \<rho>"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type numeric_binop.LessOp \<tau> \<sigma> Boolean[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type numeric_binop.LessEqOp \<tau> \<sigma> Boolean[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type numeric_binop.GreaterOp \<tau> \<sigma> Boolean[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
numeric_binop_type numeric_binop.GreaterEqOp \<tau> \<sigma> Boolean[1]"
subsection \<open>String Operations\<close>
inductive string_unop_type where
"string_unop_type SizeOp String[1] Integer[1]"
| "string_unop_type CharactersOp String[1] (Sequence String[1])"
| "string_unop_type ToUpperCaseOp String[1] String[1]"
| "string_unop_type ToLowerCaseOp String[1] String[1]"
| "string_unop_type ToBooleanOp String[1] Boolean[1]"
| "string_unop_type ToIntegerOp String[1] Integer[1]"
| "string_unop_type ToRealOp String[1] Real[1]"
inductive string_binop_type where
"string_binop_type ConcatOp String[1] String[1] String[1]"
| "string_binop_type EqualsIgnoreCaseOp String[1] String[1] Boolean[1]"
| "string_binop_type LessOp String[1] String[1] Boolean[1]"
| "string_binop_type LessEqOp String[1] String[1] Boolean[1]"
| "string_binop_type GreaterOp String[1] String[1] Boolean[1]"
| "string_binop_type GreaterEqOp String[1] String[1] Boolean[1]"
| "string_binop_type IndexOfOp String[1] String[1] Integer[1]"
| "\<tau> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
string_binop_type AtOp String[1] \<tau> String[1]"
inductive string_ternop_type where
"\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
\<rho> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
string_ternop_type SubstringOp String[1] \<sigma> \<rho> String[1]"
subsection \<open>Collection Operations\<close>
text \<open>
Please take a note, that @{text "flatten()"} preserves a collection kind.\<close>
inductive collection_unop_type where
"element_type \<tau> _ \<Longrightarrow>
collection_unop_type CollectionSizeOp \<tau> Integer[1]"
| "element_type \<tau> _ \<Longrightarrow>
collection_unop_type IsEmptyOp \<tau> Boolean[1]"
| "element_type \<tau> _ \<Longrightarrow>
collection_unop_type NotEmptyOp \<tau> Boolean[1]"
| "element_type \<tau> \<sigma> \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
collection_unop_type CollectionMaxOp \<tau> \<sigma>"
| "element_type \<tau> \<sigma> \<Longrightarrow> operation \<sigma> STR ''max'' [\<sigma>] oper \<Longrightarrow>
collection_unop_type CollectionMaxOp \<tau> \<sigma>"
| "element_type \<tau> \<sigma> \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
collection_unop_type CollectionMinOp \<tau> \<sigma>"
| "element_type \<tau> \<sigma> \<Longrightarrow> operation \<sigma> STR ''min'' [\<sigma>] oper \<Longrightarrow>
collection_unop_type CollectionMinOp \<tau> \<sigma>"
| "element_type \<tau> \<sigma> \<Longrightarrow> \<sigma> = UnlimitedNatural[1]\<midarrow>Real[1] \<Longrightarrow>
collection_unop_type SumOp \<tau> \<sigma>"
| "element_type \<tau> \<sigma> \<Longrightarrow> operation \<sigma> STR ''+'' [\<sigma>] oper \<Longrightarrow>
collection_unop_type SumOp \<tau> \<sigma>"
| "element_type \<tau> \<sigma> \<Longrightarrow>
collection_unop_type AsSetOp \<tau> (Set \<sigma>)"
| "element_type \<tau> \<sigma> \<Longrightarrow>
collection_unop_type AsOrderedSetOp \<tau> (OrderedSet \<sigma>)"
| "element_type \<tau> \<sigma> \<Longrightarrow>
collection_unop_type AsBagOp \<tau> (Bag \<sigma>)"
| "element_type \<tau> \<sigma> \<Longrightarrow>
collection_unop_type AsSequenceOp \<tau> (Sequence \<sigma>)"
| "update_element_type \<tau> (to_single_type \<tau>) \<sigma> \<Longrightarrow>
collection_unop_type FlattenOp \<tau> \<sigma>"
| "collection_unop_type FirstOp (OrderedSet \<tau>) \<tau>"
| "collection_unop_type FirstOp (Sequence \<tau>) \<tau>"
| "collection_unop_type LastOp (OrderedSet \<tau>) \<tau>"
| "collection_unop_type LastOp (Sequence \<tau>) \<tau>"
| "collection_unop_type ReverseOp (OrderedSet \<tau>) (OrderedSet \<tau>)"
| "collection_unop_type ReverseOp (Sequence \<tau>) (Sequence \<tau>)"
text \<open>
Please take a note that if both arguments are collections,
then an element type of the resulting collection is a super type
of element types of orginal collections. However for single-valued
operations (@{text "append()"}, @{text "insertAt()"}, ...)
this behavior looks undesirable. So we restrict such arguments
to have a subtype of the collection element type.\<close>
text \<open>
Please take a note that we allow the following expressions:\<close>
text \<open>
\<^verbatim>\<open>let nullable_value : Integer[?] = null in
Sequence{1..3}->inculdes(nullable_value) and
Sequence{1..3}->inculdes(null) and
Sequence{1..3}->inculdesAll(Set{1,null})\<close>\<close>
text \<open>
The OCL specification defines @{text "including()"} and
@{text "excluding()"} operations for the @{text Sequence} type
but does not define them for the @{text OrderedSet} type.
We define them for all collection types.
It is a good idea to prohibit including of values that
do not conform to a collection element type.
However we do not restrict it.
At first we defined the following typing rules for the
@{text "excluding()"} operation:
{\isacharbar}\ {\isachardoublequoteopen}element{\isacharunderscore}type\ {\isasymtau}\ {\isasymrho}\ {\isasymLongrightarrow}\ {\isasymsigma}\ {\isasymle}\ {\isasymrho}\ {\isasymLongrightarrow}\ {\isasymsigma}\ {\isasymnoteq}\ OclVoid{\isacharbrackleft}{\isacharquery}{\isacharbrackright}\ {\isasymLongrightarrow}\isanewline
\ \ \ collection{\isacharunderscore}binop{\isacharunderscore}type\ ExcludingOp\ {\isasymtau}\ {\isasymsigma}\ {\isasymtau}{\isachardoublequoteclose}\isanewline
{\isacharbar}\ {\isachardoublequoteopen}element{\isacharunderscore}type\ {\isasymtau}\ {\isasymrho}\ {\isasymLongrightarrow}\ {\isasymsigma}\ {\isasymle}\ {\isasymrho}\ {\isasymLongrightarrow}\ {\isasymsigma}\ {\isacharequal}\ OclVoid{\isacharbrackleft}{\isacharquery}{\isacharbrackright}\ {\isasymLongrightarrow}\isanewline
\ \ \ update{\isacharunderscore}element{\isacharunderscore}type\ {\isasymtau}\ {\isacharparenleft}to{\isacharunderscore}required{\isacharunderscore}type\ {\isasymrho}{\isacharparenright}\ {\isasymupsilon}\ {\isasymLongrightarrow}\isanewline
\ \ \ collection{\isacharunderscore}binop{\isacharunderscore}type\ ExcludingOp\ {\isasymtau}\ {\isasymsigma}\ {\isasymupsilon}{\isachardoublequoteclose}\isanewline
This operation could play a special role in a definition
of safe navigation operations:\<close>
text \<open>
\<^verbatim>\<open>Sequence{1,2,null}->exculding(null) : Integer[1]\<close>\<close>
text \<open>
However it is more natural to use a @{text "selectByKind(T[1])"}
operation instead.\<close>
inductive collection_binop_type where
"element_type \<tau> \<rho> \<Longrightarrow> \<sigma> \<le> to_optional_type_nested \<rho> \<Longrightarrow>
collection_binop_type IncludesOp \<tau> \<sigma> Boolean[1]"
| "element_type \<tau> \<rho> \<Longrightarrow> \<sigma> \<le> to_optional_type_nested \<rho> \<Longrightarrow>
collection_binop_type ExcludesOp \<tau> \<sigma> Boolean[1]"
| "element_type \<tau> \<rho> \<Longrightarrow> \<sigma> \<le> to_optional_type_nested \<rho> \<Longrightarrow>
collection_binop_type CountOp \<tau> \<sigma> Integer[1]"
| "element_type \<tau> \<rho> \<Longrightarrow> element_type \<sigma> \<upsilon> \<Longrightarrow> \<upsilon> \<le> to_optional_type_nested \<rho> \<Longrightarrow>
collection_binop_type IncludesAllOp \<tau> \<sigma> Boolean[1]"
| "element_type \<tau> \<rho> \<Longrightarrow> element_type \<sigma> \<upsilon> \<Longrightarrow> \<upsilon> \<le> to_optional_type_nested \<rho> \<Longrightarrow>
collection_binop_type ExcludesAllOp \<tau> \<sigma> Boolean[1]"
| "element_type \<tau> \<rho> \<Longrightarrow> element_type \<sigma> \<upsilon> \<Longrightarrow>
collection_binop_type ProductOp \<tau> \<sigma>
(Set (Tuple (fmap_of_list [(STR ''first'', \<rho>), (STR ''second'', \<upsilon>)])))"
| "collection_binop_type UnionOp (Set \<tau>) (Set \<sigma>) (Set (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type UnionOp (Set \<tau>) (Bag \<sigma>) (Bag (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type UnionOp (Bag \<tau>) (Set \<sigma>) (Bag (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type UnionOp (Bag \<tau>) (Bag \<sigma>) (Bag (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type IntersectionOp (Set \<tau>) (Set \<sigma>) (Set (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type IntersectionOp (Set \<tau>) (Bag \<sigma>) (Set (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type IntersectionOp (Bag \<tau>) (Set \<sigma>) (Set (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type IntersectionOp (Bag \<tau>) (Bag \<sigma>) (Bag (\<tau> \<squnion> \<sigma>))"
| "collection_binop_type SetMinusOp (Set \<tau>) (Set \<sigma>) (Set \<tau>)"
| "collection_binop_type SymmetricDifferenceOp (Set \<tau>) (Set \<sigma>) (Set (\<tau> \<squnion> \<sigma>))"
| "element_type \<tau> \<rho> \<Longrightarrow> update_element_type \<tau> (\<rho> \<squnion> \<sigma>) \<upsilon> \<Longrightarrow>
collection_binop_type IncludingOp \<tau> \<sigma> \<upsilon>"
| "element_type \<tau> \<rho> \<Longrightarrow> \<sigma> \<le> \<rho> \<Longrightarrow>
collection_binop_type ExcludingOp \<tau> \<sigma> \<tau>"
| "\<sigma> \<le> \<tau> \<Longrightarrow>
collection_binop_type AppendOp (OrderedSet \<tau>) \<sigma> (OrderedSet \<tau>)"
| "\<sigma> \<le> \<tau> \<Longrightarrow>
collection_binop_type AppendOp (Sequence \<tau>) \<sigma> (Sequence \<tau>)"
| "\<sigma> \<le> \<tau> \<Longrightarrow>
collection_binop_type PrependOp (OrderedSet \<tau>) \<sigma> (OrderedSet \<tau>)"
| "\<sigma> \<le> \<tau> \<Longrightarrow>
collection_binop_type PrependOp (Sequence \<tau>) \<sigma> (Sequence \<tau>)"
| "\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
collection_binop_type CollectionAtOp (OrderedSet \<tau>) \<sigma> \<tau>"
| "\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
collection_binop_type CollectionAtOp (Sequence \<tau>) \<sigma> \<tau>"
| "\<sigma> \<le> \<tau> \<Longrightarrow>
collection_binop_type CollectionIndexOfOp (OrderedSet \<tau>) \<sigma> Integer[1]"
| "\<sigma> \<le> \<tau> \<Longrightarrow>
collection_binop_type CollectionIndexOfOp (Sequence \<tau>) \<sigma> Integer[1]"
inductive collection_ternop_type where
"\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow> \<rho> \<le> \<tau> \<Longrightarrow>
collection_ternop_type InsertAtOp (OrderedSet \<tau>) \<sigma> \<rho> (OrderedSet \<tau>)"
| "\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow> \<rho> \<le> \<tau> \<Longrightarrow>
collection_ternop_type InsertAtOp (Sequence \<tau>) \<sigma> \<rho> (Sequence \<tau>)"
| "\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
\<rho> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
collection_ternop_type SubOrderedSetOp (OrderedSet \<tau>) \<sigma> \<rho> (OrderedSet \<tau>)"
| "\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
\<rho> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
collection_ternop_type SubSequenceOp (Sequence \<tau>) \<sigma> \<rho> (Sequence \<tau>)"
subsection \<open>Coercions\<close>
inductive unop_type where
"any_unop_type op \<tau> \<sigma> \<Longrightarrow>
unop_type (Inl op) DotCall \<tau> \<sigma>"
| "boolean_unop_type op \<tau> \<sigma> \<Longrightarrow>
unop_type (Inr (Inl op)) DotCall \<tau> \<sigma>"
| "numeric_unop_type op \<tau> \<sigma> \<Longrightarrow>
unop_type (Inr (Inr (Inl op))) DotCall \<tau> \<sigma>"
| "string_unop_type op \<tau> \<sigma> \<Longrightarrow>
unop_type (Inr (Inr (Inr (Inl op)))) DotCall \<tau> \<sigma>"
| "collection_unop_type op \<tau> \<sigma> \<Longrightarrow>
unop_type (Inr (Inr (Inr (Inr op)))) ArrowCall \<tau> \<sigma>"
inductive binop_type where
"super_binop_type op \<tau> \<sigma> \<rho> \<Longrightarrow>
binop_type (Inl op) DotCall \<tau> \<sigma> \<rho>"
| "boolean_binop_type op \<tau> \<sigma> \<rho> \<Longrightarrow>
binop_type (Inr (Inl op)) DotCall \<tau> \<sigma> \<rho>"
| "numeric_binop_type op \<tau> \<sigma> \<rho> \<Longrightarrow>
binop_type (Inr (Inr (Inl op))) DotCall \<tau> \<sigma> \<rho>"
| "string_binop_type op \<tau> \<sigma> \<rho> \<Longrightarrow>
binop_type (Inr (Inr (Inr (Inl op)))) DotCall \<tau> \<sigma> \<rho>"
| "collection_binop_type op \<tau> \<sigma> \<rho> \<Longrightarrow>
binop_type (Inr (Inr (Inr (Inr op)))) ArrowCall \<tau> \<sigma> \<rho>"
inductive ternop_type where
"string_ternop_type op \<tau> \<sigma> \<rho> \<upsilon> \<Longrightarrow>
ternop_type (Inl op) DotCall \<tau> \<sigma> \<rho> \<upsilon>"
| "collection_ternop_type op \<tau> \<sigma> \<rho> \<upsilon> \<Longrightarrow>
ternop_type (Inr op) ArrowCall \<tau> \<sigma> \<rho> \<upsilon>"
inductive op_type where
"unop_type op k \<tau> \<upsilon> \<Longrightarrow>
op_type (Inl op) k \<tau> [] \<upsilon>"
| "binop_type op k \<tau> \<sigma> \<upsilon> \<Longrightarrow>
op_type (Inr (Inl op)) k \<tau> [\<sigma>] \<upsilon>"
| "ternop_type op k \<tau> \<sigma> \<rho> \<upsilon> \<Longrightarrow>
op_type (Inr (Inr (Inl op))) k \<tau> [\<sigma>, \<rho>] \<upsilon>"
| "operation \<tau> op \<pi> oper \<Longrightarrow>
op_type (Inr (Inr (Inr op))) DotCall \<tau> \<pi> (oper_type oper)"
(*** Simplification Rules ***************************************************)
subsection \<open>Simplification Rules\<close>
inductive_simps op_type_alt_simps:
"mataop_type \<tau> op \<sigma>"
"typeop_type k op \<tau> \<sigma> \<rho>"
"op_type op k \<tau> \<pi> \<sigma>"
"unop_type op k \<tau> \<sigma>"
"binop_type op k \<tau> \<sigma> \<rho>"
"ternop_type op k \<tau> \<sigma> \<rho> \<upsilon>"
"any_unop_type op \<tau> \<sigma>"
"boolean_unop_type op \<tau> \<sigma>"
"numeric_unop_type op \<tau> \<sigma>"
"string_unop_type op \<tau> \<sigma>"
"collection_unop_type op \<tau> \<sigma>"
"super_binop_type op \<tau> \<sigma> \<rho>"
"boolean_binop_type op \<tau> \<sigma> \<rho>"
"numeric_binop_type op \<tau> \<sigma> \<rho>"
"string_binop_type op \<tau> \<sigma> \<rho>"
"collection_binop_type op \<tau> \<sigma> \<rho>"
"string_ternop_type op \<tau> \<sigma> \<rho> \<upsilon>"
"collection_ternop_type op \<tau> \<sigma> \<rho> \<upsilon>"
(*** Determinism ************************************************************)
subsection \<open>Determinism\<close>
lemma typeop_type_det:
"typeop_type op k \<tau> \<sigma> \<rho>\<^sub>1 \<Longrightarrow>
typeop_type op k \<tau> \<sigma> \<rho>\<^sub>2 \<Longrightarrow> \<rho>\<^sub>1 = \<rho>\<^sub>2"
by (induct rule: typeop_type.induct;
auto simp add: typeop_type.simps update_element_type_det)
lemma any_unop_type_det:
"any_unop_type op \<tau> \<sigma>\<^sub>1 \<Longrightarrow>
any_unop_type op \<tau> \<sigma>\<^sub>2 \<Longrightarrow> \<sigma>\<^sub>1 = \<sigma>\<^sub>2"
by (induct rule: any_unop_type.induct; simp add: any_unop_type.simps)
lemma boolean_unop_type_det:
"boolean_unop_type op \<tau> \<sigma>\<^sub>1 \<Longrightarrow>
boolean_unop_type op \<tau> \<sigma>\<^sub>2 \<Longrightarrow> \<sigma>\<^sub>1 = \<sigma>\<^sub>2"
by (induct rule: boolean_unop_type.induct; simp add: boolean_unop_type.simps)
lemma numeric_unop_type_det:
"numeric_unop_type op \<tau> \<sigma>\<^sub>1 \<Longrightarrow>
numeric_unop_type op \<tau> \<sigma>\<^sub>2 \<Longrightarrow> \<sigma>\<^sub>1 = \<sigma>\<^sub>2"
by (induct rule: numeric_unop_type.induct; auto simp add: numeric_unop_type.simps)
lemma string_unop_type_det:
"string_unop_type op \<tau> \<sigma>\<^sub>1 \<Longrightarrow>
string_unop_type op \<tau> \<sigma>\<^sub>2 \<Longrightarrow> \<sigma>\<^sub>1 = \<sigma>\<^sub>2"
by (induct rule: string_unop_type.induct; simp add: string_unop_type.simps)
lemma collection_unop_type_det:
"collection_unop_type op \<tau> \<sigma>\<^sub>1 \<Longrightarrow>
collection_unop_type op \<tau> \<sigma>\<^sub>2 \<Longrightarrow> \<sigma>\<^sub>1 = \<sigma>\<^sub>2"
apply (induct rule: collection_unop_type.induct)
by (erule collection_unop_type.cases;
auto simp add: element_type_det update_element_type_det)+
lemma unop_type_det:
"unop_type op k \<tau> \<sigma>\<^sub>1 \<Longrightarrow>
unop_type op k \<tau> \<sigma>\<^sub>2 \<Longrightarrow> \<sigma>\<^sub>1 = \<sigma>\<^sub>2"
by (induct rule: unop_type.induct;
simp add: unop_type.simps any_unop_type_det
boolean_unop_type_det numeric_unop_type_det
string_unop_type_det collection_unop_type_det)
lemma super_binop_type_det:
"super_binop_type op \<tau> \<sigma> \<rho>\<^sub>1 \<Longrightarrow>
super_binop_type op \<tau> \<sigma> \<rho>\<^sub>2 \<Longrightarrow> \<rho>\<^sub>1 = \<rho>\<^sub>2"
by (induct rule: super_binop_type.induct; auto simp add: super_binop_type.simps)
lemma boolean_binop_type_det:
"boolean_binop_type op \<tau> \<sigma> \<rho>\<^sub>1 \<Longrightarrow>
boolean_binop_type op \<tau> \<sigma> \<rho>\<^sub>2 \<Longrightarrow> \<rho>\<^sub>1 = \<rho>\<^sub>2"
by (induct rule: boolean_binop_type.induct; simp add: boolean_binop_type.simps)
lemma numeric_binop_type_det:
"numeric_binop_type op \<tau> \<sigma> \<rho>\<^sub>1 \<Longrightarrow>
numeric_binop_type op \<tau> \<sigma> \<rho>\<^sub>2 \<Longrightarrow> \<rho>\<^sub>1 = \<rho>\<^sub>2"
by (induct rule: numeric_binop_type.induct;
auto simp add: numeric_binop_type.simps split: if_splits)
lemma string_binop_type_det:
"string_binop_type op \<tau> \<sigma> \<rho>\<^sub>1 \<Longrightarrow>
string_binop_type op \<tau> \<sigma> \<rho>\<^sub>2 \<Longrightarrow> \<rho>\<^sub>1 = \<rho>\<^sub>2"
by (induct rule: string_binop_type.induct; simp add: string_binop_type.simps)
lemma collection_binop_type_det:
"collection_binop_type op \<tau> \<sigma> \<rho>\<^sub>1 \<Longrightarrow>
collection_binop_type op \<tau> \<sigma> \<rho>\<^sub>2 \<Longrightarrow> \<rho>\<^sub>1 = \<rho>\<^sub>2"
apply (induct rule: collection_binop_type.induct; simp add: collection_binop_type.simps)
using element_type_det update_element_type_det by blast+
lemma binop_type_det:
"binop_type op k \<tau> \<sigma> \<rho>\<^sub>1 \<Longrightarrow>
binop_type op k \<tau> \<sigma> \<rho>\<^sub>2 \<Longrightarrow> \<rho>\<^sub>1 = \<rho>\<^sub>2"
by (induct rule: binop_type.induct;
simp add: binop_type.simps super_binop_type_det
boolean_binop_type_det numeric_binop_type_det
string_binop_type_det collection_binop_type_det)
lemma string_ternop_type_det:
"string_ternop_type op \<tau> \<sigma> \<rho> \<upsilon>\<^sub>1 \<Longrightarrow>
string_ternop_type op \<tau> \<sigma> \<rho> \<upsilon>\<^sub>2 \<Longrightarrow> \<upsilon>\<^sub>1 = \<upsilon>\<^sub>2"
by (induct rule: string_ternop_type.induct; simp add: string_ternop_type.simps)
lemma collection_ternop_type_det:
"collection_ternop_type op \<tau> \<sigma> \<rho> \<upsilon>\<^sub>1 \<Longrightarrow>
collection_ternop_type op \<tau> \<sigma> \<rho> \<upsilon>\<^sub>2 \<Longrightarrow> \<upsilon>\<^sub>1 = \<upsilon>\<^sub>2"
by (induct rule: collection_ternop_type.induct; simp add: collection_ternop_type.simps)
lemma ternop_type_det:
"ternop_type op k \<tau> \<sigma> \<rho> \<upsilon>\<^sub>1 \<Longrightarrow>
ternop_type op k \<tau> \<sigma> \<rho> \<upsilon>\<^sub>2 \<Longrightarrow> \<upsilon>\<^sub>1 = \<upsilon>\<^sub>2"
by (induct rule: ternop_type.induct;
simp add: ternop_type.simps string_ternop_type_det collection_ternop_type_det)
lemma op_type_det:
"op_type op k \<tau> \<pi> \<sigma> \<Longrightarrow>
op_type op k \<tau> \<pi> \<rho> \<Longrightarrow> \<sigma> = \<rho>"
apply (induct rule: op_type.induct)
apply (erule op_type.cases; simp add: unop_type_det)
apply (erule op_type.cases; simp add: binop_type_det)
apply (erule op_type.cases; simp add: ternop_type_det)
by (erule op_type.cases; simp; metis operation_det)
(*** Expressions Typing *****************************************************)
section \<open>Expressions Typing\<close>
text \<open>
The following typing rules are preliminary. The final rules are given at
the end of the next chapter.\<close>
inductive typing :: "('a :: ocl_object_model) type env \<Rightarrow> 'a expr \<Rightarrow> 'a type \<Rightarrow> bool"
("(1_/ \<turnstile>\<^sub>E/ (_ :/ _))" [51,51,51] 50)
and collection_parts_typing ("(1_/ \<turnstile>\<^sub>C/ (_ :/ _))" [51,51,51] 50)
and collection_part_typing ("(1_/ \<turnstile>\<^sub>P/ (_ :/ _))" [51,51,51] 50)
and iterator_typing ("(1_/ \<turnstile>\<^sub>I/ (_ :/ _))" [51,51,51] 50)
and expr_list_typing ("(1_/ \<turnstile>\<^sub>L/ (_ :/ _))" [51,51,51] 50) where
\<comment> \<open>Primitive Literals\<close>
NullLiteralT:
"\<Gamma> \<turnstile>\<^sub>E NullLiteral : OclVoid[?]"
|BooleanLiteralT:
"\<Gamma> \<turnstile>\<^sub>E BooleanLiteral c : Boolean[1]"
|RealLiteralT:
"\<Gamma> \<turnstile>\<^sub>E RealLiteral c : Real[1]"
|IntegerLiteralT:
"\<Gamma> \<turnstile>\<^sub>E IntegerLiteral c : Integer[1]"
|UnlimitedNaturalLiteralT:
"\<Gamma> \<turnstile>\<^sub>E UnlimitedNaturalLiteral c : UnlimitedNatural[1]"
|StringLiteralT:
"\<Gamma> \<turnstile>\<^sub>E StringLiteral c : String[1]"
|EnumLiteralT:
"has_literal enum lit \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E EnumLiteral enum lit : (Enum enum)[1]"
\<comment> \<open>Collection Literals\<close>
|SetLiteralT:
"\<Gamma> \<turnstile>\<^sub>C prts : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E CollectionLiteral SetKind prts : Set \<tau>"
|OrderedSetLiteralT:
"\<Gamma> \<turnstile>\<^sub>C prts : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E CollectionLiteral OrderedSetKind prts : OrderedSet \<tau>"
|BagLiteralT:
"\<Gamma> \<turnstile>\<^sub>C prts : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E CollectionLiteral BagKind prts : Bag \<tau>"
|SequenceLiteralT:
"\<Gamma> \<turnstile>\<^sub>C prts : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E CollectionLiteral SequenceKind prts : Sequence \<tau>"
\<comment> \<open>We prohibit empty collection literals, because their type is unclear.
We could use @{text "OclVoid[1]"} element type for empty collections, but
the typing rules will give wrong types for nested collections, because,
for example, @{text "OclVoid[1] \<squnion> Set(Integer[1]) = OclSuper"}\<close>
|CollectionPartsSingletonT:
"\<Gamma> \<turnstile>\<^sub>P x : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>C [x] : \<tau>"
|CollectionPartsListT:
"\<Gamma> \<turnstile>\<^sub>P x : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>C y # xs : \<sigma> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>C x # y # xs : \<tau> \<squnion> \<sigma>"
|CollectionPartItemT:
"\<Gamma> \<turnstile>\<^sub>E a : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>P CollectionItem a : \<tau>"
|CollectionPartRangeT:
"\<Gamma> \<turnstile>\<^sub>E a : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E b : \<sigma> \<Longrightarrow>
\<tau> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
\<sigma> = UnlimitedNatural[1]\<midarrow>Integer[1] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>P CollectionRange a b : Integer[1]"
\<comment> \<open>Tuple Literals\<close>
\<comment> \<open>We do not prohibit empty tuples, because it could be useful.
@{text "Tuple()"} is a supertype of all other tuple types.\<close>
|TupleLiteralNilT:
"\<Gamma> \<turnstile>\<^sub>E TupleLiteral [] : Tuple fmempty"
|TupleLiteralConsT:
"\<Gamma> \<turnstile>\<^sub>E TupleLiteral elems : Tuple \<xi> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E tuple_element_expr el : \<tau> \<Longrightarrow>
tuple_element_type el = Some \<sigma> \<Longrightarrow>
\<tau> \<le> \<sigma> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E TupleLiteral (el # elems) : Tuple (\<xi>(tuple_element_name el \<mapsto>\<^sub>f \<sigma>))"
\<comment> \<open>Misc Expressions\<close>
|LetT:
"\<Gamma> \<turnstile>\<^sub>E init : \<sigma> \<Longrightarrow>
\<sigma> \<le> \<tau> \<Longrightarrow>
\<Gamma>(v \<mapsto>\<^sub>f \<tau>) \<turnstile>\<^sub>E body : \<rho> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E Let v (Some \<tau>) init body : \<rho>"
|VarT:
"fmlookup \<Gamma> v = Some \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E Var v : \<tau>"
|IfT:
"\<Gamma> \<turnstile>\<^sub>E a : Boolean[1] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E b : \<sigma> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E c : \<rho> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E If a b c : \<sigma> \<squnion> \<rho>"
\<comment> \<open>Call Expressions\<close>
|MetaOperationCallT:
"mataop_type \<tau> op \<sigma> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E MetaOperationCall \<tau> op : \<sigma>"
|StaticOperationCallT:
"\<Gamma> \<turnstile>\<^sub>L params : \<pi> \<Longrightarrow>
static_operation \<tau> op \<pi> oper \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E StaticOperationCall \<tau> op params : oper_type oper"
|TypeOperationCallT:
"\<Gamma> \<turnstile>\<^sub>E a : \<tau> \<Longrightarrow>
typeop_type k op \<tau> \<sigma> \<rho> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E TypeOperationCall a k op \<sigma> : \<rho>"
|AttributeCallT:
"\<Gamma> \<turnstile>\<^sub>E src : \<langle>\<C>\<rangle>\<^sub>\<T>[1] \<Longrightarrow>
attribute \<C> prop \<D> \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E AttributeCall src DotCall prop : \<tau>"
|AssociationEndCallT:
"\<Gamma> \<turnstile>\<^sub>E src : \<langle>\<C>\<rangle>\<^sub>\<T>[1] \<Longrightarrow>
association_end \<C> from role \<D> end \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E AssociationEndCall src DotCall from role : assoc_end_type end"
|AssociationClassCallT:
"\<Gamma> \<turnstile>\<^sub>E src : \<langle>\<C>\<rangle>\<^sub>\<T>[1] \<Longrightarrow>
referred_by_association_class \<C> from \<A> \<D> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E AssociationClassCall src DotCall from \<A> : class_assoc_type \<A>"
|AssociationClassEndCallT:
"\<Gamma> \<turnstile>\<^sub>E src : \<langle>\<A>\<rangle>\<^sub>\<T>[1] \<Longrightarrow>
association_class_end \<A> role end \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E AssociationClassEndCall src DotCall role : class_assoc_end_type end"
|OperationCallT:
"\<Gamma> \<turnstile>\<^sub>E src : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>L params : \<pi> \<Longrightarrow>
op_type op k \<tau> \<pi> \<sigma> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E OperationCall src k op params : \<sigma>"
|TupleElementCallT:
"\<Gamma> \<turnstile>\<^sub>E src : Tuple \<pi> \<Longrightarrow>
fmlookup \<pi> elem = Some \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E TupleElementCall src DotCall elem : \<tau>"
\<comment> \<open>Iterator Expressions\<close>
|IteratorT:
"\<Gamma> \<turnstile>\<^sub>E src : \<tau> \<Longrightarrow>
element_type \<tau> \<sigma> \<Longrightarrow>
\<sigma> \<le> its_ty \<Longrightarrow>
\<Gamma> ++\<^sub>f fmap_of_list (map (\<lambda>it. (it, its_ty)) its) \<turnstile>\<^sub>E body : \<rho> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>I (src, its, (Some its_ty), body) : (\<tau>, \<sigma>, \<rho>)"
|IterateT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, Let res (Some res_t) res_init body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
\<rho> \<le> res_t \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E IterateCall src ArrowCall its its_ty res (Some res_t) res_init body : \<rho>"
|AnyIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
\<rho> \<le> Boolean[?] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E AnyIteratorCall src ArrowCall its its_ty body : \<sigma>"
|ClosureIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
to_single_type \<rho> \<le> \<sigma> \<Longrightarrow>
to_unique_collection \<tau> \<upsilon> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E ClosureIteratorCall src ArrowCall its its_ty body : \<upsilon>"
|CollectIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
to_nonunique_collection \<tau> \<upsilon> \<Longrightarrow>
update_element_type \<upsilon> (to_single_type \<rho>) \<phi> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E CollectIteratorCall src ArrowCall its its_ty body : \<phi>"
|CollectNestedIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
to_nonunique_collection \<tau> \<upsilon> \<Longrightarrow>
update_element_type \<upsilon> \<rho> \<phi> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E CollectNestedIteratorCall src ArrowCall its its_ty body : \<phi>"
|ExistsIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
\<rho> \<le> Boolean[?] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E ExistsIteratorCall src ArrowCall its its_ty body : \<rho>"
|ForAllIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
\<rho> \<le> Boolean[?] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E ForAllIteratorCall src ArrowCall its its_ty body : \<rho>"
|OneIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
\<rho> \<le> Boolean[?] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E OneIteratorCall src ArrowCall its its_ty body : Boolean[1]"
|IsUniqueIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E IsUniqueIteratorCall src ArrowCall its its_ty body : Boolean[1]"
|SelectIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
\<rho> \<le> Boolean[?] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E SelectIteratorCall src ArrowCall its its_ty body : \<tau>"
|RejectIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
\<rho> \<le> Boolean[?] \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E RejectIteratorCall src ArrowCall its its_ty body : \<tau>"
|SortedByIteratorT:
"\<Gamma> \<turnstile>\<^sub>I (src, its, its_ty, body) : (\<tau>, \<sigma>, \<rho>) \<Longrightarrow>
length its \<le> 1 \<Longrightarrow>
to_ordered_collection \<tau> \<upsilon> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E SortedByIteratorCall src ArrowCall its its_ty body : \<upsilon>"
\<comment> \<open>Expression Lists\<close>
|ExprListNilT:
"\<Gamma> \<turnstile>\<^sub>L [] : []"
|ExprListConsT:
"\<Gamma> \<turnstile>\<^sub>E expr : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>L exprs : \<pi> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>L expr # exprs : \<tau> # \<pi>"
(*** Elimination Rules ******************************************************)
section \<open>Elimination Rules\<close>
inductive_cases NullLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E NullLiteral : \<tau>"
inductive_cases BooleanLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E BooleanLiteral c : \<tau>"
inductive_cases RealLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E RealLiteral c : \<tau>"
inductive_cases IntegerLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E IntegerLiteral c : \<tau>"
inductive_cases UnlimitedNaturalLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E UnlimitedNaturalLiteral c : \<tau>"
inductive_cases StringLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E StringLiteral c : \<tau>"
inductive_cases EnumLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E EnumLiteral enm lit : \<tau>"
inductive_cases CollectionLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E CollectionLiteral k prts : \<tau>"
inductive_cases TupleLiteralTE [elim]: "\<Gamma> \<turnstile>\<^sub>E TupleLiteral elems : \<tau>"
inductive_cases LetTE [elim]: "\<Gamma> \<turnstile>\<^sub>E Let v \<tau> init body : \<sigma>"
inductive_cases VarTE [elim]: "\<Gamma> \<turnstile>\<^sub>E Var v : \<tau>"
inductive_cases IfTE [elim]: "\<Gamma> \<turnstile>\<^sub>E If a b c : \<tau>"
inductive_cases MetaOperationCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E MetaOperationCall \<tau> op : \<sigma>"
inductive_cases StaticOperationCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E StaticOperationCall \<tau> op as : \<sigma>"
inductive_cases TypeOperationCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E TypeOperationCall a k op \<sigma> : \<tau>"
inductive_cases AttributeCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E AttributeCall src k prop : \<tau>"
inductive_cases AssociationEndCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E AssociationEndCall src k role from : \<tau>"
inductive_cases AssociationClassCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E AssociationClassCall src k a from : \<tau>"
inductive_cases AssociationClassEndCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E AssociationClassEndCall src k role : \<tau>"
inductive_cases OperationCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E OperationCall src k op params : \<tau>"
inductive_cases TupleElementCallTE [elim]: "\<Gamma> \<turnstile>\<^sub>E TupleElementCall src k elem : \<tau>"
inductive_cases IteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>I (src, its, body) : ys"
inductive_cases IterateTE [elim]: "\<Gamma> \<turnstile>\<^sub>E IterateCall src k its its_ty res res_t res_init body : \<tau>"
inductive_cases AnyIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E AnyIteratorCall src k its its_ty body : \<tau>"
inductive_cases ClosureIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E ClosureIteratorCall src k its its_ty body : \<tau>"
inductive_cases CollectIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E CollectIteratorCall src k its its_ty body : \<tau>"
inductive_cases CollectNestedIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E CollectNestedIteratorCall src k its its_ty body : \<tau>"
inductive_cases ExistsIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E ExistsIteratorCall src k its its_ty body : \<tau>"
inductive_cases ForAllIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E ForAllIteratorCall src k its its_ty body : \<tau>"
inductive_cases OneIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E OneIteratorCall src k its its_ty body : \<tau>"
inductive_cases IsUniqueIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E IsUniqueIteratorCall src k its its_ty body : \<tau>"
inductive_cases SelectIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E SelectIteratorCall src k its its_ty body : \<tau>"
inductive_cases RejectIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E RejectIteratorCall src k its its_ty body : \<tau>"
inductive_cases SortedByIteratorTE [elim]: "\<Gamma> \<turnstile>\<^sub>E SortedByIteratorCall src k its its_ty body : \<tau>"
inductive_cases CollectionPartsNilTE [elim]: "\<Gamma> \<turnstile>\<^sub>C [x] : \<tau>"
inductive_cases CollectionPartsItemTE [elim]: "\<Gamma> \<turnstile>\<^sub>C x # y # xs : \<tau>"
inductive_cases CollectionItemTE [elim]: "\<Gamma> \<turnstile>\<^sub>P CollectionItem a : \<tau>"
inductive_cases CollectionRangeTE [elim]: "\<Gamma> \<turnstile>\<^sub>P CollectionRange a b : \<tau>"
inductive_cases ExprListTE [elim]: "\<Gamma> \<turnstile>\<^sub>L exprs : \<pi>"
(*** Simplification Rules ***************************************************)
section \<open>Simplification Rules\<close>
inductive_simps typing_alt_simps:
"\<Gamma> \<turnstile>\<^sub>E NullLiteral : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E BooleanLiteral c : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E RealLiteral c : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E UnlimitedNaturalLiteral c : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E IntegerLiteral c : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E StringLiteral c : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E EnumLiteral enm lit : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E CollectionLiteral k prts : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E TupleLiteral [] : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E TupleLiteral (x # xs) : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E Let v \<tau> init body : \<sigma>"
"\<Gamma> \<turnstile>\<^sub>E Var v : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E If a b c : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E MetaOperationCall \<tau> op : \<sigma>"
"\<Gamma> \<turnstile>\<^sub>E StaticOperationCall \<tau> op as : \<sigma>"
"\<Gamma> \<turnstile>\<^sub>E TypeOperationCall a k op \<sigma> : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E AttributeCall src k prop : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E AssociationEndCall src k role from : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E AssociationClassCall src k a from : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E AssociationClassEndCall src k role : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E OperationCall src k op params : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E TupleElementCall src k elem : \<tau>"
"\<Gamma> \<turnstile>\<^sub>I (src, its, body) : ys"
"\<Gamma> \<turnstile>\<^sub>E IterateCall src k its its_ty res res_t res_init body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E AnyIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E ClosureIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E CollectIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E CollectNestedIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E ExistsIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E ForAllIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E OneIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E IsUniqueIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E SelectIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E RejectIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>E SortedByIteratorCall src k its its_ty body : \<tau>"
"\<Gamma> \<turnstile>\<^sub>C [x] : \<tau>"
"\<Gamma> \<turnstile>\<^sub>C x # y # xs : \<tau>"
"\<Gamma> \<turnstile>\<^sub>P CollectionItem a : \<tau>"
"\<Gamma> \<turnstile>\<^sub>P CollectionRange a b : \<tau>"
"\<Gamma> \<turnstile>\<^sub>L [] : \<pi>"
"\<Gamma> \<turnstile>\<^sub>L x # xs : \<pi>"
(*** Determinism ************************************************************)
section \<open>Determinism\<close>
lemma
typing_det:
"\<Gamma> \<turnstile>\<^sub>E expr : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>E expr : \<sigma> \<Longrightarrow> \<tau> = \<sigma>" and
collection_parts_typing_det:
"\<Gamma> \<turnstile>\<^sub>C prts : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>C prts : \<sigma> \<Longrightarrow> \<tau> = \<sigma>" and
collection_part_typing_det:
"\<Gamma> \<turnstile>\<^sub>P prt : \<tau> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>P prt : \<sigma> \<Longrightarrow> \<tau> = \<sigma>" and
iterator_typing_det:
"\<Gamma> \<turnstile>\<^sub>I (src, its, body) : xs \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>I (src, its, body) : ys \<Longrightarrow> xs = ys" and
expr_list_typing_det:
"\<Gamma> \<turnstile>\<^sub>L exprs : \<pi> \<Longrightarrow>
\<Gamma> \<turnstile>\<^sub>L exprs : \<xi> \<Longrightarrow> \<pi> = \<xi>"
proof (induct arbitrary: \<sigma> and \<sigma> and \<sigma> and ys and \<xi>
rule: typing_collection_parts_typing_collection_part_typing_iterator_typing_expr_list_typing.inducts)
case (NullLiteralT \<Gamma>) thus ?case by auto
next
case (BooleanLiteralT \<Gamma> c) thus ?case by auto
next
case (RealLiteralT \<Gamma> c) thus ?case by auto
next
case (IntegerLiteralT \<Gamma> c) thus ?case by auto
next
case (UnlimitedNaturalLiteralT \<Gamma> c) thus ?case by auto
next
case (StringLiteralT \<Gamma> c) thus ?case by auto
next
case (EnumLiteralT \<Gamma> \<tau> lit) thus ?case by auto
next
case (SetLiteralT \<Gamma> prts \<tau>) thus ?case by blast
next
case (OrderedSetLiteralT \<Gamma> prts \<tau>) thus ?case by blast
next
case (BagLiteralT \<Gamma> prts \<tau>) thus ?case by blast
next
case (SequenceLiteralT \<Gamma> prts \<tau>) thus ?case by blast
next
case (CollectionPartsSingletonT \<Gamma> x \<tau>) thus ?case by blast
next
case (CollectionPartsListT \<Gamma> x \<tau> y xs \<sigma>) thus ?case by blast
next
case (CollectionPartItemT \<Gamma> a \<tau>) thus ?case by blast
next
case (CollectionPartRangeT \<Gamma> a \<tau> b \<sigma>) thus ?case by blast
next
case (TupleLiteralNilT \<Gamma>) thus ?case by auto
next
case (TupleLiteralConsT \<Gamma> elems \<xi> el \<tau>) show ?case
apply (insert TupleLiteralConsT.prems)
apply (erule TupleLiteralTE, simp)
using TupleLiteralConsT.hyps by auto
next
case (LetT \<Gamma> \<M> init \<sigma> \<tau> v body \<rho>) thus ?case by blast
next
case (VarT \<Gamma> v \<tau> \<M>) thus ?case by auto
next
case (IfT \<Gamma> a \<tau> b \<sigma> c \<rho>) thus ?case
apply (insert IfT.prems)
apply (erule IfTE)
by (simp add: IfT.hyps)
next
case (MetaOperationCallT \<tau> op \<sigma> \<Gamma>) thus ?case
by (metis MetaOperationCallTE mataop_type.cases)
next
case (StaticOperationCallT \<tau> op \<pi> oper \<Gamma> as) thus ?case
apply (insert StaticOperationCallT.prems)
apply (erule StaticOperationCallTE)
using StaticOperationCallT.hyps static_operation_det by blast
next
case (TypeOperationCallT \<Gamma> a \<tau> op \<sigma> \<rho>) thus ?case
by (metis TypeOperationCallTE typeop_type_det)
next
case (AttributeCallT \<Gamma> src \<tau> \<C> "prop" \<D> \<sigma>) show ?case
apply (insert AttributeCallT.prems)
apply (erule AttributeCallTE)
using AttributeCallT.hyps attribute_det by blast
next
case (AssociationEndCallT \<Gamma> src \<C> "from" role \<D> "end") show ?case
apply (insert AssociationEndCallT.prems)
apply (erule AssociationEndCallTE)
using AssociationEndCallT.hyps association_end_det by blast
next
case (AssociationClassCallT \<Gamma> src \<C> "from" \<A>) thus ?case by blast
next
case (AssociationClassEndCallT \<Gamma> src \<tau> \<A> role "end") show ?case
apply (insert AssociationClassEndCallT.prems)
apply (erule AssociationClassEndCallTE)
using AssociationClassEndCallT.hyps association_class_end_det by blast
next
case (OperationCallT \<Gamma> src \<tau> params \<pi> op k) show ?case
apply (insert OperationCallT.prems)
apply (erule OperationCallTE)
using OperationCallT.hyps op_type_det by blast
next
case (TupleElementCallT \<Gamma> src \<pi> elem \<tau>) thus ?case
apply (insert TupleElementCallT.prems)
apply (erule TupleElementCallTE)
using TupleElementCallT.hyps by fastforce
next
case (IteratorT \<Gamma> src \<tau> \<sigma> its_ty its body \<rho>) show ?case
apply (insert IteratorT.prems)
apply (erule IteratorTE)
using IteratorT.hyps element_type_det by blast
next
case (IterateT \<Gamma> src its its_ty res res_t res_init body \<tau> \<sigma> \<rho>) show ?case
apply (insert IterateT.prems)
using IterateT.hyps by blast
next
case (AnyIteratorT \<Gamma> src its its_ty body \<tau> \<sigma> \<rho>) thus ?case
by (meson AnyIteratorTE Pair_inject)
next
case (ClosureIteratorT \<Gamma> src its its_ty body \<tau> \<sigma> \<rho> \<upsilon>) show ?case
apply (insert ClosureIteratorT.prems)
apply (erule ClosureIteratorTE)
using ClosureIteratorT.hyps to_unique_collection_det by blast
next
case (CollectIteratorT \<Gamma> src its its_ty body \<tau> \<sigma> \<rho> \<upsilon>) show ?case
apply (insert CollectIteratorT.prems)
apply (erule CollectIteratorTE)
using CollectIteratorT.hyps to_nonunique_collection_det
update_element_type_det Pair_inject by metis
next
case (CollectNestedIteratorT \<Gamma> src its its_ty body \<tau> \<sigma> \<rho> \<upsilon>) show ?case
apply (insert CollectNestedIteratorT.prems)
apply (erule CollectNestedIteratorTE)
using CollectNestedIteratorT.hyps to_nonunique_collection_det
update_element_type_det Pair_inject by metis
next
case (ExistsIteratorT \<Gamma> src its its_ty body \<tau> \<sigma> \<rho>) show ?case
apply (insert ExistsIteratorT.prems)
apply (erule ExistsIteratorTE)
using ExistsIteratorT.hyps Pair_inject by metis
next
case (ForAllIteratorT \<Gamma> \<M> src its its_ty body \<tau> \<sigma> \<rho>) show ?case
apply (insert ForAllIteratorT.prems)
apply (erule ForAllIteratorTE)
using ForAllIteratorT.hyps Pair_inject by metis
next
case (OneIteratorT \<Gamma> \<M> src its its_ty body \<tau> \<sigma> \<rho>) show ?case
apply (insert OneIteratorT.prems)
apply (erule OneIteratorTE)
by simp
next
case (IsUniqueIteratorT \<Gamma> \<M> src its its_ty body \<tau> \<sigma> \<rho>) show ?case
apply (insert IsUniqueIteratorT.prems)
apply (erule IsUniqueIteratorTE)
by simp
next
case (SelectIteratorT \<Gamma> \<M> src its its_ty body \<tau> \<sigma> \<rho>) show ?case
apply (insert SelectIteratorT.prems)
apply (erule SelectIteratorTE)
using SelectIteratorT.hyps by blast
next
case (RejectIteratorT \<Gamma> \<M> src its its_ty body \<tau> \<sigma> \<rho>) show ?case
apply (insert RejectIteratorT.prems)
apply (erule RejectIteratorTE)
using RejectIteratorT.hyps by blast
next
case (SortedByIteratorT \<Gamma> \<M> src its its_ty body \<tau> \<sigma> \<rho> \<upsilon>) show ?case
apply (insert SortedByIteratorT.prems)
apply (erule SortedByIteratorTE)
using SortedByIteratorT.hyps to_ordered_collection_det by blast
next
case (ExprListNilT \<Gamma>) thus ?case
using expr_list_typing.cases by auto
next
case (ExprListConsT \<Gamma> expr \<tau> exprs \<pi>) show ?case
apply (insert ExprListConsT.prems)
apply (erule ExprListTE)
by (simp_all add: ExprListConsT.hyps)
qed
(*** Code Setup *************************************************************)
section \<open>Code Setup\<close>
code_pred op_type .
+
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) iterator_typing .
end
diff --git a/thys/Safe_OCL/Object_Model.thy b/thys/Safe_OCL/Object_Model.thy
--- a/thys/Safe_OCL/Object_Model.thy
+++ b/thys/Safe_OCL/Object_Model.thy
@@ -1,489 +1,503 @@
(* Title: Safe OCL
Author: Denis Nikiforov, March 2019
Maintainer: Denis Nikiforov <denis.nikif at gmail.com>
License: LGPL
*)
section \<open>Object Model\<close>
theory Object_Model
imports "HOL-Library.Extended_Nat" Finite_Map_Ext
begin
text \<open>
The section defines a generic object model.\<close>
(*** Type Synonyms **********************************************************)
subsection \<open>Type Synonyms\<close>
type_synonym attr = String.literal
type_synonym assoc = String.literal
type_synonym role = String.literal
type_synonym oper = String.literal
type_synonym param = String.literal
type_synonym elit = String.literal
datatype param_dir = In | Out | InOut
type_synonym 'c assoc_end = "'c \<times> nat \<times> enat \<times> bool \<times> bool"
type_synonym 't param_spec = "param \<times> 't \<times> param_dir"
type_synonym ('t, 'e) oper_spec =
"oper \<times> 't \<times> 't param_spec list \<times> 't \<times> bool \<times> 'e option"
definition "assoc_end_class :: 'c assoc_end \<Rightarrow> 'c \<equiv> fst"
definition "assoc_end_min :: 'c assoc_end \<Rightarrow> nat \<equiv> fst \<circ> snd"
definition "assoc_end_max :: 'c assoc_end \<Rightarrow> enat \<equiv> fst \<circ> snd \<circ> snd"
definition "assoc_end_ordered :: 'c assoc_end \<Rightarrow> bool \<equiv> fst \<circ> snd \<circ> snd \<circ> snd"
definition "assoc_end_unique :: 'c assoc_end \<Rightarrow> bool \<equiv> snd \<circ> snd \<circ> snd \<circ> snd"
definition "oper_name :: ('t, 'e) oper_spec \<Rightarrow> oper \<equiv> fst"
definition "oper_context :: ('t, 'e) oper_spec \<Rightarrow> 't \<equiv> fst \<circ> snd"
definition "oper_params :: ('t, 'e) oper_spec \<Rightarrow> 't param_spec list \<equiv> fst \<circ> snd \<circ> snd"
definition "oper_result :: ('t, 'e) oper_spec \<Rightarrow> 't \<equiv> fst \<circ> snd \<circ> snd \<circ> snd"
definition "oper_static :: ('t, 'e) oper_spec \<Rightarrow> bool \<equiv> fst \<circ> snd \<circ> snd \<circ> snd \<circ> snd"
definition "oper_body :: ('t, 'e) oper_spec \<Rightarrow> 'e option \<equiv> snd \<circ> snd \<circ> snd \<circ> snd \<circ> snd"
definition "param_name ::'t param_spec \<Rightarrow> param \<equiv> fst"
definition "param_type ::'t param_spec \<Rightarrow> 't \<equiv> fst \<circ> snd"
definition "param_dir ::'t param_spec \<Rightarrow> param_dir \<equiv> snd \<circ> snd"
definition "oper_in_params op \<equiv>
filter (\<lambda>p. param_dir p = In \<or> param_dir p = InOut) (oper_params op)"
definition "oper_out_params op \<equiv>
filter (\<lambda>p. param_dir p = Out \<or> param_dir p = InOut) (oper_params op)"
subsection \<open>Attributes\<close>
inductive owned_attribute' where
"\<C> |\<in>| fmdom attributes \<Longrightarrow>
fmlookup attributes \<C> = Some attrs\<^sub>\<C> \<Longrightarrow>
fmlookup attrs\<^sub>\<C> attr = Some \<tau> \<Longrightarrow>
owned_attribute' attributes \<C> attr \<tau>"
inductive attribute_not_closest where
"owned_attribute' attributes \<D>' attr \<tau>' \<Longrightarrow>
\<C> \<le> \<D>' \<Longrightarrow>
\<D>' < \<D> \<Longrightarrow>
attribute_not_closest attributes \<C> attr \<D> \<tau>"
inductive closest_attribute where
"owned_attribute' attributes \<D> attr \<tau> \<Longrightarrow>
\<C> \<le> \<D> \<Longrightarrow>
\<not> attribute_not_closest attributes \<C> attr \<D> \<tau> \<Longrightarrow>
closest_attribute attributes \<C> attr \<D> \<tau>"
inductive closest_attribute_not_unique where
"closest_attribute attributes \<C> attr \<D>' \<tau>' \<Longrightarrow>
\<D> \<noteq> \<D>' \<or> \<tau> \<noteq> \<tau>' \<Longrightarrow>
closest_attribute_not_unique attributes \<C> attr \<D> \<tau>"
inductive unique_closest_attribute where
"closest_attribute attributes \<C> attr \<D> \<tau> \<Longrightarrow>
\<not> closest_attribute_not_unique attributes \<C> attr \<D> \<tau> \<Longrightarrow>
unique_closest_attribute attributes \<C> attr \<D> \<tau>"
subsection \<open>Association Ends\<close>
inductive role_refer_class where
"role |\<in>| fmdom ends \<Longrightarrow>
fmlookup ends role = Some end \<Longrightarrow>
assoc_end_class end = \<C> \<Longrightarrow>
role_refer_class ends \<C> role"
inductive association_ends' where
"\<C> |\<in>| classes \<Longrightarrow>
assoc |\<in>| fmdom associations \<Longrightarrow>
fmlookup associations assoc = Some ends \<Longrightarrow>
role_refer_class ends \<C> from \<Longrightarrow>
role |\<in>| fmdom ends \<Longrightarrow>
fmlookup ends role = Some end \<Longrightarrow>
role \<noteq> from \<Longrightarrow>
association_ends' classes associations \<C> from role end"
inductive association_ends_not_unique' where
"association_ends' classes associations \<C> from role end\<^sub>1 \<Longrightarrow>
association_ends' classes associations \<C> from role end\<^sub>2 \<Longrightarrow>
end\<^sub>1 \<noteq> end\<^sub>2 \<Longrightarrow>
association_ends_not_unique' classes associations"
inductive owned_association_end' where
"association_ends' classes associations \<C> from role end \<Longrightarrow>
owned_association_end' classes associations \<C> None role end"
| "association_ends' classes associations \<C> from role end \<Longrightarrow>
owned_association_end' classes associations \<C> (Some from) role end"
inductive association_end_not_closest where
"owned_association_end' classes associations \<D>' from role end' \<Longrightarrow>
\<C> \<le> \<D>' \<Longrightarrow>
\<D>' < \<D> \<Longrightarrow>
association_end_not_closest classes associations \<C> from role \<D> end"
inductive closest_association_end where
"owned_association_end' classes associations \<D> from role end \<Longrightarrow>
\<C> \<le> \<D> \<Longrightarrow>
\<not> association_end_not_closest classes associations \<C> from role \<D> end \<Longrightarrow>
closest_association_end classes associations \<C> from role \<D> end"
inductive closest_association_end_not_unique where
"closest_association_end classes associations \<C> from role \<D>' end' \<Longrightarrow>
\<D> \<noteq> \<D>' \<or> end \<noteq> end' \<Longrightarrow>
closest_association_end_not_unique classes associations \<C> from role \<D> end"
inductive unique_closest_association_end where
"closest_association_end classes associations \<C> from role \<D> end \<Longrightarrow>
\<not> closest_association_end_not_unique classes associations \<C> from role \<D> end \<Longrightarrow>
unique_closest_association_end classes associations \<C> from role \<D> end"
subsection \<open>Association Classes\<close>
inductive referred_by_association_class'' where
"fmlookup association_classes \<A> = Some assoc \<Longrightarrow>
fmlookup associations assoc = Some ends \<Longrightarrow>
role_refer_class ends \<C> from \<Longrightarrow>
referred_by_association_class'' association_classes associations \<C> from \<A>"
inductive referred_by_association_class' where
"referred_by_association_class'' association_classes associations \<C> from \<A> \<Longrightarrow>
referred_by_association_class' association_classes associations \<C> None \<A>"
| "referred_by_association_class'' association_classes associations \<C> from \<A> \<Longrightarrow>
referred_by_association_class' association_classes associations \<C> (Some from) \<A>"
inductive association_class_not_closest where
"referred_by_association_class' association_classes associations \<D>' from \<A> \<Longrightarrow>
\<C> \<le> \<D>' \<Longrightarrow>
\<D>' < \<D> \<Longrightarrow>
association_class_not_closest association_classes associations \<C> from \<A> \<D>"
inductive closest_association_class where
"referred_by_association_class' association_classes associations \<D> from \<A> \<Longrightarrow>
\<C> \<le> \<D> \<Longrightarrow>
\<not> association_class_not_closest association_classes associations \<C> from \<A> \<D> \<Longrightarrow>
closest_association_class association_classes associations \<C> from \<A> \<D>"
inductive closest_association_class_not_unique where
"closest_association_class association_classes associations \<C> from \<A> \<D>' \<Longrightarrow>
\<D> \<noteq> \<D>' \<Longrightarrow>
closest_association_class_not_unique
association_classes associations \<C> from \<A> \<D>"
inductive unique_closest_association_class where
"closest_association_class association_classes associations \<C> from \<A> \<D> \<Longrightarrow>
\<not> closest_association_class_not_unique
association_classes associations \<C> from \<A> \<D> \<Longrightarrow>
unique_closest_association_class association_classes associations \<C> from \<A> \<D>"
subsection \<open>Association Class Ends\<close>
inductive association_class_end' where
"fmlookup association_classes \<A> = Some assoc \<Longrightarrow>
fmlookup associations assoc = Some ends \<Longrightarrow>
fmlookup ends role = Some end \<Longrightarrow>
association_class_end' association_classes associations \<A> role end"
inductive association_class_end_not_unique where
"association_class_end' association_classes associations \<A> role end' \<Longrightarrow>
end \<noteq> end' \<Longrightarrow>
association_class_end_not_unique association_classes associations \<A> role end"
inductive unique_association_class_end where
"association_class_end' association_classes associations \<A> role end \<Longrightarrow>
\<not> association_class_end_not_unique
association_classes associations \<A> role end \<Longrightarrow>
unique_association_class_end association_classes associations \<A> role end"
subsection \<open>Operations\<close>
inductive any_operation' where
"op |\<in>| fset_of_list operations \<Longrightarrow>
oper_name op = name \<Longrightarrow>
\<tau> \<le> oper_context op \<Longrightarrow>
list_all2 (\<lambda>\<sigma> param. \<sigma> \<le> param_type param) \<pi> (oper_in_params op) \<Longrightarrow>
any_operation' operations \<tau> name \<pi> op"
inductive operation' where
"any_operation' operations \<tau> name \<pi> op \<Longrightarrow>
\<not> oper_static op \<Longrightarrow>
operation' operations \<tau> name \<pi> op"
inductive operation_not_unique where
"operation' operations \<tau> name \<pi> oper' \<Longrightarrow>
oper \<noteq> oper' \<Longrightarrow>
operation_not_unique operations \<tau> name \<pi> oper"
inductive unique_operation where
"operation' operations \<tau> name \<pi> oper \<Longrightarrow>
\<not> operation_not_unique operations \<tau> name \<pi> oper \<Longrightarrow>
unique_operation operations \<tau> name \<pi> oper"
inductive operation_defined' where
"unique_operation operations \<tau> name \<pi> oper \<Longrightarrow>
operation_defined' operations \<tau> name \<pi>"
inductive static_operation' where
"any_operation' operations \<tau> name \<pi> op \<Longrightarrow>
oper_static op \<Longrightarrow>
static_operation' operations \<tau> name \<pi> op"
inductive static_operation_not_unique where
"static_operation' operations \<tau> name \<pi> oper' \<Longrightarrow>
oper \<noteq> oper' \<Longrightarrow>
static_operation_not_unique operations \<tau> name \<pi> oper"
inductive unique_static_operation where
"static_operation' operations \<tau> name \<pi> oper \<Longrightarrow>
\<not> static_operation_not_unique operations \<tau> name \<pi> oper \<Longrightarrow>
unique_static_operation operations \<tau> name \<pi> oper"
inductive static_operation_defined' where
"unique_static_operation operations \<tau> name \<pi> oper \<Longrightarrow>
static_operation_defined' operations \<tau> name \<pi>"
subsection \<open>Literals\<close>
inductive has_literal' where
"fmlookup literals e = Some lits \<Longrightarrow>
lit |\<in>| lits \<Longrightarrow>
has_literal' literals e lit"
(*** Definition *************************************************************)
subsection \<open>Definition\<close>
locale object_model =
fixes classes :: "'a :: semilattice_sup fset"
and attributes :: "'a \<rightharpoonup>\<^sub>f attr \<rightharpoonup>\<^sub>f 't :: order"
and associations :: "assoc \<rightharpoonup>\<^sub>f role \<rightharpoonup>\<^sub>f 'a assoc_end"
and association_classes :: "'a \<rightharpoonup>\<^sub>f assoc"
and operations :: "('t, 'e) oper_spec list"
and literals :: "'n \<rightharpoonup>\<^sub>f elit fset"
assumes assoc_end_min_less_eq_max:
"assoc |\<in>| fmdom associations \<Longrightarrow>
fmlookup associations assoc = Some ends \<Longrightarrow>
role |\<in>| fmdom ends \<Longrightarrow>
fmlookup ends role = Some end \<Longrightarrow>
assoc_end_min end \<le> assoc_end_max end"
assumes association_ends_unique:
"association_ends' classes associations \<C> from role end\<^sub>1 \<Longrightarrow>
association_ends' classes associations \<C> from role end\<^sub>2 \<Longrightarrow> end\<^sub>1 = end\<^sub>2"
begin
abbreviation "owned_attribute \<equiv>
owned_attribute' attributes"
abbreviation "attribute \<equiv>
unique_closest_attribute attributes"
abbreviation "association_ends \<equiv>
association_ends' classes associations"
abbreviation "owned_association_end \<equiv>
owned_association_end' classes associations"
abbreviation "association_end \<equiv>
unique_closest_association_end classes associations"
abbreviation "referred_by_association_class \<equiv>
unique_closest_association_class association_classes associations"
abbreviation "association_class_end \<equiv>
unique_association_class_end association_classes associations"
abbreviation "operation \<equiv>
unique_operation operations"
abbreviation "operation_defined \<equiv>
operation_defined' operations"
abbreviation "static_operation \<equiv>
unique_static_operation operations"
abbreviation "static_operation_defined \<equiv>
static_operation_defined' operations"
abbreviation "has_literal \<equiv>
has_literal' literals"
end
declare operation_defined'.simps [simp]
declare static_operation_defined'.simps [simp]
declare has_literal'.simps [simp]
(*** Properties *************************************************************)
subsection \<open>Properties\<close>
lemma (in object_model) attribute_det:
"attribute \<C> attr \<D>\<^sub>1 \<tau>\<^sub>1 \<Longrightarrow>
attribute \<C> attr \<D>\<^sub>2 \<tau>\<^sub>2 \<Longrightarrow> \<D>\<^sub>1 = \<D>\<^sub>2 \<and> \<tau>\<^sub>1 = \<tau>\<^sub>2"
by (meson closest_attribute_not_unique.intros unique_closest_attribute.cases)
lemma (in object_model) attribute_self_or_inherited:
"attribute \<C> attr \<D> \<tau> \<Longrightarrow> \<C> \<le> \<D>"
by (meson closest_attribute.cases unique_closest_attribute.cases)
lemma (in object_model) attribute_closest:
"attribute \<C> attr \<D> \<tau> \<Longrightarrow>
owned_attribute \<D>' attr \<tau> \<Longrightarrow>
\<C> \<le> \<D>' \<Longrightarrow> \<not> \<D>' < \<D>"
by (meson attribute_not_closest.intros closest_attribute.cases
unique_closest_attribute.cases)
lemma (in object_model) association_end_det:
"association_end \<C> from role \<D>\<^sub>1 end\<^sub>1 \<Longrightarrow>
association_end \<C> from role \<D>\<^sub>2 end\<^sub>2 \<Longrightarrow> \<D>\<^sub>1 = \<D>\<^sub>2 \<and> end\<^sub>1 = end\<^sub>2"
by (meson closest_association_end_not_unique.intros
unique_closest_association_end.cases)
lemma (in object_model) association_end_self_or_inherited:
"association_end \<C> from role \<D> end \<Longrightarrow> \<C> \<le> \<D>"
by (meson closest_association_end.cases unique_closest_association_end.cases)
lemma (in object_model) association_end_closest:
"association_end \<C> from role \<D> end \<Longrightarrow>
owned_association_end \<D>' from role end \<Longrightarrow>
\<C> \<le> \<D>' \<Longrightarrow> \<not> \<D>' < \<D>"
by (meson association_end_not_closest.intros closest_association_end.cases
unique_closest_association_end.cases)
lemma (in object_model) association_class_end_det:
"association_class_end \<A> role end\<^sub>1 \<Longrightarrow>
association_class_end \<A> role end\<^sub>2 \<Longrightarrow> end\<^sub>1 = end\<^sub>2"
by (meson association_class_end_not_unique.intros unique_association_class_end.cases)
lemma (in object_model) operation_det:
"operation \<tau> name \<pi> oper\<^sub>1 \<Longrightarrow>
operation \<tau> name \<pi> oper\<^sub>2 \<Longrightarrow> oper\<^sub>1 = oper\<^sub>2"
by (meson operation_not_unique.intros unique_operation.cases)
lemma (in object_model) static_operation_det:
"static_operation \<tau> name \<pi> oper\<^sub>1 \<Longrightarrow>
static_operation \<tau> name \<pi> oper\<^sub>2 \<Longrightarrow> oper\<^sub>1 = oper\<^sub>2"
by (meson static_operation_not_unique.intros unique_static_operation.cases)
(*** Code Setup *************************************************************)
subsection \<open>Code Setup\<close>
-lemma fmember_code_predI [code_pred_intro]:
- "x |\<in>| xs" if "Predicate_Compile.contains (fset xs) x"
- using that by (simp add: Predicate_Compile.contains_def fmember_iff_member_fset)
-
-code_pred fmember
- by (simp add: Predicate_Compile.contains_def fmember_iff_member_fset)
+declare owned_attribute'.intros[folded Predicate_Compile.contains_def, code_pred_intro]
+code_pred (modes:
+ i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
+ i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
+ i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
+ i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) owned_attribute'
+ by (elim owned_attribute'.cases) (simp add: Predicate_Compile.contains_def)
code_pred unique_closest_attribute .
+declare role_refer_class.intros[folded Predicate_Compile.contains_def, code_pred_intro]
+code_pred (modes:
+ i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
+ i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
+ i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
+ i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool) role_refer_class
+ by (elim role_refer_class.cases) (simp add: Predicate_Compile.contains_def)
+
+declare association_ends'.intros[folded Predicate_Compile.contains_def, code_pred_intro]
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
- i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool) association_ends' .
+ i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool) association_ends'
+ by (auto simp: Predicate_Compile.contains_def elim: association_ends'.cases)
code_pred association_ends_not_unique' .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool) owned_association_end' .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool) closest_association_end .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool ) closest_association_end_not_unique .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool) unique_closest_association_end .
code_pred unique_closest_association_class .
code_pred association_class_end' .
code_pred association_class_end_not_unique .
code_pred unique_association_class_end .
-code_pred (modes:
- i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
- i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) any_operation' .
+declare any_operation'.intros[folded Predicate_Compile.contains_def, code_pred_intro]
+code_pred [show_modes] any_operation'
+ by (elim any_operation'.cases) (simp add: Predicate_Compile.contains_def)
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) operation' .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool) operation_not_unique .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) unique_operation .
code_pred operation_defined' .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) static_operation' .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool) static_operation_not_unique .
code_pred (modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool,
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) unique_static_operation .
code_pred static_operation_defined' .
-code_pred has_literal' .
+declare has_literal'.intros[folded Predicate_Compile.contains_def, code_pred_intro]
+code_pred (modes:
+ i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool) has_literal'
+ by (elim has_literal'.cases) (simp add: Predicate_Compile.contains_def)
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,10462 +1,10454 @@
(***********************************************************************************
* 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\<open>The Shadow DOM\<close>
theory Shadow_DOM
imports
"monads/ShadowRootMonad"
Core_DOM.Core_DOM
begin
abbreviation "safe_shadow_root_element_types \<equiv> {''article'', ''aside'', ''blockquote'', ''body'',
''div'', ''footer'', ''h1'', ''h2'', ''h3'', ''h4'', ''h5'', ''h6'', ''header'', ''main'',
''nav'', ''p'', ''section'', ''span''}"
subsection \<open>Function Definitions\<close>
subsubsection \<open>get\_child\_nodes\<close>
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 \<Rightarrow> unit
\<Rightarrow> (_, (_) 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 \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> unit
\<Rightarrow> (_, (_) node_ptr list) dom_prog)) list" where
"a_get_child_nodes_tups \<equiv> [(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 \<circ> the \<circ> cast)]"
definition a_get_child_nodes :: "(_) object_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" where
"a_get_child_nodes_locs ptr \<equiv>
(if is_shadow_root_ptr_kind ptr
then {preserved (get_M (the (cast ptr)) RShadowRoot.child_nodes)} else {}) \<union>
CD.a_get_child_nodes_locs ptr"
definition first_child :: "(_) object_ptr \<Rightarrow> (_, (_) node_ptr option) dom_prog"
where
"first_child ptr = do {
children \<leftarrow> a_get_child_nodes ptr;
return (case children of [] \<Rightarrow> None | child#_ \<Rightarrow> 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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 |\<in>| object_ptr_kinds h"
shows "h \<turnstile> 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 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
shows "ptr |\<in>| 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 "\<not> 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 \<circ> the \<circ> 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 \<open>new\_document\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_child_nodes_new_document:
"ptr' \<noteq> cast new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^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 \<open>new\_shadow\_root\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_child_nodes_new_shadow_root:
"ptr' \<noteq> cast new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> r h h'"
apply(auto simp add: get_child_nodes_locs_def)[1]
apply (metis document_ptr_casts_commute3 insert_absorb insert_not_empty is_document_ptr_kind_none
new_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD)
apply(auto simp add: CD.get_child_nodes_locs_def)[1]
using new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t apply blast
apply (metis empty_iff insertE new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
apply (metis empty_iff new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t singletonD)
done
lemma new_shadow_root_no_child_nodes:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^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' \<noteq> cast new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> r h h'"
assumes new_shadow_root_no_child_nodes:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^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 \<open>new\_element\<close>
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' \<noteq> cast new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
assume "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assume "\<not> 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 \<open>h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr\<close>
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 \<open>delete\_shadow\_root\<close>
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' \<noteq> cast shadow_root_ptr \<Longrightarrow> h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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' \<noteq> cast shadow_root_ptr \<Longrightarrow> h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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 \<open>set\_child\_nodes\<close>
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 \<Rightarrow> (_) node_ptr list
\<Rightarrow> (_, 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 \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> (_) node_ptr list
\<Rightarrow> (_, unit) dom_prog)) list" where
"a_set_child_nodes_tups \<equiv> [(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 \<circ> the \<circ> cast)]"
definition a_set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, 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 \<Rightarrow> (_, unit) dom_prog set"
where
"a_set_child_nodes_locs ptr \<equiv>
(if is_shadow_root_ptr_kind ptr
then all_args (put_M (the (cast ptr)) RShadowRoot.child_nodes_update)
else {}) \<union> 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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, unit) dom_prog"
and set_child_nodes_locs :: "(_) object_ptr \<Rightarrow> (_, 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 \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, 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 \<Rightarrow> (_, 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 \<in> set_child_nodes_locs object_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_child_nodes_locs object_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)]
by(auto simp add: all_args_def type_wf_impl a_set_child_nodes_tups_def set_child_nodes_locs_def
CD.set_child_nodes_locs_def
split: if_splits option.splits)
end
interpretation
i_set_child_nodes?: l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf
DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes
Core_DOM_Functions.set_child_nodes_locs
apply(unfold_locales)
by (auto simp add: set_child_nodes_def set_child_nodes_locs_def)
declare l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma set_child_nodes_is_l_set_child_nodes [instances]: "l_set_child_nodes type_wf set_child_nodes
set_child_nodes_locs"
using instances Shadow_DOM.i_set_child_nodes.set_child_nodes_pointers_preserved Shadow_DOM.i_set_child_nodes.set_child_nodes_writes i_set_child_nodes.set_child_nodes_types_preserved
unfolding l_set_child_nodes_def
by blast
paragraph \<open>get\_child\_nodes\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap, exception, unit) prog set"
begin
lemma set_child_nodes_get_child_nodes:
assumes "known_ptr ptr"
assumes "type_wf h"
assumes "h \<turnstile> set_child_nodes ptr children \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
proof -
have "h \<turnstile> check_in_heap ptr \<rightarrow>\<^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 |\<in>| 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' \<turnstile> check_in_heap ptr \<rightarrow>\<^sub>r ()"
using check_in_heap_reads set_child_nodes_writes assms(3) \<open>h \<turnstile> check_in_heap ptr \<rightarrow>\<^sub>r ()\<close>
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 |\<in>| object_ptr_kinds h'"
using check_in_heap_ptr_in_heap by blast
with assms ptr_in_h \<open>type_wf h'\<close> 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 \<noteq> ptr'"
assumes "w \<in> set_child_nodes_locs ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> 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 \<open>set\_tag\_type\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_tag_name :: "(_) element_ptr \<Rightarrow> tag_name \<Rightarrow> (_, unit) dom_prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> (_, 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<in> set_tag_name_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_tag_name_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<open>get\_child\_nodes\<close>
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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_shadow\_root\<close>
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 \<Rightarrow> (_, (_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
where
"a_get_shadow_root_locs element_ptr \<equiv> {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 \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
fixes get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r children"
shows "element_ptr |\<in>| 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (get_shadow_root element_ptr)"
assumes get_shadow_root_ptr_in_heap:
"h \<turnstile> ok (get_shadow_root element_ptr) \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h"
assumes get_shadow_root_pure [simp]:
"pure (get_shadow_root element_ptr) h"
lemma get_shadow_root_is_l_get_shadow_root [instances]:
"l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs"
using instances
unfolding l_get_shadow_root_def
by (metis (no_types, lifting) ElementClass.l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms get_shadow_root_ok get_shadow_root_pure get_shadow_root_reads l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.intro l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_shadow_root_def)
paragraph \<open>set\_disconnected\_nodes\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_disconnected_nodes
:: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma set_disconnected_nodes_get_shadow_root:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_tag\_type\<close>
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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_child\_nodes\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma set_child_nodes_get_shadow_root: "\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow>
(\<forall>r \<in> 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:
"\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>delete\_shadow\_root\<close>
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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>new\_character\_data\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_character_data:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>new\_document\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_document:
"h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>new\_element\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_element:
"ptr' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_shadow_root new_element_ptr \<rightarrow>\<^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' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr
\<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> r h h'"
assumes new_element_no_shadow_root:
"h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_shadow_root new_element_ptr \<rightarrow>\<^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 \<open>new\_shadow\_root\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_shadow_root:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>set\_shadow\_root\<close>
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 \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, 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 \<Rightarrow> ((_, unit) dom_prog) set"
where
"a_set_shadow_root_locs element_ptr \<equiv> 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 \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, unit) dom_prog"
fixes set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_, 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 \<Rightarrow> bool"
and set_shadow_root :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, unit) dom_prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_, 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<turnstile> ok (set_shadow_root element_ptr shadow_root) \<Longrightarrow> element_ptr |\<in>| 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 \<in> set_shadow_root_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_shadow_root_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow>
h \<turnstile> ok (set_shadow_root element_ptr shadow_root)"
assumes set_shadow_root_ptr_in_heap:
"h \<turnstile> ok (set_shadow_root element_ptr shadow_root) \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h"
assumes set_shadow_root_pointers_preserved:
"w \<in> set_shadow_root_locs element_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
object_ptr_kinds h = object_ptr_kinds h'"
assumes set_shadow_root_types_preserved:
"w \<in> set_shadow_root_locs element_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> 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 \<open>get\_shadow\_root\<close>
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 \<Longrightarrow> h \<turnstile> set_shadow_root ptr shadow_root_ptr_opt \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_shadow_root ptr \<rightarrow>\<^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 \<noteq> ptr' \<Longrightarrow>
\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<Longrightarrow> h \<turnstile> set_shadow_root ptr shadow_root_ptr_opt \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r shadow_root_ptr_opt"
assumes set_shadow_root_get_shadow_root_different_pointers:
"ptr \<noteq> ptr' \<Longrightarrow> w \<in> set_shadow_root_locs ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>set\_mode\<close>
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 \<Rightarrow> shadow_root_mode \<Rightarrow> (_, 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 \<Rightarrow> ((_, unit) dom_prog) set"
where
"a_set_mode_locs shadow_root_ptr \<equiv> 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 \<Rightarrow> shadow_root_mode \<Rightarrow> (_, unit) dom_prog"
fixes set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> (_, 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 \<Rightarrow> bool"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, unit) dom_prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> (_, 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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<turnstile> ok (set_mode shadow_root_ptr shadow_root_mode) \<Longrightarrow>
shadow_root_ptr |\<in>| 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 \<in> set_mode_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_mode_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow>
h \<turnstile> ok (set_mode shadow_root_ptr shadow_root_mode)"
assumes set_mode_ptr_in_heap:
"h \<turnstile> ok (set_mode shadow_root_ptr shadow_root_mode) \<Longrightarrow>
shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
assumes set_mode_pointers_preserved:
"w \<in> set_mode_locs shadow_root_ptr \<Longrightarrow>
h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h'"
assumes set_mode_types_preserved:
"w \<in> set_mode_locs shadow_root_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> 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 \<open>get\_child\_nodes\<close>
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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_shadow\_root\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_child\_nodes\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_host\<close>
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 \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_get_host :: "(_) shadow_root_ptr \<Rightarrow> (_, (_) element_ptr) dom_prog" where
"a_get_host shadow_root_ptr = do {
host_ptrs \<leftarrow> element_ptr_kinds_M \<bind> filter_M (\<lambda>element_ptr. do {
shadow_root_opt \<leftarrow> get_shadow_root element_ptr;
return (shadow_root_opt = Some shadow_root_ptr)
});
(case host_ptrs of host_ptr#[] \<Rightarrow> return host_ptr | _ \<Rightarrow> error HierarchyRequestError)
}"
definition "a_get_host_locs \<equiv> (\<Union>element_ptr. (get_shadow_root_locs element_ptr)) \<union>
(\<Union>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 \<Rightarrow> (_, (_) element_ptr) dom_prog"
fixes get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<open>get\_mode\<close>
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 \<Rightarrow> (_, 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
where
"a_get_mode_locs shadow_root_ptr \<equiv> {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 \<Rightarrow> (_, shadow_root_mode) dom_prog"
fixes get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> ((_) heap, exception, shadow_root_mode) prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf :: "(_) heap \<Rightarrow> 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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<turnstile> get_mode shadow_root_ptr \<rightarrow>\<^sub>r children"
shows "shadow_root_ptr |\<in>| 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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (get_mode shadow_root_ptr)"
assumes get_mode_ptr_in_heap:
"h \<turnstile> ok (get_mode shadow_root_ptr) \<Longrightarrow> shadow_root_ptr |\<in>| 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 \<open>get\_shadow\_root\_safe\<close>
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 \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_mode :: "(_) shadow_root_ptr \<Rightarrow> (_, shadow_root_mode) dom_prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_get_shadow_root_safe :: "(_) element_ptr \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
where
"a_get_shadow_root_safe element_ptr = do {
shadow_root_ptr_opt \<leftarrow> get_shadow_root element_ptr;
(case shadow_root_ptr_opt of
Some shadow_root_ptr \<Rightarrow> do {
mode \<leftarrow> get_mode shadow_root_ptr;
(if mode = Open then
return (Some shadow_root_ptr)
else
return None
)
} | None \<Rightarrow> return None)
}"
definition a_get_shadow_root_safe_locs
:: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" where
"a_get_shadow_root_safe_locs element_ptr shadow_root_ptr \<equiv>
(get_shadow_root_locs element_ptr) \<union> (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 \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
fixes get_shadow_root_safe_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> bool"
and get_shadow_root_safe ::
"(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_safe_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_mode :: "(_) shadow_root_ptr \<Rightarrow> (_, shadow_root_mode) dom_prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
assumes get_shadow_root_safe_impl: "get_shadow_root_safe = a_get_shadow_root_safe"
assumes get_shadow_root_safe_locs_impl: "get_shadow_root_safe_locs = a_get_shadow_root_safe_locs"
begin
lemmas get_shadow_root_safe_def =
get_shadow_root_safe_impl[unfolded get_shadow_root_safe_def a_get_shadow_root_safe_def]
lemmas get_shadow_root_safe_locs_def =
get_shadow_root_safe_locs_impl[unfolded get_shadow_root_safe_locs_def a_get_shadow_root_safe_locs_def]
lemma get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h"
by (auto simp add: get_shadow_root_safe_def bind_pure_I option.case_eq_if)
end
interpretation i_get_shadow_root_safe?: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe
get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs
using instances
by (auto simp add: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def
get_shadow_root_safe_def get_shadow_root_safe_locs_def)
declare l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_get_shadow_root_safe = l_get_shadow_root_safe_defs +
assumes get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h"
lemma get_shadow_root_safe_is_l_get_shadow_root_safe [instances]:
"l_get_shadow_root_safe get_shadow_root_safe"
using instances
apply(auto simp add: l_get_shadow_root_safe_def)[1]
done
subsubsection \<open>set\_disconnected\_nodes\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_disconnected_nodes ::
"(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma set_disconnected_nodes_ok:
"type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<in> set_disconnected_nodes_locs object_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<open>get\_child\_nodes\<close>
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 \<Rightarrow> bool"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit)
prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma set_disconnected_nodes_get_child_nodes:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_host\<close>
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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_tag\_name\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, tag_name) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma get_tag_name_ok:
"type_wf h \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<open>set\_disconnected\_nodes\<close>
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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_child\_nodes\<close>
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:
"\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>delete\_shadow\_root\<close>
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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>set\_shadow\_root\<close>
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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>new\_element\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, tag_name) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_tag_name_new_element:
"ptr' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr
\<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> r h h'"
assumes new_element_empty_tag_name:
"h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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 \<open>get\_shadow\_root\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>new\_document\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, tag_name) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_tag_name_new_document:
"h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>new\_shadow\_root\<close>
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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>new\_character\_data\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, tag_name) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_tag_name_new_character_data:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>get\_tag\_type\<close>
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 \<turnstile> CD.a_set_tag_name element_ptr tag \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> CD.a_get_tag_name element_ptr \<rightarrow>\<^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 \<noteq> ptr'"
assumes "w \<in> CD.a_set_tag_name_locs ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> 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 \<open>attach\_shadow\_root\<close>
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 \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> ((_) heap, exception, unit) prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, char list) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_attach_shadow_root ::
"(_) element_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, (_) shadow_root_ptr) dom_prog"
where
"a_attach_shadow_root element_ptr shadow_root_mode = do {
tag \<leftarrow> get_tag_name element_ptr;
(if tag \<notin> safe_shadow_root_element_types then error NotSupportedError else return ());
prev_shadow_root \<leftarrow> get_shadow_root element_ptr;
(if prev_shadow_root \<noteq> None then error NotSupportedError else return ());
new_shadow_root_ptr \<leftarrow> 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 \<Rightarrow> shadow_root_mode \<Rightarrow> (_, (_) 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 \<Rightarrow> bool"
and set_shadow_root ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> ((_) heap, exception, unit) prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and attach_shadow_root ::
"(_) element_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr) prog"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, char list) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<turnstile> ok (attach_shadow_root element_ptr shadow_root_mode)"
shows "element_ptr |\<in>| element_ptr_kinds h"
proof -
obtain h' where "h \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^sub>h h'"
using assms by auto
then
obtain h2 h3 new_shadow_root_ptr where
h2: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h2" and
new_shadow_root_ptr: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr" and
h3: "h2 \<turnstile> set_mode new_shadow_root_ptr shadow_root_mode \<rightarrow>\<^sub>h h3" and
"h3 \<turnstile> set_shadow_root element_ptr (Some new_shadow_root_ptr) \<rightarrow>\<^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 |\<in>| element_ptr_kinds h3"
using set_shadow_root_ptr_in_heap by blast
moreover
have "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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="\<lambda>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 \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^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 \<open>get\_parent\<close>
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 \<open>set\_disconnected\_nodes\<close>
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]:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_root\_node\<close>
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 \<open>get\_root\_node\_si\<close>
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 \<Rightarrow> ((_) heap, exception, (_::linorder) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
partial_function (dom_prog) a_get_ancestors_si ::
"(_::linorder) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_get_ancestors_si ptr = do {
check_in_heap ptr;
ancestors \<leftarrow> (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 \<Rightarrow> do {
parent_ptr_opt \<leftarrow> get_parent node_ptr;
(case parent_ptr_opt of
Some parent_ptr \<Rightarrow> a_get_ancestors_si parent_ptr
| None \<Rightarrow> return [])
}
| None \<Rightarrow> (case cast ptr of
Some shadow_root_ptr \<Rightarrow> do {
host \<leftarrow> get_host shadow_root_ptr;
a_get_ancestors_si (cast host)
} |
None \<Rightarrow> return []));
return (ptr # ancestors)
}"
definition "a_get_ancestors_si_locs = get_parent_locs \<union> get_host_locs"
definition a_get_root_node_si :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr) dom_prog"
where
"a_get_root_node_si ptr = do {
ancestors \<leftarrow> 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 \<Rightarrow> (_, (_) object_ptr list) dom_prog"
fixes get_ancestors_si_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
locale l_get_root_node_si_defs =
fixes get_root_node_si :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr) dom_prog"
fixes get_root_node_si_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 "\<forall>ptr h h' x. h \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r x \<longrightarrow> h \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>h h' \<longrightarrow> 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 \<turnstile> ok (get_ancestors_si ptr)"
shows "ptr |\<in>| 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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
shows "ptr \<in> 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 \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors"
shows "ancestors \<noteq> []"
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 \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None \<Longrightarrow> h \<turnstile> get_root_node_si (cast node_ptr) \<rightarrow>\<^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 \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r root"
shows "\<not> 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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^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 \<open>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)\<close> \<open>root = last y\<close> 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 \<open>set\_disconnected\_nodes\<close>
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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_attribute\<close>
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 \<open>to\_tree\_order\<close>
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 \<open>to\_tree\_order\_si\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
partial_function (dom_prog) a_to_tree_order_si :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_to_tree_order_si ptr = (do {
children \<leftarrow> get_child_nodes ptr;
shadow_root_part \<leftarrow> (case cast ptr of
Some element_ptr \<Rightarrow> do {
shadow_root_opt \<leftarrow> get_shadow_root element_ptr;
(case shadow_root_opt of
Some shadow_root_ptr \<Rightarrow> return [cast shadow_root_ptr]
| None \<Rightarrow> return [])
} |
None \<Rightarrow> return []);
treeorders \<leftarrow> 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 \<Rightarrow> (_, (_) 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 "\<forall>ptr h h' x. h \<turnstile> to_tree_order_si ptr \<rightarrow>\<^sub>r x \<longrightarrow> h \<turnstile> to_tree_order_si ptr \<rightarrow>\<^sub>h h' \<longrightarrow> 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 "\<And>x h. pure (f x) h"
by (metis is_OK_returns_heap_E is_OK_returns_result_E pure_def)
then have "\<And>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 \<open>first\_in\_tree\_order\<close>
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 \<open>first\_in\_tree\_order\<close>
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 \<open>get\_element\_by\<close>
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 \<open>get\_element\_by\_si\<close>
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 \<open>find\_slot\<close>
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 \<Rightarrow> ((_) heap, exception, (_::linorder) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_mode :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, shadow_root_mode) prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_attribute :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, char list option) prog"
and get_attribute_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and first_in_tree_order ::
"(_) object_ptr \<Rightarrow> ((_) object_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog) \<Rightarrow>
((_) heap, exception, (_) element_ptr option) prog"
begin
definition a_find_slot :: "bool \<Rightarrow> (_) node_ptr \<Rightarrow> (_, (_) element_ptr option) dom_prog"
where
"a_find_slot open_flag slotable = do {
parent_opt \<leftarrow> get_parent slotable;
(case parent_opt of
Some parent \<Rightarrow>
if is_element_ptr_kind parent
then do {
shadow_root_ptr_opt \<leftarrow> get_shadow_root (the (cast parent));
(case shadow_root_ptr_opt of
Some shadow_root_ptr \<Rightarrow> do {
shadow_root_mode \<leftarrow> get_mode shadow_root_ptr;
if open_flag \<and> shadow_root_mode \<noteq> Open
then return None
else first_in_tree_order (cast shadow_root_ptr) (\<lambda>ptr. if is_element_ptr_kind ptr
then do {
tag \<leftarrow> get_tag_name (the (cast ptr));
name_attr \<leftarrow> get_attribute (the (cast ptr)) ''name'';
slotable_name_attr \<leftarrow> (if is_element_ptr_kind slotable
then get_attribute (the (cast slotable)) ''slot''
else return None);
(if (tag = ''slot'' \<and> (name_attr = slotable_name_attr \<or>
(name_attr = None \<and> slotable_name_attr = Some '''') \<or>
(name_attr = Some '''' \<and> slotable_name_attr = None)))
then return (Some (the (cast ptr)))
else return None)}
else return None)}
| None \<Rightarrow> return None)}
else return None
| _ \<Rightarrow> return None)}"
definition a_assigned_slot :: "(_) node_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> (_) node_ptr \<Rightarrow> (_, (_) element_ptr option) dom_prog"
and assigned_slot :: "(_) node_ptr \<Rightarrow> (_, (_) 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 \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r slot_opt"
shows "slotable |\<in>| 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 \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r Some slot"
shows "slot |\<in>| 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 \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r slot_opt \<Longrightarrow> slotable |\<in>| node_ptr_kinds h"
assumes find_slot_slot_in_heap:
"h \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r Some slot \<Longrightarrow> slot |\<in>| 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 \<open>get\_disconnected\_nodes\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma get_disconnected_nodes_ok:
"type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<open>set\_child\_nodes\<close>
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:
"\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_disconnected\_nodes\<close>
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 \<open>delete\_shadow\_root\<close>
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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> 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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> 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 \<open>set\_shadow\_root\<close>
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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_mode\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>new\_shadow\_root\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_disconnected_nodes_new_shadow_root:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> 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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> 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 \<open>remove\_shadow\_root\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_shadow_root ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_remove_shadow_root :: "(_) element_ptr \<Rightarrow> (_, unit) dom_prog" where
"a_remove_shadow_root element_ptr = do {
shadow_root_ptr_opt \<leftarrow> get_shadow_root element_ptr;
(case shadow_root_ptr_opt of
Some shadow_root_ptr \<Rightarrow> do {
children \<leftarrow> 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 \<Rightarrow> error HierarchyRequestError)
}"
definition a_remove_shadow_root_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_, unit) dom_prog) set"
where
"a_remove_shadow_root_locs element_ptr shadow_root_ptr \<equiv>
set_shadow_root_locs element_ptr \<union> {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 \<Rightarrow> (_, unit) dom_prog"
fixes remove_shadow_root_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_, 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 \<turnstile> get_shadow_root element_ptr|\<^sub>r))
(remove_shadow_root element_ptr) h h'"
apply(auto simp add: remove_shadow_root_locs_def remove_shadow_root_def all_args_def
writes_union_right_I writes_union_left_I set_shadow_root_writes
intro!: writes_bind writes_bind_pure[OF get_shadow_root_pure]
writes_bind_pure[OF get_child_nodes_pure]
intro: writes_subset[OF set_shadow_root_writes] writes_subset[OF writes_singleton2]
split: option.splits)[1]
using writes_union_left_I[OF set_shadow_root_writes]
apply (metis inf_sup_aci(5) insert_is_Un)
using writes_union_right_I[OF writes_singleton[of delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M]]
by (smt (verit, best) insert_is_Un writes_singleton2 writes_union_left_I)
end
interpretation i_remove_shadow_root?: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes
get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs
get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs
type_wf known_ptr
by(auto simp add: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def
remove_shadow_root_def remove_shadow_root_locs_def instances)
declare l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
paragraph \<open>get\_child\_nodes\<close>
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 \<noteq> cast shadow_root_ptr"
assumes "w \<in> remove_shadow_root_locs element_ptr shadow_root_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> 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 \<noteq> cast shadow_root_ptr \<Longrightarrow> w \<in> remove_shadow_root_locs element_ptr shadow_root_ptr \<Longrightarrow>
h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_child_nodes_locs ptr \<Longrightarrow> 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 \<open>get\_tag\_name\<close>
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 \<in> remove_shadow_root_locs element_ptr shadow_root_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> 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 \<in> remove_shadow_root_locs element_ptr shadow_root_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
r \<in> get_tag_name_locs ptr \<Longrightarrow> 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 \<open>get\_owner\_document\<close>
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 \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> unit \<Rightarrow> (_, (_) 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 \<leftarrow> 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 \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> unit
\<Rightarrow> (_, (_) 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 \<circ> the \<circ> cast)]"
definition a_get_owner_document :: "(_::linorder) object_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \<Rightarrow> bool"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node_si :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_si_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_owner_document :: "(_) object_ptr \<Rightarrow> ((_) 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 "\<And>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 \<turnstile> ok (get_owner_document ptr)"
shows "ptr |\<in>| 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 \<open>remove\_child\<close>
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 \<open>get\_disconnected\_document\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_get_disconnected_document :: "(_) node_ptr \<Rightarrow> (_, (_) document_ptr) dom_prog"
where
"a_get_disconnected_document node_ptr = do {
check_in_heap (cast node_ptr);
ptrs \<leftarrow> document_ptr_kinds_M;
candidates \<leftarrow> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> get_disconnected_nodes document_ptr;
return (node_ptr \<in> set disconnected_nodes)
}) ptrs;
(case candidates of
Cons document_ptr [] \<Rightarrow> return document_ptr |
_ \<Rightarrow> error HierarchyRequestError
)
}"
definition "a_get_disconnected_document_locs =
(\<Union>document_ptr. get_disconnected_nodes_locs document_ptr) \<union>
(\<Union>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 \<Rightarrow> (_, (_::linorder) document_ptr) dom_prog"
fixes get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<turnstile> ok (get_disconnected_document node_ptr) \<Longrightarrow> node_ptr |\<in>| 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 \<turnstile> get_disconnected_document child_node \<rightarrow>\<^sub>r disconnected_document"
shows "disconnected_document |\<in>| 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 \<turnstile> ok (get_disconnected_document node_ptr) \<Longrightarrow> node_ptr |\<in>| 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 \<turnstile> get_disconnected_document child_node \<rightarrow>\<^sub>r disconnected_document \<Longrightarrow>
disconnected_document |\<in>| 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 \<open>get\_disconnected\_nodes\<close>
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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>adopt\_node\<close>
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 \<open>get\_shadow\_root\<close>
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:
"\<forall>w \<in> adopt_node_locs parent owner_document document_ptr.
(h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> adopt_node_locs parent owner_document document_ptr.
(h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>insert\_before\<close>
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 \<open>get\_assigned\_nodes\<close>
fun map_filter_M2 :: "('x \<Rightarrow> ('heap, 'e, 'y option) prog) \<Rightarrow> 'x list
\<Rightarrow> ('heap, 'e, 'y list) prog"
where
"map_filter_M2 f [] = return []" |
"map_filter_M2 f (x # xs) = do {
res \<leftarrow> f x;
remainder \<leftarrow> map_filter_M2 f xs;
return ((case res of Some r \<Rightarrow> [r] | None \<Rightarrow> []) @ remainder)
}"
lemma map_filter_M2_pure [simp]:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> 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 "\<And>x. x \<in> set xs \<Longrightarrow> pure (f x) h"
assumes "h \<turnstile> map_filter_M2 f xs \<rightarrow>\<^sub>r ys"
shows
"ys = map the (filter (\<lambda>x. x \<noteq> None) (map (\<lambda>x. |h \<turnstile> f x|\<^sub>r) xs))" and
"\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> ok (f x)"
using assms
apply(induct xs arbitrary: h ys)
by(auto elim!: bind_returns_result_E2)
lemma map_filter_pure_foo:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> pure (f x) h"
assumes "h \<turnstile> map_filter_M2 f xs \<rightarrow>\<^sub>r ys"
assumes "y \<in> set ys"
obtains x where "h \<turnstile> f x \<rightarrow>\<^sub>r Some y" and "x \<in> set xs"
using assms
apply(induct xs arbitrary: ys)
by(auto elim!: bind_returns_result_E2)
lemma map_filter_M2_in_result:
assumes "h \<turnstile> map_filter_M2 P xs \<rightarrow>\<^sub>r ys"
assumes "a \<in> set xs"
assumes "\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h"
assumes "h \<turnstile> P a \<rightarrow>\<^sub>r Some b"
shows "b \<in> 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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and find_slot :: "bool \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and assigned_slot :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and remove :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before ::
"(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> (_) node_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before_locs :: "(_) object_ptr \<Rightarrow> (_) object_ptr option \<Rightarrow> (_) document_ptr \<Rightarrow>
(_) document_ptr \<Rightarrow> (_, unit) dom_prog set"
and append_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow>
((_) heap, exception, unit) prog set"
begin
definition a_assigned_nodes :: "(_) element_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
where
"a_assigned_nodes slot = do {
tag \<leftarrow> get_tag_name slot;
(if tag \<noteq> ''slot''
then error HierarchyRequestError
else return ());
root \<leftarrow> get_root_node (cast slot);
if is_shadow_root_ptr_kind root
then do {
host \<leftarrow> get_host (the (cast root));
children \<leftarrow> get_child_nodes (cast host);
filter_M (\<lambda>slotable. do {
found_slot \<leftarrow> find_slot False slotable;
return (found_slot = Some slot)}) children}
else return []}"
partial_function (dom_prog) a_assigned_nodes_flatten ::
"(_) element_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
where
"a_assigned_nodes_flatten slot = do {
tag \<leftarrow> get_tag_name slot;
(if tag \<noteq> ''slot''
then error HierarchyRequestError
else return ());
root \<leftarrow> get_root_node (cast slot);
(if is_shadow_root_ptr_kind root
then do {
slotables \<leftarrow> a_assigned_nodes slot;
slotables_or_child_nodes \<leftarrow> (if slotables = []
then do {
get_child_nodes (cast slot)
} else do {
return slotables
});
list_of_lists \<leftarrow> map_M (\<lambda>node_ptr. do {
(case cast node_ptr of
Some element_ptr \<Rightarrow> do {
tag \<leftarrow> get_tag_name element_ptr;
(if tag = ''slot''
then do {
root \<leftarrow> 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 \<Rightarrow> 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 \<leftarrow> element_ptr_kinds_M \<bind> map_filter_M2 (\<lambda>element_ptr. do {
tag \<leftarrow> get_tag_name element_ptr;
assigned_nodes \<leftarrow> a_assigned_nodes element_ptr;
(if tag = ''slot'' \<and> assigned_nodes \<noteq> []
then return (Some (element_ptr, assigned_nodes))
else return None)});
forall_M (\<lambda>(slot, assigned_nodes). do {
get_child_nodes (cast slot) \<bind> forall_M remove;
forall_M (append_child (cast slot)) assigned_nodes
}) tups;
shadow_root_ptr_kinds_M \<bind> forall_M (\<lambda>shadow_root_ptr. do {
host \<leftarrow> get_host shadow_root_ptr;
get_child_nodes (cast host) \<bind> forall_M remove;
get_child_nodes (cast shadow_root_ptr) \<bind> 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 \<Rightarrow> (_, (_) node_ptr list) dom_prog"
fixes assigned_nodes_flatten :: "(_) element_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> bool"
and assigned_nodes :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and assigned_nodes_flatten :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and flatten_dom :: "((_) heap, exception, unit) prog"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and find_slot :: "bool \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and assigned_slot :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and remove :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before ::
"(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> (_) node_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before_locs :: "(_) object_ptr \<Rightarrow> (_) object_ptr option \<Rightarrow> (_) document_ptr \<Rightarrow>
(_) document_ptr \<Rightarrow> (_, unit) dom_prog set"
and append_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow>
((_) heap, exception, unit) prog set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_shadow_root ::
"(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_shadow_root ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) 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 \<turnstile> ok (assigned_nodes slot)"
shows "slot |\<in>| 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 \<turnstile> ok (assigned_nodes slot)"
shows "h \<turnstile> get_tag_name slot \<rightarrow>\<^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 \<turnstile> assigned_nodes slot \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> assigned_nodes slot' \<rightarrow>\<^sub>r nodes'"
assumes "slot \<noteq> slot'"
shows "set nodes \<inter> set nodes' = {}"
proof (rule ccontr)
assume "set nodes \<inter> set nodes' \<noteq> {} "
then obtain common_ptr where "common_ptr \<in> set nodes" and "common_ptr \<in> set nodes'"
by auto
have "h \<turnstile> find_slot False common_ptr \<rightarrow>\<^sub>r Some slot"
using \<open>common_ptr \<in> set nodes\<close>
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 \<turnstile> find_slot False common_ptr \<rightarrow>\<^sub>r Some slot'"
using \<open>common_ptr \<in> set nodes'\<close>
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 \<turnstile> ok (assigned_nodes slot) \<Longrightarrow> slot |\<in>| element_ptr_kinds h"
assumes assigned_nodes_slot_is_slot:
"h \<turnstile> ok (assigned_nodes slot) \<Longrightarrow> h \<turnstile> get_tag_name slot \<rightarrow>\<^sub>r ''slot''"
assumes assigned_nodes_different_ptr:
"h \<turnstile> assigned_nodes slot \<rightarrow>\<^sub>r nodes \<Longrightarrow> h \<turnstile> assigned_nodes slot' \<rightarrow>\<^sub>r nodes' \<Longrightarrow>
slot \<noteq> slot' \<Longrightarrow> set nodes \<inter> 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 \<open>set\_val\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> (_, unit) dom_prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> (_, unit) dom_prog set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma set_val_ok:
"type_wf h \<Longrightarrow> character_data_ptr |\<in>| character_data_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<in> set_val_locs character_data_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_val_locs character_data_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<open>get\_shadow\_root\<close>
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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_tag\_type\<close>
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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>create\_character\_data\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool" +
assumes known_ptr_impl: "known_ptr = a_known_ptr"
begin
lemma create_character_data_document_in_heap:
assumes "h \<turnstile> ok (create_character_data document_ptr text)"
shows "document_ptr |\<in>| document_ptr_kinds h"
using assms CD.create_character_data_document_in_heap by simp
lemma create_character_data_known_ptr:
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<open>create\_element\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes ::
"(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and create_element ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> 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 \<turnstile> ok (create_element document_ptr tag)"
shows "document_ptr |\<in>| document_ptr_kinds h"
using CD.create_element_document_in_heap assms .
lemma create_element_known_ptr:
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<open>A wellformed heap (Core DOM)\<close>
subsubsection \<open>wellformed\_heap\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_host_shadow_root_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
where
"a_host_shadow_root_rel h = (\<lambda>(x, y). (cast x, cast y)) ` {(host, shadow_root).
host |\<in>| element_ptr_kinds h \<and> |h \<turnstile> 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
(\<lambda>host. (case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root \<Rightarrow> [(cast host, cast shadow_root)] |
None \<Rightarrow> []))
(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 \<Rightarrow> bool" where
"a_all_ptrs_in_heap h = ((\<forall>host shadow_root_ptr.
(h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr) \<longrightarrow>
shadow_root_ptr |\<in>| shadow_root_ptr_kinds h))"
definition a_distinct_lists :: "(_) heap \<Rightarrow> bool"
where
"a_distinct_lists h = distinct (concat (
map (\<lambda>element_ptr. (case |h \<turnstile> get_shadow_root element_ptr|\<^sub>r of
Some shadow_root_ptr \<Rightarrow> [shadow_root_ptr] | None \<Rightarrow> []))
|h \<turnstile> element_ptr_kinds_M|\<^sub>r
))"
definition a_shadow_root_valid :: "(_) heap \<Rightarrow> bool" where
"a_shadow_root_valid h = (\<forall>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h).
(\<exists>host \<in> fset(element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr))"
definition a_heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
where
"a_heap_is_wellformed h \<longleftrightarrow> CD.a_heap_is_wellformed h \<and>
acyclic (CD.a_parent_child_rel h \<union> a_host_shadow_root_rel h) \<and>
a_all_ptrs_in_heap h \<and>
a_distinct_lists h \<and>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 = ((\<forall>host \<in> fset (element_ptr_kinds h).
h \<turnstile> ok (get_shadow_root host) \<longrightarrow> (case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root_ptr \<Rightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h |
None \<Rightarrow> 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)
+ by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap select_result_I2)
lemma get_shadow_root_shadow_root_ptr_in_heap:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr"
shows "shadow_root_ptr |\<in>| 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 \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r host"
shows "shadow_root_ptr |\<in>| 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 \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr"
assumes "h \<turnstile> get_shadow_root host' \<rightarrow>\<^sub>r Some shadow_root_ptr"
shows "host = host'"
proof (rule ccontr)
assume " host \<noteq> host'"
have "host |\<in>| element_ptr_kinds h"
using assms(3)
by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap)
moreover
have "host' |\<in>| 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 \<open>host \<noteq> host'\<close> apply(simp)
apply(auto)[1]
done
qed
lemma shadow_root_host_dual:
assumes "h \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r host"
shows "h \<turnstile> get_shadow_root host \<rightarrow>\<^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 \<turnstile> get_disconnected_document disc_node \<rightarrow>\<^sub>r disc_doc"
obtains disc_nodes where
"h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes" and
"disc_node \<in> 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 \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r host"
assumes "h \<turnstile> get_tag_name host \<rightarrow>\<^sub>r tag"
shows "tag \<in> safe_shadow_root_element_types"
proof -
obtain host' where
"host' |\<in>| element_ptr_kinds h" and
"|h \<turnstile> get_tag_name host'|\<^sub>r \<in> safe_shadow_root_element_types"
and "h \<turnstile> get_shadow_root host' \<rightarrow>\<^sub>r Some shadow_root_ptr"
using assms
- by (metis finite_set_in get_host_ptr_in_heap local.a_shadow_root_valid_def
+ by (metis get_host_ptr_in_heap local.a_shadow_root_valid_def
local.get_shadow_root_ok local.heap_is_wellformed_def returns_result_select_result)
then have "host = host'"
by (meson assms(1) assms(2) assms(3) shadow_root_host_dual shadow_root_same_host)
then show ?thesis
using \<open>|h \<turnstile> get_tag_name host'|\<^sub>r \<in> safe_shadow_root_element_types\<close> assms(4) by auto
qed
lemma a_host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)"
proof -
have "a_host_shadow_root_rel h = (\<Union>host \<in> fset (element_ptr_kinds h).
(case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root \<Rightarrow> {(cast host, cast shadow_root)} | None \<Rightarrow> {}))"
by(auto simp add: a_host_shadow_root_rel_def split: option.splits)
moreover have "finite (\<Union>host \<in> fset (element_ptr_kinds h).
(case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root \<Rightarrow> {(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 \<Rightarrow> {}))"
by(auto split: option.splits)
ultimately show ?thesis
by auto
qed
lemma heap_is_wellformed_children_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children \<Longrightarrow>
child |\<in>| 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 \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
node \<in> set disc_nodes \<Longrightarrow> node |\<in>| 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 \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow>
h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children' \<Longrightarrow> set children \<inter> set children' \<noteq> {} \<Longrightarrow> 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 \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes' \<Longrightarrow> set disc_nodes \<inter> set disc_nodes' \<noteq> {}
\<Longrightarrow> 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 \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow>
h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow> set children \<inter> 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 \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
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 \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> 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 \<Longrightarrow> node_ptr |\<in>| node_ptr_kinds h \<Longrightarrow>
\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r) \<Longrightarrow>
(\<exists>document_ptr \<in> fset (document_ptr_kinds h).
node_ptr \<in> set |h \<turnstile> 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 \<Longrightarrow> finite (parent_child_rel h)"
using CD.parent_child_rel_finite by blast
lemma parent_child_rel_acyclic: "heap_is_wellformed h \<Longrightarrow> 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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptr parent \<Longrightarrow>
(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> child_ptr |\<in>| 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 \<open>get\_parent\<close>
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 \<open>get\_disconnected\_nodes\<close>
subsubsection \<open>set\_disconnected\_nodes\<close>
paragraph \<open>get\_disconnected\_nodes\<close>
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 \<open>get\_root\_node\<close>
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 \<open>to\_tree\_order\<close>
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 \<open>get\_root\_node\<close>
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 \<open>remove\_child\<close>
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 \<open>A wellformed heap\<close>
subsubsection \<open>get\_parent\<close>
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 \<open>remove\_shadow\_root\<close>
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 \<turnstile> remove_shadow_root ptr \<rightarrow>\<^sub>h h'"
shows "known_ptrs h'" and "type_wf h'" "heap_is_wellformed h'"
proof -
obtain shadow_root_ptr h2 where
"h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr" and
"h \<turnstile> get_child_nodes (cast shadow_root_ptr) \<rightarrow>\<^sub>r []" and
h2: "h \<turnstile> set_shadow_root ptr None \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> delete_M shadow_root_ptr \<rightarrow>\<^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="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_shadow_root_writes h2]
using \<open>type_wf h\<close> 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="\<lambda>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 \<open>known_ptrs h\<close> object_ptr_kinds_eq_h known_ptrs_subset
by blast
have object_ptr_kinds_eq_h2: "object_ptr_kinds h' |\<subseteq>| 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' |\<union>| {|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' |\<subseteq>| 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' |\<union>| {|shadow_root_ptr|}"
using object_ptr_kinds_eq2_h2
- apply (auto simp add: shadow_root_ptr_kinds_def)[1]
- by (metis \<open>h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> assms(1)
- fset.map_comp local.get_shadow_root_shadow_root_ptr_in_heap object_ptr_kinds_eq_h
- shadow_root_ptr_kinds_def)
+ by (auto simp add: shadow_root_ptr_kinds_def)
show "known_ptrs h'"
using object_ptr_kinds_eq_h2 \<open>known_ptrs h2\<close> known_ptrs_subset
by blast
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h2 \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. |h \<turnstile> get_tag_name doc_ptr|\<^sub>r = |h2 \<turnstile> get_tag_name doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h' \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. |h2 \<turnstile> get_tag_name doc_ptr|\<^sub>r = |h' \<turnstile> get_tag_name doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h:
"\<And>ptr' children. h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. ptr' \<noteq> cast shadow_root_ptr \<Longrightarrow> h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children =
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr' \<noteq> cast shadow_root_ptr \<Longrightarrow> |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r =
|h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "cast shadow_root_ptr |\<notin>| 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:
"\<And>shadow_root_opt ptr'. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>shadow_root_opt ptr'. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_shadow_root ptr'|\<^sub>r =
|h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
using select_result_eq by force
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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' \<subseteq> parent_child_rel h2"
using object_ptr_kinds_eq_h2
apply(auto simp add: CD.parent_child_rel_def)[1]
by (metis \<open>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 |\<notin>| object_ptr_kinds h'\<close> 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 \<open>heap_is_wellformed h\<close>
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 \<open>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 |\<notin>| object_ptr_kinds h'\<close>
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
+ apply (metis (no_types, opaque_lifting) children_eq2_h2 fin_mono object_ptr_kinds_eq_h2
subsetD)
by (metis (full_types) assms(1) assms(2) disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok
local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h node_ptr_kinds_eq_h2
returns_result_select_result)
moreover
have "CD.a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
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 \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
using children_eq_h2 concat_map_all_distinct[of "(\<lambda>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r)"]
- apply (metis (no_types, lifting) children_eq2_h2 finite_fset fmember_iff_member_fset fset_mp
+ apply (metis (no_types, lifting) children_eq2_h2 finite_fset fset_mp
object_ptr_kinds_eq_h2 set_sorted_list_of_set)
apply(case_tac "x = cast shadow_root_ptr")
using \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
apply(case_tac "y = cast shadow_root_ptr")
using \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
using children_eq_h2 distinct_concat_map_E(1)[of "(\<lambda>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r)"]
apply (metis (no_types, lifting) IntI \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> children_eq2_h2 empty_iff
- finite_fset finite_set_in is_OK_returns_result_I l_get_child_nodes.get_child_nodes_ptr_in_heap
+ finite_fset is_OK_returns_result_I l_get_child_nodes.get_child_nodes_ptr_in_heap
local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_get_child_nodes_axioms
returns_result_select_result sorted_list_of_set.set_sorted_key_list_of_set)
apply(case_tac "xa = cast shadow_root_ptr")
using \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
by (metis (mono_tags, lifting) CD.distinct_lists_no_parent \<open>known_ptrs h'\<close>
\<open>local.CD.a_distinct_lists h2\<close> \<open>type_wf h'\<close> children_eq2_h2 children_eq_h2
disconnected_nodes_eq_h2 is_OK_returns_result_I l_get_child_nodes.get_child_nodes_ptr_in_heap
local.get_child_nodes_ok local.get_disconnected_nodes_ok local.known_ptrs_known_ptr
local.l_get_child_nodes_axioms returns_result_select_result)
moreover
have "CD.a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close>
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: "\<forall>node_ptr\<in>fset (node_ptr_kinds h').
(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' \<and>
node_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h2 \<and>
node_ptr \<in> set |h2 \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
and 1: "node_ptr |\<in>| node_ptr_kinds h'"
and 2: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<longrightarrow>
node_ptr \<notin> set |h' \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
then have "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h2 \<longrightarrow>
node_ptr \<notin> set |h2 \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
apply(auto)[1]
apply(case_tac "parent_ptr = cast shadow_root_ptr")
using \<open>h \<turnstile> get_child_nodes (cast shadow_root_ptr) \<rightarrow>\<^sub>r []\<close> 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 "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' \<and>
node_ptr \<in> set |h' \<turnstile> 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 \<union> a_host_shadow_root_rel h)"
using \<open>heap_is_wellformed h\<close>
by(simp add: heap_is_wellformed_def)
then
have "acyclic (parent_child_rel h2 \<union> a_host_shadow_root_rel h2)"
proof -
have "a_host_shadow_root_rel h2 \<subseteq> 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) \<open>type_wf h2\<close> 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) \<open>type_wf h2\<close> 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 \<open>parent_child_rel h = parent_child_rel h2\<close>
by (metis (no_types, opaque_lifting) \<open>acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h)\<close>
acyclic_subset order_refl sup_mono)
qed
then
have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h')"
proof -
have "a_host_shadow_root_rel h' \<subseteq> 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 \<open>parent_child_rel h' \<subseteq> parent_child_rel h2\<close>
\<open>acyclic (parent_child_rel h2 \<union> a_host_shadow_root_rel h2)\<close>
using acyclic_subset sup_mono
by (metis (no_types, opaque_lifting))
qed
moreover
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close>
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) \<open>h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> 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 \<open>heap_is_wellformed h\<close>
by(simp add: heap_is_wellformed_def)
then
have "a_distinct_lists h2"
apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1]
apply(auto intro!: distinct_concat_map_I split: option.splits)[1]
apply(case_tac "x = ptr")
apply(simp)
apply (metis (no_types, opaque_lifting) assms(2) h2 is_OK_returns_result_I
l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root
l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI
returns_result_eq returns_result_select_result)
apply(case_tac "y = ptr")
apply(simp)
apply (metis (no_types, opaque_lifting) assms(2) h2 is_OK_returns_result_I
l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root
l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI
returns_result_eq returns_result_select_result)
by (metis \<open>type_wf h2\<close> 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 \<open>heap_is_wellformed h\<close>
by(simp add: heap_is_wellformed_def)
then
have "a_shadow_root_valid h'"
by (smt (verit) \<open>h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> assms(2)
delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap element_ptr_kinds_eq_h element_ptr_kinds_eq_h2
- finite_set_in finsert_iff funion_finsert_right get_shadow_root_eq2_h2 get_shadow_root_eq_h h'
+ finsert_iff funion_finsert_right get_shadow_root_eq2_h2 get_shadow_root_eq_h h'
local.a_shadow_root_valid_def local.get_shadow_root_ok object_ptr_kinds_eq2_h2
object_ptr_kinds_eq_h option.sel returns_result_select_result select_result_I2
shadow_root_ptr_kinds_commutes sup_bot.right_neutral tag_name_eq2_h tag_name_eq2_h2)
ultimately show "heap_is_wellformed h'"
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_remove_shadow_root_wf?: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf get_tag_name get_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs
set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root
get_shadow_root_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host
get_host_locs get_disconnected_document get_disconnected_document_locs remove_shadow_root
remove_shadow_root_locs known_ptrs get_parent get_parent_locs
by(auto simp add: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_root\_node\<close>
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 \<open>get\_parent\_get\_host\<close>
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 = (\<Union>host \<in> fset (element_ptr_kinds h).
(case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root \<Rightarrow> {(cast host, cast shadow_root)} | None \<Rightarrow> {}))"
by(auto simp add: a_host_shadow_root_rel_def split: option.splits)
moreover have "finite (\<Union>host \<in> fset (element_ptr_kinds h).
(case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root \<Rightarrow> {(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 \<Rightarrow> {}))"
by(auto split: option.splits)
ultimately show ?thesis
by auto
qed
lemma host_shadow_root_rel_shadow_root:
"h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r shadow_root_option \<Longrightarrow>
shadow_root_option = Some shadow_root \<longleftrightarrow> ((cast host, cast shadow_root) \<in> 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 \<Longrightarrow> h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host \<Longrightarrow>
(cast host, cast shadow_root) \<in> 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 "\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> P (cast child)) \<Longrightarrow> (\<And>shadow_root host. parent = cast host \<Longrightarrow>
h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root \<Longrightarrow> P (cast shadow_root)) \<Longrightarrow> P parent"
shows "P ptr"
proof -
fix ptr
have "finite (parent_child_rel h \<union> 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 \<union> a_host_shadow_root_rel h)\<inverse>)"
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 "\<And>child. (\<And>parent child_node. child = cast child_node \<Longrightarrow>
h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent) \<Longrightarrow>
(\<And>host shadow_root. child = cast shadow_root \<Longrightarrow> h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host \<Longrightarrow>
P (cast host)) \<Longrightarrow> P child"
shows "P ptr"
proof -
fix ptr
have "finite (parent_child_rel h \<union> 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 \<union> 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
\<Longrightarrow> (\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children
\<Longrightarrow> P (cast child))
\<Longrightarrow> (\<And>shadow_root host. parent = cast host \<Longrightarrow> h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root
\<Longrightarrow> P (cast shadow_root))
\<Longrightarrow> P parent)
\<Longrightarrow> P ptr"
assumes heap_wellformed_induct_rev_si [consumes 1, case_names step]:
"heap_is_wellformed h
\<Longrightarrow> (\<And>child. (\<And>parent child_node. child = cast child_node \<Longrightarrow>
h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent)
\<Longrightarrow> (\<And>host shadow_root. child = cast shadow_root \<Longrightarrow> h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host
\<Longrightarrow> P (cast host))
\<Longrightarrow> P child)
\<Longrightarrow> 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 \<open>get\_host\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
begin
lemma get_host_ok [simp]:
assumes "heap_is_wellformed h"
assumes "type_wf h"
assumes "known_ptrs h"
assumes "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
shows "h \<turnstile> ok (get_host shadow_root_ptr)"
proof -
obtain host where host: "host |\<in>| element_ptr_kinds h"
and "|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types"
and shadow_root: "h \<turnstile> get_shadow_root host \<rightarrow>\<^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)
+ by (smt returns_result_select_result)
obtain host_candidates where
host_candidates: "h \<turnstile> filter_M (\<lambda>element_ptr. Heap_Error_Monad.bind (get_shadow_root element_ptr)
(\<lambda>shadow_root_opt. return (shadow_root_opt = Some shadow_root_ptr)))
(sorted_list_of_set (fset (element_ptr_kinds h)))
\<rightarrow>\<^sub>r host_candidates"
apply(drule is_OK_returns_result_E[rotated])
using get_shadow_root_ok assms(2)
by(auto intro!: filter_M_is_OK_I bind_pure_I bind_is_OK_I2)
then have "host_candidates = [host]"
apply(rule filter_M_ex1)
apply (simp add: host)
- apply (smt (verit) assms(1) assms(2) bind_pure_returns_result_I2 finite_fset finite_set_in
+ apply (smt (verit) assms(1) assms(2) bind_pure_returns_result_I2 finite_fset
host local.get_shadow_root_ok local.get_shadow_root_pure local.shadow_root_same_host
return_returns_result returns_result_eq shadow_root sorted_list_of_set(1))
by (simp_all add: assms(2) bind_pure_I bind_pure_returns_result_I2 host local.get_shadow_root_ok
returns_result_eq shadow_root)
then
show ?thesis
using host_candidates host assms(1) get_shadow_root_ok
apply(auto simp add: get_host_def known_ptrs_known_ptr
intro!: bind_is_OK_pure_I filter_M_pure_I filter_M_is_OK_I bind_pure_I split: list.splits)[1]
using assms(2) apply blast
apply (meson list.distinct(1) returns_result_eq)
by (meson list.distinct(1) list.inject returns_result_eq)
qed
end
interpretation i_get_host_wf?: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document
get_disconnected_document_locs known_ptr known_ptrs type_wf get_host get_host_locs get_shadow_root
get_shadow_root_locs get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel
heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
by(auto simp add: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_get_host_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_host_defs +
assumes get_host_ok:
"heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow>
shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<open>get\_root\_node\_si\<close>
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 \<turnstile> ok (get_root_node_si ptr)"
shows "ptr |\<in>| 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 |\<in>| object_ptr_kinds h"
shows "h \<turnstile> 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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
and "h' \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors'"
and "\<And>p children children'. h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children
\<Longrightarrow> h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children' \<Longrightarrow> set children' \<subseteq> set children"
and "\<And>p shadow_root_option shadow_root_option'. h \<turnstile> get_shadow_root p \<rightarrow>\<^sub>r shadow_root_option \<Longrightarrow>
h' \<turnstile> get_shadow_root p \<rightarrow>\<^sub>r shadow_root_option' \<Longrightarrow> (if shadow_root_option = None
then shadow_root_option' = None else shadow_root_option' = None \<or>
shadow_root_option' = shadow_root_option)"
and "node \<notin> 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 \<notin> set ancestors'"
proof -
have object_ptr_kinds_M_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
using object_ptr_kinds_eq3
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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 \<open>h \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors\<close> 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 \<open>h' \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors'\<close> 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 |\<in>| object_ptr_kinds h"
using local.get_ancestors_si_ptr_in_heap object_ptr_kinds_eq3 step.prems(2) by fastforce
have "node \<noteq> child"
using ancestors_remains step.prems(3) by auto
have 1: "\<And>p parent. h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent \<Longrightarrow> h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
proof -
fix p parent
assume "h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
then obtain children' where
children': "h' \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children'" and
p_in_children': "p \<in> set children'"
using get_parent_child_dual by blast
obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^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 \<open>h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent\<close> get_parent_parent_in_heap is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
have "p \<in> set children"
using assms(5) children children' p_in_children'
by blast
then show "h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
using child_parent_dual assms(1) children known_ptrs type_wf by blast
qed
have 2: "\<And>p host. h' \<turnstile> get_host p \<rightarrow>\<^sub>r host \<Longrightarrow> h \<turnstile> get_host p \<rightarrow>\<^sub>r host"
proof -
fix p host
assume "h' \<turnstile> get_host p \<rightarrow>\<^sub>r host"
then have "h' \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some p"
using local.shadow_root_host_dual by blast
then have "h \<turnstile> get_shadow_root host \<rightarrow>\<^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 \<turnstile> get_host p \<rightarrow>\<^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) \<open>node \<noteq> child\<close>
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) \<open>node \<noteq> child\<close>
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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
shows "ptr' |\<in>| 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 \<turnstile> get_ancestors_si x \<rightarrow>\<^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 \<open>x = a\<close> by blast
next
case False
obtain ptr'' where ptr'': "h \<turnstile> get_ancestors_si ptr'' \<rightarrow>\<^sub>r ancestors"
using \<open> h \<turnstile> get_ancestors_si x \<rightarrow>\<^sub>r a # ancestors\<close> 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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
and "ancestor \<in> set ancestors"
and "h \<turnstile> get_ancestors_si ancestor \<rightarrow>\<^sub>r ancestor_ancestors"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "set ancestor_ancestors \<subseteq> 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 |\<in>| object_ptr_kinds h"
using get_ancestors_si_ptr_in_heap step(3) by auto
(* then have "h \<turnstile> check_in_heap child \<rightarrow>\<^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) \<open>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\<close>
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 |\<in>| shadow_root_ptr_kinds h"
using \<open>child |\<in>| object_ptr_kinds h\<close>
by (metis (no_types, lifting) shadow_root_ptr_casts_commute shadow_root_ptr_kinds_commutes)
obtain host where host: "h \<turnstile> get_host shadow_root_child \<rightarrow>\<^sub>r host"
using get_host_ok assms
by (meson \<open>shadow_root_child |\<in>| shadow_root_ptr_kinds h\<close> is_OK_returns_result_E)
then
have "h \<turnstile> get_ancestors_si (cast host) \<rightarrow>\<^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 \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using \<open>child |\<in>| object_ptr_kinds h\<close> 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: \<open>ancestors = [child]\<close>)[1]
using assms(4) returns_result_eq by fastforce
next
case (Some parent)
then
have "h \<turnstile> get_ancestors_si parent \<rightarrow>\<^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 \<open>h \<turnstile> get_ancestors_si parent \<rightarrow>\<^sub>r tl_ancestors\<close>
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 \<turnstile> get_ancestors_si some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast child \<in> set ancestors"
and "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "parent \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors:
"h \<turnstile> get_ancestors_si (cast child) \<rightarrow>\<^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 \<in> 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 \<turnstile> get_ancestors_si some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast shadow_root \<in> set ancestors"
and "h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "cast host \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors:
"h \<turnstile> get_ancestors_si (cast shadow_root) \<rightarrow>\<^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 \<in> 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 \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors"
assumes "((ptr, child) \<in> (parent_child_rel h)\<^sup>*)"
shows "ptr \<in> 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) \<in> (parent_child_rel h) \<and> (ptr_child, child) \<in> (parent_child_rel h)\<^sup>*"
using converse_rtranclE[OF 1(3)] \<open>ptr \<noteq> child\<close>
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 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children"
proof -
assume a1: "\<And>children. \<lbrakk>h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children; ptr_child_node \<in> set children\<rbrakk>
\<Longrightarrow> thesis"
have "ptr |\<in>| object_ptr_kinds h"
using CD.parent_child_rel_parent_in_heap ptr_child by blast
moreover have "ptr_child_node \<in> set |h \<turnstile> get_child_nodes ptr|\<^sub>r"
by (metis calculation \<open>known_ptrs h\<close> 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 \<open>type_wf h\<close>)
ultimately show ?thesis
using a1 get_child_nodes_ok \<open>type_wf h\<close> \<open>known_ptrs h\<close>
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) \<in> (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 \<in> set ancestors"
using 1 by auto
moreover have "h \<turnstile> get_parent ptr_child_node \<rightarrow>\<^sub>r Some ptr"
using assms(1) children ptr_child_node child_parent_dual
using \<open>known_ptrs h\<close> \<open>type_wf h\<close> by blast
ultimately show ?thesis
using get_ancestors_si_also_parent assms \<open>type_wf h\<close> 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 \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors"
assumes "((ptr, child) \<in> (parent_child_rel h \<union> a_host_shadow_root_rel h)\<^sup>*)"
shows "ptr \<in> 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) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h) \<and>
(ptr_child, child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<^sup>*"
using converse_rtranclE[OF 1(3)] \<open>ptr \<noteq> child\<close>
by metis
then show ?thesis
proof(cases "(ptr, ptr_child) \<in> 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 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children"
proof -
assume a1: "\<And>children. \<lbrakk>h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children; ptr_child_node \<in> set children\<rbrakk>
\<Longrightarrow> thesis"
have "ptr |\<in>| object_ptr_kinds h"
using CD.parent_child_rel_parent_in_heap True by blast
moreover have "ptr_child_node \<in> set |h \<turnstile> 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 \<open>type_wf h\<close> \<open>known_ptrs h\<close>
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) \<in>
(parent_child_rel h \<union> 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 \<in> set ancestors"
using 1 by auto
moreover have "h \<turnstile> get_parent ptr_child_node \<rightarrow>\<^sub>r Some ptr"
using assms(1) children ptr_child_node child_parent_dual
using \<open>known_ptrs h\<close> \<open>type_wf h\<close> by blast
ultimately show ?thesis
using get_ancestors_si_also_parent assms \<open>type_wf h\<close> 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 \<turnstile> get_shadow_root host \<rightarrow>\<^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) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<^sup>*"
using ptr_child ptr_child_shadow_root by blast
ultimately have "cast shadow_root \<in> set ancestors"
using "1.hyps"(2) host by blast
moreover have "h \<turnstile> get_host shadow_root \<rightarrow>\<^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 |\<in>| object_ptr_kinds h"
shows "h \<turnstile> 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 \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r root"
shows "root |\<in>| 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 \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r cast child"
shows "h \<turnstile> get_parent child \<rightarrow>\<^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 \<turnstile> get_parent child_node \<rightarrow>\<^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 \<open>get\_disconnected\_document\<close>
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 \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
shows "h \<turnstile> ok (get_disconnected_document node_ptr)"
proof -
have "node_ptr |\<in>| node_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap)
have "\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> 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 "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). node_ptr \<in>
set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
using heap_is_wellformed_children_disc_nodes
using \<open>node_ptr |\<in>| node_ptr_kinds h\<close> assms(1) by blast
then obtain some_owner_document where
"some_owner_document \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))" and
"node_ptr \<in> set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
have h5: "\<exists>!x. x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h))) \<and>
h \<turnstile> Heap_Error_Monad.bind (get_disconnected_nodes x)
(\<lambda>children. return (node_ptr \<in> set children)) \<rightarrow>\<^sub>r True"
apply(auto intro!: bind_pure_returns_result_I)[1]
apply (smt (verit, best) CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure
\<open>\<exists>document_ptr\<in>fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes
- document_ptr|\<^sub>r\<close> assms(2) bind_pure_returns_result_I finite_set_in return_returns_result
+ document_ptr|\<^sub>r\<close> assms(2) bind_pure_returns_result_I return_returns_result
returns_result_select_result)
apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1]
using heap_is_wellformed_one_disc_parent assms(1)
by blast
let ?filter_M = "filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (node_ptr \<in> set disconnected_nodes)))
(sorted_list_of_set (fset (document_ptr_kinds h)))"
have "h \<turnstile> ok (?filter_M)"
using CD.get_disconnected_nodes_ok
by (smt (verit) CD.get_disconnected_nodes_pure DocumentMonad.ptr_kinds_M_ptr_kinds
DocumentMonad.ptr_kinds_ptr_kinds_M assms(2) bind_is_OK_pure_I bind_pure_I
document_ptr_kinds_M_def filter_M_is_OK_I l_ptr_kinds_M.ptr_kinds_M_ok return_ok
return_pure returns_result_select_result)
then
obtain candidates where candidates: "h \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (node_ptr \<in> set disconnected_nodes)))
(sorted_list_of_set (fset (document_ptr_kinds h)))
\<rightarrow>\<^sub>r candidates"
by auto
have "candidates = [some_owner_document]"
apply(rule filter_M_ex1[OF candidates \<open>some_owner_document \<in>
set (sorted_list_of_set (fset (document_ptr_kinds h)))\<close> h5])
using \<open>node_ptr \<in> set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
\<open>some_owner_document \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))\<close>
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 \<open>node_ptr |\<in>| node_ptr_kinds h\<close>
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 \<open>get\_owner\_document\<close>
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 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "node_ptr \<in> set disc_nodes"
assumes known_ptrs: "known_ptrs h"
assumes type_wf: "type_wf h"
shows "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r document_ptr"
proof -
have 2: "node_ptr |\<in>| 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 |\<in>| document_ptr_kinds h"
using assms(2) get_disconnected_nodes_ptr_in_heap by blast
then have 4: "\<not>(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> 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 "(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
using assms(1) 2 "3" assms(2) assms(3) by auto
ultimately have 0: "\<exists>!document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r.
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
by (meson DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) disjoint_iff
local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent
returns_result_select_result type_wf)
have "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
by (metis (mono_tags, lifting) "2" "4" known_ptrs local.get_parent_child_dual
local.get_parent_ok local.get_parent_parent_in_heap returns_result_select_result
select_result_I2 split_option_ex type_wf)
then have 4: "h \<turnstile> get_root_node_si (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
using get_root_node_si_no_parent
by simp
obtain document_ptrs where document_ptrs: "h \<turnstile> document_ptr_kinds_M \<rightarrow>\<^sub>r document_ptrs"
by simp
then have "h \<turnstile> ok (filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> 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 \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^sub>r candidates"
by auto
have filter: "filter (\<lambda>document_ptr. |h \<turnstile> do {
disconnected_nodes \<leftarrow> 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 \<in> cast ` set disconnected_nodes)
}|\<^sub>r) document_ptrs = [document_ptr]"
apply(rule filter_ex1)
using 0 document_ptrs
apply (smt (verit) DocumentMonad.ptr_kinds_ptr_kinds_M bind_pure_returns_result_I2
local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure node_ptr_inclusion
return_returns_result select_result_I2 type_wf)
using assms(2) assms(3)
apply (metis (no_types, lifting) bind_pure_returns_result_I2 is_OK_returns_result_I
local.get_disconnected_nodes_pure node_ptr_inclusion return_returns_result select_result_I2)
using document_ptrs 3 apply(simp)
using document_ptrs
by simp
have "h \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^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 \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^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 \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
assumes "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document"
assumes "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
assumes "known_ptrs h"
assumes "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
have "\<And>parent. parent |\<in>| object_ptr_kinds h \<Longrightarrow> node_ptr \<notin> set |h \<turnstile> get_child_nodes parent|\<^sub>r"
using assms(2)
by (meson get_child_nodes_ok assms(1) assms(5) assms(6) local.child_parent_dual
local.known_ptrs_known_ptr option.distinct(1) returns_result_eq returns_result_select_result)
then show ?thesis
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
- finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok
+ is_OK_returns_result_I local.get_disconnected_nodes_ok
local.get_owner_document_disconnected_nodes local.get_parent_ptr_in_heap
local.heap_is_wellformed_children_disc_nodes returns_result_eq returns_result_select_result)
qed
lemma get_owner_document_owner_document_in_heap_node:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
shows "owner_document |\<in>| document_ptr_kinds h"
proof -
obtain root where
root: "h \<turnstile> get_root_node_si (cast node_ptr) \<rightarrow>\<^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 |\<in>| object_ptr_kinds h"
using root
using assms local.get_root_node_si_root_in_heap
by blast
have "\<not>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 \<open>known_ptr root\<close> \<open>root |\<in>| object_ptr_kinds h\<close>
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 "(\<exists>document_ptr \<in> fset (document_ptr_kinds h).
root \<in> cast ` set |h \<turnstile> 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
+ node_ptr_inclusion node_ptr_kinds_commutes option.distinct(1) returns_result_eq
returns_result_select_result root
- by (metis (no_types, lifting) assms \<open>root |\<in>| object_ptr_kinds h\<close>)
+ by (metis (no_types, opaque_lifting) assms \<open>root |\<in>| object_ptr_kinds h\<close>)
then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
then
obtain candidates where
candidates: "h \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (root \<in> 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)))
\<rightarrow>\<^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))
+ return_ok return_pure sorted_list_of_set(1))
then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto simp add: assms local.get_disconnected_nodes_ok
intro!: bind_pure_I bind_pure_returns_result_I)[1]
done
then have "candidates \<noteq> []"
by auto
then have "owner_document \<in> 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 \<open>is_node_ptr_kind root\<close> 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 \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
shows "owner_document |\<in>| 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 \<turnstile> invoke [] ptr () \<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| document_ptr_kinds h"
by (meson invoke_empty is_OK_returns_result_I)
next
assume "h \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (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 \<circ> the \<circ> 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 ())
\<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| 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: "\<not> 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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^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: "\<not> 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: "\<not> 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: "\<not> 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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (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 \<circ> the \<circ> 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 ())
\<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| 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 |\<in>| object_ptr_kinds h"
shows "h \<turnstile> 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: "\<not> 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: "\<not> 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: "\<not> 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: "\<not> 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 \<turnstile> 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: "\<not> 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: "\<not> 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: "\<not> 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 \<longrightarrow> h \<turnstile> ok Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (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 \<circ> the \<circ> 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 \<longrightarrow> h \<turnstile> ok Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (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 \<circ> the \<circ> 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 \<longrightarrow> h \<turnstile> ok Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 \<longrightarrow> h \<turnstile> ok Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 \<open>remove\_child\<close>
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 \<turnstile> remove_child ptr child \<rightarrow>\<^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 \<turnstile> remove_child ptr child \<rightarrow>\<^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 \<turnstile> remove_child ptr child \<rightarrow>\<^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: "\<And>ptr' shadow_root_ptr_opt. h \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq: "\<And>ptr' tag. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> 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="\<lambda>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' \<subseteq> parent_child_rel h"
using \<open>heap_is_wellformed h\<close> heap_is_wellformed_def
using CD.remove_child_parent_child_rel_subset
using \<open>known_ptrs h\<close> \<open>type_wf h\<close> assms(4)
by simp
show ?thesis
using \<open>heap_is_wellformed h\<close>
using \<open>heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\<close> \<open>parent_child_rel h' \<subseteq> parent_child_rel h\<close>
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 \<turnstile> remove child \<rightarrow>\<^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 \<turnstile> remove child \<rightarrow>\<^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 \<turnstile> remove child \<rightarrow>\<^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 \<Longrightarrow> h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> child \<notin> 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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow>
h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children \<Longrightarrow> h \<turnstile> remove_child ptr node_ptr \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow>
h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children \<Longrightarrow> h \<turnstile> remove node_ptr \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow>
h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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 \<open>adopt\_node\<close>
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 \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h2"
and children: "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<notin> set children"
proof -
obtain old_document parent_opt h' where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h': "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return () ) \<rightarrow>\<^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' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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
\<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close> returns_result_eq by fastforce
next
case (Some option)
then show ?case
using remove_child_removes_child \<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close> known_ptrs type_wf
wellformed
by auto
qed
qed
lemma adopt_node_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^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 \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent child | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> (if document_ptr \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 child old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes document_ptr;
set_disconnected_nodes document_ptr (child # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^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="\<lambda>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:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> 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' \<and> known_ptrs h' \<and> type_wf h'"
proof(cases "document_ptr = old_document")
case True
then show "heap_is_wellformed h' \<and> known_ptrs h' \<and> type_wf h'"
using h' wellformed_h2 \<open>known_ptrs h2\<close> \<open>type_wf h2\<close> by auto
next
case False
then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where
docs_neq: "document_ptr \<noteq> old_document" and
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 child old_disc_nodes) \<rightarrow>\<^sub>h h3" and
disc_nodes_document_ptr_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \<rightarrow>\<^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="\<lambda>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:
"\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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:
"\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r = |h3 \<turnstile> 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="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h3: "|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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:
"\<And>ptr children. h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r = |h' \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2: "\<And>doc_ptr disc_nodes. old_document \<noteq> doc_ptr \<Longrightarrow>
h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. old_document \<noteq> doc_ptr \<Longrightarrow>
|h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
using old_disc_nodes by blast
then have disc_nodes_old_document_h3:
"h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr \<Longrightarrow>
h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr \<Longrightarrow>
|h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2:
"h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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 |\<in>| 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 |\<in>| 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 \<in> 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 \<notin> set (remove1 child disc_nodes_old_document_h2)"
using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2
\<open>distinct disc_nodes_old_document_h2\<close> by auto
have "child \<notin> 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 (\<lambda>document_ptr.
|h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) |h2 \<turnstile> document_ptr_kinds_M|\<^sub>r))"
by(simp add: CD.a_distinct_lists_def)
show ?thesis
using distinct_concat_map_E(1)[OF 0] \<open>child \<in> set disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h2 disc_nodes_document_ptr_h2
by (meson \<open>type_wf h2\<close> 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 |\<in>| 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' \<subseteq> parent_child_rel h2"
proof
fix x
assume "x \<in> parent_child_rel h'"
then show "x \<in> 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 \<open> CD.a_acyclic_heap h2\<close> 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 \<open>type_wf h'\<close> children_eq2_h3 children_eq_h2 children_eq_h3 known_ptrs
l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok
local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms node_ptr_kinds_eq3_h2
object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3
returns_result_select_result wellformed_h2)
by (metis (no_types, opaque_lifting) disc_nodes_old_document_h2 disc_nodes_old_document_h3
- disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 finite_set_in select_result_I2
+ disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 select_result_I2
set_remove1_subset subsetD)
then have "CD.a_all_ptrs_in_heap h'"
apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1]
- apply (metis (no_types, opaque_lifting) children_eq2_h3 finite_set_in object_ptr_kinds_h3_eq3
- subsetD)
+ apply (metis (no_types, opaque_lifting) children_eq2_h3 object_ptr_kinds_h3_eq3 subsetD)
by (metis (no_types, opaque_lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> 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
+ document_ptr_kinds_eq3_h3 local.heap_is_wellformed_disc_nodes_in_heap
node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 select_result_I2 set_ConsD subsetD wellformed_h2)
moreover have "CD.a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
apply(simp add: CD.a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
children_eq2_h2 children_eq2_h3 )
by (metis (no_types) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2
disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3
in_set_remove1 list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2)
have a_distinct_lists_h2: "CD.a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h'"
apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2
children_eq2_h2 children_eq2_h3)[1]
proof -
assume 1: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 3: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> 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 \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 4: "distinct |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "old_document \<noteq> x")
case True
then show ?thesis
proof (cases "document_ptr \<noteq> x")
case True
then show ?thesis
using disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>]
disconnected_nodes_eq2_h3[OF \<open>document_ptr \<noteq> x\<close>] 4
by(auto)
next
case False
then show ?thesis
using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4
\<open>child \<notin> set disc_nodes_document_ptr_h3\<close>
by(auto simp add: disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>] )
qed
next
case False
then show ?thesis
by (metis (no_types, opaque_lifting) \<open>distinct disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 distinct_remove1 docs_neq select_result_I2)
qed
next
fix x y
assume a0: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a1: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a2: "x \<noteq> y"
moreover have 5: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter>
set |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter>
set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
proof(cases "old_document = x")
case True
have "old_document \<noteq> y"
using \<open>x \<noteq> y\<close> \<open>old_document = x\<close> by simp
have "document_ptr \<noteq> x"
using docs_neq \<open>old_document = x\<close> 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] \<open>old_document = x\<close>
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
\<open>document_ptr \<noteq> x\<close> 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 \<open>old_document = x\<close> docs_neq
\<open>old_document \<noteq> y\<close>
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]
\<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr = x\<close>
apply(simp)
by (metis (no_types, lifting)
\<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
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] \<open>old_document \<noteq> x\<close>
\<open>old_document = y\<close> \<open>document_ptr \<noteq> x\<close>
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 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False
\<open>type_wf h2\<close> 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 \<noteq> y"
using \<open>x \<noteq> y\<close> by auto
have "set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
using \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
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] \<open>old_document \<noteq> x\<close>
\<open>old_document \<noteq> y\<close> \<open>document_ptr = x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3 \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
by(auto)
next
case False
then show ?thesis
proof(cases "document_ptr = y")
case True
have f1: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set disc_nodes_document_ptr_h3 = {}"
using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2
document_ptr_kinds_eq2_h3 \<open>document_ptr \<noteq> x\<close>
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 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter>
set |h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = {}"
using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2
document_ptr_kinds_eq2_h3 \<open>old_document \<noteq> x\<close>
by (metis (no_types, lifting) a0 distinct_concat_map_E(1)
- document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 finite_fset fmember_iff_member_fset
+ document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 finite_fset
set_sorted_list_of_set)
ultimately show ?thesis
using 5 select_result_I2[OF disc_nodes_document_ptr_h']
select_result_I2[OF disc_nodes_old_document_h2] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr = y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> 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] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3
by (metis \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter>
set disc_nodes_old_document_h2 = {}\<close> empty_iff inf.idem)
qed
qed
qed
qed
qed
next
fix x xa xb
assume 0: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 2: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h'"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h'"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
then show False
using \<open>child \<in> set disc_nodes_old_document_h2\<close> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a3: "h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r remove1 child disc_nodes_old_document_h2"
assume a4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a5: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f6: "old_document |\<in>| document_ptr_kinds h'"
using a1 \<open>xb |\<in>| document_ptr_kinds h'\<close> by blast
have f7: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a2 by simp
have "x \<in> set disc_nodes_old_document_h2"
using f6 a3 a1 by (metis (no_types) \<open>type_wf h'\<close>
\<open>x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r\<close> 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' \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using f7 f6 a5 a4 \<open>xa |\<in>| object_ptr_kinds h'\<close>
by fastforce
then show ?thesis
using \<open>x \<in> set disc_nodes_old_document_h2\<close> a1 a4 f7 by blast
next
assume a1: "xb \<noteq> old_document"
assume a2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3"
assume a3: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a4: "xa |\<in>| object_ptr_kinds h'"
assume a5: "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r child # disc_nodes_document_ptr_h3"
assume a6: "old_document |\<in>| document_ptr_kinds h'"
assume a7: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
assume a8: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a10: "\<And>doc_ptr. old_document \<noteq> doc_ptr \<Longrightarrow>
|h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a11: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr \<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r =
|h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a12: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f13: "\<And>d. d \<notin> set |h' \<turnstile> document_ptr_kinds_M|\<^sub>r \<or> h2 \<turnstile> ok get_disconnected_nodes d"
using a9 \<open>type_wf h2\<close> get_disconnected_nodes_ok
by simp
then have f14: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a6 a3 by simp
have "x \<notin> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
using a12 a8 a4 \<open>xb |\<in>| document_ptr_kinds h'\<close>
- by (meson UN_I disjoint_iff_not_equal fmember_iff_member_fset)
+ by (meson UN_I disjoint_iff_not_equal)
then have "x = child"
using f13 a11 a10 a7 a5 a2 a1
by (metis (no_types, lifting) select_result_I2 set_ConsD)
then have "child \<notin> set disc_nodes_old_document_h2"
using f14 a12 a8 a6 a4
by (metis \<open>type_wf h'\<close> 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 \<open>child \<in> set disc_nodes_old_document_h2\<close> 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 \<open>CD.a_owner_document_valid h'\<close> CD.heap_is_wellformed_def
by simp
have shadow_root_eq_h2:
"\<And>ptr' shadow_root_ptr_opt. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have shadow_root_eq_h3:
"\<And>ptr' shadow_root_ptr_opt. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h2: "\<And>ptr' tag. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h3: "\<And>ptr' tag. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> 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="\<lambda>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="\<lambda>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' \<and> known_ptrs h' \<and> type_wf h'"
using \<open>heap_is_wellformed h2\<close>
using \<open>heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\<close> \<open>known_ptrs h'\<close> \<open>type_wf h'\<close>
using \<open>parent_child_rel h' \<subseteq> parent_child_rel h2\<close>
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 \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'"
and "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node_ptr # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
assume a2: "\<And>h document_ptr disc_nodes h'.
h \<turnstile> set_disconnected_nodes document_ptr disc_nodes \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume "h'a \<turnstile> set_disconnected_nodes owner_document (node_ptr # xb) \<rightarrow>\<^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 \<open>insert\_before\<close>
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 \<turnstile> insert_before ptr node child \<rightarrow>\<^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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
reference_child:
"h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
(* children: "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and *)
(* h': "h3 \<turnstile> set_child_nodes ptr (insert_before_list node reference_child children) \<rightarrow>\<^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 \<open>type_wf h\<close>
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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 "\<dots> = 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 "\<dots> = 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 \<open>known_ptrs h\<close> known_ptrs_preserved
by blast
have "known_ptrs h2"
using \<open>known_ptrs h\<close> known_ptrs_preserved \<open>object_ptr_kinds h = object_ptr_kinds h2\<close>
by blast
then
have "known_ptrs h3"
using known_ptrs_preserved \<open>object_ptr_kinds h2 = object_ptr_kinds h3\<close>
by blast
have "known_ptr ptr"
by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I \<open>known_ptrs h\<close>
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="\<lambda>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:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs )
then have object_ptr_kinds_M_eq2_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> 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 \<open>heap_is_wellformed h\<close> h2] \<open>known_ptrs h\<close>
\<open>type_wf h\<close>
.
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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:
"\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h2: "|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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="\<lambda>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:
"\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h3: "|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h3: "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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:
"\<And>ptr' shadow_root. h \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root =
h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. owner_document \<noteq> doc_ptr \<Longrightarrow>
h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. doc_ptr \<noteq> owner_document \<Longrightarrow>
|h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r =
|h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_h3:
"h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r remove1 node disconnected_nodes_h2"
using h3 set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h3:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow>
h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr \<noteq> ptr' \<Longrightarrow> |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
obtain children_h3 where children_h3: "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h3"
using h' a_insert_node_def by auto
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r insert_before_list node reference_child children_h3"
using h' \<open>type_wf h3\<close> \<open>known_ptr ptr\<close>
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 |\<in>| object_ptr_kinds h3"
using children_h3 get_child_nodes_ptr_in_heap by blast
have node_in_heap: "node |\<in>| node_ptr_kinds h"
using h2 adopt_node_child_in_heap by fast
have child_not_in_any_children:
"\<And>p children. h2 \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children \<Longrightarrow> node \<notin> set children"
using \<open>heap_is_wellformed h\<close> h2 adopt_node_removes_child \<open>type_wf h\<close> \<open>known_ptrs h\<close> by auto
have "node \<in> set disconnected_nodes_h2"
using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1)
\<open>type_wf h\<close> \<open>known_ptrs h\<close> by blast
have node_not_in_disconnected_nodes:
"\<And>d. d |\<in>| document_ptr_kinds h3 \<Longrightarrow> node \<notin> set |h3 \<turnstile> get_disconnected_nodes d|\<^sub>r"
proof -
fix d
assume "d |\<in>| document_ptr_kinds h3"
show "node \<notin> set |h3 \<turnstile> 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 \<open>d |\<in>| document_ptr_kinds h3\<close> disconnected_nodes_h3
by fastforce
next
case False
then have "set |h2 \<turnstile> get_disconnected_nodes d|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes owner_document|\<^sub>r = {}"
using distinct_concat_map_E(1) wellformed_h2
by (metis (no_types, lifting) \<open>d |\<in>| document_ptr_kinds h3\<close> \<open>type_wf h2\<close>
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] \<open>node \<in> set disconnected_nodes_h2\<close>
disconnected_nodes_h2 by fastforce
qed
qed
have "cast node \<noteq> ptr"
using ancestors node_not_in_ancestors get_ancestors_ptr
by fast
obtain ancestors_h2 where ancestors_h2: "h2 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_si_ok
by (metis \<open>known_ptrs h2\<close> \<open>type_wf h2\<close> is_OK_returns_result_E
object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2)
have ancestors_h3: "h3 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors_h2"
using get_ancestors_si_reads set_disconnected_nodes_writes h3
apply(rule reads_writes_separate_forwards)
using \<open>heap_is_wellformed h2\<close> 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 \<notin> set ancestors_h2"
using \<open>heap_is_wellformed h\<close> \<open>heap_is_wellformed h2\<close> 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)
\<open>type_wf h2\<close>
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 \<notin> {x. (x, ptr) \<in> (parent_child_rel h2)\<^sup>*}"
using get_ancestors_si_parent_child_rel
using \<open>known_ptrs h2\<close> \<open>type_wf h2\<close> ancestors_h2 node_not_in_ancestors_h2 wellformed_h2
by blast
then have "cast node \<notin> {x. (x, ptr) \<in> (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 \<open>parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\<close>
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 \<open>CD.a_all_ptrs_in_heap h2\<close>
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)
+ apply (metis (no_types, opaque_lifting) children_eq2_h2 in_mono object_ptr_kinds_M_eq3_h2)
by (metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h2
- disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in
+ disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes
node_ptr_kinds_eq2_h2 object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD)
have "set children_h3 \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using children_h3 \<open>CD.a_all_ptrs_in_heap h3\<close>
apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1]
using CD.parent_child_rel_child_nodes2 \<open>known_ptr ptr\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> \<open>type_wf h2\<close>
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) \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_in_heap
apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1]
- by (metis (no_types, opaque_lifting) contra_subsetD finite_set_in insert_before_list_in_set
+ by (metis (no_types, opaque_lifting) contra_subsetD insert_before_list_in_set
node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2)
then show ?thesis
using \<open>CD.a_all_ptrs_in_heap h3\<close>
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)
+ apply (metis (no_types, lifting) children_eq2_h3 select_result_I2 subsetD)
by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h3
- document_ptr_kinds_eq2_h3 finite_set_in subsetD)
+ document_ptr_kinds_eq2_h3 subsetD)
qed
moreover have "CD.a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def )
then have "CD.a_distinct_lists h3"
proof(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2
children_eq2_h2 intro!: distinct_concat_map_I)
fix x
assume 1: "x |\<in>| document_ptr_kinds h3"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
show "distinct |h3 \<turnstile> 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_iff_member_fset set_sorted_list_of_set)
+ by (metis (full_types) distinct_remove1 finite_fset set_sorted_list_of_set)
next
fix x y xa
assume 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and 2: "x |\<in>| document_ptr_kinds h3"
and 3: "y |\<in>| document_ptr_kinds h3"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r"
and 6: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r"
show False
proof (cases "x = owner_document")
case True
then have "y \<noteq> 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 \<open>y \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal
notin_set_remove1)
next
case False
then show ?thesis
proof (cases "y = owner_document")
case True
then show ?thesis
using distinct_concat_map_E(1)[OF 1]
using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3]
select_result_I2[OF disconnected_nodes_h2]
apply(auto simp add: True disconnected_nodes_eq2_h2[OF \<open>x \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal
notin_set_remove1)
next
case False
then show ?thesis
using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6
using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
- disjoint_iff_not_equal finite_fset fmember_iff_member_fset notin_set_remove1 select_result_I2
- set_sorted_list_of_set
- by (metis (no_types, lifting))
+ disjoint_iff_not_equal notin_set_remove1 select_result_I2
+ by (metis (no_types, opaque_lifting))
qed
qed
next
fix x xa xb
assume 1: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h3 \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h3). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 2: "xa |\<in>| object_ptr_kinds h3"
and 3: "x \<in> set |h3 \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h3"
and 5: "x \<in> set |h3 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 4
by (metis \<open>type_wf h2\<close> children_eq2_h2 document_ptr_kinds_commutes \<open>known_ptrs h\<close>
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 (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))" and
2: "x |\<in>| object_ptr_kinds h'"
have 3: "\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> distinct |h3 \<turnstile> get_child_nodes p|\<^sub>r"
using 1 by (auto elim: distinct_concat_map_E)
show "distinct |h' \<turnstile> 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 (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "x |\<in>| object_ptr_kinds h'"
and 3: "y |\<in>| object_ptr_kinds h'"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h' \<turnstile> get_child_nodes x|\<^sub>r"
and 6: "xa \<in> set |h' \<turnstile> get_child_nodes y|\<^sub>r"
have 7:"set |h3 \<turnstile> get_child_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 \<noteq> 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 \<open>ptr \<noteq> y\<close>])[1]
by (metis (no_types, opaque_lifting) "3" "7" \<open>type_wf h3\<close> children_eq2_h3 disjoint_iff_not_equal
get_child_nodes_ok insert_before_list_in_set \<open>known_ptrs h\<close> 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 \<open>ptr \<noteq> x\<close>])[1]
by (metis (no_types, opaque_lifting) "2" "4" "7" IntI \<open>known_ptrs h3\<close> \<open>type_wf h'\<close>
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 \<open>ptr \<noteq> x\<close>] children_eq2_h3[OF \<open>ptr \<noteq> y\<close>] 5 6 7 by auto
qed
qed
next
fix x xa xb
assume 1: " (\<Union>x\<in>fset (object_ptr_kinds h'). set |h3 \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h'). set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r) = {} "
and 2: "xa |\<in>| object_ptr_kinds h'"
and 3: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h'"
and 5: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 3 4 5
proof -
have "\<forall>h d. \<not> type_wf h \<or> d |\<notin>| document_ptr_kinds h \<or> h \<turnstile> ok get_disconnected_nodes d"
using local.get_disconnected_nodes_ok by satx
then have "h' \<turnstile> ok get_disconnected_nodes xb"
using "4" \<open>type_wf h'\<close> by fastforce
then have f1: "h3 \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
by (simp add: disconnected_nodes_eq_h3)
have "xa |\<in>| object_ptr_kinds h3"
using "2" object_ptr_kinds_M_eq3_h' by blast
then show ?thesis
using f1 \<open>local.CD.a_distinct_lists h3\<close> 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
\<open>CD.a_distinct_lists h3\<close> \<open>type_wf h'\<close> 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
+ by (metis (no_types, opaque_lifting) Core_DOM_Functions.i_insert_before.insert_before_list_in_set
children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2
- disconnected_nodes_h3 finite_set_in in_set_remove1 is_OK_returns_result_I object_ptr_kinds_M_eq3_h'
+ disconnected_nodes_h3 in_set_remove1 is_OK_returns_result_I object_ptr_kinds_M_eq3_h'
ptr_in_heap returns_result_eq returns_result_select_result)
ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'"
by (simp add: CD.heap_is_wellformed_def)
have shadow_root_eq_h2:
"\<And>ptr' shadow_root_ptr_opt. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have shadow_root_eq_h3:
"\<And>ptr' shadow_root_ptr_opt. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h2: "\<And>ptr' tag. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h3: "\<And>ptr' tag. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> 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="\<lambda>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="\<lambda>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 \<notin> {x. (x, ptr) \<in> (parent_child_rel h3 \<union> a_host_shadow_root_rel h3)\<^sup>*}"
using get_ancestors_si_parent_child_host_shadow_root_rel
using \<open>known_ptrs h2\<close> \<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> \<open>type_wf h2\<close> ancestors_h2 node_not_in_ancestors_h2
wellformed_h2
by auto
have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3)"
using \<open>heap_is_wellformed h2\<close>
by(auto simp add: heap_is_wellformed_def \<open>parent_child_rel h2 = parent_child_rel h3\<close>
\<open>a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\<close>)
then
have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h')"
apply(auto simp add: \<open>a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'\<close>
\<open>parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\<close>)[1]
using \<open>cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h3 \<union> a_host_shadow_root_rel h3)\<^sup>*}\<close>
by (simp add: \<open>local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\<close>)
then
show "heap_is_wellformed h'"
using \<open>heap_is_wellformed h2\<close>
using \<open>heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\<close>
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 \<open>append\_child\<close>
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 \<open>to\_tree\_order\<close>
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 \<open>get\_root\_node\<close>
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 \<open>to\_tree\_order\_si\<close>
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 |\<in>| object_ptr_kinds h"
shows "h \<turnstile> 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 \<open>get\_assigned\_nodes\<close>
lemma forall_M_small_big: "h \<turnstile> forall_M f xs \<rightarrow>\<^sub>h h' \<Longrightarrow> P h \<Longrightarrow>
(\<And>h h' x. x \<in> set xs \<Longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h' \<Longrightarrow> P h \<Longrightarrow> P h') \<Longrightarrow> 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 \<turnstile> assigned_nodes slot \<rightarrow>\<^sub>r nodes"
shows "distinct nodes"
proof -
have "\<And>ptr children. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> 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 \<turnstile> flatten_dom \<rightarrow>\<^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 \<turnstile> element_ptr_kinds_M \<rightarrow>\<^sub>r element_ptrs" and
tups: "h \<turnstile> map_filter_M2 (\<lambda>element_ptr. do {
tag \<leftarrow> get_tag_name element_ptr;
assigned_nodes \<leftarrow> assigned_nodes element_ptr;
(if tag = ''slot'' \<and> assigned_nodes \<noteq> [] then
return (Some (element_ptr, assigned_nodes)) else return None)}) element_ptrs \<rightarrow>\<^sub>r tups"
(is "h \<turnstile> map_filter_M2 ?f element_ptrs \<rightarrow>\<^sub>r tups") and
h2: "h \<turnstile> forall_M (\<lambda>(slot, assigned_nodes). do {
get_child_nodes (cast slot) \<bind> forall_M remove;
forall_M (append_child (cast slot)) assigned_nodes
}) tups \<rightarrow>\<^sub>h h2" and
"h2 \<turnstile> shadow_root_ptr_kinds_M \<rightarrow>\<^sub>r shadow_root_ptrs" and
h': "h2 \<turnstile> forall_M (\<lambda>shadow_root_ptr. do {
host \<leftarrow> get_host shadow_root_ptr;
get_child_nodes (cast host) \<bind> forall_M remove;
get_child_nodes (cast shadow_root_ptr) \<bind> forall_M (append_child (cast host));
remove_shadow_root host
}) shadow_root_ptrs \<rightarrow>\<^sub>h h'"
using \<open>h \<turnstile> flatten_dom \<rightarrow>\<^sub>h h'\<close>
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 \<and> known_ptrs h2 \<and> type_wf h2"
using h2 \<open>heap_is_wellformed h\<close> \<open>known_ptrs h\<close> \<open>type_wf h\<close>
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 = "\<lambda>h. heap_is_wellformed h \<and> known_ptrs h \<and> 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 = "\<lambda>h. heap_is_wellformed h \<and> known_ptrs h \<and> 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 \<open>get\_shadow\_root\_safe\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root ::
"(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root_safe ::
"(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_safe_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_mode :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, shadow_root_mode) prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
end
subsubsection \<open>create\_element\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes ::
"(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
begin
lemma create_element_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^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 \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>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="\<lambda>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 \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr
by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| 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 \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^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:
"\<And>ptr' disc_nodes. ptr' \<noteq> new_element_ptr
\<Longrightarrow> h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> new_element_ptr
\<Longrightarrow> |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h2 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. ptr' \<noteq> new_element_ptr
\<Longrightarrow> h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> new_element_ptr
\<Longrightarrow> |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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 \<open>type_wf h\<close> new_element_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. ptr' \<noteq> new_element_ptr
\<Longrightarrow> h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> new_element_ptr
\<Longrightarrow> |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_element_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close>
using \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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:
"\<And>ptr' disc_nodes. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting)
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = 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 \<open>heap_is_wellformed h\<close> 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 \<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> 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
+ by (metis (no_types, opaque_lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> assms(3) assms(4)
- children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I
+ children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h is_OK_returns_result_I
local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subsetD)
then have "CD.a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "CD.a_all_ptrs_in_heap h'"
by (smt (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2
- disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in finsertCI funion_finsert_right
+ disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finsertCI funion_finsert_right
h' insert_iff list.simps(15) local.CD.a_all_ptrs_in_heap_def
local.set_disconnected_nodes_get_disconnected_nodes node_ptr_kinds_eq_h node_ptr_kinds_eq_h2
node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2 subsetD subsetI)
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_element_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_element_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_element_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_element_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "CD.a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: CD.heap_is_wellformed_def heap_is_wellformed_def)
then have "CD.a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_element_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
by (metis \<open> CD.a_distinct_lists h\<close> \<open>type_wf h2\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> 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) \<open>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 \<notin> set disc_nodes_h3\<close>
\<open> CD.a_distinct_lists h3\<close> \<open>type_wf h'\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using calculation by(auto dest: distinct_concat_map_E(1))
ultimately show "False"
by (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>local.CD.a_all_ptrs_in_heap h\<close>
disc_nodes_document_ptr_h disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
- disconnected_nodes_eq2_h3 disjoint_iff document_ptr_kinds_eq_h document_ptr_kinds_eq_h2
- finite_set_in h'
+ disconnected_nodes_eq2_h3 disjoint_iff document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.CD.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply -
apply(cases "xb = document_ptr")
apply (metis (no_types, opaque_lifting) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3
\<Longrightarrow> 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 \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
\<open> CD.a_distinct_lists h3\<close> 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" \<open> CD.a_distinct_lists h3\<close> \<open>type_wf h3\<close> 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 \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
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
+ by(metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> 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'
+ document_ptr_kinds_eq_h h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
node_ptr_kinds_commutes select_result_I2)
have "CD.a_heap_is_wellformed h'"
using \<open>CD.a_acyclic_heap h'\<close> \<open>CD.a_all_ptrs_in_heap h'\<close> \<open>CD.a_distinct_lists h'\<close>
\<open>CD.a_owner_document_valid h'\<close>
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: "\<And>element_ptr shadow_root_opt. element_ptr \<noteq> new_element_ptr
\<Longrightarrow> h \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt"
proof -
fix element_ptr shadow_root_opt
assume "element_ptr \<noteq> new_element_ptr "
have "\<forall>P \<in> get_shadow_root_locs element_ptr. P h h2"
using get_shadow_root_new_element new_element_ptr h2
using \<open>element_ptr \<noteq> new_element_ptr\<close> 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 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt"
by (simp add: preserved_def)
qed
have shadow_root_none: "h2 \<turnstile> get_shadow_root (new_element_ptr) \<rightarrow>\<^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:
"\<And>ptr' children. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>ptr' children. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children"
using get_shadow_root_reads set_disconnected_nodes_writes h'
apply(rule reads_writes_preserved)
using set_disconnected_nodes_get_shadow_root
by(auto simp add: set_disconnected_nodes_get_shadow_root)
have "a_all_ptrs_in_heap h"
by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap)
then have "a_all_ptrs_in_heap h2"
apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1]
using returns_result_eq shadow_root_eq_h shadow_root_none by fastforce
then have "a_all_ptrs_in_heap h3"
apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1]
using shadow_root_eq_h2 by blast
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1]
by (simp add: shadow_root_eq_h3)
have "a_distinct_lists h"
using assms(1)
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h2"
apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1]
apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1]
apply(case_tac "x = new_element_ptr")
using shadow_root_none apply auto[1]
using shadow_root_eq_h
by (smt (verit) Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds
- ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) finite_set_in h2 insort_split
+ ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) h2 insort_split
local.get_shadow_root_ok local.shadow_root_same_host new_element_ptr new_element_ptr_not_in_heap
option.distinct(1) returns_result_select_result select_result_I2 shadow_root_none)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2])
then have "a_distinct_lists h'"
by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])
have "a_shadow_root_valid h"
using assms(1)
by (simp add: heap_is_wellformed_def)
then have "a_shadow_root_valid h2"
proof (unfold a_shadow_root_valid_def; safe)
fix shadow_root_ptr
assume "\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h). \<exists>host\<in>fset (element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and> |h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
assume "shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h2)"
obtain previous_host where
"previous_host \<in> fset (element_ptr_kinds h)" and
"|h \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr"
by (metis \<open>local.a_shadow_root_valid h\<close> \<open>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h2)\<close>
local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h)
moreover have "previous_host \<noteq> new_element_ptr"
using calculation(1) h2 new_element_ptr new_element_ptr_not_in_heap by auto
ultimately have "|h2 \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h2 \<turnstile> 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 \<open>previous_host \<noteq> new_element_ptr\<close>
\<open>|h \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\<close>
select_result_eq shadow_root_eq_h)
then
show "\<exists>host\<in>fset (element_ptr_kinds h2).
|h2 \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h2 \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
by (meson \<open>previous_host \<in> fset (element_ptr_kinds h)\<close> \<open>previous_host \<noteq> new_element_ptr\<close>
- assms(3) local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap notin_fset
+ assms(3) local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap
returns_result_select_result shadow_root_eq_h)
qed
then have "a_shadow_root_valid h3"
proof (unfold a_shadow_root_valid_def; safe)
fix shadow_root_ptr
assume "\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h2). \<exists>host\<in>fset (element_ptr_kinds h2).
|h2 \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h2 \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
assume "shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h3)"
obtain previous_host where
"previous_host \<in> fset (element_ptr_kinds h2)" and
"|h2 \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h2 \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr"
by (metis \<open>local.a_shadow_root_valid h2\<close> \<open>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h3)\<close>
local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2)
moreover have "previous_host \<noteq> 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 \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h2 \<turnstile> 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 \<open>previous_host \<noteq> new_element_ptr\<close>
\<open>|h2 \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\<close> select_result_eq
shadow_root_eq_h)
then
show "\<exists>host\<in>fset (element_ptr_kinds h3).
|h3 \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h3 \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
by (metis \<open>previous_host \<in> fset (element_ptr_kinds h2)\<close> \<open>previous_host \<noteq> new_element_ptr\<close>
element_ptr_kinds_eq_h2 select_result_eq shadow_root_eq_h2 tag_name_eq2_h2)
qed
then have "a_shadow_root_valid h'"
- by (smt (verit) \<open>type_wf h3\<close> element_ptr_kinds_eq_h3 finite_set_in local.a_shadow_root_valid_def
+ by (smt (verit) \<open>type_wf h3\<close> element_ptr_kinds_eq_h3 local.a_shadow_root_valid_def
local.get_shadow_root_ok returns_result_select_result select_result_I2 shadow_root_eq_h3
shadow_root_ptr_kinds_eq_h3 tag_name_eq2_h3)
have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2"
apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq_h)[1]
apply (smt (verit, del_insts) assms(3) case_prodI h2 local.get_shadow_root_ok mem_Collect_eq
new_element_ptr new_element_ptr_not_in_heap pair_imageI returns_result_select_result
select_result_I2 shadow_root_eq_h)
using shadow_root_none apply auto[1]
by (metis (no_types, lifting) assms(3) h2 host_shadow_root_rel_def
host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl
local.get_shadow_root_ok local.new_element_no_shadow_root new_element_ptr option.distinct(1)
returns_result_select_result select_result_I2 shadow_root_eq_h)
have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3"
apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1]
apply (metis (mono_tags, lifting) \<open>type_wf h2\<close> case_prodI local.get_shadow_root_ok
mem_Collect_eq pair_imageI returns_result_select_result select_result_I2 shadow_root_eq_h2)
apply (metis (mono_tags, lifting) case_prodI mem_Collect_eq pair_imageI select_result_eq shadow_root_eq_h2)
done
have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'"
apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1]
apply (metis (mono_tags, lifting) \<open>type_wf h'\<close> case_prodI element_ptr_kinds_eq_h2
element_ptr_kinds_eq_h3 local.get_shadow_root_ok mem_Collect_eq pair_imageI
returns_result_select_result select_result_I2 shadow_root_eq_h3)
apply (metis (mono_tags, lifting) \<open>type_wf h'\<close> case_prodI element_ptr_kinds_eq_h2
element_ptr_kinds_eq_h3 local.get_shadow_root_ok mem_Collect_eq pair_imageI
returns_result_select_result select_result_I2 shadow_root_eq_h3)
done
have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
have "parent_child_rel h \<union> a_host_shadow_root_rel h =
parent_child_rel h2 \<union> a_host_shadow_root_rel h2"
using \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>parent_child_rel h = parent_child_rel h2\<close> by auto
have "parent_child_rel h2 \<union> a_host_shadow_root_rel h2 =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3"
using \<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> by auto
have "parent_child_rel h' \<union> a_host_shadow_root_rel h' =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3"
by (simp add: \<open>local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\<close>
\<open>parent_child_rel h3 = parent_child_rel h'\<close>)
have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3)"
using \<open>acyclic (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<close>
\<open>parent_child_rel h \<union> local.a_host_shadow_root_rel h =
parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2\<close>
\<open>parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 =
parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3\<close>
by auto
then have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h')"
by(simp add: \<open>parent_child_rel h' \<union> a_host_shadow_root_rel h' =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3\<close>)
show " heap_is_wellformed h' "
using \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h')\<close>
by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl
\<open>local.CD.a_heap_is_wellformed h'\<close> \<open>local.a_all_ptrs_in_heap h'\<close> \<open>local.a_distinct_lists h'\<close>
\<open>local.a_shadow_root_valid h'\<close>)
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 \<open>create\_character\_data\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) 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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
begin
lemma create_character_data_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^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 \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>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="\<lambda>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 \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
local.create_character_data_known_ptr by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| 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 \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>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="\<lambda>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 |\<in>| 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 \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h:
"\<And>ptr' disc_nodes. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h2 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_character_data_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h3:
"\<And>ptr' disc_nodes. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_character_data_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close> using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = 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 \<open>heap_is_wellformed h\<close> 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 \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
- apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M
+ apply (metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>parent_child_rel h = parent_child_rel h2\<close>
- children_eq2_h finite_set_in finsert_iff funion_finsert_right CD.parent_child_rel_child
+ children_eq2_h finsert_iff funion_finsert_right CD.parent_child_rel_child
CD.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h
select_result_I2 subsetD sup_bot.right_neutral)
- by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap
+ by (metis (no_types, opaque_lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> assms(3) assms(4)
- children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in
+ children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h
is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes
returns_result_select_result subset_code(1))
then have "CD.a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "CD.a_all_ptrs_in_heap h'"
by (smt (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2
- disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in finsertCI funion_finsert_right
+ disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finsertCI funion_finsert_right
h' local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes
node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3
select_result_I2 set_ConsD subsetD subsetI)
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_character_data_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow>
cast new_character_data_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow>
cast new_character_data_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_character_data_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "CD.a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_character_data_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr
returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
thm children_eq2_h
using \<open>CD.a_distinct_lists h\<close> \<open>type_wf h2\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h'
set_disconnected_nodes_get_disconnected_nodes
by (metis (no_types, opaque_lifting)
\<open>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 \<notin> set disc_nodes_h3\<close> \<open>type_wf h2\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
by (smt (verit) \<open>local.CD.a_all_ptrs_in_heap h\<close> disc_nodes_document_ptr_h
disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff
- document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h'
+ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.CD.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> 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
\<open>local.CD.a_distinct_lists h2\<close> \<open>type_wf h'\<close> 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 \<open>local.CD.a_distinct_lists h3\<close>
\<open>type_wf h3\<close> children_eq2_h3 local.get_disconnected_nodes_ok returns_result_select_result)
qed
have "CD.a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
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)
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
- children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h'
+ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h h'
l_ptr_kinds_M.ptr_kinds_ptr_kinds_M
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
object_ptr_kinds_M_def
select_result_I2)
have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h"
using object_ptr_kinds_eq_h
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_eq_h:
"\<And>character_data_ptr shadow_root_opt. h \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^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:
"\<And>ptr' children. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>ptr' children. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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 \<open>type_wf h2\<close> 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 \<union> a_host_shadow_root_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
have "parent_child_rel h \<union> a_host_shadow_root_rel h =
parent_child_rel h2 \<union> a_host_shadow_root_rel h2"
using \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>parent_child_rel h = parent_child_rel h2\<close> by auto
have "parent_child_rel h2 \<union> a_host_shadow_root_rel h2 =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3"
using \<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> by auto
have "parent_child_rel h' \<union> a_host_shadow_root_rel h' =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3"
by (simp add: \<open>local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\<close>
\<open>parent_child_rel h3 = parent_child_rel h'\<close>)
have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3)"
using \<open>acyclic (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<close>
\<open>parent_child_rel h \<union> local.a_host_shadow_root_rel h = parent_child_rel h2 \<union>
local.a_host_shadow_root_rel h2\<close> \<open>parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 =
parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3\<close> by auto
then have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h')"
by(simp add: \<open>parent_child_rel h' \<union> a_host_shadow_root_rel h' =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3\<close>)
have "CD.a_heap_is_wellformed h'"
apply(simp add: CD.a_heap_is_wellformed_def)
by (simp add: \<open>local.CD.a_acyclic_heap h'\<close> \<open>local.CD.a_all_ptrs_in_heap h'\<close>
\<open>local.CD.a_distinct_lists h'\<close> \<open>local.CD.a_owner_document_valid h'\<close>)
show "heap_is_wellformed h' "
using \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h')\<close>
by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl
\<open>local.CD.a_heap_is_wellformed h'\<close> \<open>local.a_all_ptrs_in_heap h'\<close> \<open>local.a_distinct_lists h'\<close>
\<open>local.a_shadow_root_valid h'\<close>)
qed
end
subsubsection \<open>create\_document\<close>
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 \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes ::
"(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_document :: "((_) heap, exception, (_) document_ptr) prog"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
begin
lemma create_document_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_document \<rightarrow>\<^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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr" and
h': "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
using assms(2)
apply(simp add: create_document_def)
using new_document_ok by blast
have "new_document_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have "new_document_ptr |\<notin>| 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 |\<notin>| object_ptr_kinds h"
by simp
have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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)
+ by (auto simp add: document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h"
using object_ptr_kinds_eq
apply(simp add: shadow_root_ptr_kinds_def)
by force
have children_eq:
"\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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) \<open>\<And>thesis. (\<And>new_document_ptr.
\<lbrakk>h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr; h \<turnstile> new_document \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+
then have disconnected_nodes_eq2_h: "\<And>doc_ptr. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
using h' local.new_document_no_disconnected_nodes new_document_ptr by blast
have "type_wf h'"
using \<open>type_wf h\<close> new_document_types_preserved h' by blast
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h'"
by (simp add: object_ptr_kinds_eq)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_document_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
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 \<open>heap_is_wellformed h\<close> 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
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close> 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)
+ apply (metis (no_types, opaque_lifting)
\<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
- children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq
+ children_eq2 finsert_iff funion_finsert_right object_ptr_kinds_eq
select_result_I2 subsetD sup_bot.right_neutral)
by (metis (no_types, lifting) \<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close> \<open>type_wf h'\<close>
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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h'"
using \<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close>
\<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: children_eq2[symmetric] CD.a_distinct_lists_def insort_split
object_ptr_kinds_eq
document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(auto simp add: dest: distinct_concat_map_E)[1]
apply(auto simp add: dest: distinct_concat_map_E)[1]
using \<open>new_document_ptr |\<notin>| document_ptr_kinds h\<close>
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 \<noteq> y"
assume a2: "x |\<in>| document_ptr_kinds h"
assume a3: "x \<noteq> new_document_ptr"
assume a4: "y |\<in>| document_ptr_kinds h"
assume a5: "y \<noteq> new_document_ptr"
assume a6: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
assume a7: "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
assume a8: "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
have f9: "xa \<in> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using a7 a3 disconnected_nodes_eq2_h by presburger
have f10: "xa \<in> set |h \<turnstile> get_disconnected_nodes y|\<^sub>r"
using a8 a5 disconnected_nodes_eq2_h by presburger
have f11: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))"
using a4 by simp
have "x \<in> 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' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
and 1: "h' \<turnstile> 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) \<rightarrow>\<^sub>r []"
and 2: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
and 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
and 4: "(\<Union>x\<in>fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h). set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 5: "x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
and 7: "xa |\<in>| object_ptr_kinds h"
and 8: "xa \<noteq> 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 |\<in>| document_ptr_kinds h"
and 10: "xb \<noteq> new_document_ptr"
then show "False"
by (metis \<open>CD.a_distinct_lists h\<close> 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 \<open>heap_is_wellformed h\<close> 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 \<open>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 |\<notin>| object_ptr_kinds h\<close>
- children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff
+ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes funion_iff
node_ptr_kinds_eq object_ptr_kinds_eq)
have shadow_root_eq_h: "\<And>character_data_ptr shadow_root_opt.
h \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h' \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^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 \<open>type_wf h'\<close> 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:
"\<And>ptr' disc_nodes. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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 \<union> a_host_shadow_root_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
moreover
have "parent_child_rel h \<union> a_host_shadow_root_rel h =
parent_child_rel h' \<union> a_host_shadow_root_rel h'"
by (simp add: \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h'\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close>)
ultimately have "acyclic (parent_child_rel h' \<union> 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: \<open>local.CD.a_acyclic_heap h'\<close> \<open>local.CD.a_all_ptrs_in_heap h'\<close>
\<open>local.CD.a_distinct_lists h'\<close> \<open>local.CD.a_owner_document_valid h'\<close>)
show "heap_is_wellformed h'"
using CD.heap_is_wellformed_impl \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h')\<close>
\<open>local.CD.a_heap_is_wellformed h'\<close> \<open>local.a_all_ptrs_in_heap h'\<close> \<open>local.a_distinct_lists h'\<close>
\<open>local.a_shadow_root_valid h'\<close> 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 \<open>attach\_shadow\_root\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) 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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_shadow_root :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, unit) dom_prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_, unit) dom_prog set"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, unit) dom_prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> (_, unit) dom_prog set"
and attach_shadow_root :: "(_) element_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, (_) 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 \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^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 \<turnstile> get_tag_name element_ptr \<rightarrow>\<^sub>r element_tag_name" and
"element_tag_name \<in> safe_shadow_root_element_types" and
prev_shadow_root: "h \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r None" and
h2: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h2" and
new_shadow_root_ptr: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr" and
h3: "h2 \<turnstile> set_mode new_shadow_root_ptr new_mode \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> set_shadow_root element_ptr (Some new_shadow_root_ptr) \<rightarrow>\<^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 \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^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 \<open>element_tag_name \<in> safe_shadow_root_element_types\<close> apply(simp)
using \<open>element_tag_name \<in> safe_shadow_root_element_types\<close> 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 \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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="\<lambda>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="\<lambda>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 \<open>h \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^sub>r new_shadow_root_ptr\<close>
create_shadow_root_known_ptr by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| element_ptr_kinds h"
by (meson \<open>h \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^sub>r new_shadow_root_ptr\<close>
is_OK_returns_result_I local.attach_shadow_root_element_ptr_in_heap)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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="\<lambda>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="\<lambda>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: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
using h2 local.new_shadow_root_no_child_nodes new_shadow_root_ptr by auto
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h:
"\<And>ptr' disc_nodes. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h2 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_shadow_root_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h3:
"\<And>ptr' disc_nodes. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_shadow_root_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = 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 \<open>heap_is_wellformed h\<close> 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
\<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
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
\<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> 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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close> children_eq2_h
apply(auto simp add: select_result_eq[OF disconnected_nodes_eq_h] CD.a_distinct_lists_def
insort_split object_ptr_kinds_eq_h
document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I
dest: distinct_concat_map_E)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(auto simp add: dest: distinct_concat_map_E)[1]
apply(case_tac "x = cast new_shadow_root_ptr")
using \<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close> children_eq2_h apply blast
apply(case_tac "y = cast new_shadow_root_ptr")
using \<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close> children_eq2_h apply blast
proof -
fix x y :: "(_) object_ptr"
fix xa :: "(_) node_ptr"
assume a1: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
assume "x \<noteq> y"
assume "xa \<in> set |h2 \<turnstile> get_child_nodes x|\<^sub>r"
assume "xa \<in> set |h2 \<turnstile> get_child_nodes y|\<^sub>r"
assume "x |\<in>| object_ptr_kinds h"
assume "x \<noteq> 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 |\<in>| object_ptr_kinds h"
assume "y \<noteq> 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 \<open>x |\<in>| object_ptr_kinds h\<close> \<open>y |\<in>| object_ptr_kinds h\<close>
using \<open>xa \<in> set |h2 \<turnstile> get_child_nodes x|\<^sub>r\<close> \<open>xa \<in> set |h2 \<turnstile> get_child_nodes y|\<^sub>r\<close>
using \<open>x \<noteq> y\<close>
by(auto simp add: children_eq2_h[OF \<open>x \<noteq> 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\<close>]
children_eq2_h[OF \<open>y \<noteq> 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\<close>])
next
fix x :: "(_) node_ptr"
fix xa :: "(_) object_ptr"
fix xb :: "(_) document_ptr"
assume "(\<Union>x\<in>fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
assume "x \<in> set |h2 \<turnstile> get_child_nodes xa|\<^sub>r"
assume "xb |\<in>| document_ptr_kinds h"
assume "x \<in> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
assume "xa |\<in>| object_ptr_kinds h"
assume "xa \<noteq> 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 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
by (metis (no_types, lifting) CD.get_child_nodes_ok \<open>xa |\<in>| object_ptr_kinds h\<close>
\<open>xb |\<in>| document_ptr_kinds h\<close> 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 \<open>x \<in> set |h2 \<turnstile> get_child_nodes xa|\<^sub>r\<close> \<open>x \<in> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r\<close>
\<open>xa \<noteq> 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\<close> 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 \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
(* using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close> *)
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
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> assms(2) assms(3)
- children_eq2_h children_eq_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I
+ children_eq2_h children_eq_h document_ptr_kinds_eq_h is_OK_returns_result_I
l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.known_ptrs_known_ptr object_ptr_kinds_M_def
returns_result_select_result)
have shadow_root_eq_h:
"\<And>character_data_ptr shadow_root_opt. h \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^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:
"\<And>ptr' children. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>ptr' children. element_ptr \<noteq> ptr' \<Longrightarrow> h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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' \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r Some new_shadow_root_ptr"
using \<open>type_wf h3\<close> 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 \<open>type_wf h3\<close> 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 \<open>type_wf h2\<close> assms(1) assms(2) local.get_shadow_root_ok local.shadow_root_same_host
returns_result_select_result shadow_root_eq_h)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2
select_result_eq[OF shadow_root_eq_h2])
then have "a_distinct_lists h'"
apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3
select_result_eq[OF shadow_root_eq_h3])[1]
apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1]
by (smt (verit, best) ShadowRootMonad.ptr_kinds_ptr_kinds_M \<open>new_shadow_root_ptr \<notin> set |h \<turnstile>
shadow_root_ptr_kinds_M|\<^sub>r\<close> \<open>type_wf h'\<close> assms(1) assms(2) element_ptr_kinds_eq_h3
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.shadow_root_same_host local.get_shadow_root_ok
local.get_shadow_root_shadow_root_ptr_in_heap
local.l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms returns_result_select_result
select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_h3)
have "a_shadow_root_valid h"
using assms(1)
by (simp add: heap_is_wellformed_def)
then
have "a_shadow_root_valid h'"
proof(unfold a_shadow_root_valid_def; safe)
fix shadow_root_ptr
assume "\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h). \<exists>host\<in>fset (element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
assume "a_shadow_root_valid h"
assume "shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h')"
show "\<exists>host\<in>fset (element_ptr_kinds h').
|h' \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and> |h' \<turnstile> get_shadow_root host|\<^sub>r =
Some shadow_root_ptr"
proof (cases "shadow_root_ptr = new_shadow_root_ptr")
case True
have "element_ptr \<in> fset (element_ptr_kinds h')"
by (simp add: \<open>element_ptr |\<in>| element_ptr_kinds h\<close> element_ptr_kinds_eq_h
element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3)
moreover have "|h' \<turnstile> get_tag_name element_ptr|\<^sub>r \<in> safe_shadow_root_element_types"
by (smt (verit, best) \<open>\<And>thesis. (\<And>h2 h3 new_shadow_root_ptr element_tag_name. \<lbrakk>h \<turnstile>
get_tag_name element_ptr \<rightarrow>\<^sub>r element_tag_name; element_tag_name \<in>
safe_shadow_root_element_types; h \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r None; h \<turnstile>
new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h2; h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr; h2
\<turnstile> set_mode new_shadow_root_ptr new_mode \<rightarrow>\<^sub>h h3; h3 \<turnstile> set_shadow_root element_ptr (Some
new_shadow_root_ptr) \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> select_result_I2 tag_name_eq2_h2
tag_name_eq2_h3 tag_name_eq_h)
moreover have "|h' \<turnstile> 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 \<in> fset (element_ptr_kinds h)" and
"|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
using \<open>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h')\<close>
using \<open>\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h). \<exists>host\<in>fset (element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr\<close>
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)
+ by meson
moreover have "host \<noteq> element_ptr"
using calculation(3) prev_shadow_root by auto
ultimately show ?thesis
using element_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h
- by (metis (no_types, lifting) assms(2)
+ by (metis (no_types, opaque_lifting) assms(2)
l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_shadow_root_ok
- local.l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms notin_fset returns_result_select_result
+ local.l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms returns_result_select_result
select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 tag_name_eq2_h
tag_name_eq2_h2 tag_name_eq2_h3)
qed
qed
have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h
select_result_eq[OF shadow_root_eq_h])
have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2
select_result_eq[OF shadow_root_eq_h2])
have "a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \<union>
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 \<noteq> 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)
\<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close> \<open>type_wf h3\<close>
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) \<open>\<And>ptr'. (\<And>x. element_ptr \<noteq> ptr') \<Longrightarrow>
|h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r\<close> 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 \<open>\<And>ptr'. (\<And>x. element_ptr \<noteq> ptr') \<Longrightarrow>
|h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r\<close> \<open>type_wf h3\<close> case_prodE case_prodI
host_shadow_root_rel_shadow_root image_iff local.get_shadow_root_impl local.get_shadow_root_ok
mem_Collect_eq option.discI prev_shadow_root returns_result_select_result select_result_I2
shadow_root_eq_h shadow_root_eq_h2
by (smt (verit))
have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
have "parent_child_rel h \<union> a_host_shadow_root_rel h =
parent_child_rel h2 \<union> a_host_shadow_root_rel h2"
using \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>parent_child_rel h = parent_child_rel h2\<close> by auto
have "parent_child_rel h2 \<union> a_host_shadow_root_rel h2 =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3"
using \<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> by auto
have "parent_child_rel h' \<union> a_host_shadow_root_rel h' =
{(cast element_ptr, cast new_shadow_root_ptr)} \<union> parent_child_rel h3 \<union> a_host_shadow_root_rel h3"
by (simp add: \<open>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)} \<union>
local.a_host_shadow_root_rel h3\<close> \<open>parent_child_rel h3 = parent_child_rel h'\<close>)
have "\<And>a b. (a, b) \<in> parent_child_rel h3 \<Longrightarrow> a \<noteq> cast new_shadow_root_ptr"
using CD.parent_child_rel_parent_in_heap \<open>parent_child_rel h = parent_child_rel h2\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> 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 "\<And>a b. (a, b) \<in> a_host_shadow_root_rel h3 \<Longrightarrow> a \<noteq> 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 \<notin> {x. (x, cast element_ptr) \<in> (parent_child_rel h3 \<union>
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 \<union> a_host_shadow_root_rel h3)"
using \<open>acyclic (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<close>
\<open>parent_child_rel h \<union> local.a_host_shadow_root_rel h =
parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2\<close> \<open>parent_child_rel h2 \<union>
local.a_host_shadow_root_rel h2 = parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3\<close>
by auto
ultimately have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h')"
by(simp add: \<open>parent_child_rel h' \<union> a_host_shadow_root_rel h' =
{(cast element_ptr, cast new_shadow_root_ptr)} \<union> parent_child_rel h3 \<union>
a_host_shadow_root_rel h3\<close>)
have "CD.a_heap_is_wellformed h'"
apply(simp add: CD.a_heap_is_wellformed_def)
by (simp add: \<open>local.CD.a_acyclic_heap h'\<close> \<open>local.CD.a_all_ptrs_in_heap h'\<close>
\<open>local.CD.a_distinct_lists h'\<close> \<open>local.CD.a_owner_document_valid h'\<close>)
show "heap_is_wellformed h' "
using \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h')\<close>
by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \<open>local.CD.a_heap_is_wellformed h'\<close>
\<open>local.a_all_ptrs_in_heap h'\<close> \<open>local.a_distinct_lists h'\<close> \<open>local.a_shadow_root_valid h'\<close>)
qed
end
interpretation l_attach_shadow_root_wf?: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes
set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name
get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document
get_disconnected_document_locs set_val set_val_locs create_character_data DocumentClass.known_ptr
DocumentClass.type_wf set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root
by(auto simp add: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
end
diff --git a/thys/Shadow_DOM/classes/ShadowRootClass.thy b/thys/Shadow_DOM/classes/ShadowRootClass.thy
--- a/thys/Shadow_DOM/classes/ShadowRootClass.thy
+++ b/thys/Shadow_DOM/classes/ShadowRootClass.thy
@@ -1,460 +1,462 @@
(***********************************************************************************
* 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.
********************************************************************************\***)
section \<open>The Shadow DOM Data Model\<close>
theory ShadowRootClass
imports
Core_DOM.ShadowRootPointer
Core_DOM.DocumentClass
begin
subsection \<open>ShadowRoot\<close>
datatype shadow_root_mode = Open | Closed
record ('node_ptr, 'element_ptr, 'character_data_ptr) RShadowRoot = RObject +
nothing :: unit
mode :: shadow_root_mode
child_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'ShadowRoot) ShadowRoot
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'ShadowRoot option) RShadowRoot_scheme"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'ShadowRoot) ShadowRoot"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData, 'Document,
'ShadowRoot) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, ('node_ptr, 'element_ptr,
'character_data_ptr, 'ShadowRoot option)
RShadowRoot_ext + 'Object, 'Node, 'Element, 'CharacterData, 'Document) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData,
'Document, 'ShadowRoot) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData, 'Document, 'ShadowRoot) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
('node_ptr, 'element_ptr,
'character_data_ptr, 'ShadowRoot option) RShadowRoot_ext + 'Object, 'Node, 'Element,
'CharacterData, 'Document) heap"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot) heap"
type_synonym "heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l" = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition shadow_root_ptr_kinds :: "(_) heap \<Rightarrow> (_) shadow_root_ptr fset"
where
"shadow_root_ptr_kinds heap =
the |`| (cast |`| (ffilter is_shadow_root_ptr_kind (object_ptr_kinds heap)))"
lemma shadow_root_ptr_kinds_simp [simp]:
"shadow_root_ptr_kinds (Heap (fmupd (cast shadow_root_ptr) shadow_root (the_heap h))) =
{|shadow_root_ptr|} |\<union>| shadow_root_ptr_kinds h"
- apply(auto simp add: shadow_root_ptr_kinds_def)[1]
- by force
+ by (auto simp add: shadow_root_ptr_kinds_def)
definition shadow_root_ptrs :: "(_) heap \<Rightarrow> (_) shadow_root_ptr fset"
where
"shadow_root_ptrs heap = ffilter is_shadow_root_ptr (shadow_root_ptr_kinds heap)"
definition cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) Object \<Rightarrow> (_) ShadowRoot option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t obj = (case RObject.more obj of
Inr (Inr (Inl shadow_root)) \<Rightarrow> Some (RObject.extend (RObject.truncate obj) shadow_root)
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
definition cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:: "(_) ShadowRoot \<Rightarrow> (_) Object"
where
"cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t shadow_root =
(RObject.extend (RObject.truncate shadow_root) (Inr (Inr (Inl (RObject.more shadow_root)))))"
adhoc_overloading cast cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
definition is_shadow_root_kind :: "(_) Object \<Rightarrow> bool"
where
"is_shadow_root_kind ptr \<longleftrightarrow> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr \<noteq> None"
lemma shadow_root_ptr_kinds_heap_upd [simp]:
"shadow_root_ptr_kinds (Heap (fmupd (cast shadow_root_ptr) shadow_root (the_heap h))) =
{|shadow_root_ptr|} |\<union>| shadow_root_ptr_kinds h"
- apply(auto simp add: shadow_root_ptr_kinds_def)[1]
- by force
+ by (auto simp add: shadow_root_ptr_kinds_def)
lemma shadow_root_ptr_kinds_commutes [simp]:
"cast shadow_root_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
- apply(auto simp add: object_ptr_kinds_def shadow_root_ptr_kinds_def)[1]
- by (metis (no_types, lifting) shadow_root_ptr_casts_commute2 shadow_root_ptr_shadow_root_ptr_cast
- ffmember_filter fimage_eqI
- fset.map_comp option.sel)
+proof (rule iffI)
+ show "cast shadow_root_ptr |\<in>| object_ptr_kinds h \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
+ by (metis (no_types, opaque_lifting) finsertI1 finsert_absorb funion_finsert_left
+ funion_finsert_right put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs shadow_root_ptr_kinds_def
+ shadow_root_ptr_kinds_heap_upd sup_bot.right_neutral)
+next
+ show "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow> cast shadow_root_ptr |\<in>| object_ptr_kinds h"
+ by (auto simp add: object_ptr_kinds_def shadow_root_ptr_kinds_def)
+qed
definition get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) shadow_root_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) ShadowRoot option"
where
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Option.bind (get (cast shadow_root_ptr) h) cast"
adhoc_overloading get get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
locale l_type_wf_def\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (DocumentClass.type_wf h \<and> (\<forall>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h).
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t: "type_wf h \<Longrightarrow> ShadowRootClass.type_wf h"
sublocale l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t \<subseteq> l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
by (metis (full_types) ShadowRootClass.type_wf_def a_type_wf_def l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_axioms
l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma type_wf_implies_previous: "type_wf h \<Longrightarrow> DocumentClass.type_wf h"
by(simp add: type_wf_defs)
locale l_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas = l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
sublocale l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
lemma get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf:
assumes "type_wf h"
shows "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<longleftrightarrow> get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h \<noteq> None"
using l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
- by (metis is_none_bind is_none_simps(1) is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf notin_fset
+ by (metis is_none_bind is_none_simps(1) is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
shadow_root_ptr_kinds_commutes)
end
global_interpretation l_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas type_wf by unfold_locales
definition put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) shadow_root_ptr \<Rightarrow> (_) ShadowRoot \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root = put (cast shadow_root_ptr) (cast shadow_root)"
adhoc_overloading put put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
lemma put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap:
assumes "put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h = h'"
shows "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h'"
using assms
unfolding put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
by (metis shadow_root_ptr_kinds_commutes put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap)
lemma put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_put_ptrs:
assumes "put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast shadow_root_ptr|}"
using assms
by (simp add: put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs)
lemma cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_inject [simp]: "cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x = cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
by (metis (full_types) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_none [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t obj = None \<longleftrightarrow> \<not> (\<exists>shadow_root. cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t shadow_root = obj)"
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
split: sum.splits)[1]
by (metis (full_types) RObject.select_convs(2) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_some [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t obj = Some shadow_root \<longleftrightarrow> cast shadow_root = obj"
by(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def split: sum.splits)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_inv [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t (cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t shadow_root) = Some shadow_root"
by simp
lemma cast_shadow_root_not_node [simp]:
"cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t shadow_root \<noteq> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node"
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node \<noteq> cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t shadow_root"
by(auto simp add: cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
lemma get_shadow_root_ptr_simp1 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h) = Some shadow_root"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_shadow_root_ptr_simp2 [simp]:
"shadow_root_ptr \<noteq> shadow_root_ptr'
\<Longrightarrow> get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' shadow_root h) =
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_shadow_root_ptr_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_shadow_root_ptr_simp4 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_shadow_root_ptr_simp5 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_shadow_root_ptr_simp6 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_shadow_root_ptr_simp7 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_shadow_root_ptr_simp8 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
abbreviation "create_shadow_root_obj mode_arg child_nodes_arg
\<equiv> \<lparr> RObject.nothing = (), RShadowRoot.nothing = (), mode = mode_arg,
RShadowRoot.child_nodes = child_nodes_arg, \<dots> = None \<rparr>"
definition new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_)heap \<Rightarrow> ((_) shadow_root_ptr \<times> (_) heap)"
where
"new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (let new_shadow_root_ptr =
shadow_root_ptr.Ref (Suc (fMax (shadow_root_ptr.the_ref |`| (shadow_root_ptrs h)))) in
(new_shadow_root_ptr, put new_shadow_root_ptr (create_shadow_root_obj Open []) h))"
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "new_shadow_root_ptr |\<in>| shadow_root_ptr_kinds h'"
using assms
unfolding new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def
using put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap by blast
lemma new_shadow_root_ptr_new:
"shadow_root_ptr.Ref (Suc (fMax (finsert 0 (shadow_root_ptr.the_ref |`| shadow_root_ptrs h)))) |\<notin>|
shadow_root_ptrs h"
by (metis Suc_n_not_le_n shadow_root_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_not_in_heap:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "new_shadow_root_ptr |\<notin>| shadow_root_ptr_kinds h"
using assms
unfolding new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
by (metis Pair_inject shadow_root_ptrs_def fMax_finsert fempty_iff ffmember_filter
fimage_is_fempty is_shadow_root_ptr_ref max_0L new_shadow_root_ptr_new)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_shadow_root_ptr|}"
using assms
by (metis Pair_inject new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_put_ptrs)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_is_shadow_root_ptr:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "is_shadow_root_ptr new_shadow_root_ptr"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
assumes "ptr \<noteq> cast new_shadow_root_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
apply(simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
apply(simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
assumes "ptr \<noteq> new_shadow_root_ptr"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
locale l_known_ptr\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_shadow_root_ptr ptr)"
lemma known_ptr_not_shadow_root_ptr: "a_known_ptr ptr \<Longrightarrow> \<not>is_shadow_root_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
lemma known_ptr_new_shadow_root_ptr: "a_known_ptr ptr \<Longrightarrow> \<not>known_ptr ptr \<Longrightarrow> is_shadow_root_ptr ptr"
using l_known_ptr\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t.known_ptr_not_shadow_root_ptr by blast
end
global_interpretation l_known_ptr\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by (simp add: less_eq_fset.rep_eq local.a_known_ptrs_def subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs [instances]: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset
known_ptrs_new_ptr
by blast
lemma shadow_root_get_put_1 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h) = Some shadow_root"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma shadow_root_different_get_put [simp]:
"shadow_root_ptr \<noteq> shadow_root_ptr' \<Longrightarrow>
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' shadow_root h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma shadow_root_get_put_2 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma shadow_root_get_put_3 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t shadow_root_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma shadow_root_get_put_4 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma shadow_root_get_put_5 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t character_data_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a shadow_root_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t character_data_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma shadow_root_get_put_6 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma shadow_root_get_put_7 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t document_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t shadow_root_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma known_ptrs_implies: "DocumentClass.known_ptrs h \<Longrightarrow> ShadowRootClass.known_ptrs h"
by(auto simp add: DocumentClass.known_ptrs_defs DocumentClass.known_ptr_defs
ShadowRootClass.known_ptrs_defs ShadowRootClass.known_ptr_defs)
definition delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) shadow_root_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) heap option" where
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr = delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast shadow_root_ptr)"
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_pointer_removed:
assumes "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = Some h'"
shows "ptr |\<notin>| shadow_root_ptr_kinds h'"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_removed delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def shadow_root_ptr_kinds_def
split: if_splits)
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_pointer_ptr_in_heap:
assumes "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = Some h'"
shows "ptr |\<in>| shadow_root_ptr_kinds h"
using assms
apply(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_ptr_in_heap delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def split: if_splits)[1]
using delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_ptr_in_heap
by fastforce
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok:
assumes "ptr |\<in>| shadow_root_ptr_kinds h"
shows "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h \<noteq> None"
using assms
by (simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma shadow_root_delete_get_1 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h' = None"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
lemma shadow_root_different_delete_get [simp]:
"shadow_root_ptr \<noteq> shadow_root_ptr' \<Longrightarrow> delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' h = Some h' \<Longrightarrow>
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h' = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
lemma shadow_root_delete_get_2 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> object_ptr \<noteq> cast shadow_root_ptr \<Longrightarrow>
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h' = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
lemma shadow_root_delete_get_3 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h' = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma shadow_root_delete_get_4 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h' = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma shadow_root_delete_get_5 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow>
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h' = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma shadow_root_delete_get_6 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h' = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma shadow_root_delete_get_7 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> shadow_root_ptr' \<noteq> shadow_root_ptr \<Longrightarrow>
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' h' = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' h"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
end
diff --git a/thys/Shadow_DOM/monads/ShadowRootMonad.thy b/thys/Shadow_DOM/monads/ShadowRootMonad.thy
--- a/thys/Shadow_DOM/monads/ShadowRootMonad.thy
+++ b/thys/Shadow_DOM/monads/ShadowRootMonad.thy
@@ -1,711 +1,711 @@
(***********************************************************************************
* 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\<open>Shadow Root Monad\<close>
theory ShadowRootMonad
imports
"Core_DOM.DocumentMonad"
"../classes/ShadowRootClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot, 'result) dom_prog"
global_interpretation l_ptr_kinds_M shadow_root_ptr_kinds defines shadow_root_ptr_kinds_M = a_ptr_kinds_M .
lemmas shadow_root_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma shadow_root_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> shadow_root_ptr_kinds_M|\<^sub>r = |h' \<turnstile> shadow_root_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: shadow_root_ptr_kinds_M_defs object_ptr_kinds_M_defs shadow_root_ptr_kinds_def)
global_interpretation l_dummy defines get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = "l_get_M.a_get_M get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t type_wf shadow_root_ptr_kinds"
apply(simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf l_get_M_def)
by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs bind_eq_None_conv
shadow_root_ptr_kinds_commutes get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def option.simps(3))
lemmas get_M_defs = get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
locale l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas = l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
sublocale l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_get_M get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t type_wf shadow_root_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf local.type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t)
by (meson ShadowRootMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok = get_M_ok[folded get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def]
lemmas 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 = get_M_ptr_in_heap[folded get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf shadow_root_ptr_kinds get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t rewrites
"a_get_M = get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t" defines put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
locale l_put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas = l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
sublocale l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_put_M type_wf shadow_root_ptr_kinds get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
apply(unfold_locales)
apply (simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf local.type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t)
by (meson ShadowRootMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok = put_M_ok[folded put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas type_wf by unfold_locales
lemma shadow_root_put_get [simp]: "h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Mshadow_root_preserved1 [simp]:
"shadow_root_ptr \<noteq> shadow_root_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma shadow_root_put_get_preserved [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
apply(cases "shadow_root_ptr = shadow_root_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved2 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs NodeMonad.get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved3 [simp]:
"cast shadow_root_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved4 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast shadow_root_ptr \<noteq> object_ptr")[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved5 [simp]:
"cast shadow_root_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved7 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved8 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved10 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast shadow_root_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma new_element_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
definition 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 \<Rightarrow> (_, unit) dom_prog" where
"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 = do {
h \<leftarrow> get_heap;
(case delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h of
Some h \<Rightarrow> return_heap h |
None \<Rightarrow> error HierarchyRequestError)
}"
adhoc_overloading delete_M delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ok [simp]:
assumes "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
shows "h \<turnstile> ok (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)"
using assms
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: prod.splits)
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
shows "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
using assms
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
lemma 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:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
shows "shadow_root_ptr |\<notin>| shadow_root_ptr_kinds h'"
using assms
apply(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)[1]
by (metis comp_apply fmdom_notI fmdrop_lookup heap.sel object_ptr_kinds_def shadow_root_ptr_kinds_commutes)
lemma delete_shadow_root_pointers:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
shows "object_ptr_kinds h = object_ptr_kinds h' |\<union>| {|cast shadow_root_ptr|}"
using assms
apply(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def split: option.splits)[1]
apply (metis (no_types, lifting) ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_pointer_ptr_in_heap fmlookup_drop get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel option.sel
shadow_root_ptr_kinds_commutes)
using delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_pointer_ptr_in_heap apply blast
by (metis (no_types, lifting) ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_pointer_ptr_in_heap fmlookup_drop get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel option.sel
shadow_root_ptr_kinds_commutes)
lemma delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t: "h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
ptr \<noteq> cast shadow_root_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e: "h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def NodeMonad.get_M_defs ObjectMonad.get_M_defs
preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t: "h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def ElementMonad.get_M_defs NodeMonad.get_M_defs
ObjectMonad.get_M_defs preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a: "h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def CharacterDataMonad.get_M_defs
NodeMonad.get_M_defs ObjectMonad.get_M_defs preserved_def split: prod.splits option.splits if_splits
elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t: "h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def DocumentMonad.get_M_defs ObjectMonad.get_M_defs
preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma 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: "h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
shadow_root_ptr \<noteq> shadow_root_ptr' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get_M_defs ObjectMonad.get_M_defs
preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma shadow_root_put_get_1 [simp]: "shadow_root_ptr \<noteq> shadow_root_ptr' \<Longrightarrow>
h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma shadow_root_put_get_2 [simp]: "(\<And>x. getter (setter (\<lambda>_. v) x) = getter x) \<Longrightarrow>
h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
by (cases "shadow_root_ptr = shadow_root_ptr'") (auto simp add: put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma shadow_root_put_get_3 [simp]: "h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_4 [simp]: "h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_5 [simp]: "h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_6 [simp]: "h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_7 [simp]: "h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: put_M_defs DocumentMonad.get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_8 [simp]: "h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow>
preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: DocumentMonad.put_M_defs get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_9 [simp]: "(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x)) \<Longrightarrow>
h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by (cases "cast shadow_root_ptr = object_ptr") (auto simp add: put_M_defs get_M_defs
ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv split: option.splits)
subsection \<open>new\_M\<close>
definition new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M :: "(_, (_) shadow_root_ptr) dom_prog"
where
"new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h);
return_heap h';
return new_ptr
}"
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ok [simp]:
"h \<turnstile> ok new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def split: prod.splits)
lemma 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:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "new_shadow_root_ptr |\<in>| shadow_root_ptr_kinds h'"
using assms
unfolding new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def
by(auto simp add: 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 put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma 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:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "new_shadow_root_ptr |\<notin>| shadow_root_ptr_kinds h"
using assms new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_not_in_heap
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_shadow_root_ptr|}"
using assms new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma 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:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "is_shadow_root_ptr new_shadow_root_ptr"
using assms new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_is_shadow_root_ptr
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def elim!: bind_returns_result_E split: prod.splits)
lemma new_shadow_root_mode:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "h' \<turnstile> get_M new_shadow_root_ptr mode \<rightarrow>\<^sub>r Open"
using assms
by(auto simp add: get_M_defs 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
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_children:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "h' \<turnstile> get_M new_shadow_root_ptr child_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs 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
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> ptr \<noteq> cast new_shadow_root_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def CharacterDataMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def DocumentMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma 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:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> ptr \<noteq> new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection \<open>modified heaps\<close>
lemma shadow_root_get_put_1 [simp]: "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = cast shadow_root_ptr
then cast obj else get shadow_root_ptr h)"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def split: option.splits Option.bind_splits)
lemma shadow_root_ptr_kinds_new[simp]: "shadow_root_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = shadow_root_ptr_kinds h |\<union>|
(if is_shadow_root_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: shadow_root_ptr_kinds_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "DocumentClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_shadow_root_ptr_kind ptr \<Longrightarrow> is_shadow_root_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs is_shadow_root_kind_def split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: DocumentMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "DocumentClass.type_wf h"
assumes "is_shadow_root_ptr_kind ptr \<Longrightarrow> is_shadow_root_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: DocumentMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1]
by (metis (no_types, opaque_lifting) ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
- bind.bind_lunit finite_set_in get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def option.exhaust_sel)
+ bind.bind_lunit get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def option.exhaust_sel)
subsection \<open>type\_wf\<close>
lemma new_element_type_wf_preserved [simp]:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_element_ptr where "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using new_element_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by auto
with assms show ?thesis
by(auto simp add: ElementMonad.new_element_def type_wf_defs Let_def elim!: bind_returns_heap_E
split: prod.splits)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr RElement.child_nodes_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma new_character_data_type_wf_preserved [simp]:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_character_data_ptr where "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using new_character_data_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by auto
with assms show ?thesis
by(auto simp add: CharacterDataMonad.new_character_data_def type_wf_defs Let_def
elim!: bind_returns_heap_E split: prod.splits)
qed
lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_val_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: CharacterDataMonad.put_M_defs type_wf_defs)
qed
lemma new_document_type_wf_preserved [simp]:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_document_ptr where "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_document_ptr|}"
using new_document_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by auto
with assms show ?thesis
by(auto simp add: DocumentMonad.new_document_def type_wf_defs Let_def elim!: bind_returns_heap_E
split: prod.splits)
qed
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doctype_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M document_ptr doctype_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: DocumentMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M document_ptr document_element_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: DocumentMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M document_ptr disconnected_nodes_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: DocumentMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_mode_type_wf_preserved [simp]:
"h \<turnstile> put_M shadow_root_ptr mode_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
by(auto simp add: get_M_defs is_shadow_root_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs DocumentClass.type_wf_defs put_M_defs
put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: option.splits)
lemma put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M shadow_root_ptr RShadowRoot.child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
by(auto simp add: get_M_defs is_shadow_root_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs DocumentClass.type_wf_defs put_M_defs
put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: option.splits)
lemma shadow_root_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(simp add: shadow_root_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms])
lemma shadow_root_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow>
(\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def shadow_root_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma new_shadow_root_known_ptr:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "known_ptr (cast new_shadow_root_ptr)"
using assms
apply(auto simp add: 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 a_known_ptr_def
elim!: bind_returns_result_E2 split: prod.splits)[1]
using assms 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 by blast
lemma new_shadow_root_type_wf_preserved [simp]: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: 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 put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ShadowRootClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ShadowRootClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
ShadowRootClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e ShadowRootClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_not_in_heap
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I DocumentMonad.type_wf_put_I ElementMonad.type_wf_put_I
CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
by(auto simp add: type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_shadow_root_kind_def is_document_kind_def
split: option.splits)[1]
locale l_new_shadow_root = l_type_wf +
assumes new_shadow_root_types_preserved: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_shadow_root_is_l_new_shadow_root [instances]: "l_new_shadow_root type_wf"
using l_new_shadow_root.intro new_shadow_root_type_wf_preserved
by blast
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
assumes "\<And>shadow_root_ptr. preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr RShadowRoot.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3) assms(4) assms(5)]
allI[OF assms(6), of id, simplified] shadow_root_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs shadow_root_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply(force)
apply(force)
done
lemma new_element_is_l_new_element [instances]:
"l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma new_character_data_is_l_new_character_data [instances]:
"l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma new_document_is_l_new_document [instances]:
"l_new_document type_wf"
using l_new_document.intro new_document_type_wf_preserved
by blast
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
\<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
\<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
\<forall>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
\<forall>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
\<forall>shadow_root_ptr. preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr RShadowRoot.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs)[1]
using type_wf_drop
apply blast
- by (metis (no_types, lifting) DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
- ElementMonad.type_wf_drop fmember_iff_member_fset fmlookup_drop get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
+ by (metis (no_types, opaque_lifting) DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
+ ElementMonad.type_wf_drop fmlookup_drop get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
object_ptr_kinds_code5 shadow_root_ptr_kinds_commutes)
lemma delete_shadow_root_type_wf_preserved [simp]:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
assumes "type_wf h"
shows "type_wf h'"
using assms
using type_wf_drop
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
end
diff --git a/thys/Shadow_SC_DOM/Shadow_DOM.thy b/thys/Shadow_SC_DOM/Shadow_DOM.thy
--- a/thys/Shadow_SC_DOM/Shadow_DOM.thy
+++ b/thys/Shadow_SC_DOM/Shadow_DOM.thy
@@ -1,12913 +1,12903 @@
(***********************************************************************************
* 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\<open>The Shadow DOM\<close>
theory Shadow_DOM
imports
"monads/ShadowRootMonad"
Core_SC_DOM.Core_DOM
begin
abbreviation "safe_shadow_root_element_types \<equiv> {''article'', ''aside'', ''blockquote'', ''body'',
''div'', ''footer'', ''h1'', ''h2'', ''h3'', ''h4'', ''h5'', ''h6'', ''header'', ''main'',
''nav'', ''p'', ''section'', ''span''}"
subsection \<open>Function Definitions\<close>
subsubsection \<open>get\_child\_nodes\<close>
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 \<Rightarrow> unit
\<Rightarrow> (_, (_) 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 \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> unit
\<Rightarrow> (_, (_) node_ptr list) dom_prog)) list" where
"a_get_child_nodes_tups \<equiv> [(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 \<circ> the \<circ> cast)]"
definition a_get_child_nodes :: "(_) object_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" where
"a_get_child_nodes_locs ptr \<equiv>
(if is_shadow_root_ptr_kind ptr
then {preserved (get_M (the (cast ptr)) RShadowRoot.child_nodes)} else {}) \<union>
CD.a_get_child_nodes_locs ptr"
definition first_child :: "(_) object_ptr \<Rightarrow> (_, (_) node_ptr option) dom_prog"
where
"first_child ptr = do {
children \<leftarrow> a_get_child_nodes ptr;
return (case children of [] \<Rightarrow> None | child#_ \<Rightarrow> 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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_child_nodes ptr)"
using assms[unfolded known_ptr_impl type_wf_impl]
apply(auto simp add: get_child_nodes_def)[1]
apply(split CD.get_child_nodes_splits, rule conjI)+
using ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t CD.get_child_nodes_ok CD.known_ptr_impl CD.type_wf_impl
apply blast
apply(auto simp add: CD.known_ptr_impl a_get_child_nodes_tups_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok
dest!: known_ptr_new_shadow_root_ptr intro!: bind_is_OK_I2)[1]
by(auto dest: get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok split: option.splits)
lemma get_child_nodes_ptr_in_heap:
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
by(auto simp add: get_child_nodes_def invoke_ptr_in_heap dest: is_OK_returns_result_I)
lemma get_child_nodes_pure [simp]:
"pure (get_child_nodes ptr) h"
apply (auto simp add: get_child_nodes_def a_get_child_nodes_tups_def)[1]
apply(split CD.get_child_nodes_splits, rule conjI)+
apply(simp)
apply(split invoke_splits, rule conjI)+
apply(simp)
by(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I)
lemma get_child_nodes_reads: "reads (get_child_nodes_locs ptr) (get_child_nodes ptr) h h'"
apply (simp add: get_child_nodes_def a_get_child_nodes_tups_def get_child_nodes_locs_def
CD.get_child_nodes_locs_def)
apply(split CD.get_child_nodes_splits, rule conjI)+
apply(auto intro!: reads_subset[OF CD.get_child_nodes_reads[unfolded CD.get_child_nodes_locs_def]]
split: if_splits)[1]
apply(split invoke_splits, rule conjI)+
apply(auto)[1]
apply(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro: reads_subset[OF reads_singleton] reads_subset[OF check_in_heap_reads]
intro!: reads_bind_pure reads_subset[OF return_reads] split: option.splits)[1]
done
end
interpretation i_get_child_nodes?: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf
DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes
Core_DOM_Functions.get_child_nodes_locs
by(simp add: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)
declare l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_child_nodes_is_l_get_child_nodes [instances]: "l_get_child_nodes type_wf known_ptr
get_child_nodes get_child_nodes_locs"
apply(auto simp add: l_get_child_nodes_def instances)[1]
using get_child_nodes_reads get_child_nodes_ok get_child_nodes_ptr_in_heap get_child_nodes_pure
by blast+
paragraph \<open>new\_document\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_child_nodes_new_document:
"ptr' \<noteq> cast new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^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 \<open>new\_shadow\_root\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_child_nodes_new_shadow_root:
"ptr' \<noteq> cast new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
apply(auto simp add: get_child_nodes_def )[1]
apply(split CD.get_child_nodes_splits, rule conjI)+
apply(auto simp add: CD.get_child_nodes_def CD.a_get_child_nodes_tups_def)[1]
apply(split invoke_splits, rule conjI)+
using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr
known_ptr_not_element_ptr local.CD.known_ptr_impl apply blast
apply(auto simp add: is_document_ptr_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: option.splits document_ptr.splits)[1]
apply(auto simp add: is_character_data_ptr_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: option.splits document_ptr.splits)[1]
apply(auto simp add: is_element_ptr_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: option.splits document_ptr.splits)[1]
apply(auto simp add: a_get_child_nodes_tups_def)[1]
apply(split invoke_splits, rule conjI)+
apply(auto simp add: is_shadow_root_ptr_def split: shadow_root_ptr.splits
dest!: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_is_shadow_root_ptr)[1]
apply(auto intro!: bind_pure_returns_result_I)[1]
apply(drule(1) new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap)
apply(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)[1]
apply (metis check_in_heap_ptr_in_heap is_OK_returns_result_E old.unit.exhaust)
using new_shadow_root_children
by (simp add: new_shadow_root_children get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
end
interpretation i_new_shadow_root_get_child_nodes?:
l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs
DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes
Core_DOM_Functions.get_child_nodes_locs
by(unfold_locales)
declare l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances]
locale l_new_shadow_root_get_child_nodes = l_get_child_nodes +
assumes get_child_nodes_new_shadow_root:
"ptr' \<noteq> cast new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> r h h'"
assumes new_shadow_root_no_child_nodes:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^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 \<open>new\_element\<close>
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' \<noteq> cast new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []"
apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def
split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)[1]
apply(split CD.get_child_nodes_splits, rule conjI)+
using local.new_element_no_child_nodes apply auto[1]
apply(auto simp add: invoke_def)[1]
using case_optionE apply fastforce
apply(auto simp add: new_element_ptr_in_heap get_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def check_in_heap_def
new_element_child_nodes intro!: bind_pure_returns_result_I
intro: new_element_is_element_ptr elim!: new_element_ptr_in_heap)[1]
proof -
assume " h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
assume "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
assume "\<not> is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr)"
assume "\<not> 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 \<open>h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr\<close>
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 \<open>delete\_shadow\_root\<close>
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' \<noteq> cast shadow_root_ptr \<Longrightarrow> h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
r \<in> get_child_nodes_locs ptr' \<Longrightarrow> r h h'"
by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
split: if_splits intro: is_shadow_root_ptr_kind_obtains
intro: is_shadow_root_ptr_kind_obtains delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
simp add: shadow_root_ptr_casts_commute3 delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
intro!: delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t dest: document_ptr_casts_commute3
split: option.splits)
end
locale l_delete_shadow_root_get_child_nodes = l_get_child_nodes_defs +
assumes get_child_nodes_delete_shadow_root:
"ptr' \<noteq> cast shadow_root_ptr \<Longrightarrow> h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow>
r \<in> get_child_nodes_locs ptr' \<Longrightarrow> 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 \<open>set\_child\_nodes\<close>
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 \<Rightarrow> (_) node_ptr list
\<Rightarrow> (_, 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 \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> (_) node_ptr list
\<Rightarrow> (_, unit) dom_prog)) list" where
"a_set_child_nodes_tups \<equiv> [(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 \<circ> the \<circ> cast)]"
definition a_set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, 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 \<Rightarrow> (_, unit) dom_prog set"
where
"a_set_child_nodes_locs ptr \<equiv>
(if is_shadow_root_ptr_kind ptr then all_args (put_M (the (cast ptr)) RShadowRoot.child_nodes_update) else {}) \<union>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, unit) dom_prog"
and set_child_nodes_locs :: "(_) object_ptr \<Rightarrow> (_, 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 \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, 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 \<Rightarrow> (_, unit) dom_prog set" +
assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr"
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
assumes set_child_nodes_impl: "set_child_nodes = a_set_child_nodes"
assumes set_child_nodes_locs_impl: "set_child_nodes_locs = a_set_child_nodes_locs"
begin
lemmas set_child_nodes_def = set_child_nodes_impl[unfolded a_set_child_nodes_def set_child_nodes_def]
lemmas set_child_nodes_locs_def =set_child_nodes_locs_impl[unfolded a_set_child_nodes_locs_def
set_child_nodes_locs_def, folded CD.set_child_nodes_locs_impl]
lemma set_child_nodes_writes: "writes (set_child_nodes_locs ptr) (set_child_nodes ptr children) h h'"
apply (simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes_locs_def)
apply(split CD.set_child_nodes_splits, rule conjI)+
apply (simp add: CD.set_child_nodes_writes writes_union_right_I)
apply(split invoke_splits, rule conjI)+
apply(auto simp add: a_set_child_nodes_def)[1]
apply(auto simp add: set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: writes_bind_pure
intro: writes_union_right_I writes_union_left_I split: list.splits)[1]
by (metis is_shadow_root_ptr_implies_kind option.case_eq_if)
lemma set_child_nodes_pointers_preserved:
assumes "w \<in> set_child_nodes_locs object_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_child_nodes_locs object_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] type_wf_impl
by(auto simp add: all_args_def a_set_child_nodes_tups_def set_child_nodes_locs_def CD.set_child_nodes_locs_def
split: if_splits option.splits)
end
interpretation
i_set_child_nodes?: l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf
DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes
Core_DOM_Functions.set_child_nodes_locs
apply(unfold_locales)
by (auto simp add: set_child_nodes_def set_child_nodes_locs_def)
declare l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma set_child_nodes_is_l_set_child_nodes [instances]: "l_set_child_nodes type_wf
set_child_nodes set_child_nodes_locs"
apply(auto simp add: l_set_child_nodes_def instances)[1]
using set_child_nodes_writes apply fast
using set_child_nodes_pointers_preserved apply(fast, fast)
using set_child_nodes_types_preserved apply(fast, fast)
done
paragraph \<open>get\_child\_nodes\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap, exception, unit) prog set"
begin
lemma set_child_nodes_get_child_nodes:
assumes "known_ptr ptr"
assumes "type_wf h"
assumes "h \<turnstile> set_child_nodes ptr children \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
proof -
have "h \<turnstile> check_in_heap ptr \<rightarrow>\<^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 |\<in>| 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' \<turnstile> check_in_heap ptr \<rightarrow>\<^sub>r ()"
using check_in_heap_reads set_child_nodes_writes assms(3) \<open>h \<turnstile> check_in_heap ptr \<rightarrow>\<^sub>r ()\<close>
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 |\<in>| object_ptr_kinds h'"
using check_in_heap_ptr_in_heap by blast
with assms ptr_in_h \<open>type_wf h'\<close> 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 \<noteq> ptr'"
assumes "w \<in> set_child_nodes_locs ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> 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 \<open>set\_tag\_type\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_tag_name :: "(_) element_ptr \<Rightarrow> tag_name \<Rightarrow> (_, unit) dom_prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> (_, 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<in> set_tag_name_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_tag_name_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<open>get\_child\_nodes\<close>
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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_shadow\_root\<close>
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 \<Rightarrow> (_, (_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
where
"a_get_shadow_root_locs element_ptr \<equiv> {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 \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
fixes get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r children"
shows "element_ptr |\<in>| 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (get_shadow_root element_ptr)"
assumes get_shadow_root_ptr_in_heap: "h \<turnstile> ok (get_shadow_root element_ptr) \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h"
assumes get_shadow_root_pure [simp]: "pure (get_shadow_root element_ptr) h"
lemma get_shadow_root_is_l_get_shadow_root [instances]: "l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs"
using instances
apply(auto simp add: l_get_shadow_root_def)[1]
using get_shadow_root_reads apply blast
using get_shadow_root_ok apply blast
using get_shadow_root_ptr_in_heap apply blast
done
paragraph \<open>set\_disconnected\_nodes\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma set_disconnected_nodes_get_shadow_root:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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: "\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_tag\_type\<close>
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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_child\_nodes\<close>
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 \<Rightarrow> bool"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma set_child_nodes_get_shadow_root: "\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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: "\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>delete\_shadow\_root\<close>
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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>new\_character\_data\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_character_data:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>new\_document\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_document:
"h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>new\_element\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_element:
"ptr' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_shadow_root new_element_ptr \<rightarrow>\<^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' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr
\<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> r h h'"
assumes new_element_no_shadow_root:
"h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_shadow_root new_element_ptr \<rightarrow>\<^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 \<open>new\_shadow\_root\<close>
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 \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_shadow_root_new_shadow_root:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow> 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 \<open>set\_shadow\_root\<close>
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 \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, 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 \<Rightarrow> ((_, unit) dom_prog) set"
where
"a_set_shadow_root_locs element_ptr \<equiv> 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 \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, unit) dom_prog"
fixes set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_, 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 \<Rightarrow> bool"
and set_shadow_root :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, unit) dom_prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_, 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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<turnstile> ok (set_shadow_root element_ptr shadow_root) \<Longrightarrow> element_ptr |\<in>| 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 \<in> set_shadow_root_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_shadow_root_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (set_shadow_root element_ptr shadow_root)"
assumes set_shadow_root_ptr_in_heap:
"h \<turnstile> ok (set_shadow_root element_ptr shadow_root) \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h"
assumes set_shadow_root_pointers_preserved:
"w \<in> set_shadow_root_locs element_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h'"
assumes set_shadow_root_types_preserved:
"w \<in> set_shadow_root_locs element_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> 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 \<open>get\_shadow\_root\<close>
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 \<Longrightarrow> h \<turnstile> set_shadow_root ptr shadow_root_ptr_opt \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_shadow_root ptr \<rightarrow>\<^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 \<noteq> ptr' \<Longrightarrow> \<forall>w \<in> set_shadow_root_locs ptr.
(h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<Longrightarrow> h \<turnstile> set_shadow_root ptr shadow_root_ptr_opt \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r shadow_root_ptr_opt"
assumes set_shadow_root_get_shadow_root_different_pointers:
"ptr \<noteq> ptr' \<Longrightarrow> w \<in> set_shadow_root_locs ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_shadow_root_locs ptr' \<Longrightarrow>
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 \<open>set\_mode\<close>
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 \<Rightarrow> shadow_root_mode \<Rightarrow> (_, 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 \<Rightarrow> ((_, unit) dom_prog) set"
where
"a_set_mode_locs shadow_root_ptr \<equiv> 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 \<Rightarrow> shadow_root_mode \<Rightarrow> (_, unit) dom_prog"
fixes set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> (_, 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 \<Rightarrow> bool"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, unit) dom_prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> (_, 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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<turnstile> ok (set_mode shadow_root_ptr shadow_root_mode) \<Longrightarrow> shadow_root_ptr |\<in>| 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 \<in> set_mode_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_mode_locs element_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (set_mode shadow_root_ptr shadow_root_mode)"
assumes set_mode_ptr_in_heap:
"h \<turnstile> ok (set_mode shadow_root_ptr shadow_root_mode) \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
assumes set_mode_pointers_preserved:
"w \<in> set_mode_locs shadow_root_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h'"
assumes set_mode_types_preserved:
"w \<in> set_mode_locs shadow_root_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> 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 \<open>get\_child\_nodes\<close>
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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_shadow\_root\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_child\_nodes\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_host\<close>
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 \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_get_host :: "(_) shadow_root_ptr \<Rightarrow> (_, (_) element_ptr) dom_prog"
where
"a_get_host shadow_root_ptr = do {
host_ptrs \<leftarrow> element_ptr_kinds_M \<bind> filter_M (\<lambda>element_ptr. do {
shadow_root_opt \<leftarrow> get_shadow_root element_ptr;
return (shadow_root_opt = Some shadow_root_ptr)
});
(case host_ptrs of host_ptr#[] \<Rightarrow> return host_ptr | _ \<Rightarrow> error HierarchyRequestError)
}"
definition "a_get_host_locs \<equiv> (\<Union>element_ptr. (get_shadow_root_locs element_ptr)) \<union>
(\<Union>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 \<Rightarrow> (_, (_) element_ptr) dom_prog"
fixes get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<open>get\_mode\<close>
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 \<Rightarrow> (_, 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
where
"a_get_mode_locs shadow_root_ptr \<equiv> {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 \<Rightarrow> (_, shadow_root_mode) dom_prog"
fixes get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> ((_) heap, exception, shadow_root_mode) prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf :: "(_) heap \<Rightarrow> 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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<turnstile> get_mode shadow_root_ptr \<rightarrow>\<^sub>r children"
shows "shadow_root_ptr |\<in>| 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 \<Longrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow>
h \<turnstile> ok (get_mode shadow_root_ptr)"
assumes get_mode_ptr_in_heap: "h \<turnstile> ok (get_mode shadow_root_ptr) \<Longrightarrow> shadow_root_ptr |\<in>| 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 \<open>get\_shadow\_root\_safe\<close>
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 \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_mode :: "(_) shadow_root_ptr \<Rightarrow> (_, shadow_root_mode) dom_prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_get_shadow_root_safe :: "(_) element_ptr \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
where
"a_get_shadow_root_safe element_ptr = do {
shadow_root_ptr_opt \<leftarrow> get_shadow_root element_ptr;
(case shadow_root_ptr_opt of
Some shadow_root_ptr \<Rightarrow> do {
mode \<leftarrow> get_mode shadow_root_ptr;
(if mode = Open then
return (Some shadow_root_ptr)
else
return None
)
} | None \<Rightarrow> return None)
}"
definition a_get_shadow_root_safe_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
where
"a_get_shadow_root_safe_locs element_ptr shadow_root_ptr \<equiv>
(get_shadow_root_locs element_ptr) \<union> (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 \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
fixes get_shadow_root_safe_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> bool"
and get_shadow_root_safe :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_safe_locs :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> (_, (_) shadow_root_ptr option) dom_prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_mode :: "(_) shadow_root_ptr \<Rightarrow> (_, shadow_root_mode) dom_prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<open>set\_disconnected\_nodes\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma set_disconnected_nodes_ok:
"type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<in> set_disconnected_nodes_locs object_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)]
apply(unfold type_wf_impl)
by(auto simp add: all_args_def CD.set_disconnected_nodes_locs_def
intro: put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved split: if_splits)
end
interpretation i_set_disconnected_nodes?: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf
set_disconnected_nodes set_disconnected_nodes_locs
by(auto simp add: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)
declare l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma set_disconnected_nodes_is_l_set_disconnected_nodes [instances]:
"l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs"
apply(auto simp add: l_set_disconnected_nodes_def)[1]
apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_writes)
using set_disconnected_nodes_ok apply blast
apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_ptr_in_heap)
using i_set_disconnected_nodes.set_disconnected_nodes_pointers_preserved apply (blast, blast)
using set_disconnected_nodes_typess_preserved apply(blast, blast)
done
paragraph \<open>get\_child\_nodes\<close>
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 \<Rightarrow> bool"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) 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 \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma set_disconnected_nodes_get_child_nodes:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_host\<close>
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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_tag\_name\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, tag_name) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma get_tag_name_ok:
"type_wf h \<Longrightarrow> element_ptr |\<in>| element_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<open>set\_disconnected\_nodes\<close>
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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_child\_nodes\<close>
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:
"\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>delete\_shadow\_root\<close>
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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>set\_shadow\_root\<close>
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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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: "\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>new\_element\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, tag_name) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_tag_name_new_element:
"ptr' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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' \<noteq> new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr
\<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> r h h'"
assumes new_element_empty_tag_name:
"h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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 \<open>get\_shadow\_root\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>new\_document\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, tag_name) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_tag_name_new_document:
"h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>new\_shadow\_root\<close>
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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>new\_character\_data\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, tag_name) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_tag_name_new_character_data:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_tag_name_locs ptr' \<Longrightarrow> 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 \<open>get\_tag\_type\<close>
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 \<turnstile> CD.a_set_tag_name element_ptr tag \<rightarrow>\<^sub>h h'"
shows "h' \<turnstile> CD.a_get_tag_name element_ptr \<rightarrow>\<^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 \<noteq> ptr'"
assumes "w \<in> CD.a_set_tag_name_locs ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> 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 \<open>attach\_shadow\_root\<close>
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 \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> ((_) heap, exception, unit) prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, char list) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_attach_shadow_root :: "(_) element_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, (_) shadow_root_ptr) dom_prog"
where
"a_attach_shadow_root element_ptr shadow_root_mode = do {
tag \<leftarrow> get_tag_name element_ptr;
(if tag \<notin> safe_shadow_root_element_types then error HierarchyRequestError else return ());
prev_shadow_root \<leftarrow> get_shadow_root element_ptr;
(if prev_shadow_root \<noteq> None then error HierarchyRequestError else return ());
new_shadow_root_ptr \<leftarrow> 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 \<Rightarrow> shadow_root_mode \<Rightarrow> (_, (_) 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 \<Rightarrow> bool"
and set_shadow_root :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> ((_) heap, exception, unit) prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and attach_shadow_root :: "(_) element_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr) prog"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_tag_name :: "(_) element_ptr \<Rightarrow> (_, char list) dom_prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<turnstile> ok (attach_shadow_root element_ptr shadow_root_mode)"
shows "element_ptr |\<in>| element_ptr_kinds h"
proof -
obtain h' where "h \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^sub>h h'"
using assms by auto
then
obtain h2 h3 new_shadow_root_ptr where
h2: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h2" and
new_shadow_root_ptr: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr" and
h3: "h2 \<turnstile> set_mode new_shadow_root_ptr shadow_root_mode \<rightarrow>\<^sub>h h3" and
"h3 \<turnstile> set_shadow_root element_ptr (Some new_shadow_root_ptr) \<rightarrow>\<^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 |\<in>| element_ptr_kinds h3"
using set_shadow_root_ptr_in_heap by blast
moreover
have "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_mode_writes h3])
using set_mode_pointers_preserved
apply blast
by (auto simp add: reflp_def transp_def)
ultimately
show ?thesis
by (metis (no_types, lifting) cast_document_ptr_not_node_ptr(2) element_ptr_kinds_commutes
finsertE funion_finsert_right node_ptr_kinds_commutes sup_bot.right_neutral)
qed
lemma create_shadow_root_known_ptr:
assumes "h \<turnstile> attach_shadow_root element_ptr shadow_root_mode \<rightarrow>\<^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 \<open>get\_parent\<close>
global_interpretation l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs
defines get_parent = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent get_child_nodes"
and get_parent_locs = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent_locs get_child_nodes_locs"
.
interpretation i_get_parent?: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes
get_child_nodes_locs known_ptrs get_parent get_parent_locs
by(simp add: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_parent_def
get_parent_locs_def instances)
declare l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_parent_is_l_get_parent [instances]: "l_get_parent type_wf known_ptr known_ptrs get_parent
get_parent_locs get_child_nodes get_child_nodes_locs"
apply(simp add: l_get_parent_def l_get_parent_axioms_def instances)
using get_parent_reads get_parent_ok get_parent_ptr_in_heap get_parent_pure get_parent_parent_in_heap get_parent_child_dual
using get_parent_reads_pointers
by blast
paragraph \<open>set\_disconnected\_nodes\<close>
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]: "\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_root\_node\<close>
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 \<open>get\_root\_node\_si\<close>
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 \<Rightarrow> ((_) heap, exception, (_::linorder) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
partial_function (dom_prog) a_get_ancestors_si :: "(_::linorder) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_get_ancestors_si ptr = do {
check_in_heap ptr;
ancestors \<leftarrow> (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 \<Rightarrow> do {
parent_ptr_opt \<leftarrow> get_parent node_ptr;
(case parent_ptr_opt of
Some parent_ptr \<Rightarrow> a_get_ancestors_si parent_ptr
| None \<Rightarrow> return [])
}
| None \<Rightarrow> (case cast ptr of
Some shadow_root_ptr \<Rightarrow> do {
host \<leftarrow> get_host shadow_root_ptr;
a_get_ancestors_si (cast host)
} |
None \<Rightarrow> return []));
return (ptr # ancestors)
}"
definition "a_get_ancestors_si_locs = get_parent_locs \<union> get_host_locs"
definition a_get_root_node_si :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr) dom_prog"
where
"a_get_root_node_si ptr = do {
ancestors \<leftarrow> 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 \<Rightarrow> (_, (_) object_ptr list) dom_prog"
fixes get_ancestors_si_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
locale l_get_root_node_si_defs =
fixes get_root_node_si :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr) dom_prog"
fixes get_root_node_si_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 "\<forall>ptr h h' x. h \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r x \<longrightarrow> h \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>h h' \<longrightarrow> 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 \<turnstile> ok (get_ancestors_si ptr)"
shows "ptr |\<in>| 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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
shows "ptr \<in> 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 \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors"
shows "ancestors \<noteq> []"
using assms
apply(simp add: get_ancestors_si_def)
by(auto elim!: bind_returns_result_E2 split: option.splits)
(*
lemma get_ancestors_si_not_node:
assumes "h \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
assumes "\<not>is_node_ptr_kind ptr"
shows "ancestors = [ptr]"
using assms
by (simp add: get_ancestors_si_def) (auto elim!: bind_returns_result_E2 split: option.splits)
*)
lemma get_root_node_si_no_parent:
"h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None \<Longrightarrow> h \<turnstile> get_root_node_si (cast node_ptr) \<rightarrow>\<^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 \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r root"
shows "\<not> 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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^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 \<open>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)\<close> \<open>root = last y\<close> 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 \<open>set\_disconnected\_nodes\<close>
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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_disconnected_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_attribute\<close>
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 \<open>to\_tree\_order\<close>
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 \<open>to\_tree\_order\_si\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
partial_function (dom_prog) a_to_tree_order_si :: "(_) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_to_tree_order_si ptr = (do {
children \<leftarrow> get_child_nodes ptr;
shadow_root_part \<leftarrow> (case cast ptr of
Some element_ptr \<Rightarrow> do {
shadow_root_opt \<leftarrow> get_shadow_root element_ptr;
(case shadow_root_opt of
Some shadow_root_ptr \<Rightarrow> return [cast shadow_root_ptr]
| None \<Rightarrow> return [])
} |
None \<Rightarrow> return []);
treeorders \<leftarrow> 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 \<Rightarrow> (_, (_) 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 "\<forall>ptr h h' x. h \<turnstile> to_tree_order_si ptr \<rightarrow>\<^sub>r x \<longrightarrow> h \<turnstile> to_tree_order_si ptr \<rightarrow>\<^sub>h h' \<longrightarrow> 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 "\<And>x h. pure (f x) h"
by (metis is_OK_returns_heap_E is_OK_returns_result_E pure_def)
then have "\<And>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 \<open>first\_in\_tree\_order\<close>
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 \<open>first\_in\_tree\_order\<close>
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 \<open>get\_element\_by\<close>
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 \<open>get\_element\_by\_si\<close>
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 \<open>find\_slot\<close>
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 \<Rightarrow> ((_) heap, exception, (_::linorder) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_mode :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, shadow_root_mode) prog"
and get_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_attribute :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, char list option) prog"
and get_attribute_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and first_in_tree_order ::
"(_) object_ptr \<Rightarrow> ((_) object_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog) \<Rightarrow>
((_) heap, exception, (_) element_ptr option) prog"
begin
definition a_find_slot :: "bool \<Rightarrow> (_) node_ptr \<Rightarrow> (_, (_) element_ptr option) dom_prog"
where
"a_find_slot open_flag slotable = do {
parent_opt \<leftarrow> get_parent slotable;
(case parent_opt of
Some parent \<Rightarrow>
if is_element_ptr_kind parent
then do {
shadow_root_ptr_opt \<leftarrow> get_shadow_root (the (cast parent));
(case shadow_root_ptr_opt of
Some shadow_root_ptr \<Rightarrow> do {
shadow_root_mode \<leftarrow> get_mode shadow_root_ptr;
if open_flag \<and> shadow_root_mode \<noteq> Open
then return None
else first_in_tree_order (cast shadow_root_ptr) (\<lambda>ptr. if is_element_ptr_kind ptr
then do {
tag \<leftarrow> get_tag_name (the (cast ptr));
name_attr \<leftarrow> get_attribute (the (cast ptr)) ''name'';
slotable_name_attr \<leftarrow> (if is_element_ptr_kind slotable
then get_attribute (the (cast slotable)) ''slot'' else return None);
(if (tag = ''slot'' \<and> (name_attr = slotable_name_attr \<or>
(name_attr = None \<and> slotable_name_attr = Some '''') \<or>
(name_attr = Some '''' \<and> slotable_name_attr = None)))
then return (Some (the (cast ptr)))
else return None)}
else return None)}
| None \<Rightarrow> return None)}
else return None
| _ \<Rightarrow> return None)}"
definition a_assigned_slot :: "(_) node_ptr \<Rightarrow> (_, (_) 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 \<Rightarrow> (_) node_ptr \<Rightarrow> (_, (_) element_ptr option) dom_prog"
and assigned_slot :: "(_) node_ptr \<Rightarrow> (_, (_) 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 \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r slot_opt"
shows "slotable |\<in>| 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 \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r Some slot"
shows "slot |\<in>| 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 \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r slot_opt \<Longrightarrow> slotable |\<in>| node_ptr_kinds h"
assumes find_slot_slot_in_heap: "h \<turnstile> find_slot open_flag slotable \<rightarrow>\<^sub>r Some slot \<Longrightarrow> slot |\<in>| 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 \<open>get\_disconnected\_nodes\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma get_disconnected_nodes_ok:
"type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<open>set\_child\_nodes\<close>
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:
"\<forall>w \<in> set_child_nodes_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> get_disconnected_nodes_locs ptr'. r h h'))"
by(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def
CD.get_disconnected_nodes_locs_def all_args_def elim: get_M_document_put_M_shadow_root
split: option.splits)
end
interpretation
i_set_child_nodes_get_disconnected_nodes?: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf
known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs
Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs
apply(auto simp add: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1]
by(unfold_locales)
declare l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma set_child_nodes_get_disconnected_nodes_is_l_set_child_nodes_get_disconnected_nodes [instances]:
"l_set_child_nodes_get_disconnected_nodes type_wf set_child_nodes set_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs"
using set_child_nodes_is_l_set_child_nodes get_disconnected_nodes_is_l_get_disconnected_nodes
apply(simp add: l_set_child_nodes_get_disconnected_nodes_def l_set_child_nodes_get_disconnected_nodes_axioms_def )
using set_child_nodes_get_disconnected_nodes
by fast
paragraph \<open>set\_disconnected\_nodes\<close>
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 \<open>delete\_shadow\_root\<close>
locale l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_disconnected_nodes_delete_shadow_root:
"cast shadow_root_ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> r h h'"
by(auto simp add: CD.get_disconnected_nodes_locs_def delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
end
locale l_delete_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs +
assumes get_disconnected_nodes_delete_shadow_root:
"cast shadow_root_ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> 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 \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> 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 \<open>set\_shadow\_root\<close>
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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_shadow_root_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>set\_mode\<close>
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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_mode_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>new\_shadow\_root\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma get_disconnected_nodes_new_shadow_root_different_pointers:
"cast new_shadow_root_ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> r h h'"
by(auto simp add: CD.get_disconnected_nodes_locs_def new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)
lemma new_shadow_root_no_disconnected_nodes:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
by(simp add: CD.get_disconnected_nodes_def new_shadow_root_disconnected_nodes)
end
interpretation i_new_shadow_root_get_disconnected_nodes?:
l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes
get_disconnected_nodes_locs
by unfold_locales
declare l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
locale l_new_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs +
assumes get_disconnected_nodes_new_shadow_root_different_pointers:
"cast new_shadow_root_ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> r \<in> get_disconnected_nodes_locs ptr' \<Longrightarrow> r h h'"
assumes new_shadow_root_no_disconnected_nodes:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
lemma new_shadow_root_get_disconnected_nodes_is_l_new_shadow_root_get_disconnected_nodes [instances]:
"l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs"
apply (auto simp add: l_new_shadow_root_get_disconnected_nodes_def)[1]
using get_disconnected_nodes_new_shadow_root_different_pointers apply fast
using new_shadow_root_no_disconnected_nodes apply blast
done
subsubsection \<open>remove\_shadow\_root\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_shadow_root :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_remove_shadow_root :: "(_) element_ptr \<Rightarrow> (_, unit) dom_prog" where
"a_remove_shadow_root element_ptr = do {
shadow_root_ptr_opt \<leftarrow> get_shadow_root element_ptr;
(case shadow_root_ptr_opt of
Some shadow_root_ptr \<Rightarrow> do {
children \<leftarrow> get_child_nodes (cast shadow_root_ptr);
disconnected_nodes \<leftarrow> get_disconnected_nodes (cast shadow_root_ptr);
(if children = [] \<and> disconnected_nodes = []
then do {
set_shadow_root element_ptr None;
delete_M shadow_root_ptr
} else do {
error HierarchyRequestError
})
} |
None \<Rightarrow> error HierarchyRequestError)
}"
definition a_remove_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_, unit) dom_prog) set"
where
"a_remove_shadow_root_locs element_ptr shadow_root_ptr \<equiv> set_shadow_root_locs element_ptr \<union> {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 \<Rightarrow> (_, unit) dom_prog"
fixes remove_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_, unit) dom_prog) set"
locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +
l_remove_shadow_root_defs +
l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_child_nodes +
l_get_disconnected_nodes +
assumes remove_shadow_root_impl: "remove_shadow_root = a_remove_shadow_root"
assumes remove_shadow_root_locs_impl: "remove_shadow_root_locs = a_remove_shadow_root_locs"
begin
lemmas remove_shadow_root_def =
remove_shadow_root_impl[unfolded remove_shadow_root_def a_remove_shadow_root_def]
lemmas remove_shadow_root_locs_def =
remove_shadow_root_locs_impl[unfolded remove_shadow_root_locs_def a_remove_shadow_root_locs_def]
lemma remove_shadow_root_writes:
"writes (remove_shadow_root_locs element_ptr (the |h \<turnstile> 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 \<open>get\_child\_nodes\<close>
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 \<noteq> cast shadow_root_ptr"
assumes "w \<in> remove_shadow_root_locs element_ptr shadow_root_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> get_child_nodes_locs ptr"
shows "r h h'"
using assms
by(auto simp add: all_args_def get_child_nodes_locs_def CD.get_child_nodes_locs_def
remove_shadow_root_locs_def set_shadow_root_locs_def
delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
intro: is_shadow_root_ptr_kind_obtains
simp add: delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t[rotated] element_put_get_preserved[where setter=shadow_root_opt_update]
elim: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains
split: if_splits option.splits)[1]
end
locale l_remove_shadow_root_get_child_nodes = l_get_child_nodes_defs + l_remove_shadow_root_defs +
assumes remove_shadow_root_get_child_nodes_different_pointers:
"ptr \<noteq> cast shadow_root_ptr \<Longrightarrow> w \<in> remove_shadow_root_locs element_ptr shadow_root_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
r \<in> get_child_nodes_locs ptr \<Longrightarrow> 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 \<open>get\_tag\_name\<close>
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 \<in> remove_shadow_root_locs element_ptr shadow_root_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^sub>h h'"
assumes "r \<in> 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 \<in> remove_shadow_root_locs element_ptr shadow_root_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> r \<in> get_tag_name_locs ptr \<Longrightarrow>
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 \<open>get\_owner\_document\<close>
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 \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<Rightarrow> unit \<Rightarrow> (_, (_) document_ptr) dom_prog"
where
"a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast shadow_root_ptr)"
definition a_get_owner_document_tups :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> unit
\<Rightarrow> (_, (_) 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 \<circ> the \<circ> cast)]"
definition a_get_owner_document :: "(_::linorder) object_ptr \<Rightarrow> (_, (_) document_ptr) dom_prog"
where
"a_get_owner_document ptr = invoke (CD.a_get_owner_document_tups @ a_get_owner_document_tups) ptr ()"
end
global_interpretation l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs
get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs
defines get_owner_document_tups = "l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document_tups"
and get_owner_document =
"l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document get_root_node get_disconnected_nodes"
and get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r =
"l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r"
and get_owner_document_tups\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
"l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document_tups get_root_node get_disconnected_nodes"
and get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r =
"l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r get_root_node get_disconnected_nodes"
.
locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs get_disconnected_nodes
get_disconnected_nodes_locs get_host get_host_locs +
l_get_owner_document_defs get_owner_document +
l_get_host get_host get_host_locs +
CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_parent get_parent_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes
get_disconnected_nodes_locs get_root_node get_root_node_locs get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
for known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \<Rightarrow> bool"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_owner_document :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" +
assumes get_owner_document_impl: "get_owner_document = a_get_owner_document"
begin
lemmas get_owner_document_def = a_get_owner_document_def[folded get_owner_document_impl]
lemma get_owner_document_pure [simp]:
"pure (get_owner_document ptr) h"
proof -
have "\<And>shadow_root_ptr. pure (a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr ()) h"
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I
split: option.splits)[1]
by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I
split: option.splits)
then show ?thesis
apply(auto simp add: get_owner_document_def)[1]
apply(split CD.get_owner_document_splits, rule conjI)+
apply(simp)
apply(auto simp add: a_get_owner_document_tups_def)[1]
apply(split invoke_splits, rule conjI)+
apply simp
by(auto intro!: bind_pure_I)
qed
lemma get_owner_document_ptr_in_heap:
assumes "h \<turnstile> ok (get_owner_document ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
by(auto simp add: get_owner_document_def invoke_ptr_in_heap dest: is_OK_returns_heap_I)
end
interpretation
i_get_owner_document?: l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M DocumentClass.known_ptr get_parent get_parent_locs
DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node get_root_node_locs CD.a_get_owner_document get_host get_host_locs get_owner_document get_child_nodes get_child_nodes_locs
using get_child_nodes_is_l_get_child_nodes[unfolded ShadowRootClass.known_ptr_defs]
by(auto simp add: instances l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def
l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_owner_document_def Core_DOM_Functions.get_owner_document_def)
declare l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_owner_document_is_l_get_owner_document [instances]: "l_get_owner_document get_owner_document"
apply(auto simp add: l_get_owner_document_def)[1]
using get_owner_document_ptr_in_heap apply fast
done
subsubsection \<open>remove\_child\<close>
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 \<open>get\_disconnected\_document\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_get_disconnected_document :: "(_) node_ptr \<Rightarrow> (_, (_) document_ptr) dom_prog"
where
"a_get_disconnected_document node_ptr = do {
check_in_heap (cast node_ptr);
ptrs \<leftarrow> document_ptr_kinds_M;
candidates \<leftarrow> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> get_disconnected_nodes document_ptr;
return (node_ptr \<in> set disconnected_nodes)
}) ptrs;
(case candidates of
Cons document_ptr [] \<Rightarrow> return document_ptr |
_ \<Rightarrow> error HierarchyRequestError
)
}"
definition "a_get_disconnected_document_locs =
(\<Union>document_ptr. get_disconnected_nodes_locs document_ptr) \<union> (\<Union>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 \<Rightarrow> (_, (_::linorder) document_ptr) dom_prog"
fixes get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 \<turnstile> ok (get_disconnected_document node_ptr) \<Longrightarrow> node_ptr |\<in>| 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 \<turnstile> get_disconnected_document child_node \<rightarrow>\<^sub>r disconnected_document"
shows "disconnected_document |\<in>| 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 \<turnstile> ok (get_disconnected_document node_ptr) \<Longrightarrow> node_ptr |\<in>| 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 \<turnstile> get_disconnected_document child_node \<rightarrow>\<^sub>r disconnected_document \<Longrightarrow>
disconnected_document |\<in>| document_ptr_kinds h"
global_interpretation l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes
get_disconnected_nodes_locs defines
get_disconnected_document = "l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_disconnected_document
get_disconnected_nodes" and
get_disconnected_document_locs = "l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_disconnected_document_locs
get_disconnected_nodes_locs" .
interpretation i_get_disconnected_document?: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf
by(auto simp add: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def
get_disconnected_document_def get_disconnected_document_locs_def instances)
declare l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_disconnected_document_is_l_get_disconnected_document [instances]:
"l_get_disconnected_document get_disconnected_document get_disconnected_document_locs"
apply(auto simp add: l_get_disconnected_document_def instances)[1]
using get_disconnected_document_ptr_in_heap get_disconnected_document_pure
get_disconnected_document_disconnected_document_in_heap get_disconnected_document_reads
by blast+
paragraph \<open>get\_disconnected\_nodes\<close>
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:
"\<forall>w \<in> set_tag_name_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_ancestors\_di\<close>
locale l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs =
l_get_parent_defs get_parent get_parent_locs +
l_get_host_defs get_host get_host_locs +
l_get_disconnected_document_defs get_disconnected_document get_disconnected_document_locs
for get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_::linorder) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
partial_function (dom_prog) a_get_ancestors_di :: "(_::linorder) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
where
"a_get_ancestors_di ptr = do {
check_in_heap ptr;
ancestors \<leftarrow> (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 \<Rightarrow> do {
parent_ptr_opt \<leftarrow> get_parent node_ptr;
(case parent_ptr_opt of
Some parent_ptr \<Rightarrow> a_get_ancestors_di parent_ptr
| None \<Rightarrow> do {
document_ptr \<leftarrow> get_disconnected_document node_ptr;
a_get_ancestors_di (cast document_ptr)
})
}
| None \<Rightarrow> (case cast ptr of
Some shadow_root_ptr \<Rightarrow> do {
host \<leftarrow> get_host shadow_root_ptr;
a_get_ancestors_di (cast host)
} |
None \<Rightarrow> return []));
return (ptr # ancestors)
}"
definition "a_get_ancestors_di_locs = get_parent_locs \<union> get_host_locs \<union> get_disconnected_document_locs"
end
locale l_get_ancestors_di_defs =
fixes get_ancestors_di :: "(_::linorder) object_ptr \<Rightarrow> (_, (_) object_ptr list) dom_prog"
fixes get_ancestors_di_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
locale l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_parent +
l_get_host +
l_get_disconnected_document +
l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +
l_get_ancestors_di_defs +
assumes get_ancestors_di_impl: "get_ancestors_di = a_get_ancestors_di"
assumes get_ancestors_di_locs_impl: "get_ancestors_di_locs = a_get_ancestors_di_locs"
begin
lemmas get_ancestors_di_def = a_get_ancestors_di.simps[folded get_ancestors_di_impl]
lemmas get_ancestors_di_locs_def = a_get_ancestors_di_locs_def[folded get_ancestors_di_locs_impl]
lemma get_ancestors_di_pure [simp]:
"pure (get_ancestors_di ptr) h"
proof -
have "\<forall>ptr h h' x. h \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>r x \<longrightarrow> h \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>h h' \<longrightarrow> h = h'"
proof (induct rule: a_get_ancestors_di.fixp_induct[folded get_ancestors_di_impl])
case 1
then show ?case
by(rule admissible_dom_prog)
next
case 2
then show ?case
by simp
next
case (3 f)
then show ?case
using get_parent_pure get_host_pure get_disconnected_document_pure
apply(auto simp add: pure_returns_heap_eq pure_def split: option.splits elim!: bind_returns_heap_E
bind_returns_result_E dest!: pure_returns_heap_eq[rotated, OF check_in_heap_pure])[1]
apply (metis is_OK_returns_result_I returns_heap_eq returns_result_eq)
apply (meson option.simps(3) returns_result_eq)
apply (meson option.simps(3) returns_result_eq)
apply(metis get_parent_pure pure_returns_heap_eq)
apply(metis get_host_pure pure_returns_heap_eq)
done
qed
then show ?thesis
by (meson pure_eq_iff)
qed
lemma get_ancestors_di_ptr:
assumes "h \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>r ancestors"
shows "ptr \<in> set ancestors"
using assms
by(simp add: get_ancestors_di_def) (auto elim!: bind_returns_result_E2 split: option.splits
intro!: bind_pure_I)
lemma get_ancestors_di_ptr_in_heap:
assumes "h \<turnstile> ok (get_ancestors_di ptr)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
by(auto simp add: get_ancestors_di_def check_in_heap_ptr_in_heap elim!: bind_is_OK_E
dest: is_OK_returns_result_I)
lemma get_ancestors_di_never_empty:
assumes "h \<turnstile> get_ancestors_di child \<rightarrow>\<^sub>r ancestors"
shows "ancestors \<noteq> []"
using assms
apply(simp add: get_ancestors_di_def)
by(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)
end
global_interpretation l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_host get_host_locs
get_disconnected_document get_disconnected_document_locs
defines get_ancestors_di = a_get_ancestors_di
and get_ancestors_di_locs = a_get_ancestors_di_locs .
declare a_get_ancestors_di.simps [code]
interpretation i_get_ancestors_di?: l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_ancestors_di get_ancestors_di_locs
by(auto simp add: l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def
get_ancestors_di_def get_ancestors_di_locs_def instances)
declare l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_ancestors_di_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors_di"
apply(auto simp add: l_get_ancestors_def)[1]
using get_ancestors_di_ptr_in_heap apply fast
using get_ancestors_di_ptr apply fast
done
subsubsection \<open>adopt\_node\<close>
locale l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w_\<^sub>D\<^sub>O\<^sub>M_defs =
CD: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs remove_child
remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs +
l_get_ancestors_di_defs get_ancestors_di get_ancestors_di_locs
for get_owner_document :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and remove_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_child_locs :: "(_) object_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_ancestors_di :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_di_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_adopt_node :: "(_) document_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> (_, unit) dom_prog"
where
"a_adopt_node document node = do {
ancestors \<leftarrow> get_ancestors_di (cast document);
(if cast node \<in> set ancestors
then error HierarchyRequestError
else CD.a_adopt_node document node)}"
definition a_adopt_node_locs ::
"(_) object_ptr option \<Rightarrow> (_) document_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> (_, unit) dom_prog set"
where "a_adopt_node_locs = CD.a_adopt_node_locs"
end
locale l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs remove_child
remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs +
l_adopt_node_defs adopt_node adopt_node_locs +
l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs
get_child_nodes get_child_nodes_locs get_host get_host_locs get_disconnected_document
get_disconnected_document_locs get_ancestors_di get_ancestors_di_locs +
CD: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child
remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes
set_child_nodes_locs remove
for get_owner_document :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and remove_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_child_locs :: "(_) object_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_ancestors_di :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and get_ancestors_di_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and adopt_node :: "(_) document_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and adopt_node_locs ::
"(_) object_ptr option \<Rightarrow> (_) document_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) document_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ::
"(_) object_ptr option \<Rightarrow> (_) document_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and set_child_nodes :: "(_) object_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and remove :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog" +
assumes adopt_node_impl: "adopt_node = a_adopt_node"
assumes adopt_node_locs_impl: "adopt_node_locs = a_adopt_node_locs"
begin
lemmas adopt_node_def = a_adopt_node_def[folded adopt_node_impl CD.adopt_node_impl]
lemmas adopt_node_locs_def = a_adopt_node_locs_def[folded adopt_node_locs_impl CD.adopt_node_locs_impl]
lemma adopt_node_writes:
"writes (adopt_node_locs |h \<turnstile> get_parent node|\<^sub>r
|h \<turnstile> get_owner_document (cast node)|\<^sub>r document_ptr) (adopt_node document_ptr node) h h'"
by(auto simp add: CD.adopt_node_writes adopt_node_def CD.adopt_node_impl[symmetric]
adopt_node_locs_def CD.adopt_node_locs_impl[symmetric]
intro!: writes_bind_pure[OF get_ancestors_di_pure])
lemma adopt_node_pointers_preserved:
"w \<in> adopt_node_locs parent owner_document document_ptr
\<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h'"
using CD.adopt_node_locs_impl CD.adopt_node_pointers_preserved local.adopt_node_locs_def by blast
lemma adopt_node_types_preserved:
"w \<in> adopt_node_locs parent owner_document document_ptr
\<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using CD.adopt_node_locs_impl CD.adopt_node_types_preserved local.adopt_node_locs_def by blast
lemma adopt_node_child_in_heap:
"h \<turnstile> ok (adopt_node document_ptr child) \<Longrightarrow> child |\<in>| node_ptr_kinds h"
by (smt CD.adopt_node_child_in_heap CD.adopt_node_impl bind_is_OK_E error_returns_heap
is_OK_returns_heap_E l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.adopt_node_def l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms
local.get_ancestors_di_pure pure_returns_heap_eq)
lemma adopt_node_children_subset:
"h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'
\<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> set children' \<subseteq> set children"
by (smt CD.adopt_node_children_subset CD.adopt_node_impl bind_returns_heap_E error_returns_heap
local.adopt_node_def local.get_ancestors_di_pure pure_returns_heap_eq)
end
global_interpretation l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs
remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs
defines adopt_node = "a_adopt_node"
and adopt_node_locs = "a_adopt_node_locs"
and adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = "CD.a_adopt_node"
and adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = "CD.a_adopt_node_locs"
.
interpretation i_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document get_parent get_parent_locs remove_child remove_child_locs
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs
adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs
known_ptrs set_child_nodes set_child_nodes_locs remove
by(auto simp add: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def
adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare i_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
interpretation i_adopt_node?: l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs get_ancestors_di
get_ancestors_di_locs adopt_node adopt_node_locs CD.a_adopt_node CD.a_adopt_node_locs known_ptr
type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs
get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove
by(auto simp add: l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def adopt_node_def
adopt_node_locs_def instances)
declare l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma adopt_node_is_l_adopt_node [instances]: "l_adopt_node type_wf known_ptr known_ptrs get_parent
adopt_node adopt_node_locs get_child_nodes get_owner_document"
apply(auto simp add: l_adopt_node_def l_adopt_node_axioms_def instances)[1]
using adopt_node_writes apply fast
using adopt_node_pointers_preserved apply (fast, fast)
using adopt_node_types_preserved apply (fast, fast)
using adopt_node_child_in_heap apply fast
using adopt_node_children_subset apply fast
done
paragraph \<open>get\_shadow\_root\<close>
locale l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma adopt_node_get_shadow_root:
"\<forall>w \<in> adopt_node_locs parent owner_document document_ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow>
(\<forall>r \<in> get_shadow_root_locs ptr'. r h h'))"
by(auto simp add: adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def
all_args_def set_disconnected_nodes_get_shadow_root set_child_nodes_get_shadow_root)
end
locale l_adopt_node_get_shadow_root = l_adopt_node_defs + l_get_shadow_root_defs +
assumes adopt_node_get_shadow_root:
"\<forall>w \<in> adopt_node_locs parent owner_document document_ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow>
(\<forall>r \<in> get_shadow_root_locs ptr'. r h h'))"
interpretation i_adopt_node_get_shadow_root?: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs
Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root
get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_owner_document
get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes
get_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs
adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs known_ptrs get_host
get_host_locs get_disconnected_document get_disconnected_document_locs remove
by(auto simp add: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
interpretation i_adopt_node_get_shadow_root?: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes
set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs
get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs
get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes
get_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs
adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs known_ptrs get_host
get_host_locs get_disconnected_document get_disconnected_document_locs remove
by(auto simp add: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma adopt_node_get_shadow_root_is_l_adopt_node_get_shadow_root [instances]:
"l_adopt_node_get_shadow_root adopt_node_locs get_shadow_root_locs"
apply(auto simp add: l_adopt_node_get_shadow_root_def)[1]
using adopt_node_get_shadow_root apply fast
done
subsubsection \<open>insert\_before\<close>
global_interpretation l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_child_nodes
get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_di get_ancestors_di_locs
adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_owner_document
defines
next_sibling = a_next_sibling and
insert_node = a_insert_node and
ensure_pre_insertion_validity = a_ensure_pre_insertion_validity and
insert_before = a_insert_before and
insert_before_locs = a_insert_before_locs
.
global_interpretation l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs insert_before
defines append_child = "l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_append_child insert_before"
.
interpretation i_insert_before?: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes
get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_di get_ancestors_di_locs
adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf
known_ptr known_ptrs
by(auto simp add: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def insert_before_def
insert_before_locs_def instances)
declare l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
interpretation i_append_child?: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M append_child insert_before insert_before_locs
by(simp add: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances append_child_def)
declare l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
subsubsection \<open>get\_assigned\_nodes\<close>
fun map_filter_M2 :: "('x \<Rightarrow> ('heap, 'e, 'y option) prog) \<Rightarrow> 'x list
\<Rightarrow> ('heap, 'e, 'y list) prog"
where
"map_filter_M2 f [] = return []" |
"map_filter_M2 f (x # xs) = do {
res \<leftarrow> f x;
remainder \<leftarrow> map_filter_M2 f xs;
return ((case res of Some r \<Rightarrow> [r] | None \<Rightarrow> []) @ remainder)
}"
lemma map_filter_M2_pure [simp]:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> 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 "\<And>x. x \<in> set xs \<Longrightarrow> pure (f x) h"
assumes "h \<turnstile> map_filter_M2 f xs \<rightarrow>\<^sub>r ys"
shows
"ys = map the (filter (\<lambda>x. x \<noteq> None) (map (\<lambda>x. |h \<turnstile> f x|\<^sub>r) xs))" and
"\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> ok (f x)"
using assms
apply(induct xs arbitrary: h ys)
by(auto elim!: bind_returns_result_E2)
lemma map_filter_pure_foo:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> pure (f x) h"
assumes "h \<turnstile> map_filter_M2 f xs \<rightarrow>\<^sub>r ys"
assumes "y \<in> set ys"
obtains x where "h \<turnstile> f x \<rightarrow>\<^sub>r Some y" and "x \<in> set xs"
using assms
apply(induct xs arbitrary: ys)
by(auto elim!: bind_returns_result_E2)
lemma map_filter_M2_in_result:
assumes "h \<turnstile> map_filter_M2 P xs \<rightarrow>\<^sub>r ys"
assumes "a \<in> set xs"
assumes "\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h"
assumes "h \<turnstile> P a \<rightarrow>\<^sub>r Some b"
shows "b \<in> 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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and find_slot :: "bool \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and assigned_slot :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and remove :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before ::
"(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> (_) node_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before_locs ::
"(_) object_ptr \<Rightarrow> (_) object_ptr option \<Rightarrow> (_) document_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> (_, unit) dom_prog set"
and append_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root_locs ::
"(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
begin
definition a_assigned_nodes :: "(_) element_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
where
"a_assigned_nodes slot = do {
tag \<leftarrow> get_tag_name slot;
(if tag \<noteq> ''slot''
then error HierarchyRequestError
else return ());
root \<leftarrow> get_root_node (cast slot);
if is_shadow_root_ptr_kind root
then do {
host \<leftarrow> get_host (the (cast root));
children \<leftarrow> get_child_nodes (cast host);
filter_M (\<lambda>slotable. do {
found_slot \<leftarrow> find_slot False slotable;
return (found_slot = Some slot)}) children}
else return []}"
partial_function (dom_prog) a_assigned_nodes_flatten ::
"(_) element_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
where
"a_assigned_nodes_flatten slot = do {
tag \<leftarrow> get_tag_name slot;
(if tag \<noteq> ''slot''
then error HierarchyRequestError
else return ());
root \<leftarrow> get_root_node (cast slot);
(if is_shadow_root_ptr_kind root
then do {
slotables \<leftarrow> a_assigned_nodes slot;
slotables_or_child_nodes \<leftarrow> (if slotables = []
then do {
get_child_nodes (cast slot)
} else do {
return slotables
});
list_of_lists \<leftarrow> map_M (\<lambda>node_ptr. do {
(case cast node_ptr of
Some element_ptr \<Rightarrow> do {
tag \<leftarrow> get_tag_name element_ptr;
(if tag = ''slot''
then do {
root \<leftarrow> 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 \<Rightarrow> 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 \<leftarrow> element_ptr_kinds_M \<bind> map_filter_M2 (\<lambda>element_ptr. do {
tag \<leftarrow> get_tag_name element_ptr;
assigned_nodes \<leftarrow> a_assigned_nodes element_ptr;
(if tag = ''slot'' \<and> assigned_nodes \<noteq> []
then return (Some (element_ptr, assigned_nodes)) else return None)});
forall_M (\<lambda>(slot, assigned_nodes). do {
get_child_nodes (cast slot) \<bind> forall_M remove;
forall_M (append_child (cast slot)) assigned_nodes
}) tups;
shadow_root_ptr_kinds_M \<bind> forall_M (\<lambda>shadow_root_ptr. do {
host \<leftarrow> get_host shadow_root_ptr;
get_child_nodes (cast host) \<bind> forall_M remove;
get_child_nodes (cast shadow_root_ptr) \<bind> forall_M (append_child (cast host));
remove_shadow_root host
});
return ()
}"
end
global_interpretation l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs
get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs
find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root
remove_shadow_root_locs
defines assigned_nodes =
"l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_assigned_nodes get_child_nodes get_tag_name get_root_node get_host
find_slot"
and assigned_nodes_flatten =
"l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_assigned_nodes_flatten get_child_nodes get_tag_name get_root_node
get_host find_slot"
and flatten_dom =
"l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_flatten_dom get_child_nodes get_tag_name get_root_node get_host
find_slot remove append_child remove_shadow_root"
.
declare a_assigned_nodes_flatten.simps [code]
locale l_assigned_nodes_defs =
fixes assigned_nodes :: "(_) element_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
fixes assigned_nodes_flatten :: "(_) element_ptr \<Rightarrow> (_, (_) node_ptr list) dom_prog"
fixes flatten_dom :: "(_, unit) dom_prog"
locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_assigned_nodes_defs
assigned_nodes assigned_nodes_flatten flatten_dom
+ l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs
get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node
get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before
insert_before_locs append_child remove_shadow_root remove_shadow_root_locs
(* + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *)
+ l_get_shadow_root
type_wf get_shadow_root get_shadow_root_locs
+ l_set_shadow_root
type_wf set_shadow_root set_shadow_root_locs
+ l_remove
+ l_insert_before
insert_before insert_before_locs
+ l_find_slot
find_slot assigned_slot
+ l_get_tag_name
type_wf get_tag_name get_tag_name_locs
+ l_get_root_node
get_root_node get_root_node_locs get_parent get_parent_locs
+ l_get_host
get_host get_host_locs
+ l_get_child_nodes
type_wf known_ptr get_child_nodes get_child_nodes_locs
+ l_to_tree_order
to_tree_order
for known_ptr :: "(_::linorder) object_ptr \<Rightarrow> bool"
and assigned_nodes :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and assigned_nodes_flatten :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and flatten_dom :: "((_) heap, exception, unit) prog"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_root_node :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr) prog"
and get_root_node_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and find_slot :: "bool \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and assigned_slot :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and remove :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> (_) node_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before_locs ::
"(_) object_ptr \<Rightarrow> (_) object_ptr option \<Rightarrow> (_) document_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> (_, unit) dom_prog set"
and append_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
and remove_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_shadow_root :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and to_tree_order :: "(_) object_ptr \<Rightarrow> ((_) 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 \<turnstile> ok (assigned_nodes slot)"
shows "slot |\<in>| 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 \<turnstile> ok (assigned_nodes slot)"
shows "h \<turnstile> get_tag_name slot \<rightarrow>\<^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 \<turnstile> assigned_nodes slot \<rightarrow>\<^sub>r nodes"
assumes "h \<turnstile> assigned_nodes slot' \<rightarrow>\<^sub>r nodes'"
assumes "slot \<noteq> slot'"
shows "set nodes \<inter> set nodes' = {}"
proof (rule ccontr)
assume "set nodes \<inter> set nodes' \<noteq> {} "
then obtain common_ptr where "common_ptr \<in> set nodes" and "common_ptr \<in> set nodes'"
by auto
have "h \<turnstile> find_slot False common_ptr \<rightarrow>\<^sub>r Some slot"
using \<open>common_ptr \<in> set nodes\<close>
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 \<turnstile> find_slot False common_ptr \<rightarrow>\<^sub>r Some slot'"
using \<open>common_ptr \<in> set nodes'\<close>
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 \<turnstile> ok (assigned_nodes slot) \<Longrightarrow> slot |\<in>| element_ptr_kinds h"
assumes assigned_nodes_slot_is_slot: "h \<turnstile> ok (assigned_nodes slot) \<Longrightarrow> h \<turnstile> get_tag_name slot \<rightarrow>\<^sub>r ''slot''"
assumes assigned_nodes_different_ptr:
"h \<turnstile> assigned_nodes slot \<rightarrow>\<^sub>r nodes \<Longrightarrow> h \<turnstile> assigned_nodes slot' \<rightarrow>\<^sub>r nodes' \<Longrightarrow> slot \<noteq> slot' \<Longrightarrow>
set nodes \<inter> 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 \<open>set\_val\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> (_, unit) dom_prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> (_, unit) dom_prog set" +
assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf"
begin
lemma set_val_ok:
"type_wf h \<Longrightarrow> character_data_ptr |\<in>| character_data_ptr_kinds h \<Longrightarrow>
h \<turnstile> 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 \<in> set_val_locs character_data_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<in> set_val_locs character_data_ptr"
assumes "h \<turnstile> w \<rightarrow>\<^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 \<open>get\_shadow\_root\<close>
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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>get\_tag\_type\<close>
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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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:
"\<forall>w \<in> set_val_locs ptr. (h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>r \<in> 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 \<open>create\_character\_data\<close>
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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> bool" +
assumes known_ptr_impl: "known_ptr = a_known_ptr"
begin
lemma create_character_data_document_in_heap:
assumes "h \<turnstile> ok (create_character_data document_ptr text)"
shows "document_ptr |\<in>| document_ptr_kinds h"
using assms CD.create_character_data_document_in_heap by simp
lemma create_character_data_known_ptr:
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<open>create\_element\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \<Rightarrow> 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 \<turnstile> ok (create_element document_ptr tag)"
shows "document_ptr |\<in>| document_ptr_kinds h"
using CD.create_element_document_in_heap assms .
lemma create_element_known_ptr:
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<open>A wellformed heap (Core DOM)\<close>
subsubsection \<open>wellformed\_heap\<close>
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 \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
definition a_host_shadow_root_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
where
"a_host_shadow_root_rel h = (\<lambda>(x, y). (cast x, cast y)) ` {(host, shadow_root).
host |\<in>| element_ptr_kinds h \<and> |h \<turnstile> 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
(\<lambda>host. (case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root \<Rightarrow> [(cast host, cast shadow_root)] |
None \<Rightarrow> []))
(sorted_list_of_fset (element_ptr_kinds h)))
)"
by(auto simp add: a_host_shadow_root_rel_def)
definition a_ptr_disconnected_node_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
where
"a_ptr_disconnected_node_rel h = (\<lambda>(x, y). (cast x, cast y)) ` {(document_ptr, disconnected_node).
document_ptr |\<in>| document_ptr_kinds h \<and> disconnected_node \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r}"
lemma a_ptr_disconnected_node_rel_code [code]: "a_ptr_disconnected_node_rel h = set (concat (map
(\<lambda>ptr. map
(\<lambda>node. (cast ptr, cast node))
|h \<turnstile> get_disconnected_nodes ptr|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h)))
)"
by(auto simp add: a_ptr_disconnected_node_rel_def)
definition a_all_ptrs_in_heap :: "(_) heap \<Rightarrow> bool" where
"a_all_ptrs_in_heap h = ((\<forall>host shadow_root_ptr.
(h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr) \<longrightarrow>
shadow_root_ptr |\<in>| shadow_root_ptr_kinds h))"
definition a_distinct_lists :: "(_) heap \<Rightarrow> bool"
where
"a_distinct_lists h = distinct (concat (
map (\<lambda>element_ptr. (case |h \<turnstile> get_shadow_root element_ptr|\<^sub>r of
Some shadow_root_ptr \<Rightarrow> [shadow_root_ptr] | None \<Rightarrow> []))
|h \<turnstile> element_ptr_kinds_M|\<^sub>r
))"
definition a_shadow_root_valid :: "(_) heap \<Rightarrow> bool" where
"a_shadow_root_valid h = (\<forall>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h).
(\<exists>host \<in> fset(element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr))"
definition a_heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
where
"a_heap_is_wellformed h \<longleftrightarrow> CD.a_heap_is_wellformed h \<and>
acyclic (CD.a_parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h) \<and>
a_all_ptrs_in_heap h \<and>
a_distinct_lists h \<and>
a_shadow_root_valid h"
end
global_interpretation l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs
get_tag_name get_tag_name_locs
defines heap_is_wellformed = a_heap_is_wellformed
and parent_child_rel = CD.a_parent_child_rel
and host_shadow_root_rel = a_host_shadow_root_rel
and ptr_disconnected_node_rel = a_ptr_disconnected_node_rel
and all_ptrs_in_heap = a_all_ptrs_in_heap
and distinct_lists = a_distinct_lists
and shadow_root_valid = a_shadow_root_valid
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_heap_is_wellformed
and parent_child_rel\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_parent_child_rel
and acyclic_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_acyclic_heap
and all_ptrs_in_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_all_ptrs_in_heap
and distinct_lists\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_distinct_lists
and owner_document_valid\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_owner_document_valid
.
interpretation i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel
by (auto simp add: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def
heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def instances)
declare i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
lemma heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_heap_is_wellformed [instances]:
"l_heap_is_wellformed type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes
get_disconnected_nodes"
apply(auto simp add: l_heap_is_wellformed_def)[1]
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_in_heap apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disc_nodes_in_heap apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_parent apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_disc_parent apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes_different apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disconnected_nodes_distinct apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_distinct apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply (blast, blast)
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_acyclic apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast
using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child_in_heap apply blast
done
locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs
+ CD: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel
+ l_heap_is_wellformed_defs
heap_is_wellformed parent_child_rel
+ l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf
+ l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs
get_disconnected_document get_disconnected_document_locs type_wf
+ l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs
for get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> 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 = ((\<forall>host \<in> fset (element_ptr_kinds h).
h \<turnstile> ok (get_shadow_root host) \<longrightarrow> (case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root_ptr \<Rightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h |
None \<Rightarrow> 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)
+ by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap select_result_I2)
lemma get_shadow_root_shadow_root_ptr_in_heap:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr"
shows "shadow_root_ptr |\<in>| 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 \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r host"
shows "shadow_root_ptr |\<in>| 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 \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root_ptr"
assumes "h \<turnstile> get_shadow_root host' \<rightarrow>\<^sub>r Some shadow_root_ptr"
shows "host = host'"
proof (rule ccontr)
assume " host \<noteq> host'"
have "host |\<in>| element_ptr_kinds h"
using assms(3)
by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap)
moreover
have "host' |\<in>| 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 \<open>host \<noteq> host'\<close> apply(simp)
apply(auto)[1]
done
qed
lemma shadow_root_host_dual:
assumes "h \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r host"
shows "h \<turnstile> get_shadow_root host \<rightarrow>\<^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 \<turnstile> get_disconnected_document disc_node \<rightarrow>\<^sub>r disc_doc"
obtains disc_nodes where "h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes" and
"disc_node \<in> 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 \<turnstile> get_host shadow_root_ptr \<rightarrow>\<^sub>r host"
assumes "h \<turnstile> get_tag_name host \<rightarrow>\<^sub>r tag"
shows "tag \<in> safe_shadow_root_element_types"
proof -
obtain host' where "host' |\<in>| element_ptr_kinds h" and
"|h \<turnstile> get_tag_name host'|\<^sub>r \<in> safe_shadow_root_element_types"
and "h \<turnstile> get_shadow_root host' \<rightarrow>\<^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)
+ by (smt (z3) assms(1) get_host_ptr_in_heap local.get_shadow_root_ok returns_result_select_result)
then have "host = host'"
by (meson assms(1) assms(2) assms(3) shadow_root_host_dual shadow_root_same_host)
then show ?thesis
by (smt \<open>\<And>thesis. (\<And>host'. \<lbrakk>host' |\<in>| element_ptr_kinds h; |h \<turnstile> get_tag_name host'|\<^sub>r \<in>
safe_shadow_root_element_types; h \<turnstile> get_shadow_root host' \<rightarrow>\<^sub>r Some shadow_root_ptr\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow>
thesis\<close> \<open>h \<turnstile> get_shadow_root host' \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> 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 = (\<Union>host \<in> fset (element_ptr_kinds h).
(case |h \<turnstile> get_shadow_root host|\<^sub>r of Some shadow_root \<Rightarrow> {(cast host, cast shadow_root)} | None \<Rightarrow> {}))"
by(auto simp add: a_host_shadow_root_rel_def split: option.splits)
moreover have "finite (\<Union>host \<in> fset (element_ptr_kinds h). (case |h \<turnstile> get_shadow_root host|\<^sub>r of
Some shadow_root \<Rightarrow> {(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 \<Rightarrow> {}))"
by(auto split: option.splits)
ultimately show ?thesis
by auto
qed
lemma a_ptr_disconnected_node_rel_finite: "finite (a_ptr_disconnected_node_rel h)"
proof -
have "a_ptr_disconnected_node_rel h = (\<Union>owner_document \<in> set |h \<turnstile> document_ptr_kinds_M|\<^sub>r.
(\<Union>disconnected_node \<in> set |h \<turnstile> get_disconnected_nodes owner_document|\<^sub>r.
{(cast owner_document, cast disconnected_node)}))"
by(auto simp add: a_ptr_disconnected_node_rel_def)
moreover have "finite (\<Union>owner_document \<in> set |h \<turnstile> document_ptr_kinds_M|\<^sub>r.
(\<Union>disconnected_node \<in> set |h \<turnstile> get_disconnected_nodes owner_document|\<^sub>r.
{(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disconnected_node)}))"
by simp
ultimately show ?thesis
by simp
qed
lemma heap_is_wellformed_children_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children \<Longrightarrow>
child |\<in>| 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 \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
node \<in> set disc_nodes \<Longrightarrow> node |\<in>| 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 \<Longrightarrow>
h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children' \<Longrightarrow>
set children \<inter> set children' \<noteq> {} \<Longrightarrow> 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 \<Longrightarrow>
h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
h \<turnstile> get_disconnected_nodes document_ptr' \<rightarrow>\<^sub>r disc_nodes' \<Longrightarrow> set disc_nodes \<inter> set disc_nodes' \<noteq> {} \<Longrightarrow>
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 \<Longrightarrow>
h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow>
set children \<inter> 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 \<Longrightarrow>
h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow> 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 \<Longrightarrow>
h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> 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 \<Longrightarrow>
node_ptr |\<in>| node_ptr_kinds h \<Longrightarrow> \<not>(\<exists>parent \<in> fset (object_ptr_kinds h).
node_ptr \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r) \<Longrightarrow> (\<exists>document_ptr \<in> fset (document_ptr_kinds h).
node_ptr \<in> set |h \<turnstile> 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 \<Longrightarrow> finite (parent_child_rel h)"
using CD.parent_child_rel_finite by blast
lemma parent_child_rel_acyclic: "heap_is_wellformed h \<Longrightarrow> 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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptr parent \<Longrightarrow>
(parent, child_ptr) \<in> parent_child_rel h \<Longrightarrow> child_ptr |\<in>| 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 \<open>get\_parent\<close>
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 \<open>get\_disconnected\_nodes\<close>
subsubsection \<open>set\_disconnected\_nodes\<close>
paragraph \<open>get\_disconnected\_nodes\<close>
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 \<open>get\_root\_node\<close>
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 \<open>to\_tree\_order\<close>
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 \<open>get\_root\_node\<close>
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 \<open>remove\_child\<close>
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 \<open>A wellformed heap\<close>
subsubsection \<open>get\_parent\<close>
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 \<open>remove\_shadow\_root\<close>
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 \<turnstile> remove_shadow_root ptr \<rightarrow>\<^sub>h h'"
shows "known_ptrs h'" and "type_wf h'" "heap_is_wellformed h'"
proof -
obtain shadow_root_ptr h2 where
"h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr" and
"h \<turnstile> get_child_nodes (cast shadow_root_ptr) \<rightarrow>\<^sub>r []" and
"h \<turnstile> get_disconnected_nodes (cast shadow_root_ptr) \<rightarrow>\<^sub>r []" and
h2: "h \<turnstile> set_shadow_root ptr None \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> delete_M shadow_root_ptr \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: remove_shadow_root_def elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
split: option.splits if_splits)
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_shadow_root_writes h2]
using \<open>type_wf h\<close> 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="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_shadow_root_writes h2])
using set_shadow_root_pointers_preserved
apply blast
by (auto simp add: reflp_def transp_def)
have node_ptr_kinds_eq_h: "node_ptr_kinds h = node_ptr_kinds h2"
using object_ptr_kinds_eq_h
by (simp add: node_ptr_kinds_def)
have element_ptr_kinds_eq_h: "element_ptr_kinds h = element_ptr_kinds h2"
using node_ptr_kinds_eq_h
by (simp add: element_ptr_kinds_def)
have document_ptr_kinds_eq_h: "document_ptr_kinds h = document_ptr_kinds h2"
using object_ptr_kinds_eq_h
by (simp add: document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h2"
using object_ptr_kinds_eq_h
by (simp add: document_ptr_kinds_eq_h shadow_root_ptr_kinds_def)
have "known_ptrs h2"
using \<open>known_ptrs h\<close> object_ptr_kinds_eq_h known_ptrs_subset
by blast
have object_ptr_kinds_eq_h2: "object_ptr_kinds h' |\<subseteq>| 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' |\<union>| {|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' |\<union>| {|cast shadow_root_ptr|}"
using object_ptr_kinds_eq_h2
- apply(auto simp add: document_ptr_kinds_def delete_shadow_root_pointers[OF h'])[1]
- using document_ptr_kinds_def by fastforce
+ by (auto simp add: document_ptr_kinds_def delete_shadow_root_pointers[OF h'])
then
have document_ptr_kinds_eq2_h2: "document_ptr_kinds h' |\<subseteq>| document_ptr_kinds h2"
using h' delete_shadow_root_pointers
by auto
have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h' |\<subseteq>| shadow_root_ptr_kinds h2"
using object_ptr_kinds_eq_h2
apply(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)[1]
by auto
have shadow_root_ptr_kinds_eq2_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h' |\<union>| {|shadow_root_ptr|}"
- using object_ptr_kinds_eq2_h2
- apply (auto simp add: shadow_root_ptr_kinds_def)[1]
- using document_ptr_kinds_eq_h2 apply auto[1]
- apply (metis \<open>h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> assms(1) document_ptr_kinds_eq_h
- fset.map_comp local.get_shadow_root_shadow_root_ptr_in_heap shadow_root_ptr_kinds_def)
- using document_ptr_kinds_eq_h2 by auto
+ using object_ptr_kinds_eq2_h2 document_ptr_kinds_eq_h2
+ by (auto simp add: shadow_root_ptr_kinds_def)
show "known_ptrs h'"
using object_ptr_kinds_eq_h2 \<open>known_ptrs h2\<close> known_ptrs_subset
by blast
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. doc_ptr \<noteq> cast shadow_root_ptr \<Longrightarrow> h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads get_disconnected_nodes_delete_shadow_root[rotated, OF h']
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by(metis (no_types, lifting))+
then have disconnected_nodes_eq2_h2:
"\<And>doc_ptr. doc_ptr \<noteq> cast shadow_root_ptr \<Longrightarrow> |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r =
|h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^sub>r disc_nodes =
h2 \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. |h \<turnstile> get_tag_name doc_ptr|\<^sub>r = |h2 \<turnstile> get_tag_name doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_tag_name doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. |h2 \<turnstile> get_tag_name doc_ptr|\<^sub>r = |h' \<turnstile> get_tag_name doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h:
"\<And>ptr' children. h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. ptr' \<noteq> cast shadow_root_ptr \<Longrightarrow> h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children =
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr' \<noteq> cast shadow_root_ptr \<Longrightarrow> |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "cast shadow_root_ptr |\<notin>| 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:
"\<And>shadow_root_opt ptr'. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>shadow_root_opt ptr'. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
using select_result_eq by force
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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' \<subseteq> parent_child_rel h2"
using object_ptr_kinds_eq_h2
apply(auto simp add: CD.parent_child_rel_def)[1]
by (metis \<open>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 |\<notin>| object_ptr_kinds h'\<close> 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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_all_ptrs_in_heap h2"
apply(auto simp add: CD.a_all_ptrs_in_heap_def object_ptr_kinds_eq_h node_ptr_kinds_def
children_eq_h disconnected_nodes_eq_h)[1]
- apply (metis (no_types, lifting) children_eq2_h finite_set_in subsetD)
+ apply (metis (no_types, lifting) children_eq2_h subsetD)
by (metis (no_types, lifting) disconnected_nodes_eq2_h document_ptr_kinds_eq_h
- finite_set_in in_mono)
+ in_mono)
then have "CD.a_all_ptrs_in_heap h'"
apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 children_eq_h2
disconnected_nodes_eq_h2)[1]
apply(case_tac "ptr = cast shadow_root_ptr")
using object_ptr_kinds_eq_h2 children_eq_h2
apply (meson \<open>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 |\<notin>| object_ptr_kinds h'\<close>
is_OK_returns_result_I local.get_child_nodes_ptr_in_heap)
apply(auto dest!: children_eq_h2)[1]
using assms(1) children_eq_h local.heap_is_wellformed_children_in_heap node_ptr_kinds_eq_h
node_ptr_kinds_eq_h2 apply blast
apply (meson \<open>known_ptrs h'\<close> \<open>type_wf h'\<close> local.get_child_nodes_ok local.known_ptrs_known_ptr
returns_result_select_result)
by (metis (no_types, lifting) \<open>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 |\<notin>| object_ptr_kinds h'\<close>
\<open>type_wf h2\<close> assms(1) disconnected_nodes_eq2_h2 disconnected_nodes_eq_h document_ptr_kinds_commutes
document_ptr_kinds_eq2_h2 fin_mono local.get_disconnected_nodes_ok
local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h node_ptr_kinds_eq_h2
returns_result_select_result)
moreover
have "CD.a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
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 \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
using children_eq_h2 concat_map_all_distinct[of "(\<lambda>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r)"]
- apply (metis (no_types, lifting) children_eq2_h2 finite_fset fmember_iff_member_fset fset_mp
+ apply (metis (no_types, lifting) children_eq2_h2 finite_fset fset_mp
object_ptr_kinds_eq_h2 set_sorted_list_of_set)
apply(case_tac "x = cast shadow_root_ptr")
using \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
apply(case_tac "y = cast shadow_root_ptr")
using \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
using children_eq_h2 distinct_concat_map_E(1)[of "(\<lambda>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r)"]
- apply (smt IntI children_eq2_h2 empty_iff finite_fset fmember_iff_member_fset fset_mp
+ apply (smt (verit) IntI children_eq2_h2 empty_iff finite_fset fset_mp
object_ptr_kinds_eq_h2 set_sorted_list_of_set)
apply(auto simp add: intro!: distinct_concat_map_I)[1]
apply(case_tac "x = cast shadow_root_ptr")
using \<open>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 |\<notin>| object_ptr_kinds h'\<close> document_ptr_kinds_commutes
apply blast
apply (metis (mono_tags, lifting) \<open>local.CD.a_distinct_lists h2\<close> \<open>type_wf h'\<close>
disconnected_nodes_eq_h2 is_OK_returns_result_E local.CD.distinct_lists_disconnected_nodes
local.get_disconnected_nodes_ok select_result_I2)
apply(case_tac "x = cast shadow_root_ptr")
using \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
apply(case_tac "y = cast shadow_root_ptr")
using \<open>cast shadow_root_ptr |\<notin>| object_ptr_kinds h'\<close> apply simp
proof -
fix x and y and xa
assume a1: "x |\<in>| document_ptr_kinds h'"
assume a2: "y |\<in>| document_ptr_kinds h'"
assume a3: "x \<noteq> y"
assume a4: "x \<noteq> cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr"
assume a5: "y \<noteq> cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr"
assume a6: "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
assume a7: "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
assume "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(insort (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) (sorted_list_of_set (fset (document_ptr_kinds h') -
{cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr})))))"
then show False
using a7 a6 a5 a4 a3 a2 a1 by (metis (no_types) IntI
distinct_concat_map_E(1)[of "(\<lambda>ptr. |h2 \<turnstile> get_disconnected_nodes ptr|\<^sub>r)"] disconnected_nodes_eq2_h2
- empty_iff finite_fset finsert.rep_eq fmember_iff_member_fset insert_iff set_sorted_list_of_set
+ empty_iff finite_fset finsert.rep_eq insert_iff set_sorted_list_of_set
sorted_list_of_set_insert_remove)
next
fix x xa xb
assume 0: "distinct (concat (map (\<lambda>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h2)))))"
and 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(insort (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) (sorted_list_of_set (fset (document_ptr_kinds h') -
{cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr})))))"
and 2: "(\<Union>x\<in>fset (object_ptr_kinds h2). set |h2 \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>set (insort (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) (sorted_list_of_set (fset (document_ptr_kinds h') -
{cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr}))). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h'"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h'"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
then show "False"
apply(cases "xa = cast shadow_root_ptr")
using \<open>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 |\<notin>| object_ptr_kinds h'\<close> apply blast
apply(cases "xb = cast shadow_root_ptr")
using \<open>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 |\<notin>| object_ptr_kinds h'\<close> document_ptr_kinds_commutes
apply blast
by (metis (no_types, opaque_lifting) \<open>local.CD.a_distinct_lists h2\<close> \<open>type_wf h'\<close> children_eq2_h2
disconnected_nodes_eq_h2 fset_rev_mp is_OK_returns_result_E local.CD.distinct_lists_no_parent
local.get_disconnected_nodes_ok object_ptr_kinds_eq_h2 select_result_I2)
qed
moreover
have "CD.a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h2"
by(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
node_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h)
then have "CD.a_owner_document_valid h'"
apply(auto simp add: CD.a_owner_document_valid_def document_ptr_kinds_eq_h2 node_ptr_kinds_eq_h2
disconnected_nodes_eq2_h2)[1]
- by (smt \<open>h \<turnstile> get_child_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
+ by (smt (z3) \<open>h \<turnstile> get_child_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
\<open>h \<turnstile> get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \<rightarrow>\<^sub>r []\<close> \<open>local.CD.a_distinct_lists h\<close>
- children_eq2_h children_eq2_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 finite_set_in finsert_iff
+ children_eq2_h children_eq2_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 finsert_iff
funion_finsert_right local.CD.distinct_lists_no_parent object_ptr_kinds_eq2_h2 object_ptr_kinds_eq_h
select_result_I2 sup_bot.comm_neutral)
ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'"
by(simp add: CD.heap_is_wellformed_def)
moreover
have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using \<open>heap_is_wellformed h\<close>
by(simp add: heap_is_wellformed_def)
then
have "acyclic (parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2)"
proof -
have "a_host_shadow_root_rel h2 \<subseteq> 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) \<open>type_wf h2\<close> 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) \<open>type_wf h2\<close> image_eqI is_OK_returns_result_E
local.get_shadow_root_ok mem_Collect_eq prod.simps(2) select_result_I2)
moreover have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2"
by (simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq2_h document_ptr_kinds_eq_h)
ultimately show ?thesis
using \<open>parent_child_rel h = parent_child_rel h2\<close>
by (smt \<open>acyclic (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union>
local.a_ptr_disconnected_node_rel h)\<close> acyclic_subset subset_refl sup_mono)
qed
then
have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')"
proof -
have "a_host_shadow_root_rel h' \<subseteq> a_host_shadow_root_rel h2"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2)
moreover have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h'"
apply(simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2)
by (metis (no_types, lifting) \<open>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 |\<notin>| object_ptr_kinds h'\<close>
\<open>h \<turnstile> get_child_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
\<open>h \<turnstile> get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \<rightarrow>\<^sub>r []\<close> \<open>local.CD.a_distinct_lists h\<close>
disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 document_ptr_kinds_commutes is_OK_returns_result_I
local.CD.distinct_lists_no_parent local.get_disconnected_nodes_ptr_in_heap select_result_I2)
ultimately show ?thesis
using \<open>parent_child_rel h' \<subseteq> parent_child_rel h2\<close>
\<open>acyclic (parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2)\<close>
using acyclic_subset order_refl sup_mono
by (metis (no_types, opaque_lifting))
qed
moreover
have "a_all_ptrs_in_heap h"
using \<open>heap_is_wellformed h\<close>
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) \<open>h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> 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 \<open>heap_is_wellformed h\<close>
by(simp add: heap_is_wellformed_def)
then
have "a_distinct_lists h2"
apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1]
apply(auto intro!: distinct_concat_map_I split: option.splits)[1]
apply(case_tac "x = ptr")
apply(simp)
apply (metis (no_types, opaque_lifting) assms(2) h2 is_OK_returns_result_I
l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root
l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq
returns_result_select_result)
apply(case_tac "y = ptr")
apply(simp)
apply (metis (no_types, opaque_lifting) assms(2) h2 is_OK_returns_result_I
l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root
l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq
returns_result_select_result)
by (metis \<open>type_wf h2\<close> 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 \<open>heap_is_wellformed h\<close>
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 \<open>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 |\<notin>| object_ptr_kinds h'\<close>
+ by (smt (z3) \<open>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 |\<notin>| object_ptr_kinds h'\<close>
\<open>h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r Some shadow_root_ptr\<close> assms(2) document_ptr_kinds_commutes
- element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok
+ element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 local.get_shadow_root_ok
option.inject returns_result_select_result select_result_I2 shadow_root_ptr_kinds_commutes)
ultimately show "heap_is_wellformed h'"
by(simp add: heap_is_wellformed_def)
qed
end
interpretation i_remove_shadow_root_wf?: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf get_tag_name get_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs
set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root
get_shadow_root_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host
get_host_locs get_disconnected_document get_disconnected_document_locs remove_shadow_root
remove_shadow_root_locs known_ptrs get_parent get_parent_locs
by(auto simp add: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_root\_node\<close>
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 \<open>get\_parent\_get\_host\_get\_disconnected\_document\<close>
locale l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs
known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs
get_disconnected_document get_disconnected_document_locs +
l_get_disconnected_document get_disconnected_document get_disconnected_document_locs +
l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes
get_child_nodes_locs get_parent get_parent_locs +
l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs +
l_get_host get_host get_host_locs +
l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs
for get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and get_parent :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr option) prog"
and get_parent_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
begin
lemma a_host_shadow_root_rel_shadow_root:
"h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r shadow_root_option \<Longrightarrow> shadow_root_option = Some shadow_root \<longleftrightarrow>
((cast host, cast shadow_root) \<in> a_host_shadow_root_rel h)"
apply(auto simp add: a_host_shadow_root_rel_def)[1]
by(metis (mono_tags, lifting) case_prodI is_OK_returns_result_I
l_get_shadow_root.get_shadow_root_ptr_in_heap local.l_get_shadow_root_axioms mem_Collect_eq
pair_imageI select_result_I2)
lemma a_host_shadow_root_rel_host:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host \<Longrightarrow>
((cast host, cast shadow_root) \<in> a_host_shadow_root_rel h)"
apply(auto simp add: a_host_shadow_root_rel_def)[1]
using shadow_root_host_dual
by (metis (no_types, lifting) Collect_cong a_host_shadow_root_rel_shadow_root
local.a_host_shadow_root_rel_def split_cong)
lemma a_ptr_disconnected_node_rel_disconnected_node:
"h \<turnstile> get_disconnected_nodes document \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow> node_ptr \<in> set disc_nodes \<longleftrightarrow>
(cast document, cast node_ptr) \<in> a_ptr_disconnected_node_rel h"
apply(auto simp add: a_ptr_disconnected_node_rel_def)[1]
by (smt CD.get_disconnected_nodes_ptr_in_heap case_prodI is_OK_returns_result_I mem_Collect_eq
pair_imageI select_result_I2)
lemma a_ptr_disconnected_node_rel_document:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_document node_ptr \<rightarrow>\<^sub>r document \<Longrightarrow>
(cast document, cast node_ptr) \<in> a_ptr_disconnected_node_rel h"
apply(auto simp add: a_ptr_disconnected_node_rel_def)[1]
using disc_doc_disc_node_dual
by (metis (no_types, lifting) local.a_ptr_disconnected_node_rel_def
a_ptr_disconnected_node_rel_disconnected_node)
lemma heap_wellformed_induct_si [consumes 1, case_names step]:
assumes "heap_is_wellformed h"
assumes "\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children \<Longrightarrow>
P (cast child))
\<Longrightarrow> (\<And>shadow_root host. parent = cast host \<Longrightarrow> h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root \<Longrightarrow>
P (cast shadow_root))
\<Longrightarrow> (\<And>owner_document disc_nodes node_ptr. parent = cast owner_document \<Longrightarrow>
h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow> node_ptr \<in> set disc_nodes \<Longrightarrow> P (cast node_ptr))
\<Longrightarrow> P parent"
shows "P ptr"
proof -
fix ptr
have "finite (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using a_host_shadow_root_rel_finite a_ptr_disconnected_node_rel_finite
using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl
by auto
then
have "wf ((parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<inverse>)"
using assms(1)
apply(simp add: heap_is_wellformed_def)
by (simp add: finite_acyclic_wf_converse local.CD.parent_child_rel_impl)
then show "?thesis"
proof (induct rule: wf_induct_rule)
case (less parent)
then show ?case
apply(auto)[1]
using assms a_ptr_disconnected_node_rel_disconnected_node a_host_shadow_root_rel_shadow_root
local.CD.parent_child_rel_child
by blast
qed
qed
lemma heap_wellformed_induct_rev_si [consumes 1, case_names step]:
assumes "heap_is_wellformed h"
assumes "\<And>child. (\<And>parent child_node. child = cast child_node \<Longrightarrow>
h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent)
\<Longrightarrow> (\<And>host shadow_root. child = cast shadow_root \<Longrightarrow> h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host \<Longrightarrow>
P (cast host))
\<Longrightarrow> (\<And>disc_doc disc_node. child = cast disc_node \<Longrightarrow>
h \<turnstile> get_disconnected_document disc_node \<rightarrow>\<^sub>r disc_doc\<Longrightarrow> P (cast disc_doc))
\<Longrightarrow> P child"
shows "P ptr"
proof -
fix ptr
have "finite (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using a_host_shadow_root_rel_finite a_ptr_disconnected_node_rel_finite
using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl
by auto
then
have "wf (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using assms(1)
apply(simp add: heap_is_wellformed_def)
by (simp add: finite_acyclic_wf)
then show "?thesis"
proof (induct rule: wf_induct_rule)
case (less parent)
then show ?case
apply(auto)[1]
using parent_child_rel_parent a_host_shadow_root_rel_host a_ptr_disconnected_node_rel_document
using assms(1) assms(2) by auto
qed
qed
end
interpretation i_get_parent_get_host_get_disconnected_document_wf?:
l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed
parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document
get_disconnected_document_locs known_ptrs get_parent get_parent_locs
by(auto simp add: l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_get_parent_get_host_wf =
l_heap_is_wellformed_defs +
l_get_parent_defs +
l_get_shadow_root_defs +
l_get_host_defs +
l_get_child_nodes_defs +
l_get_disconnected_document_defs +
l_get_disconnected_nodes_defs +
assumes heap_wellformed_induct_si [consumes 1, case_names step]:
"heap_is_wellformed h
\<Longrightarrow> (\<And>parent. (\<And>children child. h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> P (cast child))
\<Longrightarrow> (\<And>shadow_root host. parent = cast host \<Longrightarrow>
h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root \<Longrightarrow> P (cast shadow_root))
\<Longrightarrow> (\<And>owner_document disc_nodes node_ptr. parent = cast owner_document \<Longrightarrow>
h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes \<Longrightarrow> node_ptr \<in> set disc_nodes \<Longrightarrow>
P (cast node_ptr))
\<Longrightarrow> P parent)
\<Longrightarrow> P ptr"
assumes heap_wellformed_induct_rev_si [consumes 1, case_names step]:
"heap_is_wellformed h
\<Longrightarrow> (\<And>child. (\<And>parent child_node. child = cast child_node \<Longrightarrow>
h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent \<Longrightarrow> P parent)
\<Longrightarrow> (\<And>host shadow_root. child = cast shadow_root \<Longrightarrow>
h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host \<Longrightarrow> P (cast host))
\<Longrightarrow> (\<And>disc_doc disc_node. child = cast disc_node \<Longrightarrow>
h \<turnstile> get_disconnected_document disc_node \<rightarrow>\<^sub>r disc_doc \<Longrightarrow> P (cast disc_doc))
\<Longrightarrow> P child)
\<Longrightarrow> P ptr"
lemma l_get_parent_get_host_wf_is_get_parent_get_host_wf [instances]:
"l_get_parent_get_host_wf heap_is_wellformed get_parent get_shadow_root get_host get_child_nodes
get_disconnected_document get_disconnected_nodes"
using heap_wellformed_induct_si heap_wellformed_induct_rev_si
using l_get_parent_get_host_wf_def by blast
subsubsection \<open>get\_host\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
begin
lemma get_host_ok [simp]:
assumes "heap_is_wellformed h"
assumes "type_wf h"
assumes "known_ptrs h"
assumes "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
shows "h \<turnstile> ok (get_host shadow_root_ptr)"
proof -
obtain host where host: "host |\<in>| element_ptr_kinds h"
and "|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types"
and shadow_root: "h \<turnstile> get_shadow_root host \<rightarrow>\<^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)
+ by (smt (z3) returns_result_select_result)
obtain host_candidates where
host_candidates: "h \<turnstile> filter_M (\<lambda>element_ptr. Heap_Error_Monad.bind (get_shadow_root element_ptr)
(\<lambda>shadow_root_opt. return (shadow_root_opt = Some shadow_root_ptr)))
(sorted_list_of_set (fset (element_ptr_kinds h)))
\<rightarrow>\<^sub>r host_candidates"
apply(drule is_OK_returns_result_E[rotated])
using get_shadow_root_ok assms(2)
by(auto intro!: filter_M_is_OK_I bind_pure_I bind_is_OK_I2)
then have "host_candidates = [host]"
apply(rule filter_M_ex1)
using host apply(auto)[1]
- apply (smt assms(1) assms(2) bind_pure_returns_result_I2 bind_returns_result_E finite_set_in host
+ apply (smt (verit) assms(1) assms(2) bind_pure_returns_result_I2 bind_returns_result_E host
local.get_shadow_root_ok local.get_shadow_root_pure local.shadow_root_same_host return_returns_result
returns_result_eq shadow_root sorted_list_of_fset.rep_eq sorted_list_of_fset_simps(1))
apply (simp add: bind_pure_I)
apply(auto intro!: bind_pure_returns_result_I)[1]
- apply (smt assms(2) bind_pure_returns_result_I2 host local.get_shadow_root_ok
+ apply (smt (verit) assms(2) bind_pure_returns_result_I2 host local.get_shadow_root_ok
local.get_shadow_root_pure return_returns_result returns_result_eq shadow_root)
done
then
show ?thesis
using host_candidates host assms(1) get_shadow_root_ok
apply(auto simp add: get_host_def known_ptrs_known_ptr
intro!: bind_is_OK_pure_I filter_M_pure_I filter_M_is_OK_I bind_pure_I split: list.splits)[1]
using assms(2) apply blast
apply (meson list.distinct(1) returns_result_eq)
by (meson list.distinct(1) list.inject returns_result_eq)
qed
end
interpretation i_get_host_wf?: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_disconnected_document get_disconnected_document_locs known_ptr known_ptrs type_wf get_host
get_host_locs get_shadow_root get_shadow_root_locs get_child_nodes get_child_nodes_locs
get_disconnected_nodes get_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed
parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
by(auto simp add: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
locale l_get_host_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_host_defs +
assumes get_host_ok: "heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow>
shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<Longrightarrow> h \<turnstile> 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 \<open>get\_root\_node\_si\<close>
locale l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf +
l_get_parent_get_host_wf +
l_get_host_wf
begin
lemma get_root_node_ptr_in_heap:
assumes "h \<turnstile> ok (get_root_node_si ptr)"
shows "ptr |\<in>| 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 |\<in>| object_ptr_kinds h"
shows "h \<turnstile> 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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
and "h' \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors'"
and "\<And>p children children'. h \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children
\<Longrightarrow> h' \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children' \<Longrightarrow> set children' \<subseteq> set children"
and "\<And>p shadow_root_option shadow_root_option'. h \<turnstile> get_shadow_root p \<rightarrow>\<^sub>r shadow_root_option \<Longrightarrow>
h' \<turnstile> get_shadow_root p \<rightarrow>\<^sub>r shadow_root_option' \<Longrightarrow> (if shadow_root_option = None
then shadow_root_option' = None else shadow_root_option' = None \<or> shadow_root_option' = shadow_root_option)"
and "node \<notin> 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 \<notin> set ancestors'"
proof -
have object_ptr_kinds_M_eq:
"\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
using object_ptr_kinds_eq3
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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 \<open>h \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors\<close> 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 \<open>h' \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors'\<close> 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 |\<in>| object_ptr_kinds h"
using local.get_ancestors_si_ptr_in_heap object_ptr_kinds_eq3 step.prems(2) by fastforce
have "node \<noteq> child"
using ancestors_remains step.prems(3) by auto
have 1: "\<And>p parent. h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent \<Longrightarrow> h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
proof -
fix p parent
assume "h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
then obtain children' where
children': "h' \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children'" and
p_in_children': "p \<in> set children'"
using get_parent_child_dual by blast
obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^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 \<open>h' \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent\<close> get_parent_parent_in_heap is_OK_returns_result_E
local.known_ptrs_known_ptr object_ptr_kinds_eq3)
have "p \<in> set children"
using assms(5) children children' p_in_children'
by blast
then show "h \<turnstile> get_parent p \<rightarrow>\<^sub>r Some parent"
using child_parent_dual assms(1) children known_ptrs type_wf by blast
qed
have 2: "\<And>p host. h' \<turnstile> get_host p \<rightarrow>\<^sub>r host \<Longrightarrow> h \<turnstile> get_host p \<rightarrow>\<^sub>r host"
proof -
fix p host
assume "h' \<turnstile> get_host p \<rightarrow>\<^sub>r host"
then have "h' \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some p"
using local.shadow_root_host_dual by blast
then have "h \<turnstile> get_shadow_root host \<rightarrow>\<^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 \<turnstile> get_host p \<rightarrow>\<^sub>r host"
by (metis assms(1) is_OK_returns_result_E known_ptrs local.get_host_ok
local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host
type_wf)
qed
show ?case
proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then show ?thesis
using step(4) step(5) \<open>node \<noteq> child\<close>
apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2
split: option.splits)[1]
by (metis "2" assms(1) shadow_root_same_host list.set_intros(2) shadow_root_host_dual
step.hyps(2) step.prems(3) type_wf)
next
case (Some node_child)
then
show ?thesis
using step(4) step(5) \<open>node \<noteq> child\<close>
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 \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
shows "ptr' |\<in>| 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 \<turnstile> get_ancestors_si x \<rightarrow>\<^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 \<open>x = a\<close> by blast
next
case False
obtain ptr'' where ptr'': "h \<turnstile> get_ancestors_si ptr'' \<rightarrow>\<^sub>r ancestors"
using \<open> h \<turnstile> get_ancestors_si x \<rightarrow>\<^sub>r a # ancestors\<close> Cons.prems(2) False
apply(auto simp add: get_ancestors_si_def elim!: bind_returns_result_E2)[1]
apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1]
apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1]
apply (metis local.get_ancestors_si_def)
by (simp add: local.get_ancestors_si_def)
then show ?thesis
using Cons.hyps Cons.prems(2) False by auto
qed
qed
lemma get_ancestors_si_reads:
assumes "heap_is_wellformed h"
shows "reads get_ancestors_si_locs (get_ancestors_si node_ptr) h h'"
proof (insert assms(1), induct rule: heap_wellformed_induct_rev_si)
case (step child)
then show ?case
using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def]
get_host_reads[unfolded reads_def]
apply(simp (no_asm) add: get_ancestors_si_def)
by(auto simp add: get_ancestors_si_locs_def get_parent_reads_pointers
intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF return_reads]
reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads]
reads_subset[OF get_host_reads]
split: option.splits)
qed
lemma get_ancestors_si_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
and "ancestor \<in> set ancestors"
and "h \<turnstile> get_ancestors_si ancestor \<rightarrow>\<^sub>r ancestor_ancestors"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "set ancestor_ancestors \<subseteq> 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 |\<in>| object_ptr_kinds h"
using get_ancestors_si_ptr_in_heap step(4) by auto
(* then have "h \<turnstile> check_in_heap child \<rightarrow>\<^sub>r ()"
using returns_result_select_result by force *)
obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
using step(4)
by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_I
elim!: bind_returns_result_E2 split: option.splits)
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then show ?case
using step(4) \<open>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\<close>
apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2)[1]
by (metis (no_types, lifting) assms(4) empty_iff empty_set select_result_I2 set_ConsD
step.prems(1) step.prems(2))
next
case (Some shadow_root_child)
then
have "cast shadow_root_child |\<in>| document_ptr_kinds h"
using \<open>child |\<in>| object_ptr_kinds h\<close>
- apply(auto simp add: document_ptr_kinds_def split: option.splits)[1]
- by (metis (mono_tags, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes
- document_ptr_kinds_def fset.map_comp shadow_root_ptr_casts_commute)
+ apply(auto simp add: document_ptr_kinds_def image_iff Bex_def split: option.splits)[1]
+ by (metis (mono_tags) shadow_root_ptr_casts_commute)
then
have "shadow_root_child |\<in>| shadow_root_ptr_kinds h"
using shadow_root_ptr_kinds_commutes by blast
obtain host where host: "h \<turnstile> get_host shadow_root_child \<rightarrow>\<^sub>r host"
using get_host_ok assms
by (meson \<open>shadow_root_child |\<in>| shadow_root_ptr_kinds h\<close> is_OK_returns_result_E)
then
have "h \<turnstile> get_ancestors_si (cast host) \<rightarrow>\<^sub>r tl_ancestors"
using Some step(4) tl_ancestors None
by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_returns_result_I
elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
then
show ?case
using step(2) Some host step(5) tl_ancestors
using assms(4) dual_order.trans eq_iff returns_result_eq set_ConsD set_subset_Cons
shadow_root_ptr_casts_commute document_ptr_casts_commute step.prems(1)
by (smt case_optionE local.shadow_root_host_dual option.case_distrib option.distinct(1))
qed
next
case (Some child_node)
note s1 = Some
obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using \<open>child |\<in>| object_ptr_kinds h\<close> assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs]
by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes)
then show ?case
proof (induct parent_opt)
case None
then have "ancestors = [child]"
using step(4) s1
apply(simp add: get_ancestors_si_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
show ?case
using step(4) step(5)
apply(auto simp add: \<open>ancestors = [child]\<close>)[1]
using assms(4) returns_result_eq by fastforce
next
case (Some parent)
then
have "h \<turnstile> get_ancestors_si parent \<rightarrow>\<^sub>r tl_ancestors"
using s1 tl_ancestors step(4)
by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2
split: option.splits dest: returns_result_eq)
show ?case
by (metis (no_types, lifting) Some.prems \<open>h \<turnstile> get_ancestors_si parent \<rightarrow>\<^sub>r tl_ancestors\<close>
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 \<turnstile> get_ancestors_si some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast child \<in> set ancestors"
and "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "parent \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors: "h \<turnstile> get_ancestors_si (cast child) \<rightarrow>\<^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 \<in> 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 \<turnstile> get_ancestors_si some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast shadow_root \<in> set ancestors"
and "h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "cast host \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors: "h \<turnstile> get_ancestors_si (cast shadow_root) \<rightarrow>\<^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 \<in> set child_ancestors"
apply(simp add: get_ancestors_si_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)]
get_ancestors_si_ptr)
then show ?thesis
using assms child_ancestors get_ancestors_si_subset by blast
qed
lemma get_ancestors_si_obtains_children_or_shadow_root:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
and "h \<turnstile> get_ancestors_si ptr \<rightarrow>\<^sub>r ancestors"
and "ancestor \<noteq> ptr"
and "ancestor \<in> set ancestors"
shows "((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast ancestor_child \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)
\<or> ((\<forall>ancestor_element shadow_root. ancestor = cast ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow> cast shadow_root \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis)"
proof (insert assms(4) assms(5) assms(6), induct ptr arbitrary: ancestors
rule: heap_wellformed_induct_rev_si[OF assms(1)])
case (1 child)
then show ?case
proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then obtain shadow_root where shadow_root: "child = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root"
using 1(4) 1(5) 1(6)
by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2
split: option.splits)
then obtain host where host: "h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host"
by (metis "1.prems"(1) assms(1) assms(2) assms(3) document_ptr_kinds_commutes
get_ancestors_si_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_si_ptr local.get_host_ok
shadow_root_ptr_kinds_commutes)
then obtain host_ancestors where host_ancestors: "h \<turnstile> get_ancestors_si (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host) \<rightarrow>\<^sub>r host_ancestors"
by (metis "1.prems"(1) assms(1) assms(2) assms(3) get_ancestors_si_also_host get_ancestors_si_ok
get_ancestors_si_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_si_ptr shadow_root)
then have "ancestors = cast shadow_root # host_ancestors"
using 1(4) 1(5) 1(3) None shadow_root host
by(auto simp add: get_ancestors_si_def[of child, simplified shadow_root]
elim!: bind_returns_result_E2 dest!: returns_result_eq[OF host] split: option.splits)
then show ?thesis
proof (cases "ancestor = cast host")
case True
then show ?thesis
using "1.prems"(1) host local.get_ancestors_si_ptr local.shadow_root_host_dual shadow_root
by blast
next
case False
have "ancestor \<in> set ancestors"
using host host_ancestors 1(3) get_ancestors_si_also_host assms(1) assms(2) assms(3)
using "1.prems"(3) by blast
then have "((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set host_ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis) \<or>
((\<forall>ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow>
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 \<in> set host_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)"
using "1.hyps"(2) "1.prems"(2) False \<open>ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\<close>
host host_ancestors shadow_root
by auto
then show ?thesis
using \<open>ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\<close> by auto
qed
next
case (Some child_node)
then obtain parent where parent: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent"
using 1(4) 1(5) 1(6)
by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2
split: option.splits)
then obtain parent_ancestors where parent_ancestors: "h \<turnstile> get_ancestors_si parent \<rightarrow>\<^sub>r parent_ancestors"
by (meson assms(1) assms(2) assms(3) get_ancestors_si_ok is_OK_returns_result_E
local.get_parent_parent_in_heap)
then have "ancestors = cast child_node # parent_ancestors"
using 1(4) 1(5) 1(3) Some
by(auto simp add: get_ancestors_si_def[of child, simplified Some]
elim!: bind_returns_result_E2 dest!: returns_result_eq[OF parent] split: option.splits)
then show ?thesis
proof (cases "ancestor = parent")
case True
then show ?thesis
by (metis (no_types, lifting) "1.prems"(1) Some local.get_ancestors_si_ptr
local.get_parent_child_dual node_ptr_casts_commute parent)
next
case False
have "ancestor \<in> set ancestors"
by (simp add: "1.prems"(3))
then have "((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis) \<or>
((\<forall>ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow>
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 \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)"
using "1.hyps"(1) "1.prems"(2) False Some \<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close>
parent parent_ancestors
by auto
then show ?thesis
using \<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close> by auto
qed
qed
qed
lemma a_host_shadow_root_rel_shadow_root:
"h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root \<Longrightarrow> (cast host, cast shadow_root) \<in> a_host_shadow_root_rel h"
by(auto simp add: is_OK_returns_result_I get_shadow_root_ptr_in_heap a_host_shadow_root_rel_def)
lemma get_ancestors_si_parent_child_a_host_shadow_root_rel:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> get_ancestors_si child \<rightarrow>\<^sub>r ancestors"
shows "(ptr, child) \<in> (parent_child_rel h \<union> a_host_shadow_root_rel h)\<^sup>* \<longleftrightarrow> ptr \<in> set ancestors"
proof
assume "(ptr, child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<^sup>* "
then show "ptr \<in> set ancestors"
proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
using assms(4) local.get_ancestors_si_ptr by blast
next
case False
obtain ptr_child where
ptr_child: "(ptr, ptr_child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h) \<and>
(ptr_child, child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<^sup>*"
using converse_rtranclE[OF 1(4)] \<open>ptr \<noteq> child\<close>
by metis
then show ?thesis
proof(cases "(ptr, ptr_child) \<in> 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 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children"
proof -
assume a1: "\<And>children. \<lbrakk>h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children; ptr_child_node \<in> set children\<rbrakk>
\<Longrightarrow> thesis"
have "ptr |\<in>| object_ptr_kinds h"
using CD.parent_child_rel_parent_in_heap True by blast
moreover have "ptr_child_node \<in> set |h \<turnstile> 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 \<open>type_wf h\<close> \<open>known_ptrs h\<close>
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) \<in> (parent_child_rel h \<union> 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 \<in> set ancestors"
using 1 by auto
moreover have "h \<turnstile> get_parent ptr_child_node \<rightarrow>\<^sub>r Some ptr"
using assms(1) children ptr_child_node child_parent_dual
using \<open>known_ptrs h\<close> \<open>type_wf h\<close> by blast
ultimately show ?thesis
using get_ancestors_si_also_parent assms \<open>type_wf h\<close> 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 \<turnstile> get_shadow_root host \<rightarrow>\<^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) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<^sup>*"
using ptr_child ptr_child_shadow_root by blast
ultimately have "cast shadow_root \<in> set ancestors"
using "1.hyps"(2) host by blast
moreover have "h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host"
by (metis assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_host_ok
local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host
shadow_root)
ultimately show ?thesis
using get_ancestors_si_also_host assms(1) assms(2) assms(3) assms(4) host
by blast
qed
qed
qed
next
assume "ptr \<in> set ancestors"
then show "(ptr, child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h)\<^sup>*"
proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by simp
next
case False
have "\<And>thesis. ((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast ancestor_child \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)
\<or> ((\<forall>ancestor_element shadow_root. ptr = cast ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow> cast shadow_root \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis)"
using "1.prems" False assms(1) assms(2) assms(3) assms(4) get_ancestors_si_obtains_children_or_shadow_root
by blast
then show ?thesis
proof (cases "\<forall>thesis. ((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast ancestor_child \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)")
case True
then obtain children ancestor_child
where "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and "ancestor_child \<in> set children"
and "cast ancestor_child \<in> set ancestors"
by blast
then show ?thesis
by (meson "1.hyps"(1) in_rtrancl_UnI local.CD.parent_child_rel_child r_into_rtrancl rtrancl_trans)
next
case False
obtain ancestor_element shadow_root
where "ptr = cast ancestor_element"
and "h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root"
and "cast shadow_root \<in> set ancestors"
using False \<open>\<And>thesis. ((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis) \<or>
((\<forall>ancestor_element shadow_root. ptr = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow>
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 \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)\<close>
by blast
then show ?thesis
using 1(2) a_host_shadow_root_rel_shadow_root
apply(simp)
by (meson Un_iff converse_rtrancl_into_rtrancl)
qed
qed
qed
qed
lemma get_root_node_si_root_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r root"
shows "root |\<in>| 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 \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r cast child"
shows "h \<turnstile> get_parent child \<rightarrow>\<^sub>r None"
proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev_si)
case (step c)
then show ?case
proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c")
case None
then show ?thesis
using step(4)
by(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2
split: if_splits option.splits intro!: step(2) bind_pure_returns_result_I)
next
case (Some child_node)
note s = this
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using step(4)
apply(auto simp add: get_root_node_si_def get_ancestors_si_def intro!: bind_pure_I
elim!: bind_returns_result_E2)[1]
by(auto split: option.splits)
then show ?thesis
proof(induct parent_opt)
case None
then show ?case
using Some get_root_node_si_no_parent returns_result_eq step.prems by fastforce
next
case (Some parent)
then show ?case
using step(4) s
apply(auto simp add: get_root_node_si_def get_ancestors_si_def[of c]
elim!: bind_returns_result_E2 split: option.splits list.splits if_splits)[1]
using assms(1) get_ancestors_si_never_empty apply blast
by(auto simp add: get_root_node_si_def dest: returns_result_eq
intro!: step(1) bind_pure_returns_result_I)
qed
qed
qed
lemma get_root_node_si_parent_child_a_host_shadow_root_rel:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> get_root_node_si ptr \<rightarrow>\<^sub>r root"
shows "(root, ptr) \<in> (parent_child_rel h \<union> a_host_shadow_root_rel h)\<^sup>*"
using assms
using get_ancestors_si_parent_child_a_host_shadow_root_rel get_ancestors_si_never_empty
by(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)
end
interpretation i_get_root_node_si_wf?: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs
get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name
get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document
get_disconnected_document_locs
by(auto simp add: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_disconnected\_document\<close>
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 \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
shows "h \<turnstile> ok (get_disconnected_document node_ptr)"
proof -
have "node_ptr |\<in>| node_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap)
have "\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> 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 "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
using heap_is_wellformed_children_disc_nodes
using \<open>node_ptr |\<in>| node_ptr_kinds h\<close> assms(1) by blast
then obtain some_owner_document where
"some_owner_document \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))" and
"node_ptr \<in> set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
have h5: "\<exists>!x. x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h))) \<and> h \<turnstile> Heap_Error_Monad.bind (get_disconnected_nodes x)
(\<lambda>children. return (node_ptr \<in> set children)) \<rightarrow>\<^sub>r True"
apply(auto intro!: bind_pure_returns_result_I)[1]
- apply (smt CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure
+ apply (smt (verit) CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure
\<open>\<exists>document_ptr\<in>fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r\<close>
- assms(2) bind_pure_returns_result_I2 notin_fset return_returns_result select_result_I2)
+ assms(2) bind_pure_returns_result_I2 return_returns_result select_result_I2)
apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1]
using heap_is_wellformed_one_disc_parent assms(1)
by blast
let ?filter_M = "filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (node_ptr \<in> set disconnected_nodes)))
(sorted_list_of_set (fset (document_ptr_kinds h)))"
have "h \<turnstile> 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 \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (node_ptr \<in> set disconnected_nodes)))
(sorted_list_of_set (fset (document_ptr_kinds h)))
\<rightarrow>\<^sub>r candidates"
by auto
have "candidates = [some_owner_document]"
apply(rule filter_M_ex1[OF candidates \<open>some_owner_document \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))\<close> h5])
using \<open>node_ptr \<in> set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
\<open>some_owner_document \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))\<close>
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 \<open>node_ptr |\<in>| node_ptr_kinds h\<close>
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 \<open>get\_ancestors\_di\<close>
locale l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf +
l_get_parent_get_host_wf +
l_get_host_wf +
l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma get_ancestors_di_ok:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
and "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_ancestors_di ptr)"
proof (insert assms(1) assms(4), induct rule: heap_wellformed_induct_rev_si)
case (step child)
then show ?case
using assms(2) assms(3)
apply(auto simp add: get_ancestors_di_def[of child] assms(1) get_parent_parent_in_heap
intro!: bind_is_OK_pure_I bind_pure_I split: option.splits)[1]
using local.get_parent_ok apply blast
using assms(1) get_disconnected_document_ok apply blast
apply(simp add: get_ancestors_di_def )
apply(auto intro!: bind_is_OK_pure_I split: option.splits)[1]
apply (metis (no_types, lifting) bind_is_OK_E document_ptr_kinds_commutes is_OK_returns_heap_I
local.get_ancestors_di_def local.get_disconnected_document_disconnected_document_in_heap step.hyps(3))
apply (metis (no_types, lifting) bind_is_OK_E document_ptr_kinds_commutes is_OK_returns_heap_I
local.get_ancestors_di_def local.get_disconnected_document_disconnected_document_in_heap step.hyps(3))
using assms(1) local.get_disconnected_document_disconnected_document_in_heap local.get_host_ok
shadow_root_ptr_kinds_commutes apply blast
apply (smt assms(1) bind_returns_heap_E document_ptr_casts_commute2 is_OK_returns_heap_E
is_OK_returns_heap_I l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.shadow_root_same_host
local.get_disconnected_document_disconnected_document_in_heap local.get_host_pure
local.l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.shadow_root_host_dual option.simps(4) option.simps(5)
pure_returns_heap_eq shadow_root_ptr_casts_commute2)
using get_host_ok assms(1) apply blast
by (meson assms(1) is_OK_returns_result_I local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual)
qed
lemma get_ancestors_di_ptrs_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>r ancestors"
assumes "ptr' \<in> set ancestors"
shows "ptr' |\<in>| 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 \<turnstile> get_ancestors_di x \<rightarrow>\<^sub>r a # ancestors"
by(auto simp add: get_ancestors_di_def[of a] elim!: bind_returns_result_E2 split: option.splits)
then have "x = a"
by(auto simp add: get_ancestors_di_def[of x] intro!: bind_pure_I elim!: bind_returns_result_E2
split: option.splits)
then show ?case
proof (cases "ptr' = a")
case True
then show ?thesis
using Cons.hyps Cons.prems(2) get_ancestors_di_ptr_in_heap x
using \<open>x = a\<close> by blast
next
case False
obtain ptr'' where ptr'': "h \<turnstile> get_ancestors_di ptr'' \<rightarrow>\<^sub>r ancestors"
using \<open> h \<turnstile> get_ancestors_di x \<rightarrow>\<^sub>r a # ancestors\<close> Cons.prems(2) False
apply(auto simp add: get_ancestors_di_def elim!: bind_returns_result_E2)[1]
apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1]
apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1]
apply (metis (no_types, lifting) local.get_ancestors_di_def)
apply (metis (no_types, lifting) local.get_ancestors_di_def)
by (simp add: local.get_ancestors_di_def)
then show ?thesis
using Cons.hyps Cons.prems(2) False by auto
qed
qed
lemma get_ancestors_di_reads:
assumes "heap_is_wellformed h"
shows "reads get_ancestors_di_locs (get_ancestors_di node_ptr) h h'"
proof (insert assms(1), induct rule: heap_wellformed_induct_rev_si)
case (step child)
then show ?case
using (* [[simproc del: Product_Type.unit_eq]] *) get_parent_reads[unfolded reads_def]
get_host_reads[unfolded reads_def] get_disconnected_document_reads[unfolded reads_def]
apply(auto simp add: get_ancestors_di_def[of child])[1]
by(auto simp add: get_ancestors_di_locs_def get_parent_reads_pointers
intro!: bind_pure_I reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF return_reads]
reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads]
reads_subset[OF get_host_reads] reads_subset[OF get_disconnected_document_reads]
split: option.splits list.splits
)
qed
lemma get_ancestors_di_subset:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>r ancestors"
and "ancestor \<in> set ancestors"
and "h \<turnstile> get_ancestors_di ancestor \<rightarrow>\<^sub>r ancestor_ancestors"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "set ancestor_ancestors \<subseteq> 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 |\<in>| object_ptr_kinds h"
using get_ancestors_di_ptr_in_heap step(4) by auto
(* then have "h \<turnstile> check_in_heap child \<rightarrow>\<^sub>r ()"
using returns_result_select_result by force *)
obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
using step(4)
by(auto simp add: get_ancestors_di_def[of child] intro!: bind_pure_I
elim!: bind_returns_result_E2 split: option.splits)
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
show ?case
proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then show ?case
using step(4) \<open>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\<close>
apply(auto simp add: get_ancestors_di_def[of child] elim!: bind_returns_result_E2)[1]
by (metis (no_types, lifting) assms(4) empty_iff empty_set select_result_I2 set_ConsD
step.prems(1) step.prems(2))
next
case (Some shadow_root_child)
then
have "cast shadow_root_child |\<in>| document_ptr_kinds h"
using \<open>child |\<in>| object_ptr_kinds h\<close>
- apply(auto simp add: document_ptr_kinds_def split: option.splits)[1]
- by (metis (mono_tags, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes
- document_ptr_kinds_def fset.map_comp shadow_root_ptr_casts_commute)
+ apply(auto simp add: document_ptr_kinds_def image_iff Bex_def split: option.splits)[1]
+ by (metis (mono_tags, lifting) shadow_root_ptr_casts_commute)
then
have "shadow_root_child |\<in>| shadow_root_ptr_kinds h"
using shadow_root_ptr_kinds_commutes by blast
obtain host where host: "h \<turnstile> get_host shadow_root_child \<rightarrow>\<^sub>r host"
using get_host_ok assms
by (meson \<open>shadow_root_child |\<in>| shadow_root_ptr_kinds h\<close> is_OK_returns_result_E)
then
have "h \<turnstile> get_ancestors_di (cast host) \<rightarrow>\<^sub>r tl_ancestors"
using Some step(4) tl_ancestors None
by(auto simp add: get_ancestors_di_def[of child] intro!: bind_pure_returns_result_I
elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
then
show ?case
using step(2) Some host step(5) tl_ancestors
using assms(4) dual_order.trans eq_iff returns_result_eq set_ConsD set_subset_Cons
shadow_root_ptr_casts_commute document_ptr_casts_commute step.prems(1)
by (smt case_optionE local.shadow_root_host_dual option.case_distrib option.distinct(1))
qed
next
case (Some child_node)
note s1 = Some
obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
using \<open>child |\<in>| object_ptr_kinds h\<close> assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs]
by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes)
then show ?case
proof (induct parent_opt)
case None
then obtain disc_doc where disc_doc: "h \<turnstile> get_disconnected_document child_node \<rightarrow>\<^sub>r disc_doc"
and "h \<turnstile> get_ancestors_di (cast disc_doc) \<rightarrow>\<^sub>r tl_ancestors"
using step(4) s1 tl_ancestors
apply(simp add: get_ancestors_di_def[of child])
by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I split: option.splits dest: returns_result_eq)
then show ?thesis
using step(3) step(4) step(5)
by (metis (no_types, lifting) assms(4) dual_order.trans eq_iff node_ptr_casts_commute s1
select_result_I2 set_ConsD set_subset_Cons tl_ancestors)
next
case (Some parent)
then
have "h \<turnstile> get_ancestors_di parent \<rightarrow>\<^sub>r tl_ancestors"
using s1 tl_ancestors step(4)
by(auto simp add: get_ancestors_di_def[of child] elim!: bind_returns_result_E2
split: option.splits dest: returns_result_eq)
show ?case
by (metis (no_types, lifting) Some.prems \<open>h \<turnstile> get_ancestors_di parent \<rightarrow>\<^sub>r tl_ancestors\<close>
assms(4) eq_iff node_ptr_casts_commute order_trans s1 select_result_I2 set_ConsD set_subset_Cons
step.hyps(1) step.prems(1) step.prems(2) tl_ancestors)
qed
qed
qed
lemma get_ancestors_di_also_parent:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors_di some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast child \<in> set ancestors"
and "h \<turnstile> get_parent child \<rightarrow>\<^sub>r Some parent"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "parent \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors: "h \<turnstile> get_ancestors_di (cast child) \<rightarrow>\<^sub>r child_ancestors"
by (meson assms(1) assms(4) get_ancestors_di_ok is_OK_returns_result_I known_ptrs
local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result
type_wf)
then have "parent \<in> set child_ancestors"
apply(simp add: get_ancestors_di_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)]
get_ancestors_di_ptr)
then show ?thesis
using assms child_ancestors get_ancestors_di_subset by blast
qed
lemma get_ancestors_di_also_host:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors_di some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast shadow_root \<in> set ancestors"
and "h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
shows "cast host \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors: "h \<turnstile> get_ancestors_di (cast shadow_root) \<rightarrow>\<^sub>r child_ancestors"
by (meson assms(1) assms(2) assms(3) get_ancestors_di_ok get_ancestors_di_ptrs_in_heap
is_OK_returns_result_E known_ptrs type_wf)
then have "cast host \<in> set child_ancestors"
apply(simp add: get_ancestors_di_def)
by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)]
get_ancestors_di_ptr)
then show ?thesis
using assms child_ancestors get_ancestors_di_subset by blast
qed
lemma get_ancestors_di_also_disconnected_document:
assumes "heap_is_wellformed h"
and "h \<turnstile> get_ancestors_di some_ptr \<rightarrow>\<^sub>r ancestors"
and "cast disc_node \<in> set ancestors"
and "h \<turnstile> get_disconnected_document disc_node \<rightarrow>\<^sub>r disconnected_document"
and type_wf: "type_wf h"
and known_ptrs: "known_ptrs h"
and "h \<turnstile> get_parent disc_node \<rightarrow>\<^sub>r None"
shows "cast disconnected_document \<in> set ancestors"
proof -
obtain child_ancestors where child_ancestors: "h \<turnstile> get_ancestors_di (cast disc_node) \<rightarrow>\<^sub>r child_ancestors"
by (meson assms(1) assms(2) assms(3) get_ancestors_di_ok get_ancestors_di_ptrs_in_heap
is_OK_returns_result_E known_ptrs type_wf)
then have "cast disconnected_document \<in> set child_ancestors"
apply(simp add: get_ancestors_di_def)
by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I split: option.splits
dest!: returns_result_eq[OF assms(4)] returns_result_eq[OF assms(7)]
get_ancestors_di_ptr)
then show ?thesis
using assms child_ancestors get_ancestors_di_subset by blast
qed
lemma disc_node_disc_doc_dual:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
assumes "h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes"
assumes "node_ptr \<in> set disc_nodes"
shows "h \<turnstile> get_disconnected_document node_ptr \<rightarrow>\<^sub>r disc_doc"
proof -
have "node_ptr |\<in>| node_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap)
then
have "\<not>(\<exists>parent \<in> fset (object_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)"
apply(auto)[1]
using child_parent_dual[OF assms(1)]
assms known_ptrs_known_ptr option.simps(3)
returns_result_eq returns_result_select_result
by (metis (no_types, lifting) get_child_nodes_ok)
then
have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
using heap_is_wellformed_children_disc_nodes
using \<open>node_ptr |\<in>| node_ptr_kinds h\<close> assms(1) by blast
(* then obtain some_owner_document where
"some_owner_document \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))" and
"node_ptr \<in> set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto *)
then have "disc_doc \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))" and
"node_ptr \<in> set |h \<turnstile> get_disconnected_nodes disc_doc|\<^sub>r"
using CD.get_disconnected_nodes_ptr_in_heap assms(5)
assms(6) by auto
have h5: "\<exists>!x. x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h))) \<and>
h \<turnstile> Heap_Error_Monad.bind (get_disconnected_nodes x)
(\<lambda>children. return (node_ptr \<in> set children)) \<rightarrow>\<^sub>r True"
apply(auto intro!: bind_pure_returns_result_I)[1]
- apply (smt CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure
+ apply (smt (verit) CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure
\<open>\<exists>document_ptr\<in>fset (document_ptr_kinds h). node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r\<close>
- assms(2) bind_pure_returns_result_I2 notin_fset return_returns_result select_result_I2)
+ assms(2) bind_pure_returns_result_I2 return_returns_result select_result_I2)
apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1]
using heap_is_wellformed_one_disc_parent assms(1)
by blast
let ?filter_M = "filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (node_ptr \<in> set disconnected_nodes)))
(sorted_list_of_set (fset (document_ptr_kinds h)))"
have "h \<turnstile> 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 \<turnstile> ?filter_M \<rightarrow>\<^sub>r candidates"
by auto
have "candidates = [disc_doc]"
apply(rule filter_M_ex1[OF candidates \<open>disc_doc \<in>
set (sorted_list_of_set (fset (document_ptr_kinds h)))\<close> h5])
using \<open>node_ptr \<in> set |h \<turnstile> get_disconnected_nodes disc_doc|\<^sub>r\<close>
\<open>disc_doc \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))\<close>
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 \<open>node_ptr |\<in>| node_ptr_kinds h\<close> candidates
by(auto simp add: bind_pure_I get_disconnected_document_def
elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I)
qed
lemma get_ancestors_di_obtains_children_or_shadow_root_or_disconnected_node:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
and "h \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>r ancestors"
and "ancestor \<noteq> ptr"
and "ancestor \<in> set ancestors"
shows "((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast ancestor_child \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)
\<or> ((\<forall>ancestor_element shadow_root. ancestor = cast ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow> cast shadow_root \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis)
\<or> ((\<forall>disc_doc disc_nodes disc_node. ancestor = cast disc_doc \<longrightarrow>
h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes \<longrightarrow> disc_node \<in> set disc_nodes \<longrightarrow>
cast disc_node \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)"
proof (insert assms(4) assms(5) assms(6), induct ptr arbitrary: ancestors
rule: heap_wellformed_induct_rev_si[OF assms(1)])
case (1 child)
then show ?case
proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child")
case None
then obtain shadow_root where shadow_root: "child = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root"
using 1(4) 1(5) 1(6)
by(auto simp add: get_ancestors_di_def[of child] elim!: bind_returns_result_E2
split: option.splits)
then obtain host where host: "h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host"
by (metis "1.prems"(1) assms(1) assms(2) assms(3) document_ptr_kinds_commutes
get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_di_ptr local.get_host_ok
shadow_root_ptr_kinds_commutes)
then obtain host_ancestors where host_ancestors:
"h \<turnstile> get_ancestors_di (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host) \<rightarrow>\<^sub>r host_ancestors"
by (metis "1.prems"(1) assms(1) assms(2) assms(3) get_ancestors_di_also_host get_ancestors_di_ok
get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_di_ptr shadow_root)
then have "ancestors = cast shadow_root # host_ancestors"
using 1(4) 1(5) 1(3) None shadow_root host
by(auto simp add: get_ancestors_di_def[of child, simplified shadow_root]
elim!: bind_returns_result_E2 dest!: returns_result_eq[OF host] split: option.splits)
then show ?thesis
proof (cases "ancestor = cast host")
case True
then show ?thesis
using "1.prems"(1) host local.get_ancestors_di_ptr local.shadow_root_host_dual shadow_root
by blast
next
case False
have "ancestor \<in> set ancestors"
using host host_ancestors 1(3) get_ancestors_di_also_host assms(1) assms(2) assms(3)
using "1.prems"(3) by blast
then have "((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set host_ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis) \<or>
((\<forall>ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow>
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 \<in> set host_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)
\<or> ((\<forall>disc_doc disc_nodes disc_node. ancestor = cast disc_doc \<longrightarrow>
h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes \<longrightarrow> disc_node \<in> set disc_nodes \<longrightarrow>
cast disc_node \<in> set host_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)"
using "1.hyps"(2) "1.prems"(2) False \<open>ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\<close>
host host_ancestors shadow_root
by auto
then show ?thesis
using \<open>ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\<close>
by auto
qed
next
case (Some child_node)
then obtain parent_opt where parent_opt: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r parent_opt"
by (metis (no_types, lifting) "1.prems"(1) assms(1) assms(2) assms(3)
get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_di_ptr
local.get_parent_ok node_ptr_casts_commute node_ptr_kinds_commutes)
then show ?thesis
proof (induct parent_opt)
case None
then obtain disc_doc where disc_doc: "h \<turnstile> get_disconnected_document child_node \<rightarrow>\<^sub>r disc_doc"
by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_disconnected_document_ok)
then obtain parent_ancestors where parent_ancestors: "h \<turnstile> get_ancestors_di (cast disc_doc) \<rightarrow>\<^sub>r parent_ancestors"
by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E
l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_di_ok l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms
local.get_disconnected_document_disconnected_document_in_heap)
then have "ancestors = cast child_node # parent_ancestors"
using 1(4) 1(5) 1(6) Some \<open>cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\<close>
apply(auto simp add: get_ancestors_di_def[of child,
simplified \<open>cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\<close>] intro!: bind_pure_I
elim!: bind_returns_result_E2 dest!: returns_result_eq[OF disc_doc] split: option.splits)[1]
apply (simp add: returns_result_eq)
by (meson None.prems option.distinct(1) returns_result_eq)
then show ?thesis
proof (cases "ancestor = cast disc_doc")
case True
then show ?thesis
by (metis \<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close> disc_doc
list.set_intros(1) local.disc_doc_disc_node_dual)
next
case False
have "ancestor \<in> set ancestors"
by (simp add: "1.prems"(3))
then have "((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis) \<or>
((\<forall>ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow>
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 \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)
\<or> ((\<forall>disc_doc disc_nodes disc_node. ancestor = cast disc_doc \<longrightarrow>
h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes \<longrightarrow> disc_node \<in> set disc_nodes \<longrightarrow>
cast disc_node \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)"
using "1.hyps"(3) "1.prems"(2) False Some \<open>cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\<close>
\<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close> disc_doc parent_ancestors
by auto
then show ?thesis
using \<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close> by auto
qed
next
case (Some option)
then obtain parent where parent: "h \<turnstile> get_parent child_node \<rightarrow>\<^sub>r Some parent"
using 1(4) 1(5) 1(6)
by(auto simp add: get_ancestors_di_def[of child] intro!: bind_pure_I
elim!: bind_returns_result_E2 split: option.splits)
then obtain parent_ancestors where parent_ancestors:
"h \<turnstile> get_ancestors_di parent \<rightarrow>\<^sub>r parent_ancestors"
by (meson assms(1) assms(2) assms(3) get_ancestors_di_ok is_OK_returns_result_E
local.get_parent_parent_in_heap)
then have "ancestors = cast child_node # parent_ancestors"
using 1(4) 1(5) 1(6) Some \<open>cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\<close>
by(auto simp add: get_ancestors_di_def[of child, simplified
\<open>cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\<close>] dest!: elim!: bind_returns_result_E2
dest!: returns_result_eq[OF parent] split: option.splits)
then show ?thesis
proof (cases "ancestor = parent")
case True
then show ?thesis
by (metis \<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close>
list.set_intros(1) local.get_parent_child_dual parent)
next
case False
have "ancestor \<in> set ancestors"
by (simp add: "1.prems"(3))
then have "((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ancestor \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis) \<or>
((\<forall>ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow> 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 \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)
\<or> ((\<forall>disc_doc disc_nodes disc_node. ancestor = cast disc_doc \<longrightarrow>
h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes \<longrightarrow> disc_node \<in> set disc_nodes \<longrightarrow>
cast disc_node \<in> set parent_ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)"
using "1.hyps"(1) "1.prems"(2) False Some \<open>cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\<close>
\<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close> parent parent_ancestors
by auto
then show ?thesis
using \<open>ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\<close>
by auto
qed
qed
qed
qed
lemma get_ancestors_di_parent_child_a_host_shadow_root_rel:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
assumes "h \<turnstile> get_ancestors_di child \<rightarrow>\<^sub>r ancestors"
shows "(ptr, child) \<in> (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>* \<longleftrightarrow> ptr \<in> set ancestors"
proof
assume "(ptr, child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>* "
then show "ptr \<in> set ancestors"
proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
using assms(4) get_ancestors_di_ptr by blast
next
case False
obtain ptr_child where
ptr_child: "(ptr, ptr_child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h) \<and>
(ptr_child, child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using converse_rtranclE[OF 1(4)] \<open>ptr \<noteq> child\<close>
by metis
then show ?thesis
proof(cases "(ptr, ptr_child) \<in> 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 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children" and
ptr_child_node: "ptr_child_node \<in> set children"
proof -
assume a1: "\<And>children. \<lbrakk>h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children; ptr_child_node \<in> set children\<rbrakk>
\<Longrightarrow> thesis"
have "ptr |\<in>| object_ptr_kinds h"
using CD.parent_child_rel_parent_in_heap True by blast
moreover have "ptr_child_node \<in> set |h \<turnstile> 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 \<open>type_wf h\<close> \<open>known_ptrs h\<close>
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) \<in>
(parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using ptr_child True ptr_child_ptr_child_node by auto
ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \<in> set ancestors"
using 1 by auto
moreover have "h \<turnstile> get_parent ptr_child_node \<rightarrow>\<^sub>r Some ptr"
using assms(1) children ptr_child_node child_parent_dual
using \<open>known_ptrs h\<close> \<open>type_wf h\<close> by blast
ultimately show ?thesis
using get_ancestors_di_also_parent assms \<open>type_wf h\<close> by blast
next
case False
then show ?thesis
proof (cases "(ptr, ptr_child) \<in> a_host_shadow_root_rel h")
case True
then
obtain host where host: "ptr = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host"
using ptr_child
by(auto simp add: a_host_shadow_root_rel_def)
then obtain shadow_root where shadow_root: "h \<turnstile> get_shadow_root host \<rightarrow>\<^sub>r Some shadow_root"
and ptr_child_shadow_root: "ptr_child = cast shadow_root"
using False True
apply(auto simp add: a_host_shadow_root_rel_def)[1]
by (metis (no_types, lifting) assms(3) local.get_shadow_root_ok select_result_I)
moreover have "(cast shadow_root, child) \<in>
(parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using ptr_child ptr_child_shadow_root by blast
ultimately have "cast shadow_root \<in> set ancestors"
using "1.hyps"(2) host by blast
moreover have "h \<turnstile> get_host shadow_root \<rightarrow>\<^sub>r host"
by (metis assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_host_ok
local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host
shadow_root)
ultimately show ?thesis
using get_ancestors_di_also_host assms(1) assms(2) assms(3) assms(4) host
by blast
next
case False
then
obtain disc_doc where disc_doc: "ptr = cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_doc"
using ptr_child \<open>(ptr, ptr_child) \<notin> parent_child_rel h\<close>
by(auto simp add: a_ptr_disconnected_node_rel_def)
then obtain disc_node disc_nodes where disc_nodes:
"h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes"
and disc_node: "disc_node \<in> set disc_nodes"
and ptr_child_disc_node: "ptr_child = cast disc_node"
using False \<open>(ptr, ptr_child) \<notin> parent_child_rel h\<close> ptr_child
apply(auto simp add: a_ptr_disconnected_node_rel_def)[1]
by (metis (no_types, lifting) CD.get_disconnected_nodes_ok assms(3)
returns_result_select_result)
moreover have "(cast disc_node, child) \<in>
(parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using ptr_child ptr_child_disc_node by blast
ultimately have "cast disc_node \<in> set ancestors"
using "1.hyps"(3) disc_doc by blast
moreover have "h \<turnstile> get_parent disc_node \<rightarrow>\<^sub>r None"
using \<open>(ptr, ptr_child) \<notin> parent_child_rel h\<close>
apply(auto simp add: parent_child_rel_def ptr_child_disc_node)[1]
by (smt assms(1) assms(2) assms(3) assms(4) calculation disc_node disc_nodes
get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.CD.a_heap_is_wellformed_def
local.CD.distinct_lists_no_parent local.CD.heap_is_wellformed_impl local.get_parent_child_dual
local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_def
node_ptr_kinds_commutes select_result_I2 split_option_ex)
then
have "h \<turnstile> get_disconnected_document disc_node \<rightarrow>\<^sub>r disc_doc"
using disc_node_disc_doc_dual disc_nodes disc_node assms(1) assms(2) assms(3)
by blast
ultimately show ?thesis
using \<open>h \<turnstile> get_parent disc_node \<rightarrow>\<^sub>r None\<close> assms(1) assms(2) assms(3) assms(4)
disc_doc get_ancestors_di_also_disconnected_document
by blast
qed
qed
qed
qed
next
assume "ptr \<in> set ancestors"
then show "(ptr, child) \<in> (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)])
case (1 ptr)
then show ?case
proof (cases "ptr = child")
case True
then show ?thesis
by simp
next
case False
have 2: "\<And>thesis. ((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast ancestor_child \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)
\<or> ((\<forall>ancestor_element shadow_root. ptr = cast ancestor_element \<longrightarrow>
h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root \<longrightarrow> cast shadow_root \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow>
thesis)
\<or> ((\<forall>disc_doc disc_nodes disc_node. ptr = cast disc_doc \<longrightarrow>
h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes \<longrightarrow> disc_node \<in> set disc_nodes \<longrightarrow>
cast disc_node \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)"
using "1.prems" False assms(1) assms(2) assms(3) assms(4)
get_ancestors_di_obtains_children_or_shadow_root_or_disconnected_node by blast
then show ?thesis
proof (cases "\<forall>thesis. ((\<forall>children ancestor_child. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<longrightarrow>
ancestor_child \<in> set children \<longrightarrow> cast ancestor_child \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)")
case True
then obtain children ancestor_child
where "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and "ancestor_child \<in> set children"
and "cast ancestor_child \<in> set ancestors"
by blast
then show ?thesis
by (meson "1.hyps"(1) in_rtrancl_UnI local.CD.parent_child_rel_child r_into_rtrancl
rtrancl_trans)
next
case False
note f1 = this
then show ?thesis
proof (cases "\<forall>thesis. ((\<forall>disc_doc disc_nodes disc_node. ptr = cast disc_doc \<longrightarrow>
h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes \<longrightarrow> disc_node \<in> set disc_nodes \<longrightarrow>
cast disc_node \<in> set ancestors \<longrightarrow> thesis) \<longrightarrow> thesis)")
case True
then obtain disc_doc disc_nodes disc_node
where "ptr = cast disc_doc"
and "h \<turnstile> get_disconnected_nodes disc_doc \<rightarrow>\<^sub>r disc_nodes"
and "disc_node \<in> set disc_nodes"
and "cast disc_node \<in> set ancestors"
by blast
then show ?thesis
by (meson "1.hyps"(3) in_rtrancl_UnI
local.a_ptr_disconnected_node_rel_disconnected_node r_into_rtrancl rtrancl_trans)
next
case False
then
obtain ancestor_element shadow_root
where "ptr = cast ancestor_element"
and "h \<turnstile> get_shadow_root ancestor_element \<rightarrow>\<^sub>r Some shadow_root"
and "cast shadow_root \<in> set ancestors"
using f1 2 by smt
then show ?thesis
by (meson "1.hyps"(2) in_rtrancl_UnI local.a_host_shadow_root_rel_shadow_root
r_into_rtrancl rtrancl_trans)
qed
qed
qed
qed
qed
end
interpretation i_get_ancestors_di_wf?: l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_ancestors_di
get_ancestors_di_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root
get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel
heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
by(auto simp add: l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
subsubsection \<open>get\_owner\_document\<close>
locale l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_disconnected_nodes +
l_get_child_nodes +
l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf +
l_known_ptrs +
l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr"
begin
lemma get_owner_document_disconnected_nodes:
assumes "heap_is_wellformed h"
assumes "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assumes "node_ptr \<in> set disc_nodes"
assumes known_ptrs: "known_ptrs h"
assumes type_wf: "type_wf h"
shows "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r document_ptr"
proof -
have 2: "node_ptr |\<in>| 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 |\<in>| document_ptr_kinds h"
using assms(2) get_disconnected_nodes_ptr_in_heap by blast
then have 4: "\<not>(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> 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 "(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
using assms(1) 2 "3" assms(2) assms(3) by auto
ultimately have 0: "\<exists>!document_ptr\<in>set |h \<turnstile> document_ptr_kinds_M|\<^sub>r.
node_ptr \<in> set |h \<turnstile> 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 \<turnstile> get_parent node_ptr \<rightarrow>\<^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 \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r cast node_ptr"
using get_root_node_no_parent
by simp
obtain document_ptrs where document_ptrs: "h \<turnstile> document_ptr_kinds_M \<rightarrow>\<^sub>r document_ptrs"
by simp
then have "h \<turnstile> ok (filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> 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 \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^sub>r candidates"
by auto
have filter: "filter (\<lambda>document_ptr. |h \<turnstile> do {
disconnected_nodes \<leftarrow> 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 \<in> cast ` set disconnected_nodes)
}|\<^sub>r) document_ptrs = [document_ptr]"
apply(rule filter_ex1)
using 0 document_ptrs apply(simp)[1]
apply (smt "0" "3" assms bind_is_OK_pure_I bind_pure_returns_result_I bind_pure_returns_result_I2
bind_returns_result_E2 bind_returns_result_E3 document_ptr_kinds_M_def get_disconnected_nodes_ok
get_disconnected_nodes_pure image_eqI is_OK_returns_result_E l_ptr_kinds_M.ptr_kinds_ptr_kinds_M return_ok
return_returns_result returns_result_eq select_result_E select_result_I select_result_I2 select_result_I2)
using assms(2) assms(3)
apply (smt bind_is_OK_I2 bind_returns_result_E3 get_disconnected_nodes_pure image_eqI
is_OK_returns_result_I return_ok return_returns_result select_result_E)
using document_ptrs 3 apply(simp)
using document_ptrs
by simp
have "h \<turnstile> filter_M (\<lambda>document_ptr. do {
disconnected_nodes \<leftarrow> 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)) \<in> cast ` set disconnected_nodes)
}) document_ptrs \<rightarrow>\<^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 \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^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 \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r None"
assumes "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r owner_document"
assumes "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
assumes "known_ptrs h"
assumes "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
have "\<And>parent. parent |\<in>| object_ptr_kinds h \<Longrightarrow> node_ptr \<notin> set |h \<turnstile> 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
+ by (smt (verit) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
is_OK_returns_result_I local.get_disconnected_nodes_ok local.get_owner_document_disconnected_nodes
local.get_parent_ptr_in_heap local.heap_is_wellformed_children_disc_nodes returns_result_select_result
select_result_I2)
qed
lemma get_owner_document_owner_document_in_heap_node:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
shows "owner_document |\<in>| document_ptr_kinds h"
proof -
obtain root where
root: "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r root"
using assms(4)
by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
split: option.splits)
then show ?thesis
proof (cases "is_document_ptr root")
case True
then show ?thesis
using assms(4) root
apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply(drule(1) returns_result_eq) apply(auto)[1]
using assms document_ptr_kinds_commutes get_root_node_root_in_heap
by blast
next
case False
have "known_ptr root"
using assms local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast
have "root |\<in>| object_ptr_kinds h"
using root
using assms local.get_root_node_root_in_heap
by blast
show ?thesis
proof (cases "is_shadow_root_ptr root")
case True
then show ?thesis
using assms(4) root
apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
apply(drule(1) returns_result_eq) apply(auto)[1]
using assms document_ptr_kinds_commutes get_root_node_root_in_heap
by blast
next
case False
then have "is_node_ptr_kind root"
using \<open>\<not> is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\<close> \<open>known_ptr root\<close> \<open>root |\<in>| object_ptr_kinds h\<close>
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 "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
using local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent
local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3
- node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq
+ node_ptr_inclusion node_ptr_kinds_commutes option.distinct(1) returns_result_eq
returns_result_select_result root
- by (metis (no_types, lifting) assms \<open>root |\<in>| object_ptr_kinds h\<close>)
+ by (metis (no_types, opaque_lifting) assms \<open>root |\<in>| object_ptr_kinds h\<close>)
then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
by auto
then
obtain candidates where
candidates: "h \<turnstile> filter_M
(\<lambda>document_ptr.
Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
(\<lambda>disconnected_nodes. return (root \<in> 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)))
\<rightarrow>\<^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
+ is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure
return_ok return_pure sorted_list_of_set(1))
then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto simp add: assms local.get_disconnected_nodes_ok
intro!: bind_pure_I bind_pure_returns_result_I)[1]
done
then have "candidates \<noteq> []"
by auto
then have "owner_document \<in> 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 \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq)
then show ?thesis
using candidates
by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I
local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure)
qed
qed
qed
lemma get_owner_document_owner_document_in_heap:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
shows "owner_document |\<in>| 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 \<turnstile> invoke [] ptr () \<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| document_ptr_kinds h"
by (meson invoke_empty is_OK_returns_result_I)
next
assume "h \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (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 \<circ> the \<circ> 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 ())
\<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| 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: "\<not> 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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^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: "\<not> 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: "\<not> 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: "\<not> 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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (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 \<circ> the \<circ> 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 ()) \<rightarrow>\<^sub>r owner_document"
then show "owner_document |\<in>| document_ptr_kinds h"
by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
elim!: bind_returns_result_E2 split: if_splits)
qed
lemma get_owner_document_ok:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
assumes "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> ok (get_owner_document ptr)"
proof -
have "known_ptr ptr"
using assms(2) assms(4) local.known_ptrs_known_ptr
by blast
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(auto simp add: known_ptr_impl)[1]
using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr
known_ptr_not_shadow_root_ptr known_ptr_not_element_ptr apply blast
using assms(4)
apply(auto simp add: get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I
split: option.splits)[1]
using assms(4)
apply(auto simp add: get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I
filter_M_pure_I bind_pure_I filter_M_is_OK_I split: option.splits)[1]
using assms(4)
apply(auto simp add: assms(1) assms(2) assms(3) local.get_ancestors_ok get_disconnected_nodes_ok
get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
filter_M_is_OK_I split: option.splits)[1]
using assms(4)
apply(auto simp add: assms(1) assms(2) assms(3) local.get_ancestors_ok get_disconnected_nodes_ok
get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
filter_M_is_OK_I split: option.splits)[1]
done
qed
lemma get_owner_document_child_same:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
assumes "child \<in> set children"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document"
proof -
have "ptr |\<in>| object_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_child_nodes_ptr_in_heap)
then have "known_ptr ptr"
using assms(2) local.known_ptrs_known_ptr by blast
have "cast child |\<in>| object_ptr_kinds h"
using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes
by blast
then
have "known_ptr (cast child)"
using assms(2) local.known_ptrs_known_ptr by blast
then have "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast child) \<or> is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast child)"
by(auto simp add: known_ptr_impl NodeClass.a_known_ptr_def ElementClass.a_known_ptr_def
CharacterDataClass.a_known_ptr_def DocumentClass.a_known_ptr_def a_known_ptr_def
split: option.splits)
obtain root where root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.get_root_node_ok)
then have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root"
using assms(1) assms(2) assms(3) assms(4) assms(5) local.child_parent_dual local.get_root_node_parent_same
by blast
have "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr ptr")
case True
then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = ptr"
using case_optionE document_ptr_casts_commute by blast
then have "root = cast document_ptr"
using root
by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
split: option.splits)
then have "h \<turnstile> CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using document_ptr \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close>
document_ptr]
apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF
\<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> document_ptr], rotated]
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1]
using \<open>ptr |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes by blast
then show ?thesis
using \<open>known_ptr ptr\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_document_ptr)
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>ptr |\<in>| object_ptr_kinds h\<close> True
by(auto simp add: document_ptr[symmetric] intro!: bind_pure_returns_result_I
split: option.splits)
next
case False
then show ?thesis
proof (cases "is_shadow_root_ptr ptr")
case True
then obtain shadow_root_ptr where shadow_root_ptr: "cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = ptr"
using case_optionE shadow_root_ptr_casts_commute
by (metis (no_types, lifting) document_ptr_casts_commute3 is_document_ptr_kind_none option.case_eq_if)
then have "root = cast shadow_root_ptr"
using root
by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
split: option.splits)
then have "h \<turnstile> 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 () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using shadow_root_ptr \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast shadow_root_ptr\<close>
shadow_root_ptr]
apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified
\<open>root = cast shadow_root_ptr\<close> shadow_root_ptr], rotated] intro!: bind_pure_returns_result_I
filter_M_pure_I bind_pure_I split: if_splits option.splits)[1]
using \<open>ptr |\<in>| object_ptr_kinds h\<close> shadow_root_ptr_kinds_commutes document_ptr_kinds_commutes
by blast
then show ?thesis
using \<open>known_ptr ptr\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_document_ptr)
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>ptr |\<in>| object_ptr_kinds h\<close> True
using False
by(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def shadow_root_ptr[symmetric]
intro!: bind_pure_returns_result_I split: option.splits)
next
case False
then obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = ptr"
using \<open>known_ptr ptr\<close> \<open>\<not> is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then have "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using root \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>
unfolding CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure)
then show ?thesis
using \<open>known_ptr ptr\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False \<open>\<not> is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\<close>
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False \<open>\<not> is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\<close>
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False \<open>\<not> is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\<close>
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False \<open>\<not> is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\<close>
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> False \<open>\<not> is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\<close>
apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
by(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I dest!: is_OK_returns_result_I)
qed
qed
then show ?thesis
using \<open>is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<or> is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)\<close>
using \<open>cast child |\<in>| object_ptr_kinds h\<close>
by(auto simp add: get_owner_document_def[of "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child"]
a_get_owner_document_tups_def CD.a_get_owner_document_tups_def split: invoke_splits)
qed
lemma get_owner_document_rel:
assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
assumes "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
assumes "ptr \<noteq> cast owner_document"
shows "(cast owner_document, ptr) \<in> (parent_child_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
proof -
have "ptr |\<in>| object_ptr_kinds h"
using assms
by (meson is_OK_returns_result_I local.get_owner_document_ptr_in_heap)
then
have "known_ptr ptr"
using known_ptrs_known_ptr[OF assms(2)] by simp
have "is_node_ptr_kind ptr"
proof (rule ccontr)
assume "\<not> is_node_ptr_kind ptr"
then
show False
using assms(4) \<open>known_ptr ptr\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply(drule(1) known_ptr_not_shadow_root_ptr)
apply(drule(1) known_ptr_not_document_ptr)
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(5)
by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
elim!: bind_returns_result_E2 split: if_splits option.splits)
qed
then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
using assms(4) \<open>known_ptr ptr\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply(drule(1) known_ptr_not_shadow_root_ptr)
apply(drule(1) known_ptr_not_document_ptr)
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>ptr |\<in>| object_ptr_kinds h\<close>
by (auto simp add: is_document_ptr_kind_none)
then obtain root where root: "h \<turnstile> get_root_node (cast node_ptr) \<rightarrow>\<^sub>r root"
by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2)
then have "root |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap by blast
then
have "known_ptr root"
using \<open>known_ptrs h\<close> local.known_ptrs_known_ptr by blast
have "(root, cast node_ptr) \<in> (parent_child_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using root
by (simp add: assms(1) assms(2) assms(3) in_rtrancl_UnI local.get_root_node_parent_child_rel)
show ?thesis
proof (cases "is_document_ptr_kind root")
case True
then have "root = cast owner_document"
using \<open>h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> root
by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def
dest!: bind_returns_result_E3[rotated, OF root, rotated] split: option.splits)
then have "(root, cast node_ptr) \<in> (parent_child_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using assms(1) assms(2) assms(3) in_rtrancl_UnI local.get_root_node_parent_child_rel root
by blast
then show ?thesis
using \<open>root = cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\<close> node_ptr by blast
next
case False
then obtain root_node where root_node: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node"
using assms(2) \<open>root |\<in>| object_ptr_kinds h\<close>
by(auto simp add: known_ptr_impl ShadowRootClass.known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
dest!: known_ptrs_known_ptr split: option.splits)
have "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> root False
apply(auto simp add: root_node CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF root, rotated] split: option.splits
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I)[1]
by (simp add: assms(1) assms(2) assms(3) local.get_root_node_no_parent local.get_root_node_same_no_parent)
then
have "h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
using \<open>known_ptr root\<close>
apply(auto simp add: get_owner_document_def CD.a_get_owner_document_tups_def
a_get_owner_document_tups_def known_ptr_impl)[1]
apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_document_ptr)
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
using \<open>h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> root
False \<open>root |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> root
False \<open>root |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> root
False \<open>root |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: root_node intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> root
False \<open>root |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: root_node intro!: bind_pure_returns_result_I split: option.splits)[1]
done
have "\<not> (\<exists>parent\<in>fset (object_ptr_kinds h). root_node \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)"
using root_node
- by (metis (no_types, lifting) assms(1) assms(2) assms(3) local.child_parent_dual
+ by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3) local.child_parent_dual
local.get_child_nodes_ok local.get_root_node_same_no_parent local.known_ptrs_known_ptr
- notin_fset option.distinct(1) returns_result_eq returns_result_select_result root)
+ option.distinct(1) returns_result_eq returns_result_select_result root)
have "root_node |\<in>| node_ptr_kinds h"
using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap node_ptr_kinds_commutes root root_node
by blast
then have "\<exists>document_ptr\<in>fset (document_ptr_kinds h). root_node \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using \<open>\<not> (\<exists>parent\<in>fset (object_ptr_kinds h). root_node \<in> set |h \<turnstile> get_child_nodes parent|\<^sub>r)\<close> assms(1)
local.heap_is_wellformed_children_disc_nodes by blast
then obtain disc_nodes document_ptr where "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
and "root_node \<in> set disc_nodes"
- by (meson assms(3) local.get_disconnected_nodes_ok notin_fset returns_result_select_result)
+ by (meson assms(3) local.get_disconnected_nodes_ok returns_result_select_result)
then have "document_ptr |\<in>| document_ptr_kinds h"
by (meson is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
then have "document_ptr = owner_document"
by (metis \<open>h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes\<close>
\<open>h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document\<close> \<open>root_node \<in> set disc_nodes\<close> assms(1) assms(2)
assms(3) local.get_owner_document_disconnected_nodes returns_result_eq root_node)
then have "(cast owner_document, cast root_node) \<in> a_ptr_disconnected_node_rel h"
apply(auto simp add: a_ptr_disconnected_node_rel_def)[1]
using \<open>h \<turnstile> local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> assms(1)
assms(2) assms(3) get_owner_document_owner_document_in_heap_node
by (metis (no_types, lifting) \<open>h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes\<close>
\<open>root_node \<in> set disc_nodes\<close> case_prodI mem_Collect_eq pair_imageI select_result_I2)
moreover have "(cast root_node, cast node_ptr) \<in>
(parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
by (metis assms(1) assms(2) assms(3) in_rtrancl_UnI local.get_root_node_parent_child_rel
root root_node)
ultimately show ?thesis
by (metis (no_types, lifting) assms(1) assms(2) assms(3) in_rtrancl_UnI
local.get_root_node_parent_child_rel node_ptr r_into_rtrancl root root_node rtrancl_trans)
qed
qed
end
interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf get_disconnected_nodes get_disconnected_nodes_locs known_ptr get_child_nodes
get_child_nodes_locs DocumentClass.known_ptr get_parent get_parent_locs DocumentClass.type_wf
get_root_node get_root_node_locs CD.a_get_owner_document get_host get_host_locs get_owner_document
get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed
parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs
known_ptrs get_ancestors get_ancestors_locs
by(auto simp add: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)
declare l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]:
"l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes
get_owner_document get_parent get_child_nodes"
apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def instances)[1]
using get_owner_document_disconnected_nodes apply fast
using in_disconnected_nodes_no_parent apply fast
using get_owner_document_owner_document_in_heap apply fast
using get_owner_document_ok apply fast
using get_owner_document_child_same apply (fast, fast)
done
paragraph \<open>get\_owner\_document\<close>
locale l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node_wf +
l_get_owner_document_wf +
assumes known_ptr_impl: "known_ptr = a_known_ptr"
begin
lemma get_root_node_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
assumes "is_document_ptr_kind root"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r the (cast root)"
proof -
have "ptr |\<in>| object_ptr_kinds h"
using assms(4)
by (meson is_OK_returns_result_I local.get_root_node_ptr_in_heap)
then have "known_ptr ptr"
using assms(3) local.known_ptrs_known_ptr by blast
{
assume "is_document_ptr_kind ptr"
then have "ptr = root"
using assms(4)
by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
split: option.splits)
then have ?thesis
using \<open>is_document_ptr_kind ptr\<close> \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
intro!: bind_pure_returns_result_I split: option.splits)
}
moreover
{
assume "is_node_ptr_kind ptr"
then have ?thesis
using \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs)
apply(auto split: option.splits)[1]
using \<open>h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root\<close> assms(5)
by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def
intro!: bind_pure_returns_result_I split: option.splits)
}
ultimately
show ?thesis
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
qed
lemma get_root_node_same_owner_document:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
shows "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
proof -
have "ptr |\<in>| object_ptr_kinds h"
by (meson assms(4) is_OK_returns_result_I local.get_root_node_ptr_in_heap)
have "root |\<in>| object_ptr_kinds h"
using assms(1) assms(2) assms(3) assms(4) local.get_root_node_root_in_heap by blast
have "known_ptr ptr"
using \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(3) local.known_ptrs_known_ptr by blast
have "known_ptr root"
using \<open>root |\<in>| object_ptr_kinds h\<close> assms(3) local.known_ptrs_known_ptr by blast
show ?thesis
proof (cases "is_document_ptr_kind ptr")
case True
then
have "ptr = root"
using assms(4)
apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
by (metis document_ptr_casts_commute3 last_ConsL local.get_ancestors_not_node node_ptr_no_document_ptr_cast)
then show ?thesis
by auto
next
case False
then have "is_node_ptr_kind ptr"
using \<open>known_ptr ptr\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
by (metis node_ptr_casts_commute3)
show ?thesis
proof
assume "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
then have "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
using node_ptr
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
by(auto elim!: bind_returns_result_E2 split: option.splits)
show "h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr_kind root")
case True
then show ?thesis
proof (cases "is_shadow_root_ptr root")
case True
then
have "is_shadow_root_ptr root"
using True \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
have "root = cast owner_document"
using \<open>is_document_ptr_kind root\<close>
by (smt \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3) assms(4)
document_ptr_casts_commute3 get_root_node_document returns_result_eq)
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_shadow_root_ptr root\<close> apply blast
using \<open>root |\<in>| object_ptr_kinds h\<close>
apply(simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none)
apply (metis \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3)
case_optionE document_ptr_kinds_def is_shadow_root_ptr_kind_none l_get_owner_document_wf.get_owner_document_owner_document_in_heap local.l_get_owner_document_wf_axioms not_None_eq return_bind shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes shadow_root_ptr_kinds_def)
using \<open>root |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1]
using \<open>root |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1]
using \<open>root |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1]
done
next
case False
then
have "is_document_ptr root"
using True \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
have "root = cast owner_document"
using True
by (smt \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3) assms(4)
document_ptr_casts_commute3 get_root_node_document returns_result_eq)
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\<close> apply blast
using \<open>root |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none)[1]
apply (metis \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3)
case_optionE document_ptr_kinds_def is_shadow_root_ptr_kind_none
l_get_owner_document_wf.get_owner_document_owner_document_in_heap local.l_get_owner_document_wf_axioms
not_None_eq return_bind shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes shadow_root_ptr_kinds_def)
using \<open>root |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1]
using \<open>root |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1]
using \<open>root |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes
apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1]
done
qed
next
case False
then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)
then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> assms(4)
apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
apply (metis assms(1) assms(2) assms(3) local.get_root_node_no_parent
local.get_root_node_same_no_parent node_ptr returns_result_eq)
using \<open>is_node_ptr_kind root\<close> node_ptr returns_result_eq by fastforce
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_node_ptr_kind root\<close> \<open>known_ptr root\<close>
apply(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)[1]
using \<open>is_node_ptr_kind root\<close> \<open>known_ptr root\<close>
apply(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)[1]
using \<open>is_node_ptr_kind root\<close> \<open>known_ptr root\<close>
apply(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
split: option.splits)[1]
using \<open>root |\<in>| object_ptr_kinds h\<close>
by(auto simp add: root_node_ptr)
qed
next
assume "h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
show "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr_kind root")
case True
have "root = cast owner_document"
using \<open>h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
apply(auto simp add: True CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits option.splits)[1]
apply(auto simp add: True CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits option.splits)[1]
apply (metis True cast_document_ptr_not_node_ptr(2) is_document_ptr_kind_obtains
is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
by (metis True cast_document_ptr_not_node_ptr(1) document_ptr_casts_commute3
is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
then show ?thesis
using assms(1) assms(2) assms(3) assms(4) get_root_node_document
by fastforce
next
case False
then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr"
by (metis node_ptr_casts_commute3)
then have "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document\<close>
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I)
by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2)
then have "h \<turnstile> CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
using assms(1) assms(2) assms(3) assms(4) local.get_root_node_no_parent
local.get_root_node_same_no_parent node_ptr returns_result_eq root_node_ptr
by fastforce+
then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def
CD.a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+
using node_ptr \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
intro!: bind_pure_returns_result_I split: option.splits)
qed
qed
qed
qed
end
interpretation i_get_owner_document_wf_get_root_node_wf?: l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
DocumentClass.known_ptr get_parent get_parent_locs DocumentClass.type_wf get_disconnected_nodes
get_disconnected_nodes_locs get_root_node get_root_node_locs CD.a_get_owner_document
get_host get_host_locs get_owner_document get_child_nodes get_child_nodes_locs type_wf known_ptr
known_ptrs get_ancestors get_ancestors_locs heap_is_wellformed parent_child_rel
by(auto simp add: l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def
l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)
declare l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]:
"l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs get_root_node get_owner_document"
apply(auto simp add: l_get_owner_document_wf_get_root_node_wf_def l_get_owner_document_wf_get_root_node_wf_axioms_def instances)[1]
using get_root_node_document apply blast
using get_root_node_same_owner_document apply (blast, blast)
done
subsubsection \<open>remove\_child\<close>
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 \<turnstile> remove_child ptr child \<rightarrow>\<^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 \<turnstile> remove_child ptr child \<rightarrow>\<^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 \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'"
shows "heap_is_wellformed h'"
proof -
obtain owner_document children_h h2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \<rightarrow>\<^sub>r owner_document" and
children_h: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h2: "h \<turnstile> set_disconnected_nodes owner_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> set_child_nodes ptr (remove1 child children_h) \<rightarrow>\<^sub>h h'"
using assms(4)
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1]
using pure_returns_heap_eq by fastforce
have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'"
using CD.remove_child_heap_is_wellformed_preserved(3) assms
unfolding heap_is_wellformed_def
by auto
have "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document"
using owner_document children_h child_in_children_h
using local.get_owner_document_child_same assms by blast
have shadow_root_eq: "\<And>ptr' shadow_root_ptr_opt. h \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq: "\<And>ptr' tag. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> 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="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF remove_child_writes assms(4)])
unfolding remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
have document_ptr_kinds_eq: "document_ptr_kinds h = document_ptr_kinds h'"
using object_ptr_kinds_eq
by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
using object_ptr_kinds_eq
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
have element_ptr_kinds_eq: "element_ptr_kinds h = element_ptr_kinds h'"
using object_ptr_kinds_eq
by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def)
have "parent_child_rel h' \<subseteq> parent_child_rel h"
using \<open>heap_is_wellformed h\<close> heap_is_wellformed_def
using CD.remove_child_parent_child_rel_subset
using \<open>known_ptrs h\<close> \<open>type_wf h\<close> assms(4)
by simp
have "known_ptr ptr"
using assms(3)
using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF set_disconnected_nodes_writes h2]
using set_disconnected_nodes_types_preserved \<open>type_wf h\<close>
by(auto simp add: reflp_def transp_def)
have children_eq:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children =
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(4)])
unfolding remove_child_locs_def
using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
by fast
then have children_eq2:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h"
apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads
set_disconnected_nodes_writes h2 children_h] )
by (simp add: set_disconnected_nodes_get_child_nodes)
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r remove1 child children_h"
using assms(4) owner_document h2 disconnected_nodes_h children_h
apply(auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto split: if_splits)[1]
apply(simp)
apply(auto split: if_splits)[1]
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E3)
apply(auto)[1]
apply(simp)
apply(drule bind_returns_heap_E4)
apply(auto)[1]
apply simp
using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h'
by blast
have disconnected_nodes_eq: "\<And>ptr' disc_nodes. ptr' \<noteq> owner_document \<Longrightarrow>
h \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes = h2 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes"
using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h2
apply(rule reads_writes_preserved)
by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then
have disconnected_nodes_eq2: "\<And>ptr'. ptr' \<noteq> owner_document \<Longrightarrow>
|h \<turnstile> get_disconnected_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r"
by (meson select_result_eq)
have "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
using h2 local.set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h2:
"\<And>ptr' disc_nodes. h2 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes"
using local.get_disconnected_nodes_reads set_child_nodes_writes h'
apply(rule reads_writes_preserved)
using local.set_child_nodes_get_disconnected_nodes by blast
then
have disconnected_nodes_eq2_h2: "\<And>ptr'. |h2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r = |h' \<turnstile> get_disconnected_nodes ptr'|\<^sub>r"
by (meson select_result_eq)
have "a_host_shadow_root_rel h' = a_host_shadow_root_rel h"
by(auto simp add: a_host_shadow_root_rel_def shadow_root_eq2 element_ptr_kinds_eq)
moreover
have "(ptr, cast child) \<in> parent_child_rel h"
using child_in_children_h children_h local.CD.parent_child_rel_child by blast
moreover
have "a_ptr_disconnected_node_rel h' = insert (cast owner_document, cast child) (a_ptr_disconnected_node_rel h)"
using \<open>h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h\<close> disconnected_nodes_eq2 disconnected_nodes_h
apply(auto simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq2_h2[symmetric] document_ptr_kinds_eq[symmetric])[1]
apply(case_tac "aa = owner_document")
apply(auto)[1]
apply(auto)[1]
apply (metis (no_types, lifting) assms(4) case_prodI disconnected_nodes_eq_h2 h2
is_OK_returns_heap_I local.remove_child_in_disconnected_nodes
local.set_disconnected_nodes_ptr_in_heap mem_Collect_eq owner_document pair_imageI select_result_I2)
by (metis (no_types, lifting) case_prodI list.set_intros(2) mem_Collect_eq pair_imageI select_result_I2)
then
have "a_ptr_disconnected_node_rel h' = a_ptr_disconnected_node_rel h \<union> {(cast owner_document, cast child)}"
by auto
moreover have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using assms(1) local.heap_is_wellformed_def by blast
moreover have "parent_child_rel h' = parent_child_rel h - {(ptr, cast child)}"
apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq children_eq2)[1]
apply (metis (no_types, lifting) children_eq2 children_h children_h' notin_set_remove1 select_result_I2)
using \<open>h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h\<close>
\<open>heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\<close> disconnected_nodes_eq_h2 local.CD.distinct_lists_no_parent
local.CD.heap_is_wellformed_def apply auto[1]
by (metis (no_types, lifting) children_eq2 children_h children_h' in_set_remove1 select_result_I2)
moreover have "(cast owner_document, ptr) \<in> (parent_child_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> get_owner_document_rel
using assms(1) assms(2) assms(3) by blast
then have "(cast owner_document, ptr) \<in> (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
by (metis (no_types, lifting) in_rtrancl_UnI inf_sup_aci(5) inf_sup_aci(7))
ultimately
have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')"
by (smt Un_assoc Un_insert_left Un_insert_right acyclic_insert insert_Diff_single
insert_absorb2 mk_disjoint_insert prod.inject rtrancl_Un_separator_converseE rtrancl_trans
singletonD sup_bot.comm_neutral)
show ?thesis
using \<open>heap_is_wellformed h\<close>
using \<open>heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\<close>
using \<open>acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')\<close>
apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def
host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1]
by(auto simp add: object_ptr_kinds_eq element_ptr_kinds_eq shadow_root_ptr_kinds_eq
shadow_root_eq shadow_root_eq2 tag_name_eq tag_name_eq2)
qed
lemma remove_preserves_type_wf:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> remove child \<rightarrow>\<^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 \<turnstile> remove child \<rightarrow>\<^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 \<turnstile> remove child \<rightarrow>\<^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 \<Longrightarrow> h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
\<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> child \<notin> 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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children \<Longrightarrow>
h \<turnstile> remove_child ptr node_ptr \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r node_ptr # children \<Longrightarrow>
h \<turnstile> remove node_ptr \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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 \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow>
h \<turnstile> forall_M remove children \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
using CD.remove_for_all_empty_children local.heap_is_wellformed_def by blast
end
interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root
get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel
heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs
DocumentClass.known_ptr get_parent get_parent_locs DocumentClass.type_wf get_root_node get_root_node_locs
CD.a_get_owner_document get_owner_document known_ptrs get_ancestors get_ancestors_locs set_child_nodes
set_child_nodes_locs remove_child remove_child_locs remove
by(auto simp add: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma remove_child_wf2_is_l_remove_child_wf2 [instances]:
"l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove"
apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1]
using remove_child_preserves_type_wf apply fast
using remove_child_preserves_known_ptrs apply fast
using remove_child_heap_is_wellformed_preserved apply (fast)
using remove_preserves_type_wf apply fast
using remove_preserves_known_ptrs apply fast
using remove_heap_is_wellformed_preserved apply (fast)
using remove_child_removes_child apply fast
using remove_child_removes_first_child apply fast
using remove_removes_child apply fast
using remove_for_all_empty_children apply fast
done
subsubsection \<open>adopt\_node\<close>
locale l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
CD: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ _ _ _ adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma adopt_node_removes_first_child: "heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r node # children
\<Longrightarrow> h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
by (smt CD.adopt_node_removes_first_child bind_returns_heap_E error_returns_heap
l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.adopt_node_def local.CD.adopt_node_impl local.get_ancestors_di_pure
local.l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms pure_returns_heap_eq)
lemma adopt_node_document_in_heap: "heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h
\<Longrightarrow> h \<turnstile> ok (adopt_node owner_document node)
\<Longrightarrow> owner_document |\<in>| document_ptr_kinds h"
by (metis (no_types, lifting) bind_returns_heap_E document_ptr_kinds_commutes is_OK_returns_heap_E
is_OK_returns_result_I local.adopt_node_def local.get_ancestors_di_ptr_in_heap)
end
locale l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_child_nodes +
l_get_disconnected_nodes +
l_set_child_nodes_get_shadow_root +
l_set_disconnected_nodes_get_shadow_root +
l_set_child_nodes_get_tag_name +
l_set_disconnected_nodes_get_tag_name +
l_get_owner_document +
l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M+
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_root_node +
l_set_disconnected_nodes_get_child_nodes +
l_get_owner_document_wf +
l_remove_child_wf2 +
l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_disconnected_document +
l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma adopt_node_removes_child:
assumes wellformed: "heap_is_wellformed h"
and adopt_node: "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h2"
and children: "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<notin> set children"
proof -
obtain old_document parent_opt h' where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h': "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return () ) \<rightarrow>\<^sub>h h'"
using adopt_node
by(auto simp add: adopt_node_def CD.adopt_node_def elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_ancestors_di_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
split: if_splits)
then have "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using adopt_node
apply(auto simp add: adopt_node_def CD.adopt_node_def
dest!: bind_returns_heap_E3[rotated, OF old_document, rotated]
bind_returns_heap_E3[rotated, OF parent_opt, rotated]
elim!: bind_returns_heap_E2[rotated, OF get_ancestors_di_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
bind_returns_heap_E4[rotated, OF h', rotated] split: if_splits)[1]
apply(auto split: if_splits elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_ancestors_di_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1]
apply (simp add: set_disconnected_nodes_get_child_nodes children
reads_writes_preserved[OF get_child_nodes_reads set_disconnected_nodes_writes])
using children by blast
show ?thesis
proof(insert parent_opt h', induct parent_opt)
case None
then show ?case
using child_parent_dual wellformed known_ptrs type_wf \<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close>
returns_result_eq by fastforce
next
case (Some option)
then show ?case
using remove_child_removes_child \<open>h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children\<close> known_ptrs type_wf wellformed
by auto
qed
qed
lemma adopt_node_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> adopt_node document_ptr child \<rightarrow>\<^sub>h h'"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "heap_is_wellformed h'"
proof -
obtain old_document parent_opt h2 ancestors where
"h \<turnstile> get_ancestors_di (cast document_ptr) \<rightarrow>\<^sub>r ancestors" and
"cast child \<notin> set ancestors" and
old_document: "h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent child \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent child | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> (if document_ptr \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 child old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes document_ptr;
set_disconnected_nodes document_ptr (child # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^sub>h h'"
using assms(2)
apply(auto simp add: adopt_node_def[unfolded CD.adopt_node_def] elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_ancestors_di_pure])[1]
apply(split if_splits)
by(auto simp add: elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_parent_pure])
have object_ptr_kinds_h_eq3: "object_ptr_kinds h = object_ptr_kinds h2"
using h2 apply(simp split: option.splits)
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
unfolding object_ptr_kinds_M_defs by simp
then have object_ptr_kinds_eq_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have wellformed_h2: "heap_is_wellformed h2"
using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
then show "heap_is_wellformed h'"
proof(cases "document_ptr = old_document")
case True
then show "heap_is_wellformed h'"
using h' wellformed_h2 by auto
next
case False
then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where
docs_neq: "document_ptr \<noteq> old_document" and
old_disc_nodes: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h3: "h2 \<turnstile> set_disconnected_nodes old_document (remove1 child old_disc_nodes) \<rightarrow>\<^sub>h h3" and
disc_nodes_document_ptr_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \<rightarrow>\<^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="\<lambda>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: "\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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: "\<And>ptr children. h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h2 \<turnstile> get_child_nodes ptr|\<^sub>r = |h3 \<turnstile> 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="\<lambda>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: "\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h3: "|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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: "\<And>ptr children. h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^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: "\<And>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r = |h' \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. old_document \<noteq> doc_ptr \<Longrightarrow>
h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. old_document \<noteq> doc_ptr \<Longrightarrow>
|h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
using old_disc_nodes by blast
then have disc_nodes_old_document_h3:
"h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr \<Longrightarrow>
h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr \<Longrightarrow>
|h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2:
"h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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 |\<in>| 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 |\<in>| 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 \<in> 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 \<notin> set (remove1 child disc_nodes_old_document_h2)"
using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \<open>distinct disc_nodes_old_document_h2\<close>
by auto
have "child \<notin> 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 (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r))"
by(simp add: CD.a_distinct_lists_def)
show ?thesis
using distinct_concat_map_E(1)[OF 0] \<open>child \<in> set disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h2 disc_nodes_document_ptr_h2
by (meson \<open>type_wf h2\<close> 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 |\<in>| 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' \<subseteq> parent_child_rel h2"
proof
fix x
assume "x \<in> parent_child_rel h'"
then show "x \<in> 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 \<open> CD.a_acyclic_heap h2\<close> CD.acyclic_heap_def acyclic_subset by blast
moreover have " CD.a_all_ptrs_in_heap h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_all_ptrs_in_heap h3"
apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1]
apply (metis CD.l_heap_is_wellformed_axioms \<open>type_wf h2\<close> children_eq2_h2 known_ptrs
l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok
local.known_ptrs_known_ptr node_ptr_kinds_eq3_h2 object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3
returns_result_select_result wellformed_h2)
- by (metis (no_types, lifting) disc_nodes_old_document_h2 disc_nodes_old_document_h3
- disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 finite_set_in select_result_I2 set_remove1_subset
+ by (metis (no_types, opaque_lifting) disc_nodes_old_document_h2 disc_nodes_old_document_h3
+ disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 select_result_I2 set_remove1_subset
subsetD)
then have "CD.a_all_ptrs_in_heap h'"
by (smt \<open>child \<in> set disc_nodes_old_document_h2\<close> children_eq2_h3 disc_nodes_document_ptr_h'
disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3
- finite_set_in local.CD.a_all_ptrs_in_heap_def local.heap_is_wellformed_disc_nodes_in_heap
+ local.CD.a_all_ptrs_in_heap_def local.heap_is_wellformed_disc_nodes_in_heap
node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h3_eq3 select_result_I2 set_ConsD
subset_code(1) wellformed_h2)
moreover have "CD.a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
apply(simp add: CD.a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
children_eq2_h2 children_eq2_h3 )
by (metis (no_types) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2
disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3
in_set_remove1 list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2)
have a_distinct_lists_h2: "CD.a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h'"
apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2
children_eq2_h2 children_eq2_h3)[1]
proof -
assume 1: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r) (sorted_list_of_set
(fset (object_ptr_kinds h')))))"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 3: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
show "distinct (concat (map (\<lambda>document_ptr. |h' \<turnstile> 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 \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
have 4: "distinct |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r"
proof (cases "old_document \<noteq> x")
case True
then show ?thesis
proof (cases "document_ptr \<noteq> x")
case True
then show ?thesis
using disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>]
disconnected_nodes_eq2_h3[OF \<open>document_ptr \<noteq> x\<close>] 4
by(auto)
next
case False
then show ?thesis
using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4
\<open>child \<notin> set disc_nodes_document_ptr_h3\<close>
by(auto simp add: disconnected_nodes_eq2_h2[OF \<open>old_document \<noteq> x\<close>] )
qed
next
case False
then show ?thesis
by (metis (no_types, opaque_lifting) \<open>distinct disc_nodes_old_document_h2\<close>
disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 distinct_remove1 docs_neq select_result_I2)
qed
next
fix x y
assume a0: "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a1: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h')))"
and a2: "x \<noteq> y"
moreover have 5: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> 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' \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
proof(cases "old_document = x")
case True
have "old_document \<noteq> y"
using \<open>x \<noteq> y\<close> \<open>old_document = x\<close> by simp
have "document_ptr \<noteq> x"
using docs_neq \<open>old_document = x\<close> 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]
\<open>old_document = x\<close>
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
\<open>document_ptr \<noteq> x\<close> 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 \<open>old_document = x\<close>
docs_neq \<open>old_document \<noteq> y\<close>
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] \<open>old_document \<noteq> x\<close> \<open>old_document = y\<close> \<open>document_ptr = x\<close>
apply(simp)
by (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
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] \<open>old_document \<noteq> x\<close> \<open>old_document = y\<close>
\<open>document_ptr \<noteq> x\<close>
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 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False
\<open>type_wf h2\<close> 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 \<noteq> y"
using \<open>x \<noteq> y\<close> by auto
have "set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}"
using \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
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] \<open>old_document \<noteq> x\<close> \<open>old_document \<noteq> y\<close>
\<open>document_ptr = x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3 \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
by(auto)
next
case False
then show ?thesis
proof(cases "document_ptr = y")
case True
have f1: "set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set disc_nodes_document_ptr_h3 = {}"
using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>document_ptr \<noteq> x\<close> 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 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = {}"
using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
\<open>old_document \<noteq> x\<close>
by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2
- document_ptr_kinds_eq3_h3 finite_fset fmember_iff_member_fset set_sorted_list_of_set)
+ document_ptr_kinds_eq3_h3 finite_fset set_sorted_list_of_set)
ultimately show ?thesis
using 5 select_result_I2[OF disc_nodes_document_ptr_h']
select_result_I2[OF disc_nodes_old_document_h2] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr = y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> 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] \<open>old_document \<noteq> x\<close>
\<open>document_ptr \<noteq> x\<close> \<open>document_ptr \<noteq> y\<close>
\<open>child \<in> set disc_nodes_old_document_h2\<close> disconnected_nodes_eq2_h2
disconnected_nodes_eq2_h3
by (metis \<open>set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r \<inter> set disc_nodes_old_document_h2 = {}\<close>
empty_iff inf.idem)
qed
qed
qed
qed
qed
next
fix x xa xb
assume 0: "distinct (concat (map (\<lambda>ptr. |h' \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h2)))))"
and 2: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h2). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h'"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h'"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
then show False
using \<open>child \<in> set disc_nodes_old_document_h2\<close> 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 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a3: "h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r remove1 child disc_nodes_old_document_h2"
assume a4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a5: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f6: "old_document |\<in>| document_ptr_kinds h'"
using a1 \<open>xb |\<in>| document_ptr_kinds h'\<close> by blast
have f7: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a2 by simp
have "x \<in> set disc_nodes_old_document_h2"
using f6 a3 a1 by (metis (no_types) \<open>type_wf h'\<close> \<open>x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r\<close>
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' \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using f7 f6 a5 a4 \<open>xa |\<in>| object_ptr_kinds h'\<close>
by fastforce
then show ?thesis
using \<open>x \<in> set disc_nodes_old_document_h2\<close> a1 a4 f7 by blast
next
assume a1: "xb \<noteq> old_document"
assume a2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_document_ptr_h3"
assume a3: "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disc_nodes_old_document_h2"
assume a4: "xa |\<in>| object_ptr_kinds h'"
assume a5: "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r child # disc_nodes_document_ptr_h3"
assume a6: "old_document |\<in>| document_ptr_kinds h'"
assume a7: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
assume a8: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'"
assume a10: "\<And>doc_ptr. old_document \<noteq> doc_ptr \<Longrightarrow>
|h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a11: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr \<Longrightarrow>
|h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
assume a12: "(\<Union>x\<in>fset (object_ptr_kinds h'). set |h' \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h'). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
have f13: "\<And>d. d \<notin> set |h' \<turnstile> document_ptr_kinds_M|\<^sub>r \<or> h2 \<turnstile> ok get_disconnected_nodes d"
using a9 \<open>type_wf h2\<close> get_disconnected_nodes_ok
by simp
then have f14: "|h2 \<turnstile> get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2"
using a6 a3 by simp
have "x \<notin> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
using a12 a8 a4 \<open>xb |\<in>| document_ptr_kinds h'\<close>
- by (meson UN_I disjoint_iff_not_equal fmember_iff_member_fset)
+ by (meson UN_I disjoint_iff_not_equal)
then have "x = child"
using f13 a11 a10 a7 a5 a2 a1
by (metis (no_types, lifting) select_result_I2 set_ConsD)
then have "child \<notin> set disc_nodes_old_document_h2"
using f14 a12 a8 a6 a4
by (metis \<open>type_wf h'\<close> 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 \<open>child \<in> set disc_nodes_old_document_h2\<close> 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 \<open>type_wf h'\<close> \<open>CD.a_owner_document_valid h'\<close> CD.heap_is_wellformed_def by blast
have shadow_root_eq_h2: "\<And>ptr' shadow_root_ptr_opt. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have shadow_root_eq_h3: "\<And>ptr' shadow_root_ptr_opt. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h2: "\<And>ptr' tag. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h3: "\<And>ptr' tag. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> 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="\<lambda>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="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_disconnected_nodes_writes h'])
unfolding adopt_node_locs_def remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def split: if_splits)
have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h3"
using object_ptr_kinds_eq_h2
by(auto simp add: document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3"
using object_ptr_kinds_eq_h2
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3"
using object_ptr_kinds_eq_h2
by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def)
have document_ptr_kinds_eq_h3: "document_ptr_kinds h3 = document_ptr_kinds h'"
using object_ptr_kinds_eq_h3
by(auto simp add: document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'"
using object_ptr_kinds_eq_h3
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'"
using object_ptr_kinds_eq_h3
by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def)
have "a_host_shadow_root_rel h' = a_host_shadow_root_rel h3" and
"a_host_shadow_root_rel h3 = a_host_shadow_root_rel h2"
by(auto simp add: a_host_shadow_root_rel_def shadow_root_eq2_h2 shadow_root_eq2_h3
element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3)
have "parent_child_rel h' = parent_child_rel h3" and "parent_child_rel h3 = parent_child_rel h2"
by(auto simp add: CD.parent_child_rel_def children_eq2_h2 children_eq2_h3
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3)
have "parent_child_rel h2 \<subseteq> parent_child_rel h"
using h2 parent_opt
proof (induct parent_opt)
case None
then show ?case
by simp
next
case (Some parent)
then
have h2: "h \<turnstile> remove_child parent child \<rightarrow>\<^sub>h h2"
by auto
have child_nodes_eq_h: "\<And>ptr children. parent \<noteq> ptr \<Longrightarrow>
h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using get_child_nodes_reads remove_child_writes h2
apply(rule reads_writes_preserved)
apply(auto simp add: remove_child_locs_def)[1]
by (simp add: set_child_nodes_get_child_nodes_different_pointers)
moreover obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
using Some local.get_parent_child_dual by blast
ultimately show ?thesis
using object_ptr_kinds_eq_h h2
apply(auto simp add: CD.parent_child_rel_def split: option.splits)[1]
apply(case_tac "parent = a")
apply (metis (no_types, lifting) \<open>type_wf h3\<close> children_eq2_h2 children_eq_h2 known_ptrs
local.get_child_nodes_ok local.known_ptrs_known_ptr local.remove_child_children_subset
object_ptr_kinds_h2_eq3 returns_result_select_result subset_code(1) type_wf)
apply (metis (no_types, lifting) known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr
returns_result_select_result select_result_I2 type_wf)
done
qed
have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h"
using h2
proof (induct parent_opt)
case None
then show ?case
by simp
next
case (Some parent)
then
have h2: "h \<turnstile> remove_child parent child \<rightarrow>\<^sub>h h2"
by auto
have "\<And>ptr shadow_root. h \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r shadow_root = h2 \<turnstile> get_shadow_root ptr \<rightarrow>\<^sub>r shadow_root"
using get_shadow_root_reads remove_child_writes h2
apply(rule reads_writes_preserved)
apply(auto simp add: remove_child_locs_def)[1]
by (auto simp add: set_disconnected_nodes_get_shadow_root set_child_nodes_get_shadow_root)
then show ?case
apply(auto simp add: a_host_shadow_root_rel_def)[1]
apply (metis (mono_tags, lifting) Collect_cong \<open>type_wf h2\<close> case_prodE case_prodI
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_host_shadow_root_rel_def local.get_shadow_root_ok
local.a_host_shadow_root_rel_shadow_root returns_result_select_result)
by (metis (no_types, lifting) Collect_cong case_prodE case_prodI local.get_shadow_root_ok
local.a_host_shadow_root_rel_def local.a_host_shadow_root_rel_shadow_root returns_result_select_result type_wf)
qed
have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast old_document, cast child)}"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1]
using disconnected_nodes_eq2_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3
using \<open>distinct disc_nodes_old_document_h2\<close>
apply (metis (no_types, lifting) \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close>
case_prodI in_set_remove1 mem_Collect_eq pair_imageI select_result_I2)
using \<open>child \<notin> set (remove1 child disc_nodes_old_document_h2)\<close> disc_nodes_old_document_h3
apply auto[1]
by (metis (no_types, lifting) case_prodI disc_nodes_old_document_h2 disc_nodes_old_document_h3
disconnected_nodes_eq2_h2 in_set_remove1 mem_Collect_eq pair_imageI select_result_I2)
have "a_ptr_disconnected_node_rel h3 \<subseteq> a_ptr_disconnected_node_rel h"
using h2 parent_opt
proof (induct parent_opt)
case None
then show ?case
by(auto simp add: \<open>a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast old_document, cast child)}\<close>)
next
case (Some parent)
then
have h2: "h \<turnstile> remove_child parent child \<rightarrow>\<^sub>h h2"
by auto
then
obtain children_h h'2 disconnected_nodes_h where
children_h: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "child \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h'2: "h \<turnstile> set_disconnected_nodes old_document (child # disconnected_nodes_h) \<rightarrow>\<^sub>h h'2" and
h': "h'2 \<turnstile> set_child_nodes parent (remove1 child children_h) \<rightarrow>\<^sub>h h2"
using old_document
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure]
pure_returns_heap_eq[rotated, OF get_disconnected_nodes_pure] split: if_splits)[1]
using select_result_I2 by fastforce
have "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq_h object_ptr_kinds_eq_h2
by(auto simp add: document_ptr_kinds_def)
have disconnected_nodes_eq_h: "\<And>ptr disc_nodes. old_document \<noteq> ptr \<Longrightarrow>
h \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes = h2 \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads remove_child_writes h2
apply(rule reads_writes_preserved)
apply(auto simp add: remove_child_locs_def)[1]
using old_document
by (auto simp add:set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then
have foo: "\<And>ptr disc_nodes. old_document \<noteq> ptr \<Longrightarrow>
h \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes ptr \<rightarrow>\<^sub>r disc_nodes"
using disconnected_nodes_eq_h2 by simp
then
have foo2: "\<And>ptr. old_document \<noteq> ptr \<Longrightarrow> |h \<turnstile> get_disconnected_nodes ptr|\<^sub>r =
|h3 \<turnstile> get_disconnected_nodes ptr|\<^sub>r"
by (meson select_result_eq)
have "h'2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
using h'2
using local.set_disconnected_nodes_get_disconnected_nodes by blast
have "h2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
using get_disconnected_nodes_reads set_child_nodes_writes h'
\<open>h'2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r child # disconnected_nodes_h\<close>
apply(rule reads_writes_separate_forwards)
using local.set_child_nodes_get_disconnected_nodes by blast
then have "h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disconnected_nodes_h"
using h3
using disc_nodes_old_document_h2 disc_nodes_old_document_h3 returns_result_eq
by fastforce
have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h"
using \<open>|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h \<turnstile> document_ptr_kinds_M|\<^sub>r\<close>
apply(auto simp add: a_ptr_disconnected_node_rel_def )[1]
apply(case_tac "old_document = aa")
using disconnected_nodes_h \<open>h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disconnected_nodes_h\<close>
using foo2
apply(auto)[1]
using disconnected_nodes_h \<open>h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disconnected_nodes_h\<close>
using foo2
apply(auto)[1]
apply(case_tac "old_document = aa")
using disconnected_nodes_h \<open>h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disconnected_nodes_h\<close>
using foo2
apply(auto)[1]
using disconnected_nodes_h \<open>h3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disconnected_nodes_h\<close>
using foo2
apply(auto)[1]
done
then show ?thesis
by auto
qed
have "acyclic (parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2)"
using local.heap_is_wellformed_def wellformed_h2 by blast
then have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)"
using \<open>a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast old_document, cast child)}\<close>
by(auto simp add: \<open>parent_child_rel h3 = parent_child_rel h2\<close> \<open>a_host_shadow_root_rel h3 = a_host_shadow_root_rel h2\<close> elim!: acyclic_subset)
moreover
have "a_ptr_disconnected_node_rel h' = insert (cast document_ptr, cast child) (a_ptr_disconnected_node_rel h3)"
using disconnected_nodes_eq2_h3[symmetric] disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' document_ptr_in_heap[unfolded document_ptr_kinds_eq_h2]
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3[symmetric])[1]
apply(case_tac "document_ptr = aa")
apply(auto)[1]
apply(auto)[1]
apply(case_tac "document_ptr = aa")
apply(auto)[1]
apply(auto)[1]
done
moreover have "(cast child, cast document_ptr) \<notin> (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using \<open>h \<turnstile> get_ancestors_di (cast document_ptr) \<rightarrow>\<^sub>r ancestors\<close>
\<open>cast child \<notin> set ancestors\<close> get_ancestors_di_parent_child_a_host_shadow_root_rel
using assms(1) known_ptrs type_wf by blast
moreover have "(cast child, cast document_ptr) \<notin> (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)\<^sup>*"
proof -
have "(parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3 \<union> local.a_ptr_disconnected_node_rel h3) \<subseteq> (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> local.a_ptr_disconnected_node_rel h)"
apply(simp add: \<open>parent_child_rel h3 = parent_child_rel h2\<close> \<open>a_host_shadow_root_rel h3 = a_host_shadow_root_rel h2\<close> \<open>a_host_shadow_root_rel h2 = a_host_shadow_root_rel h\<close>)
using \<open>local.a_ptr_disconnected_node_rel h3 \<subseteq> local.a_ptr_disconnected_node_rel h\<close> \<open>parent_child_rel h2 \<subseteq> parent_child_rel h\<close>
by blast
then show ?thesis
using calculation(3) rtrancl_mono by blast
qed
ultimately have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')"
by(auto simp add: \<open>parent_child_rel h' = parent_child_rel h3\<close> \<open>a_host_shadow_root_rel h' = a_host_shadow_root_rel h3\<close>)
show "heap_is_wellformed h'"
using \<open>heap_is_wellformed h2\<close>
using \<open>heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\<close>
using \<open>acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')\<close>
apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def
a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1]
by(auto simp add: object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2
element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2
shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2
tag_name_eq2_h3)
qed
qed
lemma adopt_node_node_in_disconnected_nodes:
assumes wellformed: "heap_is_wellformed h"
and adopt_node: "h \<turnstile> adopt_node owner_document node_ptr \<rightarrow>\<^sub>h h'"
and "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "node_ptr \<in> set disc_nodes"
proof -
obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast node_ptr) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node_ptr \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node_ptr | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h2" and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node_ptr # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^sub>h h'"
using assms(2)[unfolded adopt_node_def CD.adopt_node_def]
by(auto elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_parent_pure] pure_returns_heap_eq[rotated, OF get_ancestors_di_pure]
split: option.splits if_splits)
show ?thesis
proof (cases "owner_document = old_document")
case True
then show ?thesis
proof (insert parent_opt h2, induct parent_opt)
case None
then have "h = h'"
using h2 h' by(auto)
then show ?case
using in_disconnected_nodes_no_parent assms None old_document by blast
next
case (Some parent)
then show ?case
using remove_child_in_disconnected_nodes known_ptrs True h' assms(3) old_document
by auto
qed
next
case False
then show ?thesis
using assms(3) h' list.set_intros(1) select_result_I2
set_disconnected_nodes_get_disconnected_nodes
apply(auto elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1]
proof -
fix x and h'a and xb
assume a1: "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes"
assume a2: "\<And>h document_ptr disc_nodes h'. h \<turnstile> set_disconnected_nodes document_ptr disc_nodes \<rightarrow>\<^sub>h h' \<Longrightarrow>
h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes"
assume "h'a \<turnstile> set_disconnected_nodes owner_document (node_ptr # xb) \<rightarrow>\<^sub>h h'"
then have "node_ptr # xb = disc_nodes"
using a2 a1 by (meson returns_result_eq)
then show ?thesis
by (meson list.set_intros(1))
qed
qed
qed
end
interpretation i_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes
set_child_nodes_locs remove heap_is_wellformed parent_child_rel
by(auto simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare i_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
interpretation i_adopt_node_wf?: l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs get_ancestors_di
get_ancestors_di_locs adopt_node adopt_node_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf
get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs get_host
get_host_locs get_disconnected_document get_disconnected_document_locs remove heap_is_wellformed
parent_child_rel
by(auto simp add: l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes
set_disconnected_nodes_locs get_tag_name get_tag_name_locs get_owner_document get_parent get_parent_locs
remove_child remove_child_locs remove known_ptrs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_root_node
get_root_node_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
by(auto simp add: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma adopt_node_wf_is_l_adopt_node_wf [instances]:
"l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes
get_disconnected_nodes known_ptrs adopt_node"
apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def instances)[1]
using adopt_node_preserves_wellformedness apply blast
using adopt_node_removes_child apply blast
using adopt_node_node_in_disconnected_nodes apply blast
using adopt_node_removes_first_child apply blast
using adopt_node_document_in_heap apply blast
done
subsubsection \<open>insert\_before\<close>
locale l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_child_nodes +
l_get_disconnected_nodes +
l_set_child_nodes_get_shadow_root +
l_set_disconnected_nodes_get_shadow_root +
l_set_child_nodes_get_tag_name +
l_set_disconnected_nodes_get_tag_name +
l_set_disconnected_nodes_get_disconnected_nodes +
l_set_child_nodes_get_disconnected_nodes +
l_set_disconnected_nodes_get_disconnected_nodes_wf +
(* l_set_disconnected_nodes_get_ancestors_si + *)
l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ get_ancestors_di get_ancestors_di_locs +
(* l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + *)
l_get_owner_document +
l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node_wf +
l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_adopt_node_get_shadow_root +
l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_remove_child_wf2
begin
lemma insert_before_child_preserves:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> insert_before ptr node child \<rightarrow>\<^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 \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
reference_child: "h \<turnstile> (if Some node = child then a_next_sibling node else return child) \<rightarrow>\<^sub>r reference_child" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node reference_child \<rightarrow>\<^sub>h h'"
using assms(4)
by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
obtain old_document parent_opt h'2 (* ancestors *) where
(* "h \<turnstile> get_ancestors_di (cast owner_document) \<rightarrow>\<^sub>r ancestors" and
"cast child \<notin> set ancestors" and *)
old_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node \<rightarrow>\<^sub>r parent_opt" and
h'2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> remove_child parent node | None \<Rightarrow> return ()) \<rightarrow>\<^sub>h h'2" and
h2': "h'2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
set_disconnected_nodes old_document (remove1 node old_disc_nodes);
disc_nodes \<leftarrow> get_disconnected_nodes owner_document;
set_disconnected_nodes owner_document (node # disc_nodes)
} else do {
return ()
}) \<rightarrow>\<^sub>h h2"
using h2
apply(auto simp add: adopt_node_def[unfolded CD.adopt_node_def] elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_ancestors_di_pure])[1]
apply(split if_splits)
by(auto simp add: elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure])
have "type_wf h2"
using \<open>type_wf h\<close>
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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 "\<dots> = 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 "\<dots> = 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 \<open>known_ptrs h\<close> known_ptrs_preserved
by blast
have "known_ptrs h2"
using \<open>known_ptrs h\<close> known_ptrs_preserved \<open>object_ptr_kinds h = object_ptr_kinds h2\<close>
by blast
then
have "known_ptrs h3"
using known_ptrs_preserved \<open>object_ptr_kinds h2 = object_ptr_kinds h3\<close>
by blast
have "known_ptr ptr"
by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I \<open>known_ptrs h\<close>
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="\<lambda>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: "\<And>ptrs. h \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs )
then have object_ptr_kinds_M_eq2_h: "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h: "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> 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 \<open>heap_is_wellformed h\<close> h2] \<open>known_ptrs h\<close> \<open>type_wf h\<close>
.
have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>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: "\<And>ptrs. h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h2: "|h2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h2: "|h2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h2: "|h2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h3 \<turnstile> 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="\<lambda>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: "\<And>ptrs. h3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h' \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_M_eq2_h3: "|h3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
then have node_ptr_kinds_eq2_h3: "|h3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
have document_ptr_kinds_eq2_h3: "|h3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> 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: "\<And>ptr' shadow_root. h \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root =
h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>doc_ptr disc_nodes. owner_document \<noteq> doc_ptr \<Longrightarrow>
h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. doc_ptr \<noteq> owner_document \<Longrightarrow>
|h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_h3: "h3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r remove1 node disconnected_nodes_h2"
using h3 set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h3:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr \<noteq> ptr' \<Longrightarrow> |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
obtain children_h3 where children_h3: "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h3"
using h' a_insert_node_def by auto
have children_h': "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r insert_before_list node reference_child children_h3"
using h' \<open>type_wf h3\<close> \<open>known_ptr ptr\<close>
by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2
dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3])
have shadow_root_eq_h: "\<And>ptr' shadow_root_ptr_opt. h \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt"
using get_shadow_root_reads adopt_node_writes h2
apply(rule reads_writes_preserved)
by(auto simp add: adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def
set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root)
then
have shadow_root_eq2_h: "\<And>ptr'. |h \<turnstile> get_shadow_root ptr'|\<^sub>r = |h2 \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have shadow_root_eq_h2: "\<And>ptr' shadow_root_ptr_opt. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have shadow_root_eq_h3: "\<And>ptr' shadow_root_ptr_opt. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r shadow_root_ptr_opt =
h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h2: "\<And>ptr' tag. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
by (meson select_result_eq)
have tag_name_eq_h3: "\<And>ptr' tag. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r tag = h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> get_tag_name ptr'|\<^sub>r"
by (meson select_result_eq)
have object_ptr_kinds_eq_hx: "object_ptr_kinds h = object_ptr_kinds h'2"
using h'2 apply(simp split: option.splits)
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'", OF CD.remove_child_writes])
using CD.remove_child_pointers_preserved
by (auto simp add: reflp_def transp_def)
have document_ptr_kinds_eq_hx: "document_ptr_kinds h = document_ptr_kinds h'2"
using object_ptr_kinds_eq_hx
by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_hx: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'2"
using object_ptr_kinds_eq_hx
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
have element_ptr_kinds_eq_hx: "element_ptr_kinds h = element_ptr_kinds h'2"
using object_ptr_kinds_eq_hx
by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def)
have object_ptr_kinds_eq_h: "object_ptr_kinds h = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2])
unfolding adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def split: if_splits)
have document_ptr_kinds_eq_h: "document_ptr_kinds h = document_ptr_kinds h2"
using object_ptr_kinds_eq_h
by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h2"
using object_ptr_kinds_eq_h
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
have element_ptr_kinds_eq_h: "element_ptr_kinds h = element_ptr_kinds h2"
using object_ptr_kinds_eq_h
by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def)
have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3])
unfolding adopt_node_locs_def remove_child_locs_def
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def split: if_splits)
have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h3"
using object_ptr_kinds_eq_h2
by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3"
using object_ptr_kinds_eq_h2
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3"
using object_ptr_kinds_eq_h2
by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h'])
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def split: if_splits)
have document_ptr_kinds_eq_h3: "document_ptr_kinds h3 = document_ptr_kinds h'"
using object_ptr_kinds_eq_h3
by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'"
using object_ptr_kinds_eq_h3
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'"
using object_ptr_kinds_eq_h3
by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def)
have wellformed_h'2: "heap_is_wellformed h'2"
using h'2 remove_child_heap_is_wellformed_preserved assms
by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
have "known_ptrs h'2"
using \<open>known_ptrs h\<close> known_ptrs_preserved \<open>object_ptr_kinds h = object_ptr_kinds h'2\<close>
by blast
have "type_wf h'2"
using \<open>type_wf h\<close> h'2
apply(auto split: option.splits)[1]
apply(drule writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF CD.remove_child_writes])
using CD.remove_child_types_preserved
by(auto simp add: reflp_def transp_def )
have ptr_in_heap: "ptr |\<in>| object_ptr_kinds h3"
using children_h3 get_child_nodes_ptr_in_heap by blast
have node_in_heap: "node |\<in>| node_ptr_kinds h"
using h2 adopt_node_child_in_heap by fast
have child_not_in_any_children: "\<And>p children. h2 \<turnstile> get_child_nodes p \<rightarrow>\<^sub>r children \<Longrightarrow> node \<notin> set children"
using \<open>heap_is_wellformed h\<close> h2 adopt_node_removes_child \<open>type_wf h\<close> \<open>known_ptrs h\<close> by auto
have "node \<in> set disconnected_nodes_h2"
using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1) \<open>type_wf h\<close> \<open>known_ptrs h\<close> by blast
have node_not_in_disconnected_nodes: "\<And>d. d |\<in>| document_ptr_kinds h3 \<Longrightarrow> node \<notin> set |h3 \<turnstile> get_disconnected_nodes d|\<^sub>r"
proof -
fix d
assume "d |\<in>| document_ptr_kinds h3"
show "node \<notin> set |h3 \<turnstile> 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
\<open>d |\<in>| document_ptr_kinds h3\<close> disconnected_nodes_h3
by fastforce
next
case False
then have "set |h2 \<turnstile> get_disconnected_nodes d|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes owner_document|\<^sub>r = {}"
using distinct_concat_map_E(1) wellformed_h2
by (metis (no_types, lifting) \<open>d |\<in>| document_ptr_kinds h3\<close> \<open>type_wf h2\<close> 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] \<open>node \<in> set disconnected_nodes_h2\<close> disconnected_nodes_h2 by fastforce
qed
qed
have "cast node \<noteq> ptr"
using ancestors node_not_in_ancestors get_ancestors_ptr
by fast
have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq2_h)
have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq2_h2)
have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 shadow_root_eq2_h3)
have "parent_child_rel h2 \<subseteq> parent_child_rel h"
proof -
have "parent_child_rel h'2 \<subseteq> parent_child_rel h"
using h'2 parent_opt
proof (induct parent_opt)
case None
then show ?case
by simp
next
case (Some parent)
then
have h'2: "h \<turnstile> remove_child parent node \<rightarrow>\<^sub>h h'2"
by auto
then
have "parent |\<in>| object_ptr_kinds h"
using CD.remove_child_ptr_in_heap
by blast
have child_nodes_eq_h: "\<And>ptr children. parent \<noteq> ptr \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children =
h'2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using get_child_nodes_reads CD.remove_child_writes h'2
apply(rule reads_writes_preserved)
apply(auto simp add: CD.remove_child_locs_def)[1]
by (simp add: set_child_nodes_get_child_nodes_different_pointers)
moreover obtain children where children: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children"
using Some local.get_parent_child_dual by blast
moreover obtain children_h'2 where children_h'2: "h'2 \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children_h'2"
using object_ptr_kinds_eq_hx calculation(2) \<open>parent |\<in>| object_ptr_kinds h\<close> get_child_nodes_ok
by (metis \<open>type_wf h'2\<close> assms(3) is_OK_returns_result_E local.known_ptrs_known_ptr)
ultimately show ?thesis
using object_ptr_kinds_eq_h h2
apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_hx split: option.splits)[1]
apply(case_tac "parent = a")
using CD.remove_child_children_subset
apply (metis (no_types, lifting) assms(2) assms(3) contra_subsetD h'2 select_result_I2)
by (metis select_result_eq)
qed
moreover have "parent_child_rel h2 = parent_child_rel h'2"
proof(cases "owner_document = old_document")
case True
then show ?thesis
using h2' by simp
next
case False
then obtain h'3 old_disc_nodes disc_nodes_document_ptr_h'3 where
docs_neq: "owner_document \<noteq> old_document" and
old_disc_nodes: "h'2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h'3: "h'2 \<turnstile> set_disconnected_nodes old_document (remove1 node old_disc_nodes) \<rightarrow>\<^sub>h h'3" and
disc_nodes_document_ptr_h3: "h'3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes_document_ptr_h'3" and
h2': "h'3 \<turnstile> set_disconnected_nodes owner_document (node # disc_nodes_document_ptr_h'3) \<rightarrow>\<^sub>h h2"
using h2'
by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
have object_ptr_kinds_h'2_eq3: "object_ptr_kinds h'2 = object_ptr_kinds h'3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_disconnected_nodes_writes h'3])
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_M_eq_h'2: "\<And>ptrs. h'2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h'3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h'2: "|h'2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h'3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h'2: "|h'2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h'3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
then have node_ptr_kinds_eq3_h'2: "node_ptr_kinds h'2 = node_ptr_kinds h'3"
by auto
have document_ptr_kinds_eq2_h'2: "|h'2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h'3 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3_h'2: "document_ptr_kinds h'2 = document_ptr_kinds h'3"
using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto
have children_eq_h'2: "\<And>ptr children. h'2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children = h'3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h'3
apply(rule reads_writes_preserved)
by (simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h'2: "\<And>ptr. |h'2 \<turnstile> get_child_nodes ptr|\<^sub>r = |h'3 \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have object_ptr_kinds_h'3_eq3: "object_ptr_kinds h'3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h2. object_ptr_kinds h = object_ptr_kinds h2",
OF set_disconnected_nodes_writes h2'])
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_M_eq_h'3: "\<And>ptrs. h'3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h'3: "|h'3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h'3: "|h'3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
then have node_ptr_kinds_eq3_h'3: "node_ptr_kinds h'3 = node_ptr_kinds h2"
by auto
have document_ptr_kinds_eq2_h'3: "|h'3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3_h'3: "document_ptr_kinds h'3 = document_ptr_kinds h2"
using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto
have children_eq_h'3: "\<And>ptr children. h'3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children =
h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children"
using get_child_nodes_reads set_disconnected_nodes_writes h2'
apply(rule reads_writes_preserved)
by (simp add: set_disconnected_nodes_get_child_nodes)
then have children_eq2_h'3: "\<And>ptr. |h'3 \<turnstile> get_child_nodes ptr|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
show ?thesis
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_h'2_eq3 object_ptr_kinds_h'3_eq3
children_eq2_h'3 children_eq2_h'2)
qed
ultimately
show ?thesis
by simp
qed
have "parent_child_rel h2 = parent_child_rel h3"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))"
using children_h3 children_h' ptr_in_heap
apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3
insert_before_list_node_in_set)[1]
apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2)
by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2)
have "a_ptr_disconnected_node_rel h3 \<subseteq> a_ptr_disconnected_node_rel h"
proof -
have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2)[1]
apply(case_tac "aa = owner_document")
using disconnected_nodes_h2 disconnected_nodes_h3 notin_set_remove1 apply fastforce
using disconnected_nodes_eq2_h2 apply auto[1]
using node_not_in_disconnected_nodes apply blast
apply(case_tac "aa = owner_document")
using disconnected_nodes_h2 disconnected_nodes_h3 notin_set_remove1 apply fastforce
using disconnected_nodes_eq2_h2 apply auto[1]
apply(case_tac "aa = owner_document")
using disconnected_nodes_h2 disconnected_nodes_h3 notin_set_remove1 apply fastforce
using disconnected_nodes_eq2_h2 apply auto[1]
done
then have "a_ptr_disconnected_node_rel h'2 \<subseteq> a_ptr_disconnected_node_rel h \<union> {(cast old_document, cast node)}"
using h'2 parent_opt
proof (induct parent_opt)
case None
then show ?case
by auto
next
case (Some parent)
then
have h'2: "h \<turnstile> remove_child parent node \<rightarrow>\<^sub>h h'2"
by auto
then
have "parent |\<in>| object_ptr_kinds h"
using CD.remove_child_ptr_in_heap
by blast
obtain children_h h''2 disconnected_nodes_h where
owner_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and
children_h: "h \<turnstile> get_child_nodes parent \<rightarrow>\<^sub>r children_h" and
child_in_children_h: "node \<in> set children_h" and
disconnected_nodes_h: "h \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r disconnected_nodes_h" and
h''2: "h \<turnstile> set_disconnected_nodes old_document (node # disconnected_nodes_h) \<rightarrow>\<^sub>h h''2" and
h'2: "h''2 \<turnstile> set_child_nodes parent (remove1 node children_h) \<rightarrow>\<^sub>h h'2"
using h'2 old_document
apply(auto simp add: CD.remove_child_def elim!: bind_returns_heap_E
dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1]
using pure_returns_heap_eq returns_result_eq by fastforce
have disconnected_nodes_eq: "\<And>ptr' disc_nodes. ptr' \<noteq> old_document \<Longrightarrow>
h \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes = h''2 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes"
using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h''2
apply(rule reads_writes_preserved)
by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then
have disconnected_nodes_eq2: "\<And>ptr'. ptr' \<noteq> old_document \<Longrightarrow>
|h \<turnstile> get_disconnected_nodes ptr'|\<^sub>r = |h''2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r"
by (meson select_result_eq)
have "h''2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r node # disconnected_nodes_h"
using h''2 local.set_disconnected_nodes_get_disconnected_nodes
by blast
have disconnected_nodes_eq_h2:
"\<And>ptr' disc_nodes. h''2 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes = h'2 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes"
using local.get_disconnected_nodes_reads set_child_nodes_writes h'2
apply(rule reads_writes_preserved)
by (metis local.set_child_nodes_get_disconnected_nodes)
then
have disconnected_nodes_eq2_h2:
"\<And>ptr'. |h''2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r = |h'2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r"
by (meson select_result_eq)
show ?case
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_hx
\<open>\<And>ptr'. |h''2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r = |h'2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r\<close>[symmetric])[1]
apply(case_tac "aa = old_document")
using disconnected_nodes_h
\<open>h''2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r node # disconnected_nodes_h\<close>
apply(auto)[1]
apply(auto dest!: disconnected_nodes_eq2)[1]
apply(case_tac "aa = old_document")
using disconnected_nodes_h
\<open>h''2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r node # disconnected_nodes_h\<close>
apply(auto)[1]
apply(auto dest!: disconnected_nodes_eq2)[1]
done
qed
show ?thesis
proof(cases "owner_document = old_document")
case True
then have "a_ptr_disconnected_node_rel h'2 = a_ptr_disconnected_node_rel h2"
using h2'
by(auto simp add: a_ptr_disconnected_node_rel_def)
then
show ?thesis
using \<open>a_ptr_disconnected_node_rel h3 =
a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}\<close>
using True \<open>local.a_ptr_disconnected_node_rel h'2 \<subseteq> local.a_ptr_disconnected_node_rel h \<union>
{(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r old_document, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node)}\<close> by auto
next
case False
then obtain h'3 old_disc_nodes disc_nodes_document_ptr_h'3 where
docs_neq: "owner_document \<noteq> old_document" and
old_disc_nodes: "h'2 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r old_disc_nodes" and
h'3: "h'2 \<turnstile> set_disconnected_nodes old_document (remove1 node old_disc_nodes) \<rightarrow>\<^sub>h h'3" and
disc_nodes_document_ptr_h3: "h'3 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes_document_ptr_h'3" and
h2': "h'3 \<turnstile> set_disconnected_nodes owner_document (node # disc_nodes_document_ptr_h'3) \<rightarrow>\<^sub>h h2"
using h2'
by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
have object_ptr_kinds_h'2_eq3: "object_ptr_kinds h'2 = object_ptr_kinds h'3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h = object_ptr_kinds h'",
OF set_disconnected_nodes_writes h'3])
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_M_eq_h'2:
"\<And>ptrs. h'2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h'3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h'2: "|h'2 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h'3 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h'2: "|h'2 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h'3 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
then have node_ptr_kinds_eq3_h'2: "node_ptr_kinds h'2 = node_ptr_kinds h'3"
by auto
have document_ptr_kinds_eq2_h'2: "|h'2 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h'3 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3_h'2: "document_ptr_kinds h'2 = document_ptr_kinds h'3"
using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto
have disconnected_nodes_eq: "\<And>ptr' disc_nodes. ptr' \<noteq> old_document \<Longrightarrow>
h'2 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes = h'3 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes"
using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h'3
apply(rule reads_writes_preserved)
by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then
have disconnected_nodes_eq2: "\<And>ptr'. ptr' \<noteq> old_document \<Longrightarrow>
|h'2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r = |h'3 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r"
by (meson select_result_eq)
have "h'3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r (remove1 node old_disc_nodes)"
using h'3 local.set_disconnected_nodes_get_disconnected_nodes
by blast
have object_ptr_kinds_h'3_eq3: "object_ptr_kinds h'3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h2. object_ptr_kinds h = object_ptr_kinds h2",
OF set_disconnected_nodes_writes h2'])
using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have object_ptr_kinds_M_eq_h'3: "\<And>ptrs. h'3 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs = h2 \<turnstile> object_ptr_kinds_M \<rightarrow>\<^sub>r ptrs"
by(simp add: object_ptr_kinds_M_defs)
then have object_ptr_kinds_eq_h'3: "|h'3 \<turnstile> object_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> object_ptr_kinds_M|\<^sub>r"
by(simp)
then have node_ptr_kinds_eq_h'3: "|h'3 \<turnstile> node_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_ptr_kinds_M_eq by blast
then have node_ptr_kinds_eq3_h'3: "node_ptr_kinds h'3 = node_ptr_kinds h2"
by auto
have document_ptr_kinds_eq2_h'3: "|h'3 \<turnstile> document_ptr_kinds_M|\<^sub>r = |h2 \<turnstile> document_ptr_kinds_M|\<^sub>r"
using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto
then have document_ptr_kinds_eq3_h'3: "document_ptr_kinds h'3 = document_ptr_kinds h2"
using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto
have disc_nodes_eq_h'3:
"\<And>ptr disc_nodes. h'3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r disc_nodes = h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r disc_nodes"
using get_child_nodes_reads set_disconnected_nodes_writes h2'
apply(rule reads_writes_preserved)
by (simp add: set_disconnected_nodes_get_child_nodes)
then have disc_nodes_eq2_h'3: "\<And>ptr. |h'3 \<turnstile> get_child_nodes ptr|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq: "\<And>ptr' disc_nodes. ptr' \<noteq> owner_document \<Longrightarrow>
h'3 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes = h2 \<turnstile> get_disconnected_nodes ptr' \<rightarrow>\<^sub>r disc_nodes"
using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h2'
apply(rule reads_writes_preserved)
by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers)
then
have disconnected_nodes_eq2': "\<And>ptr'. ptr' \<noteq> owner_document \<Longrightarrow>
|h'3 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes ptr'|\<^sub>r"
by (meson select_result_eq)
have "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r (node # disc_nodes_document_ptr_h'3)"
using h2' local.set_disconnected_nodes_get_disconnected_nodes
by blast
have "a_ptr_disconnected_node_rel h'3 = a_ptr_disconnected_node_rel h'2 - {(cast old_document, cast node)}"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq2_h'2[simplified])[1]
apply(case_tac "aa = old_document")
using \<open>h'3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r (remove1 node old_disc_nodes)\<close>
notin_set_remove1 old_disc_nodes
apply fastforce
apply(auto dest!: disconnected_nodes_eq2)[1]
using \<open>h'3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r remove1 node old_disc_nodes\<close> h'3
local.remove_from_disconnected_nodes_removes old_disc_nodes wellformed_h'2
apply auto[1]
defer
apply(case_tac "aa = old_document")
using \<open>h'3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r (remove1 node old_disc_nodes)\<close>
notin_set_remove1 old_disc_nodes
apply fastforce
apply(auto dest!: disconnected_nodes_eq2)[1]
apply(case_tac "aa = old_document")
using \<open>h'3 \<turnstile> get_disconnected_nodes old_document \<rightarrow>\<^sub>r (remove1 node old_disc_nodes)\<close>
notin_set_remove1 old_disc_nodes
apply fastforce
apply(auto dest!: disconnected_nodes_eq2)[1]
done
moreover
have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h'3 \<union>
{(cast owner_document, cast node)}"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq2_h'3[simplified])[1]
apply(case_tac "aa = owner_document")
apply(simp)
apply(auto dest!: disconnected_nodes_eq2')[1]
apply(case_tac "aa = owner_document")
using \<open>h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r node # disc_nodes_document_ptr_h'3\<close>
disc_nodes_document_ptr_h3 apply auto[1]
apply(auto dest!: disconnected_nodes_eq2')[1]
using \<open>node \<in> set disconnected_nodes_h2\<close> disconnected_nodes_h2 local.a_ptr_disconnected_node_rel_def
local.a_ptr_disconnected_node_rel_disconnected_node apply blast
defer
apply(case_tac "aa = owner_document")
using \<open>h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r node # disc_nodes_document_ptr_h'3\<close>
disc_nodes_document_ptr_h3 apply auto[1]
apply(auto dest!: disconnected_nodes_eq2')[1]
done
ultimately show ?thesis
using \<open>a_ptr_disconnected_node_rel h3 =
a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}\<close>
using \<open>a_ptr_disconnected_node_rel h'2 \<subseteq>
a_ptr_disconnected_node_rel h \<union> {(cast old_document, cast node)}\<close>
by blast
qed
qed
have "(cast node, ptr) \<notin> (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)\<^sup>*"
using h2
apply(auto simp add: adopt_node_def elim!: bind_returns_heap_E2 split: if_splits)[1]
using ancestors assms(1) assms(2) assms(3) local.get_ancestors_di_parent_child_a_host_shadow_root_rel node_not_in_ancestors
by blast
then
have "(cast node, ptr) \<notin> (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)\<^sup>*"
apply(simp add: \<open>a_host_shadow_root_rel h = a_host_shadow_root_rel h2\<close> \<open>a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\<close>)
apply(simp add: \<open>parent_child_rel h2 = parent_child_rel h3\<close>[symmetric])
using \<open>parent_child_rel h2 \<subseteq> parent_child_rel h\<close> \<open>a_ptr_disconnected_node_rel h3 \<subseteq> a_ptr_disconnected_node_rel h\<close>
by (smt Un_assoc in_rtrancl_UnI sup.orderE sup_left_commute)
have "CD.a_acyclic_heap h'"
proof -
have "acyclic (parent_child_rel h2)"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def)
then have "acyclic (parent_child_rel h3)"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
moreover have "cast node \<notin> {x. (x, ptr) \<in> (parent_child_rel h3)\<^sup>*}"
by (meson \<open>(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) \<notin> (parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3 \<union>
local.a_ptr_disconnected_node_rel h3)\<^sup>*\<close> in_rtrancl_UnI mem_Collect_eq)
ultimately show ?thesis
using \<open>parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\<close>
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 \<open>CD.a_all_ptrs_in_heap h2\<close>
apply(auto simp add: CD.a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2
children_eq_h2)[1]
apply (metis \<open>known_ptrs h3\<close> \<open>type_wf h3\<close> children_eq_h2
l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok
local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms node_ptr_kinds_commutes
object_ptr_kinds_eq_h2 returns_result_select_result wellformed_h2)
by (metis (mono_tags, opaque_lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2
- disconnected_nodes_h3 document_ptr_kinds_eq_h2 finite_set_in node_ptr_kinds_commutes
+ disconnected_nodes_h3 document_ptr_kinds_eq_h2 node_ptr_kinds_commutes
object_ptr_kinds_eq_h2 select_result_I2 set_remove1_subset subsetD)
have "set children_h3 \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using children_h3 \<open>CD.a_all_ptrs_in_heap h3\<close>
apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1]
using children_eq_h2 local.heap_is_wellformed_children_in_heap node_ptr_kinds_eq2_h2
node_ptr_kinds_eq2_h3 wellformed_h2 by auto
then have "set (insert_before_list node reference_child children_h3) \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_in_heap
apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1]
- by (metis (no_types, opaque_lifting) contra_subsetD finite_set_in insert_before_list_in_set
+ by (metis (no_types, opaque_lifting) contra_subsetD insert_before_list_in_set
node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2)
then show ?thesis
using \<open>CD.a_all_ptrs_in_heap h3\<close>
apply(auto simp add: object_ptr_kinds_M_eq3_h' CD.a_all_ptrs_in_heap_def node_ptr_kinds_def
node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1]
- apply (metis (no_types, lifting) children_eq2_h3 children_h' finite_set_in select_result_I2 subsetD)
- by (metis (no_types, lifting) disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 in_mono notin_fset)
+ apply (metis (no_types, lifting) children_eq2_h3 children_h' select_result_I2 subsetD)
+ by (metis (no_types, lifting) disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 in_mono)
qed
moreover have "CD.a_distinct_lists h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def )
then have "CD.a_distinct_lists h3"
proof(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2
children_eq2_h2 intro!: distinct_concat_map_I)
fix x
assume 1: "x |\<in>| document_ptr_kinds h3"
and 2: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
show "distinct |h3 \<turnstile> 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_iff_member_fset set_sorted_list_of_set)
+ by (metis (full_types) distinct_remove1 finite_fset set_sorted_list_of_set)
next
fix x y xa
assume 1: "distinct (concat (map (\<lambda>document_ptr. |h2 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and 2: "x |\<in>| document_ptr_kinds h3"
and 3: "y |\<in>| document_ptr_kinds h3"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r"
and 6: "xa \<in> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r"
show False
proof (cases "x = owner_document")
case True
then have "y \<noteq> 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 \<open>y \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
next
case False
then show ?thesis
proof (cases "y = owner_document")
case True
then show ?thesis
using distinct_concat_map_E(1)[OF 1]
using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2]
apply(auto simp add: True disconnected_nodes_eq2_h2[OF \<open>x \<noteq> owner_document\<close>])[1]
by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
next
case False
then show ?thesis
using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6
using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
- disjoint_iff_not_equal finite_fset fmember_iff_member_fset notin_set_remove1 select_result_I2
+ disjoint_iff_not_equal finite_fset notin_set_remove1 select_result_I2
set_sorted_list_of_set
by (metis (no_types, lifting))
qed
qed
next
fix x xa xb
assume 1: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h3 \<turnstile> get_child_nodes x|\<^sub>r) \<inter>
(\<Union>x\<in>fset (document_ptr_kinds h3). set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 2: "xa |\<in>| object_ptr_kinds h3"
and 3: "x \<in> set |h3 \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h3"
and 5: "x \<in> set |h3 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 4
by (metis \<open>type_wf h2\<close> children_eq2_h2 document_ptr_kinds_commutes \<open>known_ptrs h\<close>
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 (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and
2: "x |\<in>| object_ptr_kinds h'"
have 3: "\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> distinct |h3 \<turnstile> get_child_nodes p|\<^sub>r"
using 1 by (auto elim: distinct_concat_map_E)
show "distinct |h' \<turnstile> 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 (\<lambda>ptr. |h3 \<turnstile> get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))"
and 2: "x |\<in>| object_ptr_kinds h'"
and 3: "y |\<in>| object_ptr_kinds h'"
and 4: "x \<noteq> y"
and 5: "xa \<in> set |h' \<turnstile> get_child_nodes x|\<^sub>r"
and 6: "xa \<in> set |h' \<turnstile> get_child_nodes y|\<^sub>r"
have 7:"set |h3 \<turnstile> get_child_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 \<noteq> 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 \<open>ptr \<noteq> y\<close>])[1]
by (metis (no_types, opaque_lifting) "3" "7" \<open>type_wf h3\<close> children_eq2_h3 disjoint_iff_not_equal
get_child_nodes_ok insert_before_list_in_set \<open>known_ptrs h\<close> 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 \<open>ptr \<noteq> x\<close>])[1]
by (metis (no_types, opaque_lifting) "2" "4" "7" IntI \<open>known_ptrs h3\<close> \<open>type_wf h'\<close>
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 \<open>ptr \<noteq> x\<close>] children_eq2_h3[OF \<open>ptr \<noteq> y\<close>] 5 6 7 by auto
qed
qed
next
fix x xa xb
assume 1: " (\<Union>x\<in>fset (object_ptr_kinds h'). set |h3 \<turnstile> get_child_nodes x|\<^sub>r) \<inter> (\<Union>x\<in>fset (document_ptr_kinds h'). set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r) = {} "
and 2: "xa |\<in>| object_ptr_kinds h'"
and 3: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 4: "xb |\<in>| document_ptr_kinds h'"
and 5: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
have 6: "set |h3 \<turnstile> get_child_nodes xa|\<^sub>r \<inter> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r = {}"
using 1 2 3 4 5
proof -
have "\<forall>h d. \<not> type_wf h \<or> d |\<notin>| document_ptr_kinds h \<or> h \<turnstile> ok get_disconnected_nodes d"
using local.get_disconnected_nodes_ok by satx
then have "h' \<turnstile> ok get_disconnected_nodes xb"
using "4" \<open>type_wf h'\<close> by fastforce
then have f1: "h3 \<turnstile> get_disconnected_nodes xb \<rightarrow>\<^sub>r |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
by (simp add: disconnected_nodes_eq_h3)
have "xa |\<in>| object_ptr_kinds h3"
using "2" object_ptr_kinds_M_eq3_h' by blast
then show ?thesis
using f1 \<open>local.CD.a_distinct_lists h3\<close> 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 \<open>CD.a_distinct_lists h3\<close>
\<open>type_wf h'\<close> disconnected_nodes_eq_h3 CD.distinct_lists_no_parent document_ptr_kinds_eq2_h3
get_disconnected_nodes_ok insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result)
next
case False
then show ?thesis
using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce
qed
qed
moreover have "CD.a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
apply(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_M_eq2_h2 object_ptr_kinds_M_eq2_h3
node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 )[1]
apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified]
object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified] node_ptr_kinds_eq2_h2[simplified]
node_ptr_kinds_eq2_h3[simplified])[1]
apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1]
by (smt Core_DOM_Functions.i_insert_before.insert_before_list_in_set children_eq2_h3 children_h'
- children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1
+ children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 in_set_remove1
object_ptr_kinds_eq_h3 ptr_in_heap select_result_I2)
ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'"
by (simp add: CD.heap_is_wellformed_def)
have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1]
apply(case_tac "aa = owner_document")
apply (metis (no_types, lifting) case_prodI disconnected_nodes_h2 disconnected_nodes_h3
in_set_remove1 mem_Collect_eq node_not_in_disconnected_nodes pair_imageI select_result_I2)
using disconnected_nodes_eq2_h2 apply auto[1]
using node_not_in_disconnected_nodes apply blast
by (metis (no_types, lifting) case_prodI disconnected_nodes_eq2_h2 disconnected_nodes_h2
disconnected_nodes_h3 in_set_remove1 mem_Collect_eq pair_imageI select_result_I2)
have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h'"
by(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3)
have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)"
using \<open>heap_is_wellformed h2\<close>
by(auto simp add: \<open>a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}\<close>
heap_is_wellformed_def \<open>parent_child_rel h2 = parent_child_rel h3\<close> \<open>a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\<close> elim!: acyclic_subset)
then
have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> local.a_ptr_disconnected_node_rel h')"
using \<open>(cast node, ptr) \<notin> (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)\<^sup>*\<close>
by(auto simp add: \<open>a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h'\<close> \<open>a_host_shadow_root_rel h3 =
a_host_shadow_root_rel h'\<close> \<open>parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\<close>)
then
show "heap_is_wellformed h'"
using \<open>heap_is_wellformed h2\<close>
using \<open>heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\<close>
apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def
a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1]
by(auto simp add: object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2
element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2
shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2
tag_name_eq2_h3)
qed
end
interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes
get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node
adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed
parent_child_rel
by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma insert_before_wf_is_l_insert_before_wf [instances]: "l_insert_before_wf Shadow_DOM.heap_is_wellformed
ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs
Shadow_DOM.insert_before Shadow_DOM.get_child_nodes"
apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1]
using insert_before_removes_child apply fast
done
lemma l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf ShadowRootClass.type_wf
ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes
get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs"
apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1]
by (metis Diff_iff Shadow_DOM.i_heap_is_wellformed.heap_is_wellformed_disconnected_nodes_distinct Shadow_DOM.i_remove_child.set_disconnected_nodes_get_disconnected_nodes insert_iff returns_result_eq set_remove1_eq)
interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes
set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel get_parent
get_parent_locs adopt_node adopt_node_locs get_owner_document insert_before insert_before_locs append_child
known_ptrs remove_child remove_child_locs get_ancestors_di get_ancestors_di_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs
remove heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
by(auto simp add: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma insert_before_wf2_is_l_insert_before_wf2 [instances]:
"l_insert_before_wf2 ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before
Shadow_DOM.heap_is_wellformed"
apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1]
using insert_before_child_preserves apply(fast, fast, fast)
done
subsubsection \<open>append\_child\<close>
locale l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma append_child_heap_is_wellformed_preserved:
assumes wellformed: "heap_is_wellformed h"
and append_child: "h \<turnstile> append_child ptr node \<rightarrow>\<^sub>h h'"
and known_ptrs: "known_ptrs h"
and type_wf: "type_wf h"
shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
using assms
by(auto simp add: append_child_def intro: insert_before_child_preserves)
lemma append_child_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
assumes "h \<turnstile> append_child ptr node \<rightarrow>\<^sub>h h'"
assumes "node \<notin> set xs"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ [node]"
proof -
obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where
ancestors: "h \<turnstile> get_ancestors_di ptr \<rightarrow>\<^sub>r ancestors" and
node_not_in_ancestors: "cast node \<notin> set ancestors" and
owner_document: "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document" and
h2: "h \<turnstile> adopt_node owner_document node \<rightarrow>\<^sub>h h2" and
disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disconnected_nodes_h2" and
h3: "h2 \<turnstile> set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> a_insert_node ptr node None \<rightarrow>\<^sub>h h'"
using assms(5)
by(auto simp add: append_child_def insert_before_def a_ensure_pre_insertion_validity_def
elim!: bind_returns_heap_E bind_returns_result_E
bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
split: if_splits option.splits)
have "\<And>parent. |h \<turnstile> get_parent node|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr"
using assms(1) assms(4) assms(6)
by (metis (no_types, lifting) assms(2) assms(3) h2 is_OK_returns_heap_I is_OK_returns_result_E
local.adopt_node_child_in_heap local.get_parent_child_dual local.get_parent_ok
select_result_I2)
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
using get_child_nodes_reads adopt_node_writes h2 assms(4)
apply(rule reads_writes_separate_forwards)
using \<open>\<And>parent. |h \<turnstile> get_parent node|\<^sub>r = Some parent \<Longrightarrow> parent \<noteq> ptr\<close>
apply(auto simp add: adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def)[1]
by (meson local.set_child_nodes_get_child_nodes_different_pointers)
have "h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
using get_child_nodes_reads set_disconnected_nodes_writes h3 \<open>h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs\<close>
apply(rule reads_writes_separate_forwards)
by(auto)
have "ptr |\<in>| object_ptr_kinds h"
by (meson ancestors is_OK_returns_result_I local.get_ancestors_ptr_in_heap)
then
have "known_ptr ptr"
using assms(3)
using local.known_ptrs_known_ptr by blast
have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF adopt_node_writes h2]
using adopt_node_types_preserved \<open>type_wf h\<close>
by(auto simp add: adopt_node_locs_def remove_child_locs_def reflp_def transp_def split: if_splits)
then
have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h3]
using set_disconnected_nodes_types_preserved
by(auto simp add: reflp_def transp_def)
show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs@[node]"
using h'
apply(auto simp add: a_insert_node_def
dest!: bind_returns_heap_E3[rotated, OF \<open>h3 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs\<close>
get_child_nodes_pure, rotated])[1]
using \<open>type_wf h3\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close>
by metis
qed
lemma append_child_for_all_on_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
assumes "h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
assumes "set nodes \<inter> set xs = {}"
assumes "distinct nodes"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs@nodes"
using assms
apply(induct nodes arbitrary: h xs)
apply(simp)
proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a
assume 0: "(\<And>h xs. heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs \<Longrightarrow> h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'
\<Longrightarrow> set nodes \<inter> set xs = {} \<Longrightarrow> h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ nodes)"
and 1: "heap_is_wellformed h"
and 2: "type_wf h"
and 3: "known_ptrs h"
and 4: "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs"
and 5: "h \<turnstile> append_child ptr a \<rightarrow>\<^sub>r ()"
and 6: "h \<turnstile> append_child ptr a \<rightarrow>\<^sub>h h'a"
and 7: "h'a \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
and 8: "a \<notin> set xs"
and 9: "set nodes \<inter> set xs = {}"
and 10: "a \<notin> set nodes"
and 11: "distinct nodes"
then have "h'a \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ [a]"
using append_child_children 6
using "1" "2" "3" "4" "8" by blast
moreover have "heap_is_wellformed h'a" and "type_wf h'a" and "known_ptrs h'a"
using insert_before_child_preserves 1 2 3 6 append_child_def
by metis+
moreover have "set nodes \<inter> set (xs @ [a]) = {}"
using 9 10
by auto
ultimately show "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r xs @ a # nodes"
using 0 7
by fastforce
qed
lemma append_child_for_all_on_no_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
assumes "h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r []"
assumes "h \<turnstile> forall_M (append_child ptr) nodes \<rightarrow>\<^sub>h h'"
assumes "distinct nodes"
shows "h' \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r nodes"
using assms append_child_for_all_on_children
by force
end
interpretation i_append_child_wf?: l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes
get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs
set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed
parent_child_rel get_parent get_parent_locs adopt_node adopt_node_locs get_owner_document insert_before
insert_before_locs append_child known_ptrs remove_child remove_child_locs get_ancestors_di
get_ancestors_di_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document
get_disconnected_document_locs remove heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
by(auto simp add: l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
lemma append_child_wf_is_l_append_child_wf [instances]:
"l_append_child_wf type_wf known_ptr known_ptrs append_child heap_is_wellformed"
apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1]
using append_child_heap_is_wellformed_preserved by fast+
subsubsection \<open>to\_tree\_order\<close>
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 \<open>get\_root\_node\<close>
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 \<open>to\_tree\_order\_si\<close>
locale l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_get_child_nodes +
l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M +
l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M
begin
lemma to_tree_order_si_ok:
assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
and "ptr |\<in>| object_ptr_kinds h"
shows "h \<turnstile> 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 \<open>get\_assigned\_nodes\<close>
lemma forall_M_small_big: "h \<turnstile> forall_M f xs \<rightarrow>\<^sub>h h' \<Longrightarrow> P h \<Longrightarrow>
(\<And>h h' x. x \<in> set xs \<Longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h' \<Longrightarrow> P h \<Longrightarrow> P h') \<Longrightarrow> 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 \<turnstile> assigned_nodes slot \<rightarrow>\<^sub>r nodes"
shows "distinct nodes"
proof -
have "\<And>ptr children. h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> 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 \<turnstile> flatten_dom \<rightarrow>\<^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 \<turnstile> element_ptr_kinds_M \<rightarrow>\<^sub>r element_ptrs" and
tups: "h \<turnstile> map_filter_M2 (\<lambda>element_ptr. do {
tag \<leftarrow> get_tag_name element_ptr;
assigned_nodes \<leftarrow> assigned_nodes element_ptr;
(if tag = ''slot'' \<and> assigned_nodes \<noteq> []
then return (Some (element_ptr, assigned_nodes)) else return None)}) element_ptrs \<rightarrow>\<^sub>r tups"
(is "h \<turnstile> map_filter_M2 ?f element_ptrs \<rightarrow>\<^sub>r tups") and
h2: "h \<turnstile> forall_M (\<lambda>(slot, assigned_nodes). do {
get_child_nodes (cast slot) \<bind> forall_M remove;
forall_M (append_child (cast slot)) assigned_nodes
}) tups \<rightarrow>\<^sub>h h2" and
"h2 \<turnstile> shadow_root_ptr_kinds_M \<rightarrow>\<^sub>r shadow_root_ptrs" and
h': "h2 \<turnstile> forall_M (\<lambda>shadow_root_ptr. do {
host \<leftarrow> get_host shadow_root_ptr;
get_child_nodes (cast host) \<bind> forall_M remove;
get_child_nodes (cast shadow_root_ptr) \<bind> forall_M (append_child (cast host));
remove_shadow_root host
}) shadow_root_ptrs \<rightarrow>\<^sub>h h'"
using \<open>h \<turnstile> flatten_dom \<rightarrow>\<^sub>h h'\<close>
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 \<and> known_ptrs h2 \<and> type_wf h2"
using h2 \<open>heap_is_wellformed h\<close> \<open>known_ptrs h\<close> \<open>type_wf h\<close>
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 = "\<lambda>h. heap_is_wellformed h \<and> known_ptrs h \<and> 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 = "\<lambda>h. heap_is_wellformed h \<and> known_ptrs h \<and> 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 \<open>get\_shadow\_root\_safe\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root_safe :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_safe_locs :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_child_nodes :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
begin
end
subsubsection \<open>create\_element\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
begin
lemma create_element_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_name new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^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 \<turnstile> create_element document_ptr tag \<rightarrow>\<^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 \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_element_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>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="\<lambda>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 \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr
by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| 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 \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_element_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^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:
"\<And>ptr' disc_nodes. ptr' \<noteq> new_element_ptr
\<Longrightarrow> h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> new_element_ptr
\<Longrightarrow> |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h2 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. ptr' \<noteq> new_element_ptr
\<Longrightarrow> h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> new_element_ptr
\<Longrightarrow> |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_tag_name new_element_ptr \<rightarrow>\<^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 \<open>type_wf h\<close> new_element_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3:
"\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. ptr' \<noteq> new_element_ptr
\<Longrightarrow> h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> new_element_ptr
\<Longrightarrow> |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_element_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close>
using \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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:
"\<And>ptr' disc_nodes. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting)
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = 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 \<open>heap_is_wellformed h\<close> 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 \<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> 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
+ by (metis (no_types, opaque_lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> assms(3) assms(4) children_eq_h
- disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I
+ disconnected_nodes_eq2_h document_ptr_kinds_eq_h is_OK_returns_result_I
local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subsetD)
then have "CD.a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "CD.a_all_ptrs_in_heap h'"
- by (smt children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3
+ by (smt (verit) children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3
element_ptr_kinds_commutes h' h2 local.CD.a_all_ptrs_in_heap_def
local.set_disconnected_nodes_get_disconnected_nodes new_element_ptr new_element_ptr_in_heap
- node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2
+ node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2
set_ConsD subset_code(1))
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_element_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_element_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_element_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_element_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "CD.a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: CD.heap_is_wellformed_def heap_is_wellformed_def)
then have "CD.a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_element_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_element_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
by (metis \<open> CD.a_distinct_lists h\<close> \<open>type_wf h2\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> 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) \<open>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 \<notin> set disc_nodes_h3\<close>
\<open> CD.a_distinct_lists h3\<close> \<open>type_wf h'\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>CD.a_all_ptrs_in_heap h\<close>
+ apply(smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>CD.a_all_ptrs_in_heap h\<close>
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'
+ disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
set_disconnected_nodes_get_disconnected_nodes
CD.a_all_ptrs_in_heap_def
select_result_I2 set_ConsD subsetD)
- by (smt NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>CD.a_all_ptrs_in_heap h\<close>
+ by (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>CD.a_all_ptrs_in_heap h\<close>
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'
+ disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
CD.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply -
apply(cases "xb = document_ptr")
apply (metis (no_types, opaque_lifting) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3
\<Longrightarrow> 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 \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
\<open> CD.a_distinct_lists h3\<close> 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" \<open> CD.a_distinct_lists h3\<close> \<open>type_wf h3\<close> 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 \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
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
+ by(metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> 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'
+ document_ptr_kinds_eq_h h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes
select_result_I2)
have "CD.a_heap_is_wellformed h'"
using \<open>CD.a_acyclic_heap h'\<close> \<open>CD.a_all_ptrs_in_heap h'\<close> \<open>CD.a_distinct_lists h'\<close> \<open>CD.a_owner_document_valid h'\<close>
by(simp add: CD.a_heap_is_wellformed_def)
have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h"
using document_ptr_kinds_eq_h
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2"
using document_ptr_kinds_eq_h2
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3"
using document_ptr_kinds_eq_h3
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_eq_h: "\<And>element_ptr shadow_root_opt. element_ptr \<noteq> new_element_ptr
\<Longrightarrow> h \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt = h2 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt"
proof -
fix element_ptr shadow_root_opt
assume "element_ptr \<noteq> new_element_ptr "
have "\<forall>P \<in> get_shadow_root_locs element_ptr. P h h2"
using get_shadow_root_new_element new_element_ptr h2
using \<open>element_ptr \<noteq> new_element_ptr\<close> 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 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt = h2 \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r shadow_root_opt"
by (simp add: preserved_def)
qed
have shadow_root_none: "h2 \<turnstile> get_shadow_root (new_element_ptr) \<rightarrow>\<^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:
"\<And>ptr' children. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>ptr' children. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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
+ by (smt (verit) Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds
+ ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) h2 insort_split
local.get_shadow_root_ok local.shadow_root_same_host new_element_ptr new_element_ptr_not_in_heap
option.distinct(1) returns_result_select_result select_result_I2 shadow_root_none)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2])
then have "a_distinct_lists h'"
by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])
have "a_shadow_root_valid h"
using assms(1)
by (simp add: heap_is_wellformed_def)
then have "a_shadow_root_valid h2"
proof (unfold a_shadow_root_valid_def; safe)
fix shadow_root_ptr
assume "\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h). \<exists>host\<in>fset (element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and> |h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
assume "shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h2)"
obtain previous_host where
"previous_host \<in> fset (element_ptr_kinds h)" and
"|h \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr"
by (metis \<open>local.a_shadow_root_valid h\<close> \<open>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h2)\<close>
local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h)
moreover have "previous_host \<noteq> new_element_ptr"
using calculation(1) h2 new_element_ptr new_element_ptr_not_in_heap by auto
ultimately have "|h2 \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h2 \<turnstile> 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 \<open>previous_host \<noteq> new_element_ptr\<close> \<open>|h \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\<close>
select_result_eq shadow_root_eq_h)
then
show "\<exists>host\<in>fset (element_ptr_kinds h2). |h2 \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h2 \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
by (meson \<open>previous_host \<in> fset (element_ptr_kinds h)\<close> \<open>previous_host \<noteq> new_element_ptr\<close> assms(3)
- local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap notin_fset returns_result_select_result shadow_root_eq_h)
+ local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap returns_result_select_result shadow_root_eq_h)
qed
then have "a_shadow_root_valid h3"
proof (unfold a_shadow_root_valid_def; safe)
fix shadow_root_ptr
assume "\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h2). \<exists>host\<in>fset (element_ptr_kinds h2).
|h2 \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and> |h2 \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
assume "shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h3)"
obtain previous_host where
"previous_host \<in> fset (element_ptr_kinds h2)" and
"|h2 \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h2 \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr"
by (metis \<open>local.a_shadow_root_valid h2\<close> \<open>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h3)\<close>
local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2)
moreover have "previous_host \<noteq> 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 \<turnstile> get_tag_name previous_host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h2 \<turnstile> 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 \<open>previous_host \<noteq> new_element_ptr\<close> \<open>|h2 \<turnstile> get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\<close>
select_result_eq shadow_root_eq_h)
then
show "\<exists>host\<in>fset (element_ptr_kinds h3). |h3 \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h3 \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
- by (smt \<open>previous_host \<in> fset (element_ptr_kinds h2)\<close> \<open>previous_host \<noteq> new_element_ptr\<close> \<open>type_wf h2\<close>
- \<open>type_wf h3\<close> element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok returns_result_eq
+ by (smt (verit) \<open>previous_host \<in> fset (element_ptr_kinds h2)\<close> \<open>previous_host \<noteq> new_element_ptr\<close> \<open>type_wf h2\<close>
+ \<open>type_wf h3\<close> element_ptr_kinds_eq_h2 local.get_shadow_root_ok returns_result_eq
returns_result_select_result shadow_root_eq_h2 tag_name_eq2_h2)
qed
then have "a_shadow_root_valid h'"
apply(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h3 shadow_root_eq_h3
shadow_root_ptr_kinds_eq_h3 tag_name_eq2_h3)[1]
- by (smt \<open>type_wf h3\<close> finite_set_in local.get_shadow_root_ok returns_result_select_result
+ by (smt (z3) \<open>type_wf h3\<close> local.get_shadow_root_ok returns_result_select_result
select_result_I2 shadow_root_eq_h3)
have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2"
apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq_h)[1]
apply (smt assms(3) case_prod_conv h2 image_iff local.get_shadow_root_ok mem_Collect_eq new_element_ptr
new_element_ptr_not_in_heap returns_result_select_result select_result_I2 shadow_root_eq_h)
using shadow_root_none apply auto[1]
apply (metis (no_types, lifting) Collect_cong assms(3) case_prodE case_prodI h2 host_shadow_root_rel_def
i_get_parent_get_host_get_disconnected_document_wf.a_host_shadow_root_rel_shadow_root
local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok new_element_ptr
new_element_ptr_not_in_heap returns_result_select_result select_result_I2 shadow_root_eq_h)
done
have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3"
apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1]
apply (smt Collect_cong \<open>type_wf h2\<close> case_prodE case_prodI element_ptr_kinds_eq_h2 host_shadow_root_rel_def
i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl
local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h2)
by (metis (no_types, lifting) Collect_cong \<open>type_wf h3\<close> case_prodI2 case_prod_conv element_ptr_kinds_eq_h2
host_shadow_root_rel_def i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def
local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h2)
have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'"
apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1]
apply (smt Collect_cong Shadow_DOM.a_host_shadow_root_rel_def \<open>type_wf h3\<close> case_prodD case_prodI2
element_ptr_kinds_eq_h2 i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root local.get_shadow_root_impl
local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h3)
apply (smt Collect_cong \<open>type_wf h'\<close> case_prodE case_prodI element_ptr_kinds_eq_h2 host_shadow_root_rel_def
i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_host_shadow_root_rel_def
local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h3)
done
have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2"
by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h disconnected_nodes_eq2_h)
have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h3"
by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)
have "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_element_ptr # disc_nodes_h3"
using h' local.set_disconnected_nodes_get_disconnected_nodes by auto
have " document_ptr |\<in>| document_ptr_kinds h3"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h document_ptr_kinds_eq_h2)
have "cast new_element_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\<close>
by auto
have "a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_element_ptr)} \<union> a_ptr_disconnected_node_rel h3"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3)[1]
apply(case_tac "aa = document_ptr")
using disc_nodes_h3 h' \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_element_ptr # disc_nodes_h3\<close>
apply(auto)[1]
using disconnected_nodes_eq2_h3 apply auto[1]
using \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_element_ptr # disc_nodes_h3\<close>
using \<open>cast new_element_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r\<close>
using \<open>document_ptr |\<in>| document_ptr_kinds h3\<close> apply auto[1]
apply(case_tac "document_ptr = aa")
using \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\<close> disc_nodes_h3
apply auto[1]
using disconnected_nodes_eq_h3[THEN select_result_eq, simplified] by auto
have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
have "parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h =
parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2"
using \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close>
by auto
have "parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2 =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3"
using \<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close>
\<open>local.a_ptr_disconnected_node_rel h2 = local.a_ptr_disconnected_node_rel h3\<close> \<open>parent_child_rel h2 = parent_child_rel h3\<close>
by auto
have "parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h' =
{(cast document_ptr, cast new_element_ptr)} \<union> parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3"
by (simp add: \<open>local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\<close>
\<open>local.a_ptr_disconnected_node_rel h' = {(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr, cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr)} \<union>
local.a_ptr_disconnected_node_rel h3\<close> \<open>parent_child_rel h3 = parent_child_rel h'\<close>)
have "\<And>a b. (a, b) \<in> parent_child_rel h3 \<Longrightarrow> a \<noteq> cast new_element_ptr"
using CD.parent_child_rel_parent_in_heap \<open>parent_child_rel h = parent_child_rel h2\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> element_ptr_kinds_commutes h2 new_element_ptr
new_element_ptr_not_in_heap node_ptr_kinds_commutes
by blast
moreover
have "\<And>a b. (a, b) \<in> a_host_shadow_root_rel h3 \<Longrightarrow> a \<noteq> cast new_element_ptr"
using shadow_root_eq_h2 shadow_root_none
by(auto simp add: a_host_shadow_root_rel_def)
moreover
have "\<And>a b. (a, b) \<in> a_ptr_disconnected_node_rel h3 \<Longrightarrow> a \<noteq> cast new_element_ptr"
by(auto simp add: a_ptr_disconnected_node_rel_def)
moreover
have "cast new_element_ptr \<notin> {x. (x, cast document_ptr) \<in>
(parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)\<^sup>*}"
by (smt Un_iff \<open>\<And>b a. (a, b) \<in> local.a_host_shadow_root_rel h3 \<Longrightarrow>
a \<noteq> 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\<close> \<open>\<And>b a. (a, b) \<in> local.a_ptr_disconnected_node_rel h3 \<Longrightarrow>
a \<noteq> 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\<close> \<open>\<And>b a. (a, b) \<in> parent_child_rel h3 \<Longrightarrow>
a \<noteq> 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\<close> cast_document_ptr_not_node_ptr(1) converse_rtranclE mem_Collect_eq)
moreover
have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)"
using \<open>acyclic (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> local.a_ptr_disconnected_node_rel h)\<close>
\<open>parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> local.a_ptr_disconnected_node_rel h =
parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 \<union> local.a_ptr_disconnected_node_rel h2\<close>
\<open>parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 \<union> local.a_ptr_disconnected_node_rel h2 =
parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3 \<union> local.a_ptr_disconnected_node_rel h3\<close>
by auto
ultimately have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')"
by(simp add: \<open>parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h' =
{(cast document_ptr, cast new_element_ptr)} \<union> parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3\<close>)
show " heap_is_wellformed h' "
using \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h' \<union> local.a_ptr_disconnected_node_rel h')\<close>
by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \<open>local.CD.a_heap_is_wellformed h'\<close>
\<open>local.a_all_ptrs_in_heap h'\<close> \<open>local.a_distinct_lists h'\<close> \<open>local.a_shadow_root_valid h'\<close>)
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 \<open>create\_character\_data\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) 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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
begin
lemma create_character_data_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^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 \<turnstile> create_character_data document_ptr text \<rightarrow>\<^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 \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>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="\<lambda>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 \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
local.create_character_data_known_ptr by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| 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 \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|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 |\<union>| {|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 |\<union>| {|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="\<lambda>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="\<lambda>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 |\<in>| 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 \<open>type_wf h\<close> document_ptr_kinds_def
by (metis is_OK_returns_result_I)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_character_data_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h:
"\<And>ptr' disc_nodes. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h2 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_character_data_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. document_ptr \<noteq> doc_ptr
\<Longrightarrow> h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. document_ptr \<noteq> doc_ptr
\<Longrightarrow> |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h3:
"\<And>ptr' disc_nodes. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have disc_nodes_document_ptr_h2: "h2 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
then have disc_nodes_document_ptr_h: "h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3"
using disconnected_nodes_eq_h by auto
then have "cast new_character_data_ptr \<notin> set disc_nodes_h3"
using \<open>heap_is_wellformed h\<close> using \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = 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 \<open>heap_is_wellformed h\<close> 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 \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
- apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \<open>parent_child_rel h = parent_child_rel h2\<close>
- children_eq2_h finite_set_in finsert_iff funion_finsert_right CD.parent_child_rel_child
+ apply (metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M \<open>parent_child_rel h = parent_child_rel h2\<close>
+ children_eq2_h finsert_iff funion_finsert_right CD.parent_child_rel_child
CD.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h
select_result_I2 subsetD sup_bot.right_neutral)
- by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap
+ by (metis (no_types, opaque_lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap
\<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> assms(3) assms(4)
- children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I
+ children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h is_OK_returns_result_I
local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subset_code(1))
then have "CD.a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "CD.a_all_ptrs_in_heap h'"
- by (smt character_data_ptr_kinds_commutes character_data_ptr_kinds_eq_h2 children_eq2_h3
+ by (smt (verit) character_data_ptr_kinds_commutes character_data_ptr_kinds_eq_h2 children_eq2_h3
disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 h' h2 local.CD.a_all_ptrs_in_heap_def
local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr new_character_data_ptr_in_heap
- node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))
+ node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_character_data_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
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 "\<And>p. p |\<in>| object_ptr_kinds h2 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h2 \<turnstile> get_child_nodes p|\<^sub>r"
using children_eq2_h
apply(auto simp add: object_ptr_kinds_eq_h)[1]
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close> apply auto[1]
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>)
then have "\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h3 \<turnstile> 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:
"\<And>p. p |\<in>| object_ptr_kinds h' \<Longrightarrow> cast new_character_data_ptr \<notin> set |h' \<turnstile> get_child_nodes p|\<^sub>r"
using object_ptr_kinds_eq_h3 children_eq2_h3 by auto
have "CD.a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h2"
using \<open>h2 \<turnstile> get_child_nodes (cast new_character_data_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(case_tac "x=cast new_character_data_ptr")
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok
local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr
returns_result_select_result)
apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
thm children_eq2_h
using \<open>CD.a_distinct_lists h\<close> \<open>type_wf h2\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
then show "distinct |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes
by (metis (no_types, opaque_lifting) \<open>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 \<notin> set disc_nodes_h3\<close>
\<open>type_wf h2\<close> 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 (\<lambda>document_ptr. |h3 \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h3)))))"
and "x |\<in>| document_ptr_kinds h3"
and "y |\<in>| document_ptr_kinds h3"
and "x \<noteq> y"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
and "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> 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 \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
- by (smt local.CD.a_all_ptrs_in_heap_def \<open>CD.a_all_ptrs_in_heap h\<close> disc_nodes_document_ptr_h2
+ by (smt (verit) local.CD.a_all_ptrs_in_heap_def \<open>CD.a_all_ptrs_in_heap h\<close> 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'
+ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD)
next
fix x xa xb
assume 2: "(\<Union>x\<in>fset (object_ptr_kinds h3). set |h' \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h3). set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 3: "xa |\<in>| object_ptr_kinds h3"
and 4: "x \<in> set |h' \<turnstile> get_child_nodes xa|\<^sub>r"
and 5: "xb |\<in>| document_ptr_kinds h3"
and 6: "x \<in> set |h' \<turnstile> 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
\<open>local.CD.a_distinct_lists h2\<close> \<open>type_wf h'\<close> 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 \<open>local.CD.a_distinct_lists h3\<close> \<open>type_wf h3\<close>
children_eq2_h3 local.get_disconnected_nodes_ok returns_result_select_result)
qed
have "CD.a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
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) \<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
- children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h'
+ by (metis (mono_tags, opaque_lifting) \<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
+ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h h'
l_ptr_kinds_M.ptr_kinds_ptr_kinds_M
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
object_ptr_kinds_M_def
select_result_I2)
have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h"
using document_ptr_kinds_eq_h
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2"
using document_ptr_kinds_eq_h2
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3"
using document_ptr_kinds_eq_h3
by(auto simp add: shadow_root_ptr_kinds_def)
have shadow_root_eq_h: "\<And>character_data_ptr shadow_root_opt.
h \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^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:
"\<And>ptr' children. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>ptr' children. h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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 \<open>type_wf h2\<close> assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host
returns_result_select_result shadow_root_eq_h)
then have "a_distinct_lists h3"
by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2])
then have "a_distinct_lists h'"
by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])
have "a_shadow_root_valid h"
using assms(1)
by (simp add: heap_is_wellformed_def)
then have "a_shadow_root_valid h2"
by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h
select_result_eq[OF shadow_root_eq_h] tag_name_eq2_h)
then have "a_shadow_root_valid h3"
by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2 element_ptr_kinds_eq_h2
select_result_eq[OF shadow_root_eq_h2] tag_name_eq2_h2)
then have "a_shadow_root_valid h'"
by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h3 element_ptr_kinds_eq_h3
select_result_eq[OF shadow_root_eq_h3] tag_name_eq2_h3)
have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h])
have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2])
have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])
have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2"
by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h disconnected_nodes_eq2_h)
have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h3"
by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)
have "h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_character_data_ptr # disc_nodes_h3"
using h' local.set_disconnected_nodes_get_disconnected_nodes by auto
have " document_ptr |\<in>| document_ptr_kinds h3"
by (simp add: \<open>document_ptr |\<in>| document_ptr_kinds h\<close> document_ptr_kinds_eq_h document_ptr_kinds_eq_h2)
have "cast new_character_data_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_character_data_ptr # disc_nodes_h3\<close> by auto
have "a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_character_data_ptr)} \<union> a_ptr_disconnected_node_rel h3"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3)[1]
apply(case_tac "aa = document_ptr")
using disc_nodes_h3 h' \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_character_data_ptr # disc_nodes_h3\<close>
apply(auto)[1]
using disconnected_nodes_eq2_h3 apply auto[1]
using \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_character_data_ptr # disc_nodes_h3\<close>
using \<open>cast new_character_data_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r\<close>
using \<open>document_ptr |\<in>| document_ptr_kinds h3\<close> apply auto[1]
apply(case_tac "document_ptr = aa")
using \<open>h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r cast new_character_data_ptr # disc_nodes_h3\<close> disc_nodes_h3 apply auto[1]
using disconnected_nodes_eq_h3[THEN select_result_eq, simplified] by auto
have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
have "parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h =
parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2"
using \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> by auto
have "parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2 =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3"
using \<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close>
\<open>local.a_ptr_disconnected_node_rel h2 = local.a_ptr_disconnected_node_rel h3\<close> \<open>parent_child_rel h2 = parent_child_rel h3\<close> by auto
have "parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h' =
{(cast document_ptr, cast new_character_data_ptr)} \<union> parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3"
by (simp add: \<open>local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\<close>
\<open>local.a_ptr_disconnected_node_rel h' = {(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr, cast new_character_data_ptr)} \<union>
local.a_ptr_disconnected_node_rel h3\<close> \<open>parent_child_rel h3 = parent_child_rel h'\<close>)
have "\<And>a b. (a, b) \<in> parent_child_rel h3 \<Longrightarrow> a \<noteq> cast new_character_data_ptr"
using CD.parent_child_rel_parent_in_heap \<open>parent_child_rel h = parent_child_rel h2\<close>
\<open>parent_child_rel h2 = parent_child_rel h3\<close> character_data_ptr_kinds_commutes h2 new_character_data_ptr
new_character_data_ptr_not_in_heap node_ptr_kinds_commutes by blast
moreover
have "\<And>a b. (a, b) \<in> a_host_shadow_root_rel h3 \<Longrightarrow> a \<noteq> cast new_character_data_ptr"
using shadow_root_eq_h2
by(auto simp add: a_host_shadow_root_rel_def)
moreover
have "\<And>a b. (a, b) \<in> a_ptr_disconnected_node_rel h3 \<Longrightarrow> a \<noteq> cast new_character_data_ptr"
by(auto simp add: a_ptr_disconnected_node_rel_def)
moreover
have "cast new_character_data_ptr \<notin> {x. (x, cast document_ptr) \<in>
(parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)\<^sup>*}"
by (smt Un_iff calculation(1) calculation(2) calculation(3) cast_document_ptr_not_node_ptr(2)
converse_rtranclE mem_Collect_eq)
moreover
have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)"
using \<open>acyclic (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> local.a_ptr_disconnected_node_rel h)\<close>
\<open>parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> local.a_ptr_disconnected_node_rel h =
parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 \<union> local.a_ptr_disconnected_node_rel h2\<close>
\<open>parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 \<union> local.a_ptr_disconnected_node_rel h2 =
parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3 \<union> local.a_ptr_disconnected_node_rel h3\<close>
by auto
ultimately have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')"
by(simp add: \<open>parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h' =
{(cast document_ptr, cast new_character_data_ptr)} \<union> parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3\<close>)
have "CD.a_heap_is_wellformed h'"
apply(simp add: CD.a_heap_is_wellformed_def)
by (simp add: \<open>local.CD.a_acyclic_heap h'\<close> \<open>local.CD.a_all_ptrs_in_heap h'\<close>
\<open>local.CD.a_distinct_lists h'\<close> \<open>local.CD.a_owner_document_valid h'\<close>)
show " heap_is_wellformed h' "
using \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h' \<union> local.a_ptr_disconnected_node_rel h')\<close>
by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \<open>local.CD.a_heap_is_wellformed h'\<close>
\<open>local.a_all_ptrs_in_heap h'\<close> \<open>local.a_distinct_lists h'\<close> \<open>local.a_shadow_root_valid h'\<close>)
qed
end
subsubsection \<open>create\_document\<close>
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 \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_document :: "((_) heap, exception, (_) document_ptr) prog"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
begin
lemma create_document_preserves_wellformedness:
assumes "heap_is_wellformed h"
and "h \<turnstile> create_document \<rightarrow>\<^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 \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr" and
h': "h \<turnstile> new_document \<rightarrow>\<^sub>h h'"
using assms(2)
apply(simp add: create_document_def)
using new_document_ok by blast
have "new_document_ptr \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have "new_document_ptr |\<notin>| 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 |\<notin>| object_ptr_kinds h"
by simp
have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|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 |\<union>| {|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)
+ by (auto simp add: document_ptr_kinds_def)
have children_eq:
"\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_document_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^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:
"\<And>doc_ptr disc_nodes. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes = h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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) \<open>\<And>thesis. (\<And>new_document_ptr.
\<lbrakk>h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr; h \<turnstile> new_document \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+
then have disconnected_nodes_eq2_h: "\<And>doc_ptr. doc_ptr \<noteq> new_document_ptr
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
using h' local.new_document_no_disconnected_nodes new_document_ptr by blast
have "type_wf h'"
using \<open>type_wf h\<close> new_document_types_preserved h' by blast
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h'"
by (simp add: object_ptr_kinds_eq)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_document_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h'"
and 1: "x \<in> set |h' \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
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 \<open>heap_is_wellformed h\<close> 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
\<open>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 \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close> 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) \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
- children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq
+ apply (metis (no_types, opaque_lifting) \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
+ children_eq2 finsert_iff funion_finsert_right object_ptr_kinds_eq
select_result_I2 subsetD sup_bot.right_neutral)
by (metis (no_types, lifting) \<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close> \<open>type_wf h'\<close>
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 \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h'"
using \<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close>
\<open>h' \<turnstile> get_child_nodes (cast new_document_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: children_eq2[symmetric] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq
document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(auto simp add: dest: distinct_concat_map_E)[1]
apply(auto simp add: dest: distinct_concat_map_E)[1]
using \<open>new_document_ptr |\<notin>| document_ptr_kinds h\<close>
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 \<noteq> y"
assume a2: "x |\<in>| document_ptr_kinds h"
assume a3: "x \<noteq> new_document_ptr"
assume a4: "y |\<in>| document_ptr_kinds h"
assume a5: "y \<noteq> new_document_ptr"
assume a6: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
assume a7: "xa \<in> set |h' \<turnstile> get_disconnected_nodes x|\<^sub>r"
assume a8: "xa \<in> set |h' \<turnstile> get_disconnected_nodes y|\<^sub>r"
have f9: "xa \<in> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using a7 a3 disconnected_nodes_eq2_h by presburger
have f10: "xa \<in> set |h \<turnstile> get_disconnected_nodes y|\<^sub>r"
using a8 a5 disconnected_nodes_eq2_h by presburger
have f11: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))"
using a4 by simp
have "x \<in> 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' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []"
and 1: "h' \<turnstile> 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) \<rightarrow>\<^sub>r []"
and 2: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
and 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
and 4: "(\<Union>x\<in>fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h). set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 5: "x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r"
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
and 7: "xa |\<in>| object_ptr_kinds h"
and 8: "xa \<noteq> 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 |\<in>| document_ptr_kinds h"
and 10: "xb \<noteq> new_document_ptr"
then show "False"
by (metis \<open>CD.a_distinct_lists h\<close> 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 \<open>heap_is_wellformed h\<close> 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 \<open>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 |\<notin>| object_ptr_kinds h\<close>
- children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff
+ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes funion_iff
node_ptr_kinds_eq object_ptr_kinds_eq)
have shadow_root_eq_h: "\<And>character_data_ptr shadow_root_opt. h \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h' \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^sub>r shadow_root_opt"
using get_shadow_root_reads assms(2) get_shadow_root_new_document[rotated, OF h']
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
using local.get_shadow_root_locs_impl new_document_ptr apply blast
using local.get_shadow_root_locs_impl new_document_ptr by blast
have "a_all_ptrs_in_heap h"
by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap)
then have "a_all_ptrs_in_heap h'"
apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_def document_ptr_kinds_eq_h)[1]
using shadow_root_eq_h by fastforce
have "a_distinct_lists h"
using assms(1)
by (simp add: heap_is_wellformed_def)
then have "a_distinct_lists h'"
apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1]
apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1]
by (metis \<open>type_wf h'\<close> 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:
"\<And>ptr' disc_nodes. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes"
using get_tag_name_reads h'
get_tag_name_new_document[OF new_document_ptr h']
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by blast+
have "a_shadow_root_valid h"
using assms(1)
by (simp add: heap_is_wellformed_def)
then have "a_shadow_root_valid h'"
using new_document_is_document_ptr[OF new_document_ptr]
by(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h document_ptr_kinds_eq_h
shadow_root_ptr_kinds_def select_result_eq[OF shadow_root_eq_h] select_result_eq[OF tag_name_eq_h]
is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: option.splits)
have "a_host_shadow_root_rel h = a_host_shadow_root_rel h'"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h])
have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h'"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h disconnected_nodes_eq2_h)[1]
using \<open>new_document_ptr |\<notin>| document_ptr_kinds h\<close> disconnected_nodes_eq2_h apply fastforce
using new_document_disconnected_nodes[OF h' new_document_ptr]
apply(simp add: CD.get_disconnected_nodes_impl CD.a_get_disconnected_nodes_def)
using \<open>new_document_ptr |\<notin>| document_ptr_kinds h\<close> disconnected_nodes_eq2_h apply fastforce
done
have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
moreover
have "parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h =
parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h'"
by (simp add: \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h'\<close>
\<open>local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h'\<close>
\<open>parent_child_rel h = parent_child_rel h'\<close>)
ultimately have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')"
by simp
have "CD.a_heap_is_wellformed h'"
apply(simp add: CD.a_heap_is_wellformed_def)
by (simp add: \<open>local.CD.a_acyclic_heap h'\<close> \<open>local.CD.a_all_ptrs_in_heap h'\<close>
\<open>local.CD.a_distinct_lists h'\<close> \<open>local.CD.a_owner_document_valid h'\<close>)
show "heap_is_wellformed h'"
using CD.heap_is_wellformed_impl \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h' \<union>
local.a_ptr_disconnected_node_rel h')\<close> \<open>local.CD.a_heap_is_wellformed h'\<close> \<open>local.a_all_ptrs_in_heap h'\<close>
\<open>local.a_distinct_lists h'\<close> \<open>local.a_shadow_root_valid h'\<close> 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 \<open>attach\_shadow\_root\<close>
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 \<Rightarrow> bool"
and known_ptrs :: "(_) heap \<Rightarrow> bool"
and type_wf :: "(_) heap \<Rightarrow> bool"
and get_child_nodes :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_child_nodes_locs :: "(_) object_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed :: "(_) heap \<Rightarrow> bool"
and parent_child_rel :: "(_) heap \<Rightarrow> ((_) object_ptr \<times> (_) object_ptr) set"
and set_tag_name :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_shadow_root :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, (_) shadow_root_ptr option) prog"
and get_shadow_root_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_tag_name :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, char list) prog"
and get_tag_name_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and get_host :: "(_) shadow_root_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and get_host_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_disconnected_document :: "(_) node_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and get_disconnected_document_locs :: "((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and create_character_data ::
"(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) 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 \<Rightarrow> bool"
and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \<Rightarrow> bool"
and set_shadow_root :: "(_) element_ptr \<Rightarrow> (_) shadow_root_ptr option \<Rightarrow> (_, unit) dom_prog"
and set_shadow_root_locs :: "(_) element_ptr \<Rightarrow> (_, unit) dom_prog set"
and set_mode :: "(_) shadow_root_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, unit) dom_prog"
and set_mode_locs :: "(_) shadow_root_ptr \<Rightarrow> (_, unit) dom_prog set"
and attach_shadow_root :: "(_) element_ptr \<Rightarrow> shadow_root_mode \<Rightarrow> (_, (_) 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 \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^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 \<turnstile> get_tag_name element_ptr \<rightarrow>\<^sub>r element_tag_name" and
"element_tag_name \<in> safe_shadow_root_element_types" and
prev_shadow_root: "h \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r None" and
h2: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h2" and
new_shadow_root_ptr: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr" and
h3: "h2 \<turnstile> set_mode new_shadow_root_ptr new_mode \<rightarrow>\<^sub>h h3" and
h': "h3 \<turnstile> set_shadow_root element_ptr (Some new_shadow_root_ptr) \<rightarrow>\<^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 \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^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 \<open>element_tag_name \<in> safe_shadow_root_element_types\<close> apply(simp)
using \<open>element_tag_name \<in> safe_shadow_root_element_types\<close> 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 \<notin> set |h \<turnstile> 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 \<notin> set |h \<turnstile> document_ptr_kinds_M|\<^sub>r"
by simp
then have "cast new_shadow_root_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r"
by simp
have object_ptr_kinds_eq_h:
"object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|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 |\<union>| {|cast new_shadow_root_ptr|}"
apply(simp add: document_ptr_kinds_def)
by force
then have shadow_root_ptr_kinds_eq_h:
"shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h |\<union>| {|new_shadow_root_ptr|}"
apply(simp add: shadow_root_ptr_kinds_def)
by force
have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h"
using object_ptr_kinds_eq_h
by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)
have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_mode_writes h3])
using set_mode_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
by (auto simp add: document_ptr_kinds_def)
then have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2"
by (auto simp add: shadow_root_ptr_kinds_def)
have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
using object_ptr_kinds_eq_h2
by(auto simp add: node_ptr_kinds_def)
have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_shadow_root_writes h'])
using set_shadow_root_pointers_preserved
by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
by (auto simp add: document_ptr_kinds_def)
then have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3"
by (auto simp add: shadow_root_ptr_kinds_def)
have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
using object_ptr_kinds_eq_h3
by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_shadow_root_ptr)"
using \<open>h \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^sub>r new_shadow_root_ptr\<close> create_shadow_root_known_ptr
by blast
then
have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> 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 |\<in>| element_ptr_kinds h"
by (meson \<open>h \<turnstile> attach_shadow_root element_ptr new_mode \<rightarrow>\<^sub>r new_shadow_root_ptr\<close> is_OK_returns_result_I
local.attach_shadow_root_element_ptr_in_heap)
have children_eq_h: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> 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 |\<union>| {|cast new_shadow_root_ptr|}"
using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr h2 new_shadow_root_ptr object_ptr_kinds_eq_h by blast
then have node_ptr_kinds_eq_h:
"node_ptr_kinds h2 = node_ptr_kinds h"
apply(simp add: node_ptr_kinds_def)
by force
then have character_data_ptr_kinds_eq_h:
"character_data_ptr_kinds h2 = character_data_ptr_kinds h"
apply(simp add: character_data_ptr_kinds_def)
done
have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h"
using object_ptr_kinds_eq_h
by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)
have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h |\<union>| {|cast new_shadow_root_ptr|}"
using object_ptr_kinds_eq_h
- apply(auto simp add: document_ptr_kinds_def)[1]
- by (metis (full_types) document_ptr_kinds_def document_ptr_kinds_eq_h finsert_fsubset fset.map_comp funion_upper2)
+ 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="\<lambda>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="\<lambda>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: "\<And>(ptr'::(_) object_ptr) children. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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: "\<And>ptr'. ptr' \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> |h \<turnstile> get_child_nodes ptr'|\<^sub>r = |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
using h2 local.new_shadow_root_no_child_nodes new_shadow_root_ptr by auto
have disconnected_nodes_eq_h:
"\<And>doc_ptr disc_nodes. doc_ptr \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> h \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes"
using get_disconnected_nodes_reads h2
get_disconnected_nodes_new_shadow_root_different_pointers[rotated, OF new_shadow_root_ptr h2]
apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
by (metis (no_types, lifting))+
then have disconnected_nodes_eq2_h:
"\<And>doc_ptr. doc_ptr \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> |h \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have "h2 \<turnstile> get_disconnected_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
using h2 new_shadow_root_no_disconnected_nodes new_shadow_root_ptr by auto
have tag_name_eq_h:
"\<And>ptr' disc_nodes. h \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h \<turnstile> get_tag_name ptr'|\<^sub>r = |h2 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have children_eq_h2:
"\<And>ptr' children. h2 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
"\<And>ptr'. |h2 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h2:
"\<And>doc_ptr disc_nodes. h2 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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:
"\<And>doc_ptr. |h2 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h2:
"\<And>ptr' disc_nodes. h2 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h2 \<turnstile> get_tag_name ptr'|\<^sub>r = |h3 \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "type_wf h2"
using \<open>type_wf h\<close> new_shadow_root_types_preserved h2 by blast
then have "type_wf h3"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> 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="\<lambda>h h'. type_wf h \<longrightarrow> 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:
"\<And>ptr' children. h3 \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^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:
" \<And>ptr'. |h3 \<turnstile> get_child_nodes ptr'|\<^sub>r = |h' \<turnstile> get_child_nodes ptr'|\<^sub>r"
using select_result_eq by force
have disconnected_nodes_eq_h3: "\<And>doc_ptr disc_nodes. h3 \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_disconnected_nodes doc_ptr \<rightarrow>\<^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: "\<And>doc_ptr. |h3 \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r = |h' \<turnstile> get_disconnected_nodes doc_ptr|\<^sub>r"
using select_result_eq by force
have tag_name_eq_h3:
"\<And>ptr' disc_nodes. h3 \<turnstile> get_tag_name ptr' \<rightarrow>\<^sub>r disc_nodes
= h' \<turnstile> get_tag_name ptr' \<rightarrow>\<^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: "\<And>ptr'. |h3 \<turnstile> get_tag_name ptr'|\<^sub>r = |h' \<turnstile> get_tag_name ptr'|\<^sub>r"
using select_result_eq by force
have "acyclic (parent_child_rel h)"
using \<open>heap_is_wellformed h\<close>
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 |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h2"
by (simp add: object_ptr_kinds_eq_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h"
and 1: "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
\<open>cast new_shadow_root_ptr \<notin> set |h \<turnstile> object_ptr_kinds_M|\<^sub>r\<close> children_eq2_h)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "a |\<in>| object_ptr_kinds h"
using object_ptr_kinds_eq_h \<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto)
next
fix a x
assume 0: "a |\<in>| object_ptr_kinds h2"
and 1: "x \<in> set |h2 \<turnstile> get_child_nodes a|\<^sub>r"
then show "x \<in> set |h \<turnstile> get_child_nodes a|\<^sub>r"
by (metis (no_types, lifting) \<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
children_eq2_h empty_iff empty_set image_eqI select_result_I2)
qed
also have "\<dots> = parent_child_rel h3"
by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
also have "\<dots> = 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 \<open>heap_is_wellformed h\<close> 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
\<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
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 \<open>known_ptrs h2\<close>
\<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> assms(1) assms(2) l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child
local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap node_ptr_kinds_commutes returns_result_select_result)
by (metis (no_types, opaque_lifting) \<open>h2 \<turnstile> get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
- \<open>type_wf h2\<close> disconnected_nodes_eq_h empty_iff finite_set_in is_OK_returns_result_E is_OK_returns_result_I
+ \<open>type_wf h2\<close> disconnected_nodes_eq_h empty_iff is_OK_returns_result_E is_OK_returns_result_I
local.get_disconnected_nodes_ok local.get_disconnected_nodes_ptr_in_heap node_ptr_kinds_eq_h select_result_I2
set_empty subset_code(1))
then have "CD.a_all_ptrs_in_heap h3"
by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
then have "CD.a_all_ptrs_in_heap h'"
by (simp add: children_eq2_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3
CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3)
have "cast new_shadow_root_ptr |\<notin>| document_ptr_kinds h"
using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_commutes by blast
have "CD.a_distinct_lists h"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_distinct_lists h2"
using \<open>h2 \<turnstile> get_disconnected_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
\<open>h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
apply(auto simp add: children_eq2_h[symmetric] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq_h
document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove)
apply(auto simp add: dest: distinct_concat_map_E)[1]
apply(auto simp add: dest: distinct_concat_map_E)[1]
using \<open>cast new_shadow_root_ptr |\<notin>| document_ptr_kinds h\<close>
apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1]
apply (metis (no_types) DocumentMonad.ptr_kinds_M_ptr_kinds DocumentMonad.ptr_kinds_ptr_kinds_M
concat_map_all_distinct disconnected_nodes_eq2_h select_result_I2)
proof -
fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr"
assume a1: "x \<noteq> y"
assume a2: "x |\<in>| document_ptr_kinds h"
assume a3: "x \<noteq> cast new_shadow_root_ptr"
assume a4: "y |\<in>| document_ptr_kinds h"
assume a5: "y \<noteq> cast new_shadow_root_ptr"
assume a6: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
assume a7: "xa \<in> set |h2 \<turnstile> get_disconnected_nodes x|\<^sub>r"
assume a8: "xa \<in> set |h2 \<turnstile> get_disconnected_nodes y|\<^sub>r"
have f9: "xa \<in> set |h \<turnstile> get_disconnected_nodes x|\<^sub>r"
using a7 a3 disconnected_nodes_eq2_h
by (simp add: disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)
have f10: "xa \<in> set |h \<turnstile> get_disconnected_nodes y|\<^sub>r"
using a8 a5 disconnected_nodes_eq2_h
by (simp add: disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)
have f11: "y \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))"
using a4 by simp
have "x \<in> set (sorted_list_of_set (fset (document_ptr_kinds h)))"
using a2 by simp
then show False
using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1))
next
fix x xa xb
assume 0: "h2 \<turnstile> get_disconnected_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
and 1: "h2 \<turnstile> get_child_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []"
and 2: "distinct (concat (map (\<lambda>ptr. |h \<turnstile> get_child_nodes ptr|\<^sub>r)
(sorted_list_of_set (fset (object_ptr_kinds h)))))"
and 3: "distinct (concat (map (\<lambda>document_ptr. |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)
(sorted_list_of_set (fset (document_ptr_kinds h)))))"
and 4: "(\<Union>x\<in>fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes x|\<^sub>r)
\<inter> (\<Union>x\<in>fset (document_ptr_kinds h). set |h \<turnstile> get_disconnected_nodes x|\<^sub>r) = {}"
and 5: "x \<in> set |h \<turnstile> get_child_nodes xa|\<^sub>r"
and 6: "x \<in> set |h2 \<turnstile> get_disconnected_nodes xb|\<^sub>r"
and 7: "xa |\<in>| object_ptr_kinds h"
and 8: "xa \<noteq> cast new_shadow_root_ptr"
and 9: "xb |\<in>| document_ptr_kinds h"
and 10: "xb \<noteq> cast new_shadow_root_ptr"
then show "False"
by (metis CD.distinct_lists_no_parent \<open>local.CD.a_distinct_lists h\<close> assms(2) disconnected_nodes_eq2_h
local.get_disconnected_nodes_ok returns_result_select_result)
qed
then have "CD.a_distinct_lists h3"
by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
children_eq2_h2 object_ptr_kinds_eq_h2)[1]
then have "CD.a_distinct_lists h'"
by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3
object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)
have "CD.a_owner_document_valid h"
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def)
then have "CD.a_owner_document_valid h'"
(* using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close> *)
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
\<open>cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\<notin>| document_ptr_kinds h\<close> assms(2) assms(3) children_eq2_h
children_eq_h disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
- document_ptr_kinds_commutes finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr
+ document_ptr_kinds_commutes is_OK_returns_result_I local.known_ptrs_known_ptr
returns_result_select_result)
have shadow_root_eq_h: "\<And>character_data_ptr shadow_root_opt. h \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^sub>r shadow_root_opt =
h2 \<turnstile> get_shadow_root character_data_ptr \<rightarrow>\<^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:
"\<And>ptr' children. h2 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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:
"\<And>ptr' children. element_ptr \<noteq> ptr' \<Longrightarrow> h3 \<turnstile> get_shadow_root ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_shadow_root ptr' \<rightarrow>\<^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' \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r Some new_shadow_root_ptr"
using \<open>type_wf h3\<close> 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 \<open>type_wf h3\<close> 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 \<open>type_wf h2\<close> 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 \<open>type_wf h3\<close> 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 "\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h). \<exists>host\<in>fset (element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and> |h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
assume "a_shadow_root_valid h"
assume "shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h')"
show "\<exists>host\<in>fset (element_ptr_kinds h'). |h' \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and>
|h' \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
proof (cases "shadow_root_ptr = new_shadow_root_ptr")
case True
have "element_ptr \<in> fset (element_ptr_kinds h')"
by (simp add: \<open>element_ptr |\<in>| element_ptr_kinds h\<close> element_ptr_kinds_eq_h element_ptr_kinds_eq_h2
element_ptr_kinds_eq_h3)
moreover have "|h' \<turnstile> get_tag_name element_ptr|\<^sub>r \<in> safe_shadow_root_element_types"
by (smt \<open>\<And>thesis. (\<And>h2 h3 new_shadow_root_ptr element_tag_name. \<lbrakk>h \<turnstile> get_tag_name element_ptr \<rightarrow>\<^sub>r element_tag_name;
element_tag_name \<in> safe_shadow_root_element_types; h \<turnstile> get_shadow_root element_ptr \<rightarrow>\<^sub>r None; h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h2;
h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr; h2 \<turnstile> set_mode new_shadow_root_ptr new_mode \<rightarrow>\<^sub>h h3;
h3 \<turnstile> set_shadow_root element_ptr (Some new_shadow_root_ptr) \<rightarrow>\<^sub>h h'\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
select_result_I2 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3)
moreover have "|h' \<turnstile> 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 \<in> fset (element_ptr_kinds h)" and
"|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types" and
"|h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr"
using \<open>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h')\<close>
using \<open>\<forall>shadow_root_ptr\<in>fset (shadow_root_ptr_kinds h). \<exists>host\<in>fset (element_ptr_kinds h).
|h \<turnstile> get_tag_name host|\<^sub>r \<in> safe_shadow_root_element_types \<and> |h \<turnstile> get_shadow_root host|\<^sub>r = Some shadow_root_ptr\<close>
- 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)
+ by (auto simp add: shadow_root_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h)
moreover have "host \<noteq> 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 \<open>type_wf h'\<close> assms(2) finite_set_in local.get_shadow_root_ok returns_result_eq
+ by (smt (verit) \<open>type_wf h'\<close> assms(2) local.get_shadow_root_ok returns_result_eq
returns_result_select_result shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 tag_name_eq2_h
tag_name_eq2_h2 tag_name_eq2_h3)
qed
qed
have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h])
have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3"
by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2])
have "a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \<union> 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 \<noteq> 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) \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close> \<open>type_wf h3\<close> host_shadow_root_rel_def
i_get_parent_get_host_get_disconnected_document_wf.a_host_shadow_root_rel_shadow_root local.get_shadow_root_impl
local.get_shadow_root_ok option.distinct(1) prev_shadow_root returns_result_select_result)
apply (metis (mono_tags, lifting) \<open>\<And>ptr'. (\<And>x. element_ptr \<noteq> ptr') \<Longrightarrow>
|h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r\<close> case_prod_conv image_iff
is_OK_returns_result_I mem_Collect_eq option.inject returns_result_eq returns_result_select_result
shadow_root_h3)
using element_ptr_kinds_eq_h3 local.get_shadow_root_ptr_in_heap shadow_root_h3 apply fastforce
apply (smt Shadow_DOM.a_host_shadow_root_rel_def \<open>\<And>ptr'. (\<And>x. element_ptr \<noteq> ptr') \<Longrightarrow>
|h3 \<turnstile> get_shadow_root ptr'|\<^sub>r = |h' \<turnstile> get_shadow_root ptr'|\<^sub>r\<close> \<open>type_wf h3\<close> case_prodE case_prodI
i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root image_iff local.get_shadow_root_impl
local.get_shadow_root_ok mem_Collect_eq option.distinct(1) prev_shadow_root returns_result_eq
returns_result_select_result shadow_root_eq_h shadow_root_eq_h2)
done
have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2"
apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h)[1]
apply (metis (no_types, lifting) \<open>cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\<notin>| document_ptr_kinds h\<close>
case_prodI disconnected_nodes_eq2_h mem_Collect_eq pair_imageI)
using \<open>h2 \<turnstile> get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
apply auto[1]
apply(case_tac "cast new_shadow_root_ptr \<noteq> aa")
apply (simp add: disconnected_nodes_eq2_h image_iff)
using \<open>cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\<notin>| document_ptr_kinds h\<close>
apply blast
done
have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h3"
by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)
have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h'"
by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3)
have "acyclic (parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h)"
using \<open>heap_is_wellformed h\<close>
by (simp add: heap_is_wellformed_def)
have "parent_child_rel h \<union> a_host_shadow_root_rel h \<union> a_ptr_disconnected_node_rel h =
parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2"
using \<open>local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\<close>
\<open>local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close>
by auto
have "parent_child_rel h2 \<union> a_host_shadow_root_rel h2 \<union> a_ptr_disconnected_node_rel h2 =
parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3"
using \<open>local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\<close>
\<open>local.a_ptr_disconnected_node_rel h2 = local.a_ptr_disconnected_node_rel h3\<close> \<open>parent_child_rel h2 = parent_child_rel h3\<close>
by auto
have "parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h' =
{(cast element_ptr, cast new_shadow_root_ptr)} \<union> parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3"
by (simp add: \<open>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)} \<union> local.a_host_shadow_root_rel h3\<close>
\<open>local.a_ptr_disconnected_node_rel h3 = local.a_ptr_disconnected_node_rel h'\<close> \<open>parent_child_rel h3 = parent_child_rel h'\<close>)
have "\<And>a b. (a, b) \<in> parent_child_rel h3 \<Longrightarrow> a \<noteq> cast new_shadow_root_ptr"
using CD.parent_child_rel_parent_in_heap \<open>cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\<notin>| document_ptr_kinds h\<close>
\<open>parent_child_rel h = parent_child_rel h2\<close> \<open>parent_child_rel h2 = parent_child_rel h3\<close> document_ptr_kinds_commutes
by blast
moreover
have "\<And>a b. (a, b) \<in> a_host_shadow_root_rel h3 \<Longrightarrow> a \<noteq> cast new_shadow_root_ptr"
using shadow_root_eq_h2
by(auto simp add: a_host_shadow_root_rel_def)
moreover
have "\<And>a b. (a, b) \<in> a_ptr_disconnected_node_rel h3 \<Longrightarrow> a \<noteq> cast new_shadow_root_ptr"
using \<open>h2 \<turnstile> get_disconnected_nodes (cast new_shadow_root_ptr) \<rightarrow>\<^sub>r []\<close>
by(auto simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq_h2)
moreover
have "cast new_shadow_root_ptr \<notin> {x. (x, cast element_ptr) \<in>
(parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)\<^sup>*}"
by (smt Un_iff calculation(1) calculation(2) calculation(3) cast_document_ptr_not_node_ptr(2) converse_rtranclE mem_Collect_eq)
moreover
have "acyclic (parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3)"
using \<open>acyclic (parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> local.a_ptr_disconnected_node_rel h)\<close>
\<open>parent_child_rel h \<union> local.a_host_shadow_root_rel h \<union> local.a_ptr_disconnected_node_rel h =
parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 \<union> local.a_ptr_disconnected_node_rel h2\<close>
\<open>parent_child_rel h2 \<union> local.a_host_shadow_root_rel h2 \<union> local.a_ptr_disconnected_node_rel h2 =
parent_child_rel h3 \<union> local.a_host_shadow_root_rel h3 \<union> local.a_ptr_disconnected_node_rel h3\<close> by auto
ultimately have "acyclic (parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h')"
by(simp add: \<open>parent_child_rel h' \<union> a_host_shadow_root_rel h' \<union> a_ptr_disconnected_node_rel h' =
{(cast element_ptr, cast new_shadow_root_ptr)} \<union>
parent_child_rel h3 \<union> a_host_shadow_root_rel h3 \<union> a_ptr_disconnected_node_rel h3\<close>)
have "CD.a_heap_is_wellformed h'"
apply(simp add: CD.a_heap_is_wellformed_def)
by (simp add: \<open>local.CD.a_acyclic_heap h'\<close> \<open>local.CD.a_all_ptrs_in_heap h'\<close> \<open>local.CD.a_distinct_lists h'\<close>
\<open>local.CD.a_owner_document_valid h'\<close>)
show "heap_is_wellformed h' "
using \<open>acyclic (parent_child_rel h' \<union> local.a_host_shadow_root_rel h' \<union> local.a_ptr_disconnected_node_rel h')\<close>
by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \<open>local.CD.a_heap_is_wellformed h'\<close>
\<open>local.a_all_ptrs_in_heap h'\<close> \<open>local.a_distinct_lists h'\<close> \<open>local.a_shadow_root_valid h'\<close>)
qed
end
interpretation l_attach_shadow_root_wf?: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf
get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes
set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name
get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document
get_disconnected_document_locs set_val set_val_locs create_character_data DocumentClass.known_ptr
DocumentClass.type_wf set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root
by(auto simp add: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)
declare l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
end
diff --git a/thys/Shadow_SC_DOM/classes/ShadowRootClass.thy b/thys/Shadow_SC_DOM/classes/ShadowRootClass.thy
--- a/thys/Shadow_SC_DOM/classes/ShadowRootClass.thy
+++ b/thys/Shadow_SC_DOM/classes/ShadowRootClass.thy
@@ -1,499 +1,496 @@
(***********************************************************************************
* 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.
********************************************************************************\***)
section \<open>The Shadow DOM Data Model\<close>
theory ShadowRootClass
imports
Core_SC_DOM.ShadowRootPointer
Core_SC_DOM.DocumentClass
begin
subsection \<open>ShadowRoot\<close>
datatype shadow_root_mode = Open | Closed
record ('node_ptr, 'element_ptr, 'character_data_ptr) RShadowRoot =
"('node_ptr, 'element_ptr, 'character_data_ptr) RDocument" +
nothing :: unit
mode :: shadow_root_mode
child_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'ShadowRoot) ShadowRoot
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'ShadowRoot option) RShadowRoot_scheme"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'ShadowRoot) ShadowRoot"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'Document, 'ShadowRoot) Document
= "('node_ptr, 'element_ptr, 'character_data_ptr, ('node_ptr, 'element_ptr, 'character_data_ptr,
'ShadowRoot option) RShadowRoot_ext + 'Document) Document"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'Document, 'ShadowRoot) Document"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData, 'Document,
'ShadowRoot) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element,
'CharacterData, ('node_ptr, 'element_ptr, 'character_data_ptr, 'ShadowRoot option)
RShadowRoot_ext + 'Document) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData,
'Document, 'ShadowRoot) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData, 'Document, 'ShadowRoot) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, ('node_ptr, 'element_ptr,
'character_data_ptr, 'ShadowRoot option) RShadowRoot_ext + 'Document) heap"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition shadow_root_ptr_kinds :: "(_) heap \<Rightarrow> (_) shadow_root_ptr fset"
where
"shadow_root_ptr_kinds heap =
the |`| (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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_shadow_root_ptr_kind (document_ptr_kinds heap)))"
lemma shadow_root_ptr_kinds_simp [simp]:
"shadow_root_ptr_kinds (Heap (fmupd (cast shadow_root_ptr) shadow_root (the_heap h))) =
{|shadow_root_ptr|} |\<union>| shadow_root_ptr_kinds h"
- apply(auto simp add: shadow_root_ptr_kinds_def)[1]
- by force
+ by (auto simp add: shadow_root_ptr_kinds_def)
definition shadow_root_ptrs :: "(_) heap \<Rightarrow> (_) shadow_root_ptr fset"
where
"shadow_root_ptrs heap = ffilter is_shadow_root_ptr (shadow_root_ptr_kinds heap)"
definition cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) Document \<Rightarrow> (_) ShadowRoot option"
where
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t document = (case RDocument.more document of Some (Inl shadow_root) \<Rightarrow>
Some (RDocument.extend (RDocument.truncate document) shadow_root) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) Object \<Rightarrow> (_) ShadowRoot option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t obj \<equiv> (case cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj of
Some document \<Rightarrow> cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t document | None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
definition cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) ShadowRoot \<Rightarrow> (_) Document"
where
"cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t shadow_root = RDocument.extend (RDocument.truncate shadow_root)
(Some (Inl (RDocument.more shadow_root)))"
adhoc_overloading cast cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
abbreviation cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) ShadowRoot \<Rightarrow> (_) Object"
where
"cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr)"
adhoc_overloading cast cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
consts is_shadow_root_kind :: 'a
definition is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Document \<Rightarrow> bool"
where
"is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<longleftrightarrow> cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr \<noteq> None"
adhoc_overloading is_shadow_root_kind is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
lemmas is_shadow_root_kind_def = is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
abbreviation is_shadow_root_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) Object \<Rightarrow> bool"
where
"is_shadow_root_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr \<noteq> None"
adhoc_overloading is_shadow_root_kind is_shadow_root_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
definition get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) shadow_root_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) ShadowRoot option"
where
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Option.bind (get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast shadow_root_ptr) h) cast"
adhoc_overloading get get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
locale l_type_wf_def\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (DocumentClass.type_wf h \<and> (\<forall>shadow_root_ptr \<in> fset (shadow_root_ptr_kinds h)
.get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
locale l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = l_type_wf type_wf for type_wf :: "((_) heap \<Rightarrow> bool)" +
assumes type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t: "type_wf h \<Longrightarrow> ShadowRootClass.type_wf h"
sublocale l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t \<subseteq> l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
apply(unfold_locales)
using DocumentClass.a_type_wf_def
by (meson ShadowRootClass.a_type_wf_def l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_axioms l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
locale l_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas = l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
sublocale l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales
lemma get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf:
assumes "type_wf h"
shows "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h \<longleftrightarrow> get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h \<noteq> None"
proof
assume "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
then
show "get shadow_root_ptr h \<noteq> None"
using l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_axioms[unfolded l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def type_wf_defs] assms
- by (meson notin_fset)
+ by meson
next
assume "get shadow_root_ptr h \<noteq> None"
then
show "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
apply(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def shadow_root_ptr_kinds_def
document_ptr_kinds_def object_ptr_kinds_def
split: Option.bind_splits)[1]
- by (metis comp_eq_dest_lhs document_ptr_casts_commute2 document_ptr_document_ptr_cast
- ffmember_filter fimageI fmdomI is_shadow_root_ptr_kind_cast option.sel
+ by (metis (no_types, lifting) IntI document_ptr_casts_commute2 document_ptr_document_ptr_cast
+ fmdomI image_iff is_shadow_root_ptr_kind_cast mem_Collect_eq option.sel
shadow_root_ptr_casts_commute2)
qed
end
global_interpretation l_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas type_wf
by unfold_locales
definition put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) shadow_root_ptr \<Rightarrow> (_) ShadowRoot \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root = put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast shadow_root_ptr) (cast shadow_root)"
adhoc_overloading put put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
lemma put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap:
assumes "put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h = h'"
shows "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h'"
using assms
unfolding put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def shadow_root_ptr_kinds_def
by (metis ffmember_filter fimage_eqI is_shadow_root_ptr_kind_cast option.sel
put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap shadow_root_ptr_casts_commute2)
lemma put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_put_ptrs:
assumes "put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast shadow_root_ptr|}"
using assms
by (simp add: put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs)
lemma cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inject [simp]:
"cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t x = cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t y \<longleftrightarrow> x = y"
apply(auto simp add: cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RObject.extend_def RDocument.extend_def
RDocument.truncate_def)[1]
by (metis RDocument.select_convs(5) RShadowRoot.surjective old.unit.exhaust)
lemma cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_none [simp]:
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t document = None \<longleftrightarrow> \<not> (\<exists>shadow_root. cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t shadow_root = document)"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RObject.extend_def
RDocument.extend_def RDocument.truncate_def
split: sum.splits option.splits)[1]
by (metis (mono_tags, lifting) RDocument.select_convs(2) RDocument.select_convs(3)
RDocument.select_convs(4) RDocument.select_convs(5) RDocument.surjective old.unit.exhaust)
lemma cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_some [simp]:
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t document = Some shadow_root \<longleftrightarrow> cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t shadow_root = document"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RObject.extend_def
RDocument.extend_def RDocument.truncate_def
split: sum.splits option.splits)[1]
by (metis RDocument.select_convs(5) RShadowRoot.surjective old.unit.exhaust)
lemma cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_inv [simp]:
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t (cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t shadow_root) = Some shadow_root"
by simp
lemma is_shadow_root_kind_doctype [simp]:
"is_shadow_root_kind x \<longleftrightarrow> is_shadow_root_kind (doctype_update (\<lambda>_. v) x)"
apply(auto simp add: is_shadow_root_kind_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def
RDocument.truncate_def split: option.splits)[1]
apply (metis RDocument.ext_inject RDocument.select_convs(3) RDocument.surjective RObject.ext_inject)
by (smt RDocument.select_convs(2) RDocument.select_convs(3) RDocument.select_convs(4)
RDocument.select_convs(5) RDocument.surjective RDocument.update_convs(2) old.unit.exhaust)
lemma is_shadow_root_kind_document_element [simp]:
"is_shadow_root_kind x \<longleftrightarrow> is_shadow_root_kind (document_element_update (\<lambda>_. v) x)"
apply(auto simp add: is_shadow_root_kind_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def
RDocument.truncate_def split: option.splits)[1]
apply (metis RDocument.ext_inject RDocument.select_convs(3) RDocument.surjective RObject.ext_inject)
by (metis (no_types, lifting) RDocument.ext_inject RDocument.surjective RDocument.update_convs(3)
RObject.select_convs(1) RObject.select_convs(2))
lemma is_shadow_root_kind_disconnected_nodes [simp]:
"is_shadow_root_kind x \<longleftrightarrow> is_shadow_root_kind (disconnected_nodes_update (\<lambda>_. v) x)"
apply(auto simp add: is_shadow_root_kind_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def
RDocument.truncate_def split: option.splits)[1]
apply (metis RDocument.ext_inject RDocument.select_convs(3) RDocument.surjective RObject.ext_inject)
by (metis (no_types, lifting) RDocument.ext_inject RDocument.surjective RDocument.update_convs(4)
RObject.select_convs(1) RObject.select_convs(2))
lemma shadow_root_ptr_kinds_commutes [simp]:
"cast shadow_root_ptr |\<in>| document_ptr_kinds h \<longleftrightarrow> shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
apply(auto simp add: document_ptr_kinds_def shadow_root_ptr_kinds_def)[1]
- by (metis (no_types, lifting) shadow_root_ptr_casts_commute2 ffmember_filter fimage_eqI
- fset.map_comp is_shadow_root_ptr_kind_none document_ptr_casts_commute3
- document_ptr_kinds_commutes document_ptr_kinds_def option.sel option.simps(3))
+ by (metis Int_iff imageI is_shadow_root_ptr_kind_cast mem_Collect_eq option.sel
+ shadow_root_ptr_casts_commute2)
lemma get_shadow_root_ptr_simp1 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h) = Some shadow_root"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_shadow_root_ptr_simp2 [simp]:
"shadow_root_ptr \<noteq> shadow_root_ptr'
\<Longrightarrow> get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' shadow_root h) =
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_shadow_root_ptr_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_shadow_root_ptr_simp4 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_shadow_root_ptr_simp5 [simp]:
"get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr f h) = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_shadow_root_ptr_simp6 [simp]:
"get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr f h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_shadow_root_put_document [simp]:
"cast shadow_root_ptr \<noteq> document_ptr
\<Longrightarrow> get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h) = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma get_document_put_shadow_root [simp]:
"document_ptr \<noteq> cast shadow_root_ptr
\<Longrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr shadow_root h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_element_ptr, h')"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h = (new_character_data_ptr, h')"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h = (new_document_ptr, h')"
assumes "cast ptr \<noteq> new_document_ptr"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
abbreviation "create_shadow_root_obj mode_arg child_nodes_arg
\<equiv> \<lparr> RObject.nothing = (), RDocument.nothing = (), RDocument.doctype = ''html'',
RDocument.document_element = None, RDocument.disconnected_nodes = [], RShadowRoot.nothing = (),
mode = mode_arg, RShadowRoot.child_nodes = child_nodes_arg, \<dots> = None \<rparr>"
definition new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_)heap \<Rightarrow> ((_) shadow_root_ptr \<times> (_) heap)"
where
"new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (let new_shadow_root_ptr = shadow_root_ptr.Ref
(Suc (fMax (shadow_root_ptr.the_ref |`| (shadow_root_ptrs h)))) in
(new_shadow_root_ptr, put new_shadow_root_ptr (create_shadow_root_obj Open []) h))"
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "new_shadow_root_ptr |\<in>| shadow_root_ptr_kinds h'"
using assms
unfolding new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def
using put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap by blast
lemma new_shadow_root_ptr_new: "shadow_root_ptr.Ref
(Suc (fMax (finsert 0 (shadow_root_ptr.the_ref |`| shadow_root_ptrs h)))) |\<notin>| shadow_root_ptrs h"
by (metis Suc_n_not_le_n shadow_root_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2
set_finsert)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_not_in_heap:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "new_shadow_root_ptr |\<notin>| shadow_root_ptr_kinds h"
using assms
apply(auto simp add: Let_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def shadow_root_ptr_kinds_def)[1]
by (metis Suc_n_not_le_n fMax_ge ffmember_filter fimageI is_shadow_root_ptr_ref
shadow_root_ptr.disc(1) shadow_root_ptr.exhaust shadow_root_ptr.is_Ref_def shadow_root_ptr.sel(1)
shadow_root_ptr.simps(4) shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes
shadow_root_ptrs_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_shadow_root_ptr|}"
using assms
by (metis Pair_inject new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_put_ptrs)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_is_shadow_root_ptr:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "is_shadow_root_ptr new_shadow_root_ptr"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
assumes "ptr \<noteq> cast new_shadow_root_ptr"
shows "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr h'"
using assms
apply(simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
shows "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
assumes "ptr \<noteq> cast new_shadow_root_ptr"
shows "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr h'"
using assms
apply(simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t [simp]:
assumes "new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h = (new_shadow_root_ptr, h')"
assumes "ptr \<noteq> new_shadow_root_ptr"
shows "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h'"
using assms
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def)
locale l_known_ptr\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_shadow_root_ptr ptr)"
lemma known_ptr_not_shadow_root_ptr: "a_known_ptr ptr \<Longrightarrow> \<not>is_shadow_root_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
lemma known_ptr_new_shadow_root_ptr: "a_known_ptr ptr \<Longrightarrow> \<not>known_ptr ptr \<Longrightarrow> is_shadow_root_ptr ptr"
using l_known_ptr\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t.known_ptr_not_shadow_root_ptr by blast
end
global_interpretation l_known_ptr\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t defines known_ptr = a_known_ptr .
lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
- apply(simp add: a_known_ptrs_def)
- using notin_fset by fastforce
+ by (simp add: a_known_ptrs_def)
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs [instances]: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
lemma known_ptrs_implies: "DocumentClass.known_ptrs h \<Longrightarrow> ShadowRootClass.known_ptrs h"
by(auto simp add: DocumentClass.known_ptrs_defs DocumentClass.known_ptr_defs
ShadowRootClass.known_ptrs_defs ShadowRootClass.known_ptr_defs)
definition delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t :: "(_) shadow_root_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) heap option" where
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr = delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast shadow_root_ptr)"
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_pointer_removed:
assumes "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = Some h'"
shows "ptr |\<notin>| shadow_root_ptr_kinds h'"
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_removed delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def shadow_root_ptr_kinds_def
document_ptr_kinds_def split: if_splits)
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_pointer_ptr_in_heap:
assumes "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h = Some h'"
shows "ptr |\<in>| shadow_root_ptr_kinds h"
using assms
apply(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_ptr_in_heap delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def split: if_splits)[1]
using delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_ptr_in_heap
by fastforce
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok:
assumes "ptr |\<in>| shadow_root_ptr_kinds h"
shows "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr h \<noteq> None"
using assms
by (simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemma shadow_root_delete_get_1 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h' = None"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: if_splits)
lemma shadow_root_delete_get_2 [simp]:
"shadow_root_ptr \<noteq> shadow_root_ptr' \<Longrightarrow> delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' h = Some h' \<Longrightarrow>
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h' = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: if_splits)
lemma shadow_root_delete_get_3 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> object_ptr \<noteq> cast shadow_root_ptr \<Longrightarrow>
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h' = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
lemma shadow_root_delete_get_4 [simp]: "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow>
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h' = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma shadow_root_delete_get_5 [simp]: "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow>
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h' = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma shadow_root_delete_get_6 [simp]: "delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow>
get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h' = get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr h"
by(simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
lemma shadow_root_delete_get_7 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> document_ptr \<noteq> cast shadow_root_ptr \<Longrightarrow>
get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h' = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma shadow_root_delete_get_8 [simp]:
"delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h = Some h' \<Longrightarrow> shadow_root_ptr' \<noteq> shadow_root_ptr \<Longrightarrow>
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' h' = get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' h"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
end
diff --git a/thys/Shadow_SC_DOM/monads/ShadowRootMonad.thy b/thys/Shadow_SC_DOM/monads/ShadowRootMonad.thy
--- a/thys/Shadow_SC_DOM/monads/ShadowRootMonad.thy
+++ b/thys/Shadow_SC_DOM/monads/ShadowRootMonad.thy
@@ -1,1053 +1,1058 @@
(***********************************************************************************
* 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\<open>Shadow Root Monad\<close>
theory ShadowRootMonad
imports
"Core_SC_DOM.DocumentMonad"
"../classes/ShadowRootClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot, 'result) dom_prog"
global_interpretation l_ptr_kinds_M shadow_root_ptr_kinds defines shadow_root_ptr_kinds_M = a_ptr_kinds_M .
lemmas shadow_root_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma shadow_root_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> shadow_root_ptr_kinds_M|\<^sub>r = |h' \<turnstile> shadow_root_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: shadow_root_ptr_kinds_M_defs document_ptr_kinds_def object_ptr_kinds_M_defs
shadow_root_ptr_kinds_def)
global_interpretation l_dummy defines get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = "l_get_M.a_get_M get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t type_wf shadow_root_ptr_kinds"
- apply(simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf l_get_M_def)
- apply(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def shadow_root_ptr_kinds_def)[1]
- by (metis (no_types, lifting) DocumentMonad.l_get_M_axioms bind_eq_None_conv fset.map_comp
- l_get_M_def option.discI shadow_root_ptr_kinds_commutes shadow_root_ptr_kinds_def)
+proof -
+ have "\<forall>ptr h. (\<exists>y. get ptr h = Some y) \<longrightarrow> ptr |\<in>| shadow_root_ptr_kinds h"
+ apply(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def shadow_root_ptr_kinds_def bind_eq_Some_conv image_iff Bex_def)
+ by (metis (no_types, opaque_lifting) DocumentMonad.l_get_M_axioms is_shadow_root_ptr_kind_cast l_get_M_def
+ option.sel option.simps(3) shadow_root_ptr_casts_commute2)
+ thus ?thesis
+ by (simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf l_get_M_def)
+qed
+
lemmas get_M_defs = get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
locale l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas = l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
sublocale l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_get_M get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t type_wf shadow_root_ptr_kinds
apply(unfold_locales)
apply (simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf local.type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t)
by (meson ShadowRootMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok = get_M_ok[folded get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def]
lemmas 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 = get_M_ptr_in_heap[folded get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def]
end
global_interpretation l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf shadow_root_ptr_kinds get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t rewrites
"a_get_M = get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t" defines put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
locale l_put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas = l_type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
begin
sublocale l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas by unfold_locales
interpretation l_put_M type_wf shadow_root_ptr_kinds get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t
apply(unfold_locales)
apply (simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_type_wf local.type_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t)
by (meson ShadowRootMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok = put_M_ok[folded put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def]
end
global_interpretation l_put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas type_wf by unfold_locales
lemma shadow_root_put_get [simp]: "h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Mshadow_root_preserved1 [simp]:
"shadow_root_ptr \<noteq> shadow_root_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma shadow_root_put_get_preserved [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
apply(cases "shadow_root_ptr = shadow_root_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved2 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs NodeMonad.get_M_defs
put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved3 [simp]:
"cast shadow_root_ptr \<noteq> document_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def DocumentMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved4 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
apply(cases "cast shadow_root_ptr \<noteq> document_ptr")[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved3a [simp]:
"cast shadow_root_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved4a [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast shadow_root_ptr \<noteq> object_ptr")[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved5 [simp]:
"cast shadow_root_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved7 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved8 [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_shadow_root_put_M_document_different_pointers [simp]:
"cast shadow_root_ptr \<noteq> document_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
by(auto simp add: DocumentMonad.put_M_defs get_M_defs DocumentMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_shadow_root_put_M_document [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. is_shadow_root_kind x \<longleftrightarrow> is_shadow_root_kind (setter (\<lambda>_. v) x))
\<Longrightarrow> (\<And>x. getter (the (cast (((setter (\<lambda>_. v) (cast x)))))) = getter ((x)))
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr getter) h h'"
apply(cases "cast shadow_root_ptr \<noteq> document_ptr ")
apply(auto simp add: preserved_def is_shadow_root_kind_def DocumentMonad.put_M_defs
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get_M_defs DocumentMonad.get_M_defs split: option.splits)[1]
apply(auto simp add: preserved_def is_shadow_root_kind_def DocumentMonad.put_M_defs
get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get_M_defs DocumentMonad.get_M_defs split: option.splits)[1]
apply(metis cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_inv option.sel)
apply(metis cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_inv option.sel)
done
lemma get_M_document_put_M_shadow_root_different_pointers [simp]:
"document_ptr \<noteq> cast shadow_root_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs DocumentMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_document_put_M_shadow_root [simp]:
"h \<turnstile> put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. is_shadow_root_kind x \<Longrightarrow> getter ((cast (((setter (\<lambda>_. v) (the (cast x))))))) = getter ((x)))
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
apply(cases "document_ptr \<noteq> cast shadow_root_ptr ")
apply(auto simp add: preserved_def is_document_kind_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put_M_defs
get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def DocumentMonad.get_M_defs ShadowRootMonad.get_M_defs
split: option.splits Option.bind_splits)[1]
apply(auto simp add: preserved_def is_document_kind_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put_M_defs
get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def DocumentMonad.get_M_defs ShadowRootMonad.get_M_defs
split: option.splits Option.bind_splits)[1]
using is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def apply force
by (metis cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_inv is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.distinct(1) option.sel)
lemma cast_shadow_root_child_nodes_document_disconnected_nodes [simp]:
"RShadowRoot.child_nodes (the (cast (cast x\<lparr>disconnected_nodes := y\<rparr>))) = RShadowRoot.child_nodes x"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def RDocument.truncate_def
split: option.splits)[1]
by (metis RDocument.ext_inject RDocument.surjective RObject.ext_inject RShadowRoot.ext_inject
RShadowRoot.surjective)
lemma cast_shadow_root_child_nodes_document_doctype [simp]:
"RShadowRoot.child_nodes (the (cast (cast x\<lparr>doctype := y\<rparr>))) = RShadowRoot.child_nodes x"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def RDocument.truncate_def
split: option.splits)[1]
by (metis RDocument.ext_inject RDocument.surjective RObject.ext_inject RShadowRoot.ext_inject RShadowRoot.surjective)
lemma cast_shadow_root_child_nodes_document_document_element [simp]:
"RShadowRoot.child_nodes (the (cast (cast x\<lparr>document_element := y\<rparr>))) = RShadowRoot.child_nodes x"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def RDocument.truncate_def
split: option.splits)[1]
by (metis RDocument.ext_inject RDocument.surjective RObject.ext_inject RShadowRoot.ext_inject
RShadowRoot.surjective)
lemma cast_shadow_root_mode_document_disconnected_nodes [simp]:
"RShadowRoot.mode (the (cast (cast x\<lparr>disconnected_nodes := y\<rparr>))) = RShadowRoot.mode x"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def
RDocument.truncate_def split: option.splits)[1]
by (metis RDocument.ext_inject RDocument.surjective RObject.ext_inject RShadowRoot.ext_inject
RShadowRoot.surjective)
lemma cast_shadow_root_mode_document_doctype [simp]:
"RShadowRoot.mode (the (cast (cast x\<lparr>doctype := y\<rparr>))) = RShadowRoot.mode x"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def RDocument.truncate_def
split: option.splits)[1]
by (metis RDocument.ext_inject RDocument.surjective RObject.ext_inject RShadowRoot.ext_inject RShadowRoot.surjective)
lemma cast_shadow_root_mode_document_document_element [simp]:
"RShadowRoot.mode (the (cast (cast x\<lparr>document_element := y\<rparr>))) = RShadowRoot.mode x"
apply(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def RDocument.truncate_def
split: option.splits)[1]
by (metis RDocument.ext_inject RDocument.surjective RObject.ext_inject RShadowRoot.ext_inject RShadowRoot.surjective)
lemma cast_document_disconnected_nodes_shadow_root_child_nodes [simp]:
"is_shadow_root_kind x \<Longrightarrow>
disconnected_nodes (cast (the (cast x)\<lparr>RShadowRoot.child_nodes := arg\<rparr>)) = disconnected_nodes x"
by(auto simp add: is_shadow_root_kind_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
RDocument.extend_def RDocument.truncate_def split: option.splits sum.splits)
lemma cast_document_doctype_shadow_root_child_nodes [simp]:
"is_shadow_root_kind x \<Longrightarrow> doctype (cast (the (cast x)\<lparr>RShadowRoot.child_nodes := arg\<rparr>)) = doctype x"
by(auto simp add: is_shadow_root_kind_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
RDocument.extend_def RDocument.truncate_def split: option.splits sum.splits)
lemma cast_document_document_element_shadow_root_child_nodes [simp]:
"is_shadow_root_kind x \<Longrightarrow>
document_element (cast (the (cast x)\<lparr>RShadowRoot.child_nodes := arg\<rparr>)) = document_element x"
by(auto simp add: is_shadow_root_kind_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
RDocument.extend_def RDocument.truncate_def split: option.splits sum.splits)
lemma cast_document_disconnected_nodes_shadow_root_mode [simp]:
"is_shadow_root_kind x \<Longrightarrow>
disconnected_nodes (cast (the (cast x)\<lparr>RShadowRoot.mode := arg\<rparr>)) = disconnected_nodes x"
by(auto simp add: is_shadow_root_kind_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
RDocument.extend_def RDocument.truncate_def split: option.splits sum.splits)
lemma cast_document_doctype_shadow_root_mode [simp]:
"is_shadow_root_kind x \<Longrightarrow>
doctype (cast (the (cast x)\<lparr>RShadowRoot.mode := arg\<rparr>)) = doctype x"
by(auto simp add: is_shadow_root_kind_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
RDocument.extend_def RDocument.truncate_def split: option.splits sum.splits)
lemma cast_document_document_element_shadow_root_mode [simp]:
"is_shadow_root_kind x \<Longrightarrow>
document_element (cast (the (cast x)\<lparr>RShadowRoot.mode := arg\<rparr>)) = document_element x"
by(auto simp add: is_shadow_root_kind_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
RDocument.extend_def RDocument.truncate_def split: option.splits sum.splits)
lemma new_element_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> cast ptr \<noteq> new_document_ptr \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
definition 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 \<Rightarrow> (_, unit) dom_prog" where
"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 = do {
h \<leftarrow> get_heap;
(case delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr h of
Some h \<Rightarrow> return_heap h |
None \<Rightarrow> error HierarchyRequestError)
}"
adhoc_overloading delete_M delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ok [simp]:
assumes "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
shows "h \<turnstile> ok (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)"
using assms
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: prod.splits)
lemma delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
shows "shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
using assms
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
lemma 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:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
shows "shadow_root_ptr |\<notin>| shadow_root_ptr_kinds h'"
using assms
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def shadow_root_ptr_kinds_def
document_ptr_kinds_def object_ptr_kinds_def split: if_splits)
lemma delete_shadow_root_pointers:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
shows "object_ptr_kinds h = object_ptr_kinds h' |\<union>| {|cast shadow_root_ptr|}"
using assms
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def shadow_root_ptr_kinds_def
document_ptr_kinds_def object_ptr_kinds_def split: if_splits)
lemma delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow> ptr \<noteq> cast shadow_root_ptr \<Longrightarrow>
preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def NodeMonad.get_M_defs ObjectMonad.get_M_defs
preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def ElementMonad.get_M_defs NodeMonad.get_M_defs
ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def CharacterDataMonad.get_M_defs NodeMonad.get_M_defs
ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"cast shadow_root_ptr \<noteq> ptr \<Longrightarrow> h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def DocumentMonad.get_M_defs ObjectMonad.get_M_defs
preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma 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:
"h \<turnstile> 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 \<rightarrow>\<^sub>h h' \<Longrightarrow> shadow_root_ptr \<noteq> shadow_root_ptr' \<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr' getter) h h'"
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get_M_defs ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
subsection \<open>new\_M\<close>
definition new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M :: "(_, (_) shadow_root_ptr) dom_prog"
where
"new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M = do {
h \<leftarrow> get_heap;
(new_ptr, h') \<leftarrow> return (new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t h);
return_heap h';
return new_ptr
}"
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ok [simp]:
"h \<turnstile> ok new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def split: prod.splits)
lemma 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:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "new_shadow_root_ptr |\<in>| shadow_root_ptr_kinds h'"
using assms
unfolding new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def
by(auto simp add: 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 put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma 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:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "new_shadow_root_ptr |\<notin>| shadow_root_ptr_kinds h"
using assms new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_not_in_heap
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
and "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_shadow_root_ptr|}"
using assms new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma 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:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "is_shadow_root_ptr new_shadow_root_ptr"
using assms new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_is_shadow_root_ptr
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def elim!: bind_returns_result_E split: prod.splits)
lemma new_shadow_root_mode:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "h' \<turnstile> get_M new_shadow_root_ptr mode \<rightarrow>\<^sub>r Open"
using assms
by(auto simp add: get_M_defs 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
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_children:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "h' \<turnstile> get_M new_shadow_root_ptr child_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs 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
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_disconnected_nodes:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'"
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "h' \<turnstile> get_M (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr) disconnected_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: DocumentMonad.get_M_defs put_M_defs put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_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
cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def RDocument.extend_def RDocument.truncate_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> ptr \<noteq> cast new_shadow_root_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def CharacterDataMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> ptr \<noteq> cast new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def DocumentMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma 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:
"h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr \<Longrightarrow> ptr \<noteq> new_shadow_root_ptr
\<Longrightarrow> preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t ptr getter) h h'"
by(auto simp add: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection \<open>modified heaps\<close>
lemma shadow_root_get_put_1 [simp]: "get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) =
(if ptr = cast shadow_root_ptr then cast obj else get shadow_root_ptr h)"
by(auto simp add: get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def split: option.splits Option.bind_splits)
lemma shadow_root_ptr_kinds_new [simp]: "shadow_root_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) =
shadow_root_ptr_kinds h |\<union>| (if is_shadow_root_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: shadow_root_ptr_kinds_def is_document_ptr_kind_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "DocumentClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "is_shadow_root_ptr_kind ptr \<Longrightarrow> is_shadow_root_kind obj"
shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
using assms
by(auto simp add: type_wf_defs is_shadow_root_kind_def split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
- by (metis (no_types, lifting) DocumentMonad.type_wf_put_ptr_not_in_heap_E ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
+ by (metis (no_types, opaque_lifting) DocumentMonad.type_wf_put_ptr_not_in_heap_E ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
ObjectMonad.type_wf_put_ptr_not_in_heap_E ShadowRootClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ShadowRootClass.type_wf_defs
- document_ptr_kinds_commutes get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get_document_ptr_simp get_object_ptr_simp2 notin_fset
+ document_ptr_kinds_commutes get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def get_document_ptr_simp get_object_ptr_simp2
shadow_root_ptr_kinds_commutes)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "DocumentClass.type_wf h"
assumes "is_shadow_root_ptr_kind ptr \<Longrightarrow> is_shadow_root_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: DocumentMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1]
- by (metis (no_types, lifting) DocumentClass.l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_axioms assms(2) bind.bind_lunit
- cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv finite_set_in get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf option.collapse)
+ by (metis (no_types, opaque_lifting) DocumentClass.l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_axioms assms(2) bind.bind_lunit
+ cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf option.collapse)
subsection \<open>type\_wf\<close>
lemma new_element_type_wf_preserved [simp]:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_element_ptr where "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using new_element_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
with assms show ?thesis
by(auto simp add: ElementMonad.new_element_def type_wf_defs Let_def elim!: bind_returns_heap_E
split: prod.splits)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr RElement.child_nodes_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma new_character_data_type_wf_preserved [simp]:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_character_data_ptr where "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using new_character_data_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
with assms show ?thesis
by(auto simp add: CharacterDataMonad.new_character_data_def type_wf_defs Let_def
elim!: bind_returns_heap_E split: prod.splits)
qed
lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_val_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)
with assms show ?thesis
by(auto simp add: CharacterDataMonad.put_M_defs type_wf_defs)
qed
lemma new_document_type_wf_preserved [simp]:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I DocumentMonad.type_wf_put_I ElementMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(auto simp add: type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_document_kind_def
split: option.splits)[1]
apply (metis fMax_finsert fimage_is_fempty new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap)
apply(auto simp add: type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_document_kind_def
split: option.splits)[1]
apply(metis Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge ffmember_filter fimage_eqI is_document_ptr_ref)
done
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doctype_type_wf_preserved [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr doctype_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: DocumentMonad.put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
proof -
fix x
assume 0: "h \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr id \<rightarrow>\<^sub>r x"
and 1: "h' = put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>doctype := v\<rparr>)) h"
and 2: "ShadowRootClass.type_wf h"
and 3: "is_shadow_root_ptr_kind document_ptr"
obtain shadow_root_ptr where shadow_root_ptr: "document_ptr = cast shadow_root_ptr" and
"shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
by (metis "0" "3" DocumentMonad.get_M_ptr_in_heap is_OK_returns_result_I
is_shadow_root_ptr_kind_obtains shadow_root_ptr_kinds_commutes)
then have "is_shadow_root_kind x"
using 0 2
apply(auto simp add: is_document_kind_def type_wf_defs is_shadow_root_kind_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
split: option.splits Option.bind_splits)[1]
- by (metis (no_types, lifting) DocumentMonad.get_M_defs finite_set_in get_heap_returns_result
+ by (metis (no_types, lifting) DocumentMonad.get_M_defs get_heap_returns_result
id_apply option.simps(5) return_returns_result)
then show "\<exists>y. cast y = x\<lparr>doctype := v\<rparr>"
using cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_none is_shadow_root_kind_doctype is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def by blast
next
fix x
assume 0: "h \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr id \<rightarrow>\<^sub>r x"
and 1: "h' = put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>doctype := v\<rparr>)) h"
and 2: "ShadowRootClass.type_wf (put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>doctype := v\<rparr>)) h)"
have 3: "\<And>document_ptr'. document_ptr' \<noteq> document_ptr \<Longrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr') h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr') h'"
by (simp add: "1")
have "document_ptr |\<in>| document_ptr_kinds h"
by (meson "0" DocumentMonad.get_M_ptr_in_heap is_OK_returns_result_I)
show "ShadowRootClass.type_wf h"
proof (cases "is_shadow_root_ptr_kind document_ptr")
case True
then obtain shadow_root_ptr where shadow_root_ptr: "document_ptr = cast shadow_root_ptr"
using is_shadow_root_ptr_kind_obtains by blast
then
have "is_shadow_root_kind (x\<lparr>doctype := v\<rparr>)"
using 2 True
by(simp add: type_wf_defs is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: if_splits option.splits)
then
have "is_shadow_root_kind x"
using is_shadow_root_kind_doctype by blast
then
have "is_shadow_root_kind (the (get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr) h))"
using 0
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def is_shadow_root_kind_def
split: option.splits Option.bind_splits)
show ?thesis
using 0 2 \<open>is_shadow_root_kind x\<close> shadow_root_ptr
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def
is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits if_splits)
next
case False
then show ?thesis
using 0 1 2
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def
is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits if_splits)
qed
qed
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document_element_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms
apply(auto simp add: DocumentMonad.put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
proof -
fix x
assume 0: "h \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr id \<rightarrow>\<^sub>r x"
and 1: "h' = put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>document_element := v\<rparr>)) h"
and 2: "ShadowRootClass.type_wf h"
and 3: "is_shadow_root_ptr_kind document_ptr"
obtain shadow_root_ptr where shadow_root_ptr: "document_ptr = cast shadow_root_ptr" and
"shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
by (metis "0" "3" DocumentMonad.get_M_ptr_in_heap is_OK_returns_result_I
is_shadow_root_ptr_kind_obtains shadow_root_ptr_kinds_commutes)
then have "is_shadow_root_kind x"
using 0 2
apply(auto simp add: is_document_kind_def type_wf_defs is_shadow_root_kind_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
split: option.splits Option.bind_splits)[1]
- by (metis (no_types, lifting) DocumentMonad.get_M_defs finite_set_in get_heap_returns_result id_def
+ by (metis (no_types, lifting) DocumentMonad.get_M_defs get_heap_returns_result id_def
option.simps(5) return_returns_result)
then show "\<exists>y. cast y = x\<lparr>document_element := v\<rparr>"
using cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_none is_shadow_root_kind_document_element is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by blast
next
fix x
assume 0: "h \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr id \<rightarrow>\<^sub>r x"
and 1: "h' = put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>document_element := v\<rparr>)) h"
and 2: "ShadowRootClass.type_wf (put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>document_element := v\<rparr>)) h)"
have 3: "\<And>document_ptr'. document_ptr' \<noteq> document_ptr \<Longrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr') h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr') h'"
by (simp add: "1")
have "document_ptr |\<in>| document_ptr_kinds h"
by (meson "0" DocumentMonad.get_M_ptr_in_heap is_OK_returns_result_I)
show "ShadowRootClass.type_wf h"
proof (cases "is_shadow_root_ptr_kind document_ptr")
case True
then obtain shadow_root_ptr where shadow_root_ptr: "document_ptr = cast shadow_root_ptr"
using is_shadow_root_ptr_kind_obtains by blast
then
have "is_shadow_root_kind (x\<lparr>document_element := v\<rparr>)"
using 2 True
by(simp add: type_wf_defs is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: if_splits option.splits)
then
have "is_shadow_root_kind x"
using is_shadow_root_kind_document_element by blast
then
have "is_shadow_root_kind (the (get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr) h))"
using 0
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def is_shadow_root_kind_def
split: option.splits Option.bind_splits)
show ?thesis
using 0 2 \<open>is_shadow_root_kind x\<close> shadow_root_ptr
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def
is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits if_splits)
next
case False
then show ?thesis
using 0 1 2
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def
is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits if_splits)
qed
qed
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved [simp]:
assumes "h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr disconnected_nodes_update v \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms
apply(auto simp add: DocumentMonad.put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
proof -
fix x
assume 0: "h \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr id \<rightarrow>\<^sub>r x"
and 1: "h' = put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>disconnected_nodes := v\<rparr>)) h"
and 2: "ShadowRootClass.type_wf h"
and 3: "is_shadow_root_ptr_kind document_ptr"
obtain shadow_root_ptr where shadow_root_ptr: "document_ptr = cast shadow_root_ptr" and
"shadow_root_ptr |\<in>| shadow_root_ptr_kinds h"
by (metis "0" "3" DocumentMonad.get_M_ptr_in_heap is_OK_returns_result_I is_shadow_root_ptr_kind_obtains
shadow_root_ptr_kinds_commutes)
then have "is_shadow_root_kind x"
using 0 2
apply(auto simp add: is_document_kind_def type_wf_defs is_shadow_root_kind_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
split: option.splits Option.bind_splits)[1]
- by (metis (no_types, lifting) DocumentMonad.get_M_defs finite_set_in get_heap_returns_result
+ by (metis (no_types, lifting) DocumentMonad.get_M_defs get_heap_returns_result
id_def option.simps(5) return_returns_result)
then show "\<exists>y. cast y = x\<lparr>disconnected_nodes := v\<rparr>"
using cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_none is_shadow_root_kind_disconnected_nodes is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by blast
next
fix x
assume 0: "h \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr id \<rightarrow>\<^sub>r x"
and 1: "h' = put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>disconnected_nodes := v\<rparr>)) h"
and 2: "ShadowRootClass.type_wf (put (cast document_ptr) (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (x\<lparr>disconnected_nodes := v\<rparr>)) h)"
have 3: "\<And>document_ptr'. document_ptr' \<noteq> document_ptr \<Longrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr') h = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr') h'"
by (simp add: "1")
have "document_ptr |\<in>| document_ptr_kinds h"
by (meson "0" DocumentMonad.get_M_ptr_in_heap is_OK_returns_result_I)
show "ShadowRootClass.type_wf h"
proof (cases "is_shadow_root_ptr_kind document_ptr")
case True
then obtain shadow_root_ptr where shadow_root_ptr: "document_ptr = cast shadow_root_ptr"
using is_shadow_root_ptr_kind_obtains by blast
then
have "is_shadow_root_kind (x\<lparr>disconnected_nodes := v\<rparr>)"
using 2 True
by(simp add: type_wf_defs is_shadow_root_kind\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: if_splits option.splits)
then
have "is_shadow_root_kind x"
using is_shadow_root_kind_disconnected_nodes by blast
then
have "is_shadow_root_kind (the (get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (cast document_ptr) h))"
using 0
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def is_shadow_root_kind_def
split: option.splits Option.bind_splits)
show ?thesis
using 0 2 \<open>is_shadow_root_kind x\<close> shadow_root_ptr
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def
is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits if_splits)
next
case False
then show ?thesis
using 0 1 2
by(auto simp add: DocumentMonad.a_get_M_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def is_shadow_root_kind_def
is_document_kind_def type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits if_splits)
qed
qed
lemma put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_mode_type_wf_preserved [simp]:
"h \<turnstile> put_M shadow_root_ptr mode_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
by(auto simp add: get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def DocumentMonad.get_M_defs put_M_defs put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E elim!: bind_returns_heap_E2 intro!: type_wf_put_I DocumentMonad.type_wf_put_I
CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
simp add: is_shadow_root_kind_def is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs CharacterDataClass.type_wf_defs DocumentClass.type_wf_defs split: option.splits)[1]
lemma put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M shadow_root_ptr RShadowRoot.child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
by(auto simp add: get_M_defs get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def DocumentMonad.get_M_defs put_M_defs put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def
put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E elim!: bind_returns_heap_E2 intro!: type_wf_put_I
DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
simp add: is_shadow_root_kind_def is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs CharacterDataClass.type_wf_defs
DocumentClass.type_wf_defs split: option.splits)[1]
lemma shadow_root_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def preserved_def
object_ptr_kinds_preserved_small[OF assms])
lemma shadow_root_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def shadow_root_ptr_kinds_def document_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma new_shadow_root_known_ptr:
assumes "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>r new_shadow_root_ptr"
shows "known_ptr (cast new_shadow_root_ptr)"
using assms
apply(auto simp add: 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 a_known_ptr_def
elim!: bind_returns_result_E2 split: prod.splits)[1]
using assms 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 by blast
lemma new_shadow_root_type_wf_preserved [simp]: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: 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 put\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ShadowRootClass.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ShadowRootClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
ShadowRootClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e ShadowRootClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_not_in_heap
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I DocumentMonad.type_wf_put_I ElementMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
by(auto simp add: type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_shadow_root_kind_def is_document_kind_def
split: option.splits)[1]
locale l_new_shadow_root = l_type_wf +
assumes new_shadow_root_types_preserved: "h \<turnstile> new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
lemma new_shadow_root_is_l_new_shadow_root [instances]: "l_new_shadow_root type_wf"
using l_new_shadow_root.intro new_shadow_root_type_wf_preserved
by blast
lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
assumes "\<And>shadow_root_ptr. preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr RShadowRoot.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3) assms(4) assms(5)]
allI[OF assms(6), of id, simplified] shadow_root_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs )[1]
apply(auto simp add: preserved_def get_M_defs shadow_root_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply(force)
apply(auto simp add: preserved_def get_M_defs shadow_root_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply(force)
done
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> \<forall>character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> \<forall>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> \<forall>shadow_root_ptr. preserved (get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t shadow_root_ptr RShadowRoot.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs)[1]
using type_wf_drop
apply blast
- by (metis (no_types, lifting) DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t DocumentMonad.type_wf_drop
- ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf document_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
+ by (metis (no_types, opaque_lifting) DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t DocumentMonad.type_wf_drop
+ ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf document_ptr_kinds_commutes fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def heap.sel shadow_root_ptr_kinds_commutes)
lemma delete_shadow_root_type_wf_preserved [simp]:
assumes "h \<turnstile> 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 \<rightarrow>\<^sub>h h'"
assumes "type_wf h"
shows "type_wf h'"
using assms
using type_wf_drop
by(auto simp add: delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_splits)
lemma new_element_is_l_new_element [instances]:
"l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma new_character_data_is_l_new_character_data [instances]:
"l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma new_document_is_l_new_document [instances]:
"l_new_document type_wf"
using l_new_document.intro new_document_type_wf_preserved
by blast
end
diff --git a/thys/Simplex/Linear_Poly_Maps.thy b/thys/Simplex/Linear_Poly_Maps.thy
--- a/thys/Simplex/Linear_Poly_Maps.thy
+++ b/thys/Simplex/Linear_Poly_Maps.thy
@@ -1,204 +1,204 @@
(* Authors: F. Maric, M. Spasic, R. Thiemann *)
theory Linear_Poly_Maps
imports Abstract_Linear_Poly
"HOL-Library.Finite_Map"
"HOL-Library.Monad_Syntax"
begin
(* ************************************************************************** *)
(* Executable implementation *)
(* ************************************************************************** *)
definition get_var_coeff :: "(var, rat) fmap \<Rightarrow> var \<Rightarrow> rat" where
"get_var_coeff lp v == case fmlookup lp v of None \<Rightarrow> 0 | Some c \<Rightarrow> c"
definition set_var_coeff :: "var \<Rightarrow> rat \<Rightarrow> (var, rat) fmap \<Rightarrow> (var, rat) fmap" where
"set_var_coeff v c lp ==
if c = 0 then fmdrop v lp else fmupd v c lp"
lift_definition LinearPoly :: "(var, rat) fmap \<Rightarrow> linear_poly" is get_var_coeff
proof -
fix fmap
show "inv (get_var_coeff fmap)" unfolding inv_def
by (rule finite_subset[OF _ dom_fmlookup_finite[of fmap]],
auto intro: fmdom'I simp: get_var_coeff_def split: option.splits)
qed
definition ordered_keys :: "('a :: linorder, 'b)fmap \<Rightarrow> 'a list" where
"ordered_keys m = sorted_list_of_set (fset (fmdom m))"
context includes fmap.lifting lifting_syntax
begin
lemma [transfer_rule]: "(((=) ===> (=)) ===> pcr_linear_poly ===> (=)) (=) pcr_linear_poly"
by(standard,auto simp: pcr_linear_poly_def cr_linear_poly_def rel_fun_def OO_def)
lemma [transfer_rule]: "(pcr_fmap (=) (=) ===> pcr_linear_poly) (\<lambda> f x. case f x of None \<Rightarrow> 0 | Some x \<Rightarrow> x) LinearPoly"
by (standard, transfer, auto simp:get_var_coeff_def fmap.pcr_cr_eq cr_fmap_def)
lift_definition linear_poly_map :: "linear_poly \<Rightarrow> (var, rat) fmap" is
"\<lambda> lp x. if lp x = 0 then None else Some (lp x)" by (auto simp: dom_def)
lemma certificate[code abstype]:
"LinearPoly (linear_poly_map lp) = lp"
by (transfer, auto)
text\<open>Zero\<close>
definition zero :: "(var, rat)fmap" where "zero = fmempty"
lemma [code abstract]:
"linear_poly_map 0 = zero" unfolding zero_def
by (transfer, auto)
text\<open>Addition\<close>
definition add_monom :: "rat \<Rightarrow> var \<Rightarrow> (var, rat) fmap \<Rightarrow> (var, rat) fmap" where
"add_monom c v lp == set_var_coeff v (c + get_var_coeff lp v) lp"
definition add :: "(var, rat) fmap \<Rightarrow> (var, rat) fmap \<Rightarrow> (var, rat) fmap" where
"add lp1 lp2 = foldl (\<lambda> lp v. add_monom (get_var_coeff lp1 v) v lp) lp2 (ordered_keys lp1)"
lemma lookup_add_monom:
"get_var_coeff lp v + c \<noteq> 0 \<Longrightarrow>
fmlookup (add_monom c v lp) v = Some (get_var_coeff lp v + c)"
"get_var_coeff lp v + c = 0 \<Longrightarrow>
fmlookup (add_monom c v lp) v = None"
"x \<noteq> v \<Longrightarrow> fmlookup (add_monom c v lp) x = fmlookup lp x"
unfolding add_monom_def get_var_coeff_def set_var_coeff_def
by auto
lemma fmlookup_fold_not_mem: "x \<notin> set k1 \<Longrightarrow>
fmlookup (foldl (\<lambda>lp v. add_monom (get_var_coeff P1 v) v lp) P2 k1) x
= fmlookup P2 x"
by (induct k1 arbitrary: P2, auto simp: lookup_add_monom)
lemma [code abstract]:
"linear_poly_map (p1 + p2) = add (linear_poly_map p1) (linear_poly_map p2)"
proof (rule fmap_ext)
fix x :: nat (* index *)
let ?p1 = "fmlookup (linear_poly_map p1) x"
let ?p2 = "fmlookup (linear_poly_map p2) x"
define P1 where "P1 = linear_poly_map p1"
define P2 where "P2 = linear_poly_map p2"
define k1 where "k1 = ordered_keys P1"
have k1: "distinct k1 \<and> fset (fmdom P1) = set k1" unfolding k1_def P1_def ordered_keys_def
by auto
have id: "fmlookup (linear_poly_map (p1 + p2)) x = (case ?p1 of None \<Rightarrow> ?p2 | Some y1 \<Rightarrow>
(case ?p2 of None \<Rightarrow> Some y1 | Some y2 \<Rightarrow> if y1 + y2 = 0 then None else Some (y1 + y2)))"
by (transfer, auto)
show "fmlookup (linear_poly_map (p1 + p2)) x = fmlookup (add (linear_poly_map p1) (linear_poly_map p2)) x"
proof (cases "fmlookup P1 x")
case None
- from fmdom_notI[OF None] have "x \<notin> fset (fmdom P1)" by (meson notin_fset)
+ from fmdom_notI[OF None] have "x \<notin> fset (fmdom P1)" by metis
with k1 have x: "x \<notin> set k1" by auto
show ?thesis unfolding id P1_def[symmetric] P2_def[symmetric] None
unfolding add_def k1_def[symmetric] fmlookup_fold_not_mem[OF x] by auto
next
case (Some y1)
- from fmdomI[OF this] have "x \<in> fset (fmdom P1)" by (meson notin_fset)
+ from fmdomI[OF this] have "x \<in> fset (fmdom P1)" by metis
with k1 have "x \<in> set k1" by auto
from split_list[OF this] obtain bef aft where k1_id: "k1 = bef @ x # aft" by auto
with k1 have x: "x \<notin> set bef" "x \<notin> set aft" by auto
have xy1: "get_var_coeff P1 x = y1" using Some unfolding get_var_coeff_def by auto
let ?P = "foldl (\<lambda>lp v. add_monom (get_var_coeff P1 v) v lp) P2 bef"
show ?thesis unfolding id P1_def[symmetric] P2_def[symmetric] Some option.simps
unfolding add_def k1_def[symmetric] k1_id foldl_append foldl_Cons
unfolding fmlookup_fold_not_mem[OF x(2)] xy1
proof -
show "(case fmlookup P2 x of None \<Rightarrow> Some y1 | Some y2 \<Rightarrow> if y1 + y2 = 0 then None else Some (y1 + y2))
= fmlookup (add_monom y1 x ?P) x"
proof (cases "get_var_coeff ?P x + y1 = 0")
case True
from Some[unfolded P1_def] have y1: "y1 \<noteq> 0"
by (transfer, auto split: if_splits)
then show ?thesis unfolding lookup_add_monom(2)[OF True] using True
unfolding get_var_coeff_def[of _ x] fmlookup_fold_not_mem[OF x(1)]
by (auto split: option.splits)
next
case False
show ?thesis unfolding lookup_add_monom(1)[OF False] using False
unfolding get_var_coeff_def[of _ x] fmlookup_fold_not_mem[OF x(1)]
by (auto split: option.splits)
qed
qed
qed
qed
text\<open>Scaling\<close>
definition scale :: "rat \<Rightarrow> (var, rat) fmap \<Rightarrow> (var, rat) fmap" where
"scale r lp = (if r = 0 then fmempty else (fmmap ((*) r) lp))"
lemma [code abstract]:
"linear_poly_map (r *R p) = scale r (linear_poly_map p)"
proof (cases "r = 0")
case True
then have *: "(r = 0) = True" by simp
show ?thesis unfolding scale_def * if_True using True
by (transfer, auto)
next
case False
then have *: "(r = 0) = False" by simp
show ?thesis unfolding scale_def * if_False using False
by (transfer, auto)
qed
(* Coeff *)
lemma coeff_code [code]:
"coeff lp = get_var_coeff (linear_poly_map lp)"
by (rule ext, unfold get_var_coeff_def, transfer, auto)
(* Var *)
lemma Var_code[code abstract]:
"linear_poly_map (Var x) = set_var_coeff x 1 fmempty"
unfolding set_var_coeff_def
by (transfer, auto split: if_splits simp: fun_eq_iff map_upd_def)
(* vars *)
lemma vars_code[code]: "vars lp = fset (fmdom (linear_poly_map lp))"
by (transfer, auto simp: Transfer.Rel_def rel_fun_def pcr_fset_def cr_fset_def)
(* vars_list *)
lemma vars_list_code[code]: "vars_list lp = ordered_keys (linear_poly_map lp)"
unfolding ordered_keys_def vars_code[symmetric]
by (transfer, auto)
(* valuate *)
lemma valuate_code[code]: "valuate lp val = (
let lpm = linear_poly_map lp
in sum_list (map (\<lambda> x. (the (fmlookup lpm x)) *R (val x)) (vars_list lp)))"
unfolding Let_def
proof (subst sum_list_distinct_conv_sum_set)
show "distinct (vars_list lp)"
by (transfer, auto)
next
show "lp \<lbrace> val \<rbrace> =
(\<Sum>x\<in>set (vars_list lp). the (fmlookup (linear_poly_map lp) x) *R val x)"
unfolding set_vars_list
by (transfer, auto)
qed
end
lemma lp_monom_code[code]: "linear_poly_map (lp_monom c x) = (if c = 0 then fmempty else fmupd x c fmempty)"
proof (rule fmap_ext, goal_cases)
case (1 y)
include fmap.lifting
show ?case by (cases "c = 0", (transfer, auto)+)
qed
instantiation linear_poly :: equal
begin
definition "equal_linear_poly x y = (linear_poly_map x = linear_poly_map y)"
instance
proof (standard, unfold equal_linear_poly_def, standard)
fix x y
assume "linear_poly_map x = linear_poly_map y"
from arg_cong[OF this, of LinearPoly, unfolded certificate]
show "x = y" .
qed auto
end
end
\ No newline at end of file
diff --git a/thys/Solidity/Statements.thy b/thys/Solidity/Statements.thy
--- a/thys/Solidity/Statements.thy
+++ b/thys/Solidity/Statements.thy
@@ -1,3926 +1,3927 @@
section \<open>Statements\<close>
theory Statements
imports Environment StateMonad
begin
subsection \<open>Syntax\<close>
subsubsection \<open>Expressions\<close>
datatype L = Id Identifier
| Ref Identifier "E list"
and E = INT nat int
| UINT nat int
| ADDRESS String.literal
| BALANCE E
| THIS
| SENDER
| VALUE
| TRUE
| FALSE
| LVAL L
| PLUS E E
| MINUS E E
| EQUAL E E
| LESS E E
| AND E E
| OR E E
| NOT E
| CALL Identifier "E list"
| ECALL E Identifier "E list" E
subsubsection \<open>Statements\<close>
datatype S = SKIP
| BLOCK "(Identifier \<times> Type) \<times> (E option)" S
| ASSIGN L E
| TRANSFER E E
| COMP S S
| ITE E S S
| WHILE E S
| INVOKE Identifier "E list"
| EXTERNAL E Identifier "E list" E
abbreviation
"vbits\<equiv>{8,16,24,32,40,48,56,64,72,80,88,96,104,112,120,128,
136,144,152,160,168,176,184,192,200,208,216,224,232,240,248,256}"
lemma vbits_max[simp]:
assumes "b1 \<in> vbits"
and "b2 \<in> vbits"
shows "(max b1 b2) \<in> vbits"
proof -
consider (b1) "max b1 b2 = b1" | (b2) "max b1 b2 = b2" by (metis max_def)
then show ?thesis
proof cases
case b1
then show ?thesis using assms(1) by simp
next
case b2
then show ?thesis using assms(2) by simp
qed
qed
lemma vbits_ge_0: "(x::nat)\<in>vbits \<Longrightarrow> x>0" by auto
subsection \<open>Contracts\<close>
text \<open>
A contract consists of methods or storage variables.
A method is a triple consisting of
\begin{itemize}
\item A list of formal parameters
\item A statement
\item An optional return value
\end{itemize}
\<close>
datatype Member = Method "(Identifier \<times> Type) list \<times> S \<times> E option"
| Var STypes
text \<open>
A procedure environment assigns a contract to an address.
A contract consists of
\begin{itemize}
\item An assignment of members to identifiers
\item An optional fallback statement which is executed after money is beeing transferred to the contract.
\end{itemize}
\url{https://docs.soliditylang.org/en/v0.8.6/contracts.html#fallback-function}
\<close>
type_synonym Environment\<^sub>P = "(Address, (Identifier, Member) fmap \<times> S) fmap"
definition init::"(Identifier, Member) fmap \<Rightarrow> Identifier \<Rightarrow> Environment \<Rightarrow> Environment"
where "init ct i e = (case fmlookup ct i of
Some (Var tp) \<Rightarrow> updateEnvDup i (Storage tp) (Storeloc i) e
| _ \<Rightarrow> e)"
lemma init_s11[simp]:
assumes "fmlookup ct i = Some (Var tp)"
shows "init ct i e = updateEnvDup i (Storage tp) (Storeloc i) e"
using assms init_def by simp
lemma init_s12[simp]:
assumes "i |\<in>| fmdom (denvalue e)"
shows "init ct i e = e"
proof (cases "fmlookup ct i")
case None
then show ?thesis using init_def by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
with Some show ?thesis using init_def by simp
next
case (Var tp)
with Some have "init ct i e = updateEnvDup i (Storage tp) (Storeloc i) e" using init_def by simp
moreover from assms have "updateEnvDup i (Storage tp) (Storeloc i) e = e" by auto
ultimately show ?thesis by simp
qed
qed
lemma init_s13[simp]:
assumes "fmlookup ct i = Some (Var tp)"
and "\<not> i |\<in>| fmdom (denvalue e)"
shows "init ct i e = updateEnv i (Storage tp) (Storeloc i) e"
using assms init_def by auto
lemma init_s21[simp]:
assumes "fmlookup ct i = None"
shows "init ct i e = e"
using assms init_def by auto
lemma init_s22[simp]:
assumes "fmlookup ct i = Some (Method m)"
shows "init ct i e = e"
using assms init_def by auto
lemma init_commte: "comp_fun_commute (init ct)"
proof
fix x y
show "init ct y \<circ> init ct x = init ct x \<circ> init ct y"
proof
fix e
show "(init ct y \<circ> init ct x) e = (init ct x \<circ> init ct y) e"
proof (cases "fmlookup ct x")
case None
then show ?thesis by simp
next
case s1: (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
with s1 show ?thesis by simp
next
case v1: (Var tp)
then show ?thesis
proof (cases "x |\<in>| fmdom (denvalue e)")
case True
with s1 v1 have *: "init ct x e = e" by auto
then show ?thesis
proof (cases "fmlookup ct y")
case None
then show ?thesis by simp
next
case s2: (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
with s2 show ?thesis by simp
next
case v2: (Var tp')
then show ?thesis
proof (cases "y |\<in>| fmdom (denvalue e)")
case t1: True
with s1 v1 True s2 v2 show ?thesis by fastforce
next
define e' where "e' = updateEnv y (Storage tp') (Storeloc y) e"
case False
with s2 v2 have "init ct y e = e'" using e'_def by auto
with s1 v1 True e'_def * show ?thesis by auto
qed
qed
qed
next
define e' where "e' = updateEnv x (Storage tp) (Storeloc x) e"
case f1: False
with s1 v1 have *: "init ct x e = e'" using e'_def by auto
then show ?thesis
proof (cases "fmlookup ct y")
case None
then show ?thesis by simp
next
case s3: (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
with s3 show ?thesis by simp
next
case v2: (Var tp')
then show ?thesis
proof (cases "y |\<in>| fmdom (denvalue e)")
case t1: True
with e'_def have "y |\<in>| fmdom (denvalue e')" by simp
with s1 s3 v1 f1 v2 show ?thesis using e'_def by fastforce
next
define f' where "f' = updateEnv y (Storage tp') (Storeloc y) e"
define e'' where "e'' = updateEnv y (Storage tp') (Storeloc y) e'"
case f2: False
with s3 v2 have **: "init ct y e = f'" using f'_def by auto
show ?thesis
proof (cases "y = x")
case True
with s3 v2 e'_def have "init ct y e' = e'" by simp
moreover from s3 v2 True f'_def have "init ct x f' = f'" by simp
ultimately show ?thesis using True by simp
next
define f'' where "f'' = updateEnv x (Storage tp) (Storeloc x) f'"
case f3: False
with f2 have "\<not> y |\<in>| fmdom (denvalue e')" using e'_def by simp
with s3 v2 e''_def have "init ct y e' = e''" by auto
with * have "(init ct y \<circ> init ct x) e = e''" by simp
moreover have "init ct x f' = f''"
proof -
from s1 v1 have "init ct x f' = updateEnvDup x (Storage tp) (Storeloc x) f'" by simp
moreover from f1 f3 have "x |\<notin>| fmdom (denvalue f')" using f'_def by simp
ultimately show ?thesis using f''_def by auto
qed
moreover from f''_def e''_def f'_def e'_def f3 have "Some f'' = Some e''" using env_reorder_neq by simp
ultimately show ?thesis using ** by simp
qed
qed
qed
qed
qed
qed
qed
qed
qed
lemma init_address[simp]:
"address (init ct i e) = address e \<and> sender (init ct i e) = sender e"
proof (cases "fmlookup ct i")
case None
then show ?thesis by simp
next
case (Some a)
show ?thesis
proof (cases a)
case (Method x1)
with Some show ?thesis by simp
next
case (Var tp)
with Some show ?thesis using updateEnvDup_address updateEnvDup_sender by simp
qed
qed
lemma init_sender[simp]:
"sender (init ct i e) = sender e"
proof (cases "fmlookup ct i")
case None
then show ?thesis by simp
next
case (Some a)
show ?thesis
proof (cases a)
case (Method x1)
with Some show ?thesis by simp
next
case (Var tp)
with Some show ?thesis using updateEnvDup_sender by simp
qed
qed
lemma init_svalue[simp]:
"svalue (init ct i e) = svalue e"
proof (cases "fmlookup ct i")
case None
then show ?thesis by simp
next
case (Some a)
show ?thesis
proof (cases a)
case (Method x1)
with Some show ?thesis by simp
next
case (Var tp)
with Some show ?thesis using updateEnvDup_svalue by simp
qed
qed
lemma ffold_init_ad_same[rule_format]: "\<forall>e'. ffold (init ct) e xs = e' \<longrightarrow> address e' = address e \<and> sender e' = sender e \<and> svalue e' = svalue e"
proof (induct xs)
case empty
then show ?case by (simp add: ffold_def)
next
case (insert x xs)
then have *: "ffold (init ct) e (finsert x xs) =
init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp
show ?case
proof (rule allI[OF impI])
fix e' assume **: "ffold (init ct) e (finsert x xs) = e'"
with * obtain e'' where ***: "ffold (init ct) e xs = e''" by simp
with insert have "address e'' = address e \<and> sender e'' = sender e \<and> svalue e'' = svalue e" by blast
with * ** *** show "address e' = address e \<and> sender e' = sender e \<and> svalue e' = svalue e" using init_address init_sender init_svalue by metis
qed
qed
lemma ffold_init_dom:
"fmdom (denvalue (ffold (init ct) e xs)) |\<subseteq>| fmdom (denvalue e) |\<union>| xs"
proof (induct "xs")
case empty
then show ?case
proof
fix x
assume "x |\<in>| fmdom (denvalue (ffold (init ct) e {||}))"
- moreover have "ffold (init ct) e {||} = e" using FSet.comp_fun_commute.ffold_empty[OF init_commte, of "init ct" e] by simp
+ moreover have "ffold (init ct) e {||} = e"
+ using FSet.comp_fun_commute.ffold_empty[OF init_commte, of ct e] by simp
ultimately show "x |\<in>| fmdom (denvalue e) |\<union>| {||}" by simp
qed
next
case (insert x xs)
then have *: "ffold (init ct) e (finsert x xs) =
init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp
show ?case
proof
fix x' assume "x' |\<in>| fmdom (denvalue (ffold (init ct) e (finsert x xs)))"
with * have **: "x' |\<in>| fmdom (denvalue (init ct x (ffold (init ct) e xs)))" by simp
then consider "x' |\<in>| fmdom (denvalue (ffold (init ct) e xs))" | "x'=x"
proof (cases "fmlookup ct x")
case None
then show ?thesis using that ** by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
then show ?thesis using Some ** that by simp
next
case (Var x2)
show ?thesis
proof (cases "x=x'")
case True
then show ?thesis using that by simp
next
case False
then have "fmlookup (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs))) x' = fmlookup (denvalue (ffold (init ct) e xs)) x'" using updateEnvDup_dup by simp
moreover from ** Some Var have ***:"x' |\<in>| fmdom (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs)))" by simp
ultimately have "x' |\<in>| fmdom (denvalue (ffold (init ct) e xs))" by (simp add: fmlookup_dom_iff)
then show ?thesis using that by simp
qed
qed
qed
then show "x' |\<in>| fmdom (denvalue e) |\<union>| finsert x xs"
proof cases
case 1
then show ?thesis using insert.hyps by auto
next
case 2
then show ?thesis by simp
qed
qed
qed
lemma ffold_init_fmap:
assumes "fmlookup ct i = Some (Var tp)"
and "i |\<notin>| fmdom (denvalue e)"
shows "i|\<in>|xs \<Longrightarrow> fmlookup (denvalue (ffold (init ct) e xs)) i = Some (Storage tp, Storeloc i)"
proof (induct "xs")
case empty
then show ?case by simp
next
case (insert x xs)
then have *: "ffold (init ct) e (finsert x xs) =
init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp
from insert.prems consider (a) "i |\<in>| xs" | (b) "\<not> i |\<in>| xs \<and> i = x" by auto
then show "fmlookup (denvalue (ffold (init ct) e (finsert x xs))) i = Some (Storage tp, Storeloc i)"
proof cases
case a
with insert.hyps(2) have "fmlookup (denvalue (ffold (init ct) e xs)) i = Some (Storage tp, Storeloc i)" by simp
moreover have "fmlookup (denvalue (init ct x (ffold (init ct) e xs))) i = fmlookup (denvalue (ffold (init ct) e xs)) i"
proof (cases "fmlookup ct x")
case None
then show ?thesis by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
with Some show ?thesis by simp
next
case (Var x2)
with Some have "init ct x (ffold (init ct) e xs) = updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs)" using init_def[of ct x "(ffold (init ct) e xs)"] by simp
moreover from insert a have "i\<noteq>x" by auto
then have "fmlookup (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs))) i = fmlookup (denvalue (ffold (init ct) e xs)) i" using updateEnvDup_dup[of x i] by simp
ultimately show ?thesis by simp
qed
qed
ultimately show ?thesis using * by simp
next
case b
with assms(1) have "fmlookup ct x = Some (Var tp)" by simp
moreover from b assms(2) have "\<not> x |\<in>| fmdom (denvalue (ffold (init ct) e xs))" using ffold_init_dom by auto
ultimately have "init ct x (ffold (init ct) e xs) = updateEnv x (Storage tp) (Storeloc x) (ffold (init ct) e xs)" by auto
with b * show ?thesis by simp
qed
qed
text\<open>The following definition allows for a more fine-grained configuration of the
code generator.
\<close>
definition ffold_init::"(String.literal, Member) fmap \<Rightarrow> Environment \<Rightarrow> String.literal fset \<Rightarrow> Environment" where
\<open>ffold_init ct a c = ffold (init ct) a c\<close>
declare ffold_init_def [simp]
lemma ffold_init_code [code]:
\<open>ffold_init ct a c = fold (init ct) (remdups (sorted_list_of_set (fset c))) a\<close>
using comp_fun_commute_on.fold_set_fold_remdups ffold.rep_eq
ffold_init_def init_commte sorted_list_of_fset.rep_eq
sorted_list_of_fset_simps(1)
by (metis comp_fun_commute.comp_fun_commute comp_fun_commute_on.intro order_refl)
lemma bind_case_stackvalue_cong [fundef_cong]:
assumes "x = x'"
and "\<And>v. x = KValue v \<Longrightarrow> f v s = f' v s"
and "\<And>p. x = KCDptr p \<Longrightarrow> g p s = g' p s"
and "\<And>p. x = KMemptr p \<Longrightarrow> h p s = h' p s"
and "\<And>p. x = KStoptr p \<Longrightarrow> i p s = i' p s"
shows "(case x of KValue v \<Rightarrow> f v | KCDptr p \<Rightarrow> g p | KMemptr p \<Rightarrow> h p | KStoptr p \<Rightarrow> i p) s
= (case x' of KValue v \<Rightarrow> f' v | KCDptr p \<Rightarrow> g' p | KMemptr p \<Rightarrow> h' p | KStoptr p \<Rightarrow> i' p) s"
using assms by (cases x, auto)
lemma bind_case_type_cong [fundef_cong]:
assumes "x = x'"
and "\<And>t. x = Value t \<Longrightarrow> f t s = f' t s"
and "\<And>t. x = Calldata t \<Longrightarrow> g t s = g' t s"
and "\<And>t. x = Memory t \<Longrightarrow> h t s = h' t s"
and "\<And>t. x = Storage t \<Longrightarrow> i t s = i' t s"
shows "(case x of Value t \<Rightarrow> f t | Calldata t \<Rightarrow> g t | Memory t \<Rightarrow> h t | Storage t \<Rightarrow> i t ) s
= (case x' of Value t \<Rightarrow> f' t | Calldata t \<Rightarrow> g' t | Memory t \<Rightarrow> h' t | Storage t \<Rightarrow> i' t) s"
using assms by (cases x, auto)
lemma bind_case_denvalue_cong [fundef_cong]:
assumes "x = x'"
and "\<And>a. x = (Stackloc a) \<Longrightarrow> f a s = f' a s"
and "\<And>a. x = (Storeloc a) \<Longrightarrow> g a s = g' a s"
shows "(case x of (Stackloc a) \<Rightarrow> f a | (Storeloc a) \<Rightarrow> g a) s
= (case x' of (Stackloc a) \<Rightarrow> f' a | (Storeloc a) \<Rightarrow> g' a) s"
using assms by (cases x, auto)
lemma bind_case_mtypes_cong [fundef_cong]:
assumes "x = x'"
and "\<And>a t. x = (MTArray a t) \<Longrightarrow> f a t s = f' a t s"
and "\<And>p. x = (MTValue p) \<Longrightarrow> g p s = g' p s"
shows "(case x of (MTArray a t) \<Rightarrow> f a t | (MTValue p) \<Rightarrow> g p) s
= (case x' of (MTArray a t) \<Rightarrow> f' a t | (MTValue p) \<Rightarrow> g' p) s"
using assms by (cases x, auto)
lemma bind_case_stypes_cong [fundef_cong]:
assumes "x = x'"
and "\<And>a t. x = (STArray a t) \<Longrightarrow> f a t s = f' a t s"
and "\<And>a t. x = (STMap a t) \<Longrightarrow> g a t s = g' a t s"
and "\<And>p. x = (STValue p) \<Longrightarrow> h p s = h' p s"
shows "(case x of (STArray a t) \<Rightarrow> f a t | (STMap a t) \<Rightarrow> g a t | (STValue p) \<Rightarrow> h p) s
= (case x' of (STArray a t) \<Rightarrow> f' a t | (STMap a t) \<Rightarrow> g' a t | (STValue p) \<Rightarrow> h' p) s"
using assms by (cases x, auto)
lemma bind_case_types_cong [fundef_cong]:
assumes "x = x'"
and "\<And>a. x = (TSInt a) \<Longrightarrow> f a s = f' a s"
and "\<And>a. x = (TUInt a) \<Longrightarrow> g a s = g' a s"
and "x = TBool \<Longrightarrow> h s = h' s"
and "x = TAddr \<Longrightarrow> i s = i' s"
shows "(case x of (TSInt a) \<Rightarrow> f a | (TUInt a) \<Rightarrow> g a | TBool \<Rightarrow> h | TAddr \<Rightarrow> i) s
= (case x' of (TSInt a) \<Rightarrow> f' a | (TUInt a) \<Rightarrow> g' a | TBool \<Rightarrow> h' | TAddr \<Rightarrow> i') s"
using assms by (cases x, auto)
lemma bind_case_contract_cong [fundef_cong]:
assumes "x = x'"
and "\<And>a. x = Method a \<Longrightarrow> f a s = f' a s"
and "\<And>a. x = Var a \<Longrightarrow> g a s = g' a s"
shows "(case x of (Method a) \<Rightarrow> f a | (Var a) \<Rightarrow> g a) s
= (case x' of (Method a) \<Rightarrow> f' a | (Var a) \<Rightarrow> g' a) s"
using assms by (cases x, auto)
lemma bind_case_memoryvalue_cong [fundef_cong]:
assumes "x = x'"
and "\<And>a. x = MValue a \<Longrightarrow> f a s = f' a s"
and "\<And>a. x = MPointer a \<Longrightarrow> g a s = g' a s"
shows "(case x of (MValue a) \<Rightarrow> f a | (MPointer a) \<Rightarrow> g a) s
= (case x' of (MValue a) \<Rightarrow> f' a | (MPointer a) \<Rightarrow> g' a) s"
using assms by (cases x, auto)
abbreviation lift ::
"(E \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (Stackvalue * Type, Ex, State) state_monad)
\<Rightarrow> (Types \<Rightarrow> Types \<Rightarrow> Valuetype \<Rightarrow> Valuetype \<Rightarrow> (Valuetype * Types) option)
\<Rightarrow> E \<Rightarrow> E \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (Stackvalue * Type, Ex, State) state_monad"
where
"lift expr f e1 e2 e\<^sub>p e cd \<equiv>
(do {
kv1 \<leftarrow> expr e1 e\<^sub>p e cd;
(case kv1 of
(KValue v1, Value t1) \<Rightarrow> (do
{
kv2 \<leftarrow> expr e2 e\<^sub>p e cd;
(case kv2 of
(KValue v2, Value t2) \<Rightarrow>
(case f t1 t2 v1 v2 of
Some (v, t) \<Rightarrow> return (KValue v, Value t)
| None \<Rightarrow> throw Err)
| _ \<Rightarrow> (throw Err::(Stackvalue * Type, Ex, State) state_monad))
})
| _ \<Rightarrow> (throw Err::(Stackvalue * Type, Ex, State) state_monad))
})"
abbreviation gascheck ::
"(State \<Rightarrow> Gas) \<Rightarrow> (unit, Ex, State) state_monad"
where
"gascheck check \<equiv>
do {
g \<leftarrow> (applyf check::(Gas, Ex, State) state_monad);
(assert Gas (\<lambda>st. gas st \<le> g) (modify (\<lambda>st. st \<lparr>gas:=gas st - g\<rparr>)::(unit, Ex, State) state_monad))
}"
subsection \<open>Semantics\<close>
datatype LType = LStackloc Location
| LMemloc Location
| LStoreloc Location
locale statement_with_gas =
fixes costs :: "S\<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> State \<Rightarrow> Gas"
and costs\<^sub>e :: "E\<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> State \<Rightarrow> Gas"
assumes while_not_zero[termination_simp]: "\<And>e e\<^sub>p cd st ex s0. 0 < (costs (WHILE ex s0) e\<^sub>p e cd st) "
and call_not_zero[termination_simp]: "\<And>e e\<^sub>p cd st i ix. 0 < (costs\<^sub>e (CALL i ix) e\<^sub>p e cd st)"
and ecall_not_zero[termination_simp]: "\<And>e e\<^sub>p cd st a i ix val. 0 < (costs\<^sub>e (ECALL a i ix val) e\<^sub>p e cd st)"
and invoke_not_zero[termination_simp]: "\<And>e e\<^sub>p cd st i xe. 0 < (costs (INVOKE i xe) e\<^sub>p e cd st)"
and external_not_zero[termination_simp]: "\<And>e e\<^sub>p cd st ad i xe val. 0 < (costs (EXTERNAL ad i xe val) e\<^sub>p e cd st)"
and transfer_not_zero[termination_simp]: "\<And>e e\<^sub>p cd st ex ad. 0 < (costs (TRANSFER ad ex) e\<^sub>p e cd st)"
begin
function msel::"bool \<Rightarrow> MTypes \<Rightarrow> Location \<Rightarrow> E list \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (Location * MTypes, Ex, State) state_monad"
and ssel::"STypes \<Rightarrow> Location \<Rightarrow> E list \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (Location * STypes, Ex, State) state_monad"
and lexp :: "L \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (LType * Type, Ex, State) state_monad"
and expr::"E \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (Stackvalue * Type, Ex, State) state_monad"
and load :: "bool \<Rightarrow> (Identifier \<times> Type) list \<Rightarrow> E list \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> State \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (Environment \<times> CalldataT \<times> State, Ex, State) state_monad"
and rexp::"L \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (Stackvalue * Type, Ex, State) state_monad"
and stmt :: "S \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> (unit, Ex, State) state_monad"
where
"msel _ _ _ [] _ _ _ st = throw Err st"
| "msel _ (MTValue _) _ _ _ _ _ st = throw Err st"
| "msel _ (MTArray al t) loc [x] e\<^sub>p env cd st =
(do {
kv \<leftarrow> expr x e\<^sub>p env cd;
(case kv of
(KValue v, Value t') \<Rightarrow>
(if less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)
then return (hash loc v, t)
else throw Err)
| _ \<Rightarrow> throw Err)
}) st"
(*
Note that it is indeed possible to modify the state while evaluating the expression
to determine the index of the array to look up.
*)
| "msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st =
(do {
kv \<leftarrow> expr x e\<^sub>p env cd;
(case kv of
(KValue v, Value t') \<Rightarrow>
(if less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)
then do {
s \<leftarrow> applyf (\<lambda>st. if mm then memory st else cd);
(case accessStore (hash loc v) s of
Some (MPointer l) \<Rightarrow> msel mm t l (y#ys) e\<^sub>p env cd
| _ \<Rightarrow> throw Err)
} else throw Err)
| _ \<Rightarrow> throw Err)
}) st"
| "ssel tp loc Nil _ _ _ st = return (loc, tp) st"
| "ssel (STValue _) _ (_ # _) _ _ _ st = throw Err st"
| "ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st =
(do {
kv \<leftarrow> expr x e\<^sub>p env cd;
(case kv of
(KValue v, Value t') \<Rightarrow>
(if less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)
then ssel t (hash loc v) xs e\<^sub>p env cd
else throw Err)
| _ \<Rightarrow> throw Err)
}) st"
| "ssel (STMap _ t) loc (x # xs) e\<^sub>p env cd st =
(do {
kv \<leftarrow> expr x e\<^sub>p env cd;
(case kv of
(KValue v, _) \<Rightarrow> ssel t (hash loc v) xs e\<^sub>p env cd
| _ \<Rightarrow> throw Err)
}) st"
| "lexp (Id i) _ e _ st =
(case fmlookup (denvalue e) i of
Some (tp, (Stackloc l)) \<Rightarrow> return (LStackloc l, tp)
| Some (tp, (Storeloc l)) \<Rightarrow> return (LStoreloc l, tp)
| _ \<Rightarrow> throw Err) st"
| "lexp (Ref i r) e\<^sub>p e cd st =
(case fmlookup (denvalue e) i of
Some (tp, Stackloc l) \<Rightarrow>
do {
k \<leftarrow> applyf (\<lambda>st. accessStore l (stack st));
(case k of
Some (KCDptr _) \<Rightarrow> throw Err
| Some (KMemptr l') \<Rightarrow>
(case tp of
Memory t \<Rightarrow>
do {
(l'', t') \<leftarrow> msel True t l' r e\<^sub>p e cd;
return (LMemloc l'', Memory t')
}
| _ \<Rightarrow> throw Err)
| Some (KStoptr l') \<Rightarrow>
(case tp of
Storage t \<Rightarrow>
do {
(l'', t') \<leftarrow> ssel t l' r e\<^sub>p e cd;
return (LStoreloc l'', Storage t')
}
| _ \<Rightarrow> throw Err)
| Some (KValue _) \<Rightarrow> throw Err
| None \<Rightarrow> throw Err)
}
| Some (tp, Storeloc l) \<Rightarrow>
(case tp of
Storage t \<Rightarrow>
do {
(l', t') \<leftarrow> ssel t l r e\<^sub>p e cd;
return (LStoreloc l', Storage t')
}
| _ \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err) st"
| "expr (E.INT b x) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (E.INT b x) e\<^sub>p e cd);
(if (b \<in> vbits)
then (return (KValue (createSInt b x), Value (TSInt b)))
else (throw Err))
}) st"
| "expr (UINT b x) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (UINT b x) e\<^sub>p e cd);
(if (b \<in> vbits)
then (return (KValue (createUInt b x), Value (TUInt b)))
else (throw Err))
}) st"
| "expr (ADDRESS ad) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (ADDRESS ad) e\<^sub>p e cd);
return (KValue ad, Value TAddr)
}) st"
| "expr (BALANCE ad) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (BALANCE ad) e\<^sub>p e cd);
kv \<leftarrow> expr ad e\<^sub>p e cd;
(case kv of
(KValue adv, Value TAddr) \<Rightarrow>
return (KValue (accessBalance (accounts st) adv), Value (TUInt 256))
| _ \<Rightarrow> throw Err)
}) st"
| "expr THIS e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e THIS e\<^sub>p e cd);
return (KValue (address e), Value TAddr)
}) st"
| "expr SENDER e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e SENDER e\<^sub>p e cd);
return (KValue (sender e), Value TAddr)
}) st"
| "expr VALUE e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e VALUE e\<^sub>p e cd);
return (KValue (svalue e), Value (TUInt 256))
}) st"
| "expr TRUE e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e TRUE e\<^sub>p e cd);
return (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True), Value TBool)
}) st"
| "expr FALSE e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e FALSE e\<^sub>p e cd);
return (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False), Value TBool)
}) st"
| "expr (NOT x) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (NOT x) e\<^sub>p e cd);
kv \<leftarrow> expr x e\<^sub>p e cd;
(case kv of
(KValue v, Value t) \<Rightarrow>
(if v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True
then expr FALSE e\<^sub>p e cd
else (if v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False
then expr TRUE e\<^sub>p e cd
else throw Err))
| _ \<Rightarrow> throw Err)
}) st"
| "expr (PLUS e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (PLUS e1 e2) e\<^sub>p e cd) \<bind> (\<lambda>_. lift expr add e1 e2 e\<^sub>p e cd)) st"
| "expr (MINUS e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (MINUS e1 e2) e\<^sub>p e cd) \<bind> (\<lambda>_. lift expr sub e1 e2 e\<^sub>p e cd)) st"
| "expr (LESS e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (LESS e1 e2) e\<^sub>p e cd) \<bind> (\<lambda>_. lift expr less e1 e2 e\<^sub>p e cd)) st"
| "expr (EQUAL e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (EQUAL e1 e2) e\<^sub>p e cd) \<bind> (\<lambda>_. lift expr equal e1 e2 e\<^sub>p e cd)) st"
| "expr (AND e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (AND e1 e2) e\<^sub>p e cd) \<bind> (\<lambda>_. lift expr vtand e1 e2 e\<^sub>p e cd)) st"
| "expr (OR e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (OR e1 e2) e\<^sub>p e cd) \<bind> (\<lambda>_. lift expr vtor e1 e2 e\<^sub>p e cd)) st"
| "expr (LVAL i) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (LVAL i) e\<^sub>p e cd);
rexp i e\<^sub>p e cd
}) st"
(* Notes about method calls:
- Internal method calls use a fresh environment and stack but keep the memory [1]
- External method calls use a fresh environment and stack and memory [2]
[1]: https://docs.soliditylang.org/en/v0.8.5/control-structures.html#internal-function-calls
[2]: https://docs.soliditylang.org/en/v0.8.5/control-structures.html#external-function-calls
TODO: Functions with no return value should be able to execute
*)
| "expr (CALL i xe) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (CALL i xe) e\<^sub>p e cd);
(case fmlookup e\<^sub>p (address e) of
Some (ct, _) \<Rightarrow>
(case fmlookup ct i of
Some (Method (fp, f, Some x)) \<Rightarrow>
let e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)
in (do {
st' \<leftarrow> applyf (\<lambda>st. st\<lparr>stack:=emptyStore\<rparr>);
(e'', cd', st'') \<leftarrow> load False fp xe e\<^sub>p e' emptyStore st' e cd;
st''' \<leftarrow> get;
put st'';
stmt f e\<^sub>p e'' cd';
rv \<leftarrow> expr x e\<^sub>p e'' cd';
modify (\<lambda>st. st\<lparr>stack:=stack st''', memory := memory st'''\<rparr>);
return rv
})
| _ \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err)
}) st"
| "expr (ECALL ad i xe val) e\<^sub>p e cd st =
(do {
gascheck (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd);
kad \<leftarrow> expr ad e\<^sub>p e cd;
(case kad of
(KValue adv, Value TAddr) \<Rightarrow>
(case fmlookup e\<^sub>p adv of
Some (ct, _) \<Rightarrow>
(case fmlookup ct i of
Some (Method (fp, f, Some x)) \<Rightarrow>
(do {
kv \<leftarrow> expr val e\<^sub>p e cd;
(case kv of
(KValue v, Value t) \<Rightarrow>
let e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)
in (do {
st' \<leftarrow> applyf (\<lambda>st. st\<lparr>stack:=emptyStore, memory:=emptyStore\<rparr>);
(e'', cd', st'') \<leftarrow> load True fp xe e\<^sub>p e' emptyStore st' e cd;
st''' \<leftarrow> get;
(case transfer (address e) adv v (accounts st'') of
Some acc \<Rightarrow>
do {
put (st''\<lparr>accounts := acc\<rparr>);
stmt f e\<^sub>p e'' cd';
rv \<leftarrow> expr x e\<^sub>p e'' cd';
modify (\<lambda>st. st\<lparr>stack:=stack st''', memory := memory st'''\<rparr>);
return rv
}
| None \<Rightarrow> throw Err)
})
| _ \<Rightarrow> throw Err)
})
| _ \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err)
| _ \<Rightarrow> throw Err)
}) st"
| "load cp ((i\<^sub>p, t\<^sub>p)#pl) (e#el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st =
(do {
(v, t) \<leftarrow> expr e e\<^sub>p e\<^sub>v cd;
st'' \<leftarrow> get;
put st';
(cd'', e\<^sub>v'') \<leftarrow> decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st'') cd' e\<^sub>v';
st''' \<leftarrow> get;
put st'';
load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd
}) st"
| "load _ [] (_#_) _ _ _ _ _ _ st = throw Err st"
| "load _ (_#_) [] _ _ _ _ _ _ st = throw Err st"
| "load _ [] [] _ e\<^sub>v' cd' st' e\<^sub>v cd st = return (e\<^sub>v', cd', st') st"
(*TODO: Should be possible to simplify*)
| "rexp (Id i) e\<^sub>p e cd st =
(case fmlookup (denvalue e) i of
Some (tp, Stackloc l) \<Rightarrow>
do {
s \<leftarrow> applyf (\<lambda>st. accessStore l (stack st));
(case s of
Some (KValue v) \<Rightarrow> return (KValue v, tp)
| Some (KCDptr p) \<Rightarrow> return (KCDptr p, tp)
| Some (KMemptr p) \<Rightarrow> return (KMemptr p, tp)
| Some (KStoptr p) \<Rightarrow> return (KStoptr p, tp)
| _ \<Rightarrow> throw Err)
}
| Some (Storage (STValue t), Storeloc l) \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address e));
(case so of
Some s \<Rightarrow> return (KValue (accessStorage t l s), Value t)
| None \<Rightarrow> throw Err)
}
| Some (Storage (STArray x t), Storeloc l) \<Rightarrow> return (KStoptr l, Storage (STArray x t))
| _ \<Rightarrow> throw Err) st"
| "rexp (Ref i r) e\<^sub>p e cd st =
(case fmlookup (denvalue e) i of
Some (tp, (Stackloc l)) \<Rightarrow>
do {
kv \<leftarrow> applyf (\<lambda>st. accessStore l (stack st));
(case kv of
Some (KCDptr l') \<Rightarrow>
(case tp of
Calldata t \<Rightarrow>
do {
(l'', t') \<leftarrow> msel False t l' r e\<^sub>p e cd;
(case t' of
MTValue t'' \<Rightarrow>
(case accessStore l'' cd of
Some (MValue v) \<Rightarrow> return (KValue v, Value t'')
| _ \<Rightarrow> throw Err)
| MTArray x t'' \<Rightarrow>
(case accessStore l'' cd of
Some (MPointer p) \<Rightarrow> return (KCDptr p, Calldata (MTArray x t''))
| _ \<Rightarrow> throw Err))
}
| _ \<Rightarrow> throw Err)
| Some (KMemptr l') \<Rightarrow>
(case tp of
Memory t \<Rightarrow>
do {
(l'', t') \<leftarrow> msel True t l' r e\<^sub>p e cd;
(case t' of
MTValue t'' \<Rightarrow>
do {
mv \<leftarrow> applyf (\<lambda>st. accessStore l'' (memory st));
(case mv of
Some (MValue v) \<Rightarrow> return (KValue v, Value t'')
| _ \<Rightarrow> throw Err)
}
| MTArray x t'' \<Rightarrow>
do {
mv \<leftarrow> applyf (\<lambda>st. accessStore l'' (memory st));
(case mv of
Some (MPointer p) \<Rightarrow> return (KMemptr p, Memory (MTArray x t''))
| _ \<Rightarrow> throw Err)
}
)
}
| _ \<Rightarrow> throw Err)
| Some (KStoptr l') \<Rightarrow>
(case tp of
Storage t \<Rightarrow>
do {
(l'', t') \<leftarrow> ssel t l' r e\<^sub>p e cd;
(case t' of
STValue t'' \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address e));
(case so of
Some s \<Rightarrow> return (KValue (accessStorage t'' l'' s), Value t'')
| None \<Rightarrow> throw Err)
}
| STArray _ _ \<Rightarrow> return (KStoptr l'', Storage t')
| STMap _ _ \<Rightarrow> return (KStoptr l'', Storage t'))
}
| _ \<Rightarrow> throw Err)
| _ \<Rightarrow> throw Err)
}
| Some (tp, (Storeloc l)) \<Rightarrow>
(case tp of
Storage t \<Rightarrow>
do {
(l', t') \<leftarrow> ssel t l r e\<^sub>p e cd;
(case t' of
STValue t'' \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address e));
(case so of
Some s \<Rightarrow> return (KValue (accessStorage t'' l' s), Value t'')
| None \<Rightarrow> throw Err)
}
| STArray _ _ \<Rightarrow> return (KStoptr l', Storage t')
| STMap _ _ \<Rightarrow> return (KStoptr l', Storage t'))
}
| _ \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err) st"
| "stmt SKIP e\<^sub>p e cd st = gascheck (costs SKIP e\<^sub>p e cd) st"
| "stmt (ASSIGN lv ex) e\<^sub>p env cd st =
(do {
gascheck (costs (ASSIGN lv ex) e\<^sub>p env cd);
re \<leftarrow> expr ex e\<^sub>p env cd;
(case re of
(KValue v, Value t) \<Rightarrow>
do {
rl \<leftarrow> lexp lv e\<^sub>p env cd;
(case rl of
(LStackloc l, Value t') \<Rightarrow>
(case convert t t' v of
Some (v', _) \<Rightarrow> modify (\<lambda>st. st \<lparr>stack := updateStore l (KValue v') (stack st)\<rparr>)
| None \<Rightarrow> throw Err)
| (LStoreloc l, Storage (STValue t')) \<Rightarrow>
(case convert t t' v of
Some (v', _) \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow> modify (\<lambda>st. st\<lparr>storage := fmupd (address env) (fmupd l v' s) (storage st)\<rparr>)
| None \<Rightarrow> throw Err)
}
| None \<Rightarrow> throw Err)
| (LMemloc l, Memory (MTValue t')) \<Rightarrow>
(case convert t t' v of
Some (v', _) \<Rightarrow> modify (\<lambda>st. st\<lparr>memory := updateStore l (MValue v') (memory st)\<rparr>)
| None \<Rightarrow> throw Err)
| _ \<Rightarrow> throw Err)
}
| (KCDptr p, Calldata (MTArray x t)) \<Rightarrow>
do {
rl \<leftarrow> lexp lv e\<^sub>p env cd;
(case rl of
(LStackloc l, Memory _) \<Rightarrow> modify (\<lambda>st. st \<lparr>stack := updateStore l (KCDptr p) (stack st)\<rparr>)
| (LStackloc l, Storage _) \<Rightarrow>
do {
sv \<leftarrow> applyf (\<lambda>st. accessStore l (stack st));
(case sv of
Some (KStoptr p') \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow>
(case cpm2s p p' x t cd s of
Some s' \<Rightarrow> modify (\<lambda>st. st \<lparr>storage := fmupd (address env) s' (storage st)\<rparr>)
| None \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err)
}
| _ \<Rightarrow> throw Err)
}
| (LStoreloc l, _) \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow>
(case cpm2s p l x t cd s of
Some s' \<Rightarrow> modify (\<lambda>st. st \<lparr>storage := fmupd (address env) s' (storage st)\<rparr>)
| None \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err)
}
| (LMemloc l, _) \<Rightarrow>
do {
cs \<leftarrow> applyf (\<lambda>st. cpm2m p l x t cd (memory st));
(case cs of
Some m \<Rightarrow> modify (\<lambda>st. st \<lparr>memory := m\<rparr>)
| None \<Rightarrow> throw Err)
}
| _ \<Rightarrow> throw Err)
}
| (KMemptr p, Memory (MTArray x t)) \<Rightarrow>
do {
rl \<leftarrow> lexp lv e\<^sub>p env cd;
(case rl of
(LStackloc l, Memory _) \<Rightarrow> modify (\<lambda>st. st\<lparr>stack := updateStore l (KMemptr p) (stack st)\<rparr>)
| (LStackloc l, Storage _) \<Rightarrow>
do {
sv \<leftarrow> applyf (\<lambda>st. accessStore l (stack st));
(case sv of
Some (KStoptr p') \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow>
do {
cs \<leftarrow> applyf (\<lambda>st. cpm2s p p' x t (memory st) s);
(case cs of
Some s' \<Rightarrow> modify (\<lambda>st. st \<lparr>storage := fmupd (address env) s' (storage st)\<rparr>)
| None \<Rightarrow> throw Err)
}
| None \<Rightarrow> throw Err)
}
| _ \<Rightarrow> throw Err)
}
| (LStoreloc l, _) \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow>
do {
cs \<leftarrow> applyf (\<lambda>st. cpm2s p l x t (memory st) s);
(case cs of
Some s' \<Rightarrow> modify (\<lambda>st. st \<lparr>storage := fmupd (address env) s' (storage st)\<rparr>)
| None \<Rightarrow> throw Err)
}
| None \<Rightarrow> throw Err)
}
| (LMemloc l, _) \<Rightarrow> modify (\<lambda>st. st \<lparr>memory := updateStore l (MPointer p) (memory st)\<rparr>)
| _ \<Rightarrow> throw Err)
}
| (KStoptr p, Storage (STArray x t)) \<Rightarrow>
do {
rl \<leftarrow> lexp lv e\<^sub>p env cd;
(case rl of
(LStackloc l, Memory _) \<Rightarrow>
do {
sv \<leftarrow> applyf (\<lambda>st. accessStore l (stack st));
(case sv of
Some (KMemptr p') \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow>
do {
cs \<leftarrow> applyf (\<lambda>st. cps2m p p' x t s (memory st));
(case cs of
Some m \<Rightarrow> modify (\<lambda>st. st\<lparr>memory := m\<rparr>)
| None \<Rightarrow> throw Err)
}
| None \<Rightarrow> throw Err)
}
| _ \<Rightarrow> throw Err)
}
| (LStackloc l, Storage _) \<Rightarrow> modify (\<lambda>st. st\<lparr>stack := updateStore l (KStoptr p) (stack st)\<rparr>)
| (LStoreloc l, _) \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow>
(case copy p l x t s of
Some s' \<Rightarrow> modify (\<lambda>st. st \<lparr>storage := fmupd (address env) s' (storage st)\<rparr>)
| None \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err)
}
| (LMemloc l, _) \<Rightarrow>
do {
so \<leftarrow> applyf (\<lambda>st. fmlookup (storage st) (address env));
(case so of
Some s \<Rightarrow>
do {
cs \<leftarrow> applyf (\<lambda>st. cps2m p l x t s (memory st));
(case cs of
Some m \<Rightarrow> modify (\<lambda>st. st\<lparr>memory := m\<rparr>)
| None \<Rightarrow> throw Err)
}
| None \<Rightarrow> throw Err)
}
| _ \<Rightarrow> throw Err)
}
| (KStoptr p, Storage (STMap t t')) \<Rightarrow>
do {
rl \<leftarrow> lexp lv e\<^sub>p env cd;
(case rl of
(LStackloc l, _) \<Rightarrow> modify (\<lambda>st. st\<lparr>stack := updateStore l (KStoptr p) (stack st)\<rparr>)
| _ \<Rightarrow> throw Err)
}
| _ \<Rightarrow> throw Err)
}) st"
| "stmt (COMP s1 s2) e\<^sub>p e cd st =
(do {
gascheck (costs (COMP s1 s2) e\<^sub>p e cd);
stmt s1 e\<^sub>p e cd;
stmt s2 e\<^sub>p e cd
}) st"
| "stmt (ITE ex s1 s2) e\<^sub>p e cd st =
(do {
gascheck (costs (ITE ex s1 s2) e\<^sub>p e cd);
v \<leftarrow> expr ex e\<^sub>p e cd;
(case v of
(KValue b, Value TBool) \<Rightarrow>
(if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True
then stmt s1 e\<^sub>p e cd
else stmt s2 e\<^sub>p e cd)
| _ \<Rightarrow> throw Err)
}) st"
| "stmt (WHILE ex s0) e\<^sub>p e cd st =
(do {
gascheck (costs (WHILE ex s0) e\<^sub>p e cd);
v \<leftarrow> expr ex e\<^sub>p e cd;
(case v of
(KValue b, Value TBool) \<Rightarrow>
(if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True
then do {
stmt s0 e\<^sub>p e cd;
stmt (WHILE ex s0) e\<^sub>p e cd
}
else return ())
| _ \<Rightarrow> throw Err)
}) st"
| "stmt (INVOKE i xe) e\<^sub>p e cd st =
(do {
gascheck (costs (INVOKE i xe) e\<^sub>p e cd);
(case fmlookup e\<^sub>p (address e) of
Some (ct, _) \<Rightarrow>
(case fmlookup ct i of
Some (Method (fp, f, None)) \<Rightarrow>
(let e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)
in (do {
st' \<leftarrow> applyf (\<lambda>st. (st\<lparr>stack:=emptyStore\<rparr>));
(e'', cd', st'') \<leftarrow> load False fp xe e\<^sub>p e' emptyStore st' e cd;
st''' \<leftarrow> get;
put st'';
stmt f e\<^sub>p e'' cd';
modify (\<lambda>st. st\<lparr>stack:=stack st''', memory := memory st'''\<rparr>)
}))
| _ \<Rightarrow> throw Err)
| None \<Rightarrow> throw Err)
}) st"
(*External Method calls allow to send some money val with it*)
(*However this transfer does NOT trigger a fallback*)
| "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st =
(do {
gascheck (costs (EXTERNAL ad i xe val) e\<^sub>p e cd);
kad \<leftarrow> expr ad e\<^sub>p e cd;
(case kad of
(KValue adv, Value TAddr) \<Rightarrow>
(case fmlookup e\<^sub>p adv of
Some (ct, fb) \<Rightarrow>
(do {
kv \<leftarrow> expr val e\<^sub>p e cd;
(case kv of
(KValue v, Value t) \<Rightarrow>
(case fmlookup ct i of
Some (Method (fp, f, None)) \<Rightarrow>
let e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)
in (do {
st' \<leftarrow> applyf (\<lambda>st. st\<lparr>stack:=emptyStore, memory:=emptyStore\<rparr>);
(e'', cd', st'') \<leftarrow> load True fp xe e\<^sub>p e' emptyStore st' e cd;
st''' \<leftarrow> get;
(case transfer (address e) adv v (accounts st'') of
Some acc \<Rightarrow>
do {
put (st''\<lparr>accounts := acc\<rparr>);
stmt f e\<^sub>p e'' cd';
modify (\<lambda>st. st\<lparr>stack:=stack st''', memory := memory st'''\<rparr>)
}
| None \<Rightarrow> throw Err)
})
| None \<Rightarrow>
do {
st' \<leftarrow> get;
(case transfer (address e) adv v (accounts st') of
Some acc \<Rightarrow>
do {
st'' \<leftarrow> get;
modify (\<lambda>st. st\<lparr>accounts := acc,stack:=emptyStore, memory:=emptyStore\<rparr>);
stmt fb e\<^sub>p (emptyEnv adv (address e) v) cd;
modify (\<lambda>st. st\<lparr>stack:=stack st'', memory := memory st''\<rparr>)
}
| None \<Rightarrow> throw Err)
}
| _ \<Rightarrow> throw Err)
| _ \<Rightarrow> throw Err)
})
| None \<Rightarrow> throw Err)
| _ \<Rightarrow> throw Err)
}) st"
| "stmt (TRANSFER ad ex) e\<^sub>p e cd st =
(do {
gascheck (costs (TRANSFER ad ex) e\<^sub>p e cd);
kv \<leftarrow> expr ex e\<^sub>p e cd;
(case kv of
(KValue v, Value t) \<Rightarrow>
(do {
kv' \<leftarrow> expr ad e\<^sub>p e cd;
(case kv' of
(KValue adv, Value TAddr) \<Rightarrow>
(do {
acs \<leftarrow> applyf accounts;
(case transfer (address e) adv v acs of
Some acc \<Rightarrow> (case fmlookup e\<^sub>p adv of
Some (ct, f) \<Rightarrow>
let e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)
in (do {
st' \<leftarrow> get;
modify (\<lambda>st. (st\<lparr>accounts := acc, stack:=emptyStore, memory:=emptyStore\<rparr>));
stmt f e\<^sub>p e' emptyStore;
modify (\<lambda>st. st\<lparr>stack:=stack st', memory := memory st'\<rparr>)
})
| None \<Rightarrow> modify (\<lambda>st. (st\<lparr>accounts := acc\<rparr>)))
| None \<Rightarrow> throw Err)
})
| _ \<Rightarrow> throw Err)
})
| _ \<Rightarrow> throw Err)
}) st"
| "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st =
(do {
gascheck (costs (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd);
(case ex of
None \<Rightarrow> (do {
mem \<leftarrow> applyf memory;
(cd', e') \<leftarrow> decl id0 tp None False cd mem cd e\<^sub>v;
stmt s e\<^sub>p e' cd'
})
| Some ex' \<Rightarrow> (do {
(v, t) \<leftarrow> expr ex' e\<^sub>p e\<^sub>v cd;
mem \<leftarrow> applyf memory;
(cd', e') \<leftarrow> decl id0 tp (Some (v, t)) False cd mem cd e\<^sub>v;
stmt s e\<^sub>p e' cd'
}))
}) st"
by pat_completeness auto
subsection \<open>Gas Consumption\<close>
lemma lift_gas:
assumes "lift expr f e1 e2 e\<^sub>p e cd st = Normal ((v, t), st4')"
and "\<And>st4' v4 t4. expr e1 e\<^sub>p e cd st = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas st"
and "\<And>x1 x y xa ya x1a x1b st4' v4 t4. expr e1 e\<^sub>p e cd st = Normal (x, y)
\<Longrightarrow> (xa, ya) = x
\<Longrightarrow> xa = KValue x1a
\<Longrightarrow> ya = Value x1b
\<Longrightarrow> expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4')
\<Longrightarrow> gas st4' \<le> gas y"
shows "gas st4' \<le> gas st"
proof (cases "expr e1 e\<^sub>p e cd st")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue v1)
then show ?thesis
proof (cases c)
case (Value t1)
then show ?thesis
proof (cases "expr e2 e\<^sub>p e cd st'")
case r2: (n a' st'')
then show ?thesis
proof (cases a')
case p2: (Pair b c)
then show ?thesis
proof (cases b)
case v2: (KValue v2)
then show ?thesis
proof (cases c)
case t2: (Value t2)
then show ?thesis
proof (cases "f t1 t2 v1 v2")
case None
with assms n Pair KValue Value r2 p2 v2 t2 show ?thesis by simp
next
case (Some a'')
then show ?thesis
proof (cases a'')
case p3: (Pair v t)
with assms n Pair KValue Value r2 p2 v2 t2 Some have "gas st4'\<le>gas st''" by simp
moreover from assms n Pair KValue Value r2 p2 v2 t2 Some have "gas st''\<le>gas st'" by simp
moreover from assms n Pair KValue Value r2 p2 v2 t2 Some have "gas st'\<le>gas st" by simp
ultimately show ?thesis by arith
qed
qed
next
case (Calldata x2)
with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp
next
case (Memory x3)
with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp
next
case (Storage x4)
with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp
qed
next
case (KCDptr x2)
with assms n Pair KValue Value r2 p2 show ?thesis by simp
next
case (KMemptr x3)
with assms n Pair KValue Value r2 p2 show ?thesis by simp
next
case (KStoptr x4)
with assms n Pair KValue Value r2 p2 show ?thesis by simp
qed
qed
next
case (e x)
with assms n Pair KValue Value show ?thesis by simp
qed
next
case (Calldata x2)
with assms n Pair KValue show ?thesis by simp
next
case (Memory x3)
with assms n Pair KValue show ?thesis by simp
next
case (Storage x4)
with assms n Pair KValue show ?thesis by simp
qed
next
case (KCDptr x2)
with assms n Pair show ?thesis by simp
next
case (KMemptr x3)
with assms n Pair show ?thesis by simp
next
case (KStoptr x4)
with assms n Pair show ?thesis by simp
qed
qed
next
case (e x)
with assms show ?thesis by simp
qed
lemma msel_ssel_lexp_expr_load_rexp_stmt_dom_gas:
"msel_ssel_lexp_expr_load_rexp_stmt_dom (Inl (Inl (c1, t1, l1, xe1, ep1, ev1, cd1, st1)))
\<Longrightarrow> (\<forall>l1' t1' st1'. msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal ((l1', t1'), st1') \<longrightarrow> gas st1' \<le> gas st1)"
"msel_ssel_lexp_expr_load_rexp_stmt_dom (Inl (Inr (Inl (t2, l2, xe2, ep2, ev2, cd2, st2))))
\<Longrightarrow> (\<forall>l2' t2' st2'. ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal ((l2', t2'), st2') \<longrightarrow> gas st2' \<le> gas st2)"
"msel_ssel_lexp_expr_load_rexp_stmt_dom (Inl (Inr (Inr (l5, ep5, ev5, cd5, st5))))
\<Longrightarrow> (\<forall>l5' t5' st5'. lexp l5 ep5 ev5 cd5 st5 = Normal ((l5', t5'), st5') \<longrightarrow> gas st5' \<le> gas st5)"
"msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (e4, ep4, ev4, cd4, st4))))
\<Longrightarrow> (\<forall>st4' v4 t4. expr e4 ep4 ev4 cd4 st4 = Normal ((v4, t4), st4') \<longrightarrow> gas st4' \<le> gas st4)"
"msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (lcp, lis, lxs, lep, lev0, lcd0, lst0, lev, lcd, lst))))
\<Longrightarrow> (\<forall>ev cd st st'. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \<longrightarrow> gas st \<le> gas lst0 \<and> gas st' \<le> gas lst \<and> address ev = address lev0)"
"msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inl (l3, ep3, ev3, cd3, st3))))
\<Longrightarrow> (\<forall>l3' t3' st3'. rexp l3 ep3 ev3 cd3 st3 = Normal ((l3', t3'), st3') \<longrightarrow> gas st3' \<le> gas st3)"
"msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (s6, ep6, ev6, cd6, st6))))
\<Longrightarrow> (\<forall>st6'. stmt s6 ep6 ev6 cd6 st6 = Normal((), st6') \<longrightarrow> gas st6' \<le> gas st6)"
proof (induct rule: msel_ssel_lexp_expr_load_rexp_stmt.pinduct
[where ?P1.0="\<lambda>c1 t1 l1 xe1 ep1 ev1 cd1 st1. (\<forall>l1' t1' st1'. msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal ((l1', t1'), st1') \<longrightarrow> gas st1' \<le> gas st1)"
and ?P2.0="\<lambda>t2 l2 xe2 ep2 ev2 cd2 st2. (\<forall>l2' t2' st2'. ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal ((l2', t2'), st2') \<longrightarrow> gas st2' \<le> gas st2)"
and ?P3.0="\<lambda>l5 ep5 ev5 cd5 st5. (\<forall>l5' t5' st5'. lexp l5 ep5 ev5 cd5 st5 = Normal ((l5', t5'), st5') \<longrightarrow> gas st5' \<le> gas st5)"
and ?P4.0="\<lambda>e4 ep4 ev4 cd4 st4. (\<forall>st4' v4 t4. expr e4 ep4 ev4 cd4 st4 = Normal ((v4, t4), st4') \<longrightarrow> gas st4' \<le> gas st4)"
and ?P5.0="\<lambda>lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst. (\<forall>ev cd st st'. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \<longrightarrow> gas st \<le> gas lst0 \<and> gas st' \<le> gas lst \<and> address ev = address lev0)"
and ?P6.0="\<lambda>l3 ep3 ev3 cd3 st3. (\<forall>l3' t3' st3'. rexp l3 ep3 ev3 cd3 st3 = Normal ((l3', t3'), st3') \<longrightarrow> gas st3' \<le> gas st3)"
and ?P7.0="\<lambda>s6 ep6 ev6 cd6 st6. (\<forall>st6'. stmt s6 ep6 ev6 cd6 st6 = Normal ((), st6') \<longrightarrow> gas st6' \<le> gas st6)"
])
case (1 uu uv uw ux uy uz va)
then show ?case using msel.psimps(1) by auto
next
case (2 vb vc vd ve vf vg vh vi)
then show ?case using msel.psimps(2) by auto
next
case (3 vj al t loc x e\<^sub>p env cd st)
then show ?case using msel.psimps(3) by (auto split: if_split_asm Type.split_asm Stackvalue.split_asm prod.split_asm StateMonad.result.split_asm)
next
case (4 mm al t loc x y ys e\<^sub>p env cd st)
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix l1' t1' st1' assume a1: "msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st = Normal ((l1', t1'), st1')"
show "gas st1' \<le> gas st"
proof (cases "expr x e\<^sub>p env cd st")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue v)
then show ?thesis
proof (cases c)
case (Value t')
then show ?thesis
proof (cases)
assume l: "less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)"
then show ?thesis
proof (cases "accessStore (hash loc v) (if mm then memory st' else cd)")
case None
with 4 a1 n Pair KValue Value l show ?thesis using msel.psimps(4) by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case (MValue x1)
with 4 a1 n Pair KValue Value Some l show ?thesis using msel.psimps(4) by simp
next
case (MPointer l)
with n Pair KValue Value l Some
have "msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st = msel mm t l (y # ys) e\<^sub>p env cd st'"
using msel.psimps(4) 4(1) by simp
moreover from n Pair have "gas st' \<le> gas st" using 4(2) by simp
moreover from a1 MPointer n Pair KValue Value l Some
have "gas st1' \<le> gas st'" using msel.psimps(4) 4(3) 4(1) by simp
ultimately show ?thesis by simp
qed
qed
next
assume "\<not> less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)"
with 4 a1 n Pair KValue Value show ?thesis using msel.psimps(4) by simp
qed
next
case (Calldata x2)
with 4 a1 n Pair KValue show ?thesis using msel.psimps(4) by simp
next
case (Memory x3)
with 4 a1 n Pair KValue show ?thesis using msel.psimps(4) by simp
next
case (Storage x4)
with 4 a1 n Pair KValue show ?thesis using msel.psimps(4) by simp
qed
next
case (KCDptr x2)
with 4 a1 n Pair show ?thesis using msel.psimps(4) by simp
next
case (KMemptr x3)
with 4 a1 n Pair show ?thesis using msel.psimps(4) by simp
next
case (KStoptr x4)
with 4 a1 n Pair show ?thesis using msel.psimps(4) by simp
qed
qed
next
case (e x)
with 4 a1 show ?thesis using msel.psimps(4) by simp
qed
qed
next
case (5 tp loc vk vl vm st)
then show ?case using ssel.psimps(1) by auto
next
case (6 vn vo vp vq vr vs vt vu)
then show ?case using ssel.psimps(2) by auto
next
case (7 al t loc x xs e\<^sub>p env cd st)
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix l2' t2' st2' assume a1: "ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st = Normal ((l2', t2'), st2')"
show "gas st2' \<le> gas st"
proof (cases "expr x e\<^sub>p env cd st")
case (n a st'')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue v)
then show ?thesis
proof (cases c)
case (Value t')
then show ?thesis
proof (cases)
assume l: "less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)"
with n Pair KValue Value l
have "ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st = ssel t (hash loc v) xs e\<^sub>p env cd st''"
using ssel.psimps(3) 7(1) by simp
moreover from n Pair have "gas st'' \<le> gas st" using 7(2) by simp
moreover from a1 n Pair KValue Value l
have "gas st2' \<le> gas st''" using ssel.psimps(3) 7(3) 7(1) by simp
ultimately show ?thesis by simp
next
assume "\<not> less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)"
with 7 a1 n Pair KValue Value show ?thesis using ssel.psimps(3) by simp
qed
next
case (Calldata x2)
with 7 a1 n Pair KValue show ?thesis using ssel.psimps(3) by simp
next
case (Memory x3)
with 7 a1 n Pair KValue show ?thesis using ssel.psimps(3) by simp
next
case (Storage x4)
with 7 a1 n Pair KValue show ?thesis using ssel.psimps(3) by simp
qed
next
case (KCDptr x2)
with 7 a1 n Pair show ?thesis using ssel.psimps(3) by simp
next
case (KMemptr x3)
with 7 a1 n Pair show ?thesis using ssel.psimps(3) by simp
next
case (KStoptr x4)
with 7 a1 n Pair show ?thesis using ssel.psimps(3) by simp
qed
qed
next
case (e e)
with 7 a1 show ?thesis using ssel.psimps(3) by simp
qed
qed
next
case (8 vv t loc x xs e\<^sub>p env cd st)
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix l2' t2' st2' assume a1: "ssel (STMap vv t) loc (x # xs) e\<^sub>p env cd st = Normal ((l2', t2'), st2')"
show "gas st2' \<le> gas st"
proof (cases "expr x e\<^sub>p env cd st")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue v)
with 8 n Pair have "ssel (STMap vv t) loc (x # xs) e\<^sub>p env cd st = ssel t (hash loc v) xs e\<^sub>p env cd st'" using ssel.psimps(4) by simp
moreover from n Pair have "gas st' \<le> gas st" using 8(2) by simp
moreover from a1 n Pair KValue
have "gas st2' \<le> gas st'" using ssel.psimps(4) 8(3) 8(1) by simp
ultimately show ?thesis by simp
next
case (KCDptr x2)
with 8 a1 n Pair show ?thesis using ssel.psimps(4) by simp
next
case (KMemptr x3)
with 8 a1 n Pair show ?thesis using ssel.psimps(4) by simp
next
case (KStoptr x4)
with 8 a1 n Pair show ?thesis using ssel.psimps(4) by simp
qed
qed
next
case (e x)
with 8 a1 show ?thesis using ssel.psimps(4) by simp
qed
qed
next
case (9 i vw e vx st)
then show ?case using lexp.psimps(1)[of i vw e vx st] by (simp split: option.split_asm Denvalue.split_asm prod.split_asm)
next
case (10 i r e\<^sub>p e cd st)
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix st5' xa xaa
assume a1: "lexp (Ref i r) e\<^sub>p e cd st = Normal ((st5', xa), xaa)"
then show "gas xaa \<le> gas st"
proof (cases "fmlookup (denvalue e) i")
case None
with 10 a1 show ?thesis using lexp.psimps(2) by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case (Pair tp b)
then show ?thesis
proof (cases b)
case (Stackloc l)
then show ?thesis
proof (cases "accessStore l (stack st)")
case None
with 10 a1 Some Pair Stackloc show ?thesis using lexp.psimps(2) by simp
next
case s2: (Some a)
then show ?thesis
proof (cases a)
case (KValue x1)
with 10 a1 Some Pair Stackloc s2 show ?thesis using lexp.psimps(2) by simp
next
case (KCDptr x2)
with 10 a1 Some Pair Stackloc s2 show ?thesis using lexp.psimps(2) by simp
next
case (KMemptr l')
then show ?thesis
proof (cases tp)
case (Value x1)
with 10 a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.psimps(2) by simp
next
case (Calldata x2)
with 10 a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.psimps(2) by simp
next
case (Memory t)
then show ?thesis
proof (cases "msel True t l' r e\<^sub>p e cd st")
case (n a s)
with 10 a1 Some Pair Stackloc s2 KMemptr Memory show ?thesis using lexp.psimps(2) by (simp split: prod.split_asm)
next
case (e e)
with 10 a1 Some Pair Stackloc s2 KMemptr Memory show ?thesis using lexp.psimps(2) by simp
qed
next
case (Storage x4)
with 10 a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.psimps(2) by simp
qed
next
case (KStoptr l')
then show ?thesis
proof (cases tp)
case (Value x1)
with 10 a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp
next
case (Calldata x2)
with 10 a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp
next
case (Memory t)
with 10 a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp
next
case (Storage t)
then show ?thesis
proof (cases "ssel t l' r e\<^sub>p e cd st")
case (n a s)
with 10 a1 Some Pair Stackloc s2 KStoptr Storage show ?thesis using lexp.psimps(2) by (auto split: prod.split_asm)
next
case (e x)
with 10 a1 Some Pair Stackloc s2 KStoptr Storage show ?thesis using lexp.psimps(2) by simp
qed
qed
qed
qed
next
case (Storeloc l)
then show ?thesis
proof (cases tp)
case (Value x1)
with 10 a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp
next
case (Calldata x2)
with 10 a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp
next
case (Memory t)
with 10 a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp
next
case (Storage t)
then show ?thesis
proof (cases "ssel t l r e\<^sub>p e cd st")
case (n a s)
with 10 a1 Some Pair Storeloc Storage show ?thesis using lexp.psimps(2) by (auto split: prod.split_asm)
next
case (e x)
with 10 a1 Some Pair Storeloc Storage show ?thesis using lexp.psimps(2) by simp
qed
qed
qed
qed
qed
qed
next
case (11 b x e\<^sub>p e vy st)
then show ?case using expr.psimps(1) by (simp split:if_split_asm)
next
case (12 b x e\<^sub>p e vz st)
then show ?case using expr.psimps(2) by (simp split:if_split_asm)
next
case (13 ad e\<^sub>p e wa st)
then show ?case using expr.psimps(3) by simp
next
case (14 ad e\<^sub>p e wb st)
define g where "g = costs\<^sub>e (BALANCE ad) e\<^sub>p e wb st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix t4 xa xaa
assume *: "expr (BALANCE ad) e\<^sub>p e wb st = Normal ((xa, xaa), t4)"
show "gas t4 \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 14 g_def * show ?thesis using expr.psimps(4) by simp
next
assume gcost: "\<not> gas st \<le> g"
then show ?thesis
proof (cases "expr ad e\<^sub>p e wb (st\<lparr>gas := gas st - g\<rparr>)")
case (n a s)
show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue x1)
then show ?thesis
proof (cases c)
case (Value x1)
then show ?thesis
proof (cases x1)
case (TSInt x1)
with 14 g_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
next
case (TUInt x2)
with 14 g_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
next
case TBool
with 14 g_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
next
case TAddr
with 14 g_def * gcost n Pair KValue Value show "gas t4 \<le> gas st" using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
qed
next
case (Calldata x2)
with 14 g_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
next
case (Memory x3)
with 14 g_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
next
case (Storage x4)
with 14 g_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
qed
next
case (KCDptr x2)
with 14 g_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
next
case (KMemptr x3)
with 14 g_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
next
case (KStoptr x4)
with 14 g_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
qed
qed
next
case (e _)
with 14 g_def * gcost show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp
qed
qed
qed
next
case (15 e\<^sub>p e wc st)
then show ?case using expr.psimps(5) by simp
next
case (16 e\<^sub>p e wd st)
then show ?case using expr.psimps(6) by simp
next
case (17 e\<^sub>p e wd st)
then show ?case using expr.psimps(7) by simp
next
case (18 e\<^sub>p e wd st)
then show ?case using expr.psimps(8) by simp
next
case (19 e\<^sub>p e wd st)
then show ?case using expr.psimps(9) by simp
next
case (20 x e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (NOT x) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix st4' v4 t4 assume a1: "expr (NOT x) e\<^sub>p e cd st = Normal ((v4, t4), st4')"
show "gas st4' \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 20 g_def a1 show ?thesis using expr.psimps by simp
next
assume gcost: "\<not> gas st \<le> g"
then show ?thesis
proof (cases "expr x e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue v)
then show ?thesis
proof (cases c)
case (Value t)
then show ?thesis
proof (cases)
assume v_def: "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True"
with 20(1) g_def gcost n Pair KValue Value have "expr (NOT x) e\<^sub>p e cd st = expr FALSE e\<^sub>p e cd st'" using expr.psimps(10) by simp
moreover from 20(2) g_def gcost n Pair have "gas st' \<le> gas st" by simp
moreover from 20(1) 20(3) a1 g_def gcost n Pair KValue Value v_def have "gas st4' \<le> gas st'" using expr.psimps(10) by simp
ultimately show ?thesis by arith
next
assume v_def: "\<not> v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True"
then show ?thesis
proof (cases)
assume v_def2: "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False"
with 20(1) g_def gcost n Pair KValue Value v_def have "expr (NOT x) e\<^sub>p e cd st = expr TRUE e\<^sub>p e cd st'" using expr.psimps(10) by simp
moreover from 20(2) g_def gcost n Pair have "gas st' \<le> gas st" by simp
moreover from 20 a1 g_def gcost n Pair KValue Value v_def v_def2 have "gas st4' \<le> gas st'" using expr.psimps(10) by simp
ultimately show ?thesis by arith
next
assume "\<not> v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False"
with 20 a1 g_def gcost n Pair KValue Value v_def show ?thesis using expr.psimps(10) by simp
qed
qed
next
case (Calldata x2)
with 20 a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(10) by simp
next
case (Memory x3)
with 20 a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(10) by simp
next
case (Storage x4)
with 20 a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(10) by simp
qed
next
case (KCDptr x2)
with 20 a1 g_def gcost n Pair show ?thesis using expr.psimps(10) by simp
next
case (KMemptr x3)
with 20 a1 g_def gcost n Pair show ?thesis using expr.psimps(10) by simp
next
case (KStoptr x4)
with 20 a1 g_def gcost n Pair show ?thesis using expr.psimps(10) by simp
qed
qed
next
case (e e)
with 20 a1 g_def gcost show ?thesis using expr.psimps(10) by simp
qed
qed
qed
next
case (21 e1 e2 e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (PLUS e1 e2) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix t4 xa xaa assume e_def: "expr (PLUS e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)"
then show "gas t4 \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 21(1) e_def show ?thesis using expr.psimps(11) g_def by simp
next
assume "\<not> gas st \<le> g"
with 21(1) e_def g_def have "lift expr add e1 e2 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((xa, xaa), t4)" using expr.psimps(11)[of e1 e2 e\<^sub>p e cd st] by simp
moreover from 21(2) `\<not> gas st \<le> g` g_def have "(\<And>st4' v4 t4. expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas (st\<lparr>gas := gas st - g\<rparr>))" by simp
moreover from 21(3) `\<not> gas st \<le> g` g_def have "(\<And>x1 x y xa ya x1a x1b st4' v4 t4.
expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal (x, y) \<Longrightarrow>
(xa, ya) = x \<Longrightarrow>
xa = KValue x1a \<Longrightarrow>
ya = Value x1b \<Longrightarrow> expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas y)" by auto
ultimately show "gas t4 \<le> gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "add" "st\<lparr>gas := gas st - g\<rparr>" xa xaa t4] by simp
qed
qed
next
case (22 e1 e2 e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (MINUS e1 e2) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix t4 xa xaa assume e_def: "expr (MINUS e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)"
then show "gas t4 \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 22(1) e_def show ?thesis using expr.psimps(12) g_def by simp
next
assume "\<not> gas st \<le> g"
with 22(1) e_def g_def have "lift expr sub e1 e2 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((xa, xaa), t4)" using expr.psimps(12)[of e1 e2 e\<^sub>p e cd st] by simp
moreover from 22(2) `\<not> gas st \<le> g` g_def have "(\<And>st4' v4 t4. expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas (st\<lparr>gas := gas st - g\<rparr>))" by simp
moreover from 22(3) `\<not> gas st \<le> g` g_def have "(\<And>x1 x y xa ya x1a x1b st4' v4 t4.
expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal (x, y) \<Longrightarrow>
(xa, ya) = x \<Longrightarrow>
xa = KValue x1a \<Longrightarrow>
ya = Value x1b \<Longrightarrow> expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas y)" by auto
ultimately show "gas t4 \<le> gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "sub" "st\<lparr>gas := gas st - g\<rparr>" xa xaa t4] by simp
qed
qed
next
case (23 e1 e2 e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (LESS e1 e2) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix t4 xa xaa assume e_def: "expr (LESS e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)"
then show "gas t4 \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 23(1) e_def show ?thesis using expr.psimps(13) g_def by simp
next
assume "\<not> gas st \<le> g"
with 23(1) e_def g_def have "lift expr less e1 e2 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((xa, xaa), t4)" using expr.psimps(13)[of e1 e2 e\<^sub>p e cd st] by simp
moreover from 23(2) `\<not> gas st \<le> g` g_def have "(\<And>st4' v4 t4. expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas (st\<lparr>gas := gas st - g\<rparr>))" by simp
moreover from 23(3) `\<not> gas st \<le> g` g_def have "(\<And>x1 x y xa ya x1a x1b st4' v4 t4.
expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal (x, y) \<Longrightarrow>
(xa, ya) = x \<Longrightarrow>
xa = KValue x1a \<Longrightarrow>
ya = Value x1b \<Longrightarrow> expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas y)" by auto
ultimately show "gas t4 \<le> gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "less" "st\<lparr>gas := gas st - g\<rparr>" xa xaa t4] by simp
qed
qed
next
case (24 e1 e2 e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (EQUAL e1 e2) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix t4 xa xaa assume e_def: "expr (EQUAL e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)"
then show "gas t4 \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 24(1) e_def show ?thesis using expr.psimps(14) g_def by simp
next
assume "\<not> gas st \<le> g"
with 24(1) e_def g_def have "lift expr equal e1 e2 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((xa, xaa), t4)" using expr.psimps(14)[of e1 e2 e\<^sub>p e cd st] by simp
moreover from 24(2) `\<not> gas st \<le> g` g_def have "(\<And>st4' v4 t4. expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas (st\<lparr>gas := gas st - g\<rparr>))" by simp
moreover from 24(3) `\<not> gas st \<le> g` g_def have "(\<And>x1 x y xa ya x1a x1b st4' v4 t4.
expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal (x, y) \<Longrightarrow>
(xa, ya) = x \<Longrightarrow>
xa = KValue x1a \<Longrightarrow>
ya = Value x1b \<Longrightarrow> expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas y)" by auto
ultimately show "gas t4 \<le> gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "equal" "st\<lparr>gas := gas st - g\<rparr>" xa xaa t4] by simp
qed
qed
next
case (25 e1 e2 e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (AND e1 e2) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix t4 xa xaa assume e_def: "expr (AND e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)"
then show "gas t4 \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 25(1) e_def show ?thesis using expr.psimps(15) g_def by simp
next
assume "\<not> gas st \<le> g"
with 25(1) e_def g_def have "lift expr vtand e1 e2 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((xa, xaa), t4)" using expr.psimps(15)[of e1 e2 e\<^sub>p e cd st] by simp
moreover from 25(2) `\<not> gas st \<le> g` g_def have "(\<And>st4' v4 t4. expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas (st\<lparr>gas := gas st - g\<rparr>))" by simp
moreover from 25(3) `\<not> gas st \<le> g` g_def have "(\<And>x1 x y xa ya x1a x1b st4' v4 t4.
expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal (x, y) \<Longrightarrow>
(xa, ya) = x \<Longrightarrow>
xa = KValue x1a \<Longrightarrow>
ya = Value x1b \<Longrightarrow> expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas y)" by auto
ultimately show "gas t4 \<le> gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "vtand" "st\<lparr>gas := gas st - g\<rparr>" xa xaa t4] by simp
qed
qed
next
case (26 e1 e2 e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (OR e1 e2) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix t4 xa xaa assume e_def: "expr (OR e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)"
then show "gas t4 \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 26(1) e_def show ?thesis using expr.psimps(16) g_def by simp
next
assume "\<not> gas st \<le> g"
with 26(1) e_def g_def have "lift expr vtor e1 e2 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((xa, xaa), t4)" using expr.psimps(16)[of e1 e2 e\<^sub>p e cd st] by simp
moreover from 26(2) `\<not> gas st \<le> g` g_def have "(\<And>st4' v4 t4. expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas (st\<lparr>gas := gas st - g\<rparr>))" by simp
moreover from 26(3) `\<not> gas st \<le> g` g_def have "(\<And>x1 x y xa ya x1a x1b st4' v4 t4.
expr e1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal (x, y) \<Longrightarrow>
(xa, ya) = x \<Longrightarrow>
xa = KValue x1a \<Longrightarrow>
ya = Value x1b \<Longrightarrow> expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \<Longrightarrow> gas st4' \<le> gas y)" by auto
ultimately show "gas t4 \<le> gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "vtor" "st\<lparr>gas := gas st - g\<rparr>" xa xaa t4] by simp
qed
qed
next
case (27 i e\<^sub>p e cd st)
then show ?case using expr.psimps(17) by (auto split: prod.split_asm option.split_asm StateMonad.result.split_asm)
next
case (28 i xe e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (CALL i xe) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix st4' v4 t4 assume a1: "expr (CALL i xe) e\<^sub>p e cd st = Normal ((v4, t4), st4')"
show "gas st4' \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 28 g_def a1 show ?thesis using expr.psimps by simp
next
assume gcost: "\<not> gas st \<le> g"
then show ?thesis
proof (cases "fmlookup e\<^sub>p (address e)")
case None
with 28(1) a1 g_def gcost show ?thesis using expr.psimps(18) by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case (Pair ct _)
then show ?thesis
proof (cases "fmlookup ct i")
case None
with 28(1) a1 g_def gcost Some Pair show ?thesis using expr.psimps(18) by simp
next
case s1: (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
then show ?thesis
proof (cases x1)
case (fields fp f c)
then show ?thesis
proof (cases c)
case None
with 28(1) a1 g_def gcost Some Pair s1 Method fields show ?thesis using expr.psimps(18) by simp
next
case s2: (Some x)
define st' e'
where "st' = st\<lparr>gas := gas st - g\<rparr>\<lparr>stack:=emptyStore\<rparr>"
and "e' = ffold (init ct) (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)"
then show ?thesis
proof (cases "load False fp xe e\<^sub>p e' emptyStore st' e cd (st\<lparr>gas := gas st - g\<rparr>)")
case s4: (n a st''')
then show ?thesis
proof (cases a)
case f2: (fields e'' cd' st'')
then show ?thesis
proof (cases "stmt f e\<^sub>p e'' cd' st''")
case n2: (n a st'''')
then show ?thesis
proof (cases "expr x e\<^sub>p e'' cd' st''''")
case n3: (n a st''''')
then show ?thesis
proof (cases a)
case p1: (Pair sv tp)
with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2 n3
have "expr (CALL i xe) e\<^sub>p e cd st = Normal ((sv, tp), st'''''\<lparr>stack:=stack st''', memory := memory st'''\<rparr>)" and *: "gas st' \<le> gas (st\<lparr>gas := gas st - g\<rparr>)"
using expr.psimps(18)[of i xe e\<^sub>p e cd st] by (auto simp add: Let_def split: unit.split_asm)
with a1 have "gas st4' \<le> gas st'''''" by auto
also from 28(4)[of "()" "st\<lparr>gas := gas st - g\<rparr>" _ ct] g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2 n3
have "\<dots> \<le> gas st''''" by auto
also from 28(3)[of "()" "st\<lparr>gas := gas st - g\<rparr>" _ ct _ _ x1 fp _ f c x e' st' "st\<lparr>gas := gas st - g\<rparr>" _ st''' e'' _ cd' st'' st''' st''' "()" st''] a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2
have "\<dots> \<le> gas st''" by auto
also have "\<dots> \<le> gas st - g"
proof -
from g_def gcost have "(applyf (costs\<^sub>e (CALL i xe) e\<^sub>p e cd) \<bind> (\<lambda>g. assert Gas (\<lambda>st. gas st \<le> g) (modify (\<lambda>st. st\<lparr>gas := gas st - g\<rparr>)))) st = Normal ((), st\<lparr>gas := gas st - g\<rparr>)" by simp
moreover from e'_def have "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" by simp
moreover from st'_def have "applyf (\<lambda>st. st\<lparr>stack := emptyStore\<rparr>) (st\<lparr>gas := gas st - g\<rparr>) = Normal (st', st\<lparr>gas := gas st - g\<rparr>)" by simp
ultimately have "\<forall>ev cda sta st'a. load False fp xe e\<^sub>p e' emptyStore st' e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((ev, cda, sta), st'a) \<longrightarrow> gas sta \<le> gas st' \<and> gas st'a \<le> gas (st\<lparr>gas := gas st - g\<rparr>) \<and> address ev = address e'" using 28(2)[of "()" "st\<lparr>gas := gas st - g\<rparr>" _ ct _ _ x1 fp "(f,c)" f c x e' st' "st\<lparr>gas := gas st - g\<rparr>"] using Some Pair s1 Method fields s2 by blast
thus ?thesis using st'_def s4 f2 by auto
qed
finally show ?thesis by simp
qed
next
case (e x)
with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2 show ?thesis using expr.psimps(18)[of i xe e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm)
qed
next
case (e x)
with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 show ?thesis using expr.psimps(18)[of i xe e\<^sub>p e cd st] by (auto split:unit.split_asm)
qed
qed
next
case (e x)
with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def show ?thesis using expr.psimps(18)[of i xe e\<^sub>p e cd st] by auto
qed
qed
qed
next
case (Var x2)
with 28(1) a1 g_def gcost Some Pair s1 show ?thesis using expr.psimps(18) by simp
qed
qed
qed
qed
qed
qed
next
case (29 ad i xe val e\<^sub>p e cd st)
define g where "g = costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd st"
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix st4' v4 t4 assume a1: "expr (ECALL ad i xe val) e\<^sub>p e cd st = Normal ((v4, t4), st4')"
show "gas st4' \<le> gas st"
proof (cases)
assume "gas st \<le> g"
with 29 g_def a1 show ?thesis using expr.psimps by simp
next
assume gcost: "\<not> gas st \<le> g"
then show ?thesis
proof (cases "expr ad e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair a b)
then show ?thesis
proof (cases a)
case (KValue adv)
then show ?thesis
proof (cases b)
case (Value x1)
then show ?thesis
proof (cases x1)
case (TSInt x1)
with 29(1) a1 g_def gcost n Pair KValue Value show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (TUInt x2)
with 29(1) a1 g_def gcost n Pair KValue Value show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case TBool
with 29(1) a1 g_def gcost n Pair KValue Value show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case TAddr
then show ?thesis
proof (cases "fmlookup e\<^sub>p adv")
case None
with 29(1) a1 g_def gcost n Pair KValue Value TAddr show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case p2: (Pair ct _)
then show ?thesis
proof (cases "fmlookup ct i")
case None
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 show ?thesis using expr.psimps(19) by simp
next
case s1: (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
then show ?thesis
proof (cases x1)
case (fields fp f c)
then show ?thesis
proof (cases c)
case None
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields show ?thesis using expr.psimps(19) by simp
next
case s2: (Some x)
then show ?thesis
proof (cases "expr val e\<^sub>p e cd st'")
case n1: (n kv st'')
then show ?thesis
proof (cases kv)
case p3: (Pair a b)
then show ?thesis
proof (cases a)
case k1: (KValue v)
then show ?thesis
proof (cases b)
case v1: (Value t)
define stl e'
where "stl = st''\<lparr>stack:=emptyStore, memory:=emptyStore\<rparr>"
and "e' = ffold (init ct) (emptyEnv adv (address e) v) (fmdom ct)"
then show ?thesis
proof (cases "load True fp xe e\<^sub>p e' emptyStore stl e cd st''")
case s4: (n a st''')
then show ?thesis
proof (cases a)
case f2: (fields e'' cd' st'''')
then show ?thesis
proof (cases "transfer (address e) adv v (accounts st'''')")
case n2: None
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 stl_def e'_def s4 f2 show ?thesis using expr.psimps(19) by simp
next
case s3: (Some acc)
show ?thesis
proof (cases "stmt f e\<^sub>p e'' cd' (st''''\<lparr>accounts:=acc\<rparr>)")
case n2: (n a st''''')
then show ?thesis
proof (cases "expr x e\<^sub>p e'' cd' st'''''")
case n3: (n a st'''''')
then show ?thesis
proof (cases a)
case p1: (Pair sv tp)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3
have "expr (ECALL ad i xe val) e\<^sub>p e cd st = Normal ((sv, tp), st''''''\<lparr>stack:=stack st''', memory := memory st'''\<rparr>)"
using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by (auto simp add: Let_def split: unit.split_asm)
with a1 have "gas st4' \<le> gas st''''''" by auto
also from 29(6)[of "()" "st\<lparr>gas := gas st - g\<rparr>" _ st' _ _ adv _ _ ct _ _ x1 fp "(f,c)" f c x kv st'' _ b v t] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3
have "\<dots> \<le> gas st'''''" by auto
also from 29(5)[OF _ n Pair KValue Value TAddr Some p2 s1 Method fields _ s2 n1 p3 k1 v1 _ _ s4 f2 _ _ _, of "()" f cd' st'''' st''' st''' acc] f2 s3 stl_def e'_def n2 n3 a1 g_def gcost
have "\<dots> \<le> gas (st''''\<lparr>accounts:=acc\<rparr>)" by auto
also have "\<dots> \<le> gas stl"
proof -
from g_def gcost have "(applyf (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd) \<bind> (\<lambda>g. assert Gas (\<lambda>st. gas st \<le> g) (modify (\<lambda>st. st\<lparr>gas := gas st - g\<rparr>)))) st = Normal ((), st\<lparr>gas := gas st - g\<rparr>)" by simp
moreover from e'_def have "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" by simp
moreover from n1 have "expr val e\<^sub>p e cd st' = Normal (kv, st'')" by simp
moreover from stl_def have "applyf (\<lambda>st. st\<lparr>stack := emptyStore, memory := emptyStore\<rparr>) st'' = Normal (stl, st'')" by simp
moreover have "applyf accounts st'' = Normal ((accounts st''), st'')" by simp
ultimately have "\<forall>ev cda sta st'a. load True fp xe e\<^sub>p e' emptyStore stl e cd st'' = Normal ((ev, cda, sta), st'a) \<longrightarrow> gas sta \<le> gas stl \<and> gas st'a \<le> gas st'' \<and> address ev = address e'" using 29(4)[of "()" "st\<lparr>gas := gas st - g\<rparr>" _ st' _ _ adv _ _ ct _ _ x1 fp "(f,c)" f c x kv st'' _ b v t] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3 by blast
thus ?thesis using stl_def s4 f2 by auto
qed
also from stl_def have "\<dots> \<le> gas st''" by simp
also from 29(3)[of "()" "st\<lparr>gas := gas st - g\<rparr>" _ st' _ _ adv _ _ ct _ _ x1 fp "(f,c)" f c x] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3
have "\<dots> \<le> gas st'" by (auto split:unit.split_asm)
also from 29(2)[of "()" "st\<lparr>gas := gas st - g\<rparr>"] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3
have "\<dots> \<le> gas (st\<lparr>gas := gas st - g\<rparr>)" by simp
finally show ?thesis by simp
qed
next
case (e x)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
next
case (e x)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
qed
next
case (e x)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 stl_def e'_def show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
next
case (Calldata x2)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Memory x3)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Storage x4)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
next
case (KCDptr x2)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KMemptr x3)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KStoptr x4)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
next
case n2: (e x)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
qed
next
case (Var x2)
with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
qed
qed
qed
next
case (Calldata x2)
with 29(1) a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Memory x3)
with 29(1) a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Storage x4)
with 29(1) a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
next
case (KCDptr x2)
with 29(1) a1 g_def gcost n Pair show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KMemptr x3)
with 29(1) a1 g_def gcost n Pair show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KStoptr x4)
with 29(1) a1 g_def gcost n Pair show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
next
case (e _)
with 29(1) a1 g_def gcost show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
qed
next
case (30 cp i\<^sub>p t\<^sub>p pl e el e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st)
then show ?case
proof (cases "expr e e\<^sub>p e\<^sub>v cd st")
case (n a st'')
then show ?thesis
proof (cases a)
case (Pair v t)
then show ?thesis
proof (cases "decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st'') cd' e\<^sub>v' st'")
case n2: (n a' st''')
then show ?thesis
proof (cases a')
case f2: (Pair cd'' e\<^sub>v'')
show ?thesis
proof (rule allI[THEN allI, THEN allI, THEN allI, OF impI])
fix ev xa xaa xaaa assume load_def: "load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st = Normal ((ev, xa, xaa), xaaa)"
with 30(1) n Pair n2 f2 have "load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st = load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd st''" using load.psimps(1)[of cp i\<^sub>p t\<^sub>p pl e el e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st] by simp
with load_def have "load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd st'' = Normal ((ev, xa, xaa), xaaa)" by simp
with n Pair n2 f2 have "gas xaa \<le> gas st''' \<and> gas xaaa \<le> gas st'' \<and> address ev = address e\<^sub>v''" using 30(3)[of a st'' v t st'' st'' "()" st' a' st''' cd'' e\<^sub>v'' st''' st''' "()" st''] by simp
moreover from n Pair have "gas st'' \<le> gas st" using 30(2) by simp
moreover from n2 f2 have " address e\<^sub>v'' = address e\<^sub>v'" and "gas st''' \<le> gas st'" using decl_gas_address by auto
ultimately show "gas xaa \<le> gas st' \<and> gas xaaa \<le> gas st \<and> address ev = address e\<^sub>v'" by simp
qed
qed
next
case (e x)
with 30(1) n Pair show ?thesis using load.psimps(1) by simp
qed
qed
next
case (e x)
with 30(1) show ?thesis using load.psimps(1) by simp
qed
next
case (31 we wf wg wh wi wj wk st)
then show ?case using load.psimps(2) by auto
next
case (32 wl wm wn wo wp wq wr st)
then show ?case using load.psimps(3)[of wl wm wn wo wp wq wr] by auto
next
case (33 ws wt wu wv cd e\<^sub>v s st)
then show ?case using load.psimps(4)[of ws wt wu wv cd e\<^sub>v s st] by auto
next
case (34 i e\<^sub>p e cd st)
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix st3' xa xaa assume "rexp (L.Id i) e\<^sub>p e cd st = Normal ((st3', xa), xaa)"
then show "gas xaa \<le> gas st" using 34(1) rexp.psimps(1) by (simp split: option.split_asm Denvalue.split_asm Stackvalue.split_asm prod.split_asm Type.split_asm STypes.split_asm)
qed
next
case (35 i r e\<^sub>p e cd st)
show ?case
proof (rule allI[THEN allI, THEN allI, OF impI])
fix st3' xa xaa assume rexp_def: "rexp (Ref i r) e\<^sub>p e cd st = Normal ((st3', xa), xaa)"
show "gas xaa \<le> gas st"
proof (cases "fmlookup (denvalue e) i")
case None
with 35(1) show ?thesis using rexp.psimps rexp_def by simp
next
case (Some a)
then show ?thesis
proof (cases a)
case (Pair tp b)
then show ?thesis
proof (cases b)
case (Stackloc l)
then show ?thesis
proof (cases "accessStore l (stack st)")
case None
with 35(1) Some Pair Stackloc show ?thesis using rexp.psimps(2) rexp_def by simp
next
case s1: (Some a)
then show ?thesis
proof (cases a)
case (KValue x1)
with 35(1) Some Pair Stackloc s1 show ?thesis using rexp.psimps(2) rexp_def by simp
next
case (KCDptr l')
with 35 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e\<^sub>p e cd st] rexp_def by (simp split: option.split_asm Memoryvalue.split_asm MTypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm)
next
case (KMemptr x3)
with 35 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e\<^sub>p e cd st] rexp_def by (simp split: option.split_asm Memoryvalue.split_asm MTypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm)
next
case (KStoptr x4)
with 35 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e\<^sub>p e cd st] rexp_def by (simp split: option.split_asm STypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm)
qed
qed
next
case (Storeloc x2)
with 35 Some Pair show ?thesis using rexp.psimps rexp_def by (simp split: option.split_asm STypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm)
qed
qed
qed
qed
next
case (36 e\<^sub>p e cd st)
then show ?case using stmt.psimps(1) by simp
next
case (37 lv ex e\<^sub>p env cd st)
define g where "g = costs (ASSIGN lv ex) e\<^sub>p env cd st"
show ?case
proof (rule allI[OF impI])
fix st6'
assume stmt_def: "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st6')"
then show "gas st6' \<le> gas st"
proof cases
assume "gas st \<le> g"
with 37 stmt_def show ?thesis using stmt.psimps(2) g_def by simp
next
assume "\<not> gas st \<le> g"
show ?thesis
proof (cases "expr ex e\<^sub>p env cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue v)
then show ?thesis
proof (cases c)
case (Value t)
then show ?thesis
proof (cases "lexp lv e\<^sub>p env cd st'")
case n2: (n a st'')
then show ?thesis
proof (cases a)
case p1: (Pair a b)
then show ?thesis
proof (cases a)
case (LStackloc l)
then show ?thesis
proof (cases b)
case v2: (Value t')
then show ?thesis
proof (cases "convert t t' v ")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStackloc v2 show ?thesis using stmt.psimps(2) g_def by simp
next
case s3: (Some a)
then show ?thesis
proof (cases a)
case p2: (Pair v' b)
with 37(1) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStackloc v2 s3
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>stack := updateStore l (KValue v') (stack st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>stack := updateStore l (KValue v') (stack st'')\<rparr>)" by simp
moreover from 37(3) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair KValue Value n2 p2 have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
next
case (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Storage x4)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (LMemloc l)
then show ?thesis
proof (cases b)
case v2: (Value t')
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
then show ?thesis
proof (cases x3)
case (MTArray x11 x12)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LMemloc Memory show ?thesis using stmt.psimps(2) g_def by simp
next
case (MTValue t')
then show ?thesis
proof (cases "convert t t' v ")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LMemloc Memory MTValue show ?thesis using stmt.psimps(2) g_def by simp
next
case s3: (Some a)
then show ?thesis
proof (cases a)
case p2: (Pair v' b)
with 37(1) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LMemloc Memory MTValue s3
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>memory := updateStore l (MValue v') (memory st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>memory := updateStore l (MValue v') (memory st'')\<rparr>)" by simp
moreover from 37(3) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
qed
next
case (Storage x4)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (LStoreloc l)
then show ?thesis
proof (cases b)
case v2: (Value t')
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Storage x4)
then show ?thesis
proof (cases x4)
case (STArray x11 x12)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc Storage show ?thesis using stmt.psimps(2) g_def by simp
next
case (STMap x21 x22)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc Storage show ?thesis using stmt.psimps(2) g_def by simp
next
case (STValue t')
then show ?thesis
proof (cases "convert t t' v ")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc Storage STValue show ?thesis using stmt.psimps(2) g_def by simp
next
case s3: (Some a)
then show ?thesis
proof (cases a)
case p2: (Pair v' b)
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc Storage STValue s3 p2 show ?thesis using stmt.psimps(2) g_def by simp
next
case s4: (Some s)
with 37(1) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 LStoreloc Storage STValue s3 p2
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>storage := fmupd (address env) (fmupd l v' s) (storage st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= st'' \<lparr>storage := fmupd (address env) (fmupd l v' s) (storage st'')\<rparr>" by simp
moreover from 37(3) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair KValue Value n2 p1 have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
qed
qed
qed
qed
qed
next
case (e x)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue show ?thesis using stmt.psimps(2) g_def by simp
next
case (Storage x4)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KValue show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (KCDptr p)
then show ?thesis
proof (cases c)
case (Value x1)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def by simp
next
case (Calldata x2)
then show ?thesis
proof (cases x2)
case (MTArray x t)
then show ?thesis
proof (cases "lexp lv e\<^sub>p env cd st'")
case n2: (n a st'')
then show ?thesis
proof (cases a)
case p2: (Pair a b)
then show ?thesis
proof (cases a)
case (LStackloc l)
then show ?thesis
proof (cases b)
case v2: (Value t')
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case c2: (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
with 37(1) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>stack := updateStore l (KCDptr p) (stack st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>stack := updateStore l (KCDptr p) (stack st'')\<rparr>)" by simp
moreover from 37(4) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
next
case (Storage x4)
then show ?thesis
proof (cases "accessStore l (stack st'')")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage show ?thesis using stmt.psimps(2) g_def by simp
next
case s3: (Some a)
then show ?thesis
proof (cases a)
case (KValue x1)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case c3: (KCDptr x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case (KMemptr x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case (KStoptr p')
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr show ?thesis using stmt.psimps(2) g_def by simp
next
case s4: (Some s)
then show ?thesis
proof (cases "cpm2s p p' x t cd s")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr s4 show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some s')
with 37(1) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr s4
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= st'' \<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>" by simp
moreover from 37(4) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
qed
qed
qed
next
case (LMemloc l)
then show ?thesis
proof (cases "cpm2m p l x t cd (memory st'')")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LMemloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some m)
with 37(1) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LMemloc
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>memory := m\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>memory := m\<rparr>)" by simp
moreover from 37(4) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
next
case (LStoreloc l)
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp
next
case s4: (Some s)
then show ?thesis
proof (cases "cpm2s p l x t cd s")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc s4 show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some s')
with 37(1) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc s4
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>)" by simp
moreover from 37(4) `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
qed
qed
next
case (e x)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata MTArray show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (MTValue x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr Calldata show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (Memory x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def by simp
next
case (Storage x4)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (KMemptr p)
then show ?thesis
proof (cases c)
case (Value x1)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def by simp
next
case (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
then show ?thesis
proof (cases x3)
case (MTArray x t)
then show ?thesis
proof (cases "lexp lv e\<^sub>p env cd st'")
case n2: (n a st'')
then show ?thesis
proof (cases a)
case p2: (Pair a b)
then show ?thesis
proof (cases a)
case (LStackloc l)
then show ?thesis
proof (cases b)
case v2: (Value t')
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case c2: (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case m2: (Memory x3)
with 37(1) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>stack := updateStore l (KMemptr p) (stack st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>stack := updateStore l (KMemptr p) (stack st'')\<rparr>)" by simp
moreover from 37(5) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
next
case (Storage x4)
then show ?thesis
proof (cases "accessStore l (stack st'')")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage show ?thesis using stmt.psimps(2) g_def by simp
next
case s3: (Some a)
then show ?thesis
proof (cases a)
case (KValue x1)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case c3: (KCDptr x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case m3: (KMemptr x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case (KStoptr p')
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr show ?thesis using stmt.psimps(2) g_def by simp
next
case s4: (Some s)
then show ?thesis
proof (cases "cpm2s p p' x t (memory st'') s")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr s4 show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some s')
with 37(1) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr s4
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>)" by simp
moreover from 37(5) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
qed
qed
qed
next
case (LMemloc l)
with 37(1) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LMemloc
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>memory := updateStore l (MPointer p) (memory st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= st'' \<lparr>memory := updateStore l (MPointer p) (memory st'')\<rparr>" by simp
moreover from 37(5) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
next
case (LStoreloc l)
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp
next
case s3: (Some s)
then show ?thesis
proof (cases "cpm2s p l x t (memory st'') s")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some s')
with 37(1) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc s3
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= st''\<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>" by simp
moreover from 37(5) `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
qed
qed
next
case (e x)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory MTArray show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (MTValue x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr Memory show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (Storage x4)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (KStoptr p)
then show ?thesis
proof (cases c)
case (Value x1)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def by simp
next
case (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def by simp
next
case (Storage x4)
then show ?thesis
proof (cases x4)
case (STArray x t)
then show ?thesis
proof (cases "lexp lv e\<^sub>p env cd st'")
case n2: (n a st'')
then show ?thesis
proof (cases a)
case p2: (Pair a b)
then show ?thesis
proof (cases a)
case (LStackloc l)
then show ?thesis
proof (cases b)
case v2: (Value t')
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case c2: (Calldata x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp
next
case (Memory x3)
then show ?thesis
proof (cases "accessStore l (stack st'')")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory show ?thesis using stmt.psimps(2) g_def by simp
next
case s3: (Some a)
then show ?thesis
proof (cases a)
case (KValue x1)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case c3: (KCDptr x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def by simp
next
case (KMemptr p')
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr show ?thesis using stmt.psimps(2) g_def by simp
next
case s4: (Some s)
then show ?thesis
proof (cases "cps2m p p' x t s (memory st'')")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr s4 show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some m)
with 37(1) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr s4
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>memory := m\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>memory := m\<rparr>)" by simp
moreover from 37(6) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
next
case sp2: (KStoptr p')
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def by simp
qed
qed
next
case st2: (Storage x4)
with 37(1) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStackloc
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>stack := updateStore l (KStoptr p) (stack st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>stack := updateStore l (KStoptr p) (stack st'')\<rparr>)" by simp
moreover from 37(6) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
next
case (LMemloc l)
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LMemloc show ?thesis using stmt.psimps(2) g_def by simp
next
case s4: (Some s)
then show ?thesis
proof (cases "cps2m p l x t s (memory st'')")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LMemloc s4 show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some m)
with 37(1) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LMemloc s4
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>memory := m\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= (st''\<lparr>memory := m\<rparr>)" by simp
moreover from 37(6) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
next
case (LStoreloc l)
then show ?thesis
proof (cases "fmlookup (storage st'') (address env)")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp
next
case s4: (Some s)
then show ?thesis
proof (cases "copy p l x t s")
case None
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStoreloc s4 show ?thesis using stmt.psimps(2) g_def by simp
next
case (Some s')
with 37(1) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 LStoreloc s4
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= st''\<lparr>storage := fmupd (address env) s' (storage st'')\<rparr>" by simp
moreover from 37(6) `\<not> gas st \<le> g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
qed
qed
qed
next
case (e x)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STArray show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (STMap t t')
then show ?thesis
proof (cases "lexp lv e\<^sub>p env cd st'")
case n2: (n a st'')
then show ?thesis
proof (cases a)
case p2: (Pair a b)
then show ?thesis
proof (cases a)
case (LStackloc l)
with 37(1) `\<not> gas st \<le> g` n Pair KStoptr Storage STMap n2 p2
have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \<lparr>stack := updateStore l (KStoptr p) (stack st'')\<rparr>)"
using stmt.psimps(2) g_def by simp
with stmt_def have "st6'= st''\<lparr>stack := updateStore l (KStoptr p) (stack st'')\<rparr>" by simp
moreover from 37(7) `\<not> gas st \<le> g` n Pair KStoptr Storage STMap n2 p2 have "gas st'' \<le> gas st'" using g_def by simp
moreover from 37(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
next
case (LMemloc x2)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STMap n2 p2 show ?thesis using stmt.psimps(2) g_def by simp
next
case (LStoreloc x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STMap n2 p2 show ?thesis using stmt.psimps(2) g_def by simp
qed
qed
next
case (e x)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage STMap show ?thesis using stmt.psimps(2) g_def by simp
qed
next
case (STValue x3)
with 37(1) stmt_def `\<not> gas st \<le> g` n Pair KStoptr Storage show ?thesis using stmt.psimps(2) g_def by simp
qed
qed
qed
qed
next
case (e x)
with 37(1) stmt_def `\<not> gas st \<le> g` show ?thesis using stmt.psimps(2) g_def by (simp split: Ex.split_asm)
qed
qed
qed
next
case (38 s1 s2 e\<^sub>p e cd st)
define g where "g = costs (COMP s1 s2) e\<^sub>p e cd st"
show ?case
proof (rule allI[OF impI])
fix st6'
assume stmt_def: "stmt (COMP s1 s2) e\<^sub>p e cd st = Normal ((), st6')"
then show "gas st6' \<le> gas st"
proof cases
assume "gas st \<le> g"
with 38 stmt_def g_def show ?thesis using stmt.psimps(3) by simp
next
assume "\<not> gas st \<le> g"
show ?thesis
proof (cases "stmt s1 e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
with 38(1) stmt_def `\<not> gas st \<le> g` have "stmt (COMP s1 s2) e\<^sub>p e cd st = stmt s2 e\<^sub>p e cd st'" using stmt.psimps(3)[of s1 s2 e\<^sub>p e cd st] g_def by (simp add:Let_def split:unit.split_asm)
with 38(3)[of _ "(st\<lparr>gas := gas st - g\<rparr>)" _ st'] stmt_def \<open>\<not> gas st \<le> g\<close> n have "gas st6' \<le> gas st'" using g_def by fastforce
moreover from 38(2) \<open>\<not> gas st \<le> g\<close> n have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
next
case (e x)
with 38 stmt_def `\<not> gas st \<le> g` show ?thesis using stmt.psimps(3)[of s1 s2 e\<^sub>p e cd st] g_def by (simp split: Ex.split_asm)
qed
qed
qed
next
case (39 ex s1 s2 e\<^sub>p e cd st)
define g where "g = costs (ITE ex s1 s2) e\<^sub>p e cd st"
show ?case
proof (rule allI[OF impI])
fix st6'
assume stmt_def: "stmt (ITE ex s1 s2) e\<^sub>p e cd st = Normal ((), st6')"
then show "gas st6' \<le> gas st"
proof cases
assume "gas st \<le> g"
with 39 stmt_def show ?thesis using stmt.psimps(4) g_def by simp
next
assume "\<not> gas st \<le> g"
show ?thesis
proof (cases "expr ex e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue b)
then show ?thesis
proof (cases c)
case (Value x1)
then show ?thesis
proof (cases x1)
case (TSInt x1)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def by simp
next
case (TUInt x2)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def by simp
next
case TBool
then show ?thesis
proof cases
assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True"
with 39(1) `\<not> gas st \<le> g` n Pair KValue Value TBool have "stmt (ITE ex s1 s2) e\<^sub>p e cd st = stmt s1 e\<^sub>p e cd st'" using stmt.psimps(4) g_def by simp
with 39(3) stmt_def `\<not> gas st \<le> g` n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "gas st6' \<le> gas st'" using g_def by simp
moreover from 39(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by arith
next
assume "\<not> b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True"
with 39(1) `\<not> gas st \<le> g` n Pair KValue Value TBool have "stmt (ITE ex s1 s2) e\<^sub>p e cd st = stmt s2 e\<^sub>p e cd st'" using stmt.psimps(4) g_def by simp
with 39(4) stmt_def `\<not> gas st \<le> g` n Pair KValue Value TBool `\<not> b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "gas st6' \<le> gas st'" using g_def by simp
moreover from 39(2) `\<not> gas st \<le> g` n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by arith
qed
next
case TAddr
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def by simp
qed
next
case (Calldata x2)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair KValue show ?thesis using stmt.psimps(4) g_def by simp
next
case (Memory x3)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair KValue show ?thesis using stmt.psimps(4) g_def by simp
next
case (Storage x4)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair KValue show ?thesis using stmt.psimps(4) g_def by simp
qed
next
case (KCDptr x2)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair show ?thesis using stmt.psimps(4) g_def by simp
next
case (KMemptr x3)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair show ?thesis using stmt.psimps(4) g_def by simp
next
case (KStoptr x4)
with 39(1) stmt_def `\<not> gas st \<le> g` n Pair show ?thesis using stmt.psimps(4) g_def by simp
qed
qed
next
case (e e)
with 39(1) stmt_def `\<not> gas st \<le> g` show ?thesis using stmt.psimps(4) g_def by simp
qed
qed
qed
next
case (40 ex s0 e\<^sub>p e cd st)
define g where "g = costs (WHILE ex s0) e\<^sub>p e cd st"
show ?case
proof (rule allI[OF impI])
fix st6'
assume stmt_def: "stmt (WHILE ex s0) e\<^sub>p e cd st = Normal ((), st6')"
then show "gas st6' \<le> gas st"
proof cases
assume "gas st \<le> costs (WHILE ex s0) e\<^sub>p e cd st"
with 40(1) stmt_def show ?thesis using stmt.psimps(5) by simp
next
assume gcost: "\<not> gas st \<le> costs (WHILE ex s0) e\<^sub>p e cd st"
show ?thesis
proof (cases "expr ex e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue b)
then show ?thesis
proof (cases c)
case (Value x1)
then show ?thesis
proof (cases x1)
case (TSInt x1)
with 40(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def by simp
next
case (TUInt x2)
with 40(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def by simp
next
case TBool
then show ?thesis
proof cases
assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True"
then show ?thesis
proof (cases "stmt s0 e\<^sub>p e cd st'")
case n2: (n a st'')
with 40(1) gcost n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "stmt (WHILE ex s0) e\<^sub>p e cd st = stmt (WHILE ex s0) e\<^sub>p e cd st''" using stmt.psimps(5)[of ex s0 e\<^sub>p e cd st] g_def by (simp add: Let_def split:unit.split_asm)
with 40(4) stmt_def gcost n2 Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` n have "gas st6' \<le> gas st''" using g_def by simp
moreover from 40(3) gcost n2 Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` n have "gas st'' \<le> gas st'" using g_def by simp
moreover from 40(2)[of _ "st\<lparr>gas := gas st - g\<rparr>"] gcost n Pair have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
next
case (e x)
with 40(1) stmt_def gcost n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` show ?thesis using stmt.psimps(5) g_def by (simp split: Ex.split_asm)
qed
next
assume "\<not> b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True"
with 40(1) gcost n Pair KValue Value TBool have "stmt (WHILE ex s0) e\<^sub>p e cd st = return () st'" using stmt.psimps(5) g_def by simp
with stmt_def have "gas st6' \<le> gas st'" by simp
moreover from 40(2)[of _ "st\<lparr>gas := gas st - g\<rparr>"] gcost n have "gas st' \<le> gas st" using g_def by simp
ultimately show ?thesis by simp
qed
next
case TAddr
with 40(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def by simp
qed
next
case (Calldata x2)
with 40(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def by simp
next
case (Memory x3)
with 40(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def by simp
next
case (Storage x4)
with 40(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def by simp
qed
next
case (KCDptr x2)
with 40(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def by simp
next
case (KMemptr x3)
with 40(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def by simp
next
case (KStoptr x4)
with 40(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def by simp
qed
qed
next
case (e e)
with 40(1) stmt_def gcost show ?thesis using stmt.psimps(5) g_def by simp
qed
qed
qed
next
case (41 i xe e\<^sub>p e cd st)
define g where "g = costs (INVOKE i xe) e\<^sub>p e cd st"
show ?case
proof (rule allI[OF impI])
fix st6' assume a1: "stmt (INVOKE i xe) e\<^sub>p e cd st = Normal ((), st6')"
show "gas st6' \<le> gas st"
proof (cases)
assume "gas st \<le> costs (INVOKE i xe) e\<^sub>p e cd st"
with 41(1) a1 show ?thesis using stmt.psimps(6) by simp
next
assume gcost: "\<not> gas st \<le> costs (INVOKE i xe) e\<^sub>p e cd st"
then show ?thesis
proof (cases "fmlookup e\<^sub>p (address e)")
case None
with 41(1) a1 gcost show ?thesis using stmt.psimps(6) by simp
next
case (Some x)
then show ?thesis
proof (cases x)
case (Pair ct _)
then show ?thesis
proof (cases "fmlookup ct i")
case None
with 41(1) g_def a1 gcost Some Pair show ?thesis using stmt.psimps(6) by simp
next
case s1: (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
then show ?thesis
proof (cases x1)
case (fields fp f c)
then show ?thesis
proof (cases c)
case None
define st' e'
where "st' = st\<lparr>gas := gas st - g\<rparr>\<lparr>stack:=emptyStore\<rparr>"
and "e' = ffold (init ct) (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)"
then show ?thesis
proof (cases "load False fp xe e\<^sub>p e' emptyStore st' e cd (st\<lparr>gas := gas st - g\<rparr>)")
case s3: (n a st''')
then show ?thesis
proof (cases a)
case f1: (fields e'' cd' st'')
then show ?thesis
proof (cases "stmt f e\<^sub>p e'' cd' st''")
case n2: (n a st'''')
with 41(1) g_def a1 gcost Some Pair s1 Method fields None st'_def e'_def s3 f1
have "stmt (INVOKE i xe) e\<^sub>p e cd st = Normal ((), st''''\<lparr>stack:=stack st''', memory := memory st'''\<rparr>)" and *: "gas st' \<le> gas (st\<lparr>gas := gas st - g\<rparr>)"
using stmt.psimps(6)[of i xe e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm)
with a1 have "gas st6' \<le> gas st''''" by auto
also from 41(3) gcost g_def Some Pair s1 Method fields None st'_def e'_def s3 f1 n2
have "\<dots> \<le> gas st''" by (auto split:unit.split_asm)
also have "\<dots> \<le> gas st'"
proof -
from g_def gcost have "(applyf (costs (INVOKE i xe) e\<^sub>p e cd) \<bind> (\<lambda>g. assert Gas (\<lambda>st. gas st \<le> g) (modify (\<lambda>st. st\<lparr>gas := gas st - g\<rparr>)))) st = Normal ((), st\<lparr>gas := gas st - g\<rparr>)" by simp
moreover from e'_def have "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" by simp
moreover from st'_def have "applyf (\<lambda>st. st\<lparr>stack := emptyStore\<rparr>) (st\<lparr>gas := gas st - g\<rparr>) = Normal (st', st\<lparr>gas := gas st - g\<rparr>)" by simp
ultimately have "\<forall>ev cda sta st'a. load False fp xe e\<^sub>p e' emptyStore st' e cd (st\<lparr>gas := gas st - g\<rparr>) = Normal ((ev, cda, sta), st'a) \<longrightarrow> gas sta \<le> gas st' \<and> gas st'a \<le> gas (st\<lparr>gas := gas st - g\<rparr>) \<and> address ev = address e'" using a1 g_def gcost Some Pair s1 Method fields None st'_def e'_def s3 f1 41(2)[of _ "st\<lparr>gas := gas st - g\<rparr>" x ct _ _ x1 fp _ f c e' st' "st\<lparr>gas := gas st - g\<rparr>"] by blast
then show ?thesis using s3 f1 by auto
qed
also from * have "\<dots> \<le> gas (st\<lparr>gas := gas st - g\<rparr>)" .
finally show ?thesis by simp
next
case (e x)
with 41(1) g_def a1 gcost Some Pair s1 Method fields None st'_def e'_def s3 f1 show ?thesis using stmt.psimps(6)[of i xe e\<^sub>p e cd st] by auto
qed
qed
next
case n2: (e x)
with 41(1) g_def a1 gcost Some Pair s1 Method fields None st'_def e'_def show ?thesis using stmt.psimps(6) by auto
qed
next
case s2: (Some a)
with 41(1) g_def a1 gcost Some Pair s1 Method fields show ?thesis using stmt.psimps(6) by simp
qed
qed
next
case (Var x2)
with 41(1) g_def a1 gcost Some Pair s1 show ?thesis using stmt.psimps(6) by simp
qed
qed
qed
qed
qed
qed
next
case (42 ad i xe val e\<^sub>p e cd st)
define g where "g = costs (EXTERNAL ad i xe val) e\<^sub>p e cd st"
show ?case
proof (rule allI[OF impI])
fix st6' assume a1: "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st6')"
show "gas st6' \<le> gas st"
proof (cases)
assume "gas st \<le> costs (EXTERNAL ad i xe val) e\<^sub>p e cd st"
with 42(1) a1 show ?thesis using stmt.psimps(7) by simp
next
assume gcost: "\<not> gas st \<le> costs (EXTERNAL ad i xe val) e\<^sub>p e cd st"
then show ?thesis
proof (cases "expr ad e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue adv)
then show ?thesis
proof (cases c)
case (Value x1)
then show ?thesis
proof (cases x1)
case (TSInt x1)
with 42(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) by auto
next
case (TUInt x2)
with 42(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) by simp
next
case TBool
with 42(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) by simp
next
case TAddr
then show ?thesis
proof (cases "fmlookup e\<^sub>p adv")
case None
with 42(1) g_def a1 gcost n Pair KValue Value TAddr show ?thesis using stmt.psimps(7) by simp
next
case (Some x)
then show ?thesis
proof (cases x)
case p2: (Pair ct fb)
then show ?thesis
proof (cases "expr val e\<^sub>p e cd st'")
case n1: (n kv st'')
then show ?thesis
proof (cases kv)
case p3: (Pair a b)
then show ?thesis
proof (cases a)
case k2: (KValue v)
then show ?thesis
proof (cases b)
case v: (Value t)
show ?thesis
proof (cases "fmlookup ct i")
case None
show ?thesis
proof (cases "transfer (address e) adv v (accounts st'')")
case n2: None
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case s4: (Some acc)
show ?thesis
proof (cases "stmt fb e\<^sub>p (emptyEnv adv (address e) v) cd (st''\<lparr>accounts := acc,stack:=emptyStore, memory:=emptyStore\<rparr>)")
case n2: (n a st''')
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4
have "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st'''\<lparr>stack:=stack st'', memory := memory st''\<rparr>)"
using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm)
with a1 have "gas st6' \<le> gas st'''" by auto
also from 42(6)[OF _ n Pair KValue Value TAddr Some p2 n1 p3 k2 v None _ s4, of _ st'' st'' st'' "()"] n2 g_def gcost
have "\<dots> \<le> gas (st''\<lparr>accounts := acc,stack:=emptyStore, memory:=emptyStore\<rparr>)" by auto
also from 42(3)[of _ "st\<lparr>gas := gas st - g\<rparr>" _ st' _ _ adv _ x ct] g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4 n2
have "\<dots> \<le> gas st'" by auto
also from 42(2)[of _ "st\<lparr>gas := gas st - g\<rparr>"] g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4 n2
have "\<dots> \<le> gas (st\<lparr>gas := gas st - g\<rparr>)" by auto
finally show ?thesis by simp
next
case (e x)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
next
case s1: (Some a)
then show ?thesis
proof (cases a)
case (Method x1)
then show ?thesis
proof (cases x1)
case (fields fp f c)
then show ?thesis
proof (cases c)
case None
define stl e'
where "stl = st''\<lparr>stack:=emptyStore, memory:=emptyStore\<rparr>"
and "e' = ffold (init ct) (emptyEnv adv (address e) v) (fmdom ct)"
then show ?thesis
proof (cases "load True fp xe e\<^sub>p e' emptyStore stl e cd st''")
case s3: (n a st''')
then show ?thesis
proof (cases a)
case f1: (fields e'' cd' st'''')
show ?thesis
proof (cases "transfer (address e) adv v (accounts st'''')")
case n2: None
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v s3 f1 stl_def e'_def show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case s4: (Some acc)
show ?thesis
proof (cases "stmt f e\<^sub>p e'' cd' (st''''\<lparr>accounts := acc\<rparr>)")
case n2: (n a st''''')
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4
have "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st'''''\<lparr>stack:=stack st''', memory := memory st'''\<rparr>)"
using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm)
with a1 have "gas st6' \<le> gas (st''''')" by auto
also from 42(5)[OF _ n Pair KValue Value TAddr Some p2 n1 p3 k2 v s1 Method fields _ None _ _ s3 _ _ _ _, of "()" f e'' "(cd', st'''')" cd' st'''' st''' st''' acc "()" "st''''\<lparr>accounts := acc\<rparr>"] s4 stl_def e'_def f1 n2 g_def gcost
have "\<dots> \<le> gas (st''''\<lparr>accounts := acc\<rparr>)" by auto
also have "\<dots> \<le> gas stl"
proof -
from g_def gcost have "(applyf (costs (EXTERNAL ad i xe val) e\<^sub>p e cd) \<bind> (\<lambda>g. assert Gas (\<lambda>st. gas st \<le> g) (modify (\<lambda>st. st\<lparr>gas := gas st - g\<rparr>)))) st = Normal ((), st\<lparr>gas := gas st - g\<rparr>)" by simp
moreover from e'_def have "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" by simp
moreover from n1 have "expr val e\<^sub>p e cd st' = Normal (kv, st'')" by simp
moreover from stl_def have "applyf (\<lambda>st. st\<lparr>stack := emptyStore, memory := emptyStore\<rparr>) st'' = Normal (stl, st'')" by simp
moreover have "applyf accounts st'' = Normal ((accounts st''), st'')" by simp
ultimately have "\<forall>ev cda sta st'a. load True fp xe e\<^sub>p e' emptyStore stl e cd st'' = Normal ((ev, cda, sta), st'a) \<longrightarrow> gas sta \<le> gas stl \<and> gas st'a \<le> gas st'' \<and> address ev = address e'" using 42(4)[of _ "st\<lparr>gas := gas st - g\<rparr>" _ st' _ _ adv _ x ct _ _ st'' _ b v t _ x1 fp "(f,c)" f c e'] g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 n2 by blast
then show ?thesis using s3 f1 by auto
qed
also from stl_def have "\<dots> \<le> gas st''" by simp
also from 42(3)[of _ "st\<lparr>gas := gas st - g\<rparr>" _ st' _ _ adv _ x ct] g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 n2
have "\<dots> \<le> gas st'" by auto
also from 42(2)[of _ "st\<lparr>gas := gas st - g\<rparr>"] g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 n2
have "\<dots> \<le> gas (st\<lparr>gas := gas st - g\<rparr>)" by auto
finally show ?thesis by simp
next
case (e x)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
qed
next
case (e x)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
next
case s2: (Some a)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields n1 p3 k2 v show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
next
case (Var x2)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 n1 p3 k2 v show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
next
case (Calldata x2)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 k2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Memory x3)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 k2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Storage x4)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 k2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
next
case (KCDptr x2)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KMemptr x3)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KStoptr x4)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
next
case n2: (e x)
with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
qed
qed
next
case (Calldata x2)
with 42(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Memory x3)
with 42(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (Storage x4)
with 42(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
next
case (KCDptr x2)
with 42(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KMemptr x3)
with 42(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
next
case (KStoptr x4)
with 42(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
next
case (e _)
with 42(1) g_def a1 gcost show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp
qed
qed
qed
next
case (43 ad ex e\<^sub>p e cd st)
define g where "g = costs (TRANSFER ad ex) e\<^sub>p e cd st"
show ?case
proof (rule allI[OF impI])
fix st6' assume stmt_def: "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((), st6')"
show "gas st6' \<le> gas st"
proof cases
assume "gas st \<le> g"
with 43 stmt_def g_def show ?thesis using stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp
next
assume "\<not> gas st \<le> g"
show ?thesis
proof (cases "expr ex e\<^sub>p e cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair b c)
then show ?thesis
proof (cases b)
case (KValue v)
then show ?thesis
proof (cases c)
case (Value t)
then show ?thesis
proof (cases "expr ad e\<^sub>p e cd st'")
case n2: (n a st'')
then show ?thesis
proof (cases a)
case p2: (Pair b c)
then show ?thesis
proof (cases b)
case k2: (KValue adv)
then show ?thesis
proof (cases c)
case v2: (Value x1)
then show ?thesis
proof (cases x1)
case (TSInt x1)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 g_def show ?thesis using stmt.psimps(8) by simp
next
case (TUInt x2)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 g_def show ?thesis using stmt.psimps(8) by simp
next
case TBool
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 g_def show ?thesis using stmt.psimps(8) by simp
next
case TAddr
then show ?thesis
proof (cases "transfer (address e) adv v (accounts st'')")
case None
with 43(1) stmt_def g_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 TAddr show ?thesis using stmt.psimps(8) by simp
next
case (Some acc)
then show ?thesis
proof (cases "fmlookup e\<^sub>p adv")
case None
with 43(1) `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 TAddr Some g_def
have "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((),st''\<lparr>accounts:=acc\<rparr>)" using stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp
with stmt_def have "gas st6' \<le> gas st''" by auto
also from 43(3)[of "()" "(st\<lparr>gas := gas st - g\<rparr>)" _ st'] `\<not> gas st \<le> g` n Pair KValue Value n2 g_def have "\<dots> \<le> gas st'" by simp
also from 43(2)[of "()" "(st\<lparr>gas := gas st - g\<rparr>)"] `\<not> gas st \<le> g` n g_def have "\<dots> \<le> gas st" by simp
finally show ?thesis .
next
case s2: (Some a)
then show ?thesis
proof (cases a)
case p3: (Pair ct f)
define e' where "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)"
show ?thesis
proof (cases "stmt f e\<^sub>p e' emptyStore (st''\<lparr>accounts := acc, stack:=emptyStore, memory:=emptyStore\<rparr>)")
case n3: (n a st''')
with 43(1) `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def
have "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((),st'''\<lparr>stack:=stack st'', memory := memory st''\<rparr>)" using e'_def stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp
with stmt_def have "gas st6' \<le> gas st'''" by auto
also from 43(4)[of "()" "st\<lparr>gas := gas st - g\<rparr>" _ st' _ _ v t _ st'' _ _ adv x1 "accounts st''" st'' acc _ ct f e' _ st'' "()" "st''\<lparr>accounts := acc, stack:=emptyStore, memory:=emptyStore\<rparr>"] `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def n2 e'_def n3
have "\<dots> \<le> gas (st''\<lparr>accounts := acc, stack := emptyStore, memory := emptyStore\<rparr>)" by simp
also from 43(3)[of "()" "(st\<lparr>gas := gas st - g\<rparr>)" _ st'] `\<not> gas st \<le> g` n Pair KValue Value n2 g_def have "\<dots> \<le> gas st'" by simp
also from 43(2)[of "()" "(st\<lparr>gas := gas st - g\<rparr>)"] `\<not> gas st \<le> g` n g_def have "\<dots> \<le> gas st" by simp
finally show ?thesis .
next
case (e x)
with 43(1) `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def e'_def stmt_def show ?thesis using stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp
qed
qed
qed
qed
qed
next
case (Calldata x2)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 g_def show ?thesis using stmt.psimps(8) by simp
next
case (Memory x3)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 g_def show ?thesis using stmt.psimps(8) by simp
next
case (Storage x4)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 k2 g_def show ?thesis using stmt.psimps(8) by simp
qed
next
case (KCDptr x2)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 g_def show ?thesis using stmt.psimps(8) by simp
next
case (KMemptr x3)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 g_def show ?thesis using stmt.psimps(8) by simp
next
case (KStoptr x4)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value n2 p2 g_def show ?thesis using stmt.psimps(8) by simp
qed
qed
next
case (e e)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue Value g_def show ?thesis using stmt.psimps(8) by simp
qed
next
case (Calldata x2)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue g_def show ?thesis using stmt.psimps(8) by simp
next
case (Memory x3)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue g_def show ?thesis using stmt.psimps(8) by simp
next
case (Storage x4)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair KValue g_def show ?thesis using stmt.psimps(8) by simp
qed
next
case (KCDptr x2)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair g_def show ?thesis using stmt.psimps(8) by simp
next
case (KMemptr x3)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair g_def show ?thesis using stmt.psimps(8) by simp
next
case (KStoptr x4)
with 43(1) stmt_def `\<not> gas st \<le> g` n Pair g_def show ?thesis using stmt.psimps(8) by simp
qed
qed
next
case (e e)
with 43(1) stmt_def `\<not> gas st \<le> g` g_def show ?thesis using stmt.psimps(8) by simp
qed
qed
qed
next
case (44 id0 tp ex s e\<^sub>p e\<^sub>v cd st)
define g where "g = costs (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st"
show ?case
proof (rule allI[OF impI])
fix st6' assume stmt_def: "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st = Normal ((), st6')"
show "gas st6' \<le> gas st"
proof cases
assume "gas st \<le> g"
with 44 stmt_def g_def show ?thesis using stmt.psimps(9) by simp
next
assume "\<not> gas st \<le> g"
show ?thesis
proof (cases ex)
case None
then show ?thesis
proof (cases "decl id0 tp None False cd (memory (st\<lparr>gas := gas st - g\<rparr>)) cd e\<^sub>v (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair cd' e')
with 44(1) stmt_def `\<not> gas st \<le> g` None n g_def have "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st = stmt s e\<^sub>p e' cd' st'" using stmt.psimps(9)[of id0 tp ex s e\<^sub>p e\<^sub>v cd st] by (simp split:unit.split_asm)
with 44(4)[of "()" "st\<lparr>gas := gas st - g\<rparr>"] stmt_def `\<not> gas st \<le> g` None n Pair g_def have "gas st6' \<le> gas st'" by simp
moreover from n Pair have "gas st' \<le> gas st" using decl_gas_address by simp
ultimately show ?thesis by simp
qed
next
case (e x)
with 44 stmt_def `\<not> gas st \<le> g` None g_def show ?thesis using stmt.psimps(9) by simp
qed
next
case (Some ex')
then show ?thesis
proof (cases "expr ex' e\<^sub>p e\<^sub>v cd (st\<lparr>gas := gas st - g\<rparr>)")
case (n a st')
then show ?thesis
proof (cases a)
case (Pair v t)
then show ?thesis
proof (cases "decl id0 tp (Some (v, t)) False cd (memory st') cd e\<^sub>v st'")
case s2: (n a st'')
then show ?thesis
proof (cases a)
case f2: (Pair cd' e')
with 44(1) stmt_def `\<not> gas st \<le> g` Some n Pair s2 g_def have "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st = stmt s e\<^sub>p e' cd' st''" using stmt.psimps(9)[of id0 tp ex s e\<^sub>p e\<^sub>v cd st] by (simp split:unit.split_asm)
with 44(3)[of "()" "st\<lparr>gas := gas st - g\<rparr>" ex' _ st'] stmt_def `\<not> gas st \<le> g` Some n Pair s2 f2 g_def have "gas st6' \<le> gas st''" by simp
moreover from Some n Pair s2 f2 g_def have "gas st'' \<le> gas st'" using decl_gas_address by simp
moreover from 44(2)[of "()" "st\<lparr>gas := gas st - g\<rparr>" ex'] `\<not> gas st \<le> g` Some n Pair g_def have "gas st' \<le> gas st" by simp
ultimately show ?thesis by simp
qed
next
case (e x)
with 44(1) stmt_def `\<not> gas st \<le> g` Some n Pair g_def show ?thesis using stmt.psimps(9) by simp
qed
qed
next
case (e e)
with 44 stmt_def `\<not> gas st \<le> g` Some g_def show ?thesis using stmt.psimps(9) by simp
qed
qed
qed
qed
qed
subsection \<open>Termination\<close>
lemma x1:
assumes "expr x e\<^sub>p env cd st = Normal (val, s')"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (x, e\<^sub>p, env, cd, st))))"
shows "gas s' < gas st \<or> gas s' = gas st"
using assms msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of x e\<^sub>p env cd st] by auto
lemma x2:
assumes "(if gas st \<le> c then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - c\<rparr>))) st) = Normal ((), s')"
and "expr x e\<^sub>p e cd s' = Normal (val, s'a)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (x, e\<^sub>p, e, cd, s'))))"
shows "gas s'a < gas st \<or> gas s'a = gas st"
proof -
from assms have "gas s' < gas st \<or> gas s' = gas st" by (auto split:if_split_asm)
with assms show ?thesis using x1[of x e\<^sub>p e cd s' val s'a] by auto
qed
lemma x2sub:
assumes "(if gas st \<le> costs (TRANSFER ad ex) e\<^sub>p e cd st then throw Gas st
else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - costs (TRANSFER ad ex) e\<^sub>p e cd st\<rparr>))) st) =
Normal ((), s')" and
" expr ex e\<^sub>p e cd s' = Normal ((KValue vb, Value t), s'a)"
and " msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ex, e\<^sub>p, e, cd, s'))))"
and "(\<And>ad i xe val e\<^sub>p e cd st. 0 < costs (EXTERNAL ad i xe val) e\<^sub>p e cd st)" and "gas s'a \<noteq> gas st" shows "gas s'a < gas st"
proof -
from assms have "gas s' < gas st \<or> gas s' = gas st" by (auto split:if_split_asm)
with assms show ?thesis using x1[of ex e\<^sub>p e cd s' "(KValue vb, Value t)" s'a] by auto
qed
lemma x3:
assumes "(if gas st \<le> c then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - c\<rparr>))) st) = Normal ((), s')"
and "s'\<lparr>stack := emptyStore\<rparr> = va"
and "load False ad xe e\<^sub>p(ffold (init aa) \<lparr>address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\<rparr> (fmdom aa)) emptyStore va e cd s' = Normal ((ag, ah, s'd), vc)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (False, ad, xe, e\<^sub>p, ffold (init aa) \<lparr>address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\<rparr> (fmdom aa), emptyStore, va, e, cd, s'))))"
and "c>0"
shows "gas s'd < gas st"
proof -
from assms have "gas s'd \<le> gas va" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(5)[of False ad xe e\<^sub>p "ffold (init aa) \<lparr>address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\<rparr> (fmdom aa)" emptyStore va e cd s'] by blast
also from assms(2) have "\<dots> = gas s'" by auto
also from assms(1,5) have "\<dots> < gas st" by (auto split:if_split_asm)
finally show ?thesis .
qed
lemma x4:
assumes "(if gas st \<le> c then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - c\<rparr>))) st) = Normal ((), s')"
and "s'\<lparr>stack := emptyStore\<rparr> = va"
and "load False ad xe e\<^sub>p (ffold (init aa) \<lparr>address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\<rparr> (fmdom aa)) emptyStore va e cd s' = Normal ((ag, ah, s'd), vc)"
and "stmt ae e\<^sub>p ag ah s'd = Normal ((), s'e)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (ae, e\<^sub>p, ag, ah, s'd))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (False, ad, xe, e\<^sub>p, ffold (init aa) \<lparr>address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\<rparr> (fmdom aa), emptyStore, va, e, cd, s'))))"
and "c>0"
shows "gas s'e < gas st"
proof -
from assms have "gas s'e \<le> gas s'd" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7) by blast
with assms show ?thesis using x3[OF assms(1) assms(2) assms(3) assms(6)] by simp
qed
lemma x5:
assumes "(if gas st \<le> costs (COMP s1 s2) e\<^sub>p e cd st then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - costs (COMP s1 s2) e\<^sub>p e cd st\<rparr>))) st) = Normal ((), s')"
and "stmt s1 e\<^sub>p e cd s' = Normal ((), s'a)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (s1, e\<^sub>p, e, cd, s'))))"
shows "gas s'a < gas st \<or> gas s'a = gas st"
using assms msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7)[of s1 e\<^sub>p e cd s'] by (auto split:if_split_asm)
lemma x6:
assumes "(if gas st \<le> costs (WHILE ex s0) e\<^sub>p e cd st then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - costs (WHILE ex s0) e\<^sub>p e cd st\<rparr>))) st) = Normal ((), s')"
and "expr ex e\<^sub>p e cd s' = Normal (val, s'a)"
and "stmt s0 e\<^sub>p e cd s'a = Normal ((), s'b)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (s0, e\<^sub>p, e, cd, s'a))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ex, e\<^sub>p, e, cd, s'))))"
shows "gas s'b < gas st"
proof -
from assms have "gas s'b \<le> gas s'a" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7)[of s0 e\<^sub>p e cd s'a] by blast
also from assms have "\<dots> \<le> gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ex e\<^sub>p e cd s'] by auto
also from assms(1) have "\<dots> < gas st" using while_not_zero by (auto split:if_split_asm)
finally show ?thesis .
qed
lemma x7:
assumes "(if gas st \<le> c then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - c\<rparr>))) st) = Normal ((), s')"
and "expr ad e\<^sub>p e cd s' = Normal ((KValue vb, Value TAddr), s'a)"
and "expr val e\<^sub>p e cd s'a = Normal ((KValue vd, Value ta), s'b)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (val, e\<^sub>p, e, cd, s'a))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'))))"
and "c>0"
shows "gas s'b < gas st"
proof -
from assms(4,3) have "gas s'b \<le> gas s'a" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of val e\<^sub>p e cd s'a] by simp
also from assms(5,2) have "\<dots> \<le> gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ad e\<^sub>p e cd s'] by simp
also from assms(1,6) have "\<dots> < gas st" by (auto split:if_split_asm)
finally show ?thesis .
qed
lemma x8:
assumes "(if gas st \<le> costs (TRANSFER ad ex) e\<^sub>p e cd st then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - costs (TRANSFER ad ex) e\<^sub>p e cd st\<rparr>))) st) = Normal ((), s')"
and "expr ex e\<^sub>p e cd s' = Normal ((KValue vb, Value t), s'a)"
and "expr ad e\<^sub>p e cd s'a = Normal ((KValue vd, Value TAddr), s'b)"
and "s'b\<lparr>accounts := ab, stack := emptyStore, memory := emptyStore\<rparr> = s'e"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'a))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ex, e\<^sub>p, e, cd, s'))))"
shows "gas s'e < gas st"
proof -
from assms(4) have "gas s'e = gas s'b" by auto
also from assms(5,3) have "\<dots> \<le> gas s'a" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ad e\<^sub>p e cd s'a] by simp
also from assms(6,2) have "\<dots> \<le> gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ex e\<^sub>p e cd s'] by simp
also from assms(1) have "\<dots> < gas st" using transfer_not_zero by (auto split:if_split_asm)
finally show ?thesis .
qed
lemma x9:
assumes "(if gas st \<le> costs (BLOCK ((id0, tp), Some a) s) e\<^sub>p e\<^sub>v cd st then throw Gas st else (get \<bind> (\<lambda>sa. put (sa\<lparr>gas := gas sa - costs (BLOCK ((id0, tp), Some a) s) e\<^sub>p e\<^sub>v cd st\<rparr>))) st) = Normal ((), s')"
and "expr a e\<^sub>p e\<^sub>v cd s' = Normal ((aa, b), s'a)"
and "decl id0 tp (Some (aa, b)) False cd vb cd e\<^sub>v s'a = Normal ((ab, ba), s'c)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (a, e\<^sub>p, e\<^sub>v, cd, s'))))"
shows "gas s'c < gas st \<or> gas s'c = gas st"
proof -
from assms have "gas s'c = gas s'a" using decl_gas_address[of id0 tp "(Some (aa, b))" False cd vb cd e\<^sub>v s'a] by simp
also from assms have "\<dots> \<le> gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of a e\<^sub>p e\<^sub>v cd s'] by simp
also from assms(1) have "\<dots> \<le> gas st" by (auto split:if_split_asm)
finally show ?thesis by auto
qed
lemma x10:
assumes "(if gas st \<le> costs (BLOCK ((id0, tp), None) s) e\<^sub>p e\<^sub>v cd st then throw Gas st else (get \<bind> (\<lambda>sa. put (sa\<lparr>gas := gas sa - costs (BLOCK ((id0, tp), None) s) e\<^sub>p e\<^sub>v cd st\<rparr>))) st) = Normal ((), s')"
and "decl id0 tp None False cd va cd e\<^sub>v s' = Normal ((a, b), s'b)"
shows "gas s'b < gas st \<or> gas s'b = gas st"
proof -
from assms have "gas s'b = gas s'" using decl_gas_address[of id0 tp None False cd va cd e\<^sub>v s'] by simp
also from assms(1) have "\<dots> \<le> gas st" by (auto split:if_split_asm)
finally show ?thesis by auto
qed
lemma x11:
assumes "(if gas st \<le> c then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - c\<rparr>))) st) = Normal ((), s')"
and "expr ad e\<^sub>p e cd s' = Normal ((KValue vb, Value TAddr), s'a)"
and "expr val e\<^sub>p e cd s'a = Normal ((KValue vd, Value ta), s'b)"
and "load True af xe e\<^sub>p (ffold (init ab) \<lparr>address = vb, sender = address e, svalue = vd, denvalue = fmempty\<rparr> (fmdom ab)) emptyStore (s'b\<lparr>stack := emptyStore, memory := emptyStore\<rparr>) e cd s'b = Normal ((ak, al, s'g), vh)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (True, af, xe, e\<^sub>p, ffold (init ab) \<lparr>address = vb, sender = address e, svalue = vd, denvalue = fmempty\<rparr> (fmdom ab), emptyStore, s'b\<lparr>stack := emptyStore, memory := emptyStore\<rparr>, e, cd, s'b))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (val, e\<^sub>p, e, cd, s'a))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'))))"
and "c>0"
shows "gas s'g < gas st"
proof -
from assms have "gas s'g \<le> gas (s'b\<lparr>stack := emptyStore, memory := emptyStore\<rparr>)" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(5)[of True af xe e\<^sub>p "ffold (init ab) \<lparr>address = vb, sender = address e, svalue = vd, denvalue = fmempty\<rparr> (fmdom ab)" emptyStore "s'b\<lparr>stack := emptyStore, memory := emptyStore\<rparr>" e cd s'b] by blast
also have "\<dots> = gas s'b" by simp
also from assms have "\<dots> < gas st" using x7[OF assms(1) assms(2) assms(3) assms(6)] by auto
finally show ?thesis .
qed
lemma x12:
assumes "(if gas st \<le> c then throw Gas st else (get \<bind> (\<lambda>s. put (s\<lparr>gas := gas s - c\<rparr>))) st) = Normal ((), s')"
and "expr ad e\<^sub>p e cd s' = Normal ((KValue vb, Value TAddr), s'a)"
and "expr val e\<^sub>p e cd s'a = Normal ((KValue vd, Value ta), s'b)"
and "load True af xe e\<^sub>p (ffold (init ab) \<lparr>address = vb, sender = address e, svalue = vd, denvalue = fmempty\<rparr> (fmdom ab)) emptyStore (s'b\<lparr>stack := emptyStore, memory := emptyStore\<rparr>) e cd s'b = Normal ((ak, al, s'g), vh)"
and "stmt ag e\<^sub>p ak al (s'g\<lparr>accounts := ala\<rparr>) = Normal ((), s'h)"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (True, af, xe, e\<^sub>p, ffold (init ab) \<lparr>address = vb, sender = address e, svalue = vd, denvalue = fmempty\<rparr> (fmdom ab), emptyStore, s'b\<lparr>stack := emptyStore, memory := emptyStore\<rparr>, e, cd, s'b))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (val, e\<^sub>p, e, cd, s'a))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'))))"
and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (ag, e\<^sub>p, ak, al, (s'g\<lparr>accounts := ala\<rparr>)))))"
and "c>0"
shows "gas s'h < gas st"
proof -
from assms have "gas s'h \<le> gas (s'g\<lparr>accounts := ala\<rparr>)" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7) by blast
also from assms have "\<dots> < gas st" using x11[OF assms(1) assms(2) assms(3) assms(4)] by auto
finally show ?thesis .
qed
termination
apply (relation
"measures [\<lambda>x. case x of Inr (Inr (Inr l)) \<Rightarrow> gas (snd (snd (snd (snd l))))
| Inr (Inr (Inl l)) \<Rightarrow> gas (snd (snd (snd (snd l))))
| Inr (Inl (Inr l)) \<Rightarrow> gas (snd (snd (snd (snd (snd (snd (snd (snd (snd l)))))))))
| Inr (Inl (Inl l)) \<Rightarrow> gas (snd (snd (snd (snd l))))
| Inl (Inr (Inr l)) \<Rightarrow> gas (snd (snd (snd (snd l))))
| Inl (Inr (Inl l)) \<Rightarrow> gas (snd (snd (snd (snd (snd (snd l))))))
| Inl (Inl l) \<Rightarrow> gas (snd (snd (snd (snd (snd (snd (snd l))))))),
\<lambda>x. case x of Inr (Inr (Inr l)) \<Rightarrow> 1
| Inr (Inr (Inl l)) \<Rightarrow> 0
| Inr (Inl (Inr l)) \<Rightarrow> 0
| Inr (Inl (Inl l)) \<Rightarrow> 0
| Inl (Inr (Inr l)) \<Rightarrow> 0
| Inl (Inr (Inl l)) \<Rightarrow> 0
| Inl (Inl l) \<Rightarrow> 0,
\<lambda>x. case x of Inr (Inr (Inr l)) \<Rightarrow> size (fst l)
| Inr (Inr (Inl l)) \<Rightarrow> size (fst l)
| Inr (Inl (Inr l)) \<Rightarrow> size_list size (fst (snd (snd l)))
| Inr (Inl (Inl l)) \<Rightarrow> size (fst l)
| Inl (Inr (Inr l)) \<Rightarrow> size (fst l)
| Inl (Inr (Inl l)) \<Rightarrow> size_list size (fst (snd (snd l)))
| Inl (Inl l) \<Rightarrow> size_list size (fst (snd (snd (snd l))))]
")
apply simp_all
apply (simp only: x1)
apply (simp only: x1)
apply (simp only: x1)
apply (auto split: if_split_asm)[1]
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (auto split: if_split_asm)[1]
using call_not_zero apply (simp only: x3)
using call_not_zero apply (simp add: x4)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
using ecall_not_zero apply (simp add: x7)
using ecall_not_zero apply (auto simp add: x11)[1]
using ecall_not_zero apply (auto simp add: x12)[1]
apply (simp only: x1)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (simp only: x2)
apply (simp only: x2)
apply (simp only: x2)
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x5)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (simp only: x2)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
apply (simp only: x6)
apply (auto split: if_split_asm)[1]
using invoke_not_zero apply (simp only: x3)
apply (auto split: if_split_asm)[1]
apply (simp only: x2)
using external_not_zero apply (simp add: x7)
using external_not_zero apply (auto simp add: x11)[1]
using external_not_zero apply (auto simp add: x7)[1]
apply (auto split: if_split_asm)[1]
apply (simp add: x2)
apply (simp add: x8)
apply (auto split: if_split_asm)[1]
apply (simp only: x9)
apply (simp only: x10)
done
end
subsection \<open>A minimal cost model\<close>
fun costs_min :: "S\<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> State \<Rightarrow> Gas"
where
"costs_min SKIP e\<^sub>p e cd st = 0"
| "costs_min (ASSIGN lv ex) e\<^sub>p e cd st = 0"
| "costs_min (COMP s1 s2) e\<^sub>p e cd st = 0"
| "costs_min (ITE ex s1 s2) e\<^sub>p e cd st = 0"
| "costs_min (WHILE ex s0) e\<^sub>p e cd st = 1"
| "costs_min (TRANSFER ad ex) e\<^sub>p e cd st = 1"
| "costs_min (BLOCK ((id0, tp), ex) s) e\<^sub>p e cd st =0"
| "costs_min (INVOKE _ _) e\<^sub>p e cd st = 1"
| "costs_min (EXTERNAL _ _ _ _) e\<^sub>p e cd st = 1"
fun costs_ex :: "E \<Rightarrow> Environment\<^sub>P \<Rightarrow> Environment \<Rightarrow> CalldataT \<Rightarrow> State \<Rightarrow> Gas"
where
"costs_ex (E.INT _ _) e\<^sub>p e cd st = 0"
| "costs_ex (UINT _ _) e\<^sub>p e cd st = 0"
| "costs_ex (ADDRESS _) e\<^sub>p e cd st = 0"
| "costs_ex (BALANCE _) e\<^sub>p e cd st = 0"
| "costs_ex THIS e\<^sub>p e cd st = 0"
| "costs_ex SENDER e\<^sub>p e cd st = 0"
| "costs_ex VALUE e\<^sub>p e cd st = 0"
| "costs_ex (TRUE) e\<^sub>p e cd st = 0"
| "costs_ex (FALSE) e\<^sub>p e cd st = 0"
| "costs_ex (LVAL _) e\<^sub>p e cd st = 0"
| "costs_ex (PLUS _ _) e\<^sub>p e cd st = 0"
| "costs_ex (MINUS _ _) e\<^sub>p e cd st = 0"
| "costs_ex (EQUAL _ _) e\<^sub>p e cd st = 0"
| "costs_ex (LESS _ _) e\<^sub>p e cd st = 0"
| "costs_ex (AND _ _) e\<^sub>p e cd st = 0"
| "costs_ex (OR _ _) e\<^sub>p e cd st = 0"
| "costs_ex (NOT _) e\<^sub>p e cd st = 0"
| "costs_ex (CALL _ _) e\<^sub>p e cd st = 1"
| "costs_ex (ECALL _ _ _ _) e\<^sub>p e cd st = 1"
global_interpretation solidity: statement_with_gas costs_min costs_ex
defines stmt = "solidity.stmt"
and lexp = solidity.lexp
and expr = solidity.expr
and ssel = solidity.ssel
and rexp = solidity.rexp
and msel = solidity.msel
and load = solidity.load
by unfold_locales auto
end
diff --git a/thys/UTP/toolkit/FSet_Extra.thy b/thys/UTP/toolkit/FSet_Extra.thy
--- a/thys/UTP/toolkit/FSet_Extra.thy
+++ b/thys/UTP/toolkit/FSet_Extra.thy
@@ -1,282 +1,282 @@
(******************************************************************************)
(* Project: Isabelle/UTP Toolkit *)
(* File: FSet_Extra.thy *)
(* Authors: Frank Zeyda and Simon Foster (University of York, UK) *)
(* Emails: frank.zeyda@york.ac.uk and simon.foster@york.ac.uk *)
(******************************************************************************)
section \<open>Finite Sets: extra functions and properties\<close>
theory FSet_Extra
imports
"HOL-Library.FSet"
"HOL-Library.Countable_Set_Type"
begin
setup_lifting type_definition_fset
notation fempty ("\<lbrace>\<rbrace>")
notation fset ("\<langle>_\<rangle>\<^sub>f")
notation fminus (infixl "-\<^sub>f" 65)
syntax
"_FinFset" :: "args => 'a fset" ("\<lbrace>(_)\<rbrace>")
translations
"\<lbrace>x, xs\<rbrace>" == "CONST finsert x \<lbrace>xs\<rbrace>"
"\<lbrace>x\<rbrace>" == "CONST finsert x \<lbrace>\<rbrace>"
term "fBall"
syntax
"_fBall" :: "pttrn => 'a fset => bool => bool" ("(3\<forall> _|\<in>|_./ _)" [0, 0, 10] 10)
"_fBex" :: "pttrn => 'a fset => bool => bool" ("(3\<exists> _|\<in>|_./ _)" [0, 0, 10] 10)
translations
"\<forall> x|\<in>|A. P" == "CONST fBall A (%x. P)"
"\<exists> x|\<in>|A. P" == "CONST fBex A (%x. P)"
definition FUnion :: "'a fset fset \<Rightarrow> 'a fset" ("\<Union>\<^sub>f_" [90] 90) where
"FUnion xs = Abs_fset (\<Union>x\<in>\<langle>xs\<rangle>\<^sub>f. \<langle>x\<rangle>\<^sub>f)"
definition FInter :: "'a fset fset \<Rightarrow> 'a fset" ("\<Inter>\<^sub>f_" [90] 90) where
"FInter xs = Abs_fset (\<Inter>x\<in>\<langle>xs\<rangle>\<^sub>f. \<langle>x\<rangle>\<^sub>f)"
text \<open>Finite power set\<close>
definition FinPow :: "'a fset \<Rightarrow> 'a fset fset" where
"FinPow xs = Abs_fset (Abs_fset ` Pow \<langle>xs\<rangle>\<^sub>f)"
text \<open>Set of all finite subsets of a set\<close>
definition Fow :: "'a set \<Rightarrow> 'a fset set" where
"Fow A = {x. \<langle>x\<rangle>\<^sub>f \<subseteq> A}"
declare Abs_fset_inverse [simp]
lemma fset_intro:
"fset x = fset y \<Longrightarrow> x = y"
by (simp add:fset_inject)
lemma fset_elim:
"\<lbrakk> x = y; fset x = fset y \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
by (auto)
lemma fmember_intro:
"\<lbrakk> x \<in> fset(xs) \<rbrakk> \<Longrightarrow> x |\<in>| xs"
- by (metis fmember_iff_member_fset)
+ .
lemma fmember_elim:
"\<lbrakk> x |\<in>| xs; x \<in> fset(xs) \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
- by (metis fmember_iff_member_fset)
+ .
lemma fnmember_intro [intro]:
"\<lbrakk> x \<notin> fset(xs) \<rbrakk> \<Longrightarrow> x |\<notin>| xs"
- by (metis fmember_iff_member_fset)
+ .
lemma fnmember_elim [elim]:
"\<lbrakk> x |\<notin>| xs; x \<notin> fset(xs) \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
- by (metis fmember_iff_member_fset)
+ .
lemma fsubset_intro [intro]:
"\<langle>xs\<rangle>\<^sub>f \<subseteq> \<langle>ys\<rangle>\<^sub>f \<Longrightarrow> xs |\<subseteq>| ys"
by (metis less_eq_fset.rep_eq)
lemma fsubset_elim [elim]:
"\<lbrakk> xs |\<subseteq>| ys; \<langle>xs\<rangle>\<^sub>f \<subseteq> \<langle>ys\<rangle>\<^sub>f \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
by (metis less_eq_fset.rep_eq)
lemma fBall_intro [intro]:
"Ball \<langle>A\<rangle>\<^sub>f P \<Longrightarrow> fBall A P"
- by (metis (poly_guards_query) fBallI fmember_iff_member_fset)
+ by (metis (poly_guards_query) fBallI)
lemma fBall_elim [elim]:
"\<lbrakk> fBall A P; Ball \<langle>A\<rangle>\<^sub>f P \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> Q"
- by (metis fBallE fmember_iff_member_fset)
+ by (metis fBallE)
lift_definition finset :: "'a list \<Rightarrow> 'a fset" is set ..
context linorder
begin
lemma sorted_list_of_set_inj:
"\<lbrakk> finite xs; finite ys; sorted_list_of_set xs = sorted_list_of_set ys \<rbrakk>
\<Longrightarrow> xs = ys"
apply (simp add:sorted_list_of_set_def)
apply (induct xs rule:finite_induct)
apply (induct ys rule:finite_induct)
apply (simp_all)
apply (metis finite.insertI insert_not_empty sorted_list_of_set_def sorted_list_of_set_empty sorted_list_of_set_eq_Nil_iff)
apply (metis finite.insertI finite_list set_remdups set_sort sorted_list_of_set_def sorted_list_of_set_sort_remdups)
done
definition flist :: "'a fset \<Rightarrow> 'a list" where
"flist xs = sorted_list_of_set (fset xs)"
lemma flist_inj: "inj flist"
apply (simp add:flist_def inj_on_def)
apply (clarify)
apply (rename_tac x y)
apply (subgoal_tac "fset x = fset y")
apply (simp add:fset_inject)
apply (rule sorted_list_of_set_inj, simp_all)
done
lemma flist_props [simp]:
"sorted (flist xs)"
"distinct (flist xs)"
by (simp_all add:flist_def)
lemma flist_empty [simp]:
"flist \<lbrace>\<rbrace> = []"
by (simp add:flist_def)
lemma flist_inv [simp]: "finset (flist xs) = xs"
by (simp add:finset_def flist_def fset_inverse)
lemma flist_set [simp]: "set (flist xs) = fset xs"
by (simp add:finset_def flist_def fset_inverse)
lemma fset_inv [simp]: "\<lbrakk> sorted xs; distinct xs \<rbrakk> \<Longrightarrow> flist (finset xs) = xs"
apply (simp add:finset_def flist_def fset_inverse)
apply (metis local.sorted_list_of_set_sort_remdups local.sorted_sort_id remdups_id_iff_distinct)
done
lemma fcard_flist:
"fcard xs = length (flist xs)"
apply (simp add:fcard_def)
apply (fold flist_set)
apply (unfold distinct_card[OF flist_props(2)])
apply (rule refl)
done
lemma flist_nth:
"i < fcard vs \<Longrightarrow> flist vs ! i |\<in>| vs"
- apply (simp add: fmember_def flist_def fcard_def)
+ apply (simp add: flist_def fcard_def)
apply (metis fcard.rep_eq fcard_flist finset.rep_eq flist_def flist_inv nth_mem)
done
definition fmax :: "'a fset \<Rightarrow> 'a" where
"fmax xs = (if (xs = \<lbrace>\<rbrace>) then undefined else last (flist xs))"
end
definition flists :: "'a fset \<Rightarrow> 'a list set" where
"flists A = {xs. distinct xs \<and> finset xs = A}"
lemma flists_nonempty: "\<exists> xs. xs \<in> flists A"
apply (simp add: flists_def)
apply (metis Abs_fset_cases Abs_fset_inverse finite_distinct_list finite_fset finset.rep_eq)
done
lemma flists_elem_uniq: "\<lbrakk> x \<in> flists A; x \<in> flists B \<rbrakk> \<Longrightarrow> A = B"
by (simp add: flists_def)
definition flist_arb :: "'a fset \<Rightarrow> 'a list" where
"flist_arb A = (SOME xs. xs \<in> flists A)"
lemma flist_arb_distinct [simp]: "distinct (flist_arb A)"
by (metis (mono_tags) flist_arb_def flists_def flists_nonempty mem_Collect_eq someI_ex)
lemma flist_arb_inv [simp]: "finset (flist_arb A) = A"
by (metis (mono_tags) flist_arb_def flists_def flists_nonempty mem_Collect_eq someI_ex)
lemma flist_arb_inj:
"inj flist_arb"
by (metis flist_arb_inv injI)
lemma flist_arb_lists: "flist_arb ` Fow A \<subseteq> lists A"
- apply (auto)
- using Fow_def finset.rep_eq apply fastforce
- done
+ apply (auto simp: Fow_def flist_arb_def flists_def)
+ using finset.rep_eq
+ by (metis (mono_tags, lifting) finite_distinct_list finite_fset fset_inverse someI_ex subset_eq)
lemma countable_Fow:
fixes A :: "'a set"
assumes "countable A"
shows "countable (Fow A)"
proof -
from assms obtain to_nat_list :: "'a list \<Rightarrow> nat" where "inj_on to_nat_list (lists A)"
by blast
thus ?thesis
apply (simp add: countable_def)
apply (rule_tac x="to_nat_list \<circ> flist_arb" in exI)
apply (rule comp_inj_on)
apply (metis flist_arb_inv inj_on_def)
apply (simp add: flist_arb_lists subset_inj_on)
done
qed
definition flub :: "'a fset set \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
"flub A t = (if (\<forall> a\<in>A. a |\<subseteq>| t) then Abs_fset (\<Union>x\<in>A. \<langle>x\<rangle>\<^sub>f) else t)"
lemma finite_Union_subsets:
"\<lbrakk> \<forall> a \<in> A. a \<subseteq> b; finite b \<rbrakk> \<Longrightarrow> finite (\<Union>A)"
by (metis Sup_le_iff finite_subset)
lemma finite_UN_subsets:
"\<lbrakk> \<forall> a \<in> A. B a \<subseteq> b; finite b \<rbrakk> \<Longrightarrow> finite (\<Union>a\<in>A. B a)"
by (metis UN_subset_iff finite_subset)
lemma flub_rep_eq:
"\<langle>flub A t\<rangle>\<^sub>f = (if (\<forall> a\<in>A. a |\<subseteq>| t) then (\<Union>x\<in>A. \<langle>x\<rangle>\<^sub>f) else \<langle>t\<rangle>\<^sub>f)"
apply (subgoal_tac "(if (\<forall> a\<in>A. a |\<subseteq>| t) then (\<Union>x\<in>A. \<langle>x\<rangle>\<^sub>f) else \<langle>t\<rangle>\<^sub>f) \<in> {x. finite x}")
apply (auto simp add:flub_def)
apply (rule finite_UN_subsets[of _ _ "\<langle>t\<rangle>\<^sub>f"])
apply (auto)
done
definition fglb :: "'a fset set \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
"fglb A t = (if (A = {}) then t else Abs_fset (\<Inter>x\<in>A. \<langle>x\<rangle>\<^sub>f))"
lemma fglb_rep_eq:
"\<langle>fglb A t\<rangle>\<^sub>f = (if (A = {}) then \<langle>t\<rangle>\<^sub>f else (\<Inter>x\<in>A. \<langle>x\<rangle>\<^sub>f))"
apply (subgoal_tac "(if (A = {}) then \<langle>t\<rangle>\<^sub>f else (\<Inter>x\<in>A. \<langle>x\<rangle>\<^sub>f)) \<in> {x. finite x}")
apply (metis Abs_fset_inverse fglb_def)
apply (auto)
apply (metis finite_INT finite_fset)
done
lemma FinPow_rep_eq [simp]:
"fset (FinPow xs) = {ys. ys |\<subseteq>| xs}"
apply (subgoal_tac "finite (Abs_fset ` Pow \<langle>xs\<rangle>\<^sub>f)")
- apply (auto simp add: fmember_def FinPow_def)
+ apply (auto simp add: FinPow_def)
apply (rename_tac x' y')
apply (subgoal_tac "finite x'")
apply (auto)
apply (metis finite_fset finite_subset)
apply (metis (full_types) Pow_iff fset_inverse imageI less_eq_fset.rep_eq)
done
lemma FUnion_rep_eq [simp]:
"\<langle>\<Union>\<^sub>f xs\<rangle>\<^sub>f = (\<Union>x\<in>\<langle>xs\<rangle>\<^sub>f. \<langle>x\<rangle>\<^sub>f)"
by (simp add:FUnion_def)
lemma FInter_rep_eq [simp]:
"xs \<noteq> \<lbrace>\<rbrace> \<Longrightarrow> \<langle>\<Inter>\<^sub>f xs\<rangle>\<^sub>f = (\<Inter>x\<in>\<langle>xs\<rangle>\<^sub>f. \<langle>x\<rangle>\<^sub>f)"
apply (simp add:FInter_def)
apply (subgoal_tac "finite (\<Inter>x\<in>\<langle>xs\<rangle>\<^sub>f. \<langle>x\<rangle>\<^sub>f)")
apply (simp)
apply (metis (poly_guards_query) bot_fset.rep_eq fglb_rep_eq finite_fset fset_inverse)
done
lemma FUnion_empty [simp]:
"\<Union>\<^sub>f \<lbrace>\<rbrace> = \<lbrace>\<rbrace>"
- by (auto simp add:FUnion_def fmember_def)
+ by (auto simp add:FUnion_def)
lemma FinPow_member [simp]:
"xs |\<in>| FinPow xs"
- by (auto simp add:fmember_def)
+ by auto
lemma FUnion_FinPow [simp]:
"\<Union>\<^sub>f (FinPow x) = x"
- by (auto simp add:fmember_def less_eq_fset_def)
+ by (auto simp add: less_eq_fset_def)
lemma Fow_mem [iff]: "x \<in> Fow A \<longleftrightarrow> \<langle>x\<rangle>\<^sub>f \<subseteq> A"
by (auto simp add:Fow_def)
lemma Fow_UNIV [simp]: "Fow UNIV = UNIV"
by (simp add:Fow_def)
lift_definition FMax :: "('a::linorder) fset \<Rightarrow> 'a" is "Max" .
end
\ No newline at end of file

File Metadata

Mime Type
application/octet-stream
Expires
Sun, Apr 21, 4:45 PM (1 d, 23 h)
Storage Engine
chunks
Storage Format
Chunks
Storage Handle
6aLYXWPP8Vqu
Default Alt Text
(7 MB)

Event Timeline