Theory WellformedL
theory WellformedL
imports Wellformed "SyntaxL"
begin
chapter ‹Wellformedness Lemmas›
section ‹Prelude›
lemma b_of_subst_bb_commute:
"(b_of (τ[bv::=b]⇩τ⇩b)) = (b_of τ)[bv::=b]⇩b⇩b"
proof -
obtain z' and b' and c' where "τ = ⦃ z' : b' | c' ⦄ " using obtain_fresh_z by metis
moreover hence "(b_of (τ[bv::=b]⇩τ⇩b)) = b_of ⦃ z' : b'[bv::=b]⇩b⇩b | c' ⦄" using subst_tb.simps by simp
ultimately show ?thesis using subst_tv.simps subst_tb.simps by simp
qed
lemmas wf_intros = wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.intros wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.intros
lemmas freshers = fresh_prodN b.fresh c.fresh v.fresh ce.fresh fresh_GCons fresh_GNil fresh_at_base
section ‹Strong Elimination›
text ‹Inversion/elimination for well-formed polymorphic constructors ›
lemma wf_strong_elim:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list"
and Δ::Δ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s and tm::"'a::fs"
and cs::branch_s and css::branch_list and Θ::Θ
shows "Θ; ℬ; Γ ⊢⇩w⇩f (V_consp tyid dc b v) : b'' ⟹ (∃ bv dclist x b' c. b'' = B_app tyid b ∧
AF_typedef_poly tyid bv dclist ∈ set Θ ∧
(dc, ⦃ x : b' | c ⦄) ∈ set dclist ∧
Θ; ℬ ⊢⇩w⇩f b ∧ atom bv ♯ (Θ, ℬ, Γ, b, v) ∧ Θ; ℬ; Γ ⊢⇩w⇩f v : b'[bv::=b]⇩b⇩b ∧ atom bv ♯ tm)" and
"Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ True" and
"Θ; ℬ ⊢⇩w⇩f Γ ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f Θ ⟹True" and
"Θ; ℬ ⊢⇩w⇩f b ⟹ True " and
"Θ; ℬ; Γ ⊢⇩w⇩f ce : b' ⟹ True" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(nominal_induct
"V_consp tyid dc b v" b'' and c and Γ and τ and ts and Θ and b and b' and td
avoiding: tm
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfV_conspI bv dclist Θ x b' c ℬ Γ)
then show ?case by force
qed(auto+)
section ‹Context Extension›
definition wfExt :: "Θ ⇒ ℬ ⇒ Γ ⇒ Γ ⇒ bool" (" _ ; _ ⊢⇩w⇩f _ < _ " [50,50,50] 50) where
"wfExt T B G1 G2 = (wfG T B G2 ∧ wfG T B G1 ∧ toSet G1 ⊆ toSet G2)"
section ‹Context›
lemma wfG_cons[ms_wb]:
fixes Γ::Γ
assumes "P; ℬ ⊢⇩w⇩f (z,b,c) #⇩ΓΓ"
shows "P; ℬ ⊢⇩w⇩f Γ ∧ atom z ♯ Γ ∧ wfB P ℬ b"
using wfG_elims(2)[OF assms] by metis
lemma wfG_cons2[ms_wb]:
fixes Γ::Γ
assumes "P; ℬ ⊢⇩w⇩f zbc #⇩ΓΓ"
shows "P; ℬ ⊢⇩w⇩f Γ"
proof -
obtain z and b and c where zbc: "zbc=(z,b,c)" using prod_cases3 by blast
hence "P; ℬ ⊢⇩w⇩f (z,b,c) #⇩ΓΓ" using assms by auto
thus ?thesis using zbc wfG_cons assms by simp
qed
lemma wf_g_unique:
fixes Γ::Γ
assumes "Θ; ℬ ⊢⇩w⇩f Γ" and "(x,b,c) ∈ toSet Γ" and "(x,b',c') ∈ toSet Γ"
shows "b=b' ∧ c=c'"
using assms proof(induct Γ rule: Γ.induct)
case GNil
then show ?case by simp
next
case (GCons a Γ)
consider "(x,b,c)=a ∧ (x,b',c')=a" | "(x,b,c)=a ∧ (x,b',c')≠a" | "(x,b,c)≠a ∧ (x,b',c')=a" | "(x,b,c)≠ a ∧ (x,b',c')≠a" by blast
then show ?case proof(cases)
case 1
then show ?thesis by auto
next
case 2
hence "atom x ♯ Γ" using wfG_elims(2) GCons by blast
moreover have "(x,b',c') ∈ toSet Γ" using GCons 2 by force
ultimately show ?thesis using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem Γ.distinct subst_gv.simps 2 GCons by metis
next
case 3
hence "atom x ♯ Γ" using wfG_elims(2) GCons by blast
moreover have "(x,b,c) ∈ toSet Γ" using GCons 3 by force
ultimately show ?thesis
using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem Γ.distinct subst_gv.simps 3 GCons by metis
next
case 4
then obtain x'' and b'' and c''::c where xbc: "a=(x'',b'',c'')"
using prod_cases3 by blast
hence "Θ; ℬ ⊢⇩w⇩f ((x'',b'',c'') #⇩ΓΓ)" using GCons wfG_elims by blast
hence "Θ; ℬ ⊢⇩w⇩f Γ ∧ (x, b, c) ∈ toSet Γ ∧ (x, b', c') ∈ toSet Γ" using GCons wfG_elims 4 xbc
prod_cases3 set_GConsD using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem Γ.distinct subst_gv.simps 4 GCons by meson
thus ?thesis using GCons by auto
qed
qed
lemma lookup_if1:
fixes Γ::Γ
assumes "Θ; ℬ ⊢⇩w⇩f Γ" and "Some (b,c) = lookup Γ x"
shows "(x,b,c) ∈ toSet Γ ∧ (∀b' c'. (x,b',c') ∈ toSet Γ ⟶ b'=b ∧ c'=c)"
using assms proof(induct Γ rule: Γ.induct)
case GNil
then show ?case by auto
next
case (GCons xbc Γ)
then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
using prod_cases3 by blast
then show ?case using wf_g_unique GCons lookup_in_g xbc
lookup.simps set_GConsD wfG.cases
insertE insert_is_Un toSet.simps wfG_elims by metis
qed
lemma lookup_if2:
assumes "wfG P ℬ Γ" and "(x,b,c) ∈ toSet Γ ∧ (∀b' c'. (x,b',c') ∈ toSet Γ ⟶ b'=b ∧ c'=c)"
shows "Some (b,c) = lookup Γ x"
using assms proof(induct Γ rule: Γ.induct)
case GNil
then show ?case by auto
next
case (GCons xbc Γ)
then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
using prod_cases3 by blast
then show ?case proof(cases "x=x'")
case True
then show ?thesis using lookup.simps GCons xbc by simp
next
case False
then show ?thesis using lookup.simps GCons xbc toSet.simps Un_iff set_GConsD wfG_cons2
by (metis (full_types) Un_iff set_GConsD toSet.simps(2) wfG_cons2)
qed
qed
lemma lookup_iff:
fixes Θ::Θ and Γ::Γ
assumes "Θ; ℬ ⊢⇩w⇩f Γ"
shows "Some (b,c) = lookup Γ x ⟷ (x,b,c) ∈ toSet Γ ∧ (∀b' c'. (x,b',c') ∈ toSet Γ ⟶ b'=b ∧ c'=c)"
using assms lookup_if1 lookup_if2 by meson
lemma wfG_lookup_wf:
fixes Θ::Θ and Γ::Γ and b::b and ℬ::ℬ
assumes "Θ; ℬ ⊢⇩w⇩f Γ" and "Some (b,c) = lookup Γ x"
shows "Θ; ℬ ⊢⇩w⇩f b"
using assms proof(induct Γ rule: Γ_induct)
case GNil
then show ?case by auto
next
case (GCons x' b' c' Γ')
then show ?case proof(cases "x=x'")
case True
then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce
next
case False
then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce
qed
qed
lemma wfG_unique:
fixes Γ::Γ
assumes "wfG B Θ ((x, b, c) #⇩Γ Γ)" and "(x1,b1,c1) ∈ toSet ((x, b, c) #⇩Γ Γ)" and "x1=x"
shows "b1 = b ∧ c1 = c"
proof -
have "(x, b, c) ∈ toSet ((x, b, c) #⇩Γ Γ)" by simp
thus ?thesis using wf_g_unique assms by blast
qed
lemma wfG_unique_full:
fixes Γ::Γ
assumes "wfG Θ B (Γ'@(x, b, c) #⇩Γ Γ)" and "(x1,b1,c1) ∈ toSet (Γ'@(x, b, c) #⇩Γ Γ)" and "x1=x"
shows "b1 = b ∧ c1 = c"
proof -
have "(x, b, c) ∈ toSet (Γ'@(x, b, c) #⇩Γ Γ)" by simp
thus ?thesis using wf_g_unique assms by blast
qed
section ‹Converting between wb forms›
text ‹ We cannot prove wfB properties here for expressions and statements as need some more facts about @{term Φ}
context which we can prove without this lemma. Trying to cram everything into a single large
mutually recursive lemma is not a good idea ›
lemma wfX_wfY1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s
and css::branch_list
shows wfV_wf: "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ ⊢⇩w⇩f Θ " and
wfC_wf: "Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ ⊢⇩w⇩f Θ " and
wfG_wf :"Θ; ℬ ⊢⇩w⇩f Γ ⟹ ⊢⇩w⇩f Θ" and
wfT_wf: "Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ ⊢⇩w⇩f Θ ∧ Θ; ℬ ⊢⇩w⇩f b_of τ" and
wfTs_wf:"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ ⊢⇩w⇩f Θ" and
"⊢⇩w⇩f Θ ⟹ True" and
wfB_wf: "Θ; ℬ ⊢⇩w⇩f b ⟹ ⊢⇩w⇩f Θ" and
wfCE_wf: "Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ ⊢⇩w⇩f Θ " and
wfTD_wf: "Θ ⊢⇩w⇩f td ⟹ ⊢⇩w⇩f Θ"
proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
case (wfV_varI Θ ℬ Γ b c x)
hence "(x,b,c) ∈ toSet Γ" using lookup_iff lookup_in_g by presburger
hence "b ∈ fst`snd`toSet Γ" by force
hence "wfB Θ ℬ b" using wfV_varI using wfG_lookup_wf by auto
then show ?case using wfV_varI wfV_elims wf_intros by metis
next
case (wfV_litI Θ ℬ Γ l)
moreover have "wfTh Θ" using wfV_litI by metis
ultimately show ?case using wf_intros base_for_lit.simps l.exhaust by metis
next
case (wfV_pairI Θ ℬ Γ v1 b1 v2 b2)
then show ?case using wfB_pairI by simp
next
case (wfV_consI s dclist Θ dc x b c ℬ Γ v)
then show ?case using wf_intros by metis
next
case (wfTI z Γ Θ ℬ b c)
then show ?case using wf_intros b_of.simps wfG_cons2 by metis
qed(auto)
lemma wfX_wfY2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s
and css::branch_list
shows
wfE_wf: "Θ; Φ; ℬ; Γ; Δ ⊢⇩w⇩f e : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ Θ; ℬ; Γ ⊢⇩w⇩f Δ ∧ ⊢⇩w⇩f Θ ∧ Θ ⊢⇩w⇩f Φ " and
wfS_wf: "Θ; Φ; ℬ; Γ; Δ ⊢⇩w⇩f s : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ Θ; ℬ; Γ ⊢⇩w⇩f Δ ∧ ⊢⇩w⇩f Θ ∧ Θ ⊢⇩w⇩f Φ " and
"Θ; Φ; ℬ; Γ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ Θ; ℬ; Γ ⊢⇩w⇩f Δ ∧ ⊢⇩w⇩f Θ ∧ Θ ⊢⇩w⇩f Φ " and
"Θ; Φ; ℬ; Γ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ Θ; ℬ; Γ ⊢⇩w⇩f Δ ∧ ⊢⇩w⇩f Θ ∧ Θ ⊢⇩w⇩f Φ " and
wfPhi_wf: "Θ ⊢⇩w⇩f (Φ::Φ) ⟹ ⊢⇩w⇩f Θ" and
wfD_wf: "Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ Θ; ℬ ⊢⇩w⇩f Γ ∧ ⊢⇩w⇩f Θ " and
wfFTQ_wf: "Θ ; Φ ⊢⇩w⇩f ftq ⟹ Θ ⊢⇩w⇩f Φ ∧ ⊢⇩w⇩f Θ" and
wfFT_wf: "Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ Θ ⊢⇩w⇩f Φ ∧ ⊢⇩w⇩f Θ"
proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
case (wfS_varI Θ ℬ Γ τ v u Δ Φ s b)
then show ?case using wfD_elims by auto
next
case (wfS_assignI u τ Δ Θ ℬ Γ Φ v)
then show ?case using wf_intros by metis
next
case (wfD_emptyI Θ ℬ Γ)
then show ?case using wfX_wfY1 by auto
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
then have "Θ; ℬ ⊢⇩w⇩f Γ" using wfX_wfY1 by auto
moreover have "Θ; ℬ; Γ ⊢⇩w⇩f Δ" using wfS_assertI by auto
moreover have "⊢⇩w⇩f Θ ∧ Θ ⊢⇩w⇩f Φ " using wfS_assertI by auto
ultimately show ?case by auto
qed(auto)
lemmas wfX_wfY=wfX_wfY1 wfX_wfY2
lemma setD_ConsD:
"ut ∈ setD (ut' #⇩Δ D) = (ut = ut' ∨ ut ∈ setD D)"
proof(induct D rule: Δ_induct)
case DNil
then show ?case by auto
next
case (DCons u' t' x2)
then show ?case using setD.simps by auto
qed
lemma wfD_wfT:
fixes Δ::Δ and τ::τ
assumes "Θ; ℬ; Γ ⊢⇩w⇩f Δ"
shows "∀(u,τ) ∈ setD Δ. Θ; ℬ; Γ ⊢⇩w⇩f τ"
using assms proof(induct Δ rule: Δ_induct)
case DNil
then show ?case by auto
next
case (DCons u' t' x2)
then show ?case using wfD_elims DCons setD_ConsD
by (metis case_prodI2 set_ConsD)
qed
lemma subst_b_lookup_d:
assumes "u ∉ fst ` setD Δ"
shows "u ∉ fst ` setD Δ[bv::=b]⇩Δ⇩b"
using assms proof(induct Δ rule: Δ_induct)
case DNil
then show ?case by auto
next
case (DCons u' t' x2)
hence "u≠u'" using DCons by simp
show ?case using DCons subst_db.simps by simp
qed
lemma wfG_cons_splitI:
fixes Φ::Φ and Γ::Γ
assumes "Θ; ℬ ⊢⇩w⇩f Γ" and "atom x ♯ Γ" and "wfB Θ ℬ b" and
"c ∈ { TRUE, FALSE } ⟶ Θ; ℬ ⊢⇩w⇩f Γ " and
"c ∉ { TRUE, FALSE } ⟶ Θ ;ℬ ; (x,b,C_true) #⇩ΓΓ ⊢⇩w⇩f c"
shows "Θ; ℬ ⊢⇩w⇩f ((x,b,c) #⇩ΓΓ)"
using wfG_cons1I wfG_cons2I assms by metis
lemma wfG_consI:
fixes Φ::Φ and Γ::Γ and c::c
assumes "Θ; ℬ ⊢⇩w⇩f Γ" and "atom x ♯ Γ" and "wfB Θ ℬ b" and
"Θ ; ℬ ; (x,b,C_true) #⇩ΓΓ ⊢⇩w⇩f c"
shows "Θ ; ℬ ⊢⇩w⇩f ((x,b,c) #⇩ΓΓ)"
using wfG_cons1I wfG_cons2I wfG_cons_splitI wfC_trueI assms by metis
lemma wfG_elim2:
fixes c::c
assumes "wfG P ℬ ((x,b,c) #⇩ΓΓ)"
shows "P; ℬ ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c ∧ wfB P ℬ b"
proof(cases "c ∈ {TRUE,FALSE}")
case True
have "P; ℬ ⊢⇩w⇩f Γ ∧ atom x ♯ Γ ∧ wfB P ℬ b" using wfG_elims(2)[OF assms] by auto
hence "P; ℬ ⊢⇩w⇩f ((x,b,TRUE) #⇩ΓΓ) ∧ wfB P ℬ b" using wfG_cons2I by auto
thus ?thesis using wfC_trueI wfC_falseI True by auto
next
case False
then show ?thesis using wfG_elims(2)[OF assms] by auto
qed
lemma wfG_cons_wfC:
fixes Γ::Γ and c::c
assumes "Θ ; B ⊢⇩w⇩f (x, b, c) #⇩Γ Γ"
shows "Θ ; B ; ((x, b, TRUE) #⇩Γ Γ) ⊢⇩w⇩f c"
using assms wfG_elim2 by auto
lemma wfG_wfB:
assumes "wfG P ℬ Γ" and "b ∈ fst`snd`toSet Γ"
shows "wfB P ℬ b"
using assms proof(induct Γ rule:Γ_induct)
case GNil
then show ?case by auto
next
case (GCons x' b' c' Γ')
show ?case proof(cases "b=b'")
case True
then show ?thesis using wfG_elim2 GCons by auto
next
case False
hence "b ∈ fst`snd`toSet Γ'" using GCons by auto
moreover have "wfG P ℬ Γ'" using wfG_cons GCons by auto
ultimately show ?thesis using GCons by auto
qed
qed
lemma wfG_cons_TRUE:
fixes Γ::Γ and b::b
assumes "P; ℬ ⊢⇩w⇩f Γ" and "atom z ♯ Γ" and "P; ℬ ⊢⇩w⇩f b"
shows "P ; ℬ ⊢⇩w⇩f (z, b, TRUE) #⇩Γ Γ"
using wfG_cons2I wfG_wfB assms by simp
lemma wfG_cons_TRUE2:
assumes "P; ℬ ⊢⇩w⇩f (z,b,c) #⇩ΓΓ" and "atom z ♯ Γ"
shows "P; ℬ ⊢⇩w⇩f (z, b, TRUE) #⇩Γ Γ"
using wfG_cons wfG_cons2I assms by simp
lemma wfG_suffix:
fixes Γ::Γ
assumes "wfG P ℬ (Γ'@Γ)"
shows "wfG P ℬ Γ"
using assms proof(induct Γ' rule: Γ_induct)
case GNil
then show ?case by auto
next
case (GCons x b c Γ')
hence " P; ℬ ⊢⇩w⇩f Γ' @ Γ" using wfG_elims by auto
then show ?case using GCons wfG_elims by auto
qed
lemma wfV_wfCE:
fixes v::v
assumes "Θ; ℬ; Γ ⊢⇩w⇩f v : b"
shows " Θ ; ℬ ; Γ ⊢⇩w⇩f CE_val v : b"
proof -
have "Θ ⊢⇩w⇩f ([]::Φ) " using wfPhi_emptyI wfV_wf wfG_wf assms by metis
moreover have "Θ; ℬ; Γ ⊢⇩w⇩f []⇩Δ" using wfD_emptyI wfV_wf wfG_wf assms by metis
ultimately show ?thesis using wfCE_valI assms by auto
qed
section ‹Support›
lemma wf_supp1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list
shows wfV_supp: "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ supp v ⊆ atom_dom Γ ∪ supp ℬ" and
wfC_supp: "Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ supp c ⊆ atom_dom Γ ∪ supp ℬ" and
wfG_supp: "Θ; ℬ ⊢⇩w⇩f Γ ⟹ atom_dom Γ ⊆ supp Γ" and
wfT_supp: "Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ supp τ ⊆ atom_dom Γ ∪ supp ℬ " and
wfTs_supp: "Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ supp ts ⊆ atom_dom Γ ∪ supp ℬ" and
wfTh_supp: "⊢⇩w⇩f Θ ⟹ supp Θ = {}" and
wfB_supp: "Θ; ℬ ⊢⇩w⇩f b ⟹ supp b ⊆ supp ℬ" and
wfCE_supp: "Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ supp ce ⊆ atom_dom Γ ∪ supp ℬ" and
wfTD_supp: "Θ ⊢⇩w⇩f td ⟹ supp td ⊆ {}"
proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
case (wfB_consI Θ s dclist ℬ)
then show ?case by(auto simp add: b.supp pure_supp)
next
case (wfB_appI Θ ℬ b s bv dclist)
then show ?case by(auto simp add: b.supp pure_supp)
next
case (wfV_varI Θ ℬ Γ b c x)
then show ?case using v.supp wfV_elims
empty_subsetI insert_subset supp_at_base
fresh_dom_free2 lookup_if1
by (metis sup.coboundedI1)
next
case (wfV_litI Θ ℬ Γ l)
then show ?case using supp_l_empty v.supp by simp
next
case (wfV_pairI Θ ℬ Γ v1 b1 v2 b2)
then show ?case using v.supp wfV_elims by (metis Un_subset_iff)
next
case (wfV_consI s dclist Θ dc x b c ℬ Γ v)
then show ?case using v.supp wfV_elims
Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp by metis
next
case (wfV_conspI typid bv dclist Θ dc x b' c ℬ Γ v b)
then show ?case unfolding v.supp
using wfV_elims
Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp
by (simp add: Un_commute pure_supp sup.coboundedI1)
next
case (wfC_eqI Θ ℬ Γ e1 b e2)
hence "supp e1 ⊆ atom_dom Γ ∪ supp ℬ" using c.supp wfC_elims
image_empty list.set(1) sup_bot.right_neutral by (metis IntI UnE empty_iff subsetCE subsetI)
moreover have "supp e2 ⊆ atom_dom Γ ∪ supp ℬ" using c.supp wfC_elims
image_empty list.set(1) sup_bot.right_neutral IntI UnE empty_iff subsetCE subsetI
by (metis wfC_eqI.hyps(4))
ultimately show ?case using c.supp by auto
next
case (wfG_cons1I c Θ ℬ Γ x b)
then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis
next
case (wfG_cons2I c Θ ℬ Γ x b)
then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis
next
case wfTh_emptyI
then show ?case by (simp add: supp_Nil)
next
case (wfTh_consI Θ lst)
then show ?case using supp_Cons by fast
next
case (wfTD_simpleI Θ lst s)
then have "supp (AF_typedef s lst ) = supp lst ∪ supp s" using type_def.supp by auto
then show ?case using wfTD_simpleI pure_supp
by (simp add: pure_supp supp_Cons supp_at_base)
next
case (wfTD_poly Θ bv lst s)
then have "supp (AF_typedef_poly s bv lst ) = supp lst - { atom bv } ∪ supp s" using type_def.supp by auto
then show ?case using wfTD_poly pure_supp
by (simp add: pure_supp supp_Cons supp_at_base)
next
case (wfTs_nil Θ ℬ Γ)
then show ?case using supp_Nil by auto
next
case (wfTs_cons Θ ℬ Γ τ dc ts)
then show ?case using supp_Cons supp_Pair pure_supp[of dc] by blast
next
case (wfCE_valI Θ ℬ Γ v b)
thus ?case using ce.supp wfCE_elims by simp
next
case (wfCE_plusI Θ ℬ Γ v1 v2)
hence "supp (CE_op Plus v1 v2) ⊆ atom_dom Γ ∪ supp ℬ" using ce.supp pure_supp
by (simp add: wfCE_plusI opp.supp)
then show ?case using ce.supp wfCE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
next
case (wfCE_leqI Θ ℬ Γ v1 v2)
hence "supp (CE_op LEq v1 v2) ⊆ atom_dom Γ ∪ supp ℬ" using ce.supp pure_supp
by (simp add: wfCE_plusI opp.supp)
then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
next
case (wfCE_eqI Θ ℬ Γ v1 b v2 )
hence "supp (CE_op Eq v1 v2) ⊆ atom_dom Γ ∪ supp ℬ" using ce.supp pure_supp
by (simp add: wfCE_eqI opp.supp)
then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
next
case (wfCE_fstI Θ ℬ Γ v1 b1 b2)
thus ?case using ce.supp wfCE_elims by simp
next
case (wfCE_sndI Θ ℬ Γ v1 b1 b2)
thus ?case using ce.supp wfCE_elims by simp
next
case (wfCE_concatI Θ ℬ Γ v1 v2)
thus ?case using ce.supp wfCE_elims by simp
next
case (wfCE_lenI Θ ℬ Γ v1)
thus ?case using ce.supp wfCE_elims by simp
next
case (wfTI z Θ ℬ Γ b c)
hence "supp c ⊆ supp z ∪ atom_dom Γ ∪ supp ℬ" using supp_at_base dom_cons by metis
moreover have "supp b ⊆ supp ℬ" using wfTI by auto
ultimately have " supp ⦃ z : b | c ⦄ ⊆ atom_dom Γ ∪ supp ℬ" using τ.supp supp_at_base by force
thus ?case by auto
qed(auto)
lemma wf_supp2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and
ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ftq::fun_typ_q and
ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list
shows
wfE_supp: "Θ; Φ; ℬ; Γ; Δ ⊢⇩w⇩f e : b ⟹ (supp e ⊆ atom_dom Γ ∪ supp ℬ ∪ atom ` fst ` setD Δ)" and
wfS_supp: "Θ; Φ; ℬ; Γ; Δ ⊢⇩w⇩f s : b ⟹ supp s ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" and
"Θ; Φ; ℬ; Γ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ supp cs ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" and
"Θ; Φ; ℬ; Γ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ supp css ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" and
wfPhi_supp: "Θ ⊢⇩w⇩f (Φ::Φ) ⟹ supp Φ = {}" and
wfD_supp: "Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ supp Δ ⊆ atom`fst`(setD Δ) ∪ atom_dom Γ ∪ supp ℬ " and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ supp ftq = {}" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ supp ft ⊆ supp ℬ"
proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
case (wfE_valI Θ Φ ℬ Γ Δ v b)
hence "supp (AE_val v) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp wf_supp1 by simp
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_plusI Θ Φ ℬ Γ Δ v1 v2)
hence "supp (AE_op Plus v1 v2) ⊆ atom_dom Γ ∪ supp ℬ"
using wfE_plusI opp.supp wf_supp1 e.supp pure_supp Un_least
by (metis sup_bot.left_neutral)
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_leqI Θ Φ ℬ Γ Δ v1 v2)
hence "supp (AE_op LEq v1 v2) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp Un_least
sup_bot.left_neutral using opp.supp wf_supp1 by auto
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_eqI Θ Φ ℬ Γ Δ v1 b v2)
hence "supp (AE_op Eq v1 v2) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp Un_least
sup_bot.left_neutral using opp.supp wf_supp1 by auto
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_fstI Θ Φ ℬ Γ Δ v1 b1 b2)
hence "supp (AE_fst v1 ) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_sndI Θ Φ ℬ Γ Δ v1 b1 b2)
hence "supp (AE_snd v1 ) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp wfE_plusI opp.supp wf_supp1 by (metis Un_least)
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_concatI Θ Φ ℬ Γ Δ v1 v2)
hence "supp (AE_concat v1 v2) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp
wfE_plusI opp.supp wf_supp1 by (metis Un_least)
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_splitI Θ Φ ℬ Γ Δ v1 v2)
hence "supp (AE_split v1 v2) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp
wfE_plusI opp.supp wf_supp1 by (metis Un_least)
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_lenI Θ Φ ℬ Γ Δ v1)
hence "supp (AE_len v1 ) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp
using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto
then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
next
case (wfE_appI Θ Φ ℬ Γ Δ f x b c τ s v)
then obtain b where "Θ; ℬ; Γ ⊢⇩w⇩f v : b" using wfE_elims by metis
hence "supp v ⊆ atom_dom Γ ∪ supp ℬ" using wfE_appI wf_supp1 by metis
hence "supp (AE_app f v) ⊆ atom_dom Γ ∪ supp ℬ" using e.supp pure_supp by fast
then show ?case using e.supp(2) UnCI subsetCE subsetI wfE_appI using b.supp(3) pure_supp x_not_in_b_set by metis
next
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv v τ f xa ba ca s)
then obtain b where "Θ; ℬ; Γ ⊢⇩w⇩f v : ( b[bv::=b']⇩b)" using wfE_elims by metis
hence "supp v ⊆ atom_dom Γ ∪ supp ℬ " using wfE_appPI wf_supp1 by auto
moreover have "supp b' ⊆ supp ℬ" using wf_supp1(7) wfE_appPI by simp
ultimately show ?case unfolding e.supp using wfE_appPI pure_supp by fast
next
case (wfE_mvarI Θ Φ ℬ Γ Δ u τ)
then obtain τ where "(u,τ) ∈ setD Δ" using wfE_elims(10) by metis
hence "atom u ∈ atom`fst`setD Δ" by force
hence "supp (AE_mvar u ) ⊆ atom`fst`setD Δ" using e.supp
by (simp add: supp_at_base)
thus ?case using UnCI subsetCE subsetI e.supp wfE_mvarI supp_at_base subsetCE supp_at_base u_not_in_b_set
by (simp add: supp_at_base)
next
case (wfS_valI Θ Φ ℬ Γ v b Δ)
then show ?case using wf_supp1
by (metis s_branch_s_branch_list.supp(1) sup.coboundedI2 sup_assoc sup_commute)
next
case (wfS_letI Θ Φ ℬ Γ Δ e b' x s b)
then show ?case by auto
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 τ x s2 b)
then show ?case unfolding s_branch_s_branch_list.supp (3) using wf_supp1(4)[OF wfS_let2I(3)] by auto
next
case (wfS_ifI Θ ℬ Γ v Φ Δ s1 b s2)
then show ?case using wf_supp1(1)[OF wfS_ifI(1)] by auto
next
case (wfS_varI Θ ℬ Γ τ v u Δ Φ s b)
then show ?case using wf_supp1(1)[OF wfS_varI(2)] wf_supp1(4)[OF wfS_varI(1)] by auto
next
next
case (wfS_assignI u τ Δ Θ ℬ Γ Φ v)
hence "supp u ⊆ atom ` fst ` setD Δ" proof(induct Δ rule:Δ_induct)
case DNil
then show ?case by auto
next
case (DCons u' t' Δ')
show ?case proof(cases "u=u'")
case True
then show ?thesis using toSet.simps DCons supp_at_base by fastforce
next
case False
then show ?thesis using toSet.simps DCons supp_at_base wfS_assignI
by (metis empty_subsetI fstI image_eqI insert_subset)
qed
qed
then show ?case using s_branch_s_branch_list.supp(8) wfS_assignI wf_supp1(1)[OF wfS_assignI(6)] by auto
next
case (wfS_matchI Θ ℬ Γ v tid dclist Δ Φ cs b)
then show ?case using wf_supp1(1)[OF wfS_matchI(1)] by auto
next
case (wfS_branchI Θ Φ ℬ x τ Γ Δ s b tid dc)
moreover have "supp s ⊆ supp x ∪ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ"
using dom_cons supp_at_base wfS_branchI by auto
moreover hence "supp s - set [atom x] ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" using supp_at_base by force
ultimately have
"(supp s - set [atom x]) ∪ (supp dc ) ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ"
by (simp add: pure_supp)
thus ?case using s_branch_s_branch_list.supp(2) by auto
next
case (wfD_emptyI Θ ℬ Γ)
then show ?case using supp_DNil by auto
next
case (wfD_cons Θ ℬ Γ Δ τ u)
have "supp ((u, τ) #⇩Δ Δ) = supp u ∪ supp τ ∪ supp Δ" using supp_DCons supp_Pair by metis
also have "... ⊆ supp u ∪ atom ` fst ` setD Δ ∪ atom_dom Γ ∪ supp ℬ"
using wfD_cons wf_supp1(4)[OF wfD_cons(3)] by auto
also have "... ⊆ atom ` fst ` setD ((u, τ) #⇩Δ Δ) ∪ atom_dom Γ ∪ supp ℬ" using supp_at_base by auto
finally show ?case by auto
next
case (wfPhi_emptyI Θ)
then show ?case using supp_Nil by auto
next
case (wfPhi_consI f Θ Φ ft)
then show ?case using fun_def.supp
by (simp add: pure_supp supp_Cons)
next
case (wfFTI Θ B' b s x c τ Φ)
have " supp (AF_fun_typ x b c τ s) = supp c ∪ (supp τ ∪ supp s) - set [atom x] ∪ supp b" using fun_typ.supp by auto
thus ?case using wfFTI wf_supp1
proof -
have f1: "supp τ ⊆ {atom x} ∪ atom_dom GNil ∪ supp B'"
using dom_cons wfFTI.hyps wf_supp1(4) by blast
have "supp b ⊆ supp B'"
using wfFTI.hyps(1) wf_supp1(7) by blast
then show ?thesis
using f1 ‹supp (AF_fun_typ x b c τ s) = supp c ∪ (supp τ ∪ supp s) - set [atom x] ∪ supp b›
wfFTI.hyps(4) wfFTI.hyps by auto
qed
next
case (wfFTNone Θ Φ ft)
then show ?case by (simp add: fun_typ_q.supp(2))
next
case (wfFTSome Θ Φ bv ft)
then show ?case using fun_typ_q.supp
by (simp add: supp_at_base)
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
then have "supp c ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" using wf_supp1
by (metis Un_assoc Un_commute le_supI2)
moreover have "supp s ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" proof
fix z
assume *:"z ∈ supp s"
have **:"atom x ∉ supp s" using wfS_assertI fresh_prodN fresh_def by metis
have "z ∈ atom_dom ((x, B_bool, c) #⇩Γ Γ) ∪ atom ` fst ` setD Δ ∪ supp ℬ" using wfS_assertI * by blast
have "z ∈ atom_dom ((x, B_bool, c) #⇩Γ Γ) ⟹ z ∈ atom_dom Γ" using * ** by auto
thus "z ∈ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" using * **
using ‹z ∈ atom_dom ((x, B_bool, c) #⇩Γ Γ) ∪ atom ` fst ` setD Δ ∪ supp ℬ› by blast
qed
ultimately show ?case by auto
qed(auto)
lemmas wf_supp = wf_supp1 wf_supp2
lemma wfV_supp_nil:
fixes v::v
assumes "P ; {||} ; GNil ⊢⇩w⇩f v : b"
shows "supp v = {}"
using wfV_supp[of P " {||}" GNil v b] dom.simps toSet.simps
using assms by auto
lemma wfT_TRUE_aux:
assumes "wfG P ℬ Γ" and "atom z ♯ (P, ℬ, Γ)" and "wfB P ℬ b"
shows "wfT P ℬ Γ (⦃ z : b | TRUE ⦄)"
proof (rule)
show ‹ atom z ♯ (P, ℬ, Γ)› using assms by auto
show ‹ P; ℬ ⊢⇩w⇩f b › using assms by auto
show ‹ P ;ℬ ; (z, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f TRUE › using wfG_cons2I wfC_trueI assms by auto
qed
lemma wfT_TRUE:
assumes "wfG P ℬ Γ" and "wfB P ℬ b"
shows "wfT P ℬ Γ (⦃ z : b | TRUE ⦄)"
proof -
obtain z'::x where *:"atom z' ♯ (P, ℬ, Γ)" using obtain_fresh by metis
hence "⦃ z : b | TRUE ⦄ = ⦃ z' : b | TRUE ⦄" by auto
thus ?thesis using wfT_TRUE_aux assms * by metis
qed
lemma phi_flip_eq:
assumes "wfPhi T P"
shows "(x ↔ xa) ∙ P = P"
using wfPhi_supp[OF assms] flip_fresh_fresh fresh_def by blast
lemma wfC_supp_cons:
fixes c'::c and G::Γ
assumes "P; ℬ ; (x', b' , TRUE) #⇩ΓG ⊢⇩w⇩f c'"
shows "supp c' ⊆ atom_dom G ∪ supp x' ∪ supp ℬ" and "supp c' ⊆ supp G ∪ supp x' ∪ supp ℬ"
proof -
show "supp c' ⊆ atom_dom G ∪ supp x' ∪ supp ℬ"
using wfC_supp[OF assms] dom_cons supp_at_base by blast
moreover have "atom_dom G ⊆ supp G"
by (meson assms wfC_wf wfG_cons wfG_supp)
ultimately show "supp c' ⊆ supp G ∪ supp x' ∪ supp ℬ" using wfG_supp assms wfG_cons wfC_wf by fast
qed
lemma wfG_dom_supp:
fixes x::x
assumes "wfG P ℬ G"
shows "atom x ∈ atom_dom G ⟷ atom x ∈ supp G"
using assms proof(induct G rule: Γ_induct)
case GNil
then show ?case using dom.simps supp_of_atom_list
using supp_GNil by auto
next
case (GCons x' b' c' G)
show ?case proof(cases "x' = x")
case True
then show ?thesis using dom.simps supp_of_atom_list supp_at_base
using supp_GCons by auto
next
case False
have "(atom x ∈ atom_dom ((x', b', c') #⇩Γ G)) = (atom x ∈ atom_dom G)" using atom_dom.simps False by simp
also have "... = (atom x ∈ supp G)" using GCons wfG_elims by metis
also have "... = (atom x ∈ (supp (x', b', c') ∪ supp G))" proof
show "atom x ∈ supp G ⟹ atom x ∈ supp (x', b', c') ∪ supp G" by auto
assume "atom x ∈ supp (x', b', c') ∪ supp G"
then consider "atom x ∈ supp (x', b', c')" | "atom x ∈ supp G" by auto
then show "atom x ∈ supp G" proof(cases)
case 1
assume " atom x ∈ supp (x', b', c') "
hence "atom x ∈ supp c'" using supp_triple False supp_b_empty supp_at_base by force
moreover have "P; ℬ ; (x', b' , TRUE) #⇩ΓG ⊢⇩w⇩f c'" using wfG_elim2 GCons by simp
moreover hence "supp c' ⊆ supp G ∪ supp x' ∪ supp ℬ" using wfC_supp_cons by auto
ultimately have "atom x ∈ supp G ∪ supp x' " using x_not_in_b_set by auto
then show ?thesis using False supp_at_base by (simp add: supp_at_base)
next
case 2
then show ?thesis by simp
qed
qed
also have "... = (atom x ∈ supp ((x', b', c') #⇩Γ G))" using supp_at_base False supp_GCons by simp
finally show ?thesis by simp
qed
qed
lemma wfG_atoms_supp_eq :
fixes x::x
assumes "wfG P ℬ G"
shows "atom x ∈ atom_dom G ⟷ atom x ∈ supp G"
using wfG_dom_supp assms by auto
lemma beta_flip_eq:
fixes x::x and xa::x and ℬ::ℬ
shows "(x ↔ xa) ∙ ℬ = ℬ"
proof -
have "atom x ♯ ℬ ∧ atom xa ♯ ℬ" using x_not_in_b_set fresh_def supp_set by metis
thus ?thesis by (simp add: flip_fresh_fresh fresh_def)
qed
lemma theta_flip_eq2:
assumes "⊢⇩w⇩f Θ"
shows " (z ↔ za ) ∙ Θ = Θ"
proof -
have "supp Θ = {}" using wfTh_supp assms by simp
thus ?thesis
by (simp add: flip_fresh_fresh fresh_def)
qed
lemma theta_flip_eq:
assumes "wfTh Θ"
shows "(x ↔ xa) ∙ Θ = Θ"
using wfTh_supp flip_fresh_fresh fresh_def
by (simp add: assms theta_flip_eq2)
lemma wfT_wfC:
fixes c::c
assumes "Θ; ℬ; Γ ⊢⇩w⇩f ⦃ z : b | c ⦄" and "atom z ♯ Γ"
shows "Θ; ℬ; (z,b,TRUE) #⇩ΓΓ ⊢⇩w⇩f c"
proof -
obtain za ba ca where *:"⦃ z : b | c ⦄ = ⦃ za : ba | ca ⦄ ∧ atom za ♯ (Θ,ℬ,Γ) ∧ Θ; ℬ; (za, ba, TRUE) #⇩Γ Γ ⊢⇩w⇩f ca"
using wfT_elims[OF assms(1)] by metis
hence c1: "[[atom z]]lst. c = [[atom za]]lst. ca" using τ.eq_iff by meson
show ?thesis proof(cases "z=za")
case True
hence "ca = c" using c1 by (simp add: Abs1_eq_iff(3))
then show ?thesis using * True by simp
next
case False
have " ⊢⇩w⇩f Θ" using wfT_wf wfG_wf assms by metis
moreover have "atom za ♯ Γ" using * fresh_prodN by auto
ultimately have "Θ; ℬ; (z ↔ za ) ∙ (za, ba, TRUE) #⇩Γ Γ ⊢⇩w⇩f (z ↔ za ) ∙ ca"
using wfC.eqvt theta_flip_eq2 beta_flip_eq * GCons_eqvt assms flip_fresh_fresh by metis
moreover have "atom z ♯ ca"
proof -
have "supp ca ⊆ atom_dom Γ ∪ { atom za } ∪ supp ℬ" using * wfC_supp atom_dom.simps toSet.simps by fastforce
moreover have "atom z ∉ atom_dom Γ " using assms fresh_def wfT_wf wfG_dom_supp wfC_supp by metis
moreover hence "atom z ∉ atom_dom Γ ∪ { atom za }" using False by simp
moreover have "atom z ∉ supp ℬ" using x_not_in_b_set by simp
ultimately show ?thesis using fresh_def False by fast
qed
moreover hence "(z ↔ za ) ∙ ca = c" using type_eq_subst_eq1(3) * by metis
ultimately show ?thesis using assms G_cons_flip_fresh * by auto
qed
qed
lemma u_not_in_dom_g:
fixes u::u
shows "atom u ∉ atom_dom G"
using toSet.simps atom_dom.simps u_not_in_x_atoms by auto
lemma bv_not_in_dom_g:
fixes bv::bv
shows "atom bv ∉ atom_dom G"
using toSet.simps atom_dom.simps u_not_in_x_atoms by auto
text ‹An important lemma that confirms that @{term Γ} does not rely on mutable variables›
lemma u_not_in_g:
fixes u::u
assumes "wfG Θ B G"
shows "atom u ∉ supp G"
using assms proof(induct G rule: Γ_induct)
case GNil
then show ?case using supp_GNil fresh_def
using fresh_set_empty by fastforce
next
case (GCons x b c Γ')
moreover hence "atom u ∉ supp b" using
wfB_supp wfC_supp u_not_in_x_atoms wfG_elims wfX_wfY by auto
moreover hence "atom u ∉ supp x" using u_not_in_x_atoms supp_at_base by blast
moreover hence "atom u ∉ supp c" proof -
have "Θ ; B ; (x, b, TRUE) #⇩Γ Γ' ⊢⇩w⇩f c" using wfG_cons_wfC GCons by simp
hence "supp c ⊆ atom_dom ((x, b, TRUE) #⇩Γ Γ') ∪ supp B" using wfC_supp by blast
thus ?thesis using u_not_in_dom_g u_not_in_b_atoms
using u_not_in_b_set by auto
qed
ultimately have "atom u ∉ supp (x,b,c)" using supp_Pair by simp
thus ?case using supp_GCons GCons wfG_elims by blast
qed
text ‹An important lemma that confirms that types only depend on immutable variables›
lemma u_not_in_t:
fixes u::u
assumes "wfT Θ B G τ"
shows "atom u ∉ supp τ"
proof -
have "supp τ ⊆ atom_dom G ∪ supp B" using wfT_supp assms by auto
thus ?thesis using u_not_in_dom_g u_not_in_b_set by blast
qed
lemma wfT_supp_c:
fixes ℬ::ℬ and z::x
assumes "wfT P ℬ Γ (⦃ z : b | c ⦄)"
shows "supp c - { atom z } ⊆ atom_dom Γ ∪ supp ℬ"
using wf_supp τ.supp assms
by (metis Un_subset_iff empty_set list.simps(15))
lemma wfG_wfC[ms_wb]:
assumes "wfG P ℬ ((x,b,c) #⇩ΓΓ)"
shows "wfC P ℬ ((x,b,TRUE) #⇩ΓΓ) c"
using assms proof(cases "c ∈ {TRUE,FALSE}")
case True
have "atom x ♯ Γ ∧ wfG P ℬ Γ ∧ wfB P ℬ b" using wfG_cons assms by auto
hence "wfG P ℬ ((x,b,TRUE) #⇩ΓΓ)" using wfG_cons2I by auto
then show ?thesis using wfC_trueI wfC_falseI True by auto
next
case False
then show ?thesis using wfG_elims assms by blast
qed
lemma wfT_wf_cons:
assumes "wfT P ℬ Γ ⦃ z : b | c ⦄" and "atom z ♯ Γ"
shows "wfG P ℬ ((z,b,c) #⇩ΓΓ)"
using assms proof(cases "c ∈ { TRUE,FALSE }")
case True
then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons2I assms wfT_wf by fastforce
next
case False
then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons1I wfT_wf wfT_wfC assms by fastforce
qed
lemma wfV_b_fresh:
fixes b::b and v::v and bv::bv
assumes "Θ; ℬ; Γ ⊢⇩w⇩f v: b" and "bv |∉| ℬ"
shows "atom bv ♯ v"
using wfV_supp bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp by blast
lemma wfCE_b_fresh:
fixes b::b and ce::ce and bv::bv
assumes "Θ; ℬ; Γ ⊢⇩w⇩f ce: b" and "bv |∉| ℬ"
shows "atom bv ♯ ce"
using bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp wf_supp1(8) by fast
section ‹Freshness›
lemma wfG_fresh_x:
fixes Γ::Γ and z::x
assumes "Θ; ℬ ⊢⇩w⇩f Γ" and "atom z ♯ Γ"
shows "atom z ♯ (Θ,ℬ, Γ)"
unfolding fresh_prodN apply(intro conjI)
using wf_supp1 wfX_wfY assms fresh_def x_not_in_b_set by(metis empty_iff)+
lemma wfG_wfT:
assumes "wfG P ℬ ((x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ G)" and "atom x ♯ c"
shows "P; ℬ ; G ⊢⇩w⇩f ⦃ z : b | c ⦄"
proof -
have " P; ℬ ; (x, b, TRUE) #⇩Γ G ⊢⇩w⇩f c[z::=V_var x]⇩c⇩v ∧ wfB P ℬ b" using assms
using wfG_elim2 by auto
moreover have "atom x ♯ (P ,ℬ,G)" using wfG_elims assms wfG_fresh_x by metis
ultimately have "wfT P ℬ G ⦃ x : b | c[z::=V_var x]⇩c⇩v ⦄" using wfTI assms by metis
moreover have "⦃ x : b | c[z::=V_var x]⇩c⇩v ⦄ = ⦃ z : b | c ⦄" using type_eq_subst ‹atom x ♯ c› by auto
ultimately show ?thesis by auto
qed
lemma wfT_wfT_if:
assumes "wfT Θ ℬ Γ (⦃ z2 : b | CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]⇩c⇩v ⦄)" and "atom z2 ♯ (c,Γ)"
shows "wfT Θ ℬ Γ ⦃ z : b | c ⦄"
proof -
have *: "atom z2 ♯ (Θ, ℬ, Γ)" using wfG_fresh_x wfX_wfY assms fresh_Pair by metis
have "wfB Θ ℬ b" using assms wfT_elims by metis
have "Θ; ℬ; (GCons (z2,b,TRUE) Γ) ⊢⇩w⇩f (CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]⇩c⇩v)" using wfT_wfC assms fresh_Pair by auto
hence "Θ; ℬ; ((z2,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c[z::=V_var z2]⇩c⇩v" using wfC_elims by metis
hence "wfT Θ ℬ Γ (⦃ z2 : b | c[z::=V_var z2]⇩c⇩v⦄)" using assms fresh_Pair wfTI ‹wfB Θ ℬ b› * by auto
moreover have "⦃ z : b | c ⦄ = ⦃ z2 : b | c[z::=V_var z2]⇩c⇩v ⦄" using type_eq_subst assms fresh_Pair by auto
ultimately show ?thesis using wfTI assms by argo
qed
lemma wfT_fresh_c:
fixes x::x
assumes "wfT P ℬ Γ ⦃ z : b | c ⦄" and "atom x ♯ Γ" and "x ≠ z"
shows "atom x ♯ c"
proof(rule ccontr)
assume "¬ atom x ♯ c"
hence *:"atom x ∈ supp c" using fresh_def by auto
moreover have "supp c - set [atom z] ∪ supp b ⊆ atom_dom Γ ∪ supp ℬ"
using assms wfT_supp τ.supp by blast
moreover hence "atom x ∈ supp c - set [atom z]" using assms * by auto
ultimately have "atom x ∈ atom_dom Γ" using x_not_in_b_set by auto
thus False using assms wfG_atoms_supp_eq wfT_wf fresh_def by metis
qed
lemma wfG_x_fresh [simp]:
fixes x::x
assumes "wfG P ℬ G"
shows "atom x ∉ atom_dom G ⟷ atom x ♯ G"
using wfG_atoms_supp_eq assms fresh_def by metis
lemma wfD_x_fresh:
fixes x::x
assumes "atom x ♯ Γ" and "wfD P B Γ Δ"
shows "atom x ♯ Δ"
using assms proof(induct Δ rule: Δ_induct)
case DNil
then show ?case using supp_DNil fresh_def by auto
next
case (DCons u' t' Δ')
have wfg: "wfG P B Γ" using wfD_wf DCons by blast
hence wfd: "wfD P B Γ Δ'" using wfD_elims DCons by blast
have "supp t' ⊆ atom_dom Γ ∪ supp B" using wfT_supp DCons wfD_elims by metis
moreover have "atom x ∉ atom_dom Γ" using DCons(2) fresh_def wfG_supp wfg by blast
ultimately have "atom x ♯ t'" using fresh_def DCons wfG_supp wfg x_not_in_b_set by blast
moreover have "atom x ♯ u'" using supp_at_base fresh_def by fastforce
ultimately have "atom x ♯ (u',t')" using supp_Pair by fastforce
thus ?case using DCons fresh_DCons wfd by fast
qed
lemma wfG_fresh_x2:
fixes Γ::Γ and z::x and Δ::Δ and Φ::Φ
assumes "Θ; ℬ; Γ ⊢⇩w⇩f Δ" and "Θ ⊢⇩w⇩f Φ" and "atom z ♯ Γ"
shows "atom z ♯ (Θ,Φ,ℬ, Γ,Δ)"
unfolding fresh_prodN apply(intro conjI)
using wfG_fresh_x assms fresh_prod3 wfX_wfY apply metis
using wf_supp2(5) assms fresh_def apply blast
using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis
using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis
using wf_supp2(6) assms fresh_def wfD_x_fresh by metis
lemma wfV_x_fresh:
fixes v::v and b::b and Γ::Γ and x::x
assumes "Θ; ℬ; Γ ⊢⇩w⇩f v : b" and "atom x ♯ Γ"
shows "atom x ♯ v"
proof -
have "supp v ⊆ atom_dom Γ ∪ supp ℬ " using assms wfV_supp by auto
moreover have "atom x ∉ atom_dom Γ" using fresh_def assms
dom.simps subsetCE wfG_elims wfG_supp by (metis dom_supp_g)
moreover have "atom x ∉ supp ℬ" using x_not_in_b_set by auto
ultimately show ?thesis using fresh_def by fast
qed
lemma wfE_x_fresh:
fixes e::e and b::b and Γ::Γ and Δ::Δ and Φ::Φ and x::x
assumes "Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b" and "atom x ♯ Γ"
shows "atom x ♯ e"
proof -
have "wfG Θ ℬ Γ" using assms wfE_wf by auto
hence "supp e ⊆ atom_dom Γ ∪ supp ℬ ∪ atom`fst`setD Δ" using wfE_supp dom.simps assms by auto
moreover have "atom x ∉ atom_dom Γ" using fresh_def assms
dom.simps subsetCE ‹wfG Θ ℬ Γ› wfG_supp by (metis dom_supp_g)
moreover have "atom x ∉ atom`fst`setD Δ" by auto
ultimately show ?thesis using fresh_def x_not_in_b_set by fast
qed
lemma wfT_x_fresh:
fixes τ::τ and Γ::Γ and x::x
assumes "Θ; ℬ; Γ ⊢⇩w⇩f τ" and "atom x ♯ Γ"
shows "atom x ♯ τ"
proof -
have "wfG Θ ℬ Γ" using assms wfX_wfY by auto
hence "supp τ ⊆ atom_dom Γ ∪ supp ℬ" using wfT_supp dom.simps assms by auto
moreover have "atom x ∉ atom_dom Γ" using fresh_def assms
dom.simps subsetCE ‹wfG Θ ℬ Γ› wfG_supp by (metis dom_supp_g)
moreover have "atom x ∉ supp ℬ" using x_not_in_b_set by simp
ultimately show ?thesis using fresh_def by fast
qed
lemma wfS_x_fresh:
fixes s::s and Δ::Δ and x::x
assumes "Θ; Φ; ℬ; Γ; Δ ⊢⇩w⇩f s : b" and "atom x ♯ Γ"
shows "atom x ♯ s"
proof -
have "supp s ⊆ atom_dom Γ ∪ atom ` fst ` setD Δ ∪ supp ℬ" using wf_supp assms by metis
moreover have "atom x ∉ atom ` fst ` setD Δ" by auto
moreover have "atom x ∉ atom_dom Γ" using assms fresh_def wfG_dom_supp wfX_wfY by metis
moreover have "atom x ∉ supp ℬ" using supp_b_empty supp_fset
by (simp add: x_not_in_b_set)
ultimately show ?thesis using fresh_def by fast
qed
lemma wfTh_fresh:
fixes x
assumes "wfTh T"
shows "atom x ♯ T"
using wf_supp1 assms fresh_def by fastforce
lemmas wfTh_x_fresh = wfTh_fresh
lemma wfPhi_fresh:
fixes x
assumes "wfPhi T P"
shows "atom x ♯ P"
using wf_supp assms fresh_def by fastforce
lemmas wfPhi_x_fresh = wfPhi_fresh
lemmas wb_x_fresh = wfTh_x_fresh wfPhi_x_fresh wfD_x_fresh wfT_x_fresh wfV_x_fresh
lemma wfG_inside_fresh[ms_fresh]:
fixes Γ::Γ and x::x
assumes "wfG P ℬ (Γ'@((x,b,c) #⇩ΓΓ))"
shows "atom x ∉ atom_dom Γ'"
using assms proof(induct Γ' rule: Γ_induct)
case GNil
then show ?case by auto
next
case (GCons x1 b1 c1 Γ1)
moreover hence "atom x ∉ atom ` fst `({(x1,b1,c1)})" proof -
have *: "P; ℬ ⊢⇩w⇩f (Γ1 @ (x, b, c) #⇩Γ Γ)" using wfG_elims append_g.simps GCons by metis
have "atom x1 ♯ (Γ1 @ (x, b, c) #⇩Γ Γ)" using GCons wfG_elims append_g.simps by metis
hence "atom x1 ∉ atom_dom (Γ1 @ (x, b, c) #⇩Γ Γ)" using wfG_dom_supp fresh_def * by metis
thus ?thesis by auto
qed
ultimately show ?case using append_g.simps atom_dom.simps toSet.simps wfG_elims dom.simps
by (metis image_insert insert_iff insert_is_Un)
qed
lemma wfG_inside_x_in_atom_dom:
fixes c::c and x::x and Γ::Γ
shows "atom x ∈ atom_dom ( Γ'@ (x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ Γ)"
by(induct Γ' rule: Γ_induct, (simp add: toSet.simps atom_dom.simps)+)
lemma wfG_inside_x_neq:
fixes c::c and x::x and Γ::Γ and G::Γ and xa::x
assumes "G=( Γ'@ (x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ Γ)" and "atom xa ♯ G" and " Θ; ℬ ⊢⇩w⇩f G"
shows "xa ≠ x"
proof -
have "atom xa ∉ atom_dom G" using fresh_def wfG_atoms_supp_eq assms by metis
moreover have "atom x ∈ atom_dom G" using wfG_inside_x_in_atom_dom assms by simp
ultimately show ?thesis by auto
qed
lemma wfG_inside_x_fresh:
fixes c::c and x::x and Γ::Γ and G::Γ and xa::x
assumes "G=( Γ'@ (x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ Γ)" and "atom xa ♯ G" and " Θ; ℬ ⊢⇩w⇩f G"
shows "atom xa ♯ x"
using fresh_def supp_at_base wfG_inside_x_neq assms by auto
lemma wfT_nil_supp:
fixes t::τ
assumes "Θ ; {||} ; GNil ⊢⇩w⇩f t"
shows "supp t = {}"
using wfT_supp atom_dom.simps assms toSet.simps by force
section ‹Misc›
lemma wfG_cons_append:
fixes b'::b
assumes "Θ; ℬ ⊢⇩w⇩f ((x', b', c') #⇩Γ Γ') @ (x, b, c) #⇩Γ Γ"
shows "Θ; ℬ ⊢⇩w⇩f (Γ' @ (x, b, c) #⇩Γ Γ) ∧ atom x' ♯ (Γ' @ (x, b, c) #⇩Γ Γ) ∧ Θ; ℬ ⊢⇩w⇩f b' ∧ x' ≠ x"
proof -
have "((x', b', c') #⇩Γ Γ') @ (x, b, c) #⇩Γ Γ = (x', b', c') #⇩Γ (Γ' @ (x, b, c) #⇩Γ Γ)" using append_g.simps by auto
hence *:"Θ; ℬ ⊢⇩w⇩f (Γ' @ (x, b, c) #⇩Γ Γ) ∧ atom x' ♯ (Γ' @ (x, b, c) #⇩Γ Γ) ∧ Θ; ℬ ⊢⇩w⇩f b'" using assms wfG_cons by metis
moreover have "atom x' ♯ x" proof(rule wfG_inside_x_fresh[of "(Γ' @ (x, b, c) #⇩Γ Γ)"])
show "Γ' @ (x, b, c) #⇩Γ Γ = Γ' @ (x, b, c[x::=V_var x]⇩c⇩v) #⇩Γ Γ" by simp
show " atom x' ♯ Γ' @ (x, b, c) #⇩Γ Γ" using * by auto
show "Θ; ℬ ⊢⇩w⇩f Γ' @ (x, b, c) #⇩Γ Γ " using * by auto
qed
ultimately show ?thesis by auto
qed
lemma flip_u_eq:
fixes u::u and u'::u and Θ::Θ and τ::τ
assumes "Θ; ℬ; Γ ⊢⇩w⇩f τ"
shows "(u ↔ u') ∙ τ = τ" and "(u ↔ u') ∙ Γ = Γ" and "(u ↔ u') ∙ Θ = Θ" and "(u ↔ u') ∙ ℬ = ℬ"
proof -
show "(u ↔ u') ∙ τ = τ" using wfT_supp flip_fresh_fresh
by (metis assms(1) fresh_def u_not_in_t)
show "(u ↔ u') ∙ Γ = Γ" using u_not_in_g wfX_wfY assms flip_fresh_fresh fresh_def by metis
show "(u ↔ u') ∙ Θ = Θ" using theta_flip_eq assms wfX_wfY by metis
show "(u ↔ u') ∙ ℬ = ℬ" using u_not_in_b_set flip_fresh_fresh fresh_def by metis
qed
lemma wfT_wf_cons_flip:
fixes c::c and x::x
assumes "wfT P ℬ Γ ⦃ z : b | c ⦄" and "atom x ♯ (c,Γ)"
shows "wfG P ℬ ((x,b,c[z::=V_var x]⇩c⇩v) #⇩ΓΓ)"
proof -
have "⦃ x : b | c[z::=V_var x]⇩c⇩v ⦄ = ⦃ z : b | c ⦄" using assms freshers type_eq_subst by metis
hence *:"wfT P ℬ Γ ⦃ x : b | c[z::=V_var x]⇩c⇩v ⦄" using assms by metis
show ?thesis proof(rule wfG_consI)
show ‹ P; ℬ ⊢⇩w⇩f Γ › using assms wfT_wf by auto
show ‹atom x ♯ Γ› using assms by auto
show ‹ P; ℬ ⊢⇩w⇩f b › using assms wfX_wfY b_of.simps by metis
show ‹ P; ℬ ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c[z::=V_var x]⇩c⇩v › using wfT_wfC * assms fresh_Pair by metis
qed
qed
section ‹Context Strengthening›
text ‹We can remove an entry for a variable from the context if the variable doesn't appear in the
term and the variable is not used later in the context or any other context›
lemma fresh_restrict:
fixes y::"'a::at_base" and Γ::Γ
assumes "atom y ♯ (Γ' @ (x, b, c) #⇩Γ Γ)"
shows "atom y ♯ (Γ'@Γ)"
using assms proof(induct Γ' rule: Γ_induct)
case GNil
then show ?case using fresh_GCons fresh_GNil by auto
next
case (GCons x' b' c' Γ'')
then show ?case using fresh_GCons fresh_GNil by auto
qed
lemma wf_restrict1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ atom x ♯ v ⟹ atom x ♯ Γ⇩1 ⟹ Θ; ℬ; Γ⇩1@Γ⇩2 ⊢⇩w⇩f v : b" and
"Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ atom x ♯ c⟹ atom x ♯ Γ⇩1 ⟹ Θ ; ℬ ; Γ⇩1@Γ⇩2 ⊢⇩w⇩f c" and
"Θ; ℬ ⊢⇩w⇩f Γ ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ atom x ♯ Γ⇩1 ⟹ Θ; ℬ ⊢⇩w⇩f Γ⇩1@Γ⇩2" and
"Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ atom x ♯ τ⟹ atom x ♯ Γ⇩1 ⟹ Θ; ℬ; Γ⇩1@Γ⇩2 ⊢⇩w⇩f τ" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f Θ ⟹True" and
"Θ; ℬ ⊢⇩w⇩f b ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ atom x ♯ ce ⟹ atom x ♯ Γ⇩1 ⟹ Θ; ℬ; Γ⇩1@Γ⇩2 ⊢⇩w⇩f ce : b" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(induct arbitrary: Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
case (wfV_varI Θ ℬ Γ b c y)
hence "y≠x" using v.fresh by auto
hence "Some (b, c) = lookup (Γ⇩1@Γ⇩2) y" using lookup_restrict wfV_varI by metis
then show ?case using wfV_varI wf_intros by metis
next
case (wfV_litI Θ Γ l)
then show ?case using e.fresh wf_intros by metis
next
case (wfV_pairI Θ ℬ Γ v1 b1 v2 b2)
show ?case proof
show "Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f v1 : b1" using wfV_pairI by auto
show "Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f v2 : b2" using wfV_pairI by auto
qed
next
case (wfV_consI s dclist Θ dc x b c ℬ Γ v)
show ?case proof
show "AF_typedef s dclist ∈ set Θ" using wfV_consI by auto
show "(dc, ⦃ x : b | c ⦄) ∈ set dclist" using wfV_consI by auto
show "Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f v : b" using wfV_consI by auto
qed
next
case (wfV_conspI s bv dclist Θ dc x b' c ℬ b Γ v)
show ?case proof
show "AF_typedef_poly s bv dclist ∈ set Θ" using wfV_conspI by auto
show "(dc, ⦃ x : b' | c ⦄) ∈ set dclist" using wfV_conspI by auto
show "Θ; ℬ ⊢⇩w⇩f b" using wfV_conspI by auto
show " Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f v : b'[bv::=b]⇩b⇩b" using wfV_conspI by auto
show "atom bv ♯ (Θ, ℬ, Γ⇩1 @ Γ⇩2, b, v)" unfolding fresh_prodN fresh_append_g using wfV_conspI fresh_prodN fresh_GCons fresh_append_g by metis
qed
next
case (wfCE_valI Θ ℬ Γ v b)
then show ?case using ce.fresh wf_intros by metis
next
case (wfCE_plusI Θ ℬ Γ v1 v2)
then show ?case using ce.fresh wf_intros by metis
next
case (wfCE_leqI Θ ℬ Γ v1 v2)
then show ?case using ce.fresh wf_intros by metis
next
case (wfCE_eqI Θ ℬ Γ v1 v2)
then show ?case using ce.fresh wf_intros by metis
next
case (wfCE_fstI Θ ℬ Γ v1 b1 b2)
then show ?case using ce.fresh wf_intros by metis
next
case (wfCE_sndI Θ ℬ Γ v1 b1 b2)
then show ?case using ce.fresh wf_intros by metis
next
case (wfCE_concatI Θ ℬ Γ v1 v2)
then show ?case using ce.fresh wf_intros by metis
next
case (wfCE_lenI Θ ℬ Γ v1)
then show ?case using ce.fresh wf_intros by metis
next
case (wfTI z Θ ℬ Γ b c)
hence "x ≠ z" using wfTI
fresh_GCons fresh_prodN fresh_PairD(1) fresh_gamma_append not_self_fresh by metis
show ?case proof
show ‹atom z ♯ (Θ, ℬ, Γ⇩1 @ Γ⇩2)› using wfTI fresh_restrict[of z] using wfG_fresh_x wfX_wfY wfTI fresh_prodN by metis
show ‹ Θ; ℬ ⊢⇩w⇩f b › using wfTI by auto
have "Θ; ℬ; ((z, b, TRUE) #⇩Γ Γ⇩1) @ Γ⇩2 ⊢⇩w⇩f c " proof(rule wfTI(5)[of "(z, b, TRUE) #⇩Γ Γ⇩1" ])
show ‹(z, b, TRUE) #⇩Γ Γ = ((z, b, TRUE) #⇩Γ Γ⇩1) @ (x, b', c') #⇩Γ Γ⇩2› using wfTI by auto
show ‹atom x ♯ c› using wfTI τ.fresh ‹x ≠ z› by auto
show ‹atom x ♯ (z, b, TRUE) #⇩Γ Γ⇩1› using wfTI ‹x ≠ z› fresh_GCons by simp
qed
thus ‹ Θ; ℬ; (z, b, TRUE) #⇩Γ Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f c › by auto
qed
next
case (wfC_eqI Θ ℬ Γ e1 b e2)
show ?case proof
show "Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f e1 : b " using wfC_eqI c.fresh fresh_Nil by auto
show "Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f e2 : b " using wfC_eqI c.fresh fresh_Nil by auto
qed
next
case (wfC_trueI Θ Γ)
then show ?case using c.fresh wf_intros by metis
next
case (wfC_falseI Θ Γ)
then show ?case using c.fresh wf_intros by metis
next
case (wfC_conjI Θ Γ c1 c2)
then show ?case using c.fresh wf_intros by metis
next
case (wfC_disjI Θ Γ c1 c2)
then show ?case using c.fresh wf_intros by metis
next
case (wfC_notI Θ Γ c1)
then show ?case using c.fresh wf_intros by metis
next
case (wfC_impI Θ Γ c1 c2)
then show ?case using c.fresh wf_intros by metis
next
case (wfG_nilI Θ)
then show ?case using wfV_varI wf_intros
by (meson GNil_append Γ.simps(3))
next
case (wfG_cons1I c1 Θ ℬ G x1 b1)
show ?case proof(cases "Γ⇩1=GNil")
case True
then show ?thesis using wfG_cons1I wfG_consI by auto
next
case False
then obtain G'::Γ where *:"(x1, b1, c1) #⇩Γ G' = Γ⇩1" using GCons_eq_append_conv wfG_cons1I by auto
hence **:"G=G' @ (x, b', c') #⇩Γ Γ⇩2" using wfG_cons1I by auto
have " Θ; ℬ ⊢⇩w⇩f (x1, b1, c1) #⇩Γ (G' @ Γ⇩2)" proof(rule Wellformed.wfG_cons1I)
show ‹c1 ∉ {TRUE, FALSE}› using wfG_cons1I by auto
show ‹atom x1 ♯ G' @ Γ⇩2› using wfG_cons1I(4) ** fresh_restrict by metis
have " atom x ♯ G'" using wfG_cons1I * using fresh_GCons by blast
thus ‹ Θ; ℬ ⊢⇩w⇩f G' @ Γ⇩2 › using wfG_cons1I(3)[of G'] ** by auto
have "atom x ♯ c1 ∧ atom x ♯ (x1, b1, TRUE) #⇩Γ G'" using fresh_GCons ‹atom x ♯ Γ⇩1› * by auto
thus ‹ Θ; ℬ; (x1, b1, TRUE) #⇩Γ G' @ Γ⇩2 ⊢⇩w⇩f c1 › using wfG_cons1I(6)[of "(x1, b1, TRUE) #⇩Γ G'"] ** * wfG_cons1I by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b1 › using wfG_cons1I by auto
qed
thus ?thesis using * by auto
qed
next
case (wfG_cons2I c1 Θ ℬ G x1 b1)
show ?case proof(cases "Γ⇩1=GNil")
case True
then show ?thesis using wfG_cons2I wfG_consI by auto
next
case False
then obtain G'::Γ where *:"(x1, b1, c1) #⇩Γ G' = Γ⇩1" using GCons_eq_append_conv wfG_cons2I by auto
hence **:"G=G' @ (x, b', c') #⇩Γ Γ⇩2" using wfG_cons2I by auto
have " Θ; ℬ ⊢⇩w⇩f (x1, b1, c1) #⇩Γ (G' @ Γ⇩2)" proof(rule Wellformed.wfG_cons2I)
show ‹c1 ∈ {TRUE, FALSE}› using wfG_cons2I by auto
show ‹atom x1 ♯ G' @ Γ⇩2› using wfG_cons2I ** fresh_restrict by metis
have " atom x ♯ G'" using wfG_cons2I * using fresh_GCons by blast
thus ‹ Θ; ℬ ⊢⇩w⇩f G' @ Γ⇩2 › using wfG_cons2I ** by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b1 › using wfG_cons2I by auto
qed
thus ?thesis using * by auto
qed
qed(auto)+
lemma wf_restrict2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows "Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ atom x ♯ e ⟹ atom x ♯ Γ⇩1 ⟹ atom x ♯ Δ ⟹ Θ; Φ; ℬ; Γ⇩1@Γ⇩2 ; Δ ⊢⇩w⇩f e : b" and
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ True" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ True" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ True" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ True " and
"Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ atom x ♯ Γ⇩1 ⟹ atom x ♯ Δ ⟹ Θ; ℬ; Γ⇩1@Γ⇩2 ⊢⇩w⇩f Δ" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ True"
proof(induct arbitrary: Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
case (wfE_valI Θ Φ Γ Δ v b)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_plusI Θ Φ Γ Δ v1 v2)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_leqI Θ Φ Γ Δ v1 v2)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_eqI Θ Φ Γ Δ v1 b v2)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_fstI Θ Φ Γ Δ v1 b1 b2)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_sndI Θ Φ Γ Δ v1 b1 b2)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_concatI Θ Φ Γ Δ v1 v2)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_splitI Θ Φ Γ Δ v1 v2)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_lenI Θ Φ Γ Δ v1)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_appI Θ Φ Γ Δ f x b c τ s' v)
then show ?case using e.fresh wf_intros wf_restrict1 by metis
next
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv v τ f x b c s)
show ?case proof
show ‹ Θ ⊢⇩w⇩f Φ › using wfE_appPI by auto
show ‹ Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f Δ › using wfE_appPI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b' › using wfE_appPI by auto
have "atom bv ♯ Γ⇩1 @ Γ⇩2" using wfE_appPI fresh_prodN fresh_restrict by metis
thus ‹atom bv ♯ (Φ, Θ, ℬ, Γ⇩1 @ Γ⇩2, Δ, b', v, (b_of τ)[bv::=b']⇩b)›
using wfE_appPI fresh_prodN by auto
show ‹Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f› using wfE_appPI by auto
show ‹ Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f v : b[bv::=b']⇩b › using wfE_appPI wf_restrict1 by auto
qed
next
case (wfE_mvarI Θ Φ Γ Δ u τ)
then show ?case using e.fresh wf_intros by metis
next
case (wfD_emptyI Θ Γ)
then show ?case using c.fresh wf_intros wf_restrict1 by metis
next
case (wfD_cons Θ ℬ Γ Δ τ u)
show ?case proof
show "Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f Δ" using wfD_cons fresh_DCons by metis
show "Θ; ℬ; Γ⇩1 @ Γ⇩2 ⊢⇩w⇩f τ " using wfD_cons fresh_DCons fresh_Pair wf_restrict1 by metis
show "u ∉ fst ` setD Δ" using wfD_cons by auto
qed
next
case (wfFTNone Θ ft)
then show ?case by auto
next
case (wfFTSome Θ bv ft)
then show ?case by auto
next
case (wfFTI Θ B b Φ x c s τ)
then show ?case by auto
qed(auto)+
lemmas wf_restrict=wf_restrict1 wf_restrict2
lemma wfT_restrict2:
fixes τ::τ
assumes "wfT Θ ℬ ((x, b, c) #⇩Γ Γ) τ" and "atom x ♯ τ"
shows "Θ; ℬ; Γ ⊢⇩w⇩f τ"
using wf_restrict1(4)[of Θ ℬ "((x, b, c) #⇩Γ Γ)" τ GNil x "b" "c" Γ] assms fresh_GNil append_g.simps by auto
lemma wfG_intros2:
assumes "wfC P ℬ ((x,b,TRUE) #⇩ΓΓ) c"
shows "wfG P ℬ ((x,b,c) #⇩ΓΓ)"
proof -
have "wfG P ℬ ((x,b,TRUE) #⇩ΓΓ)" using wfC_wf assms by auto
hence *:"wfG P ℬ Γ ∧ atom x ♯ Γ ∧ wfB P ℬ b" using wfG_elims by metis
show ?thesis using assms proof(cases "c ∈ {TRUE,FALSE}")
case True
then show ?thesis using wfG_cons2I * by auto
next
case False
then show ?thesis using wfG_cons1I * assms by auto
qed
qed
section ‹Type Definitions›
lemma wf_theta_weakening1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ℬ :: ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list and t::τ
shows "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ; Γ ⊢⇩w⇩f v : b" and
"Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ; Γ ⊢⇩w⇩f c" and
"Θ; ℬ ⊢⇩w⇩f Γ ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ⊢⇩w⇩f Γ" and
"Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ; Γ ⊢⇩w⇩f τ" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ; Γ ⊢⇩w⇩f ts" and
"⊢⇩w⇩f P ⟹ True " and
"Θ; ℬ ⊢⇩w⇩f b ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ⊢⇩w⇩f b" and
"Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ; Γ ⊢⇩w⇩f ce : b" and
"Θ ⊢⇩w⇩f td ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ⊢⇩w⇩f td"
proof(nominal_induct b and c and Γ and τ and ts and P and b and b and td
avoiding: Θ'
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfV_consI s dclist Θ dc x b c ℬ Γ v)
show ?case proof
show ‹AF_typedef s dclist ∈ set Θ'› using wfV_consI by auto
show ‹(dc, ⦃ x : b | c ⦄) ∈ set dclist› using wfV_consI by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f v : b › using wfV_consI by auto
qed
next
case (wfV_conspI s bv dclist Θ dc x b' c ℬ b Γ v)
show ?case proof
show ‹AF_typedef_poly s bv dclist ∈ set Θ'› using wfV_conspI by auto
show ‹(dc, ⦃ x : b' | c ⦄) ∈ set dclist› using wfV_conspI by auto
show ‹Θ' ; ℬ ; Γ ⊢⇩w⇩f v : b'[bv::=b]⇩b⇩b › using wfV_conspI by auto
show "Θ' ; ℬ ⊢⇩w⇩f b " using wfV_conspI by auto
show "atom bv ♯ (Θ', ℬ, Γ, b, v)" using wfV_conspI fresh_prodN by auto
qed
next
case (wfTI z Θ ℬ Γ b c)
thus ?case using Wellformed.wfTI by auto
next
case (wfB_consI Θ s dclist)
show ?case proof
show ‹ ⊢⇩w⇩f Θ' › using wfB_consI by auto
show ‹AF_typedef s dclist ∈ set Θ'› using wfB_consI by auto
qed
next
case (wfB_appI Θ ℬ b s bv dclist)
show ?case proof
show ‹ ⊢⇩w⇩f Θ' › using wfB_appI by auto
show ‹AF_typedef_poly s bv dclist ∈ set Θ'› using wfB_appI by auto
show "Θ' ; ℬ ⊢⇩w⇩f b" using wfB_appI by simp
qed
qed(metis wf_intros)+
lemma wf_theta_weakening2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ℬ :: ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list and t::τ
shows
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; Φ ; ℬ ; Γ ; Δ ⊢⇩w⇩f e : b" and
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; Φ ; ℬ ; Γ ; Δ ⊢⇩w⇩f s : b" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; Φ ; ℬ ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; Φ ; ℬ ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ⊢⇩w⇩f (Φ::Φ)" and
"Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; ℬ ; Γ ⊢⇩w⇩f Δ" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; Φ ⊢⇩w⇩f ftq" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ ⊢⇩w⇩f Θ' ⟹ set Θ ⊆ set Θ' ⟹ Θ' ; Φ ; ℬ ⊢⇩w⇩f ft"
proof(nominal_induct b and b and b and b and Φ and Δ and ftq and ft
avoiding: Θ'
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv v τ f x b c s)
show ?case proof
show ‹ Θ' ⊢⇩w⇩f Φ › using wfE_appPI by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f Δ › using wfE_appPI by auto
show ‹ Θ' ; ℬ ⊢⇩w⇩f b' › using wfE_appPI wf_theta_weakening1 by auto
show ‹atom bv ♯ (Φ, Θ', ℬ, Γ, Δ, b', v, (b_of τ)[bv::=b']⇩b)› using wfE_appPI by auto
show ‹Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f› using wfE_appPI by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f v : b[bv::=b']⇩b › using wfE_appPI wf_theta_weakening1 by auto
qed
next
case (wfS_matchI Θ ℬ Γ v tid dclist Δ Φ cs b)
show ?case proof
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f v : B_id tid › using wfS_matchI wf_theta_weakening1 by auto
show ‹AF_typedef tid dclist ∈ set Θ'› using wfS_matchI by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f Δ › using wfS_matchI by auto
show ‹ Θ' ⊢⇩w⇩f Φ › using wfS_matchI by auto
show ‹Θ'; Φ; ℬ; Γ; Δ; tid; dclist ⊢⇩w⇩f cs : b › using wfS_matchI by auto
qed
next
case (wfS_varI Θ ℬ Γ τ v u Φ Δ b s)
show ?case proof
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f τ › using wfS_varI wf_theta_weakening1 by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f v : b_of τ › using wfS_varI wf_theta_weakening1 by auto
show ‹atom u ♯ (Φ, Θ', ℬ, Γ, Δ, τ, v, b)› using wfS_varI by auto
show ‹ Θ' ; Φ ; ℬ ; Γ ; (u, τ) #⇩Δ Δ ⊢⇩w⇩f s : b › using wfS_varI by auto
qed
next
case (wfS_letI Θ Φ ℬ Γ Δ e b' x s b)
show ?case proof
show ‹ Θ' ; Φ ; ℬ ; Γ ; Δ ⊢⇩w⇩f e : b' › using wfS_letI by auto
show ‹ Θ' ; Φ ; ℬ ; (x, b', TRUE) #⇩Γ Γ ; Δ ⊢⇩w⇩f s : b › using wfS_letI by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f Δ › using wfS_letI by auto
show ‹atom x ♯ (Φ, Θ', ℬ, Γ, Δ, e, b)› using wfS_letI by auto
qed
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 τ x s2 b)
show ?case proof
show ‹ Θ' ; Φ ; ℬ ; Γ ; Δ ⊢⇩w⇩f s1 : b_of τ › using wfS_let2I by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f τ › using wfS_let2I wf_theta_weakening1 by auto
show ‹ Θ' ; Φ ; ℬ ; (x, b_of τ, TRUE) #⇩Γ Γ ; Δ ⊢⇩w⇩f s2 : b › using wfS_let2I by auto
show ‹atom x ♯ (Φ, Θ', ℬ, Γ, Δ, s1, b, τ)› using wfS_let2I by auto
qed
next
case (wfS_branchI Θ Φ ℬ x τ Γ Δ s b tid dc)
show ?case proof
show ‹ Θ' ; Φ ; ℬ ; (x, b_of τ, TRUE) #⇩Γ Γ ; Δ ⊢⇩w⇩f s : b › using wfS_branchI by auto
show ‹atom x ♯ (Φ, Θ', ℬ, Γ, Δ, Γ, τ)› using wfS_branchI by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f Δ › using wfS_branchI by auto
qed
next
case (wfPhi_consI f Φ Θ ft)
show ?case proof
show "f ∉ name_of_fun ` set Φ" using wfPhi_consI by auto
show "Θ' ; Φ ⊢⇩w⇩f ft" using wfPhi_consI by auto
show "Θ' ⊢⇩w⇩f Φ" using wfPhi_consI by auto
qed
next
case (wfFTNone Θ ft)
then show ?case using wf_intros by metis
next
case (wfFTSome Θ bv ft)
then show ?case using wf_intros by metis
next
case (wfFTI Θ B b Φ x c s τ)
thus ?case using Wellformed.wfFTI wf_theta_weakening1 by simp
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
show ?case proof
show ‹ Θ' ; Φ ; ℬ ; (x, B_bool, c) #⇩Γ Γ ; Δ ⊢⇩w⇩f s : b › using wfS_assertI wf_theta_weakening1 by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f c › using wfS_assertI wf_theta_weakening1 by auto
show ‹ Θ' ; ℬ ; Γ ⊢⇩w⇩f Δ › using wfS_assertI wf_theta_weakening1 by auto
have "atom x ♯ Θ'" using wf_supp(6)[OF ‹⊢⇩w⇩f Θ' ›] fresh_def by auto
thus ‹atom x ♯ (Φ, Θ', ℬ, Γ, Δ, c, b, s)› using wfS_assertI fresh_prodN fresh_def by simp
qed
qed(metis wf_intros wf_theta_weakening1 )+
lemmas wf_theta_weakening = wf_theta_weakening1 wf_theta_weakening2
lemma lookup_wfTD:
fixes td::type_def
assumes "td ∈ set Θ" and "⊢⇩w⇩f Θ"
shows "Θ ⊢⇩w⇩f td"
using assms proof(induct Θ )
case Nil
then show ?case by auto
next
case (Cons td' Θ')
then consider "td = td'" | "td ∈ set Θ'" by auto
then have "Θ' ⊢⇩w⇩f td" proof(cases)
case 1
then show ?thesis using Cons using wfTh_elims by auto
next
case 2
then show ?thesis using Cons using wfTh_elims by auto
qed
then show ?case using wf_theta_weakening Cons by (meson set_subset_Cons)
qed
subsection ‹Simple›
lemma wfTh_dclist_unique:
assumes "wfTh Θ" and "AF_typedef tid dclist1 ∈ set Θ" and "AF_typedef tid dclist2 ∈ set Θ"
shows "dclist1 = dclist2"
using assms proof(induct Θ rule: Θ_induct)
case TNil
then show ?case by auto
next
case (AF_typedef tid' dclist' Θ')
then show ?case using wfTh_elims
by (metis image_eqI name_of_type.simps(1) set_ConsD type_def.eq_iff(1))
next
case (AF_typedef_poly tid bv dclist Θ')
then show ?case using wfTh_elims by auto
qed
lemma wfTs_ctor_unique:
fixes dclist::"(string*τ) list"
assumes "Θ ; {||} ; GNil ⊢⇩w⇩f dclist" and "(c, t1) ∈ set dclist" and "(c,t2) ∈ set dclist"
shows "t1 = t2"
using assms proof(induct dclist rule: list.inducts)
case Nil
then show ?case by auto
next
case (Cons x1 x2)
consider "x1 = (c,t1)" | "x1 = (c,t2)" | "x1 ≠ (c,t1) ∧ x1 ≠ (c,t2)" by auto
thus ?case proof(cases)
case 1
then show ?thesis using Cons wfTs_elims set_ConsD
by (metis fst_conv image_eqI prod.inject)
next
case 2
then show ?thesis using Cons wfTs_elims set_ConsD
by (metis fst_conv image_eqI prod.inject)
next
case 3
then show ?thesis using Cons wfTs_elims by (metis set_ConsD)
qed
qed
lemma wfTD_ctor_unique:
assumes "Θ ⊢⇩w⇩f (AF_typedef tid dclist)" and "(c, t1) ∈ set dclist" and "(c,t2) ∈ set dclist"
shows "t1 = t2"
using wfTD_elims wfTs_elims assms wfTs_ctor_unique by metis
lemma wfTh_ctor_unique:
assumes "wfTh Θ" and "AF_typedef tid dclist ∈ set Θ" and "(c, t1) ∈ set dclist" and "(c,t2) ∈ set dclist"
shows "t1 = t2"
using lookup_wfTD wfTD_ctor_unique assms by metis
lemma wfTs_supp_t:
fixes dclist::"(string*τ) list"
assumes "(c,t) ∈ set dclist" and "Θ ; B ; GNil ⊢⇩w⇩f dclist"
shows "supp t ⊆ supp B"
using assms proof(induct dclist arbitrary: c t rule:list.induct)
case Nil
then show ?case by auto
next
case (Cons ct dclist')
then consider "ct = (c,t)" | "(c,t) ∈ set dclist'" by auto
then show ?case proof(cases)
case 1
then have "Θ ; B ; GNil ⊢⇩w⇩f t" using Cons wfTs_elims by blast
thus ?thesis using wfT_supp atom_dom.simps by force
next
case 2
then show ?thesis using Cons wfTs_elims by metis
qed
qed
lemma wfTh_lookup_supp_empty:
fixes t::τ
assumes "AF_typedef tid dclist ∈ set Θ" and "(c,t) ∈ set dclist" and "⊢⇩w⇩f Θ"
shows "supp t = {}"
proof -
have "Θ ; {||} ; GNil ⊢⇩w⇩f dclist" using assms lookup_wfTD wfTD_elims by metis
thus ?thesis using wfTs_supp_t assms by force
qed
lemma wfTh_supp_b:
assumes "AF_typedef tid dclist ∈ set Θ" and "(dc,⦃ z : b | c ⦄ ) ∈ set dclist" and "⊢⇩w⇩f Θ"
shows "supp b = {}"
using assms wfTh_lookup_supp_empty τ.supp by blast
lemma wfTh_b_eq_iff:
fixes bva1::bv and bva2::bv and dc::string
assumes "(dc, ⦃ x1 : b1 | c1 ⦄) ∈ set dclist1" and "(dc, ⦃ x2 : b2 | c2 ⦄) ∈ set dclist2" and
"wfTs P {|bva1|} GNil dclist1" and "wfTs P {|bva2|} GNil dclist2"
"[[atom bva1]]lst.dclist1 = [[atom bva2]]lst.dclist2"
shows "[[atom bva1]]lst. (dc,⦃ x1 : b1 | c1 ⦄) = [[atom bva2]]lst. (dc,⦃ x2 : b2 | c2 ⦄)"
using assms proof(induct dclist1 arbitrary: dclist2)
case Nil
then show ?case by auto
next
case (Cons dct1' dclist1')
show ?case proof(cases "dclist2 = []")
case True
then show ?thesis using Cons by auto
next
case False
then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using list.exhaust by metis
hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' ∧ [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'"
using Cons lst_head_cons Cons cons by metis
hence **: "fst dct1' = fst dct2'" using lst_fst[THEN lst_pure]
by (metis (no_types) ‹[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' ∧ [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'›
‹⋀x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') ⟹ t1 = t2› fst_conv surj_pair)
show ?thesis proof(cases "fst dct1' = dc")
case True
have "dc ∉ fst ` set dclist1'" using wfTs_elims Cons by (metis True fstI)
hence 1:"(dc, ⦃ x1 : b1 | c1 ⦄) = dct1'" using Cons by (metis fstI image_iff set_ConsD)
have "dc ∉ fst ` set dclist2'" using wfTs_elims Cons cons
by (metis "**" True fstI)
hence 2:"(dc, ⦃ x2 : b2 | c2 ⦄) = dct2' " using Cons cons by (metis fst_conv image_eqI set_ConsD)
then show ?thesis using Cons * 1 2 by blast
next
case False
hence "fst dct2' ≠ dc" using ** by auto
hence "(dc, ⦃ x1 : b1 | c1 ⦄) ∈ set dclist1' ∧ (dc, ⦃ x2 : b2 | c2 ⦄) ∈ set dclist2' " using Cons cons False
by (metis fstI set_ConsD)
moreover have "[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2'" using * False by metis
ultimately show ?thesis using Cons ** *
using cons wfTs_elims(2) by blast
qed
qed
qed
subsection ‹Polymorphic›
lemma wfTh_wfTs_poly:
fixes dclist::"(string * τ) list"
assumes "AF_typedef_poly tyid bva dclist ∈ set P" and "⊢⇩w⇩f P"
shows "P ; {|bva|} ; GNil ⊢⇩w⇩f dclist"
proof -
have *:"P ⊢⇩w⇩f AF_typedef_poly tyid bva dclist" using lookup_wfTD assms by simp
obtain bv lst where *:"P ; {|bv|} ; GNil ⊢⇩w⇩f lst ∧
(∀c. atom c ♯ (dclist, lst) ⟶ atom c ♯ (bva, bv, dclist, lst) ⟶ (bva ↔ c) ∙ dclist = (bv ↔ c) ∙ lst)"
using wfTD_elims(2)[OF *] by metis
obtain c::bv where **:"atom c ♯ ((dclist, lst),(bva, bv, dclist, lst))" using obtain_fresh by metis
have "P ; {|bv|} ; GNil ⊢⇩w⇩f lst" using * by metis
hence "wfTs ((bv ↔ c) ∙ P) ((bv ↔ c) ∙ {|bv|}) ((bv ↔ c) ∙ GNil) ((bv ↔ c) ∙ lst)" using ** wfTs.eqvt by metis
hence "wfTs P {|c|} GNil ((bva ↔ c) ∙ dclist)" using * theta_flip_eq fresh_GNil assms
proof -
have "∀b ba. (ba::bv ↔ b) ∙ P = P" by (metis ‹⊢⇩w⇩f P› theta_flip_eq)
then show ?thesis
using "*" "**" ‹(bv ↔ c) ∙ P ; (bv ↔ c) ∙ {|bv|} ; (bv ↔ c) ∙ GNil ⊢⇩w⇩f (bv ↔ c) ∙ lst› by fastforce
qed
hence "wfTs ((bva ↔ c) ∙ P) ((bva ↔ c) ∙ {|bva|}) ((bva ↔ c) ∙ GNil) ((bva ↔ c) ∙ dclist)"
using wfTs.eqvt fresh_GNil
by (simp add: assms(2) theta_flip_eq2)
thus ?thesis using wfTs.eqvt permute_flip_cancel by metis
qed
lemma wfTh_dclist_poly_unique:
assumes "wfTh Θ" and "AF_typedef_poly tid bva dclist1 ∈ set Θ" and "AF_typedef_poly tid bva2 dclist2 ∈ set Θ"
shows "[[atom bva]]lst. dclist1 = [[atom bva2]]lst.dclist2"
using assms proof(induct Θ rule: Θ_induct)
case TNil
then show ?case by auto
next
case (AF_typedef tid' dclist' Θ')
then show ?case using wfTh_elims by auto
next
case (AF_typedef_poly tid bv dclist Θ')
then show ?case using wfTh_elims image_eqI name_of_type.simps set_ConsD type_def.eq_iff
by (metis Abs1_eq(3))
qed
lemma wfTh_poly_lookup_supp:
fixes t::τ
assumes "AF_typedef_poly tid bv dclist ∈ set Θ" and "(c,t) ∈ set dclist" and "⊢⇩w⇩f Θ"
shows "supp t ⊆ {atom bv}"
proof -
have "supp dclist ⊆ {atom bv}" using assms lookup_wfTD wf_supp1 type_def.supp
by (metis Diff_single_insert Un_subset_iff list.simps(15) supp_Nil supp_of_atom_list)
then show ?thesis using assms(2) proof(induct dclist)
case Nil
then show ?case by auto
next
case (Cons a dclist)
then show ?case using supp_Pair supp_Cons
by (metis (mono_tags, opaque_lifting) Un_empty_left Un_empty_right pure_supp subset_Un_eq subset_singletonD supp_list_member)
qed
qed
lemma wfTh_poly_supp_b:
assumes "AF_typedef_poly tid bv dclist ∈ set Θ" and "(dc,⦃ z : b | c ⦄ ) ∈ set dclist" and "⊢⇩w⇩f Θ"
shows "supp b ⊆ {atom bv}"
using assms wfTh_poly_lookup_supp τ.supp by force
lemma subst_g_inside:
fixes x::x and c::c and Γ::Γ and Γ'::Γ
assumes "wfG P ℬ (Γ' @ (x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ Γ)"
shows "(Γ' @ (x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ Γ)[x::=v]⇩Γ⇩v = (Γ'[x::=v]⇩Γ⇩v@Γ)"
using assms proof(induct Γ' rule: Γ_induct)
case GNil
then show ?case using subst_gb.simps by simp
next
case (GCons x' b' c' G)
hence wfg:"wfG P ℬ (G @ (x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ Γ) ∧ atom x' ♯ (G @ (x, b, c[z::=V_var x]⇩c⇩v) #⇩Γ Γ)" using wfG_elims(2)
using GCons.prems append_g.simps by metis
hence "atom x ∉ atom_dom ((x', b', c') #⇩Γ G)" using GCons wfG_inside_fresh by fast
hence "x≠x'"
using GCons append_Cons wfG_inside_fresh atom_dom.simps toSet.simps by simp
hence "((GCons (x', b', c') G) @ (GCons (x, b, c[z::=V_var x]⇩c⇩v) Γ))[x::=v]⇩Γ⇩v =
(GCons (x', b', c') (G @ (GCons (x, b, c[z::=V_var x]⇩c⇩v) Γ)))[x::=v]⇩Γ⇩v" by auto
also have "... = GCons (x', b', c'[x::=v]⇩c⇩v) ((G @ (GCons (x, b, c[z::=V_var x]⇩c⇩v) Γ))[x::=v]⇩Γ⇩v)"
using subst_gv.simps ‹x≠x'› by simp
also have "... = (x', b', c'[x::=v]⇩c⇩v) #⇩Γ (G[x::=v]⇩Γ⇩v @ Γ)" using GCons wfg by blast
also have "... = ((x', b', c') #⇩Γ G)[x::=v]⇩Γ⇩v @ Γ" using subst_gv.simps ‹x≠x'› by simp
finally show ?case by auto
qed
lemma wfTh_td_eq:
assumes "td1 ∈ set (td2 # P)" and "wfTh (td2 # P)" and "name_of_type td1 = name_of_type td2"
shows "td1 = td2"
proof(rule ccontr)
assume as: "td1 ≠ td2"
have "name_of_type td2 ∉ name_of_type ` set P" using wfTh_elims(2)[OF assms(2)] by metis
moreover have "td1 ∈ set P" using assms as by simp
ultimately have "name_of_type td1 ≠ name_of_type td2"
by (metis rev_image_eqI)
thus False using assms by auto
qed
lemma wfTh_td_unique:
assumes "td1 ∈ set P" and "td2 ∈ set P" and "wfTh P" and "name_of_type td1 = name_of_type td2"
shows "td1 = td2"
using assms proof(induct P rule: list.induct)
case Nil
then show ?case by auto
next
case (Cons td Θ')
consider "td = td1" | "td = td2" | "td ≠ td1 ∧ td ≠ td2" by auto
then show ?case proof(cases)
case 1
then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis
next
case 2
then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis
next
case 3
then show ?thesis using Cons wfTh_elims by auto
qed
qed
lemma wfTs_distinct:
fixes dclist::"(string * τ) list"
assumes "Θ ; B ; GNil ⊢⇩w⇩f dclist"
shows "distinct (map fst dclist)"
using assms proof(induct dclist rule: list.induct)
case Nil
then show ?case by auto
next
case (Cons x1 x2)
then show ?case
by (metis Cons.hyps Cons.prems distinct.simps(2) fst_conv list.set_map list.simps(9) wfTs_elims(2))
qed
lemma wfTh_dclist_distinct:
assumes "AF_typedef s dclist ∈ set P" and "wfTh P"
shows "distinct (map fst dclist)"
proof -
have "wfTD P (AF_typedef s dclist)" using assms lookup_wfTD by auto
hence "wfTs P {||} GNil dclist" using wfTD_elims by metis
thus ?thesis using wfTs_distinct by metis
qed
lemma wfTh_dc_t_unique2:
assumes "AF_typedef s dclist' ∈ set P" and "(dc,tc' ) ∈ set dclist'" and "AF_typedef s dclist ∈ set P" and "wfTh P" and
"(dc, tc) ∈ set dclist"
shows "tc= tc'"
proof -
have "dclist = dclist'" using assms wfTh_td_unique name_of_type.simps by force
moreover have "distinct (map fst dclist)" using wfTh_dclist_distinct assms by auto
ultimately show ?thesis using assms
by (meson eq_key_imp_eq_value)
qed
lemma wfTh_dc_t_unique:
assumes "AF_typedef s dclist' ∈ set P" and "(dc, ⦃ x' : b' | c' ⦄ ) ∈ set dclist'" and "AF_typedef s dclist ∈ set P" and "wfTh P" and
"(dc, ⦃ x : b | c ⦄) ∈ set dclist"
shows "⦃ x' : b' | c' ⦄= ⦃ x : b | c ⦄"
using assms wfTh_dc_t_unique2 by metis
lemma wfTs_wfT:
fixes dclist::"(string *τ) list" and t::τ
assumes "Θ; ℬ; GNil ⊢⇩w⇩f dclist" and "(dc,t) ∈ set dclist"
shows "Θ; ℬ; GNil ⊢⇩w⇩f t"
using assms proof(induct dclist rule:list.induct)
case Nil
then show ?case by auto
next
case (Cons x1 x2)
thus ?case using wfTs_elims(2)[OF Cons(2)] by auto
qed
lemma wfTh_wfT:
fixes t::τ
assumes "wfTh P" and "AF_typedef tid dclist ∈ set P" and "(dc,t) ∈ set dclist"
shows "P ; {||} ; GNil ⊢⇩w⇩f t"
proof -
have "P ⊢⇩w⇩f AF_typedef tid dclist" using lookup_wfTD assms by auto
hence "P ; {||} ; GNil ⊢⇩w⇩f dclist" using wfTD_elims by auto
thus ?thesis using wfTs_wfT assms by auto
qed
lemma td_lookup_eq_iff:
fixes dc :: string and bva1::bv and bva2::bv
assumes "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst. dclist2" and "(dc, ⦃ x : b | c ⦄) ∈ set dclist1"
shows "∃x2 b2 c2. (dc, ⦃ x2 : b2 | c2 ⦄) ∈ set dclist2"
using assms proof(induct dclist1 arbitrary: dclist2)
case Nil
then show ?case by auto
next
case (Cons dct1' dclist1')
then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using lst_head_cons_neq_nil[OF Cons(2)] list.exhaust by metis
hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' ∧ [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'"
using Cons lst_head_cons Cons cons by metis
show ?case proof(cases "dc=fst dct1'")
case True
hence "dc = fst dct2'" using * lst_fst[ THEN lst_pure ]
proof -
show ?thesis
by (metis (no_types) "local.*" True ‹⋀x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') ⟹ t1 = t2› prod.exhaust_sel)
qed
obtain x2 b2 and c2 where "snd dct2' = ⦃ x2 : b2 | c2 ⦄" using obtain_fresh_z by metis
hence "(dc, ⦃ x2 : b2 | c2 ⦄) = dct2'" using ‹dc = fst dct2'›
by (metis prod.exhaust_sel)
then show ?thesis using cons by force
next
case False
hence "(dc, ⦃ x : b | c ⦄) ∈ set dclist1'" using Cons by auto
then show ?thesis using Cons
by (metis "local.*" cons list.set_intros(2))
qed
qed
lemma lst_t_b_eq_iff:
fixes bva1::bv and bva2::bv
assumes "[[atom bva1]]lst. ⦃ x1 : b1 | c1 ⦄ = [[atom bva2]]lst. ⦃ x2 : b2 | c2 ⦄"
shows "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2"
proof(subst Abs1_eq_iff_all(3)[of bva1 b1 bva2 b2],rule,rule,rule)
fix c::bv
assume "atom c ♯ ( ⦃ x1 : b1 | c1 ⦄ , ⦃ x2 : b2 | c2 ⦄)" and "atom c ♯ (bva1, bva2, b1, b2)"
show "(bva1 ↔ c) ∙ b1 = (bva2 ↔ c) ∙ b2" using assms Abs1_eq_iff(3) assms
by (metis Abs1_eq_iff_fresh(3) ‹atom c ♯ (bva1, bva2, b1, b2)› τ.fresh τ.perm_simps type_eq_subst_eq2(2))
qed
lemma wfTh_typedef_poly_b_eq_iff:
assumes "AF_typedef_poly tyid bva1 dclist1 ∈ set P" and "(dc, ⦃ x1 : b1 | c1 ⦄) ∈ set dclist1"
and "AF_typedef_poly tyid bva2 dclist2 ∈ set P" and "(dc, ⦃ x2 : b2 | c2 ⦄) ∈ set dclist2" and "⊢⇩w⇩f P"
shows "b1[bva1::=b]⇩b⇩b = b2[bva2::=b]⇩b⇩b"
proof -
have "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst.dclist2" using assms wfTh_dclist_poly_unique by metis
hence "[[atom bva1]]lst. (dc,⦃ x1 : b1 | c1 ⦄) = [[atom bva2]]lst. (dc,⦃ x2 : b2 | c2 ⦄)" using wfTh_b_eq_iff assms wfTh_wfTs_poly by metis
hence "[[atom bva1]]lst. ⦃ x1 : b1 | c1 ⦄ = [[atom bva2]]lst. ⦃ x2 : b2 | c2 ⦄" using lst_snd by metis
hence "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2" using lst_t_b_eq_iff by metis
thus ?thesis using subst_b_flip_eq_two subst_b_b_def by metis
qed
section ‹Equivariance Lemmas›
lemma x_not_in_u_set[simp]:
fixes x::x and us::"u fset"
shows "atom x ∉ supp us"
by(induct us,auto, simp add: supp_finsert supp_at_base)
lemma wfS_flip_eq:
fixes s1::s and x1::x and s2::s and x2::x and Δ::Δ
assumes "[[atom x1]]lst. s1 = [[atom x2]]lst. s2" and "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "[[atom x1]]lst. c1 = [[atom x2]]lst. c2" and "atom x2 ♯ Γ" and
" Θ; ℬ; Γ ⊢⇩w⇩f Δ" and
"Θ ; Φ ; ℬ ; (x1, b, c1) #⇩Γ Γ ; Δ ⊢⇩w⇩f s1 : b_of t1"
shows "Θ ; Φ ; ℬ ; (x2, b, c2) #⇩Γ Γ ; Δ ⊢⇩w⇩f s2 : b_of t2"
proof(cases "x1=x2")
case True
hence "s1 = s2 ∧ t1 = t2 ∧ c1 = c2" using assms Abs1_eq_iff by metis
then show ?thesis using assms True by simp
next
case False
have "⊢⇩w⇩f Θ ∧ Θ ⊢⇩w⇩f Φ ∧ Θ; ℬ; Γ ⊢⇩w⇩f Δ" using wfX_wfY assms by metis
moreover have "atom x1 ♯ Γ" using wfX_wfY wfG_elims assms by metis
moreover hence "atom x1 ♯ Δ ∧ atom x2 ♯ Δ " using wfD_x_fresh assms by auto
ultimately have " Θ ; Φ ; ℬ ; (x2 ↔ x1) ∙ ((x1, b, c1) #⇩Γ Γ) ; Δ ⊢⇩w⇩f (x2 ↔ x1) ∙ s1 : (x2 ↔ x1) ∙ b_of t1"
using wfS.eqvt theta_flip_eq phi_flip_eq assms flip_base_eq beta_flip_eq flip_fresh_fresh supp_b_empty by metis
hence " Θ ; Φ ; ℬ ; ((x2, b, (x2 ↔ x1) ∙ c1) #⇩Γ ((x2 ↔ x1) ∙ Γ)) ; Δ ⊢⇩w⇩f (x2 ↔ x1) ∙ s1 : b_of ((x2 ↔ x2) ∙ t1)" by fastforce
thus ?thesis using assms Abs1_eq_iff
proof -
have f1: "x2 = x1 ∧ t2 = t1 ∨ x2 ≠ x1 ∧ t2 = (x2 ↔ x1) ∙ t1 ∧ atom x2 ♯ t1"
by (metis (full_types) Abs1_eq_iff(3) ‹[[atom x1]]lst. t1 = [[atom x2]]lst. t2›)
then have "x2 ≠ x1 ∧ s2 = (x2 ↔ x1) ∙ s1 ∧ atom x2 ♯ s1 ⟶ b_of t2 = (x2 ↔ x1) ∙ b_of t1"
by (metis b_of.eqvt)
then show ?thesis
using f1 by (metis (no_types) Abs1_eq_iff(3) G_cons_flip_fresh3 ‹[[atom x1]]lst. c1 = [[atom x2]]lst. c2› ‹[[atom x1]]lst. s1 = [[atom x2]]lst. s2› ‹Θ ; Φ ; ℬ ; (x1, b, c1) #⇩Γ Γ ; Δ ⊢⇩w⇩f s1 : b_of t1› ‹Θ ; Φ ; ℬ ; (x2 ↔ x1) ∙ ((x1, b, c1) #⇩Γ Γ) ; Δ ⊢⇩w⇩f (x2 ↔ x1) ∙ s1 : (x2 ↔ x1) ∙ b_of t1› ‹atom x1 ♯ Γ› ‹atom x2 ♯ Γ›)
qed
qed
section ‹Lookup›
lemma wf_not_in_prefix:
assumes "Θ ; B ⊢⇩w⇩f (Γ'@(x,b1,c1) #⇩ΓΓ)"
shows "x ∉ fst ` toSet Γ'"
using assms proof(induct Γ' rule: Γ.induct)
case GNil
then show ?case by simp
next
case (GCons xbc Γ')
then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
using prod_cases3 by blast
hence *:"(xbc #⇩Γ Γ') @ (x, b1, c1) #⇩Γ Γ = ((x',b',c') #⇩Γ(Γ'@ ((x, b1, c1) #⇩Γ Γ)))" by simp
hence "atom x' ♯ (Γ'@(x,b1,c1) #⇩ΓΓ)" using wfG_elims(2) GCons by metis
moreover have "Θ ; B ⊢⇩w⇩f (Γ' @ (x, b1, c1) #⇩Γ Γ)" using GCons wfG_elims * by metis
ultimately have "atom x' ∉ atom_dom (Γ'@(x,b1,c1) #⇩ΓΓ)" using wfG_dom_supp GCons append_g.simps xbc fresh_def by fast
hence "x' ≠ x" using GCons fresh_GCons xbc by fastforce
then show ?case using GCons xbc toSet.simps
using Un_commute ‹Θ ; B ⊢⇩w⇩f Γ' @ (x, b1, c1) #⇩Γ Γ› atom_dom.simps by auto
qed
lemma lookup_inside_wf[simp]:
assumes "Θ ; B ⊢⇩w⇩f (Γ'@(x,b1,c1) #⇩ΓΓ)"
shows "Some (b1,c1) = lookup (Γ'@(x,b1,c1) #⇩ΓΓ) x"
using wf_not_in_prefix lookup_inside assms by fast
lemma lookup_weakening:
fixes Θ::Θ and Γ::Γ and Γ'::Γ
assumes "Some (b,c) = lookup Γ x" and "toSet Γ ⊆ toSet Γ'" and "Θ; ℬ ⊢⇩w⇩f Γ'" and "Θ; ℬ ⊢⇩w⇩f Γ"
shows "Some (b,c) = lookup Γ' x"
proof -
have "(x,b,c) ∈ toSet Γ ∧ (∀b' c'. (x,b',c') ∈ toSet Γ ⟶ b'=b ∧ c'=c)" using assms lookup_iff toSet.simps by force
hence "(x,b,c) ∈ toSet Γ'" using assms by auto
moreover have "(∀b' c'. (x,b',c') ∈ toSet Γ' ⟶ b'=b ∧ c'=c)" using assms wf_g_unique
using calculation by auto
ultimately show ?thesis using lookup_iff
using assms(3) by blast
qed
lemma wfPhi_lookup_fun_unique:
fixes Φ::Φ
assumes "Θ ⊢⇩w⇩f Φ" and "AF_fundef f fd ∈ set Φ"
shows "Some (AF_fundef f fd) = lookup_fun Φ f"
using assms proof(induct Φ rule: list.induct )
case Nil
then show ?case using lookup_fun.simps by simp
next
case (Cons a Φ')
then obtain f' and fd' where a:"a = AF_fundef f' fd'" using fun_def.exhaust by auto
have wf: "Θ ⊢⇩w⇩f Φ' ∧ f' ∉ name_of_fun ` set Φ' " using wfPhi_elims Cons a by metis
then show ?case using Cons lookup_fun.simps using Cons lookup_fun.simps wf a
by (metis image_eqI name_of_fun.simps set_ConsD)
qed
lemma lookup_fun_weakening:
fixes Φ'::Φ
assumes "Some fd = lookup_fun Φ f" and "set Φ ⊆ set Φ'" and "Θ ⊢⇩w⇩f Φ'"
shows "Some fd = lookup_fun Φ' f"
using assms proof(induct Φ )
case Nil
then show ?case using lookup_fun.simps by simp
next
case (Cons a Φ'')
then obtain f' and fd' where a: "a = AF_fundef f' fd'" using fun_def.exhaust by auto
then show ?case proof(cases "f=f'")
case True
then show ?thesis using lookup_fun.simps Cons wfPhi_lookup_fun_unique a
by (metis lookup_fun_member subset_iff)
next
case False
then show ?thesis using lookup_fun.simps Cons
using ‹a = AF_fundef f' fd'› by auto
qed
qed
lemma fundef_poly_fresh_bv:
assumes "atom bv2 ♯ (bv1,b1,c1,τ1,s1)"
shows * : "(AF_fun_typ_some bv2 (AF_fun_typ x1 ((bv1↔bv2) ∙ b1) ((bv1↔bv2) ∙c1) ((bv1↔bv2) ∙ τ1) ((bv1↔bv2) ∙ s1)) = (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1 s1)))"
(is "(AF_fun_typ_some ?bv ?fun_typ = AF_fun_typ_some ?bva ?fun_typa)")
proof -
have 1:"atom bv2 ∉ set [atom x1]" using bv_not_in_x_atoms by simp
have 2:"bv1 ≠ bv2" using assms by auto
have 3:"(bv2 ↔ bv1) ∙ x1 = x1" using pure_fresh flip_fresh_fresh
by (simp add: flip_fresh_fresh)
have " AF_fun_typ x1 ((bv1 ↔ bv2) ∙ b1) ((bv1 ↔ bv2) ∙ c1) ((bv1 ↔ bv2) ∙ τ1) ((bv1 ↔ bv2) ∙ s1) = (bv2 ↔ bv1) ∙ AF_fun_typ x1 b1 c1 τ1 s1"
using 1 2 3 assms by (simp add: flip_commute)
moreover have "(atom bv2 ♯ c1 ∧ atom bv2 ♯ τ1 ∧ atom bv2 ♯ s1 ∨ atom bv2 ∈ set [atom x1]) ∧ atom bv2 ♯ b1"
using 1 2 3 assms fresh_prod5 by metis
ultimately show ?thesis unfolding fun_typ_q.eq_iff Abs1_eq_iff(3) fun_typ.fresh 1 2 by fastforce
qed
text ‹It is possible to collapse some of the easy to prove inductive cases into a single proof at the qed line
but this makes it fragile under change. For example, changing the lemma statement might make one of the previously
trivial cases non-trivial and so the collapsing needs to be unpacked. Is there a way to find which case
has failed in the qed line?›
lemma wb_b_weakening1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and ℬ::ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ ℬ |⊆| ℬ' ⟹ Θ; ℬ' ; Γ ⊢⇩w⇩f v : b" and
"Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ℬ |⊆| ℬ' ⟹ Θ; ℬ' ; Γ ⊢⇩w⇩f c" and
"Θ; ℬ ⊢⇩w⇩f Γ ⟹ℬ |⊆| ℬ' ⟹ Θ; ℬ' ⊢⇩w⇩f Γ " and
"Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ ℬ |⊆| ℬ' ⟹ Θ; ℬ' ; Γ ⊢⇩w⇩f τ" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ ℬ |⊆| ℬ' ⟹ Θ; ℬ' ; Γ ⊢⇩w⇩f ts" and
"⊢⇩w⇩f P ⟹ True " and
"wfB Θ ℬ b ⟹ ℬ |⊆| ℬ' ⟹ wfB Θ ℬ' b" and
"Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ ℬ |⊆| ℬ' ⟹ Θ; ℬ' ; Γ ⊢⇩w⇩f ce : b" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(nominal_induct b and c and Γ and τ and ts and P and b and b and td
avoiding: ℬ'
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfV_conspI s bv dclist Θ dc x b' c ℬ b Γ v)
show ?case proof
show ‹AF_typedef_poly s bv dclist ∈ set Θ› using wfV_conspI by metis
show ‹(dc, ⦃ x : b' | c ⦄) ∈ set dclist› using wfV_conspI by auto
show ‹ Θ ; ℬ' ⊢⇩w⇩f b › using wfV_conspI by auto
show ‹atom bv ♯ (Θ, ℬ', Γ, b, v)› using fresh_prodN wfV_conspI by auto
thus ‹ Θ; ℬ' ; Γ ⊢⇩w⇩f v : b'[bv::=b]⇩b⇩b › using wfV_conspI by simp
qed
next
case (wfTI z Θ ℬ Γ b c)
show ?case proof
show "atom z ♯ (Θ, ℬ', Γ)" using wfTI by auto
show "Θ; ℬ' ⊢⇩w⇩f b " using wfTI by auto
show "Θ; ℬ' ; (z, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c " using wfTI by auto
qed
qed( (auto simp add: wf_intros | metis wf_intros)+ )
lemma wb_b_weakening2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and ℬ::ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b ⟹ ℬ |⊆| ℬ' ⟹ Θ ; Φ ; ℬ' ; Γ ; Δ ⊢⇩w⇩f e : b" and
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ ℬ |⊆| ℬ' ⟹ Θ ; Φ ; ℬ' ; Γ ; Δ ⊢⇩w⇩f s : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ ℬ |⊆| ℬ' ⟹ Θ ; Φ ; ℬ' ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ ℬ |⊆| ℬ' ⟹ Θ ; Φ ; ℬ' ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ ℬ |⊆| ℬ' ⟹ Θ; ℬ' ; Γ ⊢⇩w⇩f Δ" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ ℬ |⊆| ℬ' ⟹ Θ ; Φ ; ℬ' ⊢⇩w⇩f ft"
proof(nominal_induct b and b and b and b and Φ and Δ and ftq and ft
avoiding: ℬ'
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_valI Θ Φ ℬ Γ Δ v b)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfE_plusI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfE_leqI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfE_eqI Θ Φ ℬ Γ Δ v1 b v2)
then show ?case using wf_intros wb_b_weakening1
by meson
next
case (wfE_fstI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using Wellformed.wfE_fstI wb_b_weakening1 by metis
next
case (wfE_sndI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfE_concatI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfE_splitI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfE_lenI Θ Φ ℬ Γ Δ v1)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfE_appI Θ Φ ℬ Γ Δ f ft v)
then show ?case using wf_intros using wb_b_weakening1 by meson
next
case (wfE_appPI Θ Φ ℬ1 Γ Δ b' bv1 v1 τ1 f1 x1 b1 c1 s1)
have "Θ ; Φ ; ℬ' ; Γ ; Δ ⊢⇩w⇩f AE_appP f1 b' v1 : (b_of τ1)[bv1::=b']⇩b"
proof
show "Θ ⊢⇩w⇩f Φ" using wfE_appPI by auto
show "Θ; ℬ' ; Γ ⊢⇩w⇩f Δ " using wfE_appPI by auto
show "Θ; ℬ' ⊢⇩w⇩f b' " using wfE_appPI wb_b_weakening1 by auto
thus " atom bv1 ♯ (Φ, Θ, ℬ', Γ, Δ, b', v1, (b_of τ1)[bv1::=b']⇩b)"
using wfE_appPI fresh_prodN by auto
show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1 s1))) = lookup_fun Φ f1" using wfE_appPI by auto
show "Θ; ℬ' ; Γ ⊢⇩w⇩f v1 : b1[bv1::=b']⇩b " using wfE_appPI wb_b_weakening1 by auto
qed
then show ?case by auto
next
case (wfE_mvarI Θ Φ ℬ Γ Δ u τ)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfS_valI Θ Φ ℬ Γ v b Δ)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfS_letI Θ Φ ℬ Γ Δ e b' x s b)
show ?case proof
show ‹ Θ ; Φ ; ℬ' ; Γ ; Δ ⊢⇩w⇩f e : b' › using wfS_letI by auto
show ‹ Θ ; Φ ; ℬ' ; (x, b', TRUE) #⇩Γ Γ ; Δ ⊢⇩w⇩f s : b › using wfS_letI by auto
show ‹ Θ; ℬ' ; Γ ⊢⇩w⇩f Δ › using wfS_letI by auto
show ‹atom x ♯ (Φ, Θ, ℬ', Γ, Δ, e, b)› using wfS_letI by auto
qed
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 τ x s2 b)
then show ?case using wb_b_weakening1 Wellformed.wfS_let2I by simp
next
case (wfS_ifI Θ ℬ Γ v Φ Δ s1 b s2)
then show ?case using wb_b_weakening1 Wellformed.wfS_ifI by simp
next
case (wfS_varI Θ ℬ Γ τ v u Δ Φ s b)
then show ?case using wb_b_weakening1 Wellformed.wfS_varI by simp
next
case (wfS_assignI u τ Δ Θ ℬ Γ Φ v)
then show ?case using wb_b_weakening1 Wellformed.wfS_assignI by simp
next
case (wfS_whileI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using wb_b_weakening1 Wellformed.wfS_whileI by simp
next
case (wfS_seqI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using Wellformed.wfS_seqI by metis
next
case (wfS_matchI Θ ℬ Γ v tid dclist Δ Φ cs b)
then show ?case using wb_b_weakening1 Wellformed.wfS_matchI by metis
next
case (wfS_branchI Θ Φ ℬ x τ Γ Δ s b tid dc)
then show ?case using Wellformed.wfS_branchI by auto
next
case (wfS_finalI Θ Φ ℬ Γ Δ tid dclist' cs b dclist)
then show ?case using wf_intros by metis
next
case (wfS_cons Θ Φ ℬ Γ Δ tid dclist' cs b css dclist)
then show ?case using wf_intros by metis
next
case (wfD_emptyI Θ ℬ Γ)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfD_cons Θ ℬ Γ Δ τ u)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfPhi_emptyI Θ)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfPhi_consI f Θ Φ ft)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfFTSome Θ bv ft)
then show ?case using wf_intros wb_b_weakening1 by metis
next
case (wfFTI Θ B b s x c τ Φ)
show ?case proof
show "Θ; ℬ' ⊢⇩w⇩f b" using wfFTI wb_b_weakening1 by auto
show "supp c ⊆ {atom x}" using wfFTI wb_b_weakening1 by auto
show "Θ; ℬ' ; (x, b, c) #⇩Γ GNil ⊢⇩w⇩f τ " using wfFTI wb_b_weakening1 by auto
show "Θ ⊢⇩w⇩f Φ " using wfFTI wb_b_weakening1 by auto
from ‹ B |⊆| ℬ'› have "supp B ⊆ supp ℬ'" proof(induct B)
case empty
then show ?case by auto
next
case (insert x B)
then show ?case
by (metis fsubset_funion_eq subset_Un_eq supp_union_fset)
qed
thus "supp s ⊆ {atom x} ∪ supp ℬ'" using wfFTI by auto
qed
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
show ?case proof
show ‹ Θ ; Φ ; ℬ' ; (x, B_bool, c) #⇩Γ Γ ; Δ ⊢⇩w⇩f s : b › using wb_b_weakening1 wfS_assertI by simp
show ‹ Θ; ℬ' ; Γ ⊢⇩w⇩f c › using wb_b_weakening1 wfS_assertI by simp
show ‹ Θ; ℬ' ; Γ ⊢⇩w⇩f Δ › using wb_b_weakening1 wfS_assertI by simp
have "atom x ♯ ℬ'" using x_not_in_b_set fresh_def by metis
thus ‹atom x ♯ (Φ, Θ, ℬ', Γ, Δ, c, b, s)› using wfS_assertI fresh_prodN by simp
qed
qed(auto)
lemmas wb_b_weakening = wb_b_weakening1 wb_b_weakening2
lemma wfG_b_weakening:
fixes Γ::Γ
assumes "ℬ |⊆| ℬ'" and "Θ; ℬ ⊢⇩w⇩f Γ"
shows "Θ; ℬ' ⊢⇩w⇩f Γ "
using wb_b_weakening assms by auto
lemma wfT_b_weakening:
fixes Γ::Γ and Θ::Θ and τ::τ
assumes "ℬ |⊆| ℬ'" and "Θ; ℬ; Γ ⊢⇩w⇩f τ"
shows "Θ; ℬ' ; Γ ⊢⇩w⇩f τ "
using wb_b_weakening assms by auto
lemma wfB_subst_wfB:
fixes τ::τ and b'::b and b::b
assumes "Θ ; {|bv|} ⊢⇩w⇩f b" and "Θ; ℬ ⊢⇩w⇩f b'"
shows "Θ; ℬ ⊢⇩w⇩f b[bv::=b']⇩b⇩b "
using assms proof(nominal_induct b rule:b.strong_induct)
case B_int
hence "Θ ; {||} ⊢⇩w⇩f B_int" using wfB_intI wfX_wfY by fast
then show ?case using subst_bb.simps wb_b_weakening by fastforce
next
case B_bool
hence "Θ ; {||} ⊢⇩w⇩f B_bool" using wfB_boolI wfX_wfY by fast
then show ?case using subst_bb.simps wb_b_weakening by fastforce
next
case (B_id x )
hence " Θ; ℬ ⊢⇩w⇩f (B_id x)" using wfB_consI wfB_elims wfX_wfY by metis
then show ?case using subst_bb.simps(4) by auto
next
case (B_pair x1 x2)
then show ?case using subst_bb.simps
by (metis wfB_elims(1) wfB_pairI)
next
case B_unit
hence "Θ ; {||} ⊢⇩w⇩f B_unit" using wfB_unitI wfX_wfY by fast
then show ?case using subst_bb.simps wb_b_weakening by fastforce
next
case B_bitvec
hence "Θ ; {||} ⊢⇩w⇩f B_bitvec" using wfB_bitvecI wfX_wfY by fast
then show ?case using subst_bb.simps wb_b_weakening by fastforce
next
case (B_var x)
then show ?case
proof -
have False
using B_var.prems(1) wfB.cases by fastforce
then show ?thesis by metis
qed
next
case (B_app s b)
then obtain bv' dclist where *:"AF_typedef_poly s bv' dclist ∈ set Θ ∧ Θ ; {|bv|} ⊢⇩w⇩f b" using wfB_elims by metis
show ?case unfolding subst_b_simps proof
show "⊢⇩w⇩f Θ " using B_app wfX_wfY by metis
show "Θ ; ℬ ⊢⇩w⇩f b[bv::=b']⇩b⇩b " using * B_app forget_subst wfB_supp fresh_def
by (metis ex_in_conv subset_empty subst_b_b_def supp_empty_fset)
show "AF_typedef_poly s bv' dclist ∈ set Θ" using * by auto
qed
qed
lemma wfT_subst_wfB:
fixes τ::τ and b'::b
assumes "Θ ; {|bv|} ; (x, b, c) #⇩Γ GNil ⊢⇩w⇩f τ" and "Θ; ℬ ⊢⇩w⇩f b'"
shows "Θ; ℬ ⊢⇩w⇩f (b_of τ)[bv::=b']⇩b⇩b "
proof -
obtain b where "Θ ; {|bv|} ⊢⇩w⇩f b ∧ b_of τ = b" using wfT_elims b_of.simps assms by metis
thus ?thesis using wfB_subst_wfB assms by auto
qed
lemma wfG_cons_unique:
assumes "(x1,b1,c1) ∈ toSet (((x,b,c) #⇩ΓΓ))" and "wfG Θ ℬ (((x,b,c) #⇩ΓΓ))" and "x = x1"
shows "b1 = b ∧ c1 = c"
proof -
have "x1 ∉ fst ` toSet Γ"
proof -
have "atom x1 ♯ Γ" using assms wfG_cons by metis
then show ?thesis
using fresh_gamma_elem
by (metis assms(2) atom_dom.simps dom.simps rev_image_eqI wfG_cons2 wfG_x_fresh)
qed
thus ?thesis using assms by force
qed
lemma wfG_member_unique:
assumes "(x1,b1,c1) ∈ toSet (Γ'@((x,b,c) #⇩ΓΓ))" and "wfG Θ ℬ (Γ'@((x,b,c) #⇩ΓΓ))" and "x = x1"
shows "b1 = b ∧ c1 = c"
using assms proof(induct Γ' rule: Γ_induct)
case GNil
then show ?case using wfG_suffix wfG_cons_unique append_g.simps by metis
next
case (GCons x' b' c' Γ')
moreover hence "(x1, b1, c1) ∈ toSet (Γ' @ (x, b, c) #⇩Γ Γ)" using wf_not_in_prefix by fastforce
ultimately show ?case using wfG_cons by fastforce
qed
section ‹Function Definitions›
lemma wb_phi_weakening:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and ℬ::ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list and Φ::Φ
shows
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b ⟹ Θ ⊢⇩w⇩f Φ' ⟹ set Φ ⊆ set Φ' ⟹ Θ ; Φ' ; ℬ ; Γ ; Δ ⊢⇩w⇩f e : b" and
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ Θ ⊢⇩w⇩f Φ' ⟹ set Φ ⊆ set Φ' ⟹ Θ ; Φ' ; ℬ ; Γ ; Δ ⊢⇩w⇩f s : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ Θ ⊢⇩w⇩f Φ' ⟹ set Φ ⊆ set Φ' ⟹ Θ ; Φ' ; ℬ ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ Θ ⊢⇩w⇩f Φ' ⟹ set Φ ⊆ set Φ' ⟹ Θ ; Φ' ; ℬ ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ True" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ Θ ⊢⇩w⇩f Φ' ⟹ set Φ ⊆ set Φ' ⟹ Θ ; Φ' ⊢⇩w⇩f ftq" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ Θ ⊢⇩w⇩f Φ' ⟹ set Φ ⊆ set Φ' ⟹ Θ ; Φ' ; ℬ ⊢⇩w⇩f ft"
proof(nominal_induct
b and b and b and b and Φ and Δ and ftq and ft
avoiding: Φ'
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_valI Θ Φ ℬ Γ Δ v b)
then show ?case using wf_intros by metis
next
case (wfE_plusI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_leqI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_eqI Θ Φ ℬ Γ Δ v1 b v2)
then show ?case using wf_intros by metis
next
case (wfE_fstI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using wf_intros by metis
next
case (wfE_sndI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using wf_intros by metis
next
case (wfE_concatI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_splitI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_lenI Θ Φ ℬ Γ Δ v1)
then show ?case using wf_intros by metis
next
case (wfE_appI Θ Φ ℬ Γ Δ f x b c τ s v)
then show ?case using wf_intros lookup_fun_weakening by metis
next
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv v τ f x b c s)
show ?case proof
show ‹ Θ ⊢⇩w⇩f Φ' › using wfE_appPI by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ › using wfE_appPI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b' › using wfE_appPI by auto
show ‹atom bv ♯ (Φ', Θ, ℬ, Γ, Δ, b', v, (b_of τ)[bv::=b']⇩b)› using wfE_appPI by auto
show ‹Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ' f›
using wfE_appPI lookup_fun_weakening by metis
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f v : b[bv::=b']⇩b › using wfE_appPI by auto
qed
next
case (wfE_mvarI Θ Φ ℬ Γ Δ u τ)
then show ?case using wf_intros by metis
next
case (wfS_valI Θ Φ ℬ Γ v b Δ)
then show ?case using wf_intros by metis
next
case (wfS_letI Θ Φ ℬ Γ Δ e b' x s b)
then show ?case using Wellformed.wfS_letI by fastforce
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 b' x s2 b)
then show ?case using Wellformed.wfS_let2I by fastforce
next
case (wfS_ifI Θ ℬ Γ v Φ Δ s1 b s2)
then show ?case using wf_intros by metis
next
case (wfS_varI Θ ℬ Γ τ v u Φ Δ b s)
show ?case proof
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f τ › using wfS_varI by simp
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f v : b_of τ › using wfS_varI by simp
show ‹atom u ♯ (Φ', Θ, ℬ, Γ, Δ, τ, v, b)› using wfS_varI by simp
show ‹ Θ ; Φ' ; ℬ ; Γ ; (u, τ) #⇩Δ Δ ⊢⇩w⇩f s : b › using wfS_varI by simp
qed
next
case (wfS_assignI u τ Δ Θ ℬ Γ Φ v)
then show ?case using wf_intros by metis
next
case (wfS_whileI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using wf_intros by metis
next
case (wfS_seqI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using wf_intros by metis
next
case (wfS_matchI Θ ℬ Γ v tid dclist Δ Φ cs b)
then show ?case using wf_intros by metis
next
case (wfS_branchI Θ Φ ℬ x τ Γ Δ s b tid dc)
then show ?case using Wellformed.wfS_branchI by fastforce
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
show ?case proof
show ‹ Θ ; Φ' ; ℬ ; (x, B_bool, c) #⇩Γ Γ ; Δ ⊢⇩w⇩f s : b › using wfS_assertI by auto
next
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f c › using wfS_assertI by auto
next
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ › using wfS_assertI by auto
have "atom x ♯ Φ'" using wfS_assertI wfPhi_supp fresh_def by blast
thus ‹atom x ♯ (Φ', Θ, ℬ, Γ, Δ, c, b, s)› using fresh_prodN wfS_assertI wfPhi_supp fresh_def by auto
qed
next
case (wfFTI Θ B b s x c τ Φ)
show ?case proof
show ‹ Θ ; B ⊢⇩w⇩f b › using wfFTI by auto
next
show ‹supp c ⊆ {atom x}› using wfFTI by auto
next
show ‹ Θ ; B ; (x, b, c) #⇩Γ GNil ⊢⇩w⇩f τ › using wfFTI by auto
next
show ‹ Θ ⊢⇩w⇩f Φ' › using wfFTI by auto
next
show ‹supp s ⊆ {atom x} ∪ supp B› using wfFTI by auto
qed
qed(auto|metis wf_intros)+
lemma wfT_fun_return_t:
fixes τa'::τ and τ'::τ
assumes "Θ; ℬ; (xa, b, ca) #⇩Γ GNil ⊢⇩w⇩f τa'" and "(AF_fun_typ x b c τ' s') = (AF_fun_typ xa b ca τa' sa')"
shows "Θ; ℬ; (x, b, c) #⇩Γ GNil ⊢⇩w⇩f τ'"
proof -
obtain cb::x where xf: "atom cb ♯ (c, τ', s', sa', τa', ca, x , xa)" using obtain_fresh by blast
hence "atom cb ♯ (c, τ', s', sa', τa', ca) ∧ atom cb ♯ (x, xa, ((c, τ'), s'), (ca, τa'), sa')" using fresh_prod6 fresh_prod4 fresh_prod8 by auto
hence *:"c[x::=V_var cb]⇩c⇩v = ca[xa::=V_var cb]⇩c⇩v ∧ τ'[x::=V_var cb]⇩τ⇩v = τa'[xa::=V_var cb]⇩τ⇩v" using assms τ.eq_iff Abs1_eq_iff_all by auto
have **: "Θ; ℬ; (xa ↔ cb ) ∙ ((xa, b, ca) #⇩Γ GNil) ⊢⇩w⇩f (xa ↔ cb ) ∙ τa'" using assms True_eqvt beta_flip_eq theta_flip_eq wfG_wf
by (metis GCons_eqvt GNil_eqvt wfT.eqvt wfT_wf)
have "Θ; ℬ; (x ↔ cb ) ∙ ((x, b, c) #⇩Γ GNil) ⊢⇩w⇩f (x ↔ cb ) ∙ τ'" proof -
have "(xa ↔ cb ) ∙ xa = (x ↔ cb ) ∙ x" using xf by auto
hence "(x ↔ cb ) ∙ ((x, b, c) #⇩Γ GNil) = (xa ↔ cb ) ∙ ((xa, b, ca) #⇩Γ GNil)" using * ** xf G_cons_flip fresh_GNil by simp
thus ?thesis using ** * xf by simp
qed
thus ?thesis using beta_flip_eq theta_flip_eq wfT_wf wfG_wf * ** True_eqvt wfT.eqvt permute_flip_cancel by metis
qed
lemma wfFT_wf_aux:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q and s::s and Δ::Δ
assumes "Θ ; Φ ; B ⊢⇩w⇩f (AF_fun_typ x b c τ s)"
shows "Θ ; B ; (x,b,c) #⇩Γ GNil ⊢⇩w⇩f τ ∧ Θ ⊢⇩w⇩f Φ ∧ supp s ⊆ { atom x } ∪ supp B"
proof -
obtain xa and ca and sa and τ' where *:"Θ ; B ⊢⇩w⇩f b ∧ (Θ ⊢⇩w⇩f Φ ) ∧
supp sa ⊆ {atom xa} ∪ supp B ∧ (Θ ; B ; (xa, b, ca) #⇩Γ GNil ⊢⇩w⇩f τ') ∧
AF_fun_typ x b c τ s = AF_fun_typ xa b ca τ' sa "
using wfFT.simps[of Θ Φ B "AF_fun_typ x b c τ s"] assms by auto
moreover hence **: "(AF_fun_typ x b c τ s) = (AF_fun_typ xa b ca τ' sa)" by simp
ultimately have "Θ ; B ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ" using wfT_fun_return_t by metis
moreover have " (Θ ⊢⇩w⇩f Φ ) " using * by auto
moreover have "supp s ⊆ { atom x } ∪ supp B" proof -
have "[[atom x]]lst.s = [[atom xa]]lst.sa" using ** fun_typ.eq_iff lst_fst lst_snd by metis
thus ?thesis using lst_supp_subset * by metis
qed
ultimately show ?thesis by auto
qed
lemma wfFT_simple_wf:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q and s::s and Δ::Δ
assumes "Θ ; Φ ⊢⇩w⇩f (AF_fun_typ_none (AF_fun_typ x b c τ s))"
shows "Θ ; {||} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ ∧ Θ ⊢⇩w⇩f Φ ∧ supp s ⊆ { atom x } "
proof -
have *:"Θ ; Φ ; {||} ⊢⇩w⇩f (AF_fun_typ x b c τ s)" using wfFTQ_elims assms by auto
thus ?thesis using wfFT_wf_aux by force
qed
lemma wfFT_poly_wf:
fixes τ::τ and Θ::Θ and Φ::Φ and ftq :: fun_typ_q and s::s and Δ::Δ
assumes "Θ ; Φ ⊢⇩w⇩f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))"
shows "Θ ; {|bv|} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ ∧ Θ ⊢⇩w⇩f Φ ∧ Θ ; Φ ; {|bv|} ⊢⇩w⇩f (AF_fun_typ x b c τ s)"
proof -
obtain bv1 ft1 where *:"Θ ; Φ ; {|bv1|} ⊢⇩w⇩f ft1 ∧ [[atom bv1]]lst. ft1 = [[atom bv]]lst. AF_fun_typ x b c τ s"
using wfFTQ_elims(3)[OF assms] by metis
show ?thesis proof(cases "bv1 = bv")
case True
then show ?thesis using * fun_typ_q.eq_iff Abs1_eq_iff by (metis (no_types, opaque_lifting) wfFT_wf_aux)
next
case False
obtain x1 b1 c1 t1 s1 where **: "ft1 = AF_fun_typ x1 b1 c1 t1 s1" using fun_typ.eq_iff
by (meson fun_typ.exhaust)
hence eqv: "(bv ↔ bv1) ∙ AF_fun_typ x1 b1 c1 t1 s1 = AF_fun_typ x b c τ s ∧ atom bv1 ♯ AF_fun_typ x b c τ s" using
Abs1_eq_iff(3) * False by metis
have "(bv ↔ bv1) ∙ Θ ; (bv ↔ bv1) ∙ Φ ; (bv ↔ bv1) ∙ {|bv1|} ⊢⇩w⇩f (bv ↔ bv1) ∙ ft1" using wfFT.eqvt * by metis
moreover have "(bv ↔ bv1) ∙ Φ = Φ" using phi_flip_eq wfX_wfY * by metis
moreover have "(bv ↔ bv1) ∙ Θ =Θ" using wfX_wfY * theta_flip_eq2 by metis
moreover have "(bv ↔ bv1) ∙ ft1 = AF_fun_typ x b c τ s" using eqv ** by metis
ultimately have "Θ ; Φ ; {|bv|} ⊢⇩w⇩f AF_fun_typ x b c τ s" by auto
thus ?thesis using wfFT_wf_aux by auto
qed
qed
lemma wfFT_poly_wfT:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Θ ; Φ ⊢⇩w⇩f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))"
shows "Θ ; {| bv |} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ"
using wfFT_poly_wf assms by simp
lemma b_of_supp:
"supp (b_of t) ⊆ supp t"
proof(nominal_induct t rule:τ.strong_induct)
case (T_refined_type x b c)
then show ?case by auto
qed
lemma wfPhi_f_simple_wf:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q and s::s and Φ'::Φ
assumes "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s)) ∈ set Φ " and "Θ ⊢⇩w⇩f Φ" and "set Φ ⊆ set Φ'" and "Θ ⊢⇩w⇩f Φ'"
shows "Θ ; {||} ; (x,b,c) #⇩Γ GNil ⊢⇩w⇩f τ ∧ Θ ⊢⇩w⇩f Φ ∧ supp s ⊆ { atom x }"
using assms proof(induct Φ rule: Φ_induct)
case PNil
then show ?case by auto
next
case (PConsSome f1 bv x1 b1 c1 τ1 s' Φ'')
hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s)) ∈ set Φ''" by auto
moreover have " Θ ⊢⇩w⇩f Φ'' ∧ set Φ'' ⊆ set Φ'" using wfPhi_elims(3) PConsSome by auto
ultimately show ?case using PConsSome wfPhi_elims wfFT_simple_wf by auto
next
case (PConsNone f' x' b' c' τ' s' Φ'')
show ?case proof(cases "f=f'")
case True
have "AF_fun_typ_none (AF_fun_typ x' b' c' τ' s') = AF_fun_typ_none (AF_fun_typ x b c τ s)"
by (metis PConsNone.prems(1) PConsNone.prems(2) True fun_def.eq_iff image_eqI name_of_fun.simps set_ConsD wfPhi_elims(2))
hence *:"Θ ; Φ'' ⊢⇩w⇩f AF_fun_typ_none (AF_fun_typ x b c τ s) " using wfPhi_elims(2)[OF PConsNone(3)] by metis
hence "Θ ; Φ'' ; {||} ⊢⇩w⇩f (AF_fun_typ x b c τ s)" using wfFTQ_elims(1) by metis
thus ?thesis using wfFT_simple_wf[OF *] wb_phi_weakening PConsNone by force
next
case False
hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s)) ∈ set Φ''" using PConsNone by simp
moreover have " Θ ⊢⇩w⇩f Φ'' ∧ set Φ'' ⊆ set Φ'" using wfPhi_elims(3) PConsNone by auto
ultimately show ?thesis using PConsNone wfPhi_elims wfFT_simple_wf by auto
qed
qed
lemma wfPhi_f_simple_wfT:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "Θ ; {||} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ"
using wfPhi_f_simple_wf assms using lookup_fun_member by blast
lemma wfPhi_f_simple_supp_b:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp b = {}"
proof -
have "Θ ; {||} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ" using wfPhi_f_simple_wfT assms by auto
thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce
qed
lemma wfPhi_f_simple_supp_t:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp τ ⊆ { atom x }"
using wfPhi_f_simple_wfT wfT_supp assms by fastforce
lemma wfPhi_f_simple_supp_c:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp c ⊆ { atom x }"
proof -
have "Θ ; {||} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ" using wfPhi_f_simple_wfT assms by auto
thus ?thesis using wfG_wfC wfC_supp wfT_wf by fastforce
qed
lemma wfPhi_f_simple_supp_s:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp s ⊆ {atom x}"
proof -
have "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c τ s)) ∈ set Φ" using lookup_fun_member assms by auto
hence "supp s ⊆ { atom x }" using wfPhi_f_simple_wf assms by blast
thus ?thesis using wf_supp(3) atom_dom.simps toSet.simps x_not_in_u_set x_not_in_b_set setD.simps
using wf_supp2(2) by fastforce
qed
lemma wfPhi_f_poly_wf:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q and s::s and Φ'::Φ
assumes "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s)) ∈ set Φ " and "Θ ⊢⇩w⇩f Φ" and "set Φ ⊆ set Φ'" and "Θ ⊢⇩w⇩f Φ'"
shows "Θ ; {|bv|} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ ∧ Θ ⊢⇩w⇩f Φ' ∧ Θ ; Φ' ; {|bv|} ⊢⇩w⇩f (AF_fun_typ x b c τ s)"
using assms proof(induct Φ rule: Φ_induct)
case PNil
then show ?case by auto
next
case (PConsNone f x b c τ s' Φ'')
moreover have " Θ ⊢⇩w⇩f Φ'' ∧ set Φ'' ⊆ set Φ'" using wfPhi_elims(3) PConsNone by auto
ultimately show ?case using PConsNone wfPhi_elims wfFT_poly_wf by auto
next
case (PConsSome f1 bv1 x1 b1 c1 τ1 s1 Φ'')
show ?case proof(cases "f=f1")
case True
have "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1 s1) = AF_fun_typ_some bv (AF_fun_typ x b c τ s)"
by (metis PConsSome.prems(1) PConsSome.prems(2) True fun_def.eq_iff list.set_intros(1) option.inject wfPhi_lookup_fun_unique)
hence *:"Θ ; Φ'' ⊢⇩w⇩f AF_fun_typ_some bv (AF_fun_typ x b c τ s) " using wfPhi_elims PConsSome by metis
thus ?thesis using wfFT_poly_wf * wb_phi_weakening PConsSome
by (meson set_subset_Cons)
next
case False
hence "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s)) ∈ set Φ''" using PConsSome
by (meson fun_def.eq_iff set_ConsD)
moreover have " Θ ⊢⇩w⇩f Φ'' ∧ set Φ'' ⊆ set Φ'" using wfPhi_elims(3) PConsSome
by (meson dual_order.trans set_subset_Cons)
ultimately show ?thesis using PConsSome wfPhi_elims wfFT_poly_wf
by blast
qed
qed
lemma wfPhi_f_poly_wfT:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "Θ ; {| bv |} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ"
using assms proof(induct Φ rule: Φ_induct)
case PNil
then show ?case by auto
next
case (PConsSome f1 bv1 x1 b1 c1 τ1 s' Φ')
then show ?case proof(cases "f1=f")
case True
hence "lookup_fun (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1 s')) # Φ') f = Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1 s')))" using
lookup_fun.simps using PConsSome.prems by simp
then show ?thesis using PConsSome.prems wfPhi_elims wfFT_poly_wfT
by (metis option.inject)
next
case False
then show ?thesis using PConsSome using lookup_fun.simps
using wfPhi_elims(3) by auto
qed
next
case (PConsNone f' x' b' c' τ' s' Φ')
then show ?case proof(cases "f'=f")
case True
then have *:"Θ ; Φ' ⊢⇩w⇩f AF_fun_typ_none (AF_fun_typ x' b' c' τ' s') " using lookup_fun.simps PConsNone wfPhi_elims by metis
thus ?thesis using PConsNone wfFT_poly_wfT wfPhi_elims lookup_fun.simps
by (metis fun_def.eq_iff fun_typ_q.distinct(1) option.inject)
next
case False
thus ?thesis using PConsNone wfPhi_elims
by (metis False lookup_fun.simps(2))
qed
qed
lemma wfPhi_f_poly_supp_b:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp b ⊆ supp bv"
proof -
have "Θ ; {|bv|} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ" using wfPhi_f_poly_wfT assms by auto
thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce
qed
lemma wfPhi_f_poly_supp_t:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp τ ⊆ { atom x , atom bv }"
using wfPhi_f_poly_wfT[OF assms, THEN wfT_supp] atom_dom.simps supp_at_base by auto
lemma wfPhi_f_poly_supp_b_of_t:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp (b_of τ) ⊆ { atom bv }"
proof -
have "atom x ∉ supp (b_of τ)" using x_fresh_b by auto
moreover have "supp (b_of τ) ⊆ { atom x , atom bv }" using wfPhi_f_poly_supp_t
using supp_at_base b_of.simps wfPhi_f_poly_supp_t τ.supp b_of_supp assms by fast
ultimately show ?thesis by blast
qed
lemma wfPhi_f_poly_supp_c:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp c ⊆ { atom x, atom bv }"
proof -
have "Θ ; {|bv|} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ" using wfPhi_f_poly_wfT assms by auto
thus ?thesis using wfG_wfC wfC_supp wfT_wf
using supp_at_base by fastforce
qed
lemma wfPhi_f_poly_supp_s:
fixes τ::τ and Θ::Θ and Φ::Φ and ft :: fun_typ_q
assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f" and "Θ ⊢⇩w⇩f Φ"
shows "supp s ⊆ {atom x, atom bv}"
proof -
have "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s)) ∈ set Φ" using lookup_fun_member assms by auto
hence *:"Θ ; Φ ; {|bv|} ⊢⇩w⇩f (AF_fun_typ x b c τ s)" using assms wfPhi_f_poly_wf by simp
thus ?thesis using wfFT_wf_aux[OF *] using supp_at_base by auto
qed
lemmas wfPhi_f_supp = wfPhi_f_poly_supp_b wfPhi_f_simple_supp_b wfPhi_f_poly_supp_c
wfPhi_f_simple_supp_t wfPhi_f_poly_supp_t wfPhi_f_simple_supp_t wfPhi_f_poly_wfT wfPhi_f_simple_wfT
wfPhi_f_poly_supp_s wfPhi_f_simple_supp_s
lemma fun_typ_eq_ret_unique:
assumes "(AF_fun_typ x1 b1 c1 τ1' s1') = (AF_fun_typ x2 b2 c2 τ2' s2')"
shows "τ1'[x1::=v]⇩τ⇩v = τ2'[x2::=v]⇩τ⇩v"
proof -
have "[[atom x1]]lst. τ1' = [[atom x2]]lst. τ2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis
thus ?thesis using subst_v_flip_eq_two[of x1 τ1' x2 τ2' v] subst_v_τ_def by metis
qed
lemma fun_typ_eq_body_unique:
fixes v::v and x1::x and x2::x and s1'::s and s2'::s
assumes "(AF_fun_typ x1 b1 c1 τ1' s1') = (AF_fun_typ x2 b2 c2 τ2' s2')"
shows "s1'[x1::=v]⇩s⇩v = s2'[x2::=v]⇩s⇩v"
proof -
have "[[atom x1]]lst. s1' = [[atom x2]]lst. s2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis
thus ?thesis using subst_v_flip_eq_two[of x1 s1' x2 s2' v] subst_v_s_def by metis
qed
lemma fun_ret_unique:
assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 τ1' s1'))) = lookup_fun Φ f" and "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 τ2' s2'))) = lookup_fun Φ f"
shows "τ1'[x1::=v]⇩τ⇩v = τ2'[x2::=v]⇩τ⇩v"
proof -
have *: " (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 τ1' s1'))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 τ2' s2')))" using option.inject assms by metis
thus ?thesis using fun_typ_eq_ret_unique fun_def.eq_iff fun_typ_q.eq_iff by metis
qed
lemma fun_poly_arg_unique:
fixes bv1::bv and bv2::bv and b::b and τ1::τ and τ2::τ
assumes "[[atom bv1]]lst. (AF_fun_typ x1 b1 c1 τ1 s1) = [[atom bv2]]lst. (AF_fun_typ x2 b2 c2 τ2 s2)" (is "[[atom ?x]]lst. ?a = [[atom ?y]]lst. ?b")
shows "⦃ x1 : b1[bv1::=b]⇩b⇩b | c1[bv1::=b]⇩c⇩b ⦄ = ⦃ x2 : b2[bv2::=b]⇩b⇩b | c2[bv2::=b]⇩c⇩b ⦄"
proof -
obtain c::bv where *:"atom c ♯ (b,b1,b2,c1,c2) ∧ atom c ♯ (bv1, bv2, AF_fun_typ x1 b1 c1 τ1 s1, AF_fun_typ x2 b2 c2 τ2 s2)" using obtain_fresh fresh_Pair by metis
hence "(bv1 ↔ c) ∙ AF_fun_typ x1 b1 c1 τ1 s1 = (bv2 ↔ c) ∙ AF_fun_typ x2 b2 c2 τ2 s2" using Abs1_eq_iff_all(3)[of ?x ?a ?y ?b] assms by metis
hence "AF_fun_typ x1 ((bv1 ↔ c) ∙ b1) ((bv1 ↔ c) ∙ c1) ((bv1 ↔ c) ∙ τ1) ((bv1 ↔ c) ∙ s1) = AF_fun_typ x2 ((bv2 ↔ c) ∙ b2) ((bv2 ↔ c) ∙ c2) ((bv2 ↔ c) ∙ τ2) ((bv2 ↔ c) ∙ s2)"
using fun_typ_flip by metis
hence **:"⦃ x1 :((bv1 ↔ c) ∙ b1) | ((bv1 ↔ c) ∙ c1) ⦄ = ⦃ x2 : ((bv2 ↔ c) ∙ b2) | ((bv2 ↔ c) ∙ c2) ⦄" (is "⦃ x1 : ?b1 | ?c1 ⦄ = ⦃ x2 : ?b2 | ?c2 ⦄") using fun_arg_unique_aux by metis
hence "⦃ x1 :((bv1 ↔ c) ∙ b1) | ((bv1 ↔ c) ∙ c1) ⦄[c::=b]⇩τ⇩b = ⦃ x2 : ((bv2 ↔ c) ∙ b2) | ((bv2 ↔ c) ∙ c2) ⦄[c::=b]⇩τ⇩b" by metis
hence "⦃ x1 :((bv1 ↔ c) ∙ b1)[c::=b]⇩b⇩b | ((bv1 ↔ c) ∙ c1)[c::=b]⇩c⇩b ⦄ = ⦃ x2 : ((bv2 ↔ c) ∙ b2)[c::=b]⇩b⇩b | ((bv2 ↔ c) ∙ c2)[c::=b]⇩c⇩b ⦄" using subst_tb.simps by metis
thus ?thesis using * flip_subst_subst subst_b_c_def subst_b_b_def fresh_prodN flip_commute by metis
qed
lemma fun_poly_ret_unique:
assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1' s1'))) = lookup_fun Φ f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 τ2' s2'))) = lookup_fun Φ f"
shows "τ1'[bv1::=b]⇩τ⇩b[x1::=v]⇩τ⇩v = τ2'[bv2::=b]⇩τ⇩b[x2::=v]⇩τ⇩v"
proof -
have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 τ2' s2')))" using option.inject assms by metis
hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 τ2' s2')"
(is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis
hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis
hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis
have "[[atom x1]]lst. τ1'[bv1::=b]⇩τ⇩b = [[atom x2]]lst. τ2'[bv2::=b]⇩τ⇩b"
apply(rule lst_snd[of _ "c1[bv1::=b]⇩c⇩b" _ _ "c2[bv2::=b]⇩c⇩b"])
apply(rule lst_fst[of _ _ "s1'[bv1::=b]⇩s⇩b" _ _ "s2'[bv2::=b]⇩s⇩b"])
using * subst_ft_b.simps fun_typ.eq_iff by metis
thus ?thesis using subst_v_flip_eq_two subst_v_τ_def by metis
qed
lemma fun_poly_body_unique:
assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1' s1'))) = lookup_fun Φ f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 τ2' s2'))) = lookup_fun Φ f"
shows "s1'[bv1::=b]⇩s⇩b[x1::=v]⇩s⇩v = s2'[bv2::=b]⇩s⇩b[x2::=v]⇩s⇩v"
proof -
have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 τ2' s2')))"
using option.inject assms by metis
hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 τ2' s2')"
(is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis
hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis
hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis
have "[[atom x1]]lst. s1'[bv1::=b]⇩s⇩b = [[atom x2]]lst. s2'[bv2::=b]⇩s⇩b"
using lst_snd lst_fst subst_ft_b.simps fun_typ.eq_iff
by (metis "local.*")
thus ?thesis using subst_v_flip_eq_two subst_v_s_def by metis
qed
lemma funtyp_eq_iff_equalities:
fixes s'::s and s::s
assumes " [[atom x']]lst. ((c', τ'), s') = [[atom x]]lst. ((c, τ), s)"
shows "⦃ x' : b | c' ⦄ = ⦃ x : b | c ⦄ ∧ s'[x'::=v]⇩s⇩v = s[x::=v]⇩s⇩v ∧ τ'[x'::=v]⇩τ⇩v = τ[x::=v]⇩τ⇩v"
proof -
have "[[atom x']]lst. s' = [[atom x]]lst. s" and "[[atom x']]lst. τ' = [[atom x]]lst. τ" and
" [[atom x']]lst. c' = [[atom x]]lst. c" using lst_snd lst_fst assms by metis+
thus ?thesis using subst_v_flip_eq_two τ.eq_iff
by (metis assms fun_typ.eq_iff fun_typ_eq_body_unique fun_typ_eq_ret_unique)
qed
section ‹Weakening›
lemma wfX_wfB1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ℬ::ℬ and Φ::Φ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows wfV_wfB: "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ Θ; ℬ ⊢⇩w⇩f b" and
"Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ True" and
"Θ; ℬ ⊢⇩w⇩f Γ ⟹ True" and
wfT_wfB: "Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ Θ; ℬ ⊢⇩w⇩f b_of τ " and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f Θ ⟹ True" and
"Θ; ℬ ⊢⇩w⇩f b ⟹ True" and
wfCE_wfB: "Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ Θ; ℬ ⊢⇩w⇩f b" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
case (wfV_varI Θ ℬ Γ b c x)
hence "(x,b,c) ∈ toSet Γ" using lookup_iff wfV_wf using lookup_in_g by presburger
hence "b ∈ fst`snd`toSet Γ" by force
hence "wfB Θ ℬ b" using wfG_wfB wfV_varI by metis
then show ?case using wfV_elims wfG_wf wf_intros by metis
next
case (wfV_litI Θ Γ l)
moreover have "wfTh Θ" using wfV_wf wfG_wf wfV_litI by metis
ultimately show ?case using wfV_wf wfG_wf wf_intros base_for_lit.simps l.exhaust by metis
next
case (wfV_pairI Θ Γ v1 b1 v2 b2)
then show ?case using wfG_wf wf_intros by metis
next
case (wfV_consI s dclist Θ dc x b c B Γ v)
then show ?case
using wfV_wf wfG_wf wfB_consI by metis
next
case (wfV_conspI s bv dclist Θ dc x b' c ℬ b Γ v)
then show ?case
using wfV_wf wfG_wf using wfB_appI by metis
next
case (wfCE_valI Θ ℬ Γ v b)
then show ?case using wfB_elims by auto
next
case (wfCE_plusI Θ ℬ Γ v1 v2)
then show ?case using wfB_elims by auto
next
case (wfCE_leqI Θ ℬ Γ v1 v2)
then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
next
case (wfCE_eqI Θ ℬ Γ v1 b v2)
then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
next
case (wfCE_fstI Θ ℬ Γ v1 b1 b2)
then show ?case using wfB_elims by metis
next
case (wfCE_sndI Θ ℬ Γ v1 b1 b2)
then show ?case using wfB_elims by metis
next
case (wfCE_concatI Θ ℬ Γ v1 v2)
then show ?case using wfB_elims by auto
next
case (wfCE_lenI Θ ℬ Γ v1)
then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
qed(auto | metis wfV_wf wfG_wf wf_intros )+
lemma wfX_wfB2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and b::b and ℬ::ℬ and Φ::Φ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows
wfE_wfB: "Θ; Φ; ℬ; Γ; Δ ⊢⇩w⇩f e : b ⟹ Θ; ℬ ⊢⇩w⇩f b" and
wfS_wfB: "Θ; Φ; ℬ; Γ; Δ ⊢⇩w⇩f s : b ⟹ Θ; ℬ ⊢⇩w⇩f b" and
wfCS_wfB: "Θ; Φ; ℬ; Γ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ Θ; ℬ ⊢⇩w⇩f b" and
wfCSS_wfB: "Θ; Φ; ℬ; Γ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ Θ; ℬ ⊢⇩w⇩f b" and
"Θ ⊢⇩w⇩f Φ ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ True" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ ℬ |⊆| ℬ' ⟹ Θ ; Φ ; ℬ' ⊢⇩w⇩f ft"
proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
case (wfE_valI Θ Φ ℬ Γ Δ v b)
then show ?case using wfB_elims wfX_wfB1 by metis
next
case (wfE_plusI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wfB_elims wfX_wfB1 by metis
next
case (wfE_eqI Θ Φ ℬ Γ Δ v1 b v2)
then show ?case using wfB_boolI wfX_wfY by metis
next
case (wfE_fstI Θ Φ Γ Δ v1 b1 b2)
then show ?case using wfB_elims wfX_wfB1 by metis
next
case (wfE_sndI Θ Φ Γ Δ v1 b1 b2)
then show ?case using wfB_elims wfX_wfB1 by metis
next
case (wfE_concatI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wfB_elims wfX_wfB1 by metis
next
case (wfE_splitI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wfB_elims wfX_wfB1
using wfB_pairI by auto
next
case (wfE_lenI Θ Φ ℬ Γ Δ v1)
then show ?case using wfB_elims wfX_wfB1
using wfB_intI wfX_wfY1(1) by auto
next
case (wfE_appI Θ Φ ℬ Γ Δ f x b c τ s v)
hence "Θ; ℬ;(x,b,c) #⇩Γ GNil ⊢⇩w⇩f τ" using wfPhi_f_simple_wfT wfT_b_weakening by fast
then show ?case using b_of.simps using wfT_b_weakening
by (metis b_of.cases bot.extremum wfT_elims(2))
next
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv v τ f x b c s)
hence "Θ ; {| bv |} ;(x,b,c) #⇩Γ GNil ⊢⇩w⇩f τ" using wfPhi_f_poly_wfT wfX_wfY by blast
then show ?case using wfE_appPI b_of.simps using wfT_b_weakening wfT_elims wfT_subst_wfB subst_b_b_def by metis
next
case (wfE_mvarI Θ Φ ℬ Γ Δ u τ)
hence "Θ; ℬ; Γ ⊢⇩w⇩f τ" using wfD_wfT by fast
then show ?case using wfT_elims b_of.simps by metis
next
case (wfFTNone Θ ft)
then show ?case by auto
next
case (wfFTSome Θ bv ft)
then show ?case by auto
next
case (wfS_valI Θ Φ ℬ Γ v b Δ)
then show ?case using wfX_wfB1 by auto
next
case (wfS_letI Θ Φ ℬ Γ Δ e b' x s b)
then show ?case using wfX_wfB1 by auto
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 τ x s2 b)
then show ?case using wfX_wfB1 by auto
next
case (wfS_ifI Θ ℬ Γ v Φ Δ s1 b s2)
then show ?case using wfX_wfB1 by auto
next
case (wfS_varI Θ ℬ Γ τ v u Φ Δ b s)
then show ?case using wfX_wfB1 by auto
next
case (wfS_assignI u τ Δ Θ ℬ Γ Φ v)
then show ?case using wfX_wfB1
using wfB_unitI wfX_wfY2(5) by auto
next
case (wfS_whileI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using wfX_wfB1 by auto
next
case (wfS_seqI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using wfX_wfB1 by auto
next
case (wfS_matchI Θ ℬ Γ v tid dclist Δ Φ cs b)
then show ?case using wfX_wfB1 by auto
next
case (wfS_branchI Θ Φ ℬ x τ Γ Δ s b tid dc)
then show ?case using wfX_wfB1 by auto
next
case (wfS_finalI Θ Φ ℬ Γ Δ tid dc t cs b)
then show ?case using wfX_wfB1 by auto
next
case (wfS_cons Θ Φ ℬ Γ Δ tid dc t cs b dclist css)
then show ?case using wfX_wfB1 by auto
next
case (wfD_emptyI Θ ℬ Γ)
then show ?case using wfX_wfB1 by auto
next
case (wfD_cons Θ ℬ Γ Δ τ u)
then show ?case using wfX_wfB1 by auto
next
case (wfPhi_emptyI Θ)
then show ?case using wfX_wfB1 by auto
next
case (wfPhi_consI f Θ Φ ft)
then show ?case using wfX_wfB1 by auto
next
case (wfFTI Θ B b Φ x c s τ)
then show ?case using wfX_wfB1
by (meson Wellformed.wfFTI wb_b_weakening2(8))
qed(metis wfV_wf wfG_wf wf_intros wfX_wfB1)
lemmas wfX_wfB = wfX_wfB1 wfX_wfB2
lemma wf_weakening1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and ℬ::ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows wfV_weakening: "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; ℬ; Γ' ⊢⇩w⇩f v : b" and
wfC_weakening: "Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; ℬ; Γ' ⊢⇩w⇩f c" and
"Θ; ℬ ⊢⇩w⇩f Γ ⟹ True" and
wfT_weakening: "Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; ℬ; Γ' ⊢⇩w⇩f τ" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f P ⟹ True " and
wfB_weakening: "wfB Θ ℬ b ⟹ ℬ |⊆| ℬ' ⟹ wfB Θ ℬ b" and
wfCE_weakening: "Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; ℬ; Γ' ⊢⇩w⇩f ce : b" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(nominal_induct
b and c and Γ and τ and ts and P and b and b and td
avoiding: Γ'
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfV_varI Θ ℬ Γ b c x)
hence "Some (b, c) = lookup Γ' x" using lookup_weakening by metis
then show ?case using Wellformed.wfV_varI wfV_varI by metis
next
case (wfTI z Θ ℬ Γ b c)
show ?case proof
show ‹atom z ♯ (Θ, ℬ, Γ')› using wfTI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b › using wfTI by auto
have *:"toSet ((z, b, TRUE) #⇩Γ Γ) ⊆ toSet ((z, b, TRUE) #⇩Γ Γ')" using toSet.simps wfTI by auto
thus ‹ Θ; ℬ; (z, b, TRUE) #⇩Γ Γ' ⊢⇩w⇩f c › using wfTI(8)[OF _ *] wfTI wfX_wfY
by (simp add: wfG_cons_TRUE)
qed
next
case (wfV_conspI s bv dclist Θ dc x b' c ℬ b Γ v)
show ?case proof
show ‹AF_typedef_poly s bv dclist ∈ set Θ› using wfV_conspI by auto
show ‹(dc, ⦃ x : b' | c ⦄) ∈ set dclist› using wfV_conspI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b › using wfV_conspI by auto
show ‹atom bv ♯ (Θ, ℬ, Γ', b, v)› using wfV_conspI by simp
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f v : b'[bv::=b]⇩b⇩b › using wfV_conspI by auto
qed
qed(metis wf_intros)+
lemma wf_weakening2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and ℬ::ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows
wfE_weakening: "Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; Φ; ℬ; Γ' ; Δ ⊢⇩w⇩f e : b" and
wfS_weakening: "Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; Φ; ℬ; Γ' ; Δ ⊢⇩w⇩f s : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; Φ; ℬ; Γ' ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; Φ; ℬ; Γ' ; Δ ; tid ; dclist ⊢⇩w⇩f css : b" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ True" and
wfD_weakning: "Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ Θ; ℬ ⊢⇩w⇩f Γ' ⟹ toSet Γ ⊆ toSet Γ' ⟹ Θ; ℬ; Γ' ⊢⇩w⇩f Δ" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ True"
proof(nominal_induct
b and b and b and b and Φ and Δ and ftq and ft
avoiding: Γ'
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv v τ f x b c s)
show ?case proof
show ‹ Θ ⊢⇩w⇩f Φ › using wfE_appPI by auto
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f Δ › using wfE_appPI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b' › using wfE_appPI by auto
show ‹atom bv ♯ (Φ, Θ, ℬ, Γ', Δ, b', v, (b_of τ)[bv::=b']⇩b)› using wfE_appPI by auto
show ‹Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f› using wfE_appPI by auto
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f v : b[bv::=b']⇩b › using wfE_appPI wf_weakening1 by auto
qed
next
case (wfS_letI Θ Φ ℬ Γ Δ e b' x s b)
show ?case proof(rule)
show ‹ Θ ; Φ ; ℬ ; Γ' ; Δ ⊢⇩w⇩f e : b' › using wfS_letI by auto
have "toSet ((x, b', TRUE) #⇩Γ Γ) ⊆ toSet ((x, b', TRUE) #⇩Γ Γ')" using wfS_letI by auto
thus ‹ Θ ; Φ ; ℬ ; (x, b', TRUE) #⇩Γ Γ' ; Δ ⊢⇩w⇩f s : b › using wfS_letI by (meson wfG_cons wfG_cons_TRUE wfS_wf)
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f Δ › using wfS_letI by auto
show ‹atom x ♯ (Φ, Θ, ℬ, Γ', Δ, e, b)› using wfS_letI by auto
qed
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 τ x s2 b)
show ?case proof
show ‹ Θ ; Φ ; ℬ ; Γ' ; Δ ⊢⇩w⇩f s1 : b_of τ › using wfS_let2I by auto
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f τ › using wfS_let2I wf_weakening1 by auto
have "toSet ((x, b_of τ, TRUE) #⇩Γ Γ) ⊆ toSet ((x, b_of τ, TRUE) #⇩Γ Γ')" using wfS_let2I by auto
thus ‹ Θ ; Φ ; ℬ ; (x, b_of τ, TRUE) #⇩Γ Γ' ; Δ ⊢⇩w⇩f s2 : b › using wfS_let2I by (meson wfG_cons wfG_cons_TRUE wfS_wf)
show ‹atom x ♯ (Φ, Θ, ℬ, Γ', Δ, s1, b, τ)› using wfS_let2I by auto
qed
next
case (wfS_varI Θ ℬ Γ τ v u Φ Δ b s)
show ?case proof
show "Θ; ℬ; Γ' ⊢⇩w⇩f τ " using wfS_varI wf_weakening1 by auto
show "Θ; ℬ; Γ' ⊢⇩w⇩f v : b_of τ " using wfS_varI wf_weakening1 by auto
show "atom u ♯ (Φ, Θ, ℬ, Γ', Δ, τ, v, b)" using wfS_varI by auto
show "Θ ; Φ ; ℬ ; Γ' ; (u, τ) #⇩Δ Δ ⊢⇩w⇩f s : b " using wfS_varI by auto
qed
next
case (wfS_branchI Θ Φ ℬ x τ Γ Δ s b tid dc)
show ?case proof
have "toSet ((x, b_of τ, TRUE) #⇩Γ Γ) ⊆ toSet ((x, b_of τ, TRUE) #⇩Γ Γ')" using wfS_branchI by auto
thus ‹ Θ ; Φ ; ℬ ; (x, b_of τ, TRUE) #⇩Γ Γ' ; Δ ⊢⇩w⇩f s : b › using wfS_branchI by (meson wfG_cons wfG_cons_TRUE wfS_wf)
show ‹atom x ♯ (Φ, Θ, ℬ, Γ', Δ, Γ', τ)› using wfS_branchI by auto
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f Δ › using wfS_branchI by auto
qed
next
case (wfS_finalI Θ Φ ℬ Γ Δ tid dclist' cs b dclist)
then show ?case using wf_intros by metis
next
case (wfS_cons Θ Φ ℬ Γ Δ tid dclist' cs b css dclist)
then show ?case using wf_intros by metis
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
show ?case proof(rule)
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f c › using wfS_assertI wf_weakening1 by auto
have "Θ; ℬ ⊢⇩w⇩f (x, B_bool, c) #⇩Γ Γ'" proof(rule wfG_consI)
show ‹ Θ; ℬ ⊢⇩w⇩f Γ' › using wfS_assertI by auto
show ‹atom x ♯ Γ'› using wfS_assertI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f B_bool › using wfS_assertI wfB_boolI wfX_wfY by metis
have "Θ; ℬ ⊢⇩w⇩f (x, B_bool, TRUE) #⇩Γ Γ'" proof
show "(TRUE) ∈ {TRUE, FALSE}" by auto
show ‹ Θ; ℬ ⊢⇩w⇩f Γ' › using wfS_assertI by auto
show ‹atom x ♯ Γ'› using wfS_assertI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f B_bool › using wfS_assertI wfB_boolI wfX_wfY by metis
qed
thus ‹ Θ; ℬ; (x, B_bool, TRUE) #⇩Γ Γ' ⊢⇩w⇩f c ›
using wf_weakening1(2)[OF ‹ Θ; ℬ; Γ' ⊢⇩w⇩f c › ‹ Θ; ℬ ⊢⇩w⇩f (x, B_bool, TRUE) #⇩Γ Γ' ›] by force
qed
thus ‹ Θ; Φ; ℬ; (x, B_bool, c) #⇩Γ Γ' ; Δ ⊢⇩w⇩f s : b › using wfS_assertI by fastforce
show ‹ Θ; ℬ; Γ' ⊢⇩w⇩f Δ › using wfS_assertI by auto
show ‹atom x ♯ (Φ, Θ, ℬ, Γ', Δ, c, b, s)› using wfS_assertI by auto
qed
qed(metis wf_intros wf_weakening1)+
lemmas wf_weakening = wf_weakening1 wf_weakening2
lemma wfV_weakening_cons:
fixes Γ::Γ and Γ'::Γ and v::v and c::c
assumes "Θ; ℬ; Γ ⊢⇩w⇩f v : b" and "atom y ♯ Γ" and "Θ; ℬ; ((y,b',TRUE) #⇩Γ Γ) ⊢⇩w⇩f c"
shows "Θ; ℬ; (y,b',c) #⇩ΓΓ ⊢⇩w⇩f v : b"
proof -
have "wfG Θ ℬ ((y,b',c) #⇩ΓΓ)" using wfG_intros2 assms by auto
moreover have "toSet Γ ⊆ toSet ((y,b',c) #⇩ΓΓ)" using toSet.simps by auto
ultimately show ?thesis using wf_weakening using assms(1) by blast
qed
lemma wfG_cons_weakening:
fixes Γ'::Γ
assumes "Θ; ℬ ⊢⇩w⇩f ((x, b, c) #⇩Γ Γ)" and "Θ; ℬ ⊢⇩w⇩f Γ'" and "toSet Γ ⊆ toSet Γ'" and "atom x ♯ Γ'"
shows "Θ; ℬ ⊢⇩w⇩f ((x, b, c) #⇩Γ Γ')"
proof(cases "c ∈ {TRUE,FALSE}")
case True
then show ?thesis using wfG_wfB wfG_cons2I assms by auto
next
case False
hence *:"Θ; ℬ ⊢⇩w⇩f Γ ∧ atom x ♯ Γ ∧ Θ; ℬ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c"
using wfG_elims(2)[OF assms(1)] by auto
have a1:"Θ; ℬ ⊢⇩w⇩f (x, b, TRUE) #⇩Γ Γ'" using wfG_wfB wfG_cons2I assms by simp
moreover have a2:"toSet ((x, b, TRUE) #⇩Γ Γ ) ⊆ toSet ((x, b, TRUE) #⇩Γ Γ')" using toSet.simps assms by blast
moreover have " Θ; ℬ ⊢⇩w⇩f (x, b, TRUE) #⇩Γ Γ'" proof
show "(TRUE) ∈ {TRUE, FALSE}" by auto
show "Θ; ℬ ⊢⇩w⇩f Γ'" using assms by auto
show "atom x ♯ Γ'" using assms by auto
show "Θ; ℬ ⊢⇩w⇩f b" using assms wfG_elims by metis
qed
hence " Θ; ℬ; (x, b, TRUE) #⇩Γ Γ' ⊢⇩w⇩f c" using wf_weakening a1 a2 * by auto
then show ?thesis using wfG_cons1I[of c Θ ℬ Γ' x b, OF False ] wfG_wfB assms by simp
qed
lemma wfT_weakening_aux:
fixes Γ::Γ and Γ'::Γ and c::c
assumes "Θ; ℬ; Γ ⊢⇩w⇩f ⦃ z : b | c ⦄" and "Θ; ℬ ⊢⇩w⇩f Γ'" and "toSet Γ ⊆ toSet Γ'" and "atom z ♯ Γ'"
shows "Θ; ℬ; Γ' ⊢⇩w⇩f ⦃ z : b | c ⦄"
proof
show ‹atom z ♯ (Θ, ℬ, Γ')›
using wf_supp wfX_wfY assms fresh_prodN fresh_def x_not_in_b_set wfG_fresh_x by metis
show ‹ Θ; ℬ ⊢⇩w⇩f b › using assms wfT_elims by metis
show ‹ Θ; ℬ; (z, b, TRUE) #⇩Γ Γ' ⊢⇩w⇩f c › proof -
have *:"Θ; ℬ; (z,b,TRUE) #⇩ΓΓ ⊢⇩w⇩f c" using wfT_wfC fresh_weakening assms by auto
moreover have a1:"Θ; ℬ ⊢⇩w⇩f (z, b, TRUE) #⇩Γ Γ'" using wfG_cons2I assms ‹Θ; ℬ ⊢⇩w⇩f b› by simp
moreover have a2:"toSet ((z, b, TRUE) #⇩Γ Γ ) ⊆ toSet ((z, b, TRUE) #⇩Γ Γ')" using toSet.simps assms by blast
moreover have " Θ; ℬ ⊢⇩w⇩f (z, b, TRUE) #⇩Γ Γ' " proof
show "(TRUE) ∈ {TRUE, FALSE}" by auto
show "Θ; ℬ ⊢⇩w⇩f Γ'" using assms by auto
show "atom z ♯ Γ'" using assms by auto
show "Θ; ℬ ⊢⇩w⇩f b" using assms wfT_elims by metis
qed
thus ?thesis using wf_weakening a1 a2 * by auto
qed
qed
lemma wfT_weakening_all:
fixes Γ::Γ and Γ'::Γ and τ::τ
assumes "Θ; ℬ; Γ ⊢⇩w⇩f τ" and "Θ; ℬ' ⊢⇩w⇩f Γ'" and "toSet Γ ⊆ toSet Γ'" and "ℬ |⊆| ℬ'"
shows "Θ; ℬ' ; Γ' ⊢⇩w⇩f τ"
using wb_b_weakening assms wfT_weakening by metis
lemma wfT_weakening_nil:
fixes Γ::Γ and Γ'::Γ and τ::τ
assumes "Θ ; {||} ; GNil ⊢⇩w⇩f τ" and "Θ; ℬ' ⊢⇩w⇩f Γ'"
shows "Θ; ℬ' ; Γ' ⊢⇩w⇩f τ"
using wfT_weakening_all
using assms(1) assms(2) toSet.simps(1) by blast
lemma wfTh_wfT2:
fixes x::x and v::v and τ::τ and G::Γ
assumes "wfTh Θ" and "AF_typedef s dclist ∈ set Θ" and
"(dc, τ) ∈ set dclist" and "Θ ; B ⊢⇩w⇩f G"
shows "supp τ = {}" and "τ[x::=v]⇩τ⇩v = τ" and "wfT Θ B G τ"
proof -
show "supp τ = {}" proof(rule ccontr)
assume a1: "supp τ ≠ {}"
have "supp Θ ≠ {}" proof -
obtain dclist where dc: "AF_typedef s dclist ∈ set Θ ∧ (dc, τ) ∈ set dclist"
using assms by auto
hence "supp (dc,τ) ≠ {}"
using a1 by (simp add: supp_Pair)
hence "supp dclist ≠ {}" using dc supp_list_member by auto
hence "supp (AF_typedef s dclist) ≠ {}" using type_def.supp by auto
thus ?thesis using supp_list_member dc by auto
qed
thus False using assms wfTh_supp by simp
qed
thus "τ[x::=v]⇩τ⇩v = τ" by (simp add: fresh_def)
have "wfT Θ {||} GNil τ" using assms wfTh_wfT by auto
thus "wfT Θ B G τ" using assms wfT_weakening_nil by simp
qed
lemma wf_d_weakening:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and s::s and ℬ::ℬ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b ⟹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' ⟹ setD Δ ⊆ setD Δ' ⟹ Θ; Φ; ℬ; Γ ; Δ' ⊢⇩w⇩f e : b" and
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' ⟹ setD Δ ⊆ setD Δ' ⟹ Θ; Φ; ℬ; Γ ; Δ' ⊢⇩w⇩f s : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' ⟹ setD Δ ⊆ setD Δ' ⟹ Θ; Φ; ℬ; Γ ; Δ' ; tid ; dc ; t ⊢⇩w⇩f cs : b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' ⟹ setD Δ ⊆ setD Δ' ⟹ Θ; Φ; ℬ; Γ ; Δ' ; tid ; dclist ⊢⇩w⇩f css : b" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ True" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ True"
proof(nominal_induct
b and b and b and b and Φ and Δ and ftq and ft
avoiding: Δ'
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_valI Θ Φ ℬ Γ Δ v b)
then show ?case using wf_intros by metis
next
case (wfE_plusI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_leqI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_eqI Θ Φ ℬ Γ Δ v1 b v2)
then show ?case using wf_intros by metis
next
case (wfE_fstI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using wf_intros by metis
next
case (wfE_sndI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using wf_intros by metis
next
case (wfE_concatI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_splitI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_intros by metis
next
case (wfE_lenI Θ Φ ℬ Γ Δ v1)
then show ?case using wf_intros by metis
next
case (wfE_appI Θ Φ ℬ Γ Δ f x b c τ s v)
then show ?case using wf_intros by metis
next
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv v τ f x b c s)
show ?case proof(rule, (rule wfE_appPI)+)
show ‹atom bv ♯ (Φ, Θ, ℬ, Γ, Δ', b', v, (b_of τ)[bv::=b']⇩b)› using wfE_appPI by auto
show ‹Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c τ s))) = lookup_fun Φ f› using wfE_appPI by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f v : b[bv::=b']⇩b › using wfE_appPI by auto
qed
next
case (wfE_mvarI Θ Φ ℬ Γ Δ u τ)
show ?case proof
show ‹ Θ ⊢⇩w⇩f Φ › using wfE_mvarI by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' › using wfE_mvarI by auto
show ‹(u, τ) ∈ setD Δ'› using wfE_mvarI by auto
qed
next
case (wfS_valI Θ Φ ℬ Γ v b Δ)
then show ?case using wf_intros by metis
next
case (wfS_letI Θ Φ ℬ Γ Δ e b' x s b)
show ?case proof(rule)
show ‹ Θ; Φ; ℬ; Γ; Δ' ⊢⇩w⇩f e : b' › using wfS_letI by auto
have "Θ; ℬ ⊢⇩w⇩f (x, b', TRUE) #⇩Γ Γ" using wfG_cons2I wfX_wfY wfS_letI by metis
hence "Θ; ℬ; (x, b', TRUE) #⇩Γ Γ ⊢⇩w⇩f Δ'" using wf_weakening2(6) wfS_letI by force
thus ‹ Θ ; Φ ; ℬ ; (x, b', TRUE) #⇩Γ Γ ; Δ' ⊢⇩w⇩f s : b › using wfS_letI by metis
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' › using wfS_letI by auto
show ‹atom x ♯ (Φ, Θ, ℬ, Γ, Δ', e, b)› using wfS_letI by auto
qed
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
show ?case proof
have "Θ; ℬ; (x, B_bool, c) #⇩Γ Γ ⊢⇩w⇩f Δ'" proof(rule wf_weakening2(6))
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' › using wfS_assertI by auto
next
show ‹ Θ; ℬ ⊢⇩w⇩f (x, B_bool, c) #⇩Γ Γ › using wfS_assertI wfX_wfY by metis
next
show ‹toSet Γ ⊆ toSet ((x, B_bool, c) #⇩Γ Γ)› using wfS_assertI by auto
qed
thus ‹ Θ; Φ; ℬ; (x, B_bool, c) #⇩Γ Γ ; Δ' ⊢⇩w⇩f s : b › using wfS_assertI wfX_wfY by metis
next
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f c › using wfS_assertI by auto
next
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' › using wfS_assertI by auto
next
show ‹atom x ♯ (Φ, Θ, ℬ, Γ, Δ', c, b, s)› using wfS_assertI by auto
qed
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 τ x s2 b)
show ?case proof
show ‹ Θ; Φ; ℬ; Γ; Δ' ⊢⇩w⇩f s1 : b_of τ › using wfS_let2I by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f τ › using wfS_let2I by auto
have "Θ; ℬ ⊢⇩w⇩f (x, b_of τ, TRUE) #⇩Γ Γ" using wfG_cons2I wfX_wfY wfS_let2I by metis
hence "Θ; ℬ; (x, b_of τ, TRUE) #⇩Γ Γ ⊢⇩w⇩f Δ'" using wf_weakening2(6) wfS_let2I by force
thus ‹ Θ ; Φ ; ℬ ; (x, b_of τ, TRUE) #⇩Γ Γ ; Δ' ⊢⇩w⇩f s2 : b › using wfS_let2I by metis
show ‹atom x ♯ (Φ, Θ, ℬ, Γ, Δ', s1, b,τ)› using wfS_let2I by auto
qed
next
case (wfS_ifI Θ ℬ Γ v Φ Δ s1 b s2)
then show ?case using wf_intros by metis
next
case (wfS_varI Θ ℬ Γ τ v u Φ Δ b s)
show ?case proof
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f τ › using wfS_varI by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f v : b_of τ › using wfS_varI by auto
show ‹atom u ♯ (Φ, Θ, ℬ, Γ, Δ', τ, v, b)› using wfS_varI setD.simps by auto
have "Θ; ℬ; Γ ⊢⇩w⇩f (u, τ) #⇩Δ Δ'" using wfS_varI wfD_cons setD.simps u_fresh_d by metis
thus ‹ Θ ; Φ ; ℬ ; Γ ; (u, τ) #⇩Δ Δ' ⊢⇩w⇩f s : b › using wfS_varI setD.simps by blast
qed
next
case (wfS_assignI u τ Δ Θ ℬ Γ Φ v)
show ?case proof
show ‹(u, τ) ∈ setD Δ'› using wfS_assignI setD.simps by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' › using wfS_assignI by auto
show ‹ Θ ⊢⇩w⇩f Φ › using wfS_assignI by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f v : b_of τ › using wfS_assignI by auto
qed
next
case (wfS_whileI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using wf_intros by metis
next
case (wfS_seqI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using wf_intros by metis
next
case (wfS_matchI Θ ℬ Γ v tid dclist Δ Φ cs b)
then show ?case using wf_intros by metis
next
case (wfS_branchI Θ Φ ℬ x τ Γ Δ s b tid dc)
show ?case proof
have "Θ; ℬ ⊢⇩w⇩f (x, b_of τ, TRUE) #⇩Γ Γ" using wfG_cons2I wfX_wfY wfS_branchI by metis
hence "Θ; ℬ; (x, b_of τ, TRUE) #⇩Γ Γ ⊢⇩w⇩f Δ'" using wf_weakening2(6) wfS_branchI by force
thus ‹ Θ ; Φ ; ℬ ; (x, b_of τ, TRUE) #⇩Γ Γ ; Δ' ⊢⇩w⇩f s : b › using wfS_branchI by simp
show ‹ atom x ♯ (Φ, Θ, ℬ, Γ, Δ', Γ, τ)› using wfS_branchI by auto
show ‹ Θ; ℬ; Γ ⊢⇩w⇩f Δ' › using wfS_branchI by auto
qed
next
case (wfS_finalI Θ Φ ℬ Γ Δ tid dclist' cs b dclist)
then show ?case using wf_intros by metis
next
case (wfS_cons Θ Φ ℬ Γ Δ tid dclist' cs b css dclist)
then show ?case using wf_intros by metis
qed(auto+)
section ‹Useful well-formedness instances›
text ‹Well-formedness for particular constructs that we will need later›
lemma wfC_e_eq:
fixes ce::ce and Γ::Γ
assumes "Θ ; ℬ ; Γ ⊢⇩w⇩f ce : b" and "atom x ♯ Γ "
shows "Θ ; ℬ ; ((x, b, TRUE) #⇩Γ Γ) ⊢⇩w⇩f (CE_val (V_var x) == ce )"
proof -
have "Θ; ℬ ⊢⇩w⇩f b" using assms wfX_wfB by auto
hence wbg: "Θ; ℬ ⊢⇩w⇩f Γ" using wfX_wfY assms by auto
show ?thesis proof
show *:"Θ ; ℬ ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f CE_val (V_var x) : b"
proof(rule)
show "Θ ; ℬ ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f V_var x : b " proof
show "Θ ; ℬ ⊢⇩w⇩f (x, b, TRUE) #⇩Γ Γ " using wfG_cons2I wfX_wfY assms ‹Θ; ℬ ⊢⇩w⇩f b› by auto
show "Some (b, TRUE) = lookup ((x, b, TRUE) #⇩Γ Γ) x" using lookup.simps by auto
qed
qed
show "Θ ; ℬ ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f ce : b"
using assms wf_weakening1(8)[OF assms(1), of "(x, b, TRUE) #⇩Γ Γ "] * toSet.simps wfX_wfY
by (metis Un_subset_iff equalityE)
qed
qed
lemma wfC_e_eq2:
fixes e1::ce and e2::ce
assumes "Θ ; ℬ ; Γ ⊢⇩w⇩f e1 : b" and "Θ ; ℬ ; Γ ⊢⇩w⇩f e2 : b" and " ⊢⇩w⇩f Θ" and "atom x ♯ Γ"
shows "Θ; ℬ; (x, b, (CE_val (V_var x)) == e1) #⇩Γ Γ ⊢⇩w⇩f (CE_val (V_var x)) == e2 "
proof(rule wfC_eqI)
have *: "Θ; ℬ ⊢⇩w⇩f (x, b, CE_val (V_var x) == e1 ) #⇩Γ Γ" proof(rule wfG_cons1I )
show "(CE_val (V_var x) == e1 ) ∉ {TRUE, FALSE}" by auto
show "Θ; ℬ ⊢⇩w⇩f Γ" using assms wfX_wfY by metis
show *:"atom x ♯ Γ" using assms by auto
show "Θ; ℬ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f CE_val (V_var x) == e1" using wfC_e_eq assms * by auto
show "Θ; ℬ ⊢⇩w⇩f b" using assms wfX_wfB by auto
qed
show "Θ; ℬ; (x, b, CE_val (V_var x) == e1 ) #⇩Γ Γ ⊢⇩w⇩f CE_val (V_var x) : b" using assms * wfCE_valI wfV_varI by auto
show "Θ; ℬ; (x, b, CE_val (V_var x) == e1 ) #⇩Γ Γ ⊢⇩w⇩f e2 : b" proof(rule wf_weakening1(8))
show "Θ; ℬ; Γ ⊢⇩w⇩f e2 : b " using assms by auto
show "Θ; ℬ ⊢⇩w⇩f (x, b, CE_val (V_var x) == e1 ) #⇩Γ Γ" using * by auto
show "toSet Γ ⊆ toSet ((x, b, CE_val (V_var x) == e1 ) #⇩Γ Γ)" by auto
qed
qed
lemma wfT_wfT_if_rev:
assumes "wfV P ℬ Γ v (base_for_lit l)" and "wfT P ℬ Γ t" and ‹atom z1 ♯ Γ›
shows "wfT P ℬ Γ (⦃ z1 : b_of t | CE_val v == CE_val (V_lit l) IMP (c_of t z1) ⦄)"
proof
show ‹ P; ℬ ⊢⇩w⇩f b_of t › using wfX_wfY assms by meson
have wfg: " P; ℬ ⊢⇩w⇩f (z1, b_of t, TRUE) #⇩Γ Γ" using assms wfV_wf wfG_cons2I wfX_wfY
by (meson wfG_cons_TRUE)
show ‹ P; ℬ ; (z1, b_of t, TRUE) #⇩Γ Γ ⊢⇩w⇩f [ v ]⇧c⇧e == [ [ l ]⇧v ]⇧c⇧e IMP c_of t z1 › proof
show *: ‹ P; ℬ ; (z1, b_of t, TRUE) #⇩Γ Γ ⊢⇩w⇩f [ v ]⇧c⇧e == [ [ l ]⇧v ]⇧c⇧e ›
proof(rule wfC_eqI[where b="base_for_lit l"])
show "P; ℬ ; (z1, b_of t, TRUE) #⇩Γ Γ ⊢⇩w⇩f [ v ]⇧c⇧e : base_for_lit l"
using assms wf_intros wf_weakening wfg by (meson wfV_weakening_cons)
show "P; ℬ ; (z1, b_of t, TRUE) #⇩Γ Γ ⊢⇩w⇩f [ [ l ]⇧v ]⇧c⇧e : base_for_lit l" using wfg assms wf_intros wf_weakening wfV_weakening_cons by meson
qed
have " t = ⦃ z1 : b_of t | c_of t z1 ⦄" using c_of_eq
using assms(2) assms(3) b_of_c_of_eq wfT_x_fresh by auto
thus ‹ P; ℬ ; (z1, b_of t, TRUE) #⇩Γ Γ ⊢⇩w⇩f c_of t z1 › using wfT_wfC assms wfG_elims * by simp
qed
show ‹atom z1 ♯ (P, ℬ, Γ)› using assms wfG_fresh_x wfX_wfY by metis
qed
lemma wfT_eq_imp:
fixes zz::x and ll::l and τ'::τ
assumes "base_for_lit ll = B_bool" and "Θ ; {||} ; GNil ⊢⇩w⇩f τ'" and
"Θ ; {||} ⊢⇩w⇩f (x, b_of ⦃ z' : B_bool | TRUE ⦄, c_of ⦃ z' : B_bool | TRUE ⦄ x) #⇩Γ GNil" and "atom zz ♯ x"
shows "Θ ; {||} ; (x, b_of ⦃ z' : B_bool | TRUE ⦄, c_of ⦃ z' : B_bool | TRUE ⦄ x) #⇩Γ
GNil ⊢⇩w⇩f ⦃ zz : b_of τ' | [ [ x ]⇧v ]⇧c⇧e == [ [ ll ]⇧v ]⇧c⇧e IMP c_of τ' zz ⦄"
proof(rule wfT_wfT_if_rev)
show ‹ Θ ; {||} ; (x, b_of ⦃ z' : B_bool | TRUE ⦄, c_of ⦃ z' : B_bool | TRUE ⦄ x) #⇩Γ GNil ⊢⇩w⇩f [ x ]⇧v : base_for_lit ll ›
using wfV_varI lookup.simps base_for_lit.simps assms by simp
show ‹ Θ ; {||} ; (x, b_of ⦃ z' : B_bool | TRUE ⦄, c_of ⦃ z' : B_bool | TRUE ⦄ x) #⇩Γ GNil ⊢⇩w⇩f τ' ›
using wf_weakening assms toSet.simps by auto
show ‹atom zz ♯ (x, b_of ⦃ z' : B_bool | TRUE ⦄, c_of ⦃ z' : B_bool | TRUE ⦄ x) #⇩Γ GNil›
unfolding fresh_GCons fresh_prod3 b_of.simps c_of_true
using x_fresh_b fresh_GNil c_of_true c.fresh assms by metis
qed
lemma wfC_v_eq:
fixes ce::ce and Γ::Γ and v::v
assumes "Θ ; ℬ ; Γ ⊢⇩w⇩f v : b" and "atom x ♯ Γ "
shows "Θ ; ℬ ; ((x, b, TRUE) #⇩Γ Γ) ⊢⇩w⇩f (CE_val (V_var x) == CE_val v )"
using wfC_e_eq wfCE_valI assms wfX_wfY by auto
lemma wfT_e_eq:
fixes ce::ce
assumes "Θ ; ℬ ; Γ ⊢⇩w⇩f ce : b" and "atom z ♯ Γ"
shows "Θ; ℬ; Γ ⊢⇩w⇩f ⦃ z : b | CE_val (V_var z) == ce ⦄"
proof
show "Θ; ℬ ⊢⇩w⇩f b" using wfX_wfB assms by auto
show " atom z ♯ (Θ, ℬ, Γ)" using assms wfG_fresh_x wfX_wfY by metis
show "Θ ; ℬ ; (z, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f CE_val (V_var z) == ce "
using wfTI wfC_e_eq assms wfTI by auto
qed
lemma wfT_v_eq:
assumes " wfB Θ ℬ b" and "wfV Θ ℬ Γ v b" and "atom z ♯ Γ"
shows "wfT Θ ℬ Γ ⦃ z : b | C_eq (CE_val (V_var z)) (CE_val v)⦄"
using wfT_e_eq wfE_valI assms wfX_wfY
by (simp add: wfCE_valI)
lemma wfC_wfG:
fixes Γ::Γ and c::c and b::b
assumes "Θ ; B ; Γ ⊢⇩w⇩f c" and "Θ ; B ⊢⇩w⇩f b" and "atom x ♯ Γ"
shows "Θ ; B ⊢⇩w⇩f (x,b,c)#⇩Γ Γ"
proof -
have " Θ ; B ⊢⇩w⇩f (x, b, TRUE) #⇩Γ Γ" using wfG_cons2I assms wfX_wfY by fast
hence " Θ ; B ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c " using wfC_weakening assms by force
thus ?thesis using wfG_consI assms wfX_wfY by metis
qed
section ‹Replacing the constraint on a variable in a context›
lemma wfG_cons_fresh2:
fixes Γ'::Γ
assumes "wfG P ℬ (( (x',b',c') #⇩Γ Γ' @ (x, b, c) #⇩Γ Γ))"
shows "x'≠x"
proof -
have "atom x' ♯ (Γ' @ (x, b, c) #⇩Γ Γ)"
using assms wfG_elims(2) by blast
thus ?thesis
using fresh_gamma_append[of "atom x'" Γ' "(x, b, c) #⇩Γ Γ"] fresh_GCons fresh_prod3[of "atom x'" x b c] by auto
qed
lemma replace_in_g_inside:
fixes Γ::Γ
assumes "wfG P ℬ (Γ'@((x,b0,c0') #⇩ΓΓ))"
shows "replace_in_g (Γ'@((x,b0,c0') #⇩ΓΓ)) x c0 = (Γ'@((x,b0,c0) #⇩ΓΓ))"
using assms proof(induct Γ' rule: Γ_induct)
case GNil
then show ?case using replace_in_g.simps by auto
next
case (GCons x' b' c' Γ'')
hence "P; ℬ ⊢⇩w⇩f ((x', b', c') #⇩Γ (Γ''@ (x, b0, c0') #⇩Γ Γ ))" by simp
hence "x ≠ x'" using wfG_cons_fresh2 by metis
then show ?case using replace_in_g.simps GCons by (simp add: wfG_cons)
qed
lemma wfG_supp_rig_eq:
fixes Γ::Γ
assumes "wfG P ℬ (Γ'' @ (x, b0, c0) #⇩Γ Γ)" and "wfG P ℬ (Γ'' @ (x, b0, c0') #⇩Γ Γ)"
shows "supp (Γ'' @ (x, b0, c0') #⇩Γ Γ) ∪ supp ℬ = supp (Γ'' @ (x, b0, c0) #⇩Γ Γ) ∪ supp ℬ"
using assms proof(induct Γ'')
case GNil
have "supp (GNil @ (x, b0, c0') #⇩Γ Γ) ∪ supp ℬ = supp ((x, b0, c0') #⇩Γ Γ) ∪ supp ℬ" using supp_Cons supp_GNil by auto
also have "... = supp x ∪ supp b0 ∪ supp c0' ∪ supp Γ ∪ supp ℬ " using supp_GCons by auto
also have "... = supp x ∪ supp b0 ∪ supp c0 ∪ supp Γ ∪ supp ℬ " using GNil wfG_wfC[THEN wfC_supp_cons(2) ] by fastforce
also have "... = (supp ((x, b0, c0) #⇩Γ Γ)) ∪ supp ℬ " using supp_GCons by auto
finally have "supp (GNil @ (x, b0, c0') #⇩Γ Γ) ∪ supp ℬ = supp (GNil @ (x, b0, c0) #⇩Γ Γ) ∪ supp ℬ" using supp_Cons supp_GNil by auto
then show ?case using supp_GCons wfG_cons2 by auto
next
case (GCons xbc Γ1)
moreover have " (xbc #⇩Γ Γ1) @ (x, b0, c0) #⇩Γ Γ = (xbc #⇩Γ (Γ1 @ (x, b0, c0) #⇩Γ Γ))" by simp
moreover have " (xbc #⇩Γ Γ1) @ (x, b0, c0') #⇩Γ Γ = (xbc #⇩Γ (Γ1 @ (x, b0, c0') #⇩Γ Γ))" by simp
ultimately have "(P; ℬ ⊢⇩w⇩f Γ1 @ ((x, b0, c0) #⇩Γ Γ)) ∧ P; ℬ ⊢⇩w⇩f Γ1 @ ((x, b0, c0') #⇩Γ Γ)" using wfG_cons2 by metis
thus ?case using GCons supp_GCons by auto
qed
lemma fresh_replace_inside[ms_fresh]:
fixes y::x and Γ::Γ
assumes "wfG P ℬ (Γ'' @ (x, b, c) #⇩Γ Γ)" and "wfG P ℬ (Γ'' @ (x, b, c') #⇩Γ Γ)"
shows "atom y ♯ (Γ'' @ (x, b, c) #⇩Γ Γ) = atom y ♯ (Γ'' @ (x, b, c') #⇩Γ Γ)"
unfolding fresh_def using wfG_supp_rig_eq assms x_not_in_b_set by fast
lemma wf_replace_inside1:
fixes Γ::Γ and Φ::Φ and Θ::Θ and Γ'::Γ and v::v and e::e and c::c and c''::c and c'::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and b'::b and b::b and s::s
and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
shows wfV_replace_inside: "Θ; ℬ; G ⊢⇩w⇩f v : b' ⟹ G = (Γ' @ (x, b, c') #⇩Γ Γ) ⟹ Θ; ℬ; ((x,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c ⟹ Θ ; ℬ ; (Γ' @ (x, b, c) #⇩Γ Γ) ⊢⇩w⇩f v : b'" and
wfC_replace_inside: "Θ; ℬ; G ⊢⇩w⇩f c'' ⟹ G = (Γ' @ (x, b, c') #⇩Γ Γ) ⟹ Θ; ℬ; ((x,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c ⟹ Θ ; ℬ ; (Γ' @ (x, b, c) #⇩Γ Γ) ⊢⇩w⇩f c''" and
wfG_replace_inside: "Θ; ℬ ⊢⇩w⇩f G ⟹ G = (Γ' @ (x, b, c') #⇩Γ Γ) ⟹ Θ; ℬ; ((x,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c ⟹ Θ; ℬ ⊢⇩w⇩f (Γ' @ (x, b, c) #⇩Γ Γ) " and
wfT_replace_inside: "Θ; ℬ; G ⊢⇩w⇩f τ ⟹ G = (Γ' @ (x, b, c') #⇩Γ Γ) ⟹ Θ; ℬ; ((x,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c ⟹ Θ ; ℬ ; (Γ' @ (x, b, c) #⇩Γ Γ) ⊢⇩w⇩f τ" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f P ⟹ True" and
"Θ; ℬ ⊢⇩w⇩f b ⟹ True" and
wfCE_replace_inside: "Θ ; ℬ ; G ⊢⇩w⇩f ce : b' ⟹ G = (Γ' @ (x, b, c') #⇩Γ Γ) ⟹ Θ; ℬ; ((x,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c ⟹ Θ ; ℬ ; (Γ' @ (x, b, c) #⇩Γ Γ) ⊢⇩w⇩f ce : b'" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(nominal_induct
b' and c'' and G and τ and ts and P and b and b' and td
avoiding: Γ' c'
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfV_varI Θ ℬ Γ2 b2 c2 x2)
then show ?case using wf_intros by (metis lookup_in_rig_eq lookup_in_rig_neq replace_in_g_inside)
next
case (wfV_conspI s bv dclist Θ dc x1 b' c1 ℬ b1 Γ1 v)
show ?case proof
show ‹AF_typedef_poly s bv dclist ∈ set Θ› using wfV_conspI by auto
show ‹(dc, ⦃ x1 : b' | c1 ⦄) ∈ set dclist› using wfV_conspI by auto
show ‹ Θ ; ℬ ⊢⇩w⇩f b1 › using wfV_conspI by auto
show *: ‹ Θ; ℬ; Γ' @ (x, b, c) #⇩Γ Γ ⊢⇩w⇩f v : b'[bv::=b1]⇩b⇩b › using wfV_conspI by auto
moreover have "Θ; ℬ ⊢⇩w⇩f Γ' @ (x, b, c') #⇩Γ Γ" using wfV_wf wfV_conspI by simp
ultimately have "atom bv ♯ Γ' @ (x, b, c) #⇩Γ Γ" unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfV_conspI
by (metis Un_iff fresh_def)
thus ‹atom bv ♯ (Θ, ℬ, Γ' @ (x, b, c) #⇩Γ Γ, b1, v)›
unfolding fresh_prodN using fresh_prodN wfV_conspI by metis
qed
next
case (wfTI z Θ ℬ G b1 c1)
show ?case proof
show ‹ Θ; ℬ ⊢⇩w⇩f b1 › using wfTI by auto
have "Θ; ℬ ⊢⇩w⇩f (x, b, c) #⇩Γ Γ" using wfG_consI wfTI wfG_cons wfX_wfY by metis
moreover hence *:"wfG Θ ℬ (Γ' @ (x, b, c) #⇩Γ Γ)" using wfX_wfY
by (metis append_g.simps(2) wfG_cons2 wfTI.hyps wfTI.prems(1) wfTI.prems(2))
hence ‹atom z ♯ Γ' @ (x, b, c) #⇩Γ Γ›
using fresh_replace_inside[of Θ ℬ Γ' x b c Γ c' z,OF *] wfTI wfX_wfY wfG_elims by metis
thus ‹atom z ♯ (Θ, ℬ, Γ' @ (x, b, c) #⇩Γ Γ)› using wfG_fresh_x[OF *] by auto
have "(z, b1, TRUE) #⇩Γ G = ((z, b1, TRUE) #⇩Γ Γ') @ (x, b, c') #⇩Γ Γ"
using wfTI append_g.simps by metis
thus ‹ Θ; ℬ; (z, b1, TRUE) #⇩Γ Γ' @ (x, b, c) #⇩Γ Γ ⊢⇩w⇩f c1 ›
using wfTI(9)[OF _ wfTI(11)] by fastforce
qed
next
case (wfG_nilI Θ)
hence "GNil = (x, b, c') #⇩Γ Γ" using append_g.simps Γ.distinct GNil_append by auto
hence "False" using Γ.distinct by auto
then show ?case by auto
next
case (wfG_cons1I c1 Θ ℬ G x1 b1)
show ?case proof(cases "Γ'=GNil")
case True
then show ?thesis using wfG_cons1I wfG_consI by auto
next
case False
then obtain G'::Γ where *:"(x1, b1, c1) #⇩Γ G' = Γ'" using wfG_cons1I wfG_cons1I(7) GCons_eq_append_conv by auto
hence **:" G = G' @ (x, b, c') #⇩Γ Γ" using wfG_cons1I by auto
hence " Θ; ℬ ⊢⇩w⇩f G' @ (x, b, c) #⇩Γ Γ" using wfG_cons1I by auto
have "Θ; ℬ ⊢⇩w⇩f (x1, b1, c1) #⇩Γ G' @ (x, b, c) #⇩Γ Γ" proof(rule Wellformed.wfG_cons1I)
show "c1 ∉ {TRUE, FALSE}" using wfG_cons1I by auto
show "Θ; ℬ ⊢⇩w⇩f G' @ (x, b, c) #⇩Γ Γ " using wfG_cons1I(3)[of G',OF **] wfG_cons1I by auto
show "atom x1 ♯ G' @ (x, b, c) #⇩Γ Γ" using wfG_cons1I * ** fresh_replace_inside by metis
show "Θ; ℬ; (x1, b1, TRUE) #⇩Γ G' @ (x, b, c) #⇩Γ Γ ⊢⇩w⇩f c1" using wfG_cons1I(6)[of " (x1, b1, TRUE) #⇩Γ G'"] wfG_cons1I ** by auto
show "Θ; ℬ ⊢⇩w⇩f b1" using wfG_cons1I by auto
qed
thus ?thesis using * by auto
qed
next
case (wfG_cons2I c1 Θ ℬ G x1 b1)
show ?case proof(cases "Γ'=GNil")
case True
then show ?thesis using wfG_cons2I wfG_consI by auto
next
case False
then obtain G'::Γ where *:"(x1, b1, c1) #⇩Γ G' = Γ'" using wfG_cons2I GCons_eq_append_conv by auto
hence **:" G = G' @ (x, b, c') #⇩Γ Γ" using wfG_cons2I by auto
moreover have " Θ; ℬ ⊢⇩w⇩f G' @ (x, b, c) #⇩Γ Γ" using wfG_cons2I * ** by auto
moreover hence "atom x1 ♯ G' @ (x, b, c) #⇩Γ Γ" using wfG_cons2I * ** fresh_replace_inside by metis
ultimately show ?thesis using Wellformed.wfG_cons2I[OF wfG_cons2I(1), of Θ ℬ "G'@ (x, b, c) #⇩Γ Γ" x1 b1] wfG_cons2I * ** by auto
qed
qed(metis wf_intros )+
lemma wf_replace_inside2:
fixes Γ::Γ and Φ::Φ and Θ::Θ and Γ'::Γ and v::v and e::e and c::c and c''::c and c'::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and b'::b and b::b and s::s
and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
shows
"Θ ; Φ ; ℬ ; G ; D ⊢⇩w⇩f e : b' ⟹ G = (Γ' @ (x, b, c') #⇩Γ Γ) ⟹ Θ; ℬ; ((x,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c ⟹ Θ ; Φ ; ℬ ; (Γ' @ (x, b, c) #⇩Γ Γ); D ⊢⇩w⇩f e : b'" and
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ True" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ True" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ True" and
"Θ ⊢⇩w⇩f Φ ⟹ True" and
"Θ; ℬ; G ⊢⇩w⇩f Δ ⟹ G = (Γ' @ (x, b, c') #⇩Γ Γ) ⟹ Θ; ℬ; ((x,b,TRUE) #⇩ΓΓ) ⊢⇩w⇩f c ⟹ Θ ; ℬ ; (Γ' @ (x, b, c) #⇩Γ Γ) ⊢⇩w⇩f Δ" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ True"
proof(nominal_induct
b' and b and b and b and Φ and Δ and ftq and ft
avoiding: Γ' c'
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_valI Θ Φ ℬ Γ Δ v b)
then show ?case using wf_replace_inside1 Wellformed.wfE_valI by auto
next
case (wfE_plusI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_replace_inside1 Wellformed.wfE_plusI by auto
next
case (wfE_leqI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_replace_inside1 Wellformed.wfE_leqI by auto
next
case (wfE_eqI Θ Φ ℬ Γ Δ v1 b v2)
then show ?case using wf_replace_inside1 Wellformed.wfE_eqI by metis
next
case (wfE_fstI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using wf_replace_inside1 Wellformed.wfE_fstI by metis
next
case (wfE_sndI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case using wf_replace_inside1 Wellformed.wfE_sndI by metis
next
case (wfE_concatI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_replace_inside1 Wellformed.wfE_concatI by auto
next
case (wfE_splitI Θ Φ ℬ Γ Δ v1 v2)
then show ?case using wf_replace_inside1 Wellformed.wfE_splitI by auto
next
case (wfE_lenI Θ Φ ℬ Γ Δ v1)
then show ?case using wf_replace_inside1 Wellformed.wfE_lenI by metis
next
case (wfE_appI Θ Φ ℬ Γ Δ f x b c τ s v)
then show ?case using wf_replace_inside1 Wellformed.wfE_appI by metis
next
case (wfE_appPI Θ Φ ℬ Γ'' Δ b' bv v τ f x1 b1 c1 s)
show ?case proof
show ‹ Θ ⊢⇩w⇩f Φ › using wfE_appPI by auto
show ‹ Θ; ℬ; Γ' @ (x, b, c) #⇩Γ Γ ⊢⇩w⇩f Δ › using wfE_appPI by auto
show ‹ Θ; ℬ ⊢⇩w⇩f b' › using wfE_appPI by auto
show *:‹ Θ; ℬ; Γ' @ (x, b, c) #⇩Γ Γ ⊢⇩w⇩f v : b1[bv::=b']⇩b › using wfE_appPI wf_replace_inside1 by auto
moreover have "Θ; ℬ ⊢⇩w⇩f Γ' @ (x, b, c') #⇩Γ Γ" using wfV_wf wfE_appPI by metis
ultimately have "atom bv ♯ Γ' @ (x, b, c) #⇩Γ Γ"
unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfE_appPI Un_iff fresh_def by metis
thus ‹atom bv ♯ (Φ, Θ, ℬ, Γ' @ (x, b, c) #⇩Γ Γ, Δ, b', v, (b_of τ)[bv::=b']⇩b)›
using wfE_appPI fresh_prodN by metis
show ‹Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b1 c1 τ s))) = lookup_fun Φ f› using wfE_appPI by auto
qed
next
case (wfE_mvarI Θ Φ ℬ Γ Δ u τ)
then show ?case using wf_replace_inside1 Wellformed.wfE_mvarI by metis
next
case (wfD_emptyI Θ ℬ Γ)
then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
next
case (wfD_cons Θ ℬ Γ Δ τ u)
then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI
by (simp add: wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfD_cons)
next
case (wfFTNone Θ Φ ft)
then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
next
case (wfFTSome Θ Φ bv ft)
then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
qed(auto)
lemmas wf_replace_inside = wf_replace_inside1 wf_replace_inside2
lemma wfC_replace_cons:
assumes "wfG P ℬ ((x,b,c1) #⇩ΓΓ)" and "wfC P ℬ ((x,b,TRUE) #⇩ΓΓ) c2"
shows "wfC P ℬ ((x,b,c1) #⇩ΓΓ) c2"
proof -
have "wfC P ℬ (GNil@((x,b,c1) #⇩ΓΓ)) c2" proof(rule wf_replace_inside1(2))
show " P; ℬ ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c2 " using wfG_elim2 assms by auto
show ‹(x, b, TRUE) #⇩Γ Γ = GNil @ (x, b, TRUE) #⇩Γ Γ› using append_g.simps by auto
show ‹P; ℬ ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c1 › using wfG_elim2 assms by auto
qed
thus ?thesis using append_g.simps by auto
qed
lemma wfC_refl:
assumes "wfG Θ ℬ ((x, b', c') #⇩ΓΓ)"
shows "wfC Θ ℬ ((x, b', c') #⇩ΓΓ) c'"
using wfG_wfC assms wfC_replace_cons by auto
lemma wfG_wfC_inside:
assumes " (x, b, c) ∈ toSet G" and "wfG Θ B G"
shows "wfC Θ B G c"
using assms proof(induct G rule: Γ_induct)
case GNil
then show ?case by auto
next
case (GCons x' b' c' Γ')
then consider (hd) "(x, b, c) = (x',b',c')" | (tail) "(x, b, c) ∈ toSet Γ'" using toSet.simps by auto
then show ?case proof(cases)
case hd
then show ?thesis using GCons wf_weakening
by (metis wfC_replace_cons wfG_cons_wfC)
next
case tail
then show ?thesis using GCons wf_weakening
by (metis insert_iff insert_is_Un subsetI toSet.simps(2) wfG_cons2)
qed
qed
lemma wfT_wf_cons3:
assumes "Θ; ℬ; Γ ⊢⇩w⇩f ⦃ z : b | c ⦄" and "atom y ♯ (c,Γ)"
shows "Θ; ℬ ⊢⇩w⇩f (y, b, c[z::=V_var y]⇩c⇩v) #⇩Γ Γ"
proof -
have "⦃ z : b | c ⦄ = ⦃ y : b | (y ↔ z) ∙ c ⦄" using type_eq_flip assms by auto
moreover hence " (y ↔ z) ∙ c = c[z::=V_var y]⇩c⇩v" using assms subst_v_c_def by auto
ultimately have "⦃ z : b | c ⦄ = ⦃ y : b | c[z::=V_var y]⇩c⇩v ⦄" by metis
thus ?thesis using assms wfT_wf_cons[of Θ ℬ Γ y b] fresh_Pair by metis
qed
lemma wfT_wfC_cons:
assumes "wfT P ℬ Γ ⦃ z1 : b | c1 ⦄" and "wfT P ℬ Γ ⦃ z2 : b | c2 ⦄" and "atom x ♯ (c1,c2,Γ)"
shows "wfC P ℬ ((x,b,c1[z1::=V_var x]⇩v) #⇩ΓΓ) (c2[z2::=V_var x]⇩v)" (is "wfC P ℬ ?G ?c")
proof -
have eq: "⦃ z2 : b | c2 ⦄ = ⦃ x : b | c2[z2::=V_var x]⇩c⇩v ⦄" using type_eq_subst assms fresh_prod3 by simp
have eq2: "⦃ z1 : b | c1 ⦄ = ⦃ x : b | c1[z1::=V_var x]⇩c⇩v ⦄" using type_eq_subst assms fresh_prod3 by simp
moreover have "wfT P ℬ Γ ⦃ x : b | c1[z1::=V_var x]⇩c⇩v ⦄" using assms eq2 by auto
moreover hence "wfG P ℬ ((x,b,c1[z1::=V_var x]⇩c⇩v) #⇩ΓΓ)" using wfT_wf_cons fresh_prod3 assms by auto
moreover have "wfT P ℬ Γ ⦃ x : b | c2[z2::=V_var x]⇩c⇩v ⦄" using assms eq by auto
moreover hence "wfC P ℬ ((x,b,TRUE) #⇩ΓΓ) (c2[z2::=V_var x]⇩c⇩v)" using wfT_wfC assms fresh_prod3 by simp
ultimately show ?thesis using wfC_replace_cons subst_v_c_def by simp
qed
lemma wfT_wfC2:
fixes c::c and x::x
assumes "Θ; ℬ; Γ ⊢⇩w⇩f ⦃ z : b | c ⦄" and "atom x ♯ Γ"
shows "Θ; ℬ; (x,b,TRUE)#⇩ΓΓ ⊢⇩w⇩f c[z::=[x]⇧v]⇩v"
proof(cases "x=z")
case True
then show ?thesis using wfT_wfC assms by auto
next
case False
hence "atom x ♯ c" using wfT_fresh_c assms by metis
hence "⦃ x : b | c[z::=[ x ]⇧v]⇩v ⦄ = ⦃ z : b | c ⦄"
using τ.eq_iff Abs1_eq_iff(3)[of x "c[z::=[ x ]⇧v]⇩v" z c]
by (metis flip_subst_v type_eq_flip)
hence " Θ; ℬ; Γ ⊢⇩w⇩f ⦃ x : b | c[z::=[ x ]⇧v]⇩v ⦄" using assms by metis
thus ?thesis using wfT_wfC assms by auto
qed
lemma wfT_wfG:
fixes x::x and Γ::Γ and z::x and c::c and b::b
assumes "Θ; ℬ; Γ ⊢⇩w⇩f ⦃ z : b | c ⦄" and "atom x ♯ Γ"
shows "Θ; ℬ ⊢⇩w⇩f (x,b, c[z::=[ x ]⇧v]⇩v) #⇩Γ Γ"
proof -
have "Θ; ℬ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c[z::=[ x ]⇧v]⇩v" using wfT_wfC2 assms by metis
thus ?thesis using wfG_consI assms wfT_wfB b_of.simps wfX_wfY by metis
qed
lemma wfG_replace_inside2:
fixes Γ::Γ
assumes "wfG P ℬ (Γ' @ (x, b, c') #⇩Γ Γ)" and "wfG P ℬ ((x,b,c) #⇩ΓΓ)"
shows "wfG P ℬ (Γ' @ (x, b, c) #⇩Γ Γ)"
proof -
have "wfC P ℬ ((x,b,TRUE) #⇩ΓΓ) c" using wfG_wfC assms by auto
thus ?thesis using wf_replace_inside1(3)[OF assms(1)] by auto
qed
lemma wfG_replace_inside_full:
fixes Γ::Γ
assumes "wfG P ℬ (Γ' @ (x, b, c') #⇩Γ Γ)" and "wfG P ℬ (Γ'@((x,b,c) #⇩ΓΓ))"
shows "wfG P ℬ (Γ' @ (x, b, c) #⇩Γ Γ)"
proof -
have "wfG P ℬ ((x,b,c) #⇩ΓΓ)" using wfG_suffix assms by auto
thus ?thesis using wfG_replace_inside assms by auto
qed
lemma wfT_replace_inside2:
assumes "wfT Θ ℬ (Γ' @ (x, b, c') #⇩Γ Γ) t" and "wfG Θ ℬ (Γ'@((x,b,c) #⇩ΓΓ))"
shows "wfT Θ ℬ (Γ' @ (x, b, c) #⇩Γ Γ) t"
proof -
have "wfG Θ ℬ (((x,b,c) #⇩ΓΓ))" using wfG_suffix assms by auto
hence "wfC Θ ℬ ((x,b,TRUE) #⇩ΓΓ) c" using wfG_wfC by auto
thus ?thesis using wf_replace_inside assms by metis
qed
lemma wfD_unique:
assumes "wfD P ℬ Γ Δ" and " (u,τ') ∈ setD Δ" and "(u,τ) ∈ setD Δ"
shows "τ'=τ"
using assms proof(induct Δ rule: Δ_induct)
case DNil
then show ?case by auto
next
case (DCons u' t' D)
hence *: "wfD P ℬ Γ ((u',t') #⇩Δ D)" using Cons by auto
show ?case proof(cases "u=u'")
case True
then have "u ∉ fst ` setD D" using wfD_elims * by blast
then show ?thesis using DCons by force
next
case False
then show ?thesis using DCons wfD_elims * by (metis fst_conv setD_ConsD)
qed
qed
lemma replace_in_g_forget:
fixes x::x
assumes "wfG P B G"
shows "atom x ∉ atom_dom G ⟹ (G[x⟼c]) = G" and
"atom x ♯ G ⟹ (G[x⟼c]) = G"
proof -
show "atom x ∉ atom_dom G ⟹ G[x⟼c] = G" by (induct G rule: Γ_induct,auto)
thus "atom x ♯ G ⟹ (G[x⟼c]) = G" using wfG_x_fresh assms by simp
qed
lemma replace_in_g_fresh_single:
fixes G::Γ and x::x
assumes ‹Θ; ℬ ⊢⇩w⇩f G[x'⟼c'']› and "atom x ♯ G" and ‹Θ; ℬ ⊢⇩w⇩f G ›
shows "atom x ♯ G[x'⟼c'']"
using rig_dom_eq wfG_dom_supp assms fresh_def atom_dom.simps dom.simps by metis
section ‹Preservation of well-formedness under substitution›
lemma wfC_cons_switch:
fixes c::c and c'::c
assumes "Θ; ℬ; (x, b, c) #⇩Γ Γ ⊢⇩w⇩f c'"
shows "Θ; ℬ; (x, b, c') #⇩Γ Γ ⊢⇩w⇩f c"
proof -
have *:"Θ; ℬ ⊢⇩w⇩f (x, b, c) #⇩Γ Γ" using wfC_wf assms by auto
hence "atom x ♯ Γ ∧ wfG Θ ℬ Γ ∧ Θ; ℬ ⊢⇩w⇩f b" using wfG_cons by auto
hence " Θ; ℬ; (x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f TRUE " using wfC_trueI wfG_cons2I by simp
hence "Θ; ℬ;(x, b, TRUE) #⇩Γ Γ ⊢⇩w⇩f c'"
using wf_replace_inside1(2)[of Θ ℬ "(x, b, c) #⇩Γ Γ" c' GNil x b c Γ TRUE] assms by auto
hence "wfG Θ ℬ ((x,b,c') #⇩ΓΓ)" using wf_replace_inside1(3)[OF *, of GNil x b c Γ c'] by auto
moreover have "wfC Θ ℬ ((x,b,TRUE) #⇩ΓΓ) c" proof(cases "c ∈ { TRUE, FALSE }")
case True
have "Θ; ℬ ⊢⇩w⇩f Γ ∧ atom x ♯ Γ ∧ Θ; ℬ ⊢⇩w⇩f b" using wfG_elims(2)[OF *] by auto
hence "Θ; ℬ ⊢⇩w⇩f (x,b,TRUE) #⇩Γ Γ" using wfG_cons_TRUE by auto
then show ?thesis using wfC_trueI wfC_falseI True by auto
next
case False
then show ?thesis using wfG_elims(2)[OF *] by auto
qed
ultimately show ?thesis using wfC_replace_cons by auto
qed
lemma subst_g_inside_simple:
fixes Γ⇩1::Γ and Γ⇩2::Γ
assumes "wfG P ℬ (Γ⇩1@((x,b,c) #⇩ΓΓ⇩2))"
shows "(Γ⇩1@((x,b,c) #⇩ΓΓ⇩2))[x::=v]⇩Γ⇩v = Γ⇩1[x::=v]⇩Γ⇩v@Γ⇩2"
using assms proof(induct Γ⇩1 rule: Γ_induct)
case GNil
then show ?case using subst_gv.simps by simp
next
case (GCons x' b' c' G)
hence *:"P; ℬ ⊢⇩w⇩f (x', b', c') #⇩Γ (G @ (x, b, c) #⇩Γ Γ⇩2)" by auto
hence "x≠x'"
using GCons append_Cons wfG_cons_fresh2[OF *] by auto
hence "((GCons (x', b', c') G) @ (GCons (x, b, c) Γ⇩2))[x::=v]⇩Γ⇩v =
(GCons (x', b', c') (G @ (GCons (x, b, c) Γ⇩2)))[x::=v]⇩Γ⇩v" by auto
also have "... = GCons (x', b', c'[x::=v]⇩c⇩v) ((G @ (GCons (x, b, c) Γ⇩2))[x::=v]⇩Γ⇩v)"
using subst_gv.simps ‹x≠x'› by simp
also have "... = (x', b', c'[x::=v]⇩c⇩v) #⇩Γ (G[x::=v]⇩Γ⇩v @ Γ⇩2)" using GCons * wfG_elims by metis
also have "... = ((x', b', c') #⇩Γ G)[x::=v]⇩Γ⇩v @ Γ⇩2" using subst_gv.simps ‹x≠x'› by simp
finally show ?case by blast
qed
lemma subst_c_TRUE_FALSE:
fixes c::c
assumes "c ∉ {TRUE,FALSE}"
shows "c[x::=v']⇩c⇩v ∉ {TRUE, FALSE}"
using assms by(nominal_induct c rule:c.strong_induct,auto simp add: subst_cv.simps)
lemma lookup_subst:
assumes "Some (b, c) = lookup Γ x" and "x ≠ x'"
shows "∃c'. Some (b,c') = lookup Γ[x'::=v']⇩Γ⇩v x"
using assms proof(induct Γ rule: Γ_induct)
case GNil
then show ?case by auto
next
case (GCons x1 b1 c1 Γ1)
then show ?case proof(cases "x1=x'")
case True
then show ?thesis using subst_gv.simps GCons by auto
next
case False
hence *:"((x1, b1, c1) #⇩Γ Γ1)[x'::=v']⇩Γ⇩v = ((x1, b1, c1[x'::=v']⇩c⇩v) #⇩Γ Γ1[x'::=v']⇩Γ⇩v)" using subst_gv.simps by auto
then show ?thesis proof(cases "x1=x")
case True
then show ?thesis using lookup.simps *
using GCons.prems(1) by auto
next
case False
then show ?thesis using lookup.simps *
using GCons.prems(1) by (simp add: GCons.hyps assms(2))
qed
qed
qed
lemma lookup_subst2:
assumes "Some (b, c) = lookup (Γ'@((x',b⇩1,c0[z0::=[x']⇧v]⇩c⇩v)#⇩ΓΓ)) x" and "x ≠ x'" and
"Θ; ℬ ⊢⇩w⇩f (Γ'@((x',b⇩1,c0[z0::=[x']⇧v]⇩c⇩v)#⇩ΓΓ))"
shows "∃c'. Some (b,c') = lookup (Γ'[x'::=v']⇩Γ⇩v@Γ) x"
using assms lookup_subst subst_g_inside by metis
lemma wf_subst1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
shows wfV_subst: "Θ; ℬ; Γ ⊢⇩w⇩f v : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ;Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ ; ℬ ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f v[x::=v']⇩v⇩v : b" and
wfC_subst: "Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f c[x::=v']⇩c⇩v" and
wfG_subst: "Θ; ℬ ⊢⇩w⇩f Γ ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ; ℬ ⊢⇩w⇩f Γ[x::=v']⇩Γ⇩v" and
wfT_subst: "Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f τ[x::=v']⇩τ⇩v" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f Θ ⟹True" and
"Θ; ℬ ⊢⇩w⇩f b ⟹ True " and
wfCE_subst: "Θ; ℬ; Γ ⊢⇩w⇩f ce : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ ; ℬ ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f ce[x::=v']⇩c⇩e⇩v : b" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(nominal_induct
b and c and Γ and τ and ts and Θ and b and b and td
avoiding: x v'
arbitrary: Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfV_varI Θ ℬ Γ b1 c1 x1)
show ?case proof(cases "x1=x")
case True
hence "(V_var x1)[x::=v']⇩v⇩v = v' " using subst_vv.simps by auto
moreover have "b' = b1" using wfV_varI True lookup_inside_wf
by (metis option.inject prod.inject)
moreover have " Θ; ℬ ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f v' : b'" using wfV_varI subst_g_inside_simple wf_weakening
append_g_toSetU sup_ge2 wfV_wf by metis
ultimately show ?thesis by auto
next
case False
hence "(V_var x1)[x::=v']⇩v⇩v = (V_var x1) " using subst_vv.simps by auto
moreover have "Θ; ℬ ⊢⇩w⇩f Γ[x::=v']⇩Γ⇩v" using wfV_varI by simp
moreover obtain c1' where "Some (b1, c1') = lookup Γ[x::=v']⇩Γ⇩v x1" using wfV_varI False lookup_subst by metis
ultimately show ?thesis using Wellformed.wfV_varI[of Θ ℬ "Γ[x::=v']⇩Γ⇩v" b1 c1' x1] by metis
qed
next
case (wfV_litI Θ Γ l)
then show ?case using subst_vv.simps wf_intros by auto
next
case (wfV_pairI Θ Γ v1 b1 v2 b2)
then show ?case using subst_vv.simps wf_intros by auto
next
case (wfV_consI s dclist Θ dc x b c Γ v)
then show ?case using subst_vv.simps wf_intros by auto
next
case (wfV_conspI s bv dclist Θ dc x' b' c ℬ b Γ va)
show ?case unfolding subst_vv.simps proof
show ‹AF_typedef_poly s bv dclist ∈ set Θ› and ‹(dc, ⦃ x' : b' | c ⦄) ∈ set dclist› using wfV_conspI by auto
show ‹ Θ ;ℬ ⊢⇩w⇩f b › using wfV_conspI by auto
have "atom bv ♯ Γ[x::=v']⇩Γ⇩v" using fresh_subst_gv_if wfV_conspI by metis
moreover have "atom bv ♯ va[x::=v']⇩v⇩v" using wfV_conspI fresh_subst_if by simp
ultimately show ‹atom bv ♯ (Θ, ℬ, Γ[x::=v']⇩Γ⇩v, b, va[x::=v']⇩v⇩v)› unfolding fresh_prodN using wfV_conspI by auto
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f va[x::=v']⇩v⇩v : b'[bv::=b]⇩b⇩b › using wfV_conspI by auto
qed
next
case (wfTI z Θ ℬ Γ b c)
have " Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f ⦃ z : b | c[x::=v']⇩c⇩v ⦄" proof
have ‹Θ; ℬ; ((z, b, TRUE) #⇩Γ Γ)[x::=v']⇩Γ⇩v ⊢⇩w⇩f c[x::=v']⇩c⇩v ›
proof(rule wfTI(9))
show ‹(z, b, TRUE) #⇩Γ Γ = ((z, b, TRUE) #⇩Γ Γ⇩1) @ (x, b', c') #⇩Γ Γ⇩2› using wfTI append_g.simps by simp
show ‹ Θ; ℬ; Γ⇩2 ⊢⇩w⇩f v' : b' › using wfTI by auto
qed
thus *:‹Θ; ℬ; (z, b, TRUE) #⇩Γ Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f c[x::=v']⇩c⇩v ›
using subst_gv.simps subst_cv.simps wfTI fresh_x_neq by auto
have "atom z ♯ Γ[x::=v']⇩Γ⇩v" using fresh_subst_gv_if wfTI by metis
moreover have "Θ; ℬ ⊢⇩w⇩f Γ[x::=v']⇩Γ⇩v" using wfTI wfX_wfY wfG_elims subst_gv.simps * by metis
ultimately show ‹atom z ♯ (Θ, ℬ, Γ[x::=v']⇩Γ⇩v)› using wfG_fresh_x by metis
show ‹ Θ; ℬ ⊢⇩w⇩f b › using wfTI by auto
qed
thus ?case using subst_tv.simps wfTI by auto
next
case (wfC_trueI Θ Γ)
then show ?case using subst_cv.simps wf_intros by auto
next
case (wfC_falseI Θ Γ)
then show ?case using subst_cv.simps wf_intros by auto
next
case (wfC_eqI Θ ℬ Γ e1 b e2)
show ?case proof(subst subst_cv.simps,rule)
show "Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f e1[x::=v']⇩c⇩e⇩v : b " using wfC_eqI subst_dv.simps by auto
show "Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f e2[x::=v']⇩c⇩e⇩v : b " using wfC_eqI by auto
qed
next
case (wfC_conjI Θ Γ c1 c2)
then show ?case using subst_cv.simps wf_intros by auto
next
case (wfC_disjI Θ Γ c1 c2)
then show ?case using subst_cv.simps wf_intros by auto
next
case (wfC_notI Θ Γ c1)
then show ?case using subst_cv.simps wf_intros by auto
next
case (wfC_impI Θ Γ c1 c2)
then show ?case using subst_cv.simps wf_intros by auto
next
case (wfG_nilI Θ)
then show ?case using subst_cv.simps wf_intros by auto
next
case (wfG_cons1I c Θ ℬ Γ y b)
show ?case proof(cases "x=y")
case True
hence "((y, b, c) #⇩Γ Γ)[x::=v']⇩Γ⇩v = Γ" using subst_gv.simps by auto
moreover have "Θ; ℬ ⊢⇩w⇩f Γ" using wfG_cons1I by auto
ultimately show ?thesis by auto
next
case False
have "Γ⇩1 ≠ GNil" using wfG_cons1I False by auto
then obtain G where "Γ⇩1 = (y, b, c) #⇩Γ G" using GCons_eq_append_conv wfG_cons1I by auto
hence *:"Γ = G @ (x, b', c') #⇩Γ Γ⇩2" using wfG_cons1I by auto
hence "((y, b, c) #⇩Γ Γ)[x::=v']⇩Γ⇩v =(y, b, c[x::=v']⇩c⇩v) #⇩ΓΓ[x::=v']⇩Γ⇩v" using subst_gv.simps False by auto
moreover have "Θ; ℬ ⊢⇩w⇩f (y, b, c[x::=v']⇩c⇩v) #⇩ΓΓ[x::=v']⇩Γ⇩v" proof(rule Wellformed.wfG_cons1I)
show ‹c[x::=v']⇩c⇩v ∉ {TRUE, FALSE}› using wfG_cons1I subst_c_TRUE_FALSE by auto
show ‹ Θ; ℬ ⊢⇩w⇩f Γ[x::=v']⇩Γ⇩v › using wfG_cons1I * by auto
have "Γ = (G @ ((x, b', c') #⇩ΓGNil)) @ Γ⇩2" using * append_g_assoc by auto
hence "atom y ♯ Γ⇩2" using fresh_suffix ‹atom y ♯ Γ› by auto
hence "atom y ♯ v'" using wfG_cons1I wfV_x_fresh by metis
thus ‹atom y ♯ Γ[x::=v']⇩Γ⇩v› using fresh_subst_gv wfG_cons1I by auto
have "((y, b, TRUE) #⇩Γ Γ)[x::=v']⇩Γ⇩v = (y, b, TRUE) #⇩Γ Γ[x::=v']⇩Γ⇩v" using subst_gv.simps subst_cv.simps False by auto
thus ‹ Θ; ℬ; (y, b, TRUE) #⇩Γ Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f c[x::=v']⇩c⇩v › using wfG_cons1I(6)[of "(y,b,TRUE) #⇩ΓG"] * subst_gv.simps
wfG_cons1I by fastforce
show "Θ; ℬ ⊢⇩w⇩f b " using wfG_cons1I by auto
qed
ultimately show ?thesis by auto
qed
next
case (wfG_cons2I c Θ ℬ Γ y b)
show ?case proof(cases "x=y")
case True
hence "((y, b, c) #⇩Γ Γ)[x::=v']⇩Γ⇩v = Γ" using subst_gv.simps by auto
moreover have "Θ; ℬ ⊢⇩w⇩f Γ" using wfG_cons2I by auto
ultimately show ?thesis by auto
next
case False
have "Γ⇩1 ≠ GNil" using wfG_cons2I False by auto
then obtain G where "Γ⇩1 = (y, b, c) #⇩Γ G" using GCons_eq_append_conv wfG_cons2I by auto
hence *:"Γ = G @ (x, b', c') #⇩Γ Γ⇩2" using wfG_cons2I by auto
hence "((y, b, c) #⇩Γ Γ)[x::=v']⇩Γ⇩v =(y, b, c[x::=v']⇩c⇩v) #⇩ΓΓ[x::=v']⇩Γ⇩v" using subst_gv.simps False by auto
moreover have "Θ; ℬ ⊢⇩w⇩f (y, b, c[x::=v']⇩c⇩v) #⇩ΓΓ[x::=v']⇩Γ⇩v" proof(rule Wellformed.wfG_cons2I)
show ‹c[x::=v']⇩c⇩v ∈ {TRUE, FALSE}› using subst_cv.simps wfG_cons2I by auto
show ‹ Θ; ℬ ⊢⇩w⇩f Γ[x::=v']⇩Γ⇩v › using wfG_cons2I * by auto
have "Γ = (G @ ((x, b', c') #⇩ΓGNil)) @ Γ⇩2" using * append_g_assoc by auto
hence "atom y ♯ Γ⇩2" using fresh_suffix wfG_cons2I by metis
hence "atom y ♯ v'" using wfG_cons2I wfV_x_fresh by metis
thus ‹atom y ♯ Γ[x::=v']⇩Γ⇩v› using fresh_subst_gv wfG_cons2I by auto
show "Θ; ℬ ⊢⇩w⇩f b " using wfG_cons2I by auto
qed
ultimately show ?thesis by auto
qed
next
case (wfCE_valI Θ ℬ Γ v b)
then show ?case using subst_vv.simps wf_intros by auto
next
case (wfCE_plusI Θ ℬ Γ v1 v2)
then show ?case using subst_vv.simps wf_intros by auto
next
case (wfCE_leqI Θ ℬ Γ v1 v2)
then show ?case using subst_vv.simps wf_intros by auto
next
case (wfCE_eqI Θ ℬ Γ v1 b v2)
then show ?case unfolding subst_cev.simps
using Wellformed.wfCE_eqI by metis
next
case (wfCE_fstI Θ ℬ Γ v1 b1 b2)
then show ?case using Wellformed.wfCE_fstI subst_cev.simps by metis
next
case (wfCE_sndI Θ ℬ Γ v1 b1 b2)
then show ?case using subst_cev.simps wf_intros by metis
next
case (wfCE_concatI Θ ℬ Γ v1 v2)
then show ?case using subst_vv.simps wf_intros by auto
next
case (wfCE_lenI Θ ℬ Γ v1)
then show ?case using subst_vv.simps wf_intros by auto
qed(metis subst_sv.simps wf_intros)+
lemma wf_subst2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
shows "Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f e : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f e[x::=v']⇩e⇩v : b" and
"Θ; Φ; ℬ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ ;ℬ ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s[x::=v']⇩s⇩v : b" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ; Φ; ℬ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ; tid ; dc ; t ⊢⇩w⇩f subst_branchv cs x v' : b" and
"Θ; Φ; ℬ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ; Φ; ℬ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ; tid ; dclist ⊢⇩w⇩f subst_branchlv css x v' : b" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ True " and
"Θ; ℬ; Γ ⊢⇩w⇩f Δ ⟹ Γ=Γ⇩1@((x,b',c') #⇩ΓΓ⇩2) ⟹ Θ; ℬ ; Γ⇩2 ⊢⇩w⇩f v' : b' ⟹ Θ ; ℬ ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ True"
proof(nominal_induct
b and b and b and b and Φ and Δ and ftq and ft
avoiding: x v'
arbitrary: Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_valI Θ Γ v b)
then show ?case using subst_vv.simps wf_intros wf_subst1
by (metis subst_ev.simps(1))
next
case (wfE_plusI Θ Γ v1 v2)
then show ?case using subst_vv.simps wf_intros wf_subst1 by auto
next
case (wfE_leqI Θ Φ Γ Δ v1 v2)
then show ?case
using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_leqI
by auto
next
case (wfE_eqI Θ Φ Γ Δ v1 b v2)
then show ?case
using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_eqI
proof -
show ?thesis
by (metis (no_types) subst_ev.simps(4) wfE_eqI.hyps(1) wfE_eqI.hyps(4) wfE_eqI.hyps(5) wfE_eqI.hyps(6) wfE_eqI.hyps(7) wfE_eqI.prems(1) wfE_eqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_eqI wfV_subst)
qed
next
case (wfE_fstI Θ Γ v1 b1 b2)
then show ?case using subst_vv.simps subst_ev.simps wf_subst1 Wellformed.wfE_fstI
proof -
show ?thesis
by (metis (full_types) subst_ev.simps(5) wfE_fstI.hyps(1) wfE_fstI.hyps(4) wfE_fstI.hyps(5) wfE_fstI.prems(1) wfE_fstI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_fstI wf_subst1(1))
qed
next
case (wfE_sndI Θ Γ v1 b1 b2)
then show ?case
by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_sndI wf_subst1(1))
next
case (wfE_concatI Θ Φ Γ Δ v1 v2)
then show ?case
by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_concatI wf_subst1(1))
next
case (wfE_splitI Θ Φ Γ Δ v1 v2)
then show ?case
by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_splitI wf_subst1(1))
next
case (wfE_lenI Θ Φ Γ Δ v1)
then show ?case
by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_lenI wf_subst1(1))
next
case (wfE_appI Θ Φ Γ Δ f x b c τ s' v)
then show ?case
by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_appI wf_subst1(1))
next
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv1 v1 τ1 f1 x1 b1 c1 s1)
show ?case proof(subst subst_ev.simps, rule)
show "Θ ⊢⇩w⇩f Φ" using wfE_appPI wfX_wfY by metis
show "Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v " using wfE_appPI by auto
show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1 s1))) = lookup_fun Φ f1" using wfE_appPI by auto
show "Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f v1[x::=v']⇩v⇩v : b1[bv1::=b']⇩b " using wfE_appPI wf_subst1 by auto
show "Θ; ℬ ⊢⇩w⇩f b' " using wfE_appPI by auto
have "atom bv1 ♯ Γ[x::=v']⇩Γ⇩v" using fresh_subst_gv_if wfE_appPI by metis
moreover have "atom bv1 ♯ v1[x::=v']⇩v⇩v" using wfE_appPI fresh_subst_if by simp
moreover have "atom bv1 ♯ Δ[x::=v']⇩Δ⇩v" using wfE_appPI fresh_subst_dv_if by simp
ultimately show "atom bv1 ♯ (Φ, Θ, ℬ, Γ[x::=v']⇩Γ⇩v, Δ[x::=v']⇩Δ⇩v, b', v1[x::=v']⇩v⇩v, (b_of τ1)[bv1::=b']⇩b)"
using wfE_appPI fresh_prodN by metis
qed
next
case (wfE_mvarI Θ Φ ℬ Γ Δ u τ)
have " Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f (AE_mvar u) : b_of τ[x::=v']⇩τ⇩v" proof
show "Θ ⊢⇩w⇩f Φ " using wfE_mvarI by auto
show "Θ; ℬ ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v " using wfE_mvarI by auto
show "(u, τ[x::=v']⇩τ⇩v) ∈ setD Δ[x::=v']⇩Δ⇩v" using wfE_mvarI subst_dv_member by auto
qed
thus ?case using subst_ev.simps b_of_subst by auto
next
case (wfD_emptyI Θ Γ)
then show ?case using subst_dv.simps wf_intros wf_subst1 by auto
next
case (wfD_cons Θ ℬ Γ Δ τ u)
moreover hence "u ∉ fst ` setD Δ[x::=v']⇩Δ⇩v" using subst_dv.simps subst_dv_iff using subst_dv_fst_eq by presburger
ultimately show ?case using subst_dv.simps Wellformed.wfD_cons wf_subst1 by auto
next
case (wfPhi_emptyI Θ)
then show ?case by auto
next
case (wfPhi_consI f Θ Φ ft)
then show ?case by auto
next
case (wfS_assertI Θ Φ ℬ x2 c Γ Δ s b)
show ?case unfolding subst_sv.simps proof
show ‹ Θ; Φ; ℬ; (x2, B_bool, c[x::=v']⇩c⇩v) #⇩Γ Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s[x::=v']⇩s⇩v : b ›
using wfS_assertI(4)[of "(x2, B_bool, c) #⇩Γ Γ⇩1" x ] wfS_assertI by auto
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f c[x::=v']⇩c⇩v › using wfS_assertI wf_subst1 by auto
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v › using wfS_assertI wf_subst1 by auto
show ‹atom x2 ♯ (Φ, Θ, ℬ, Γ[x::=v']⇩Γ⇩v, Δ[x::=v']⇩Δ⇩v, c[x::=v']⇩c⇩v, b, s[x::=v']⇩s⇩v)›
apply(unfold fresh_prodN,intro conjI)
apply(simp add: wfS_assertI )+
apply(metis fresh_subst_gv_if wfS_assertI)
apply(simp add: fresh_prodN fresh_subst_dv_if wfS_assertI)
apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_assertI)
apply(simp add: fresh_prodN fresh_subst_v_if subst_v_τ_def wfS_assertI)
by(simp add: fresh_prodN fresh_subst_v_if subst_v_s_def wfS_assertI)
qed
next
case (wfS_letI Θ Φ ℬ Γ Δ e b1 y s b2)
have "Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f LET y = (e[x::=v']⇩e⇩v) IN (s[x::=v']⇩s⇩v) : b2"
proof
show ‹ Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f e[x::=v']⇩e⇩v : b1 › using wfS_letI by auto
have ‹ Θ ; Φ ; ℬ ; ((y, b1, TRUE) #⇩Γ Γ)[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s[x::=v']⇩s⇩v : b2 ›
using wfS_letI(6) wfS_letI append_g.simps by metis
thus ‹ Θ ; Φ ; ℬ ; (y, b1, TRUE) #⇩Γ Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s[x::=v']⇩s⇩v : b2 ›
using wfS_letI subst_gv.simps by auto
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v › using wfS_letI by auto
show ‹atom y ♯ (Φ, Θ, ℬ, Γ[x::=v']⇩Γ⇩v, Δ[x::=v']⇩Δ⇩v, e[x::=v']⇩e⇩v, b2)›
apply(unfold fresh_prodN,intro conjI)
apply(simp add: wfS_letI )+
apply(metis fresh_subst_gv_if wfS_letI)
apply(simp add: fresh_prodN fresh_subst_dv_if wfS_letI)
apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_letI)
apply(simp add: fresh_prodN fresh_subst_v_if subst_v_τ_def wfS_letI)
done
qed
thus ?case using subst_sv.simps wfS_letI by auto
next
case (wfS_let2I Θ Φ ℬ Γ Δ s1 τ y s2 b)
have "Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f LET y : τ[x::=v']⇩τ⇩v = (s1[x::=v']⇩s⇩v) IN (s2[x::=v']⇩s⇩v) : b"
proof
show ‹ Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s1[x::=v']⇩s⇩v : b_of (τ[x::=v']⇩τ⇩v) › using wfS_let2I b_of_subst by simp
have ‹ Θ ; Φ ; ℬ ; ((y, b_of τ, TRUE) #⇩Γ Γ)[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s2[x::=v']⇩s⇩v : b ›
using wfS_let2I append_g.simps by metis
thus ‹ Θ ; Φ ; ℬ ; (y, b_of τ[x::=v']⇩τ⇩v, TRUE) #⇩Γ Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s2[x::=v']⇩s⇩v : b ›
using wfS_let2I subst_gv.simps append_g.simps using b_of_subst by simp
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f τ[x::=v']⇩τ⇩v › using wfS_let2I wf_subst1 by metis
show ‹atom y ♯ (Φ, Θ, ℬ, Γ[x::=v']⇩Γ⇩v, Δ[x::=v']⇩Δ⇩v, s1[x::=v']⇩s⇩v, b, τ[x::=v']⇩τ⇩v)›
apply(unfold fresh_prodN,intro conjI)
apply(simp add: wfS_let2I )+
apply(metis fresh_subst_gv_if wfS_let2I)
apply(simp add: fresh_prodN fresh_subst_dv_if wfS_let2I)
apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_let2I)
apply(simp add: fresh_prodN fresh_subst_v_if subst_v_τ_def wfS_let2I)+
done
qed
thus ?case using subst_sv.simps(3) subst_tv.simps wfS_let2I by auto
next
case (wfS_varI Θ ℬ Γ τ v u Φ Δ b s)
show ?case proof(subst subst_sv.simps, auto simp add: u_fresh_xv,rule)
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f τ[x::=v']⇩τ⇩v › using wfS_varI wf_subst1 by auto
have "b_of (τ[x::=v']⇩τ⇩v) = b_of τ" using b_of_subst by auto
thus ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f v[x::=v']⇩v⇩v : b_of τ[x::=v']⇩τ⇩v › using wfS_varI wf_subst1 by auto
have *:"atom u ♯ v'" using wfV_supp wfS_varI fresh_def by metis
show ‹atom u ♯ (Φ, Θ, ℬ, Γ[x::=v']⇩Γ⇩v, Δ[x::=v']⇩Δ⇩v, τ[x::=v']⇩τ⇩v, v[x::=v']⇩v⇩v, b)›
unfolding fresh_prodN apply(auto simp add: wfS_varI)
using wfS_varI fresh_subst_gv * fresh_subst_dv by metis+
show ‹ Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; (u, τ[x::=v']⇩τ⇩v) #⇩Δ Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s[x::=v']⇩s⇩v : b › using wfS_varI by auto
qed
next
case (wfS_assignI u τ Δ Θ ℬ Γ Φ v)
show ?case proof(subst subst_sv.simps, rule wf_intros)
show ‹(u, τ[x::=v']⇩τ⇩v) ∈ setD Δ[x::=v']⇩Δ⇩v› using subst_dv_iff wfS_assignI using subst_dv_fst_eq
using subst_dv_member by auto
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v › using wfS_assignI by auto
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f v[x::=v']⇩v⇩v : b_of τ[x::=v']⇩τ⇩v › using wfS_assignI b_of_subst wf_subst1 by auto
show "Θ ⊢⇩w⇩f Φ " using wfS_assignI by auto
qed
next
case (wfS_matchI Θ ℬ Γ v tid dclist Δ Φ cs b)
show ?case proof(subst subst_sv.simps, rule wf_intros)
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f v[x::=v']⇩v⇩v : B_id tid › using wfS_matchI wf_subst1 by auto
show ‹AF_typedef tid dclist ∈ set Θ› using wfS_matchI by auto
show ‹ Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ; tid ; dclist ⊢⇩w⇩f subst_branchlv cs x v' : b › using wfS_matchI by simp
show "Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v " using wfS_matchI by auto
show "Θ ⊢⇩w⇩f Φ " using wfS_matchI by auto
qed
next
case (wfS_branchI Θ Φ ℬ y τ Γ Δ s b tid dc)
have " Θ ; Φ ; ℬ ; Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ; tid ; dc ; τ ⊢⇩w⇩f dc y ⇒ (s[x::=v']⇩s⇩v) : b"
proof
have ‹ Θ ; Φ ; ℬ ; ((y, b_of τ, TRUE) #⇩Γ Γ)[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s[x::=v']⇩s⇩v : b ›
using wfS_branchI append_g.simps by metis
thus ‹ Θ ; Φ ; ℬ ; (y, b_of τ, TRUE) #⇩Γ Γ[x::=v']⇩Γ⇩v ; Δ[x::=v']⇩Δ⇩v ⊢⇩w⇩f s[x::=v']⇩s⇩v : b ›
using subst_gv.simps b_of_subst wfS_branchI by simp
show ‹atom y ♯ (Φ, Θ, ℬ, Γ[x::=v']⇩Γ⇩v, Δ[x::=v']⇩Δ⇩v, Γ[x::=v']⇩Γ⇩v, τ)›
apply(unfold fresh_prodN,intro conjI)
apply(simp add: wfS_branchI )+
apply(metis fresh_subst_gv_if wfS_branchI)
apply(simp add: fresh_prodN fresh_subst_dv_if wfS_branchI)
apply(metis fresh_subst_gv_if wfS_branchI)+
done
show ‹ Θ; ℬ; Γ[x::=v']⇩Γ⇩v ⊢⇩w⇩f Δ[x::=v']⇩Δ⇩v › using wfS_branchI by auto
qed
thus ?case using subst_branchv.simps wfS_branchI by auto
next
case (wfS_finalI Θ Φ ℬ Γ Δ tid dclist' cs b dclist)
then show ?case using subst_branchlv.simps wf_intros by metis
next
case (wfS_cons Θ Φ ℬ Γ Δ tid dclist' cs b css dclist)
then show ?case using subst_branchlv.simps wf_intros by metis
qed(metis subst_sv.simps wf_subst1 wf_intros)+
lemmas wf_subst = wf_subst1 wf_subst2
lemma wfG_subst_wfV:
assumes "Θ; ℬ ⊢⇩w⇩f Γ' @ (x, b, c0[z0::=V_var x]⇩c⇩v) #⇩Γ Γ" and "wfV Θ ℬ Γ v b"
shows "Θ; ℬ ⊢⇩w⇩f Γ'[x::=v]⇩Γ⇩v @ Γ "
using assms wf_subst subst_g_inside_simple by auto
lemma wfG_member_subst:
assumes "(x1,b1,c1) ∈ toSet (Γ'@Γ)" and "wfG Θ ℬ (Γ'@((x,b,c) #⇩ΓΓ))" and "x ≠ x1"
shows "∃c1'. (x1,b1,c1') ∈ toSet ((Γ'[x::=v]⇩Γ⇩v)@Γ)"
proof -
consider (lhs) "(x1,b1,c1) ∈ toSet Γ'" | (rhs) "(x1,b1,c1) ∈ toSet Γ" using append_g_toSetU assms by auto
thus ?thesis proof(cases)
case lhs
hence "(x1,b1,c1[x::=v]⇩c⇩v) ∈ toSet (Γ'[x::=v]⇩Γ⇩v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis
hence "(x1,b1,c1[x::=v]⇩c⇩v) ∈ toSet (Γ'[x::=v]⇩Γ⇩v@Γ)" using append_g_toSetU by auto
then show ?thesis by auto
next
case rhs
hence "(x1,b1,c1) ∈ toSet (Γ'[x::=v]⇩Γ⇩v@Γ)" using append_g_toSetU by auto
then show ?thesis by auto
qed
qed
lemma wfG_member_subst2:
assumes "(x1,b1,c1) ∈ toSet (Γ'@((x,b,c) #⇩ΓΓ))" and "wfG Θ ℬ (Γ'@((x,b,c) #⇩ΓΓ))" and "x ≠ x1"
shows "∃c1'. (x1,b1,c1') ∈ toSet ((Γ'[x::=v]⇩Γ⇩v)@Γ)"
proof -
consider (lhs) "(x1,b1,c1) ∈ toSet Γ'" | (rhs) "(x1,b1,c1) ∈ toSet Γ" using append_g_toSetU assms by auto
thus ?thesis proof(cases)
case lhs
hence "(x1,b1,c1[x::=v]⇩c⇩v) ∈ toSet (Γ'[x::=v]⇩Γ⇩v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis
hence "(x1,b1,c1[x::=v]⇩c⇩v) ∈ toSet (Γ'[x::=v]⇩Γ⇩v@Γ)" using append_g_toSetU by auto
then show ?thesis by auto
next
case rhs
hence "(x1,b1,c1) ∈ toSet (Γ'[x::=v]⇩Γ⇩v@Γ)" using append_g_toSetU by auto
then show ?thesis by auto
qed
qed
lemma wbc_subst:
fixes Γ::Γ and Γ'::Γ and v::v
assumes "wfC Θ ℬ (Γ'@((x,b,c') #⇩ΓΓ)) c" and "Θ; ℬ; Γ ⊢⇩w⇩f v : b"
shows "Θ; ℬ; ((Γ'[x::=v]⇩Γ⇩v)@Γ) ⊢⇩w⇩f c[x::=v]⇩c⇩v"
proof -
have "(Γ'@((x,b,c') #⇩ΓΓ))[x::=v]⇩Γ⇩v = ((Γ'[x::=v]⇩Γ⇩v)@Γ)" using assms subst_g_inside_simple wfC_wf by metis
thus ?thesis using wf_subst1(2)[OF assms(1) _ assms(2)] by metis
qed
lemma wfG_inside_fresh_suffix:
assumes "wfG P B (Γ'@(x,b,c) #⇩ΓΓ)"
shows "atom x ♯ Γ"
proof -
have "wfG P B ((x,b,c) #⇩ΓΓ)" using wfG_suffix assms by auto
thus ?thesis using wfG_elims by metis
qed
lemmas wf_b_subst_lemmas = subst_eb.simps wf_intros
forget_subst subst_b_b_def subst_b_v_def subst_b_ce_def fresh_e_opp_all subst_bb.simps wfV_b_fresh ms_fresh_all(6)
lemma wf_b_subst1:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows "Θ ; B' ; Γ ⊢⇩w⇩f v : b' ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v[bv::=b]⇩v⇩b : b'[bv::=b]⇩b⇩b" and
"Θ ; B' ; Γ ⊢⇩w⇩f c ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f c[bv::=b]⇩c⇩b" and
"Θ ; B' ⊢⇩w⇩f Γ ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; B ⊢⇩w⇩f Γ[bv::=b]⇩Γ⇩b" and
"Θ ; B' ; Γ ⊢⇩w⇩f τ ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f τ[bv::=b]⇩τ⇩b" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f Θ ⟹True" and
"Θ ; B' ⊢⇩w⇩f b' ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; B ⊢⇩w⇩f b'[bv::=b]⇩b⇩b " and
"Θ ; B' ; Γ ⊢⇩w⇩f ce : b' ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f ce[bv::=b]⇩c⇩e⇩b : b'[bv::=b]⇩b⇩b" and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(nominal_induct
b' and c and Γ and τ and ts and Θ and b' and b' and td
avoiding: bv b B
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfB_intI Θ ℬ)
then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
next
case (wfB_boolI Θ ℬ)
then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
next
case (wfB_unitI Θ ℬ)
then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
next
case (wfB_bitvecI Θ ℬ)
then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
next
case (wfB_pairI Θ ℬ b1 b2)
then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
next
case (wfB_consI Θ s dclist ℬ)
then show ?case using subst_bb.simps Wellformed.wfB_consI by simp
next
case (wfB_appI Θ ba s bva dclist ℬ)
then show ?case using subst_bb.simps Wellformed.wfB_appI forget_subst wfB_supp
by (metis bot.extremum_uniqueI ex_in_conv fresh_def subst_b_b_def supp_empty_fset)
next
case (wfV_varI Θ ℬ1 Γ b1 c x)
show ?case unfolding subst_vb.simps proof
show "Θ ; B ⊢⇩w⇩f Γ[bv::=b]⇩Γ⇩b " using wfV_varI by auto
show "Some (b1[bv::=b]⇩b⇩b, c[bv::=b]⇩c⇩b) = lookup Γ[bv::=b]⇩Γ⇩b x" using subst_b_lookup wfV_varI by simp
qed
next
case (wfV_litI Θ ℬ Γ l)
then show ?case using Wellformed.wfV_litI subst_b_base_for_lit by simp
next
case (wfV_pairI Θ ℬ1 Γ v1 b1 v2 b2)
show ?case unfolding subst_vb.simps proof(subst subst_bb.simps,rule)
show "Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v1[bv::=b]⇩v⇩b : b1[bv::=b]⇩b⇩b" using wfV_pairI by simp
show "Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v2[bv::=b]⇩v⇩b : b2[bv::=b]⇩b⇩b " using wfV_pairI by simp
qed
next
case (wfV_consI s dclist Θ dc x b' c ℬ' Γ v)
show ?case unfolding subst_vb.simps proof(subst subst_bb.simps, rule Wellformed.wfV_consI)
show 1:"AF_typedef s dclist ∈ set Θ" using wfV_consI by auto
show 2:"(dc, ⦃ x : b' | c ⦄) ∈ set dclist" using wfV_consI by auto
have "Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v[bv::=b]⇩v⇩b : b'[bv::=b]⇩b⇩b" using wfV_consI by auto
moreover hence "supp b' = {}" using 1 2 wfTh_lookup_supp_empty τ.supp wfX_wfY by blast
moreover hence "b'[bv::=b]⇩b⇩b = b'" using forget_subst subst_bb_def fresh_def by (metis empty_iff subst_b_b_def)
ultimately show "Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v[bv::=b]⇩v⇩b : b'" using wfV_consI by simp
qed
next
case (wfV_conspI s bva dclist Θ dc x b' c ℬ' ba Γ v)
have *:"atom bv ♯ b'" using wfTh_poly_supp_b[of s bva dclist Θ dc x b' c] fresh_def wfX_wfY ‹atom bva ♯ bv›
by (metis insert_iff not_self_fresh singleton_insert_inj_eq' subsetI subset_antisym wfV_conspI wfV_conspI.hyps(4) wfV_conspI.prems(2))
show ?case unfolding subst_vb.simps subst_bb.simps proof
show ‹AF_typedef_poly s bva dclist ∈ set Θ› using wfV_conspI by auto
show ‹(dc, ⦃ x : b' | c ⦄) ∈ set dclist› using wfV_conspI by auto
thus ‹ Θ ; B ⊢⇩w⇩f ba[bv::=b]⇩b⇩b › using wfV_conspI by metis
have "atom bva ♯ Γ[bv::=b]⇩Γ⇩b" using fresh_subst_if subst_b_Γ_def wfV_conspI by metis
moreover have "atom bva ♯ ba[bv::=b]⇩b⇩b" using fresh_subst_if subst_b_b_def wfV_conspI by metis
moreover have "atom bva ♯ v[bv::=b]⇩v⇩b" using fresh_subst_if subst_b_v_def wfV_conspI by metis
ultimately show ‹atom bva ♯ (Θ, B, Γ[bv::=b]⇩Γ⇩b, ba[bv::=b]⇩b⇩b, v[bv::=b]⇩v⇩b)›
unfolding fresh_prodN using wfV_conspI fresh_def supp_fset by auto
show ‹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v[bv::=b]⇩v⇩b : b'[bva::=ba[bv::=b]⇩b⇩b]⇩b⇩b ›
using wfV_conspI subst_bb_commute[of bv b' bva ba b] * wfV_conspI by metis
qed
next
case (wfTI z Θ ℬ' Γ' b' c)
show ?case proof(subst subst_tb.simps, rule Wellformed.wfTI)
show "atom z ♯ (Θ, B, Γ'[bv::=b]⇩Γ⇩b)" using wfTI subst_g_b_x_fresh by simp
show "Θ ; B ⊢⇩w⇩f b'[bv::=b]⇩b⇩b " using wfTI by auto
show "Θ ; B ; (z, b'[bv::=b]⇩b⇩b, TRUE) #⇩Γ Γ'[bv::=b]⇩Γ⇩b ⊢⇩w⇩f c[bv::=b]⇩c⇩b " using wfTI by simp
qed
next
case (wfC_eqI Θ ℬ' Γ e1 b' e2)
thus ?case using Wellformed.wfC_eqI subst_db.simps subst_cb.simps wfC_eqI by metis
next
case (wfG_nilI Θ ℬ')
then show ?case using Wellformed.wfG_nilI subst_gb.simps by simp
next
case (wfG_cons1I c' Θ ℬ' Γ' x b')
show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons1I)
show "c'[bv::=b]⇩c⇩b ∉ {TRUE, FALSE}" using wfG_cons1I(1)
by(nominal_induct c' rule: c.strong_induct,auto+)
show "Θ ; B ⊢⇩w⇩f Γ'[bv::=b]⇩Γ⇩b " using wfG_cons1I by auto
show "atom x ♯ Γ'[bv::=b]⇩Γ⇩b" using wfG_cons1I subst_g_b_x_fresh by auto
show "Θ ; B ; (x, b'[bv::=b]⇩b⇩b, TRUE) #⇩Γ Γ'[bv::=b]⇩Γ⇩b ⊢⇩w⇩f c'[bv::=b]⇩c⇩b" using wfG_cons1I by auto
show "Θ ; B ⊢⇩w⇩f b'[bv::=b]⇩b⇩b " using wfG_cons1I by auto
qed
next
case (wfG_cons2I c' Θ ℬ' Γ' x b')
show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons2I)
show "c'[bv::=b]⇩c⇩b ∈ {TRUE, FALSE}" using wfG_cons2I by auto
show "Θ ; B ⊢⇩w⇩f Γ'[bv::=b]⇩Γ⇩b " using wfG_cons2I by auto
show "atom x ♯ Γ'[bv::=b]⇩Γ⇩b" using wfG_cons2I subst_g_b_x_fresh by auto
show "Θ ; B ⊢⇩w⇩f b'[bv::=b]⇩b⇩b " using wfG_cons2I by auto
qed
next
case (wfCE_valI Θ ℬ Γ v b)
then show ?case using subst_ceb.simps wf_intros wfX_wfY
by (metis wf_b_subst_lemmas wfCE_b_fresh)
next
case (wfCE_plusI Θ ℬ Γ v1 v2)
then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
by metis
next
case (wfCE_leqI Θ ℬ Γ v1 v2)
then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
by metis
next
case (wfCE_eqI Θ ℬ Γ v1 b v2)
then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
by metis
next
case (wfCE_fstI Θ ℬ Γ v1 b1 b2)
then show ?case
by (metis (no_types) subst_bb.simps(5) subst_ceb.simps(3) wfCE_fstI.hyps(2)
wfCE_fstI.prems(1) wfCE_fstI.prems(2) Wellformed.wfCE_fstI)
next
case (wfCE_sndI Θ ℬ Γ v1 b1 b2)
then show ?case
by (metis (no_types) subst_bb.simps(5) subst_ceb.simps wfCE_sndI.hyps(2)
wfCE_sndI wfCE_sndI.prems(2) Wellformed.wfCE_sndI)
next
case (wfCE_concatI Θ ℬ Γ v1 v2)
then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh
proof -
show ?thesis
using wfCE_concatI.hyps(2) wfCE_concatI.hyps(4) wfCE_concatI.prems(1) wfCE_concatI.prems(2)
Wellformed.wfCE_concatI by auto
qed
next
case (wfCE_lenI Θ ℬ Γ v1)
then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh by metis
qed(auto simp add: wf_intros)
lemma wf_b_subst2:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def
and cs::branch_s and css::branch_list
shows "Θ ; Φ ; B' ; Γ ; Δ ⊢⇩w⇩f e : b' ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; Φ ; B ; Γ[bv::=b]⇩Γ⇩b ; Δ[bv::=b]⇩Δ⇩b ⊢⇩w⇩f e[bv::=b]⇩e⇩b : b'[bv::=b]⇩b⇩b" and
"Θ ; Φ ; ℬ ; Γ ; Δ ⊢⇩w⇩f s : b ⟹ True" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dc ; t ⊢⇩w⇩f cs : b ⟹ True" and
"Θ ; Φ ; ℬ ; Γ ; Δ ; tid ; dclist ⊢⇩w⇩f css : b ⟹ True" and
"Θ ⊢⇩w⇩f (Φ::Φ) ⟹ True " and
"Θ ; B' ; Γ ⊢⇩w⇩f Δ ⟹ {|bv|} = B' ⟹ Θ ; B ⊢⇩w⇩f b ⟹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f Δ[bv::=b]⇩Δ⇩b" and
"Θ ; Φ ⊢⇩w⇩f ftq ⟹ True" and
"Θ ; Φ ; ℬ ⊢⇩w⇩f ft ⟹ True"
proof(nominal_induct
b' and b and b and b and Φ and Δ and ftq and ft
avoiding: bv b B
rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
case (wfE_valI Θ' Φ' ℬ' Γ' Δ' v' b')
then show ?case unfolding subst_vb.simps subst_eb.simps using wf_b_subst1(1) Wellformed.wfE_valI by auto
next
case (wfE_plusI Θ Φ ℬ Γ Δ v1 v2)
then show ?case unfolding subst_eb.simps
using wf_b_subst_lemmas wf_b_subst1(1) Wellformed.wfE_plusI
proof -
have "∀b ba v g f ts. (( ts ; f ; g[bv::=ba]⇩Γ⇩b ⊢⇩w⇩f v[bv::=ba]⇩v⇩b : b[bv::=ba]⇩b⇩b) ∨ ¬ ts ; ℬ ; g ⊢⇩w⇩f v : b ) ∨ ¬ ts ; f ⊢⇩w⇩f ba"
using wfE_plusI.prems(1) wf_b_subst1(1) by force
then show "Θ ; Φ ; B ; Γ[bv::=b]⇩Γ⇩b ; Δ[bv::=b]⇩Δ⇩b ⊢⇩w⇩f [ plus v1[bv::=b]⇩v⇩b v2[bv::=b]⇩v⇩b ]⇧e : B_int[bv::=b]⇩b⇩b"
by (metis wfE_plusI.hyps(1) wfE_plusI.hyps(4) wfE_plusI.hyps(5) wfE_plusI.hyps(6) wfE_plusI.prems(1) wfE_plusI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_plusI wf_b_subst_lemmas(86))
qed
next
case (wfE_leqI Θ Φ ℬ Γ Δ v1 v2)
then show ?case unfolding subst_eb.simps
using wf_b_subst_lemmas wf_b_subst1 Wellformed.wfE_leqI
proof -
have "⋀ts f b ba g v. ¬ (ts ; f ⊢⇩w⇩f b) ∨ ¬ (ts ; {|ba|} ; g ⊢⇩w⇩f v : B_int) ∨ (ts ; f ; g[ba::=b]⇩Γ⇩b ⊢⇩w⇩f v[ba::=b]⇩v⇩b : B_int)"
by (metis wf_b_subst1(1) wf_b_subst_lemmas(86))
then show "Θ ; Φ ; B ; Γ[bv::=b]⇩Γ⇩b ; Δ[bv::=b]⇩Δ⇩b ⊢⇩w⇩f [ leq v1[bv::=b]⇩v⇩b v2[bv::=b]⇩v⇩b ]⇧e : B_bool[bv::=b]⇩b⇩b"
by (metis (no_types) wfE_leqI.hyps(1) wfE_leqI.hyps(4) wfE_leqI.hyps(5) wfE_leqI.hyps(6) wfE_leqI.prems(1) wfE_leqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_leqI wf_b_subst_lemmas(87))
qed
next
case (wfE_eqI Θ Φ ℬ Γ Δ v1 bb v2)
show ?case unfolding subst_eb.simps subst_bb.simps proof
show ‹ Θ ⊢⇩w⇩f Φ › using wfX_wfY wfE_eqI by metis
show ‹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f Δ[bv::=b]⇩Δ⇩b › using wfX_wfY wfE_eqI by metis
show ‹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v1[bv::=b]⇩v⇩b : bb › using subst_bb.simps wfE_eqI
by (metis (no_types, opaque_lifting) empty_iff insert_iff wf_b_subst1(1))
show ‹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v2[bv::=b]⇩v⇩b : bb › using wfX_wfY wfE_eqI
by (metis insert_iff singleton_iff wf_b_subst1(1) wf_b_subst_lemmas(86) wf_b_subst_lemmas(87) wf_b_subst_lemmas(90))
show ‹bb ∈ {B_bool, B_int, B_unit}› using wfE_eqI by auto
qed
next
case (wfE_fstI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(84) wf_b_subst1(1) Wellformed.wfE_fstI
by (metis wf_b_subst_lemmas(89))
next
case (wfE_sndI Θ Φ ℬ Γ Δ v1 b1 b2)
then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_sndI
by (metis wf_b_subst_lemmas(89))
next
case (wfE_concatI Θ Φ ℬ Γ Δ v1 v2)
then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_concatI
by (metis wf_b_subst_lemmas(91))
next
case (wfE_splitI Θ Φ ℬ Γ Δ v1 v2)
then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_splitI
by (metis wf_b_subst_lemmas(89) wf_b_subst_lemmas(91))
next
case (wfE_lenI Θ Φ ℬ Γ Δ v1)
then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_lenI
by (metis wf_b_subst_lemmas(91) wf_b_subst_lemmas(89))
next
case (wfE_appI Θ Φ ℬ' Γ Δ f x b' c τ s v)
hence bf: "atom bv ♯ b'" using wfPhi_f_simple_wfT wfT_supp bv_not_in_dom_g wfPhi_f_simple_supp_b fresh_def by fast
hence bseq: "b'[bv::=b]⇩b⇩b = b'" using subst_bb.simps wf_b_subst_lemmas by metis
have "Θ ; Φ ; B ; Γ[bv::=b]⇩Γ⇩b ; Δ[bv::=b]⇩Δ⇩b ⊢⇩w⇩f (AE_app f (v[bv::=b]⇩v⇩b)) : (b_of (τ[bv::=b]⇩τ⇩b))"
proof
show "Θ ⊢⇩w⇩f Φ" using wfE_appI by auto
show "Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f Δ[bv::=b]⇩Δ⇩b " using wfE_appI by simp
have "atom bv ♯ τ" using wfPhi_f_simple_wfT[OF wfE_appI(5) wfE_appI(1),THEN wfT_supp] bv_not_in_dom_g fresh_def by force
hence " τ[bv::=b]⇩τ⇩b = τ" using forget_subst subst_b_τ_def by metis
thus "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c τ[bv::=b]⇩τ⇩b s))) = lookup_fun Φ f" using wfE_appI by simp
show "Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v[bv::=b]⇩v⇩b : b'" using wfE_appI bseq wf_b_subst1 by metis
qed
then show ?case using subst_eb.simps b_of_subst_bb_commute by simp
next
case (wfE_appPI Θ Φ ℬ Γ Δ b' bv1 v1 τ1 f x1 b1 c1 s1)
then have *: "atom bv ♯ b1" using wfPhi_f_supp(1) wfE_appPI(7,11)
by (metis fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps(1))
have "Θ ; Φ ; B ; Γ[bv::=b]⇩Γ⇩b ; Δ[bv::=b]⇩Δ⇩b ⊢⇩w⇩f AE_appP f b'[bv::=b]⇩b⇩b (v1[bv::=b]⇩v⇩b) : (b_of τ1)[bv1::=b'[bv::=b]⇩b⇩b]⇩b"
proof
show ‹ Θ ⊢⇩w⇩f Φ › using wfE_appPI by auto
show ‹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f Δ[bv::=b]⇩Δ⇩b › using wfE_appPI by auto
show ‹ Θ ; B ⊢⇩w⇩f b'[bv::=b]⇩b⇩b › using wfE_appPI wf_b_subst1 by auto
have "atom bv1 ♯ Γ[bv::=b]⇩Γ⇩b" using fresh_subst_if subst_b_Γ_def wfE_appPI by metis
moreover have "atom bv1 ♯ b'[bv::=b]⇩b⇩b" using fresh_subst_if subst_b_b_def wfE_appPI by metis
moreover have "atom bv1 ♯ v1[bv::=b]⇩v⇩b" using fresh_subst_if subst_b_v_def wfE_appPI by metis
moreover have "atom bv1 ♯ Δ[bv::=b]⇩Δ⇩b" using fresh_subst_if subst_b_Δ_def wfE_appPI by metis
moreover have "atom bv1 ♯ (b_of τ1)[bv1::=b'[bv::=b]⇩b⇩b]⇩b⇩b" using fresh_subst_if subst_b_b_def wfE_appPI by metis
ultimately show "atom bv1 ♯ (Φ, Θ, B, Γ[bv::=b]⇩Γ⇩b, Δ[bv::=b]⇩Δ⇩b, b'[bv::=b]⇩b⇩b, v1[bv::=b]⇩v⇩b, (b_of τ1)[bv1::=b'[bv::=b]⇩b⇩b]⇩b)"
using wfE_appPI using fresh_def fresh_prodN subst_b_b_def by metis
show ‹Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 τ1 s1))) = lookup_fun Φ f› using wfE_appPI by auto
have ‹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v1[bv::=b]⇩v⇩b : b1[bv1::=b']⇩b[bv::=b]⇩b⇩b ›
using wfE_appPI subst_b_b_def * wf_b_subst1 by metis
thus ‹ Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f v1[bv::=b]⇩v⇩b : b1[bv1::=b'[bv::=b]⇩b⇩b]⇩b ›
using subst_bb_commute subst_b_b_def * by auto
qed
moreover have "atom bv ♯ b_of τ1" proof -
have "supp (b_of τ1) ⊆ { atom bv1 }" using wfPhi_f_poly_supp_b_of_t
using b_of.simps wfE_appPI wfPhi_f_supp(5) by simp
thus ?thesis using wfE_appPI
fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps by metis
qed
ultimately show ?case using subst_eb.simps(3) subst_bb_commute subst_b_b_def * by simp
next
case (wfE_mvarI Θ Φ ℬ' Γ Δ u τ)
have "Θ ; Φ ; B ; subst_gb Γ bv b ; subst_db Δ bv b ⊢⇩w⇩f (AE_mvar u)[bv::=b]⇩e⇩b : (b_of (τ[bv::=b]⇩τ⇩b))"
proof(subst subst_eb.simps,rule Wellformed.wfE_mvarI)
show "Θ ⊢⇩w⇩f Φ " using wfE_mvarI by simp
show "Θ ; B ; Γ[bv::=b]⇩Γ⇩b ⊢⇩w⇩f Δ[bv::=b]⇩Δ⇩b" using wfE_mvarI by metis
show "(u, τ[bv::=b]⇩τ⇩b) ∈ setD Δ[bv::=b]⇩Δ⇩b"
using wfE_mvarI subst_db.simps set_insert subst_d_b_member by simp
qed
thus ?case using b_of_subst_bb_commute by auto
next
case (wfS_seqI Θ Φ ℬ Γ Δ s1 s2 b)
then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
next
case (wfD_emptyI Θ ℬ' Γ)
then show ?case using subst_db.simps Wellformed.wfD_emptyI wf_b_subst1 by simp
next
case (wfD_cons Θ ℬ' Γ' Δ τ u)
show ?case proof(subst subst_db.simps, rule Wellformed.wfD_cons )
show "Θ ; B ; Γ'[bv::=b]⇩Γ⇩b ⊢⇩w⇩f Δ[bv::=b]⇩Δ⇩b " using wfD_cons by auto
show "Θ ; B ; Γ'[bv::=b]⇩Γ⇩b ⊢⇩w⇩f τ[bv::=b]⇩τ⇩b " using wfD_cons wf_b_subst1 by auto
show "u ∉ fst ` setD Δ[bv::=b]⇩Δ⇩b" using wfD_cons subst_b_lookup_d by metis
qed
next
case (wfS_assertI Θ Φ ℬ x c Γ Δ s b)
show ?case by auto
qed(auto)
lemmas wf_b_subst = wf_b_subst1 wf_b_subst2
lemma wfT_subst_wfT:
fixes τ::τ and b'::b and bv::bv
assumes "Θ ; {|bv|} ; (x,b,c) #⇩ΓGNil ⊢⇩w⇩f τ" and "Θ ; B ⊢⇩w⇩f b'"
shows "Θ ; B ; (x,b[bv::=b']⇩b⇩b,c[bv::=b']⇩c⇩b) #⇩ΓGNil ⊢⇩w⇩f (τ[bv::=b']⇩τ⇩b)"
proof -
have "Θ ; B ; ((x,b,c) #⇩ΓGNil)[bv::=b']⇩Γ⇩b ⊢⇩w⇩f (τ[bv::=b']⇩τ⇩b)"
using wf_b_subst assms by metis
thus ?thesis using subst_gb.simps wf_b_subst_lemmas wfCE_b_fresh by metis
qed
lemma wf_trans:
fixes Γ::Γ and Γ'::Γ and v::v and e::e and c::c and τ::τ and ts::"(string*τ) list" and Δ::Δ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s
and cs::branch_s and css::branch_list and Θ::Θ
shows "Θ; ℬ; Γ ⊢⇩w⇩f v : b' ⟹ Γ = (x, b, c2) #⇩Γ G ⟹ Θ; ℬ; (x, b, c1) #⇩Γ G ⊢⇩w⇩f c2 ⟹ Θ; ℬ; (x, b, c1) #⇩Γ G ⊢⇩w⇩f v : b'" and
"Θ; ℬ; Γ ⊢⇩w⇩f c ⟹ Γ = (x, b, c2) #⇩Γ G ⟹ Θ; ℬ; (x, b, c1) #⇩Γ G ⊢⇩w⇩f c2 ⟹ Θ; ℬ; (x, b, c1) #⇩Γ G ⊢⇩w⇩f c" and
"Θ; ℬ ⊢⇩w⇩f Γ ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f τ ⟹ True" and
"Θ; ℬ; Γ ⊢⇩w⇩f ts ⟹ True" and
"⊢⇩w⇩f Θ ⟹True" and
"Θ; ℬ ⊢⇩w⇩f b ⟹ True " and
"Θ; ℬ; Γ ⊢⇩w⇩f ce : b' ⟹ Γ = (x, b, c2) #⇩Γ G ⟹ Θ; ℬ; (x, b, c1) #⇩Γ G ⊢⇩w⇩f c2 ⟹ Θ; ℬ; (x, b, c1) #⇩Γ G ⊢⇩w⇩f ce : b' " and
"Θ ⊢⇩w⇩f td ⟹ True"
proof(nominal_induct
b' and c and Γ and τ and ts and Θ and b and b' and td
avoiding: c1
arbitrary: Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1 and Γ⇩1
rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
case (wfV_varI Θ ℬ Γ b' c' x')
have wbg: "Θ; ℬ ⊢⇩w⇩f (x, b, c1) #⇩Γ G " using wfC_wf wfV_varI by simp
show ?case proof(cases "x=x'")
case True
have "Some (b', c1) = lookup ((x, b, c1) #⇩Γ G) x'" using lookup.simps wfV_varI using True by auto
then show ?thesis using Wellformed.wfV_varI wbg by simp
next
case False
then have "Some (b', c') = lookup ((x, b, c1) #⇩Γ G) x'" using lookup.simps wfV_varI
by simp
then show ?thesis using Wellformed.wfV_varI wbg by simp
qed
next
case (wfV_conspI s bv dclist Θ dc x1 b' c ℬ b1 Γ v)
show ?case proof
show ‹AF_typedef_poly s bv dclist ∈ set Θ› using wfV_conspI by auto
show ‹(dc, ⦃ x1 : b' | c ⦄) ∈ set dclist› using wfV_conspI by auto
show ‹Θ; ℬ ⊢⇩w⇩f b1 › using wfV_conspI by auto
show ‹atom bv ♯ (Θ, ℬ, (x, b, c1) #⇩Γ G, b1, v)› unfolding fresh_prodN fresh_GCons using wfV_conspI fresh_prodN fresh_GCons by simp
show ‹Θ; ℬ; (x, b, c1) #⇩Γ G ⊢⇩w⇩f v : b'[bv::=b1]⇩b⇩b› using wfV_conspI by auto
qed
qed( (auto | metis wfC_wf wf_intros) +)
end