Theory Gram_Lang
section ‹Language of a Grammar›
theory Gram_Lang
imports DTree "HOL-Library.Infinite_Set"
begin
consts P :: "(N × (T + N) set) set"
axiomatization where
finite_N: "finite (UNIV::N set)"
and finite_in_P: "⋀ n tns. (n,tns) ∈ P ⟶ finite tns"
and used: "⋀ n. ∃ tns. (n,tns) ∈ P"
subsection‹Tree Basics: frontier, interior, etc.›
inductive inFr :: "N set ⇒ dtree ⇒ T ⇒ bool" where
Base: "⟦root tr ∈ ns; Inl t ∈ cont tr⟧ ⟹ inFr ns tr t"
|
Ind: "⟦root tr ∈ ns; Inr tr1 ∈ cont tr; inFr ns tr1 t⟧ ⟹ inFr ns tr t"
definition "Fr ns tr ≡ {t. inFr ns tr t}"
lemma inFr_root_in: "inFr ns tr t ⟹ root tr ∈ ns"
by (metis inFr.simps)
lemma inFr_mono:
assumes "inFr ns tr t" and "ns ⊆ ns'"
shows "inFr ns' tr t"
using assms apply(induct arbitrary: ns' rule: inFr.induct)
using Base Ind by (metis inFr.simps subsetD)+
lemma inFr_Ind_minus:
assumes "inFr ns1 tr1 t" and "Inr tr1 ∈ cont tr"
shows "inFr (insert (root tr) ns1) tr t"
using assms apply(induct rule: inFr.induct)
apply (metis inFr.simps insert_iff)
by (metis inFr.simps inFr_mono insertI1 subset_insertI)
inductive inFr2 :: "N set ⇒ dtree ⇒ T ⇒ bool" where
Base: "⟦root tr ∈ ns; Inl t ∈ cont tr⟧ ⟹ inFr2 ns tr t"
|
Ind: "⟦Inr tr1 ∈ cont tr; inFr2 ns1 tr1 t⟧
⟹ inFr2 (insert (root tr) ns1) tr t"
lemma inFr2_root_in: "inFr2 ns tr t ⟹ root tr ∈ ns"
apply(induct rule: inFr2.induct) by auto
lemma inFr2_mono:
assumes "inFr2 ns tr t" and "ns ⊆ ns'"
shows "inFr2 ns' tr t"
using assms apply(induct arbitrary: ns' rule: inFr2.induct)
using Base Ind
apply (metis subsetD) by (metis inFr2.simps insert_absorb insert_subset)
lemma inFr2_Ind:
assumes "inFr2 ns tr1 t" and "root tr ∈ ns" and "Inr tr1 ∈ cont tr"
shows "inFr2 ns tr t"
using assms apply(induct rule: inFr2.induct)
apply (metis inFr2.simps insert_absorb)
by (metis inFr2.simps insert_absorb)
lemma inFr_inFr2:
"inFr = inFr2"
apply (rule ext)+ apply(safe)
apply(erule inFr.induct)
apply (metis (lifting) inFr2.Base)
apply (metis (lifting) inFr2_Ind)
apply(erule inFr2.induct)
apply (metis (lifting) inFr.Base)
apply (metis (lifting) inFr_Ind_minus)
done
lemma not_root_inFr:
assumes "root tr ∉ ns"
shows "¬ inFr ns tr t"
by (metis assms inFr_root_in)
lemma not_root_Fr:
assumes "root tr ∉ ns"
shows "Fr ns tr = {}"
using not_root_inFr[OF assms] unfolding Fr_def by auto
inductive inItr :: "N set ⇒ dtree ⇒ N ⇒ bool" where
Base: "root tr ∈ ns ⟹ inItr ns tr (root tr)"
|
Ind: "⟦root tr ∈ ns; Inr tr1 ∈ cont tr; inItr ns tr1 n⟧ ⟹ inItr ns tr n"
definition "Itr ns tr ≡ {n. inItr ns tr n}"
lemma inItr_root_in: "inItr ns tr n ⟹ root tr ∈ ns"
by (metis inItr.simps)
lemma inItr_mono:
assumes "inItr ns tr n" and "ns ⊆ ns'"
shows "inItr ns' tr n"
using assms apply(induct arbitrary: ns' rule: inItr.induct)
using Base Ind by (metis inItr.simps subsetD)+
inductive subtr where
Refl: "root tr ∈ ns ⟹ subtr ns tr tr"
|
Step: "⟦root tr3 ∈ ns; subtr ns tr1 tr2; Inr tr2 ∈ cont tr3⟧ ⟹ subtr ns tr1 tr3"
lemma subtr_rootL_in:
assumes "subtr ns tr1 tr2"
shows "root tr1 ∈ ns"
using assms apply(induct rule: subtr.induct) by auto
lemma subtr_rootR_in:
assumes "subtr ns tr1 tr2"
shows "root tr2 ∈ ns"
using assms apply(induct rule: subtr.induct) by auto
lemmas subtr_roots_in = subtr_rootL_in subtr_rootR_in
lemma subtr_mono:
assumes "subtr ns tr1 tr2" and "ns ⊆ ns'"
shows "subtr ns' tr1 tr2"
using assms apply(induct arbitrary: ns' rule: subtr.induct)
using Refl Step by (metis subtr.simps subsetD)+
lemma subtr_trans_Un:
assumes "subtr ns12 tr1 tr2" and "subtr ns23 tr2 tr3"
shows "subtr (ns12 ∪ ns23) tr1 tr3"
proof-
have "subtr ns23 tr2 tr3 ⟹
(∀ ns12 tr1. subtr ns12 tr1 tr2 ⟶ subtr (ns12 ∪ ns23) tr1 tr3)"
apply(induct rule: subtr.induct, safe)
apply (metis subtr_mono sup_commute sup_ge2)
by (metis (lifting) Step UnI2)
thus ?thesis using assms by auto
qed
lemma subtr_trans:
assumes "subtr ns tr1 tr2" and "subtr ns tr2 tr3"
shows "subtr ns tr1 tr3"
using subtr_trans_Un[OF assms] by simp
lemma subtr_StepL:
assumes r: "root tr1 ∈ ns" and tr12: "Inr tr1 ∈ cont tr2" and s: "subtr ns tr2 tr3"
shows "subtr ns tr1 tr3"
apply(rule subtr_trans[OF _ s])
apply(rule Step[of tr2 ns tr1 tr1])
apply(rule subtr_rootL_in[OF s])
apply(rule Refl[OF r])
apply(rule tr12)
done
inductive subtr2 where
Refl: "root tr ∈ ns ⟹ subtr2 ns tr tr"
|
Step: "⟦root tr1 ∈ ns; Inr tr1 ∈ cont tr2; subtr2 ns tr2 tr3⟧ ⟹ subtr2 ns tr1 tr3"
lemma subtr2_rootL_in:
assumes "subtr2 ns tr1 tr2"
shows "root tr1 ∈ ns"
using assms apply(induct rule: subtr2.induct) by auto
lemma subtr2_rootR_in:
assumes "subtr2 ns tr1 tr2"
shows "root tr2 ∈ ns"
using assms apply(induct rule: subtr2.induct) by auto
lemmas subtr2_roots_in = subtr2_rootL_in subtr2_rootR_in
lemma subtr2_mono:
assumes "subtr2 ns tr1 tr2" and "ns ⊆ ns'"
shows "subtr2 ns' tr1 tr2"
using assms apply(induct arbitrary: ns' rule: subtr2.induct)
using Refl Step by (metis subtr2.simps subsetD)+
lemma subtr2_trans_Un:
assumes "subtr2 ns12 tr1 tr2" and "subtr2 ns23 tr2 tr3"
shows "subtr2 (ns12 ∪ ns23) tr1 tr3"
proof-
have "subtr2 ns12 tr1 tr2 ⟹
(∀ ns23 tr3. subtr2 ns23 tr2 tr3 ⟶ subtr2 (ns12 ∪ ns23) tr1 tr3)"
apply(induct rule: subtr2.induct, safe)
apply (metis subtr2_mono sup_commute sup_ge2)
by (metis Un_iff subtr2.simps)
thus ?thesis using assms by auto
qed
lemma subtr2_trans:
assumes "subtr2 ns tr1 tr2" and "subtr2 ns tr2 tr3"
shows "subtr2 ns tr1 tr3"
using subtr2_trans_Un[OF assms] by simp
lemma subtr2_StepR:
assumes r: "root tr3 ∈ ns" and tr23: "Inr tr2 ∈ cont tr3" and s: "subtr2 ns tr1 tr2"
shows "subtr2 ns tr1 tr3"
apply(rule subtr2_trans[OF s])
apply(rule Step[of _ _ tr3])
apply(rule subtr2_rootR_in[OF s])
apply(rule tr23)
apply(rule Refl[OF r])
done
lemma subtr_subtr2:
"subtr = subtr2"
apply (rule ext)+ apply(safe)
apply(erule subtr.induct)
apply (metis (lifting) subtr2.Refl)
apply (metis (lifting) subtr2_StepR)
apply(erule subtr2.induct)
apply (metis (lifting) subtr.Refl)
apply (metis (lifting) subtr_StepL)
done
lemma subtr_inductL[consumes 1, case_names Refl Step]:
assumes s: "subtr ns tr1 tr2" and Refl: "⋀ns tr. φ ns tr tr"
and Step:
"⋀ns tr1 tr2 tr3.
⟦root tr1 ∈ ns; Inr tr1 ∈ cont tr2; subtr ns tr2 tr3; φ ns tr2 tr3⟧ ⟹ φ ns tr1 tr3"
shows "φ ns tr1 tr2"
using s unfolding subtr_subtr2 apply(rule subtr2.induct)
using Refl Step unfolding subtr_subtr2 by auto
lemma subtr_UNIV_inductL[consumes 1, case_names Refl Step]:
assumes s: "subtr UNIV tr1 tr2" and Refl: "⋀tr. φ tr tr"
and Step:
"⋀tr1 tr2 tr3.
⟦Inr tr1 ∈ cont tr2; subtr UNIV tr2 tr3; φ tr2 tr3⟧ ⟹ φ tr1 tr3"
shows "φ tr1 tr2"
using s apply(induct rule: subtr_inductL)
apply(rule Refl) using Step subtr_mono by (metis subset_UNIV)
lemma subtr_inFr:
assumes "inFr ns tr t" and "subtr ns tr tr1"
shows "inFr ns tr1 t"
proof-
have "subtr ns tr tr1 ⟹ (∀ t. inFr ns tr t ⟶ inFr ns tr1 t)"
apply(induct rule: subtr.induct, safe) by (metis inFr.Ind)
thus ?thesis using assms by auto
qed
corollary Fr_subtr:
"Fr ns tr = ⋃{Fr ns tr' | tr'. subtr ns tr' tr}"
unfolding Fr_def proof safe
fix t assume t: "inFr ns tr t" hence "root tr ∈ ns" by (rule inFr_root_in)
thus "t ∈ ⋃{{t. inFr ns tr' t} |tr'. subtr ns tr' tr}"
apply(intro UnionI[of "{t. inFr ns tr t}" _ t]) using t subtr.Refl by auto
qed(metis subtr_inFr)
lemma inFr_subtr:
assumes "inFr ns tr t"
shows "∃ tr'. subtr ns tr' tr ∧ Inl t ∈ cont tr'"
using assms apply(induct rule: inFr.induct) apply safe
apply (metis subtr.Refl)
by (metis (lifting) subtr.Step)
corollary Fr_subtr_cont:
"Fr ns tr = ⋃{Inl -` cont tr' | tr'. subtr ns tr' tr}"
unfolding Fr_def
apply safe
apply (frule inFr_subtr)
apply auto
by (metis inFr.Base subtr_inFr subtr_rootL_in)
lemma subtr_inItr:
assumes "inItr ns tr n" and "subtr ns tr tr1"
shows "inItr ns tr1 n"
proof-
have "subtr ns tr tr1 ⟹ (∀ t. inItr ns tr n ⟶ inItr ns tr1 n)"
apply(induct rule: subtr.induct, safe) by (metis inItr.Ind)
thus ?thesis using assms by auto
qed
corollary Itr_subtr:
"Itr ns tr = ⋃{Itr ns tr' | tr'. subtr ns tr' tr}"
unfolding Itr_def apply safe
apply (metis (lifting, mono_tags) UnionI inItr_root_in mem_Collect_eq subtr.Refl)
by (metis subtr_inItr)
lemma inItr_subtr:
assumes "inItr ns tr n"
shows "∃ tr'. subtr ns tr' tr ∧ root tr' = n"
using assms apply(induct rule: inItr.induct) apply safe
apply (metis subtr.Refl)
by (metis (lifting) subtr.Step)
corollary Itr_subtr_cont:
"Itr ns tr = {root tr' | tr'. subtr ns tr' tr}"
unfolding Itr_def apply safe
apply (metis (lifting, mono_tags) inItr_subtr)
by (metis inItr.Base subtr_inItr subtr_rootL_in)
subsection‹The Immediate Subtree Function›
abbreviation "prodOf tr ≡ (id ⊕ root) ` (cont tr)"
definition "subtrOf tr n ≡ SOME tr'. Inr tr' ∈ cont tr ∧ root tr' = n"
lemma subtrOf:
assumes n: "Inr n ∈ prodOf tr"
shows "Inr (subtrOf tr n) ∈ cont tr ∧ root (subtrOf tr n) = n"
proof-
obtain tr' where "Inr tr' ∈ cont tr ∧ root tr' = n"
using n unfolding image_def by (metis (lifting) Inr_oplus_elim assms)
thus ?thesis unfolding subtrOf_def by(rule someI)
qed
lemmas Inr_subtrOf = subtrOf[THEN conjunct1]
lemmas root_subtrOf[simp] = subtrOf[THEN conjunct2]
lemma Inl_prodOf: "Inl -` (prodOf tr) = Inl -` (cont tr)"
proof safe
fix t ttr assume "Inl t = (id ⊕ root) ttr" and "ttr ∈ cont tr"
thus "t ∈ Inl -` cont tr" by(cases ttr, auto)
next
fix t assume "Inl t ∈ cont tr" thus "t ∈ Inl -` prodOf tr"
by (metis (lifting) id_def image_iff map_sum.simps(1) vimageI2)
qed
lemma root_prodOf:
assumes "Inr tr' ∈ cont tr"
shows "Inr (root tr') ∈ prodOf tr"
by (metis (lifting) assms image_iff map_sum.simps(2))
subsection‹Well-Formed Derivation Trees›
hide_const wf
coinductive wf where
dtree: "⟦(root tr, (id ⊕ root) ` (cont tr)) ∈ P; inj_on root (Inr -` cont tr);
⋀ tr'. tr' ∈ Inr -` (cont tr) ⟹ wf tr'⟧ ⟹ wf tr"
lemma wf_P:
assumes "wf tr"
shows "(root tr, (id ⊕ root) ` (cont tr)) ∈ P"
using assms wf.simps[of tr] by auto
lemma wf_inj_on:
assumes "wf tr"
shows "inj_on root (Inr -` cont tr)"
using assms wf.simps[of tr] by auto
lemma wf_inj[simp]:
assumes "wf tr" and "Inr tr1 ∈ cont tr" and "Inr tr2 ∈ cont tr"
shows "root tr1 = root tr2 ⟷ tr1 = tr2"
using assms wf_inj_on unfolding inj_on_def by auto
lemma wf_cont:
assumes "wf tr" and "Inr tr' ∈ cont tr"
shows "wf tr'"
using assms wf.simps[of tr] by auto
lemma wf_coind[elim, consumes 1, case_names Hyp]:
assumes phi: "φ tr"
and Hyp:
"⋀ tr. φ tr ⟹
(root tr, image (id ⊕ root) (cont tr)) ∈ P ∧
inj_on root (Inr -` cont tr) ∧
(∀ tr' ∈ Inr -` (cont tr). φ tr' ∨ wf tr')"
shows "wf tr"
apply(rule wf.coinduct[of φ tr, OF phi])
using Hyp by blast
lemma wf_raw_coind[elim, consumes 1, case_names Hyp]:
assumes phi: "φ tr"
and Hyp:
"⋀ tr. φ tr ⟹
(root tr, image (id ⊕ root) (cont tr)) ∈ P ∧
inj_on root (Inr -` cont tr) ∧
(∀ tr' ∈ Inr -` (cont tr). φ tr')"
shows "wf tr"
using phi apply(induct rule: wf_coind)
using Hyp by (metis (mono_tags))
lemma wf_subtr_inj_on:
assumes d: "wf tr1" and s: "subtr ns tr tr1"
shows "inj_on root (Inr -` cont tr)"
using s d apply(induct rule: subtr.induct)
apply (metis (lifting) wf_inj_on) by (metis wf_cont)
lemma wf_subtr_P:
assumes d: "wf tr1" and s: "subtr ns tr tr1"
shows "(root tr, (id ⊕ root) ` cont tr) ∈ P"
using s d apply(induct rule: subtr.induct)
apply (metis (lifting) wf_P) by (metis wf_cont)
lemma subtrOf_root[simp]:
assumes tr: "wf tr" and cont: "Inr tr' ∈ cont tr"
shows "subtrOf tr (root tr') = tr'"
proof-
have 0: "Inr (subtrOf tr (root tr')) ∈ cont tr" using Inr_subtrOf
by (metis (lifting) cont root_prodOf)
have "root (subtrOf tr (root tr')) = root tr'"
using root_subtrOf by (metis (lifting) cont root_prodOf)
thus ?thesis unfolding wf_inj[OF tr 0 cont] .
qed
lemma surj_subtrOf:
assumes "wf tr" and 0: "Inr tr' ∈ cont tr"
shows "∃ n. Inr n ∈ prodOf tr ∧ subtrOf tr n = tr'"
apply(rule exI[of _ "root tr'"])
using root_prodOf[OF 0] subtrOf_root[OF assms] by simp
lemma wf_subtr:
assumes "wf tr1" and "subtr ns tr tr1"
shows "wf tr"
proof-
have "(∃ ns tr1. wf tr1 ∧ subtr ns tr tr1) ⟹ wf tr"
proof (induct rule: wf_raw_coind)
case (Hyp tr)
then obtain ns tr1 where tr1: "wf tr1" and tr_tr1: "subtr ns tr tr1" by auto
show ?case proof safe
show "(root tr, (id ⊕ root) ` cont tr) ∈ P" using wf_subtr_P[OF tr1 tr_tr1] .
next
show "inj_on root (Inr -` cont tr)" using wf_subtr_inj_on[OF tr1 tr_tr1] .
next
fix tr' assume tr': "Inr tr' ∈ cont tr"
have tr_tr1: "subtr (ns ∪ {root tr'}) tr tr1" using subtr_mono[OF tr_tr1] by auto
have "subtr (ns ∪ {root tr'}) tr' tr1" using subtr_StepL[OF _ tr' tr_tr1] by auto
thus "∃ns' tr1. wf tr1 ∧ subtr ns' tr' tr1" using tr1 by blast
qed
qed
thus ?thesis using assms by auto
qed
subsection‹Default Trees›
definition S where "S n ≡ SOME tns. (n,tns) ∈ P"
lemma S_P: "(n, S n) ∈ P"
using used unfolding S_def by(rule someI_ex)
lemma finite_S: "finite (S n)"
using S_P finite_in_P by auto
definition deftr :: "N ⇒ dtree" where
"deftr ≡ unfold id S"
lemma deftr_simps[simp]:
"root (deftr n) = n"
"cont (deftr n) = image (id ⊕ deftr) (S n)"
using unfold(1)[of id S n] unfold(2)[of S n id, OF finite_S]
unfolding deftr_def by simp_all
lemmas root_deftr = deftr_simps(1)
lemmas cont_deftr = deftr_simps(2)
lemma root_o_deftr[simp]: "root o deftr = id"
by (rule ext, auto)
lemma wf_deftr: "wf (deftr n)"
proof-
{fix tr assume "∃ n. tr = deftr n" hence "wf tr"
apply(induct rule: wf_raw_coind) apply safe
apply (simp only: deftr_simps image_comp map_sum.comp id_comp
root_o_deftr map_sum.id image_id id_apply) apply(rule S_P)
unfolding inj_on_def by auto
}
thus ?thesis by auto
qed
subsection‹Hereditary Substitution›
definition "inFrr ns tr t ≡ ∃ tr'. Inr tr' ∈ cont tr ∧ inFr ns tr' t"
definition "Frr ns tr ≡ {t. ∃ tr'. Inr tr' ∈ cont tr ∧ t ∈ Fr ns tr'}"
context
fixes tr0 :: dtree
begin
definition "hsubst_r tr ≡ root tr"
definition "hsubst_c tr ≡ if root tr = root tr0 then cont tr0 else cont tr"
definition hsubst :: "dtree ⇒ dtree" where
"hsubst ≡ unfold hsubst_r hsubst_c"
lemma finite_hsubst_c: "finite (hsubst_c n)"
unfolding hsubst_c_def by (metis (full_types) finite_cont)
lemma root_hsubst[simp]: "root (hsubst tr) = root tr"
using unfold(1)[of hsubst_r hsubst_c tr] unfolding hsubst_def hsubst_r_def by simp
lemma root_o_subst[simp]: "root o hsubst = root"
unfolding comp_def root_hsubst ..
lemma cont_hsubst_eq[simp]:
assumes "root tr = root tr0"
shows "cont (hsubst tr) = (id ⊕ hsubst) ` (cont tr0)"
apply(subst id_comp[symmetric, of id]) unfolding id_comp
using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
unfolding hsubst_def hsubst_c_def using assms by simp
lemma hsubst_eq:
assumes "root tr = root tr0"
shows "hsubst tr = hsubst tr0"
apply(rule dtree_cong) using assms cont_hsubst_eq by auto
lemma cont_hsubst_neq[simp]:
assumes "root tr ≠ root tr0"
shows "cont (hsubst tr) = (id ⊕ hsubst) ` (cont tr)"
apply(subst id_comp[symmetric, of id]) unfolding id_comp
using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
unfolding hsubst_def hsubst_c_def using assms by simp
lemma Inl_cont_hsubst_eq[simp]:
assumes "root tr = root tr0"
shows "Inl -` cont (hsubst tr) = Inl -` (cont tr0)"
unfolding cont_hsubst_eq[OF assms] by simp
lemma Inr_cont_hsubst_eq[simp]:
assumes "root tr = root tr0"
shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr0"
unfolding cont_hsubst_eq[OF assms] by simp
lemma Inl_cont_hsubst_neq[simp]:
assumes "root tr ≠ root tr0"
shows "Inl -` cont (hsubst tr) = Inl -` (cont tr)"
unfolding cont_hsubst_neq[OF assms] by simp
lemma Inr_cont_hsubst_neq[simp]:
assumes "root tr ≠ root tr0"
shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr"
unfolding cont_hsubst_neq[OF assms] by simp
lemma wf_hsubst:
assumes tr0: "wf tr0" and tr: "wf tr"
shows "wf (hsubst tr)"
proof-
{fix tr1 have "(∃ tr. wf tr ∧ tr1 = hsubst tr) ⟹ wf tr1"
proof (induct rule: wf_raw_coind)
case (Hyp tr1) then obtain tr
where dtr: "wf tr" and tr1: "tr1 = hsubst tr" by auto
show ?case unfolding tr1 proof safe
show "(root (hsubst tr), prodOf (hsubst tr)) ∈ P"
unfolding tr1 apply(cases "root tr = root tr0")
using wf_P[OF dtr] wf_P[OF tr0]
by (auto simp add: image_comp map_sum.comp)
show "inj_on root (Inr -` cont (hsubst tr))"
apply(cases "root tr = root tr0") using wf_inj_on[OF dtr] wf_inj_on[OF tr0]
unfolding inj_on_def by (auto, blast)
fix tr' assume "Inr tr' ∈ cont (hsubst tr)"
thus "∃tra. wf tra ∧ tr' = hsubst tra"
apply(cases "root tr = root tr0", simp_all)
apply (metis wf_cont tr0)
by (metis dtr wf_cont)
qed
qed
}
thus ?thesis using assms by blast
qed
lemma Frr: "Frr ns tr = {t. inFrr ns tr t}"
unfolding inFrr_def Frr_def Fr_def by auto
lemma inFr_hsubst_imp:
assumes "inFr ns (hsubst tr) t"
shows "t ∈ Inl -` (cont tr0) ∨ inFrr (ns - {root tr0}) tr0 t ∨
inFr (ns - {root tr0}) tr t"
proof-
{fix tr1
have "inFr ns tr1 t ⟹
(⋀ tr. tr1 = hsubst tr ⟹ (t ∈ Inl -` (cont tr0) ∨ inFrr (ns - {root tr0}) tr0 t ∨
inFr (ns - {root tr0}) tr t))"
proof(induct rule: inFr.induct)
case (Base tr1 ns t tr)
hence rtr: "root tr1 ∈ ns" and t_tr1: "Inl t ∈ cont tr1" and tr1: "tr1 = hsubst tr"
by auto
show ?case
proof(cases "root tr1 = root tr0")
case True
hence "t ∈ Inl -` (cont tr0)" using t_tr1 unfolding tr1 by auto
thus ?thesis by simp
next
case False
hence "inFr (ns - {root tr0}) tr t" using t_tr1 unfolding tr1 apply simp
by (metis Base.prems Diff_iff root_hsubst inFr.Base rtr singletonE)
thus ?thesis by simp
qed
next
case (Ind tr1 ns tr1' t) note IH = Ind(4)
have rtr1: "root tr1 ∈ ns" and tr1'_tr1: "Inr tr1' ∈ cont tr1"
and t_tr1': "inFr ns tr1' t" and tr1: "tr1 = hsubst tr" using Ind by auto
have rtr1: "root tr1 = root tr" unfolding tr1 by simp
show ?case
proof(cases "root tr1 = root tr0")
case True
then obtain tr' where tr'_tr0: "Inr tr' ∈ cont tr0" and tr1': "tr1' = hsubst tr'"
using tr1'_tr1 unfolding tr1 by auto
show ?thesis using IH[OF tr1'] proof (elim disjE)
assume "inFr (ns - {root tr0}) tr' t"
thus ?thesis using tr'_tr0 unfolding inFrr_def by auto
qed auto
next
case False
then obtain tr' where tr'_tr: "Inr tr' ∈ cont tr" and tr1': "tr1' = hsubst tr'"
using tr1'_tr1 unfolding tr1 by auto
show ?thesis using IH[OF tr1'] proof (elim disjE)
assume "inFr (ns - {root tr0}) tr' t"
thus ?thesis using tr'_tr unfolding inFrr_def
by (metis Diff_iff False Ind(1) empty_iff inFr2_Ind inFr_inFr2 insert_iff rtr1)
qed auto
qed
qed
}
thus ?thesis using assms by auto
qed
lemma inFr_hsubst_notin:
assumes "inFr ns tr t" and "root tr0 ∉ ns"
shows "inFr ns (hsubst tr) t"
using assms apply(induct rule: inFr.induct)
apply (metis Inl_cont_hsubst_neq inFr2.Base inFr_inFr2 root_hsubst vimageD vimageI2)
by (metis (lifting) Inr_cont_hsubst_neq inFr.Ind rev_image_eqI root_hsubst vimageD vimageI2)
lemma inFr_hsubst_minus:
assumes "inFr (ns - {root tr0}) tr t"
shows "inFr ns (hsubst tr) t"
proof-
have 1: "inFr (ns - {root tr0}) (hsubst tr) t"
using inFr_hsubst_notin[OF assms] by simp
show ?thesis using inFr_mono[OF 1] by auto
qed
lemma inFr_self_hsubst:
assumes "root tr0 ∈ ns"
shows
"inFr ns (hsubst tr0) t ⟷
t ∈ Inl -` (cont tr0) ∨ inFrr (ns - {root tr0}) tr0 t"
(is "?A ⟷ ?B ∨ ?C")
apply(intro iffI)
apply (metis inFr_hsubst_imp Diff_iff inFr_root_in insertI1) proof(elim disjE)
assume ?B thus ?A apply(intro inFr.Base) using assms by auto
next
assume ?C then obtain tr where
tr_tr0: "Inr tr ∈ cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t"
unfolding inFrr_def by auto
define tr1 where "tr1 = hsubst tr"
have 1: "inFr ns tr1 t" using t_tr unfolding tr1_def using inFr_hsubst_minus by auto
have "Inr tr1 ∈ cont (hsubst tr0)" unfolding tr1_def using tr_tr0 by auto
thus ?A using 1 inFr.Ind assms by (metis root_hsubst)
qed
lemma Fr_self_hsubst:
assumes "root tr0 ∈ ns"
shows "Fr ns (hsubst tr0) = Inl -` (cont tr0) ∪ Frr (ns - {root tr0}) tr0"
using inFr_self_hsubst[OF assms] unfolding Frr Fr_def by auto
end
subsection‹Regular Trees›
definition "reg f tr ≡ ∀ tr'. subtr UNIV tr' tr ⟶ tr' = f (root tr')"
definition "regular tr ≡ ∃ f. reg f tr"
lemma reg_def2: "reg f tr ⟷ (∀ ns tr'. subtr ns tr' tr ⟶ tr' = f (root tr'))"
unfolding reg_def using subtr_mono by (metis subset_UNIV)
lemma regular_def2: "regular tr ⟷ (∃ f. reg f tr ∧ (∀ n. root (f n) = n))"
unfolding regular_def proof safe
fix f
assume f: "reg f tr"
define g where "g n = (if inItr UNIV tr n then f n else deftr n)" for n
show "∃g. reg g tr ∧ (∀n. root (g n) = n)"
apply(rule exI[of _ g])
using f deftr_simps(1) unfolding g_def reg_def apply safe
apply (metis (lifting) inItr.Base subtr_inItr subtr_rootL_in)
by (metis (full_types) inItr_subtr)
qed auto
lemma reg_root:
assumes "reg f tr"
shows "f (root tr) = tr"
using assms unfolding reg_def
by (metis (lifting) iso_tuple_UNIV_I subtr.Refl)
lemma reg_Inr_cont:
assumes "reg f tr" and "Inr tr' ∈ cont tr"
shows "reg f tr'"
by (metis (lifting) assms iso_tuple_UNIV_I reg_def subtr.Step)
lemma reg_subtr:
assumes "reg f tr" and "subtr ns tr' tr"
shows "reg f tr'"
using assms unfolding reg_def using subtr_trans[of UNIV tr] UNIV_I
by (metis UNIV_eq_I UnCI Un_upper1 iso_tuple_UNIV_I subtr_mono subtr_trans)
lemma regular_subtr:
assumes r: "regular tr" and s: "subtr ns tr' tr"
shows "regular tr'"
using r reg_subtr[OF _ s] unfolding regular_def by auto
lemma subtr_deftr:
assumes "subtr ns tr' (deftr n)"
shows "tr' = deftr (root tr')"
proof-
{fix tr have "subtr ns tr' tr ⟹ (∀ n. tr = deftr n ⟶ tr' = deftr (root tr'))"
apply (induct rule: subtr.induct)
proof(metis (lifting) deftr_simps(1), safe)
fix tr3 ns tr1 tr2 n
assume 1: "root (deftr n) ∈ ns" and 2: "subtr ns tr1 tr2"
and IH: "∀n. tr2 = deftr n ⟶ tr1 = deftr (root tr1)"
and 3: "Inr tr2 ∈ cont (deftr n)"
have "tr2 ∈ deftr ` UNIV"
using 3 unfolding deftr_simps image_def
by (metis (lifting, full_types) 3 CollectI Inr_oplus_iff cont_deftr
iso_tuple_UNIV_I)
then obtain n where "tr2 = deftr n" by auto
thus "tr1 = deftr (root tr1)" using IH by auto
qed
}
thus ?thesis using assms by auto
qed
lemma reg_deftr: "reg deftr (deftr n)"
unfolding reg_def using subtr_deftr by auto
lemma wf_subtrOf_Union:
assumes "wf tr"
shows "⋃{K tr' |tr'. Inr tr' ∈ cont tr} =
⋃{K (subtrOf tr n) |n. Inr n ∈ prodOf tr}"
unfolding Union_eq Bex_def mem_Collect_eq proof safe
fix x xa tr'
assume x: "x ∈ K tr'" and tr'_tr: "Inr tr' ∈ cont tr"
show "∃X. (∃n. X = K (subtrOf tr n) ∧ Inr n ∈ prodOf tr) ∧ x ∈ X"
apply(rule exI[of _ "K (subtrOf tr (root tr'))"]) apply(intro conjI)
apply(rule exI[of _ "root tr'"]) apply (metis (lifting) root_prodOf tr'_tr)
by (metis (lifting) assms subtrOf_root tr'_tr x)
next
fix x X n ttr
assume x: "x ∈ K (subtrOf tr n)" and n: "Inr n = (id ⊕ root) ttr" and ttr: "ttr ∈ cont tr"
show "∃X. (∃tr'. X = K tr' ∧ Inr tr' ∈ cont tr) ∧ x ∈ X"
apply(rule exI[of _ "K (subtrOf tr n)"]) apply(intro conjI)
apply(rule exI[of _ "subtrOf tr n"]) apply (metis imageI n subtrOf ttr)
using x .
qed
subsection ‹Paths in a Regular Tree›
inductive path :: "(N ⇒ dtree) ⇒ N list ⇒ bool" for f where
Base: "path f [n]"
|
Ind: "⟦path f (n1 # nl); Inr (f n1) ∈ cont (f n)⟧
⟹ path f (n # n1 # nl)"
lemma path_NE:
assumes "path f nl"
shows "nl ≠ Nil"
using assms apply(induct rule: path.induct) by auto
lemma path_post:
assumes f: "path f (n # nl)" and nl: "nl ≠ []"
shows "path f nl"
proof-
obtain n1 nl1 where nl: "nl = n1 # nl1" using nl by (cases nl, auto)
show ?thesis using assms unfolding nl using path.simps by (metis (lifting) list.inject)
qed
lemma path_post_concat:
assumes "path f (nl1 @ nl2)" and "nl2 ≠ Nil"
shows "path f nl2"
using assms apply (induct nl1)
apply (metis append_Nil) by (metis Nil_is_append_conv append_Cons path_post)
lemma path_concat:
assumes "path f nl1" and "path f ((last nl1) # nl2)"
shows "path f (nl1 @ nl2)"
using assms apply(induct rule: path.induct) apply simp
by (metis append_Cons last.simps list.simps(3) path.Ind)
lemma path_distinct:
assumes "path f nl"
shows "∃ nl'. path f nl' ∧ hd nl' = hd nl ∧ last nl' = last nl ∧
set nl' ⊆ set nl ∧ distinct nl'"
using assms proof(induct rule: length_induct)
case (1 nl) hence p_nl: "path f nl" by simp
then obtain n nl1 where nl: "nl = n # nl1" by (metis list.exhaust path_NE)
show ?case
proof(cases nl1)
case Nil
show ?thesis apply(rule exI[of _ nl]) using path.Base unfolding nl Nil by simp
next
case (Cons n1 nl2)
hence p1: "path f nl1" by (metis list.simps(3) nl p_nl path_post)
show ?thesis
proof(cases "n ∈ set nl1")
case False
obtain nl1' where p1': "path f nl1'" and hd_nl1': "hd nl1' = hd nl1" and
l_nl1': "last nl1' = last nl1" and d_nl1': "distinct nl1'"
and s_nl1': "set nl1' ⊆ set nl1"
using 1(1)[THEN allE[of _ nl1]] p1 unfolding nl by auto
obtain nl2' where nl1': "nl1' = n1 # nl2'" using path_NE[OF p1'] hd_nl1'
unfolding Cons by(cases nl1', auto)
show ?thesis apply(intro exI[of _ "n # nl1'"]) unfolding nl proof safe
show "path f (n # nl1')" unfolding nl1'
apply(rule path.Ind, metis nl1' p1')
by (metis (lifting) Cons list.inject nl p1 p_nl path.simps path_NE)
qed(insert l_nl1' Cons nl1' s_nl1' d_nl1' False, auto)
next
case True
then obtain nl11 nl12 where nl1: "nl1 = nl11 @ n # nl12"
by (metis split_list)
have p12: "path f (n # nl12)"
apply(rule path_post_concat[of _ "n # nl11"]) using p_nl[unfolded nl nl1] by auto
obtain nl12' where p1': "path f nl12'" and hd_nl12': "hd nl12' = n" and
l_nl12': "last nl12' = last (n # nl12)" and d_nl12': "distinct nl12'"
and s_nl12': "set nl12' ⊆ {n} ∪ set nl12"
using 1(1)[THEN allE[of _ "n # nl12"]] p12 unfolding nl nl1 by auto
thus ?thesis apply(intro exI[of _ nl12']) unfolding nl nl1 by auto
qed
qed
qed
lemma path_subtr:
assumes f: "⋀ n. root (f n) = n"
and p: "path f nl"
shows "subtr (set nl) (f (last nl)) (f (hd nl))"
using p proof (induct rule: path.induct)
case (Ind n1 nl n) let ?ns1 = "insert n1 (set nl)"
have "path f (n1 # nl)"
and "subtr ?ns1 (f (last (n1 # nl))) (f n1)"
and fn1: "Inr (f n1) ∈ cont (f n)" using Ind by simp_all
hence fn1_flast: "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n1)"
by (metis subset_insertI subtr_mono)
have 1: "last (n # n1 # nl) = last (n1 # nl)" by auto
have "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n)"
using f subtr.Step[OF _ fn1_flast fn1] by auto
thus ?case unfolding 1 by simp
qed (metis f list.sel(1) last_ConsL last_in_set not_Cons_self2 subtr.Refl)
lemma reg_subtr_path_aux:
assumes f: "reg f tr" and n: "subtr ns tr1 tr"
shows "∃ nl. path f nl ∧ f (hd nl) = tr ∧ f (last nl) = tr1 ∧ set nl ⊆ ns"
using n f proof(induct rule: subtr.induct)
case (Refl tr ns)
thus ?case
apply(intro exI[of _ "[root tr]"]) apply simp by (metis (lifting) path.Base reg_root)
next
case (Step tr ns tr2 tr1)
hence rtr: "root tr ∈ ns" and tr1_tr: "Inr tr1 ∈ cont tr"
and tr2_tr1: "subtr ns tr2 tr1" and tr: "reg f tr" by auto
have tr1: "reg f tr1" using reg_subtr[OF tr] rtr tr1_tr
by (metis (lifting) Step.prems iso_tuple_UNIV_I reg_def subtr.Step)
obtain nl where nl: "path f nl" and f_nl: "f (hd nl) = tr1"
and last_nl: "f (last nl) = tr2" and set: "set nl ⊆ ns" using Step(3)[OF tr1] by auto
have 0: "path f (root tr # nl)" apply (subst path.simps)
using f_nl nl reg_root tr tr1_tr by (metis list.sel(1) neq_Nil_conv)
show ?case apply(rule exI[of _ "(root tr) # nl"])
using 0 reg_root tr last_nl nl path_NE rtr set by auto
qed
lemma reg_subtr_path:
assumes f: "reg f tr" and n: "subtr ns tr1 tr"
shows "∃ nl. distinct nl ∧ path f nl ∧ f (hd nl) = tr ∧ f (last nl) = tr1 ∧ set nl ⊆ ns"
using reg_subtr_path_aux[OF assms] path_distinct[of f]
by (metis (lifting) order_trans)
lemma subtr_iff_path:
assumes r: "reg f tr" and f: "⋀ n. root (f n) = n"
shows "subtr ns tr1 tr ⟷
(∃ nl. distinct nl ∧ path f nl ∧ f (hd nl) = tr ∧ f (last nl) = tr1 ∧ set nl ⊆ ns)"
proof safe
fix nl assume p: "path f nl" and nl: "set nl ⊆ ns"
have "subtr (set nl) (f (last nl)) (f (hd nl))"
apply(rule path_subtr) using p f by simp_all
thus "subtr ns (f (last nl)) (f (hd nl))"
using subtr_mono nl by auto
qed(insert reg_subtr_path[OF r], auto)
lemma inFr_iff_path:
assumes r: "reg f tr" and f: "⋀ n. root (f n) = n"
shows
"inFr ns tr t ⟷
(∃ nl tr1. distinct nl ∧ path f nl ∧ f (hd nl) = tr ∧ f (last nl) = tr1 ∧
set nl ⊆ ns ∧ Inl t ∈ cont tr1)"
apply safe
apply (metis (no_types) inFr_subtr r reg_subtr_path)
by (metis f inFr.Base path_subtr subtr_inFr subtr_mono subtr_rootL_in)
subsection‹The Regular Cut of a Tree›
context fixes tr0 :: dtree
begin
definition "pick n ≡ SOME tr. subtr UNIV tr tr0 ∧ root tr = n"
lemma pick:
assumes "inItr UNIV tr0 n"
shows "subtr UNIV (pick n) tr0 ∧ root (pick n) = n"
proof-
have "∃ tr. subtr UNIV tr tr0 ∧ root tr = n"
using assms by (metis (lifting) inItr_subtr)
thus ?thesis unfolding pick_def by(rule someI_ex)
qed
lemmas subtr_pick = pick[THEN conjunct1]
lemmas root_pick = pick[THEN conjunct2]
lemma wf_pick:
assumes tr0: "wf tr0" and n: "inItr UNIV tr0 n"
shows "wf (pick n)"
using wf_subtr[OF tr0 subtr_pick[OF n]] .
definition "H_r n ≡ root (pick n)"
definition "H_c n ≡ (id ⊕ root) ` cont (pick n)"
definition H :: "N ⇒ dtree" where
"H ≡ unfold H_r H_c"
lemma finite_H_c: "finite (H_c n)"
unfolding H_c_def by (metis finite_cont finite_imageI)
lemma root_H_pick: "root (H n) = root (pick n)"
using unfold(1)[of H_r H_c n] unfolding H_def H_r_def by simp
lemma root_H[simp]:
assumes "inItr UNIV tr0 n"
shows "root (H n) = n"
unfolding root_H_pick root_pick[OF assms] ..
lemma cont_H[simp]:
"cont (H n) = (id ⊕ (H o root)) ` cont (pick n)"
apply(subst id_comp[symmetric, of id]) unfolding map_sum.comp[symmetric]
unfolding image_comp [symmetric] H_c_def [symmetric]
using unfold(2) [of H_c n H_r, OF finite_H_c]
unfolding H_def ..
lemma Inl_cont_H[simp]:
"Inl -` (cont (H n)) = Inl -` (cont (pick n))"
unfolding cont_H by simp
lemma Inr_cont_H:
"Inr -` (cont (H n)) = (H ∘ root) ` (Inr -` cont (pick n))"
unfolding cont_H by simp
lemma subtr_H:
assumes n: "inItr UNIV tr0 n" and "subtr UNIV tr1 (H n)"
shows "∃ n1. inItr UNIV tr0 n1 ∧ tr1 = H n1"
proof-
{fix tr ns assume "subtr UNIV tr1 tr"
hence "tr = H n ⟶ (∃ n1. inItr UNIV tr0 n1 ∧ tr1 = H n1)"
proof (induct rule: subtr_UNIV_inductL)
case (Step tr2 tr1 tr)
show ?case proof
assume "tr = H n"
then obtain n1 where tr2: "Inr tr2 ∈ cont tr1"
and tr1_tr: "subtr UNIV tr1 tr" and n1: "inItr UNIV tr0 n1" and tr1: "tr1 = H n1"
using Step by auto
obtain tr2' where tr2: "tr2 = H (root tr2')"
and tr2': "Inr tr2' ∈ cont (pick n1)"
using tr2 Inr_cont_H[of n1]
unfolding tr1 image_def comp_def using vimage_eq by auto
have "inItr UNIV tr0 (root tr2')"
using inItr.Base inItr.Ind n1 pick subtr_inItr tr2' by (metis iso_tuple_UNIV_I)
thus "∃n2. inItr UNIV tr0 n2 ∧ tr2 = H n2" using tr2 by blast
qed
qed(insert n, auto)
}
thus ?thesis using assms by auto
qed
lemma root_H_root:
assumes n: "inItr UNIV tr0 n" and t_tr: "t_tr ∈ cont (pick n)"
shows "(id ⊕ (root ∘ H ∘ root)) t_tr = (id ⊕ root) t_tr"
using assms apply(cases t_tr)
apply (metis (lifting) map_sum.simps(1))
using pick H_def H_r_def unfold(1)
inItr.Base comp_apply subtr_StepL subtr_inItr map_sum.simps(2)
by (metis UNIV_I)
lemma H_P:
assumes tr0: "wf tr0" and n: "inItr UNIV tr0 n"
shows "(n, (id ⊕ root) ` cont (H n)) ∈ P" (is "?L ∈ P")
proof-
have "?L = (n, (id ⊕ root) ` cont (pick n))"
unfolding cont_H image_comp map_sum.comp id_comp comp_assoc[symmetric]
unfolding prod.inject apply(rule conjI[OF refl]) apply(rule image_cong[OF refl])
by (rule root_H_root[OF n])
moreover have "... ∈ P" by (metis (lifting) wf_pick root_pick wf_P n tr0)
ultimately show ?thesis by simp
qed
lemma wf_H:
assumes tr0: "wf tr0" and "inItr UNIV tr0 n"
shows "wf (H n)"
proof-
{fix tr have "∃ n. inItr UNIV tr0 n ∧ tr = H n ⟹ wf tr"
proof (induct rule: wf_raw_coind)
case (Hyp tr)
then obtain n where n: "inItr UNIV tr0 n" and tr: "tr = H n" by auto
show ?case apply safe
apply (metis (lifting) H_P root_H n tr tr0)
unfolding tr Inr_cont_H unfolding inj_on_def apply clarsimp using root_H
apply (metis UNIV_I inItr.Base n pick subtr2.simps subtr_inItr subtr_subtr2)
by (metis n subtr.Refl subtr_StepL subtr_H tr UNIV_I)
qed
}
thus ?thesis using assms by blast
qed
definition "rcut ≡ H (root tr0)"
lemma reg_rcut: "reg H rcut"
unfolding reg_def rcut_def
by (metis inItr.Base root_H subtr_H UNIV_I)
lemma rcut_reg:
assumes "reg H tr0"
shows "rcut = tr0"
using assms unfolding rcut_def reg_def by (metis subtr.Refl UNIV_I)
lemma rcut_eq: "rcut = tr0 ⟷ reg H tr0"
using reg_rcut rcut_reg by metis
lemma regular_rcut: "regular rcut"
using reg_rcut unfolding regular_def by blast
lemma Fr_rcut: "Fr UNIV rcut ⊆ Fr UNIV tr0"
proof safe
fix t assume "t ∈ Fr UNIV rcut"
then obtain tr where t: "Inl t ∈ cont tr" and tr: "subtr UNIV tr (H (root tr0))"
using Fr_subtr[of UNIV "H (root tr0)"] unfolding rcut_def
by (metis (full_types) Fr_def inFr_subtr mem_Collect_eq)
obtain n where n: "inItr UNIV tr0 n" and tr: "tr = H n" using tr
by (metis (lifting) inItr.Base subtr_H UNIV_I)
have "Inl t ∈ cont (pick n)" using t using Inl_cont_H[of n] unfolding tr
by (metis (lifting) vimageD vimageI2)
moreover have "subtr UNIV (pick n) tr0" using subtr_pick[OF n] ..
ultimately show "t ∈ Fr UNIV tr0" unfolding Fr_subtr_cont by auto
qed
lemma wf_rcut:
assumes "wf tr0"
shows "wf rcut"
unfolding rcut_def using wf_H[OF assms inItr.Base] by simp
lemma root_rcut[simp]: "root rcut = root tr0"
unfolding rcut_def
by (metis (lifting) root_H inItr.Base reg_def reg_root subtr_rootR_in)
end
subsection‹Recursive Description of the Regular Tree Frontiers›
lemma regular_inFr:
assumes r: "regular tr" and In: "root tr ∈ ns"
and t: "inFr ns tr t"
shows "t ∈ Inl -` (cont tr) ∨
(∃ tr'. Inr tr' ∈ cont tr ∧ inFr (ns - {root tr}) tr' t)"
(is "?L ∨ ?R")
proof-
obtain f where r: "reg f tr" and f: "⋀n. root (f n) = n"
using r unfolding regular_def2 by auto
obtain nl tr1 where d_nl: "distinct nl" and p: "path f nl" and hd_nl: "f (hd nl) = tr"
and l_nl: "f (last nl) = tr1" and s_nl: "set nl ⊆ ns" and t_tr1: "Inl t ∈ cont tr1"
using t unfolding inFr_iff_path[OF r f] by auto
obtain n nl1 where nl: "nl = n # nl1" by (metis (lifting) p path.simps)
hence f_n: "f n = tr" using hd_nl by simp
have n_nl1: "n ∉ set nl1" using d_nl unfolding nl by auto
show ?thesis
proof(cases nl1)
case Nil hence "tr = tr1" using f_n l_nl unfolding nl by simp
hence ?L using t_tr1 by simp thus ?thesis by simp
next
case (Cons n1 nl2) note nl1 = Cons
have 1: "last nl1 = last nl" "hd nl1 = n1" unfolding nl nl1 by simp_all
have p1: "path f nl1" and n1_tr: "Inr (f n1) ∈ cont tr"
using path.simps[of f nl] p f_n unfolding nl nl1 by auto
have r1: "reg f (f n1)" using reg_Inr_cont[OF r n1_tr] .
have 0: "inFr (set nl1) (f n1) t" unfolding inFr_iff_path[OF r1 f]
apply(intro exI[of _ nl1], intro exI[of _ tr1])
using d_nl unfolding 1 l_nl unfolding nl using p1 t_tr1 by auto
have root_tr: "root tr = n" by (metis f f_n)
have "inFr (ns - {root tr}) (f n1) t" apply(rule inFr_mono[OF 0])
using s_nl unfolding root_tr unfolding nl using n_nl1 by auto
thus ?thesis using n1_tr by auto
qed
qed
lemma regular_Fr:
assumes r: "regular tr" and In: "root tr ∈ ns"
shows "Fr ns tr =
Inl -` (cont tr) ∪
⋃{Fr (ns - {root tr}) tr' | tr'. Inr tr' ∈ cont tr}"
unfolding Fr_def
using In inFr.Base regular_inFr[OF assms] apply safe
apply (simp, metis (full_types) mem_Collect_eq)
apply simp
by (simp, metis (lifting) inFr_Ind_minus insert_Diff)
subsection‹The Generated Languages›
definition "L ns n ≡ {Fr ns tr | tr. wf tr ∧ root tr = n}"
definition "Lr ns n ≡ {Fr ns tr | tr. wf tr ∧ root tr = n ∧ regular tr}"
lemma L_rec_notin:
assumes "n ∉ ns"
shows "L ns n = {{}}"
using assms unfolding L_def apply safe
using not_root_Fr apply force
apply(rule exI[of _ "deftr n"])
by (metis (no_types) wf_deftr not_root_Fr root_deftr)
lemma Lr_rec_notin:
assumes "n ∉ ns"
shows "Lr ns n = {{}}"
using assms unfolding Lr_def apply safe
using not_root_Fr apply force
apply(rule exI[of _ "deftr n"])
by (metis (no_types) regular_def wf_deftr not_root_Fr reg_deftr root_deftr)
lemma wf_subtrOf:
assumes "wf tr" and "Inr n ∈ prodOf tr"
shows "wf (subtrOf tr n)"
by (metis assms wf_cont subtrOf)
lemma Lr_rec_in:
assumes n: "n ∈ ns"
shows "Lr ns n ⊆
{Inl -` tns ∪ (⋃{K n' | n'. Inr n' ∈ tns}) | tns K.
(n,tns) ∈ P ∧
(∀ n'. Inr n' ∈ tns ⟶ K n' ∈ Lr (ns - {n}) n')}"
(is "Lr ns n ⊆ {?F tns K | tns K. (n,tns) ∈ P ∧ ?φ tns K}")
proof safe
fix ts assume "ts ∈ Lr ns n"
then obtain tr where dtr: "wf tr" and r: "root tr = n" and tr: "regular tr"
and ts: "ts = Fr ns tr" unfolding Lr_def by auto
define tns where "tns = (id ⊕ root) ` (cont tr)"
define K where "K n' = Fr (ns - {n}) (subtrOf tr n')" for n'
show "∃tns K. ts = ?F tns K ∧ (n, tns) ∈ P ∧ ?φ tns K"
apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI)
show "ts = Inl -` tns ∪ ⋃{K n' |n'. Inr n' ∈ tns}"
unfolding ts regular_Fr[OF tr n[unfolded r[symmetric]]]
unfolding tns_def K_def r[symmetric]
unfolding Inl_prodOf wf_subtrOf_Union[OF dtr] ..
show "(n, tns) ∈ P" unfolding tns_def r[symmetric] using wf_P[OF dtr] .
fix n' assume "Inr n' ∈ tns" thus "K n' ∈ Lr (ns - {n}) n'"
unfolding K_def Lr_def mem_Collect_eq apply(intro exI[of _ "subtrOf tr n'"])
using dtr tr apply(intro conjI refl) unfolding tns_def
apply(erule wf_subtrOf[OF dtr])
apply (metis subtrOf)
by (metis Inr_subtrOf UNIV_I regular_subtr subtr.simps)
qed
qed
lemma hsubst_aux:
fixes n ftr tns
assumes n: "n ∈ ns" and tns: "finite tns" and
1: "⋀ n'. Inr n' ∈ tns ⟹ wf (ftr n')"
defines "tr ≡ Node n ((id ⊕ ftr) ` tns)" defines "tr' ≡ hsubst tr tr"
shows "Fr ns tr' = Inl -` tns ∪ ⋃{Fr (ns - {n}) (ftr n') |n'. Inr n' ∈ tns}"
(is "_ = ?B") proof-
have rtr: "root tr = n" and ctr: "cont tr = (id ⊕ ftr) ` tns"
unfolding tr_def using tns by auto
have Frr: "Frr (ns - {n}) tr = ⋃{Fr (ns - {n}) (ftr n') |n'. Inr n' ∈ tns}"
unfolding Frr_def ctr by auto
have "Fr ns tr' = Inl -` (cont tr) ∪ Frr (ns - {n}) tr"
using Fr_self_hsubst[OF n[unfolded rtr[symmetric]]] unfolding tr'_def rtr ..
also have "... = ?B" unfolding ctr Frr by simp
finally show ?thesis .
qed
lemma L_rec_in:
assumes n: "n ∈ ns"
shows "
{Inl -` tns ∪ (⋃{K n' | n'. Inr n' ∈ tns}) | tns K.
(n,tns) ∈ P ∧
(∀ n'. Inr n' ∈ tns ⟶ K n' ∈ L (ns - {n}) n')}
⊆ L ns n"
proof safe
fix tns K
assume P: "(n, tns) ∈ P" and 0: "∀n'. Inr n' ∈ tns ⟶ K n' ∈ L (ns - {n}) n'"
{fix n' assume "Inr n' ∈ tns"
hence "K n' ∈ L (ns - {n}) n'" using 0 by auto
hence "∃ tr'. K n' = Fr (ns - {n}) tr' ∧ wf tr' ∧ root tr' = n'"
unfolding L_def mem_Collect_eq by auto
}
then obtain ftr where 0: "⋀ n'. Inr n' ∈ tns ⟹
K n' = Fr (ns - {n}) (ftr n') ∧ wf (ftr n') ∧ root (ftr n') = n'"
by metis
define tr where "tr = Node n ((id ⊕ ftr) ` tns)"
define tr' where "tr' = hsubst tr tr"
have rtr: "root tr = n" and ctr: "cont tr = (id ⊕ ftr) ` tns"
unfolding tr_def by (simp, metis P cont_Node finite_imageI finite_in_P)
have prtr: "prodOf tr = tns" apply(rule Inl_Inr_image_cong)
unfolding ctr apply simp apply simp apply safe
using 0 unfolding image_def apply force apply simp by (metis 0 vimageI2)
have 1: "{K n' |n'. Inr n' ∈ tns} = {Fr (ns - {n}) (ftr n') |n'. Inr n' ∈ tns}"
using 0 by auto
have dtr: "wf tr" apply(rule wf.dtree)
apply (metis (lifting) P prtr rtr)
unfolding inj_on_def ctr using 0 by auto
hence dtr': "wf tr'" unfolding tr'_def by (metis wf_hsubst)
have tns: "finite tns" using finite_in_P P by simp
have "Inl -` tns ∪ ⋃{Fr (ns - {n}) (ftr n') |n'. Inr n' ∈ tns} ∈ L ns n"
unfolding L_def mem_Collect_eq apply(intro exI[of _ tr'] conjI)
using dtr' 0 hsubst_aux[OF assms tns, of ftr] unfolding tr_def tr'_def by auto
thus "Inl -` tns ∪ ⋃{K n' |n'. Inr n' ∈ tns} ∈ L ns n" unfolding 1 .
qed
lemma card_N: "(n::N) ∈ ns ⟹ card (ns - {n}) < card ns"
by (metis finite_N Diff_UNIV Diff_infinite_finite card_Diff1_less finite.emptyI)
function LL where
"LL ns n =
(if n ∉ ns then {{}} else
{Inl -` tns ∪ (⋃{K n' | n'. Inr n' ∈ tns}) | tns K.
(n,tns) ∈ P ∧
(∀ n'. Inr n' ∈ tns ⟶ K n' ∈ LL (ns - {n}) n')})"
by(pat_completeness, auto)
termination apply(relation "inv_image (measure card) fst")
using card_N by auto
declare LL.simps[code]
declare LL.simps[simp del]
lemma Lr_LL: "Lr ns n ⊆ LL ns n"
proof (induct ns arbitrary: n rule: measure_induct[of card])
case (1 ns n) show ?case proof(cases "n ∈ ns")
case False thus ?thesis unfolding Lr_rec_notin[OF False] by (simp add: LL.simps)
next
case True show ?thesis apply(rule subset_trans)
using Lr_rec_in[OF True] apply assumption
unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
fix tns K
assume "n ∈ ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
assume "(n, tns) ∈ P"
and "∀n'. Inr n' ∈ tns ⟶ K n' ∈ Lr (ns - {n}) n'"
thus "∃tnsa Ka.
Inl -` tns ∪ ⋃{K n' |n'. Inr n' ∈ tns} =
Inl -` tnsa ∪ ⋃{Ka n' |n'. Inr n' ∈ tnsa} ∧
(n, tnsa) ∈ P ∧ (∀n'. Inr n' ∈ tnsa ⟶ Ka n' ∈ LL (ns - {n}) n')"
apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
qed
qed
qed
lemma LL_L: "LL ns n ⊆ L ns n"
proof (induct ns arbitrary: n rule: measure_induct[of card])
case (1 ns n) show ?case proof(cases "n ∈ ns")
case False thus ?thesis unfolding L_rec_notin[OF False] by (simp add: LL.simps)
next
case True show ?thesis apply(rule subset_trans)
prefer 2 using L_rec_in[OF True] apply assumption
unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
fix tns K
assume "n ∈ ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
assume "(n, tns) ∈ P"
and "∀n'. Inr n' ∈ tns ⟶ K n' ∈ LL (ns - {n}) n'"
thus "∃tnsa Ka.
Inl -` tns ∪ ⋃{K n' |n'. Inr n' ∈ tns} =
Inl -` tnsa ∪ ⋃{Ka n' |n'. Inr n' ∈ tnsa} ∧
(n, tnsa) ∈ P ∧ (∀n'. Inr n' ∈ tnsa ⟶ Ka n' ∈ L (ns - {n}) n')"
apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
qed
qed
qed
definition "subs L1 L2 ≡ ∀ ts2 ∈ L2. ∃ ts1 ∈ L1. ts1 ⊆ ts2"
lemma incl_subs[simp]: "L2 ⊆ L1 ⟹ subs L1 L2"
unfolding subs_def by auto
lemma subs_refl[simp]: "subs L1 L1" unfolding subs_def by auto
lemma subs_trans: "⟦subs L1 L2; subs L2 L3⟧ ⟹ subs L1 L3"
unfolding subs_def by (metis subset_trans)
definition "leqv L1 L2 ≡ subs L1 L2 ∧ subs L2 L1"
lemma subs_leqv[simp]: "leqv L1 L2 ⟹ subs L1 L2"
unfolding leqv_def by auto
lemma subs_leqv_sym[simp]: "leqv L1 L2 ⟹ subs L2 L1"
unfolding leqv_def by auto
lemma leqv_refl[simp]: "leqv L1 L1" unfolding leqv_def by auto
lemma leqv_trans:
assumes 12: "leqv L1 L2" and 23: "leqv L2 L3"
shows "leqv L1 L3"
using assms unfolding leqv_def by (metis (lifting) subs_trans)
lemma leqv_sym: "leqv L1 L2 ⟹ leqv L2 L1"
unfolding leqv_def by auto
lemma leqv_Sym: "leqv L1 L2 ⟷ leqv L2 L1"
unfolding leqv_def by auto
lemma Lr_incl_L: "Lr ns ts ⊆ L ns ts"
unfolding Lr_def L_def by auto
lemma Lr_subs_L: "subs (Lr UNIV ts) (L UNIV ts)"
unfolding subs_def proof safe
fix ts2 assume "ts2 ∈ L UNIV ts"
then obtain tr where ts2: "ts2 = Fr UNIV tr" and dtr: "wf tr" and rtr: "root tr = ts"
unfolding L_def by auto
thus "∃ts1∈Lr UNIV ts. ts1 ⊆ ts2"
apply(intro bexI[of _ "Fr UNIV (rcut tr)"])
unfolding Lr_def L_def using Fr_rcut wf_rcut root_rcut regular_rcut by auto
qed
lemma Lr_leqv_L: "leqv (Lr UNIV ts) (L UNIV ts)"
using Lr_subs_L unfolding leqv_def by (metis (lifting) Lr_incl_L incl_subs)
lemma LL_leqv_L: "leqv (LL UNIV ts) (L UNIV ts)"
by (metis (lifting) LL_L Lr_LL Lr_subs_L incl_subs leqv_def subs_trans)
lemma LL_leqv_Lr: "leqv (LL UNIV ts) (Lr UNIV ts)"
using Lr_leqv_L LL_leqv_L by (metis leqv_Sym leqv_trans)
end