diff --git a/thys/MiniSail/BTVSubst.thy b/thys/MiniSail/BTVSubst.thy --- a/thys/MiniSail/BTVSubst.thy +++ b/thys/MiniSail/BTVSubst.thy @@ -1,1550 +1,1550 @@ (*<*) theory BTVSubst imports Syntax begin (*>*) chapter \Basic Type Variable Substitution\ section \Class\ class has_subst_b = fs + fixes subst_b :: "'a::fs \ bv \ b \ 'a::fs" ("_[_::=_]\<^sub>b" [1000,50,50] 1000) assumes fresh_subst_if: "j \ (t[i::=x]\<^sub>b ) \ (atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i))" and forget_subst[simp]: "atom a \ tm \ tm[a::=x]\<^sub>b = tm" and subst_id[simp]: "tm[a::=(B_var a)]\<^sub>b = tm" and eqvt[simp,eqvt]: "(p::perm) \ (subst_b t1 x1 v ) = (subst_b (p \t1) (p \x1) (p \v) )" and flip_subst[simp]: "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" and flip_subst_subst[simp]: "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" begin lemmas flip_subst_b = flip_subst_subst lemma subst_b_simple_commute: fixes x::bv assumes "atom x \ c" shows "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = c[z::=b]\<^sub>b" proof - have "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (( x \ z) \ c)[x::=b]\<^sub>b" using flip_subst assms by simp thus ?thesis using flip_subst_subst assms by simp qed lemma subst_b_flip_eq_one: fixes z1::bv and z2::bv and x1::bv and x2::bv assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" and "atom x1 \ (z1,z2,c1,c2)" shows "(c1[z1::=B_var x1]\<^sub>b) = (c2[z2::=B_var x1]\<^sub>b)" proof - have "(c1[z1::=B_var x1]\<^sub>b) = (x1 \ z1) \ c1" using assms flip_subst by auto moreover have "(c2[z2::=B_var x1]\<^sub>b) = (x1 \ z2) \ c2" using assms flip_subst by auto ultimately show ?thesis using Abs1_eq_iff_all(3)[of z1 c1 z2 c2 z1] assms by (metis Abs1_eq_iff_fresh(3) flip_commute) qed lemma subst_b_flip_eq_two: fixes z1::bv and z2::bv and x1::bv and x2::bv assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" shows "(c1[z1::=b]\<^sub>b) = (c2[z2::=b]\<^sub>b)" proof - obtain x::bv where *:"atom x \ (z1,z2,c1,c2)" using obtain_fresh by metis hence "(c1[z1::=B_var x]\<^sub>b) = (c2[z2::=B_var x]\<^sub>b)" using subst_b_flip_eq_one[OF assms, of x] by metis hence "(c1[z1::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (c2[z2::=B_var x]\<^sub>b)[x::=b]\<^sub>b" by auto thus ?thesis using subst_b_simple_commute * fresh_prod4 by metis qed lemma subst_b_fresh_x: fixes tm::"'a::fs" and x::x shows "atom x \ tm = atom x \ tm[bv::=b']\<^sub>b" using fresh_subst_if[of "atom x" tm bv b'] using x_fresh_b by auto lemma subst_b_x_flip[simp]: fixes x'::x and x::x and bv::bv shows "((x' \ x) \ tm)[bv::=b']\<^sub>b = (x' \ x) \ (tm[bv::=b']\<^sub>b)" proof - have "(x' \ x) \ bv = bv" using pure_supp flip_fresh_fresh by force moreover have "(x' \ x) \ b' = b'" using x_fresh_b flip_fresh_fresh by auto ultimately show ?thesis using eqvt by simp qed end section \Base Type\ nominal_function subst_bb :: "b \ bv \ b \ b" where "subst_bb (B_var bv2) bv1 b = (if bv1 = bv2 then b else (B_var bv2))" | "subst_bb B_int bv1 b = B_int" | "subst_bb B_bool bv1 b = B_bool" | "subst_bb (B_id s) bv1 b = B_id s " | "subst_bb (B_pair b1 b2) bv1 b = B_pair (subst_bb b1 bv1 b) (subst_bb b2 bv1 b)" | "subst_bb B_unit bv1 b = B_unit " | "subst_bb B_bitvec bv1 b = B_bitvec " | "subst_bb (B_app s b2) bv1 b = B_app s (subst_bb b2 bv1 b)" apply (simp add: eqvt_def subst_bb_graph_aux_def ) apply (simp add: eqvt_def subst_bb_graph_aux_def ) by (auto,meson b.strong_exhaust) nominal_termination (eqvt) by lexicographic_order abbreviation subst_bb_abbrev :: "b \ bv \ b \ b" ("_[_::=_]\<^sub>b\<^sub>b" [1000,50,50] 1000) where "b[bv::=b']\<^sub>b\<^sub>b \ subst_bb b bv b' " instantiation b :: has_subst_b begin definition "subst_b = subst_bb" instance proof fix j::atom and i::bv and x::b and t::b show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" proof (induct t rule: b.induct) case (B_id x) then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto next case (B_var x) then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto next case (B_app x1 x2) then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto qed(auto simp add: subst_bb.simps fresh_def pure_fresh subst_b_b_def)+ fix a::bv and tm::b and x::b show "atom a \ tm \ tm[a::=x]\<^sub>b = tm" by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def) fix a::bv and tm::b show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def) fix p::perm and x1::bv and v::b and t1::b show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def) fix bv::bv and c::b and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+) fix bv::bv and c::b and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+) qed end lemma subst_bb_inject: assumes "b1 = b2[bv::=b]\<^sub>b\<^sub>b" and "b2 \ B_var bv" shows "b1 = B_int \ b2 = B_int" and "b1 = B_bool \ b2 = B_bool" and "b1 = B_unit \ b2 = B_unit" and "b1 = B_bitvec \ b2 = B_bitvec" and "b1 = B_pair b11 b12 \ (\b11' b12' . b11 = b11'[bv::=b]\<^sub>b\<^sub>b \ b12 = b12'[bv::=b]\<^sub>b\<^sub>b \ b2 = B_pair b11' b12')" and "b1 = B_var bv' \ b2 = B_var bv'" and "b1 = B_id tyid \ b2 = B_id tyid" and "b1 = B_app tyid b11 \ (\b11'. b11= b11'[bv::=b]\<^sub>b\<^sub>b \ b2 = B_app tyid b11')" using assms by(nominal_induct b2 rule:b.strong_induct,auto+) lemma flip_b_subst4: fixes b1::b and bv1::bv and c::bv and b::b assumes "atom c \ (b1,bv1)" shows "b1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ b1)[c ::= b]\<^sub>b\<^sub>b" using assms proof(nominal_induct b1 rule: b.strong_induct) case B_int then show ?case using subst_bb.simps b.perm_simps by auto next case B_bool then show ?case using subst_bb.simps b.perm_simps by auto next case (B_id x) hence "atom bv1 \ x \ atom c \ x" using fresh_def pure_supp by auto hence "((bv1 \ c) \ B_id x) = B_id x" using fresh_Pair b.fresh(3) flip_fresh_fresh b.perm_simps fresh_def pure_supp by metis then show ?case using subst_bb.simps by simp next case (B_pair x1 x2) hence "x1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ x1)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair by metis moreover have "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair B_pair by metis ultimately show ?case using subst_bb.simps(5) b.perm_simps(4) b.fresh(4) fresh_Pair by auto next case B_unit then show ?case using subst_bb.simps b.perm_simps by auto next case B_bitvec then show ?case using subst_bb.simps b.perm_simps by auto next case (B_var x) then show ?case proof(cases "x=bv1") case True then show ?thesis using B_var subst_bb.simps b.perm_simps by simp next case False moreover have "x\c" using B_var b.fresh fresh_def supp_at_base fresh_Pair by fastforce ultimately show ?thesis using B_var subst_bb.simps(1) b.perm_simps(7) by simp qed next case (B_app x1 x2) hence "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps b.fresh fresh_Pair by metis thus ?case using subst_bb.simps b.perm_simps b.fresh fresh_Pair B_app by (simp add: permute_pure) qed lemma subst_bb_flip_sym: fixes b1::b and b2::b assumes "atom c \ b" and "atom c \ (bv1,bv2, b1, b2)" and "(bv1 \ c) \ b1 = (bv2 \ c) \ b2" shows "b1[bv1::=b]\<^sub>b\<^sub>b = b2[bv2::=b]\<^sub>b\<^sub>b" using assms flip_b_subst4[of c b1 bv1 b] flip_b_subst4[of c b2 bv2 b] fresh_prod4 fresh_Pair by simp section \Value\ nominal_function subst_vb :: "v \ bv \ b \ v" where "subst_vb (V_lit l) x v = V_lit l" | "subst_vb (V_var y) x v = V_var y" | "subst_vb (V_cons tyid c v') x v = V_cons tyid c (subst_vb v' x v)" | "subst_vb (V_consp tyid c b v') x v = V_consp tyid c (subst_bb b x v) (subst_vb v' x v)" | "subst_vb (V_pair v1 v2) x v = V_pair (subst_vb v1 x v ) (subst_vb v2 x v )" apply (simp add: eqvt_def subst_vb_graph_aux_def) apply auto using v.strong_exhaust by meson nominal_termination (eqvt) by lexicographic_order abbreviation subst_vb_abbrev :: "v \ bv \ b \ v" ("_[_::=_]\<^sub>v\<^sub>b" [1000,50,50] 500) where "e[bv::=b]\<^sub>v\<^sub>b \ subst_vb e bv b" instantiation v :: has_subst_b begin definition "subst_b = subst_vb" instance proof fix j::atom and i::bv and x::b and t::v show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" proof (induct t rule: v.induct) case (V_lit l) have "j \ subst_b (V_lit l) i x = j \ (V_lit l)" using subst_vb.simps fresh_def pure_fresh subst_b_v_def v.supp v.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by auto also have "... = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis moreover have "(atom i \ (V_lit l) \ j \ (V_lit l) \ j \ x \ (j \ (V_lit l) \ j = atom i)) = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis ultimately show ?case by simp next case (V_var y) then show ?case using subst_b_v_def subst_vb.simps pure_fresh by force next case (V_pair x1a x2a) then show ?case using subst_b_v_def subst_vb.simps by auto next case (V_cons x1a x2a x3) then show ?case using V_cons subst_b_v_def subst_vb.simps pure_fresh by force next case (V_consp x1a x2a x3 x4) then show ?case using subst_b_v_def subst_vb.simps pure_fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by fastforce qed fix a::bv and tm::v and x::b show "atom a \ tm \ subst_b tm a x = tm" apply(induct tm rule: v.induct) apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def) using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh using has_subst_b_class.forget_subst by fastforce fix a::bv and tm::v show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def apply (induct tm rule: v.induct) apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def) using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh using has_subst_b_class.subst_id by metis fix p::perm and x1::bv and v::b and t1::v show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" apply(induct tm rule: v.induct) apply(auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) using has_subst_b_class.eqvt subst_b_b_def e.fresh using has_subst_b_class.eqvt by (simp add: subst_b_v_def)+ fix bv::bv and c::v and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+) apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2) using fresh_at_base flip_fresh_fresh[of bv x z] apply (simp add: flip_fresh_fresh) using subst_b_b_def by argo fix bv::bv and c::v and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+) apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2) using fresh_at_base flip_fresh_fresh[of bv x z] apply (simp add: flip_fresh_fresh) using subst_b_b_def flip_subst_subst by fastforce qed end section \Constraints Expressions\ nominal_function subst_ceb :: "ce \ bv \ b \ ce" where "subst_ceb ( (CE_val v') ) bv b = ( CE_val (subst_vb v' bv b) )" | "subst_ceb ( (CE_op opp v1 v2) ) bv b = ( (CE_op opp (subst_ceb v1 bv b)(subst_ceb v2 bv b)) )" | "subst_ceb ( (CE_fst v')) bv b = CE_fst (subst_ceb v' bv b)" | "subst_ceb ( (CE_snd v')) bv b = CE_snd (subst_ceb v' bv b)" | "subst_ceb ( (CE_len v')) bv b = CE_len (subst_ceb v' bv b)" | "subst_ceb ( CE_concat v1 v2) bv b = CE_concat (subst_ceb v1 bv b) (subst_ceb v2 bv b)" apply (simp add: eqvt_def subst_ceb_graph_aux_def) apply auto by (meson ce.strong_exhaust) nominal_termination (eqvt) by lexicographic_order abbreviation subst_ceb_abbrev :: "ce \ bv \ b \ ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>b" [1000,50,50] 500) where "ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b \ subst_ceb ce bv b" instantiation ce :: has_subst_b begin definition "subst_b = subst_ceb" instance proof fix j::atom and i::bv and x::b and t::ce show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" proof (induct t rule: ce.induct) case (CE_val v) then show ?case using subst_ceb.simps fresh_def pure_fresh subst_b_ce_def ce.supp v.supp ce.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by metis next case (CE_op opp v1 v2) have "(j \ v1[i::=x]\<^sub>c\<^sub>e\<^sub>b \ j \ v2[i::=x]\<^sub>c\<^sub>e\<^sub>b) = ((atom i \ v1 \ atom i \ v2) \ j \ v1 \ j \ v2 \ j \ x \ (j \ v1 \ j \ v2 \ j = atom i))" using has_subst_b_class.fresh_subst_if subst_b_v_def using CE_op.hyps(1) CE_op.hyps(2) subst_b_ce_def by auto thus ?case unfolding subst_ceb.simps subst_b_ce_def ce.fresh using fresh_def pure_fresh opp.fresh subst_b_v_def opp.exhaust fresh_e_opp_all by (metis (full_types)) next case (CE_concat x1a x2) then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by force next case (CE_fst x) then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis next case (CE_snd x) then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis next case (CE_len x) then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis qed fix a::bv and tm::ce and x::b show "atom a \ tm \ subst_b tm a x = tm" apply(induct tm rule: ce.induct) apply( auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def) using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh using has_subst_b_class.forget_subst subst_b_v_def apply metis+ done fix a::bv and tm::ce show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def apply (induct tm rule: ce.induct) apply(auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def) using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh using has_subst_b_class.subst_id subst_b_v_def apply metis+ done fix p::perm and x1::bv and v::b and t1::ce show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" apply(induct tm rule: ce.induct) apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) using has_subst_b_class.eqvt subst_b_b_def ce.fresh using has_subst_b_class.eqvt by (simp add: subst_b_ce_def)+ fix bv::bv and c::ce and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" apply (induct c rule: ce.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+) using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def apply metis using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def by (simp add: flip_fresh_fresh fresh_opp_all) fix bv::bv and c::ce and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" proof (induct c rule: ce.induct) case (CE_val x) then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce next case (CE_op x1a x2 x3) then show ?case unfolding subst_ceb.simps subst_b_ce_def ce.perm_simps using flip_subst_subst subst_b_v_def opp.perm_simps opp.strong_exhaust by (metis (full_types) ce.fresh(2)) next case (CE_concat x1a x2) then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce next case (CE_fst x) then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce next case (CE_snd x) then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce next case (CE_len x) then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce qed qed end section \Constraints\ nominal_function subst_cb :: "c \ bv \ b \ c" where "subst_cb (C_true) x v = C_true" | "subst_cb (C_false) x v = C_false" | "subst_cb (C_conj c1 c2) x v = C_conj (subst_cb c1 x v ) (subst_cb c2 x v )" | "subst_cb (C_disj c1 c2) x v = C_disj (subst_cb c1 x v ) (subst_cb c2 x v )" | "subst_cb (C_imp c1 c2) x v = C_imp (subst_cb c1 x v ) (subst_cb c2 x v )" | "subst_cb (C_eq e1 e2) x v = C_eq (subst_ceb e1 x v ) (subst_ceb e2 x v )" | "subst_cb (C_not c) x v = C_not (subst_cb c x v )" apply (simp add: eqvt_def subst_cb_graph_aux_def) apply auto using c.strong_exhaust apply metis done nominal_termination (eqvt) by lexicographic_order abbreviation subst_cb_abbrev :: "c \ bv \ b \ c" ("_[_::=_]\<^sub>c\<^sub>b" [1000,50,50] 500) where "c[bv::=b]\<^sub>c\<^sub>b \ subst_cb c bv b" instantiation c :: has_subst_b begin definition "subst_b = subst_cb" instance proof fix j::atom and i::bv and x::b and t::c show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" by (induct t rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh, (metis has_subst_b_class.fresh_subst_if subst_b_ce_def c.fresh)+ ) fix a::bv and tm::c and x::b show "atom a \ tm \ subst_b tm a x = tm" by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh, (metis has_subst_b_class.forget_subst subst_b_ce_def)+) fix a::bv and tm::c show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_c_def by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh, (metis has_subst_b_class.subst_id subst_b_ce_def)+) fix p::perm and x1::bv and v::b and t1::c show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" apply(induct tm rule: c.induct,unfold subst_cb.simps subst_b_c_def c.fresh) by( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) fix bv::bv and c::c and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+) using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def apply metis using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def apply (metis opp.perm_simps(2) opp.strong_exhaust)+ done fix bv::bv and c::c and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+) using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def using flip_subst_subst apply fastforce using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def opp.perm_simps(2) opp.strong_exhaust proof - fix x1a :: ce and x2 :: ce assume a1: "atom bv \ x2" then have "((bv \ z) \ x2)[bv::=v]\<^sub>b = x2[z::=v]\<^sub>b" by (metis flip_subst_subst) (* 0.0 ms *) then show "x2[z::=B_var bv]\<^sub>b[bv::=v]\<^sub>c\<^sub>e\<^sub>b = x2[z::=v]\<^sub>c\<^sub>e\<^sub>b" using a1 by (simp add: subst_b_ce_def) (* 0.0 ms *) qed qed end section \Types\ nominal_function subst_tb :: "\ \ bv \ b \ \" where "subst_tb (\ z : b2 | c \) bv1 b1 = \ z : b2[bv1::=b1]\<^sub>b\<^sub>b | c[bv1::=b1]\<^sub>c\<^sub>b \" proof(goal_cases) case 1 then show ?case using eqvt_def subst_tb_graph_aux_def by force next case (2 x y) then show ?case by auto next case (3 P x) then show ?case using eqvt_def subst_tb_graph_aux_def \.strong_exhaust by (metis b_of.cases prod_cases3) next case (4 z' b2' c' bv1' b1' z b2 c bv1 b1) show ?case unfolding \.eq_iff proof have *:"[[atom z']]lst. c' = [[atom z]]lst. c" using \.eq_iff 4 by auto show "[[atom z']]lst. c'[bv1'::=b1']\<^sub>c\<^sub>b = [[atom z]]lst. c[bv1::=b1]\<^sub>c\<^sub>b" proof(subst Abs1_eq_iff_all(3),rule,rule,rule) fix ca::x assume "atom ca \ z" and 1:"atom ca \ (z', z, c'[bv1'::=b1']\<^sub>c\<^sub>b, c[bv1::=b1]\<^sub>c\<^sub>b)" hence 2:"atom ca \ (c',c)" using fresh_subst_if subst_b_c_def fresh_Pair fresh_prod4 fresh_at_base subst_b_fresh_x by metis hence "(z' \ ca) \ c' = (z \ ca) \ c" using 1 2 * Abs1_eq_iff_all(3) by auto hence "((z' \ ca) \ c')[bv1'::=b1']\<^sub>c\<^sub>b = ((z \ ca) \ c)[bv1'::=b1']\<^sub>c\<^sub>b" by auto hence "(z' \ ca) \ c'[(z' \ ca) \ bv1'::=(z' \ ca) \ b1']\<^sub>c\<^sub>b = (z \ ca) \ c[(z \ ca) \ bv1'::=(z \ ca) \ b1']\<^sub>c\<^sub>b" by auto thus "(z' \ ca) \ c'[bv1'::=b1']\<^sub>c\<^sub>b = (z \ ca) \ c[bv1::=b1]\<^sub>c\<^sub>b" using 4 flip_x_b_cancel by simp qed show "b2'[bv1'::=b1']\<^sub>b\<^sub>b = b2[bv1::=b1]\<^sub>b\<^sub>b" using 4 by simp qed qed nominal_termination (eqvt) by lexicographic_order abbreviation subst_tb_abbrev :: "\ \ bv \ b \ \" ("_[_::=_]\<^sub>\\<^sub>b" [1000,50,50] 1000) where "t[bv::=b']\<^sub>\\<^sub>b \ subst_tb t bv b' " instantiation \ :: has_subst_b begin definition "subst_b = subst_tb" instance proof fix j::atom and i::bv and x::b and t::\ show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" proof (nominal_induct t avoiding: i x j rule: \.strong_induct) case (T_refined_type z b c) then show ?case unfolding subst_b_\_def subst_tb.simps \.fresh using fresh_subst_if[of j b i x ] subst_b_b_def subst_b_c_def by (metis has_subst_b_class.fresh_subst_if list.distinct(1) list.set_cases not_self_fresh set_ConsD) qed fix a::bv and tm::\ and x::b show "atom a \ tm \ subst_b tm a x = tm" proof (nominal_induct tm avoiding: a x rule: \.strong_induct) case (T_refined_type xx bb cc ) moreover hence "atom a \ bb \ atom a \ cc" using \.fresh by auto ultimately show ?case unfolding subst_b_\_def subst_tb.simps using forget_subst subst_b_b_def subst_b_c_def forget_subst \.fresh by metis qed fix a::bv and tm::\ show "subst_b tm a (B_var a) = tm" proof (nominal_induct tm rule: \.strong_induct) case (T_refined_type xx bb cc) thus ?case unfolding subst_b_\_def subst_tb.simps using subst_id subst_b_b_def subst_b_c_def by metis qed fix p::perm and x1::bv and v::b and t1::\ show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v) " by (induct tm rule: \.induct, auto simp add: fresh_at_base subst_tb.simps subst_b_\_def subst_bb.simps subst_b_b_def) fix bv::bv and c::\ and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" apply (induct c rule: \.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+) using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_c_def subst_b_b_def by (simp add: flip_fresh_fresh subst_b_\_def) fix bv::bv and c::\ and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" proof (induct c rule: \.induct) case (T_refined_type x1a x2a x3a) hence "atom bv \ x2a \ atom bv \ x3a \ atom bv \ x1a" using fresh_at_base \.fresh by simp then show ?case unfolding subst_tb.simps subst_b_\_def \.perm_simps using fresh_at_base flip_fresh_fresh[of bv x1a z] flip_subst_subst subst_b_b_def subst_b_c_def T_refined_type proof - have "atom z \ x1a" by (metis b.fresh(7) fresh_at_base(2) x_fresh_b) (* 0.0 ms *) then show "\ (bv \ z) \ x1a : ((bv \ z) \ x2a)[bv::=v]\<^sub>b\<^sub>b | ((bv \ z) \ x3a)[bv::=v]\<^sub>c\<^sub>b \ = \ x1a : x2a[z::=v]\<^sub>b\<^sub>b | x3a[z::=v]\<^sub>c\<^sub>b \" by (metis \\atom bv \ x1a; atom z \ x1a\ \ (bv \ z) \ x1a = x1a\ \atom bv \ x2a \ atom bv \ x3a \ atom bv \ x1a\ flip_subst_subst subst_b_b_def subst_b_c_def) (* 78 ms *) qed qed qed end lemma subst_bb_commute [simp]: "atom j \ A \ (subst_bb (subst_bb A i t ) j u ) = subst_bb A i (subst_bb t j u) " by (nominal_induct A avoiding: i j t u rule: b.strong_induct) (auto simp: fresh_at_base) lemma subst_vb_commute [simp]: "atom j \ A \ (subst_vb (subst_vb A i t )) j u = subst_vb A i (subst_bb t j u ) " by (nominal_induct A avoiding: i j t u rule: v.strong_induct) (auto simp: fresh_at_base) lemma subst_ceb_commute [simp]: "atom j \ A \ (subst_ceb (subst_ceb A i t )) j u = subst_ceb A i (subst_bb t j u ) " by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base) lemma subst_cb_commute [simp]: "atom j \ A \ (subst_cb (subst_cb A i t )) j u = subst_cb A i (subst_bb t j u ) " by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base) lemma subst_tb_commute [simp]: "atom j \ A \ (subst_tb (subst_tb A i t )) j u = subst_tb A i (subst_bb t j u ) " proof (nominal_induct A avoiding: i j t u rule: \.strong_induct) case (T_refined_type z b c) then show ?case using subst_tb.simps subst_bb_commute subst_cb_commute by simp qed section \Expressions\ nominal_function subst_eb :: "e \ bv \ b \ e" where "subst_eb ( (AE_val v')) bv b = ( AE_val (subst_vb v' bv b))" | "subst_eb ( (AE_app f v') ) bv b = ( (AE_app f (subst_vb v' bv b)) )" | "subst_eb ( (AE_appP f b' v') ) bv b = ( (AE_appP f (b'[bv::=b]\<^sub>b\<^sub>b) (subst_vb v' bv b)))" | "subst_eb ( (AE_op opp v1 v2) ) bv b = ( (AE_op opp (subst_vb v1 bv b) (subst_vb v2 bv b)) )" | "subst_eb ( (AE_fst v')) bv b = AE_fst (subst_vb v' bv b)" | "subst_eb ( (AE_snd v')) bv b = AE_snd (subst_vb v' bv b)" | "subst_eb ( (AE_mvar u)) bv b = AE_mvar u" | "subst_eb ( (AE_len v')) bv b = AE_len (subst_vb v' bv b)" | "subst_eb ( AE_concat v1 v2) bv b = AE_concat (subst_vb v1 bv b) (subst_vb v2 bv b)" | "subst_eb ( AE_split v1 v2) bv b = AE_split (subst_vb v1 bv b) (subst_vb v2 bv b)" apply (simp add: eqvt_def subst_eb_graph_aux_def) apply auto by (meson e.strong_exhaust) nominal_termination (eqvt) by lexicographic_order abbreviation subst_eb_abbrev :: "e \ bv \ b \ e" ("_[_::=_]\<^sub>e\<^sub>b" [1000,50,50] 500) where "e[bv::=b]\<^sub>e\<^sub>b \ subst_eb e bv b" instantiation e :: has_subst_b begin definition "subst_b = subst_eb" instance proof fix j::atom and i::bv and x::b and t::e show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" proof (induct t rule: e.induct) case (AE_val v) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp e.fresh has_subst_b_class.fresh_subst_if subst_b_e_def subst_b_v_def by metis next case (AE_app f v) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def - by (metis (mono_tags, hide_lams) e.fresh(2)) + by (metis (mono_tags, opaque_lifting) e.fresh(2)) next case (AE_appP f b' v) then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using fresh_def pure_fresh subst_b_e_def e.supp v.supp e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def) next case (AE_op opp v1 v2) then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using fresh_def pure_fresh subst_b_e_def e.supp v.supp fresh_e_opp_all e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def) next case (AE_concat x1a x2) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def by (metis subst_vb.simps(5)) next case (AE_split x1a x2) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def by (metis subst_vb.simps(5)) next case (AE_fst x) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def by metis next case (AE_snd x) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis next case (AE_mvar x) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp by auto next case (AE_len x) then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis qed fix a::bv and tm::e and x::b show "atom a \ tm \ subst_b tm a x = tm" apply(induct tm rule: e.induct) apply( auto simp add: fresh_at_base subst_eb.simps subst_b_e_def) using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh using has_subst_b_class.forget_subst subst_b_v_def apply metis+ done fix a::bv and tm::e show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def apply (induct tm rule: e.induct) apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def) using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh using has_subst_b_class.subst_id subst_b_v_def apply metis+ done fix p::perm and x1::bv and v::b and t1::e show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" apply(induct tm rule: e.induct) apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) using has_subst_b_class.eqvt subst_b_b_def e.fresh using has_subst_b_class.eqvt by (simp add: subst_b_e_def)+ fix bv::bv and c::e and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" apply (induct c rule: e.induct) apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp ) using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def flip_fresh_fresh subst_b_\_def apply metis apply (metis (full_types) opp.perm_simps opp.strong_exhaust) done fix bv::bv and c::e and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" apply (induct c rule: e.induct) apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp ) using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def flip_fresh_fresh subst_b_\_def apply simp apply (metis (full_types) opp.perm_simps opp.strong_exhaust) done qed end section \Statements\ nominal_function (default "case_sum (\x. Inl undefined) (case_sum (\x. Inl undefined) (\x. Inr undefined))") subst_sb :: "s \ bv \ b \ s" and subst_branchb :: "branch_s \ bv \ b \ branch_s" and subst_branchlb :: "branch_list \ bv \ b \ branch_list" where "subst_sb (AS_val v') bv b = (AS_val (subst_vb v' bv b))" | "subst_sb (AS_let y e s) bv b = (AS_let y (e[bv::=b]\<^sub>e\<^sub>b) (subst_sb s bv b ))" | "subst_sb (AS_let2 y t s1 s2) bv b = (AS_let2 y (subst_tb t bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b))" | "subst_sb (AS_match v' cs) bv b = AS_match (subst_vb v' bv b) (subst_branchlb cs bv b)" | "subst_sb (AS_assign y v') bv b = AS_assign y (subst_vb v' bv b)" | "subst_sb (AS_if v' s1 s2) bv b = (AS_if (subst_vb v' bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b ) )" | "subst_sb (AS_var u \ v' s) bv b = AS_var u (subst_tb \ bv b) (subst_vb v' bv b) (subst_sb s bv b )" | "subst_sb (AS_while s1 s2) bv b = AS_while (subst_sb s1 bv b ) (subst_sb s2 bv b )" | "subst_sb (AS_seq s1 s2) bv b = AS_seq (subst_sb s1 bv b ) (subst_sb s2 bv b )" | "subst_sb (AS_assert c s) bv b = AS_assert (subst_cb c bv b ) (subst_sb s bv b )" | "subst_branchb (AS_branch dc x1 s') bv b = AS_branch dc x1 (subst_sb s' bv b)" | "subst_branchlb (AS_final sb) bv b = AS_final (subst_branchb sb bv b )" | "subst_branchlb (AS_cons sb ssb) bv b = AS_cons (subst_branchb sb bv b) (subst_branchlb ssb bv b)" apply (simp add: eqvt_def subst_sb_subst_branchb_subst_branchlb_graph_aux_def ) apply (auto,metis s_branch_s_branch_list.exhaust s_branch_s_branch_list.exhaust(2) old.sum.exhaust surj_pair) proof(goal_cases) have eqvt_at_proj: "\ s xa va . eqvt_at subst_sb_subst_branchb_subst_branchlb_sumC (Inl (s, xa, va)) \ eqvt_at (\a. projl (subst_sb_subst_branchb_subst_branchlb_sumC (Inl a))) (s, xa, va)" apply(simp only: eqvt_at_def) apply(rule) apply(subst Projl_permute) apply(thin_tac _)+ apply(simp add: subst_sb_subst_branchb_subst_branchlb_sumC_def) apply(simp add: THE_default_def) apply(case_tac "Ex1 (subst_sb_subst_branchb_subst_branchlb_graph (Inl (s,xa,va)))") apply simp apply(auto)[1] apply(erule_tac x="x" in allE) apply simp apply(cases rule: subst_sb_subst_branchb_subst_branchlb_graph.cases) apply(assumption) apply(rule_tac x="Sum_Type.projl x" in exI,clarify,rule the1_equality,blast,simp (no_asm) only: sum.sel)+ apply(blast)+ apply(simp)+ done { case (1 y s bv b ya sa c) moreover have "atom y \ (bv, b) \ atom ya \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp ultimately show ?case using eqvt_triple eqvt_at_proj by metis next case (2 y s1 s2 bv b ya s2a c) moreover have "atom y \ (bv, b) \ atom ya \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp ultimately show ?case using eqvt_triple eqvt_at_proj by metis next case (3 u s bv b ua sa c) moreover have "atom u \ (bv, b) \ atom ua \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp ultimately show ?case using eqvt_triple eqvt_at_proj by metis next case (4 x1 s' bv b x1a s'a c) moreover have "atom x1 \ (bv, b) \ atom x1a \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp ultimately show ?case using eqvt_triple eqvt_at_proj by metis } qed nominal_termination (eqvt) by lexicographic_order abbreviation subst_sb_abbrev :: "s \ bv \ b \ s" ("_[_::=_]\<^sub>s\<^sub>b" [1000,50,50] 1000) where "b[bv::=b']\<^sub>s\<^sub>b \ subst_sb b bv b'" lemma fresh_subst_sb_if [simp]: "(j \ (subst_sb A i x )) = ((atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i)))" and "(j \ (subst_branchb B i x )) = ((atom i \ B \ j \ B) \ (j \ x \ (j \ B \ j = atom i)))" and "(j \ (subst_branchlb C i x )) = ((atom i \ C \ j \ C) \ (j \ x \ (j \ C \ j = atom i)))" proof (nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct) case (AS_branch x1 x2 x3) have " (j \ subst_branchb (AS_branch x1 x2 x3) i x ) = (j \ (AS_branch x1 x2 (subst_sb x3 i x)))" by auto also have "... = ((j \ x3[i::=x]\<^sub>s\<^sub>b \ j \ set [atom x2]) \ j \ x1)" using s_branch_s_branch_list.fresh by auto also have "... = ((atom i \ AS_branch x1 x2 x3 \ j \ AS_branch x1 x2 x3) \ j \ x \ (j \ AS_branch x1 x2 x3 \ j = atom i))" using subst_branchb.simps(1) s_branch_s_branch_list.fresh(1) fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\_def v.fresh AS_branch proof - have f1: "\cs b. atom (b::bv) \ (cs::char list)" using pure_fresh by auto then have "j \ x \ atom i = j \ ((j \ x3[i::=x]\<^sub>s\<^sub>b \ j \ set [atom x2]) \ j \ x1) = (atom i \ AS_branch x1 x2 x3 \ j \ AS_branch x1 x2 x3 \ j \ x \ (j \ AS_branch x1 x2 x3 \ j = atom i))" by (metis (full_types) AS_branch.hyps(3)) then have "j \ x \ ((j \ x3[i::=x]\<^sub>s\<^sub>b \ j \ set [atom x2]) \ j \ x1) = (atom i \ AS_branch x1 x2 x3 \ j \ AS_branch x1 x2 x3 \ j \ x \ (j \ AS_branch x1 x2 x3 \ j = atom i))" using AS_branch.hyps s_branch_s_branch_list.fresh by metis moreover { assume "\ j \ x" have ?thesis using f1 AS_branch.hyps(2) AS_branch.hyps(3) by force } ultimately show ?thesis by satx qed finally show ?case by auto next case (AS_cons cs css i x) show ?case unfolding subst_branchlb.simps s_branch_s_branch_list.fresh using AS_cons by auto next case (AS_val xx) then show ?case using subst_sb.simps(1) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by metis next case (AS_let x1 x2 x3) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def by fastforce next case (AS_let2 x1 x2 x3 x4) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\_def by fastforce next case (AS_if x1 x2 x3) then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using has_subst_b_class.fresh_subst_if subst_b_v_def by metis next case (AS_var u t v s) have "(((atom i \ s \ j \ s \ j \ x \ (j \ s \ j = atom i)) \ j \ set [atom u]) \ j \ t[i::=x]\<^sub>\\<^sub>b \ j \ v[i::=x]\<^sub>v\<^sub>b) = (((atom i \ s \ j \ s \ j \ x \ (j \ s \ j = atom i)) \ j \ set [atom u]) \ ((atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))) \ ((atom i \ v \ j \ v \ j \ x \ (j \ v \ j = atom i))))" using has_subst_b_class.fresh_subst_if subst_b_v_def subst_b_\_def by metis also have "... = (((atom i \ s \ atom i \ set [atom u]) \ atom i \ t \ atom i \ v) \ (j \ s \ j \ set [atom u]) \ j \ t \ j \ v \ j \ x \ ((j \ s \ j \ set [atom u]) \ j \ t \ j \ v \ j = atom i))" using u_fresh_b by auto finally show ?case using subst_sb.simps s_branch_s_branch_list.fresh AS_var by simp next case (AS_assign u v ) then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using has_subst_b_class.fresh_subst_if subst_b_v_def by force next case (AS_match v cs) have " j \ (AS_match v cs)[i::=x]\<^sub>s\<^sub>b = j \ (AS_match (subst_vb v i x) (subst_branchlb cs i x ))" using subst_sb.simps by auto also have "... = (j \ (subst_vb v i x) \ j \ (subst_branchlb cs i x ))" using s_branch_s_branch_list.fresh by simp also have "... = (j \ (subst_vb v i x) \ ((atom i \ cs \ j \ cs) \ j \ x \ (j \ cs \ j = atom i)))" using AS_match[of i x] by auto also have "... = (atom i \ AS_match v cs \ j \ AS_match v cs \ j \ x \ (j \ AS_match v cs \ j = atom i))" by (metis (no_types) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_v_def) (* 93 ms *) finally show ?case by auto next case (AS_while x1 x2) then show ?case by auto next case (AS_seq x1 x2) then show ?case by auto next case (AS_assert x1 x2) then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def by (metis subst_b_c_def) qed(auto+) lemma forget_subst_sb[simp]: "atom a \ A \ subst_sb A a x = A" and forget_subst_branchb [simp]: "atom a \ B \ subst_branchb B a x = B" and forget_subst_branchlb[simp]: "atom a \ C \ subst_branchlb C a x = C" proof (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct) case (AS_let x1 x2 x3) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_let2 x1 x2 x3 x4) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_\_def by force next case (AS_var x1 x2 x3 x4) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def using subst_b_\_def proof - have f1: "(atom a \ x4 \ atom a \ set [atom x1]) \ atom a \ x2 \ atom a \ x3" using AS_var.prems s_branch_s_branch_list.fresh by simp then have "atom a \ x4" by (metis (no_types) "Nominal-Utils.fresh_star_singleton" AS_var.hyps(1) empty_set fresh_star_def list.simps(15) not_self_fresh) (* 46 ms *) then show ?thesis using f1 by (metis AS_var.hyps(3) has_subst_b_class.forget_subst subst_b_\_def subst_b_v_def subst_sb.simps(7)) (* 62 ms *) qed next case (AS_branch x1 x2 x3) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_cons x1 x2 x3 x4) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_val x) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_if x1 x2 x3) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_assign x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_match x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_while x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_seq x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force next case (AS_assert c s) then show ?case unfolding subst_sb.simps using s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def subst_b_c_def subst_cb.simps by force qed(auto+) lemma subst_sb_id: "subst_sb A a (B_var a) = A" and subst_branchb_id [simp]: "subst_branchb B a (B_var a) = B" and subst_branchlb_id: "subst_branchlb C a (B_var a) = C" proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct) case (AS_branch x1 x2 x3) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by simp next case (AS_cons x1 x2 ) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by simp next case (AS_val x) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis next case (AS_if x1 x2 x3) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis next case (AS_assign x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis next case (AS_match x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis next case (AS_while x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis next case (AS_seq x1 x2) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis next case (AS_let x1 x2 x3) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.subst_id by metis next case (AS_let2 x1 x2 x3 x4) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id by metis next case (AS_var x1 x2 x3 x4) then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis next case (AS_assert c s ) then show ?case unfolding subst_sb.simps using s_branch_s_branch_list.fresh subst_b_c_def has_subst_b_class.subst_id by metis qed (auto) lemma flip_subst_s: fixes bv::bv and s::s and cs::branch_s and z::bv shows "atom bv \ s \ ((bv \ z) \ s) = s[z::=B_var bv]\<^sub>s\<^sub>b" and "atom bv \ cs \ ((bv \ z) \ cs) = subst_branchb cs z (B_var bv) " and "atom bv \ css \ ((bv \ z) \ css) = subst_branchlb css z (B_var bv) " proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct) case (AS_branch x1 x2 x3) hence "((bv \ z) \ x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis moreover have "((bv \ z) \ x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto next case (AS_cons x1 x2 ) hence "((bv \ z) \ x1) = subst_branchb x1 z (B_var bv)" using pure_fresh fresh_at_base flip_fresh_fresh s_branch_s_branch_list.fresh(13) by metis moreover have "((bv \ z) \ x2) = subst_branchlb x2 z (B_var bv)" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_cons s_branch_s_branch_list.fresh by metis ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_cons by auto next case (AS_val x) then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp next case (AS_let x1 x2 x3) moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def s_branch_s_branch_list.fresh by auto next case (AS_let2 x1 x2 x3 x4) moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\_def by auto next case (AS_if x1 x2 x3) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_var x1 x2 x3 x4) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def subst_b_\_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto next case (AS_assign x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto next case (AS_match x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_while x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_seq x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_assert x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_c_def subst_b_v_def s_branch_s_branch_list.fresh by simp qed(auto) lemma flip_subst_subst_s: fixes bv::bv and s::s and cs::branch_s and z::bv shows "atom bv \ s \ ((bv \ z) \ s)[bv::=v]\<^sub>s\<^sub>b = s[z::=v]\<^sub>s\<^sub>b" and "atom bv \ cs \ subst_branchb ((bv \ z) \ cs) bv v = subst_branchb cs z v" and "atom bv \ css \ subst_branchlb ((bv \ z) \ css) bv v = subst_branchlb css z v " proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct) case (AS_branch x1 x2 x3) hence "((bv \ z) \ x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis moreover have "((bv \ z) \ x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto next case (AS_cons x1 x2 ) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_branchlb.simps using s_branch_s_branch_list.fresh(1) AS_cons by auto next case (AS_val x) then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp next case (AS_let x1 x2 x3) moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst_subst subst_b_e_def s_branch_s_branch_list.fresh by force next case (AS_let2 x1 x2 x3 x4) moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\_def by auto next case (AS_if x1 x2 x3) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_var x1 x2 x3 x4) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def subst_b_\_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto next case (AS_assign x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto next case (AS_match x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_while x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_seq x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto next case (AS_assert x1 x2) thus ?case unfolding s_branch_s_branch_list.perm_simps subst_sb.simps using flip_subst subst_b_e_def subst_b_c_def s_branch_s_branch_list.fresh by auto qed(auto) instantiation s :: has_subst_b begin definition "subst_b = (\s bv b. subst_sb s bv b)" instance proof fix j::atom and i::bv and x::b and t::s show "j \ subst_b t i x = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" using fresh_subst_sb_if subst_b_s_def by metis fix a::bv and tm::s and x::b show "atom a \ tm \ subst_b tm a x = tm" using subst_b_s_def forget_subst_sb by metis fix a::bv and tm::s show "subst_b tm a (B_var a) = tm" using subst_b_s_def subst_sb_id by metis fix p::perm and x1::bv and v::b and t1::s show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" using subst_b_s_def subst_sb_subst_branchb_subst_branchlb.eqvt by metis fix bv::bv and c::s and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" using subst_b_s_def flip_subst_s by metis fix bv::bv and c::s and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" using flip_subst_subst_s subst_b_s_def by metis qed end section \Function Type\ nominal_function subst_ft_b :: "fun_typ \ bv \ b \ fun_typ" where "subst_ft_b ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z (subst_bb b x v) (subst_cb c x v) t[x::=v]\<^sub>\\<^sub>b s[x::=v]\<^sub>s\<^sub>b" apply(simp add: eqvt_def subst_ft_b_graph_aux_def ) apply(simp add:fun_typ.strong_exhaust,auto ) apply(rule_tac y=a and c="(a,b)" in fun_typ.strong_exhaust) apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) done nominal_termination (eqvt) by lexicographic_order nominal_function subst_ftq_b :: "fun_typ_q \ bv \ b \ fun_typ_q" where "atom bv \ (x,v) \ subst_ftq_b (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_b ft x v))" | "subst_ftq_b (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_b ft x v))" apply(simp add: eqvt_def subst_ftq_b_graph_aux_def ) apply(simp add:fun_typ_q.strong_exhaust,auto ) apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust) by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) nominal_termination (eqvt) by lexicographic_order instantiation fun_typ :: has_subst_b begin definition "subst_b = subst_ft_b" text \Note: Using just simp in the second apply unpacks and gives a single goal whereas auto gives 43 non-intuitive goals. These goals are easier to solve and tedious, however they that it clear if a mistake is made in the definition of the function. For example, I saw that one of the goals was going through with metis and the next wasn't. It turned out the definition of the function itself was wrong\ instance proof fix j::atom and i::bv and x::b and t::fun_typ show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" apply(nominal_induct t avoiding: i x rule:fun_typ.strong_induct) apply(auto simp add: subst_b_fun_typ_def ) by(metis fresh_subst_if subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def)+ fix a::bv and tm::fun_typ and x::b show "atom a \ tm \ subst_b tm a x = tm" apply (nominal_induct tm avoiding: a x rule: fun_typ.strong_induct) apply(simp add: subst_b_fun_typ_def Abs1_eq_iff') using subst_b_b_def subst_b_fun_typ_def subst_b_\_def subst_b_c_def subst_b_s_def forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD subst_ft_b.simps by metis fix a::bv and tm::fun_typ show "subst_b tm a (B_var a) = tm" apply (nominal_induct tm rule: fun_typ.strong_induct) apply(simp add: subst_b_fun_typ_def Abs1_eq_iff',auto) using subst_b_b_def subst_b_fun_typ_def subst_b_\_def subst_b_c_def subst_b_s_def forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD subst_ft_b.simps by (metis has_subst_b_class.subst_id)+ fix p::perm and x1::bv and v::b and t1::fun_typ show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v) " apply (nominal_induct t1 avoiding: x1 v rule: fun_typ.strong_induct) by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps) fix bv::bv and c::fun_typ and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" apply (nominal_induct c avoiding: z bv rule: fun_typ.strong_induct) by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\_def subst_b_s_def) fix bv::bv and c::fun_typ and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" apply (nominal_induct c avoiding: bv v z rule: fun_typ.strong_induct) apply(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\_def subst_b_s_def flip_subst_subst flip_subst) using subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\_def subst_b_s_def flip_subst_subst flip_subst using flip_subst_s(1) flip_subst_subst_s(1) by auto qed end instantiation fun_typ_q :: has_subst_b begin definition "subst_b = subst_ftq_b" instance proof fix j::atom and i::bv and x::b and t::fun_typ_q show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" apply (nominal_induct t avoiding: i x j rule: fun_typ_q.strong_induct,auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps) using fresh_subst_if subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def apply metis+ done fix a::bv and t::fun_typ_q and x::b show "atom a \ t \ subst_b t a x = t" apply (nominal_induct t avoiding: a x rule: fun_typ_q.strong_induct) apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+ fix p::perm and x1::bv and v::b and t::fun_typ_q show "p \ subst_b t x1 v = subst_b (p \ t) (p \ x1) (p \ v) " apply (nominal_induct t avoiding: x1 v rule: fun_typ_q.strong_induct) by(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') fix a::bv and tm::fun_typ_q show "subst_b tm a (B_var a) = tm" apply (nominal_induct tm avoiding: a rule: fun_typ_q.strong_induct) apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') using subst_id subst_b_b_def subst_b_fun_typ_def subst_b_\_def subst_b_c_def subst_b_s_def forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD subst_ft_b.simps by metis+ fix bv::bv and c::fun_typ_q and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" apply (nominal_induct c avoiding: z bv rule: fun_typ_q.strong_induct) apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+ fix bv::bv and c::fun_typ_q and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" apply (nominal_induct c avoiding: z v bv rule: fun_typ_q.strong_induct) apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') using flip_subst flip_subst_subst forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+ qed end section \Contexts\ subsection \Immutable Variables\ nominal_function subst_gb :: "\ \ bv \ b \ \" where "subst_gb GNil _ _ = GNil" | "subst_gb ((y,b',c)#\<^sub>\\) bv b = ((y,b'[bv::=b]\<^sub>b\<^sub>b,c[bv::=b]\<^sub>c\<^sub>b)#\<^sub>\ (subst_gb \ bv b))" apply (simp add: eqvt_def subst_gb_graph_aux_def )+ apply auto apply (insert \.exhaust neq_GNil_conv, force) done nominal_termination (eqvt) by lexicographic_order abbreviation subst_gb_abbrev :: "\ \ bv \ b \ \" ("_[_::=_]\<^sub>\\<^sub>b" [1000,50,50] 1000) where "g[bv::=b']\<^sub>\\<^sub>b \ subst_gb g bv b'" instantiation \ :: has_subst_b begin definition "subst_b = subst_gb" instance proof fix j::atom and i::bv and x::b and t::\ show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" proof(induct t rule: \_induct) case GNil then show ?case using fresh_GNil subst_gb.simps fresh_def pure_fresh subst_b_\_def has_subst_b_class.fresh_subst_if fresh_GNil fresh_GCons by metis next case (GCons x' b' c' \') have *: "atom i \ x' " using fresh_at_base by simp have "j \ subst_b ((x', b', c') #\<^sub>\ \') i x = j \ ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b) #\<^sub>\ (subst_b \' i x))" using subst_gb.simps subst_b_\_def by auto also have "... = (j \ ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b)) \ (j \ (subst_b \' i x)))" using fresh_GCons by auto also have "... = (((j \ x') \ (j \ b'[i::=x]\<^sub>b\<^sub>b) \ (j \ c'[i::=x]\<^sub>c\<^sub>b)) \ (j \ (subst_b \' i x)))" by auto also have "... = (((j \ x') \ ((atom i \ b' \ j \ b' \ j \ x \ (j \ b' \ j = atom i))) \ ((atom i \ c' \ j \ c' \ j \ x \ (j \ c' \ j = atom i))) \ ((atom i \ \' \ j \ \' \ j \ x \ (j \ \' \ j = atom i)))))" using fresh_subst_if[of j b' i x] fresh_subst_if[of j c' i x] GCons subst_b_b_def subst_b_c_def by simp also have "... = ((atom i \ (x', b', c') #\<^sub>\ \' \ j \ (x', b', c') #\<^sub>\ \') \ (j \ x \ (j \ (x', b', c') #\<^sub>\ \' \ j = atom i)))" using * fresh_GCons fresh_prod3 by metis finally show ?case by auto qed fix a::bv and tm::\ and x::b show "atom a \ tm \ subst_b tm a x = tm" proof (induct tm rule: \_induct) case GNil then show ?case using subst_gb.simps subst_b_\_def by auto next case (GCons x' b' c' \') have *:"b'[a::=x]\<^sub>b\<^sub>b = b' \ c'[a::=x]\<^sub>c\<^sub>b = c'" using GCons fresh_GCons[of "atom a"] fresh_prod3[of "atom a"] has_subst_b_class.forget_subst subst_b_b_def subst_b_c_def by metis have "subst_b ((x', b', c') #\<^sub>\ \') a x = ((x', b'[a::=x]\<^sub>b\<^sub>b, c'[a::=x]\<^sub>c\<^sub>b) #\<^sub>\ (subst_b \' a x))" using subst_b_\_def subst_gb.simps by auto also have "... = ((x', b', c') #\<^sub>\ \')" using * GCons fresh_GCons[of "atom a"] by auto finally show ?case using has_subst_b_class.forget_subst fresh_GCons fresh_prod3 GCons subst_b_\_def has_subst_b_class.forget_subst[of a b' x] fresh_prod3[of "atom a"] by argo qed fix a::bv and tm::\ show "subst_b tm a (B_var a) = tm" proof(induct tm rule: \_induct) case GNil then show ?case using subst_gb.simps subst_b_\_def by auto next case (GCons x' b' c' \') then show ?case using has_subst_b_class.subst_id subst_b_\_def subst_b_b_def subst_b_c_def subst_gb.simps by metis qed fix p::perm and x1::bv and v::b and t1::\ show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" proof (induct tm rule: \_induct) case GNil then show ?case using subst_b_\_def subst_gb.simps by simp next case (GCons x' b' c' \') then show ?case using subst_b_\_def subst_gb.simps has_subst_b_class.eqvt by argo qed fix bv::bv and c::\ and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" proof (induct c rule: \_induct) case GNil then show ?case using subst_b_\_def subst_gb.simps by auto next case (GCons x b c \') have *:"(bv \ z) \ (x, b, c) = (x, (bv \ z) \ b, (bv \ z) \ c)" using flip_bv_x_cancel by auto then show ?case unfolding subst_gb.simps subst_b_\_def permute_\.simps * using GCons subst_b_\_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto qed fix bv::bv and c::\ and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" proof (induct c rule: \_induct) case GNil then show ?case using subst_b_\_def subst_gb.simps by auto next case (GCons x b c \') have *:"(bv \ z) \ (x, b, c) = (x, (bv \ z) \ b, (bv \ z) \ c)" using flip_bv_x_cancel by auto then show ?case unfolding subst_gb.simps subst_b_\_def permute_\.simps * using GCons subst_b_\_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto qed qed end lemma subst_b_base_for_lit: "(base_for_lit l)[bv::=b]\<^sub>b\<^sub>b = base_for_lit l" using base_for_lit.simps l.strong_exhaust by (metis subst_bb.simps(2) subst_bb.simps(3) subst_bb.simps(6) subst_bb.simps(7)) lemma subst_b_lookup: assumes "Some (b, c) = lookup \ x" shows " Some (b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) = lookup \[bv::=b']\<^sub>\\<^sub>b x" using assms by(induct \ rule: \_induct, auto) lemma subst_g_b_x_fresh: fixes x::x and b::b and \::\ and bv::bv assumes "atom x \ \" shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" using subst_b_fresh_x subst_b_\_def assms by metis subsection \Mutable Variables\ nominal_function subst_db :: "\ \ bv \ b \ \" where "subst_db []\<^sub>\ _ _ = []\<^sub>\" | "subst_db ((u,t) #\<^sub>\ \) bv b = ((u,t[bv::=b]\<^sub>\\<^sub>b) #\<^sub>\ (subst_db \ bv b))" apply (simp add: eqvt_def subst_db_graph_aux_def,auto ) using list.exhaust delete_aux.elims using neq_DNil_conv by fastforce nominal_termination (eqvt) by lexicographic_order abbreviation subst_db_abbrev :: "\ \ bv \ b \ \" ("_[_::=_]\<^sub>\\<^sub>b" [1000,50,50] 1000) where "\[bv::=b]\<^sub>\\<^sub>b \ subst_db \ bv b" instantiation \ :: has_subst_b begin definition "subst_b = subst_db" instance proof fix j::atom and i::bv and x::b and t::\ show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" proof(induct t rule: \_induct) case DNil then show ?case using fresh_DNil subst_db.simps fresh_def pure_fresh subst_b_\_def has_subst_b_class.fresh_subst_if fresh_DNil fresh_DCons by metis next case (DCons u t \') have "j \ subst_b ((u, t) #\<^sub>\ \') i x = j \ ((u, t[i::=x]\<^sub>\\<^sub>b) #\<^sub>\ (subst_b \' i x))" using subst_db.simps subst_b_\_def by auto also have "... = (j \ ((u, t[i::=x]\<^sub>\\<^sub>b)) \ (j \ (subst_b \' i x)))" using fresh_DCons by auto also have "... = (((j \ u) \ (j \ t[i::=x]\<^sub>\\<^sub>b)) \ (j \ (subst_b \' i x)))" by auto also have "... = ((j \ u) \ ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i))) \ (atom i \ \' \ j \ \' \ j \ x \ (j \ \' \ j = atom i)))" using has_subst_b_class.fresh_subst_if[of j t i x] subst_b_\_def DCons subst_b_\_def by auto also have "... = (atom i \ (u, t) #\<^sub>\ \' \ j \ (u, t) #\<^sub>\ \' \ j \ x \ (j \ (u, t) #\<^sub>\ \' \ j = atom i))" using DCons subst_db.simps(2) has_subst_b_class.fresh_subst_if fresh_DCons subst_b_\_def pure_fresh fresh_at_base by auto finally show ?case by auto qed fix a::bv and tm::\ and x::b show "atom a \ tm \ subst_b tm a x = tm" proof (induct tm rule: \_induct) case DNil then show ?case using subst_db.simps subst_b_\_def by auto next case (DCons u t \') have *:"t[a::=x]\<^sub>\\<^sub>b = t" using DCons fresh_DCons[of "atom a"] fresh_prod2[of "atom a"] has_subst_b_class.forget_subst subst_b_\_def by metis have "subst_b ((u,t) #\<^sub>\ \') a x = ((u,t[a::=x]\<^sub>\\<^sub>b) #\<^sub>\ (subst_b \' a x))" using subst_b_\_def subst_db.simps by auto also have "... = ((u, t) #\<^sub>\ \')" using * DCons fresh_DCons[of "atom a"] by auto finally show ?case using has_subst_b_class.forget_subst fresh_DCons fresh_prod3 DCons subst_b_\_def has_subst_b_class.forget_subst[of a t x] fresh_prod3[of "atom a"] by argo qed fix a::bv and tm::\ show "subst_b tm a (B_var a) = tm" proof(induct tm rule: \_induct) case DNil then show ?case using subst_db.simps subst_b_\_def by auto next case (DCons u t \') then show ?case using has_subst_b_class.subst_id subst_b_\_def subst_b_\_def subst_db.simps by metis qed fix p::perm and x1::bv and v::b and t1::\ show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" proof (induct tm rule: \_induct) case DNil then show ?case using subst_b_\_def subst_db.simps by simp next case (DCons x' b' \') then show ?case by argo qed fix bv::bv and c::\ and z::bv show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" proof (induct c rule: \_induct) case DNil then show ?case using subst_b_\_def subst_db.simps by auto next case (DCons u t') then show ?case unfolding subst_db.simps subst_b_\_def permute_\.simps using DCons subst_b_\_def subst_db.simps flip_subst subst_b_\_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp qed fix bv::bv and c::\ and z::bv and v::b show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" proof (induct c rule: \_induct) case DNil then show ?case using subst_b_\_def subst_db.simps by auto next case (DCons u t') then show ?case unfolding subst_db.simps subst_b_\_def permute_\.simps using DCons subst_b_\_def subst_db.simps flip_subst subst_b_\_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp qed qed end lemma subst_d_b_member: assumes "(u, \) \ setD \" shows "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>b" using assms by (induct \,auto) lemmas ms_fresh_all = e.fresh s_branch_s_branch_list.fresh \.fresh c.fresh ce.fresh v.fresh l.fresh fresh_at_base opp.fresh pure_fresh ms_fresh lemmas fresh_intros[intro] = fresh_GNil x_not_in_b_set x_not_in_u_atoms x_fresh_b u_not_in_x_atoms bv_not_in_x_atoms u_not_in_b_atoms lemmas subst_b_simps = subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps subst_bb.simps subst_eb.simps subst_branchb.simps subst_sb.simps lemma subst_d_b_x_fresh: fixes x::x and b::b and \::\ and bv::bv assumes "atom x \ \" shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" using subst_b_fresh_x subst_b_\_def assms by metis lemma subst_b_fresh_x: fixes x::x shows "atom x \ v \ atom x \ v[bv::=b']\<^sub>v\<^sub>b" and "atom x \ ce \ atom x \ ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and "atom x \ e \ atom x \ e[bv::=b']\<^sub>e\<^sub>b" and "atom x \ c \ atom x \ c[bv::=b']\<^sub>c\<^sub>b" and "atom x \ t \ atom x \ t[bv::=b']\<^sub>\\<^sub>b" and "atom x \ d \ atom x \ d[bv::=b']\<^sub>\\<^sub>b" and "atom x \ g \ atom x \ g[bv::=b']\<^sub>\\<^sub>b" and "atom x \ s \ atom x \ s[bv::=b']\<^sub>s\<^sub>b" using fresh_subst_if x_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\_def subst_b_s_def subst_g_b_x_fresh subst_d_b_x_fresh by metis+ lemma subst_b_fresh_u_cls: fixes tm::"'a::has_subst_b" and x::u shows "atom x \ tm = atom x \ tm[bv::=b']\<^sub>b" using fresh_subst_if[of "atom x" tm bv b'] using u_fresh_b by auto lemma subst_g_b_u_fresh: fixes x::u and b::b and \::\ and bv::bv assumes "atom x \ \" shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" using subst_b_fresh_u_cls subst_b_\_def assms by metis lemma subst_d_b_u_fresh: fixes x::u and b::b and \::\ and bv::bv assumes "atom x \ \" shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" using subst_b_fresh_u_cls subst_b_\_def assms by metis lemma subst_b_fresh_u: fixes x::u shows "atom x \ v \ atom x \ v[bv::=b']\<^sub>v\<^sub>b" and "atom x \ ce \ atom x \ ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and "atom x \ e \ atom x \ e[bv::=b']\<^sub>e\<^sub>b" and "atom x \ c \ atom x \ c[bv::=b']\<^sub>c\<^sub>b" and "atom x \ t \ atom x \ t[bv::=b']\<^sub>\\<^sub>b" and "atom x \ d \ atom x \ d[bv::=b']\<^sub>\\<^sub>b" and "atom x \ g \ atom x \ g[bv::=b']\<^sub>\\<^sub>b" and "atom x \ s \ atom x \ s[bv::=b']\<^sub>s\<^sub>b" using fresh_subst_if u_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\_def subst_b_s_def subst_g_b_u_fresh subst_d_b_u_fresh by metis+ lemma subst_db_u_fresh: fixes u::u and b::b and D::\ assumes "atom u \ D" shows "atom u \ D[bv::=b]\<^sub>\\<^sub>b" using assms proof(induct D rule: \_induct) case DNil then show ?case by auto next case (DCons u' t' D') then show ?case using subst_db.simps fresh_def fresh_DCons fresh_subst_if subst_b_\_def by (metis fresh_Pair u_not_in_b_atoms) qed lemma flip_bt_subst4: fixes t::\ and bv::bv assumes "atom bv \ t" shows "t[bv'::=b]\<^sub>\\<^sub>b = ((bv' \ bv) \ t)[bv::=b]\<^sub>\\<^sub>b" using flip_subst_subst[OF assms,of bv' b] by (simp add: flip_commute subst_b_\_def) lemma subst_bt_flip_sym: fixes t1::\ and t2::\ assumes "atom bv \ b" and "atom bv \ (bv1, bv2, t1, t2)" and "(bv1 \ bv) \ t1 = (bv2 \ bv) \ t2" shows " t1[bv1::=b]\<^sub>\\<^sub>b = t2[bv2::=b]\<^sub>\\<^sub>b" using assms flip_bt_subst4[of bv t1 bv1 b ] flip_bt_subst4 fresh_prod4 fresh_Pair by metis end \ No newline at end of file diff --git a/thys/MiniSail/Nominal-Utils.thy b/thys/MiniSail/Nominal-Utils.thy --- a/thys/MiniSail/Nominal-Utils.thy +++ b/thys/MiniSail/Nominal-Utils.thy @@ -1,586 +1,586 @@ (*<*) theory "Nominal-Utils" imports Nominal2.Nominal2 "HOL-Library.AList" begin (*>*) chapter \Prelude\ text \Some useful Nominal lemmas. Many of these are from Launchbury.Nominal-Utils.\ section \Lemmas helping with equivariance proofs\ lemma perm_rel_lemma: assumes "\ \ x y. r (\ \ x) (\ \ y) \ r x y" shows "r (\ \ x) (\ \ y) \ r x y" (is "?l \ ?r") by (metis (full_types) assms permute_minus_cancel(2)) lemma perm_rel_lemma2: assumes "\ \ x y. r x y \ r (\ \ x) (\ \ y)" shows "r x y \ r (\ \ x) (\ \ y)" (is "?l \ ?r") by (metis (full_types) assms permute_minus_cancel(2)) lemma fun_eqvtI: assumes f_eqvt[eqvt]: "(\ p x. p \ (f x) = f (p \ x))" shows "p \ f = f" by perm_simp rule lemma eqvt_at_apply: assumes "eqvt_at f x" shows "(p \ f) x = f x" -by (metis (hide_lams, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1)) +by (metis (opaque_lifting, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1)) lemma eqvt_at_apply': assumes "eqvt_at f x" shows "p \ f x = f (p \ x)" -by (metis (hide_lams, no_types) assms eqvt_at_def) +by (metis (opaque_lifting, no_types) assms eqvt_at_def) lemma eqvt_at_apply'': assumes "eqvt_at f x" shows "(p \ f) (p \ x) = f (p \ x)" -by (metis (hide_lams, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1)) +by (metis (opaque_lifting, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1)) lemma size_list_eqvt[eqvt]: "p \ size_list f x = size_list (p \ f) (p \ x)" proof (induction x) case (Cons x xs) have "f x = p \ (f x)" by (simp add: permute_pure) also have "... = (p \ f) (p \ x)" by simp with Cons show ?case by (auto simp add: permute_pure) qed simp section \ Freshness via equivariance \ lemma eqvt_fresh_cong1: "(\p x. p \ (f x) = f (p \ x)) \ a \ x \ a \ f x " apply (rule fresh_fun_eqvt_app[of f]) apply (rule eqvtI) apply (rule eq_reflection) apply (rule ext) apply (metis permute_fun_def permute_minus_cancel(1)) apply assumption done lemma eqvt_fresh_cong2: assumes eqvt: "(\p x y. p \ (f x y) = f (p \ x) (p \ y))" and fresh1: "a \ x" and fresh2: "a \ y" shows "a \ f x y" proof- have "eqvt (\ (x,y). f x y)" using eqvt apply (- , auto simp add: eqvt_def) by (rule ext, auto, metis permute_minus_cancel(1)) moreover have "a \ (x, y)" using fresh1 fresh2 by auto ultimately have "a \ (\ (x,y). f x y) (x, y)" by (rule fresh_fun_eqvt_app) thus ?thesis by simp qed lemma eqvt_fresh_star_cong1: assumes eqvt: "(\p x. p \ (f x) = f (p \ x))" and fresh1: "a \* x" shows "a \* f x" by (metis fresh_star_def eqvt_fresh_cong1 assms) lemma eqvt_fresh_star_cong2: assumes eqvt: "(\p x y. p \ (f x y) = f (p \ x) (p \ y))" and fresh1: "a \* x" and fresh2: "a \* y" shows "a \* f x y" by (metis fresh_star_def eqvt_fresh_cong2 assms) lemma eqvt_fresh_cong3: assumes eqvt: "(\p x y z. p \ (f x y z) = f (p \ x) (p \ y) (p \ z))" and fresh1: "a \ x" and fresh2: "a \ y" and fresh3: "a \ z" shows "a \ f x y z" proof- have "eqvt (\ (x,y,z). f x y z)" using eqvt apply (- , auto simp add: eqvt_def) by(rule ext, auto, metis permute_minus_cancel(1)) moreover have "a \ (x, y, z)" using fresh1 fresh2 fresh3 by auto ultimately have "a \ (\ (x,y,z). f x y z) (x, y, z)" by (rule fresh_fun_eqvt_app) thus ?thesis by simp qed lemma eqvt_fresh_star_cong3: assumes eqvt: "(\p x y z. p \ (f x y z) = f (p \ x) (p \ y) (p \ z))" and fresh1: "a \* x" and fresh2: "a \* y" and fresh3: "a \* z" shows "a \* f x y z" by (metis fresh_star_def eqvt_fresh_cong3 assms) section \ Additional simplification rules \ lemma not_self_fresh[simp]: "atom x \ x \ False" by (metis fresh_at_base(2)) lemma fresh_star_singleton: "{ x } \* e \ x \ e" by (simp add: fresh_star_def) section \ Additional equivariance lemmas \ lemma eqvt_cases: fixes f x \ assumes eqvt: "\x. \ \ f x = f (\ \ x)" obtains "f x" "f (\ \ x)" | "\ f x " " \ f (\ \ x)" using assms[symmetric] by (cases "f x") auto lemma range_eqvt: "\ \ range Y = range (\ \ Y)" unfolding image_eqvt UNIV_eqvt .. lemma case_option_eqvt[eqvt]: "\ \ case_option d f x = case_option (\ \ d) (\ \ f) (\ \ x)" by(cases x)(simp_all) lemma supp_option_eqvt: "supp (case_option d f x) \ supp d \ supp f \ supp x" apply (cases x) apply (auto simp add: supp_Some ) apply (metis (mono_tags) Un_iff subsetCE supp_fun_app) done lemma funpow_eqvt[simp,eqvt]: "\ \ ((f :: 'a \ 'a::pt) ^^ n) = (\ \ f) ^^ (\ \ n)" by (induct n,simp, rule ext, simp, perm_simp,simp) lemma delete_eqvt[eqvt]: "\ \ AList.delete x \ = AList.delete (\ \ x) (\ \ \)" by (induct \, auto) lemma restrict_eqvt[eqvt]: "\ \ AList.restrict S \ = AList.restrict (\ \ S) (\ \ \)" unfolding AList.restrict_eq by perm_simp rule lemma supp_restrict: "supp (AList.restrict S \) \ supp \" by (induction \) (auto simp add: supp_Pair supp_Cons) lemma clearjunk_eqvt[eqvt]: "\ \ AList.clearjunk \ = AList.clearjunk (\ \ \)" by (induction \ rule: clearjunk.induct) auto lemma map_ran_eqvt[eqvt]: "\ \ map_ran f \ = map_ran (\ \ f) (\ \ \)" by (induct \, auto) lemma dom_perm: "dom (\ \ f) = \ \ (dom f)" unfolding dom_def by (perm_simp) (simp) lemmas dom_perm_rev[simp,eqvt] = dom_perm[symmetric] lemma ran_perm[simp]: "\ \ (ran f) = ran (\ \ f)" unfolding ran_def by (perm_simp) (simp) lemma map_add_eqvt[eqvt]: "\ \ (m1 ++ m2) = (\ \ m1) ++ (\ \ m2)" unfolding map_add_def by (perm_simp, rule) lemma map_of_eqvt[eqvt]: "\ \ map_of l = map_of (\ \ l)" by (induct l, simp add: permute_fun_def,simp,perm_simp,auto) lemma concat_eqvt[eqvt]: "\ \ concat l = concat (\ \ l)" by (induction l)(auto simp add: append_eqvt) lemma tranclp_eqvt[eqvt]: "\ \ tranclp P v\<^sub>1 v\<^sub>2 = tranclp (\ \ P) (\ \ v\<^sub>1) (\ \ v\<^sub>2)" unfolding tranclp_def by perm_simp rule lemma rtranclp_eqvt[eqvt]: "\ \ rtranclp P v\<^sub>1 v\<^sub>2 = rtranclp (\ \ P) (\ \ v\<^sub>1) (\ \ v\<^sub>2)" unfolding rtranclp_def by perm_simp rule lemma Set_filter_eqvt[eqvt]: "\ \ Set.filter P S = Set.filter (\ \ P) (\ \ S)" unfolding Set.filter_def by perm_simp rule lemma Sigma_eqvt'[eqvt]: "\ \ Sigma = Sigma" apply (rule ext) apply (rule ext) apply (subst permute_fun_def) apply (subst permute_fun_def) unfolding Sigma_def apply perm_simp apply (simp add: permute_self) done lemma override_on_eqvt[eqvt]: "\ \ (override_on m1 m2 S) = override_on (\ \ m1) (\ \ m2) (\ \ S)" by (auto simp add: override_on_def ) lemma card_eqvt[eqvt]: "\ \ (card S) = card (\ \ S)" by (cases "finite S", induct rule: finite_induct) (auto simp add: card_insert_if mem_permute_iff permute_pure) (* Helper lemmas provided by Christian Urban *) lemma Projl_permute: assumes a: "\y. f = Inl y" shows "(p \ (Sum_Type.projl f)) = Sum_Type.projl (p \ f)" using a by auto lemma Projr_permute: assumes a: "\y. f = Inr y" shows "(p \ (Sum_Type.projr f)) = Sum_Type.projr (p \ f)" using a by auto section \ Freshness lemmas \ lemma fresh_list_elem: assumes "a \ \" and "e \ set \" shows "a \ e" using assms by(induct \)(auto simp add: fresh_Cons) lemma set_not_fresh: "x \ set L \ \(atom x \ L)" by (metis fresh_list_elem not_self_fresh) lemma pure_fresh_star[simp]: "a \* (x :: 'a :: pure)" by (simp add: fresh_star_def pure_fresh) lemma supp_set_mem: "x \ set L \ supp x \ supp L" by (induct L) (auto simp add: supp_Cons) lemma set_supp_mono: "set L \ set L2 \ supp L \ supp L2" by (induct L)(auto simp add: supp_Cons supp_Nil dest:supp_set_mem) lemma fresh_star_at_base: fixes x :: "'a :: at_base" shows "S \* x \ atom x \ S" by (metis fresh_at_base(2) fresh_star_def) section \ Freshness and support for subsets of variables \ lemma supp_mono: "finite (B::'a::fs set) \ A \ B \ supp A \ supp B" by (metis infinite_super subset_Un_eq supp_of_finite_union) lemma fresh_subset: "finite B \ x \ (B :: 'a::at_base set) \ A \ B \ x \ A" by (auto dest:supp_mono simp add: fresh_def) lemma fresh_star_subset: "finite B \ x \* (B :: 'a::at_base set) \ A \ B \ x \* A" by (metis fresh_star_def fresh_subset) lemma fresh_star_set_subset: "x \* (B :: 'a::at_base list) \ set A \ set B \ x \* A" by (metis fresh_star_set fresh_star_subset[OF finite_set]) section \ The set of free variables of an expression \ definition fv :: "'a::pt \ 'b::at_base set" where "fv e = {v. atom v \ supp e}" lemma fv_eqvt[simp,eqvt]: "\ \ (fv e) = fv (\ \ e)" unfolding fv_def by simp lemma fv_Nil[simp]: "fv [] = {}" by (auto simp add: fv_def supp_Nil) lemma fv_Cons[simp]: "fv (x # xs) = fv x \ fv xs" by (auto simp add: fv_def supp_Cons) lemma fv_Pair[simp]: "fv (x, y) = fv x \ fv y" by (auto simp add: fv_def supp_Pair) lemma fv_append[simp]: "fv (x @ y) = fv x \ fv y" by (auto simp add: fv_def supp_append) lemma fv_at_base[simp]: "fv a = {a::'a::at_base}" by (auto simp add: fv_def supp_at_base) lemma fv_pure[simp]: "fv (a::'a::pure) = {}" by (auto simp add: fv_def pure_supp) lemma fv_set_at_base[simp]: "fv (l :: ('a :: at_base) list) = set l" by (induction l) auto lemma flip_not_fv: "a \ fv x \ b \ fv x \ (a \ b) \ x = x" by (metis flip_def fresh_def fv_def mem_Collect_eq swap_fresh_fresh) lemma fv_not_fresh: "atom x \ e \ x \ fv e" unfolding fv_def fresh_def by blast lemma fresh_fv: "finite (fv e :: 'a set) \ atom (x :: ('a::at_base)) \ (fv e :: 'a set) \ atom x \ e" unfolding fv_def fresh_def by (auto simp add: supp_finite_set_at_base) lemma finite_fv[simp]: "finite (fv (e::'a::fs) :: ('b::at_base) set)" proof- have "finite (supp e)" by (metis finite_supp) hence "finite (atom -` supp e :: 'b set)" apply (rule finite_vimageI) apply (rule inj_onI) apply (simp) done moreover have "(atom -` supp e :: 'b set) = fv e" unfolding fv_def by auto ultimately show ?thesis by simp qed definition fv_list :: "'a::fs \ 'b::at_base list" where "fv_list e = (SOME l. set l = fv e)" lemma set_fv_list[simp]: "set (fv_list e) = (fv e :: ('b::at_base) set)" proof- have "finite (fv e :: 'b set)" by (rule finite_fv) from finite_list[OF finite_fv] obtain l where "set l = (fv e :: 'b set)".. thus ?thesis unfolding fv_list_def by (rule someI) qed lemma fresh_fv_list[simp]: "a \ (fv_list e :: 'b::at_base list) \ a \ (fv e :: 'b::at_base set)" proof- have "a \ (fv_list e :: 'b::at_base list) \ a \ set (fv_list e :: 'b::at_base list)" by (rule fresh_set[symmetric]) also have "\ \ a \ (fv e :: 'b::at_base set)" by simp finally show ?thesis. qed section \ Other useful lemmas \ lemma pure_permute_id: "permute p = (\ x. (x::'a::pure))" by rule (simp add: permute_pure) lemma supp_set_elem_finite: assumes "finite S" and "(m::'a::fs) \ S" and "y \ supp m" shows "y \ supp S" using assms supp_of_finite_sets by auto lemmas fresh_star_Cons = fresh_star_list(2) lemma mem_permute_set: shows "x \ p \ S \ (- p \ x) \ S" by (metis mem_permute_iff permute_minus_cancel(2)) lemma flip_set_both_not_in: assumes "x \ S" and "x' \ S" shows "((x' \ x) \ S) = S" unfolding permute_set_def by (auto) (metis assms flip_at_base_simps(3))+ lemma inj_atom: "inj atom" by (metis atom_eq_iff injI) lemmas image_Int[OF inj_atom, simp] lemma eqvt_uncurry: "eqvt f \ eqvt (case_prod f)" unfolding eqvt_def by perm_simp simp lemma supp_fun_app_eqvt2: assumes a: "eqvt f" shows "supp (f x y) \ supp x \ supp y" proof- from supp_fun_app_eqvt[OF eqvt_uncurry [OF a]] have "supp (case_prod f (x,y)) \ supp (x,y)". thus ?thesis by (simp add: supp_Pair) qed lemma supp_fun_app_eqvt3: assumes a: "eqvt f" shows "supp (f x y z) \ supp x \ supp y \ supp z" proof- from supp_fun_app_eqvt2[OF eqvt_uncurry [OF a]] have "supp (case_prod f (x,y) z) \ supp (x,y) \ supp z". thus ?thesis by (simp add: supp_Pair) qed (* Fighting eta-contraction *) lemma permute_0[simp]: "permute 0 = (\ x. x)" by auto lemma permute_comp[simp]: "permute x \ permute y = permute (x + y)" by auto lemma map_permute: "map (permute p) = permute p" apply rule apply (induct_tac x) apply auto done lemma fresh_star_restrictA[intro]: "a \* \ \ a \* AList.restrict V \" by (induction \) (auto simp add: fresh_star_Cons) lemma Abs_lst_Nil_eq[simp]: "[[]]lst. (x::'a::fs) = [xs]lst. x' \ (([],x) = (xs, x'))" apply rule apply (frule Abs_lst_fcb2[where f = "\ x y _ . (x,y)" and as = "[]" and bs = "xs" and c = "()"]) apply (auto simp add: fresh_star_def) done lemma Abs_lst_Nil_eq2[simp]: "[xs]lst. (x::'a::fs) = [[]]lst. x' \ ((xs,x) = ([], x'))" by (subst eq_commute) auto lemma prod_cases8 [cases type]: obtains (fields) a b c d e f g h where "y = (a, b, c, d, e, f, g,h)" by (cases y, case_tac g) blast lemma prod_induct8 [case_names fields, induct type]: "(\a b c d e f g h. P (a, b, c, d, e, f, g, h)) \ P x" by (cases x) blast lemma prod_cases9 [cases type]: obtains (fields) a b c d e f g h i where "y = (a, b, c, d, e, f, g,h,i)" by (cases y, case_tac h) blast lemma prod_induct9 [case_names fields, induct type]: "(\a b c d e f g h i. P (a, b, c, d, e, f, g, h, i)) \ P x" by (cases x) blast named_theorems nominal_prod_simps named_theorems ms_fresh "Facts for helping with freshness proofs" lemma fresh_prod2[nominal_prod_simps,ms_fresh]: "x \ (a,b) = (x \ a \ x \ b )" using fresh_def supp_Pair by fastforce lemma fresh_prod3[nominal_prod_simps,ms_fresh]: "x \ (a,b,c) = (x \ a \ x \ b \ x \ c)" using fresh_def supp_Pair by fastforce lemma fresh_prod4[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d) = (x \ a \ x \ b \ x \ c \ x \ d)" using fresh_def supp_Pair by fastforce lemma fresh_prod5[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e)" using fresh_def supp_Pair by fastforce lemma fresh_prod6[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f)" using fresh_def supp_Pair by fastforce lemma fresh_prod7[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g)" using fresh_def supp_Pair by fastforce lemma fresh_prod8[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h )" using fresh_def supp_Pair by fastforce lemma fresh_prod9[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h,i) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h \ x \ i)" using fresh_def supp_Pair by fastforce lemma fresh_prod10[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h,i,j) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h \ x \ i \ x \ j)" using fresh_def supp_Pair by fastforce lemma fresh_prod12[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h,i,j,k,l) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h \ x \ i \ x \ j \ x \ k \ x \ l)" using fresh_def supp_Pair by fastforce lemmas fresh_prodN = fresh_Pair fresh_prod3 fresh_prod4 fresh_prod5 fresh_prod6 fresh_prod7 fresh_prod8 fresh_prod9 fresh_prod10 fresh_prod12 lemma fresh_prod2I: fixes x and x1 and x2 assumes "x \ x1" and "x \ x2" shows "x \ (x1,x2)" using fresh_prod2 assms by auto lemma fresh_prod3I: fixes x and x1 and x2 and x3 assumes "x \ x1" and "x \ x2" and "x \ x3" shows "x \ (x1,x2,x3)" using fresh_prod3 assms by auto lemma fresh_prod4I: fixes x and x1 and x2 and x3 and x4 assumes "x \ x1" and "x \ x2" and "x \ x3" and "x \ x4" shows "x \ (x1,x2,x3,x4)" using fresh_prod4 assms by auto lemma fresh_prod5I: fixes x and x1 and x2 and x3 and x4 and x5 assumes "x \ x1" and "x \ x2" and "x \ x3" and "x \ x4" and "x \ x5" shows "x \ (x1,x2,x3,x4,x5)" using fresh_prod5 assms by auto lemma flip_collapse[simp]: fixes b1::"'a::pt" and bv1::"'b::at" and bv2::"'b::at" assumes "atom bv2 \ b1" and "atom c \ (bv1,bv2,b1)" and "bv1 \ bv2" shows "(bv2 \ c) \ (bv1 \ bv2) \ b1 = (bv1 \ c) \ b1" proof - have "c \ bv1" and "bv2 \ bv1" using assms by auto+ hence "(bv2 \ c) + (bv1 \ bv2) + (bv2 \ c) = (bv1 \ c)" using flip_triple[of c bv1 bv2] flip_commute by metis hence "(bv2 \ c) \ (bv1 \ bv2) \ (bv2 \ c) \ b1 = (bv1 \ c) \ b1" using permute_plus by metis thus ?thesis using assms flip_fresh_fresh by force qed lemma triple_eqvt[simp]: "p \ (x, b,c) = (p \ x, p \ b , p \ c)" proof - have "(x,b,c) = (x,(b,c))" by simp thus ?thesis using Pair_eqvt by simp qed lemma lst_fst: fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs" assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))" shows " ([[atom x]]lst. t1 = [[atom x']]lst. t1')" proof - have "(\c. atom c \ (t2,t2') \ atom c \ (x, x', t1, t1') \ (x \ c) \ t1 = (x' \ c) \ t1')" proof(rule,rule,rule) fix c::'a assume "atom c \ (t2,t2')" and "atom c \ (x, x', t1, t1')" hence "atom c \ (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp thus " (x \ c) \ t1 = (x' \ c) \ t1'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp qed thus ?thesis using Abs1_eq_iff_all(3)[of x t1 x' t1' "(t2,t2')"] by simp qed lemma lst_snd: fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs" assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))" shows " ([[atom x]]lst. t2 = [[atom x']]lst. t2')" proof - have "(\c. atom c \ (t1,t1') \ atom c \ (x, x', t2, t2') \ (x \ c) \ t2 = (x' \ c) \ t2')" proof(rule,rule,rule) fix c::'a assume "atom c \ (t1,t1')" and "atom c \ (x, x', t2, t2')" hence "atom c \ (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp thus " (x \ c) \ t2 = (x' \ c) \ t2'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp qed thus ?thesis using Abs1_eq_iff_all(3)[of x t2 x' t2' "(t1,t1')"] by simp qed lemma lst_head_cons_pair: fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list" assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)" shows "[[atom y1]]lst. (x1,xs1) = [[atom y2]]lst. (x2,xs2)" proof(subst Abs1_eq_iff_all(3)[of y1 "(x1,xs1)" y2 "(x2,xs2)"],rule,rule,rule) fix c::'a assume "atom c \ (x1#xs1,x2#xs2)" and "atom c \ (y1, y2, (x1, xs1), x2, xs2)" thus "(y1 \ c) \ (x1, xs1) = (y2 \ c) \ (x2, xs2)" using assms Abs1_eq_iff_all(3) by auto qed lemma lst_head_cons_neq_nil: fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list" assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (xs2)" shows "xs2 \ []" proof assume as:"xs2 = []" thus False using Abs1_eq_iff(3)[of y1 "x1#xs1" y2 Nil] assms as by auto qed lemma lst_head_cons: fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list" assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)" shows "[[atom y1]]lst. x1 = [[atom y2]]lst. x2" and "[[atom y1]]lst. xs1 = [[atom y2]]lst. xs2" using lst_head_cons_pair lst_fst lst_snd assms by metis+ lemma lst_pure: fixes x1::"'a ::at" and t1::"'b::pure" and x2::"'a ::at" and t2::"'b::pure" assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" shows "t1=t2" using assms Abs1_eq_iff_all(3) pure_fresh flip_fresh_fresh by (metis Abs1_eq(3) permute_pure) lemma lst_supp: assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" shows "supp t1 - {atom x1} = supp t2 - {atom x2}" proof - have "supp ([[atom x1]]lst.t1) = supp ([[atom x2]]lst.t2)" using assms by auto thus ?thesis using Abs_finite_supp by (metis assms empty_set list.simps(15) supp_lst.simps) qed lemma lst_supp_subset: assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "supp t1 \ {atom x1} \ B" shows "supp t2 \ {atom x2} \ B" using assms lst_supp by fast lemma projl_inl_eqvt: fixes \ :: perm shows "\ \ (projl (Inl x)) = projl (Inl (\ \ x))" unfolding projl_def Inl_eqvt by simp end diff --git a/thys/MiniSail/Safety.thy b/thys/MiniSail/Safety.thy --- a/thys/MiniSail/Safety.thy +++ b/thys/MiniSail/Safety.thy @@ -1,1812 +1,1812 @@ (*<*) theory Safety imports Operational IVSubstTypingL BTVSubstTypingL begin (*>*) -method supp_calc = (metis (mono_tags, hide_lams) pure_supp c.supp e.supp v.supp supp_l_empty opp.supp sup_bot.right_neutral supp_at_base) +method supp_calc = (metis (mono_tags, opaque_lifting) pure_supp c.supp e.supp v.supp supp_l_empty opp.supp sup_bot.right_neutral supp_at_base) declare infer_e.intros[simp] declare infer_e.intros[intro] chapter \Safety\ text \Lemmas about the operational semantics leading up to progress and preservation and then safety.\ section \Store Lemmas\ abbreviation delta_ext (" _ \ _ ") where "delta_ext \ \' \ (setD \ \ setD \')" nominal_function dc_of :: "branch_s \ string" where "dc_of (AS_branch dc _ _) = dc" apply(auto,simp add: eqvt_def dc_of_graph_aux_def) using s_branch_s_branch_list.exhaust by metis nominal_termination (eqvt) by lexicographic_order lemma delta_sim_fresh: assumes "\ \ \ \ \" and "atom u \ \" shows "atom u \ \" using assms proof(induct rule : delta_sim.inducts) case (delta_sim_nilI \) then show ?case using fresh_def supp_DNil by blast next case (delta_sim_consI \ \ \ v \ u') hence "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" using check_v_wf by meson hence "supp \ = {}" using wfT_supp by fastforce moreover have "atom u \ u'" using delta_sim_consI fresh_Cons fresh_Pair by blast moreover have "atom u \ \" using delta_sim_consI fresh_Cons by blast ultimately show ?case using fresh_Pair fresh_DCons fresh_def by blast qed lemma delta_sim_v: fixes \::\ assumes "\ \ \ \ \ " and "(u,v) \ set \" and "(u,\) \ setD \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" shows "\ ; {||} ; GNil \ v \ \" using assms proof(induct \ arbitrary: \) case Nil then show ?case by auto next case (Cons uv \) obtain u' and v' where uv : "uv=(u',v')" by fastforce show ?case proof(cases "u'=u") case True hence *:"\ \ ((u,v')#\) \ \" using uv Cons by blast then obtain \' and \' where tt: "\ ; {||} ; GNil \ v' \ \' \ u \ fst ` set \ \ \ = (u,\')#\<^sub>\\'" using delta_sim_elims(3)[OF *] by metis moreover hence "v'=v" using Cons True by (metis Pair_inject fst_conv image_eqI set_ConsD uv) moreover have "\=\'" using wfD_unique tt Cons setD.simps list.set_intros by blast ultimately show ?thesis by metis next case False hence *:"\ \ ((u',v')#\) \ \" using uv Cons by blast then obtain \' and \' where tt: "\ \ \ \ \' \ \ ; {||} ; GNil \ v' \ \' \ u' \ fst ` set \ \ \ = (u',\')#\<^sub>\\'" using delta_sim_elims(3)[OF *] by metis moreover hence "\ ; {||} ; GNil \\<^sub>w\<^sub>f \'" using wfD_elims Cons delta_sim_elims by metis ultimately show ?thesis using Cons using False by auto qed qed lemma delta_sim_delta_lookup: assumes "\ \ \ \ \ " and "(u, \ z : b | c \) \ setD \" shows "\v. (u,v) \ set \" using assms by(induct rule: delta_sim.inducts,auto+) lemma update_d_stable: "fst ` set \ = fst ` set (update_d \ u v)" proof(induct \) case Nil then show ?case by auto next case (Cons a \) then show ?case using update_d.simps by (metis (no_types, lifting) eq_fst_iff image_cong image_insert list.simps(15) prod.exhaust_sel) qed lemma update_d_sim: fixes \::\ assumes "\ \ \ \ \" and "\ ; {||} ; GNil \ v \ \" and "(u,\) \ setD \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" shows "\ \ (update_d \ u v) \ \" using assms proof(induct \ arbitrary: \) case Nil then show ?case using delta_sim_consI by simp next case (Cons uv \) obtain u' and v' where uv : "uv=(u',v')" by fastforce hence *:"\ \ ((u',v')#\) \ \" using uv Cons by blast then obtain \' and \' where tt: "\ \ \ \ \' \ \ ; {||} ; GNil \ v' \ \' \ u' \ fst ` set \ \ \ = (u',\')#\<^sub>\\'" using delta_sim_elims * by metis show ?case proof(cases "u=u'") case True then have "(u,\') \ setD \" using tt by auto then have "\ = \'" using Cons wfD_unique by metis moreover have "update_d ((u',v')#\) u v = ((u',v)#\)" using update_d.simps True by presburger ultimately show ?thesis using delta_sim_consI tt Cons True by (simp add: tt uv) next case False have "\ \ (u',v') # (update_d \ u v) \ (u',\')#\<^sub>\\'" proof(rule delta_sim_consI) show "\ \ update_d \ u v \ \' " using Cons using delta_sim_consI delta_sim.simps update_d.simps Cons delta_sim_elims uv tt False fst_conv set_ConsD wfG_elims wfD_elims by (metis setD_ConsD) show "\ ; {||} ; GNil \ v' \ \'" using tt by auto show "u' \ fst ` set (update_d \ u v)" using update_d.simps Cons update_d_stable tt by auto qed thus ?thesis using False update_d.simps uv by (simp add: tt) qed qed section \Preservation\ text \Types are preserved under reduction step. Broken down into lemmas about different operations\ subsection \Function Application\ lemma check_s_x_fresh: fixes x::x and s::s assumes "\ ; \ ; B ; GNil ; D \ s \ \" shows "atom x \ s \ atom x \ \ \ atom x \ D" proof - have "\ ; \ ; B ; GNil ; D \\<^sub>w\<^sub>f s : b_of \" using check_s_wf[OF assms] by auto moreover have "\ ; B ; GNil \\<^sub>w\<^sub>f \ " using check_s_wf assms by auto moreover have "\ ; B ; GNil \\<^sub>w\<^sub>f D" using check_s_wf assms by auto ultimately show ?thesis using wf_supp x_fresh_u by (meson fresh_GNil wfS_x_fresh wfT_x_fresh wfD_x_fresh) qed lemma check_funtyp_subst_b: fixes b'::b assumes "check_funtyp \ \ {|bv|} (AF_fun_typ x b c \ s)" and \ \ ; {||} \\<^sub>w\<^sub>f b' \ shows "check_funtyp \ \ {||} (AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \[bv::=b']\<^sub>\\<^sub>b s[bv::=b']\<^sub>s\<^sub>b)" using assms proof (nominal_induct "{|bv|}" "AF_fun_typ x b c \ s" rule: check_funtyp.strong_induct) case (check_funtypI x' \ \ c' s' \') have "check_funtyp \ \ {||} (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \'[bv::=b']\<^sub>\\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)" proof show \atom x' \ (\, \, {||}::bv fset, b[bv::=b']\<^sub>b\<^sub>b)\ using check_funtypI fresh_prodN x_fresh_b fresh_empty_fset by metis have \ \ ; \ ; {||} ; ((x', b, c') #\<^sub>\ GNil)[bv::=b']\<^sub>\\<^sub>b ; []\<^sub>\[bv::=b']\<^sub>\\<^sub>b \ s'[bv::=b']\<^sub>s\<^sub>b \ \'[bv::=b']\<^sub>\\<^sub>b\ proof(rule subst_b_check_s) show \ \ ; {||} \\<^sub>w\<^sub>f b' \ using check_funtypI by metis show \{|bv|} = {|bv|}\ by auto show \ \ ; \ ; {|bv|} ; (x', b, c') #\<^sub>\ GNil ; []\<^sub>\ \ s' \ \'\ using check_funtypI by metis qed thus \ \ ; \ ; {||} ; (x', b[bv::=b']\<^sub>b\<^sub>b, c'[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ GNil ; []\<^sub>\ \ s'[bv::=b']\<^sub>s\<^sub>b \ \'[bv::=b']\<^sub>\\<^sub>b\ using subst_gb.simps subst_db.simps by simp qed moreover have "(AF_fun_typ x b c \ s) = (AF_fun_typ x' b c' \' s')" using fun_typ.eq_iff check_funtypI by metis moreover hence "(AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \[bv::=b']\<^sub>\\<^sub>b s[bv::=b']\<^sub>s\<^sub>b) = (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \'[bv::=b']\<^sub>\\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)" using subst_ft_b.simps by metis ultimately show ?case by metis qed lemma funtyp_simple_check: fixes s::s and \::\ and \::\ and v::v assumes "check_funtyp \ \ ({||}::bv fset) (AF_fun_typ x b c \ s)" and "\ ; {||} ; GNil \ v \ \ x : b | c \" shows "\ ; \ ; {||} ; GNil ; DNil \ s[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" using assms proof(nominal_induct " ({||}::bv fset)" "AF_fun_typ x b c \ s" avoiding: v x rule: check_funtyp.strong_induct) case (check_funtypI x' \ \ c' s' \') hence eq1: "\ x' : b | c' \ = \ x : b | c \" using funtyp_eq_iff_equalities by metis obtain x'' and c'' where xf:"\ x : b | c \ = \ x'' : b | c'' \ \ atom x'' \ (x',v) \ atom x'' \ (x,c)" using obtain_fresh_z3 by metis moreover have "atom x' \ c''" proof - have "supp \ x'' : b | c'' \ = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis hence "supp c'' \ { atom x'' }" using \.supp eq1 xf by (auto simp add: freshers) moreover have "atom x' \ atom x''" using xf fresh_Pair fresh_x_neq by metis ultimately show ?thesis using xf fresh_Pair fresh_x_neq fresh_def fresh_at_base by blast qed ultimately have eq2: "c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c'" using eq1 type_eq_subst_eq3(1)[of x' b c' x'' b c''] by metis have "atom x' \ c" proof - have "supp \ x : b | c \ = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis hence "supp c \ { atom x }" using \.supp by auto moreover have "atom x \ atom x'" using check_funtypI fresh_Pair fresh_x_neq by metis ultimately show ?thesis using fresh_def by force qed hence eq: " c[x::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c' \ s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v = \[x::=v]\<^sub>\\<^sub>v" using funtyp_eq_iff_equalities type_eq_subst_eq3 check_funtypI by metis have " \ ; \ ; {||} ; ((x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil)[x'::=v]\<^sub>\\<^sub>v ; []\<^sub>\[x'::=v]\<^sub>\\<^sub>v \ s'[x'::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v" proof(rule subst_check_check_s ) show \\ ; {||} ; GNil \ v \ \ x'' : b | c'' \\ using check_funtypI eq1 xf by metis show \atom x'' \ (x', v)\ using check_funtypI fresh_x_neq fresh_Pair xf by metis show \ \ ; \ ; {||} ; (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; []\<^sub>\ \ s' \ \'\ using check_funtypI eq2 by metis show \ (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil = GNil @ (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil\ using append_g.simps by auto qed hence " \; \; {||}; GNil; []\<^sub>\ \ s'[x'::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v" using subst_gv.simps subst_dv.simps by auto thus ?case using eq by auto qed lemma funtypq_simple_check: fixes s::s and \::\ and \::\ and v::v assumes "check_funtypq \ \ (AF_fun_typ_none (AF_fun_typ x b c t s))" and "\ ; {||} ; GNil \ v \ \ x : b | c \" shows "\; \; {||}; GNil; DNil \ s[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v" using assms proof(nominal_induct "(AF_fun_typ_none (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct) case (check_fundefq_simpleI \ \ x' c' t' s') hence eq: "\ x : b | c \ = \ x' : b | c' \ \ s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v = t'[x'::=v]\<^sub>\\<^sub>v" using funtyp_eq_iff_equalities by metis hence "\; \; {||}; GNil; []\<^sub>\ \ s'[x'::=v]\<^sub>s\<^sub>v \ t'[x'::=v]\<^sub>\\<^sub>v" using funtyp_simple_check[OF check_fundefq_simpleI(1)] check_fundefq_simpleI by metis thus ?case using eq by metis qed lemma funtyp_poly_eq_iff_equalities: assumes "[[atom bv']]lst. AF_fun_typ x' b'' c' t' s' = [[atom bv]]lst. AF_fun_typ x b c t s" shows "\ x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \ = \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ \ s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t'[bv'::=b']\<^sub>\\<^sub>b[x'::=v]\<^sub>\\<^sub>v = t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" proof - have "subst_ft_b (AF_fun_typ x' b'' c' t' s') bv' b' = subst_ft_b (AF_fun_typ x b c t s) bv b'" using subst_b_flip_eq_two subst_b_fun_typ_def assms by metis thus ?thesis using fun_typ.eq_iff subst_ft_b.simps funtyp_eq_iff_equalities subst_tb.simps by (metis (full_types) assms fun_poly_arg_unique) qed lemma funtypq_poly_check: fixes s::s and \::\ and \::\ and v::v and b'::b assumes "check_funtypq \ \ (AF_fun_typ_some bv (AF_fun_typ x b c t s))" and "\ ; {||} ; GNil \ v \ \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" and "\ ; {||} \\<^sub>w\<^sub>f b'" shows "\; \; {||}; GNil; DNil \ s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" using assms proof(nominal_induct "(AF_fun_typ_some bv (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct) case (check_funtypq_polyI bv' \ \ x' b'' c' t' s') hence **:"\ x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \ = \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ \ s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t'[bv'::=b']\<^sub>\\<^sub>b[x'::=v]\<^sub>\\<^sub>v = t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" using funtyp_poly_eq_iff_equalities by metis have *:"check_funtyp \ \ {||} (AF_fun_typ x' b''[bv'::=b']\<^sub>b\<^sub>b (c'[bv'::=b']\<^sub>c\<^sub>b) (t'[bv'::=b']\<^sub>\\<^sub>b) s'[bv'::=b']\<^sub>s\<^sub>b)" using check_funtyp_subst_b[OF check_funtypq_polyI(5) check_funtypq_polyI(8)] by metis moreover have "\ ; {||} ; GNil \ v \ \ x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \" using ** check_funtypq_polyI by metis ultimately have "\; \; {||}; GNil; []\<^sub>\ \ s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v \ t'[bv'::=b']\<^sub>\\<^sub>b[x'::=v]\<^sub>\\<^sub>v" using funtyp_simple_check[OF *] check_funtypq_polyI by metis thus ?case using ** by metis qed lemma fundef_simple_check: fixes s::s and \::\ and \::\ and v::v assumes "check_fundef \ \ (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" and "\ ; {||} ; GNil \ v \ \ x : b | c \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" shows "\; \; {||}; GNil; \ \ s[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v" using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct) case (check_fundefI \ \) then show ?case using funtypq_simple_check[THEN check_s_d_weakening(1) ] setD.simps by auto qed lemma fundef_poly_check: fixes s::s and \::\ and \::\ and v::v and b'::b assumes "check_fundef \ \ (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" and "\ ; {||} ; GNil \ v \ \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" and "\ ; {||} \\<^sub>w\<^sub>f b'" shows "\; \; {||}; GNil; \ \ s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct) case (check_fundefI \ \) then show ?case using funtypq_poly_check[THEN check_s_d_weakening(1) ] setD.simps by auto qed lemma preservation_app: assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f" and "(\fd\set \. check_fundef \ \ fd)" shows "\ ; \ ; B ; G ; \ \ ss \ \ \ B = {||} \ G = GNil \ ss = LET x = (AE_app f v) IN s \ \; \; {||}; GNil; \ \ LET x : (\1'[x1::=v]\<^sub>\\<^sub>v) = (s1'[x1::=v]\<^sub>s\<^sub>v) IN s \ \" and "check_branch_s \ \ \ GNil \ tid dc const v cs \ \ True" and "check_branch_list \ \ \ \ \ tid dclist v css \ \ True" using assms proof(nominal_induct \ and \ and \ avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_letI x2 \ \ \ \ \ e \ z s2 b c) hence eq: "e = (AE_app f v)" by simp hence *:"\ ; \ ; {||} ;GNil ; \ \ (AE_app f v) \ \ z : b | c \" using check_letI by auto then obtain x3 b3 c3 \3 s3 where **:"\ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \3 s3))) = lookup_fun \ f \ \ ; {||} ; GNil \ v \ \ x3 : b3 | c3 \ \ atom x3 \ (\, \, ({||}::bv fset), GNil, \, v, \ z : b | c \) \ \3[x3::=v]\<^sub>\\<^sub>v = \ z : b | c \" using infer_e_elims(6)[OF *] subst_defs by metis obtain z3 where z3:"\ x3 : b3 | c3 \ = \ z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \ \ atom z3 \ (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis have seq:"[[atom x3]]lst. s3 = [[atom x1]]lst. s1'" using fun_def_eq check_letI ** option.inject by metis let ?ft = "AF_fun_typ x3 b3 c3 \3 s3" have sup: "supp \3 \ { atom x3} \ supp s3 \ { atom x3}" using wfPhi_f_supp ** by metis have "\; \; {||}; GNil; \ \ AS_let2 x2 \3[x3::=v]\<^sub>\\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 \ \" proof show \atom x2 \ (\, \, {||}::bv fset, GNil, \, \3[x3::=v]\<^sub>\\<^sub>v, s3[x3::=v]\<^sub>s\<^sub>v, \)\ unfolding fresh_prodN using check_letI fresh_subst_v_if subst_v_\_def sup by (metis all_not_in_conv fresh_def fresh_empty_fset fresh_subst_sv_if fresh_subst_tv_if singleton_iff subset_singleton_iff) show \ \; \; {||}; GNil; \ \ s3[x3::=v]\<^sub>s\<^sub>v \ \3[x3::=v]\<^sub>\\<^sub>v\ proof(rule fundef_simple_check) show \check_fundef \ \ (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \3 s3)))\ using ** check_letI lookup_fun_member by metis show \\ ; {||} ; GNil \ v \ \ x3 : b3 | c3 \\ using ** by auto show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using ** by auto qed show \ \ ; \ ; {||} ; (x2, b_of \3[x3::=v]\<^sub>\\<^sub>v, c_of \3[x3::=v]\<^sub>\\<^sub>v x2) #\<^sub>\ GNil ; \ \ s2 \ \\ using check_letI ** b_of.simps c_of.simps subst_defs by metis qed moreover have "AS_let2 x2 \3[x3::=v]\<^sub>\\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\1'[x1::=v]\<^sub>\\<^sub>v) (s1'[x1::=v]\<^sub>s\<^sub>v) s" proof - have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto moreover have " \3[x3::=v]\<^sub>\\<^sub>v = \1'[x1::=v]\<^sub>\\<^sub>v" using fun_ret_unique ** check_letI by metis moreover have "s3[x3::=v]\<^sub>s\<^sub>v = (s1'[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def seq by metis ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis qed ultimately show ?case using check_letI by auto qed(auto+) lemma fresh_subst_v_subst_b: fixes x2::x and tm::"'a::{has_subst_v,has_subst_b}" and x::x assumes "supp tm \ { atom bv, atom x }" and "atom x2 \ v" shows "atom x2 \ tm[bv::=b]\<^sub>b[x::=v]\<^sub>v" using assms proof(cases "x2=x") case True then show ?thesis using fresh_subst_v_if assms by blast next case False hence "atom x2 \ tm" using assms fresh_def fresh_at_base by force hence "atom x2 \ tm[bv::=b]\<^sub>b" using assms fresh_subst_if x_fresh_b False by force then show ?thesis using fresh_subst_v_if assms by auto qed lemma preservation_poly_app: assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f" and "(\fd\set \. check_fundef \ \ fd)" shows "\ ; \ ; B ; G ; \ \ ss \ \ \ B = {||} \ G = GNil \ ss = LET x = (AE_appP f b' v) IN s \ \ ; {||} \\<^sub>w\<^sub>f b' \ \; \; {||}; GNil; \ \ LET x : (\1'[bv1::=b']\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v) = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) IN s \ \" and "check_branch_s \ \ \ GNil \ tid dc const v cs \ \ True" and "check_branch_list \ \ \ \ \ tid dclist v css \ \ True" using assms proof(nominal_induct \ and \ and \ avoiding: v x1 rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_letI x2 \ \ \ \ \ e \ z s2 b c) hence eq: "e = (AE_appP f b' v)" by simp hence *:"\ ; \ ; {||} ;GNil ; \ \ (AE_appP f b' v) \ \ z : b | c \" using check_letI by auto then obtain x3 b3 c3 \3 s3 bv3 where **:"\ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ Some (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3))) = lookup_fun \ f \ \ ; {||} ; GNil \ v \ \ x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \ \ atom x3 \ GNil \ \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v = \ z : b | c \ \ \ ; {||} \\<^sub>w\<^sub>f b'" using infer_e_elims(21)[OF *] subst_defs by metis obtain z3 where z3:"\ x3 : b3 | c3 \ = \ z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \ \ atom z3 \ (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis let ?ft = "(AF_fun_typ x3 (b3[bv3::=b']\<^sub>b\<^sub>b) (c3[bv3::=b']\<^sub>c\<^sub>b) (\3[bv3::=b']\<^sub>\\<^sub>b) (s3[bv3::=b']\<^sub>s\<^sub>b))" have *:"check_fundef \ \ (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3)))" using ** check_letI lookup_fun_member by metis hence ftq:"check_funtypq \ \ (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3))" using check_fundef_elims by auto let ?ft = "AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3)" have sup: "supp \3 \ { atom x3, atom bv3} \ supp s3 \ { atom x3, atom bv3 }" using wfPhi_f_poly_supp_t wfPhi_f_poly_supp_s ** by metis have "\; \; {||}; GNil; \ \ AS_let2 x2 \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 \ \" proof show \atom x2 \ (\, \, {||}::bv fset, GNil, \, \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v, s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v, \)\ proof - have "atom x2 \ \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v" using fresh_subst_v_subst_b subst_v_\_def subst_b_\_def \ atom x2 \ v\ sup by fastforce moreover have "atom x2 \ s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v" using fresh_subst_v_subst_b subst_v_s_def subst_b_s_def \ atom x2 \ v\ sup proof - have "\b. atom x2 = atom x3 \ atom x2 \ s3[bv3::=b]\<^sub>b" by (metis (no_types) check_letI.hyps(1) fresh_subst_sv_if(1) fresh_subst_v_subst_b insert_commute subst_v_s_def sup) (* 140 ms *) then show ?thesis by (metis check_letI.hyps(1) fresh_subst_sb_if fresh_subst_sv_if(1) has_subst_b_class.subst_b_fresh_x x_fresh_b) (* 78 ms *) qed ultimately show ?thesis using fresh_prodN check_letI by metis qed show \ \; \; {||}; GNil; \ \ s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v \ \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v\ proof( rule fundef_poly_check) show \check_fundef \ \ (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3)))\ using ** lookup_fun_member check_letI by metis show \\ ; {||} ; GNil \ v \ \ x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \\ using ** by metis show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using ** by metis show \ \ ; {||} \\<^sub>w\<^sub>f b' \ using ** by metis qed show \ \ ; \ ; {||} ; (x2, b_of \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v, c_of \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v x2) #\<^sub>\ GNil ; \ \ s2 \ \\ using check_letI ** b_of.simps c_of.simps subst_defs by metis qed moreover have "AS_let2 x2 \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\1'[bv1::=b']\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v) (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) s" proof - have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto moreover have " \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v = \1'[bv1::=b']\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v" using fun_poly_ret_unique ** check_letI by metis moreover have "s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def fun_poly_body_unique ** check_letI by metis ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis qed ultimately show ?case using check_letI by auto qed(auto+) lemma check_s_plus: assumes "\; \; {||}; GNil; \ \ LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \ \" shows "\; \; {||}; GNil; \ \ LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \ \" proof - obtain t1 where 1: "\; \; {||}; GNil; \ \ AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)) \ t1" using assms check_s_elims by metis then obtain z1 where 2: "t1 = \ z1 : B_int | CE_val (V_var z1) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \" using infer_e_plus by metis obtain z2 where 3: \\ ; \ ; {||} ; GNil ; \ \ AE_val (V_lit (L_num (n1+n2))) \ \ z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \\ using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1 by (simp add: fresh_GNil) let ?e = " (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)))" show ?thesis proof(rule subtype_let) show "\ ; \ ; {||} ; GNil ; \ \ LET x = ?e IN s' \ \" using assms by auto show "\ ; \ ; {||} ; GNil ; \ \ ?e \ t1" using 1 by auto show "\ ; \ ; {||} ; GNil ; \ \ [ [ L_num (n1 + n2) ]\<^sup>v ]\<^sup>e \ \ z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \" using 3 by auto show "\ ; {||} ; GNil \ \ z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \ \ t1" using subtype_bop_arith by (metis "1" \\thesis. (\z1. t1 = \ z1 : B_int | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ plus [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ \ thesis) \ thesis\ infer_e_wf(2) opp.distinct(1) type_for_lit.simps(3)) qed qed lemma check_s_leq: assumes "\ ; \ ; {||} ; GNil ; \ \ LET x = (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \ \" shows "\; \; {||}; GNil; \ \ LET x = (AE_val (V_lit (if (n1 \ n2) then L_true else L_false))) IN s' \ \" proof - obtain t1 where 1: "\; \; {||}; GNil; \ \ AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)) \ t1" using assms check_s_elims by metis then obtain z1 where 2: "t1 = \ z1 : B_bool | CE_val (V_var z1) == CE_op LEq ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \" using infer_e_leq by auto obtain z2 where 3: \\ ; \ ; {||} ; GNil ; \ \ AE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \\ using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1 fresh_GNil by simp show ?thesis proof(rule subtype_let) show \ \; \; {||}; GNil; \ \ AS_let x (AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v) s' \ \\ using assms by auto show \\; \; {||}; GNil; \ \ AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v \ t1\ using 1 by auto show \\; \; {||}; GNil; \ \ [ [ if n1 \ n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \\ using 3 by auto show \\ ; {||} ; GNil \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \ \ t1\ using subtype_bop_arith[where opp=LEq] check_s_wf assms 2 by (metis opp.distinct(1) subtype_bop_arith type_l_eq) qed qed lemma check_s_eq: assumes "\ ; \ ; {||} ; GNil ; \ \ LET x = (AE_op Eq (V_lit (n1)) (V_lit ( n2))) IN s' \ \" shows "\; \; {||}; GNil; \ \ LET x = (AE_val (V_lit (if (n1 = n2) then L_true else L_false))) IN s' \ \" proof - obtain t1 where 1: "\; \; {||}; GNil; \ \ AE_op Eq (V_lit (n1)) (V_lit (n2)) \ t1" using assms check_s_elims by metis then obtain z1 where 2: "t1 = \ z1 : B_bool | CE_val (V_var z1) == CE_op Eq ([V_lit (n1)]\<^sup>c\<^sup>e) ([V_lit (n2)]\<^sup>c\<^sup>e) \" using infer_e_leq by auto obtain z2 where 3: \\ ; \ ; {||} ; GNil ; \ \ AE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \\ using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1 fresh_GNil by simp show ?thesis proof(rule subtype_let) show \ \; \; {||}; GNil; \ \ AS_let x (AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v) s' \ \\ using assms by auto show \\; \; {||}; GNil; \ \ AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v \ t1\ using 1 by auto show \\; \; {||}; GNil; \ \ [ [ if n1 = n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \\ using 3 by auto show \\ ; {||} ; GNil \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \ \ t1\ proof - have " \ z2 : B_bool | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ = t1" using 2 by (metis \_fresh_c fresh_opp_all infer_l_form2 infer_l_fresh ms_fresh_all(31) ms_fresh_all(33) obtain_fresh_z type_e_eq type_l_eq) moreover have "\ ; {||} \\<^sub>w\<^sub>f GNil" using assms wfX_wfY by fastforce moreover have "base_for_lit n1 = base_for_lit n2" using 1 infer_e_wf wfE_elims(12) wfV_elims by metis ultimately show ?thesis using subtype_bop_eq[OF \\ ; {||} \\<^sub>w\<^sub>f GNil\, of n1 n2 z2] by auto qed qed qed subsection \Operators\ lemma preservation_plus: assumes "\; \; \ \ \ \ , LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \ \ \" shows "\; \; \ \ \ \ , LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \ \ \" proof - have tt: "\; \; {||}; GNil; \ \ AS_let x (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) s' \ \" and dsim: "\ \ \ \ \" and fd:"(\fd\set \. check_fundef \ \ fd)" using assms config_type_elims by blast+ hence "\; \; {||}; GNil; \ \AS_let x (AE_val (V_lit (L_num (n1+n2)))) s' \ \" using check_s_plus assms by auto hence "\; \; \ \ \ \ , AS_let x (AE_val (V_lit ( (L_num (n1+n2))))) s' \ \ \" using dsim config_typeI fd by presburger then show ?thesis using dsim config_typeI by (meson order_refl) qed lemma preservation_leq: assumes "\; \; \ \ \ \ , AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \ \ \" shows "\; \; \ \ \ \ , AS_let x (AE_val (V_lit (((if (n1 \ n2) then L_true else L_false))))) s' \ \ \" proof - have tt: "\; \; {||}; GNil; \ \ AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \ \" and dsim: "\ \ \ \ \" and fd:"(\fd\set \. check_fundef \ \ fd)" using assms config_type_elims by blast+ hence "\; \; {||}; GNil; \ \AS_let x (AE_val (V_lit ( ((if (n1 \ n2) then L_true else L_false))))) s' \ \" using check_s_leq assms by auto hence "\; \; \ \ \ \ , AS_let x (AE_val (V_lit ( (((if (n1 \ n2) then L_true else L_false)))))) s' \ \ \" using dsim config_typeI fd by presburger then show ?thesis using dsim config_typeI by (meson order_refl) qed lemma preservation_eq: assumes "\; \; \ \ \ \ , AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \ \ \" shows "\; \; \ \ \ \ , AS_let x (AE_val (V_lit (((if (n1 = n2) then L_true else L_false))))) s' \ \ \" proof - have tt: "\; \; {||}; GNil; \ \ AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \ \" and dsim: "\ \ \ \ \" and fd:"(\fd\set \. check_fundef \ \ fd)" using assms config_type_elims by blast+ hence "\; \; {||}; GNil; \ \AS_let x (AE_val (V_lit ( ((if (n1 = n2) then L_true else L_false))))) s' \ \" using check_s_eq assms by auto hence "\; \; \ \ \ \ , AS_let x (AE_val (V_lit ( (((if (n1 = n2) then L_true else L_false)))))) s' \ \ \" using dsim config_typeI fd by presburger then show ?thesis using dsim config_typeI by (meson order_refl) qed subsection \Let Statements\ lemma subst_s_abs_lst: fixes s::s and sa::s and v'::v assumes "[[atom x]]lst. s = [[atom xa]]lst. sa" and "atom xa \ v \ atom x \ v" shows "s[x::=v]\<^sub>s\<^sub>v = sa[xa::=v]\<^sub>s\<^sub>v" proof - obtain c'::x where cdash: "atom c' \ (v, x, xa, s, sa)" using obtain_fresh by blast moreover have " (x \ c') \ s = (xa \ c') \ sa" proof - have "atom c' \ (s, sa) \ atom c' \ (x, xa, s, sa)" using cdash by auto thus ?thesis using assms by auto qed ultimately show ?thesis using assms using subst_sv_flip by auto qed lemma check_let_val: fixes v::v and s::s shows "\ ; \ ; B ; G ; \ \ ss \ \ \ B = {||} \ G = GNil \ ss = AS_let x (AE_val v) s \ ss = AS_let2 x t (AS_val v) s \ \; \; {||}; GNil; \ \ (s[x::=v]\<^sub>s\<^sub>v) \ \" and "check_branch_s \ \ \ GNil \ tid dc const v cs \ \ True" and "check_branch_list \ \ \ \ \ tid dclist v css \ \ True" proof(nominal_induct \ and \ and \ avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_letI x1 \ \ \ \ \ e \ z s1 b c) hence *:"e = AE_val v" by auto let ?G = "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \" have "\ ; \ ; \ ; ?G[x1::=v]\<^sub>\\<^sub>v ; \[x1::=v]\<^sub>\\<^sub>v \ s1[x1::=v]\<^sub>s\<^sub>v \ \[x1::=v]\<^sub>\\<^sub>v" proof(rule subst_infer_check_s(1)) show **:\ \ ; \ ; \ \ v \ \ z : b | c \\ using infer_e_elims check_letI * by fast thus \\ ; \ ; \ \ \ z : b | c \ \ \ z : b | c \\ using subtype_reflI infer_v_wf by metis show \atom z \ (x1, v)\ using check_letI fresh_Pair by auto show \\ ; \ ; \ ; (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \ ; \ \ s1 \ \\ using check_letI subst_defs by auto show "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \ = GNil @ (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \" by auto qed hence "\ ; \ ; \ ; \ ; \ \ s1[x1::=v]\<^sub>s\<^sub>v \ \" using check_letI by auto moreover have " s1[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v" by (metis (full_types) check_letI fresh_GNil infer_e_elims(7) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(2) subst_s_abs_lst wfG_x_fresh_in_v_simple) ultimately show ?case using check_letI by simp next case (check_let2I x1 \ \ \ \ \ t s1 \ s2 ) hence s1eq:"s1 = AS_val v" by auto let ?G = "(x1, b_of t, c_of t x1) #\<^sub>\ \" obtain z::x where *:"atom z \ (x1 , v,t)" using obtain_fresh_z by metis hence teq:"t = \ z : b_of t | c_of t z \" using b_of_c_of_eq by auto have "\ ; \ ; \ ; ?G[x1::=v]\<^sub>\\<^sub>v ; \[x1::=v]\<^sub>\\<^sub>v \ s2[x1::=v]\<^sub>s\<^sub>v \ \[x1::=v]\<^sub>\\<^sub>v" proof(rule subst_check_check_s(1)) obtain t' where "\ ; \ ; \ \ v \ t' \ \ ; \ ; \ \ t' \ t" using check_s_elims(1) check_let2I(10) s1eq by auto thus **:\ \ ; \ ; \ \ v \ \ z : b_of t | c_of t z \\ using check_v.intros teq by auto show "atom z \ (x1, v)" using * by auto show \ \ ; \ ; \ ; (x1, b_of t, c_of t x1) #\<^sub>\ \ ; \ \ s2 \ \\ using check_let2I by auto show "(x1, b_of t , c_of t x1) #\<^sub>\ \ = GNil @ (x1, b_of t, (c_of t z)[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \" using append_g.simps c_of_switch * fresh_prodN by metis qed hence "\ ; \ ; \ ; \ ; \ \ s2[x1::=v]\<^sub>s\<^sub>v \ \" using check_let2I by auto moreover have " s2[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v" using check_let2I fresh_GNil check_s_elims s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff subst_s_abs_lst wfG_x_fresh_in_v_simple proof - have "AS_let2 x t (AS_val v) s = AS_let2 x1 t s1 s2" by (metis check_let2I.prems(3) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(3)) (* 62 ms *) then show ?thesis by (metis (no_types) check_let2I check_let2I.prems(2) check_s_elims(1) fresh_GNil s_branch_s_branch_list.eq_iff(3) subst_s_abs_lst wfG_x_fresh_in_v_simple) (* 375 ms *) qed ultimately show ?case using check_let2I by simp qed(auto+) lemma preservation_let_val: assumes "\; \; \ \ \ \ , AS_let x (AE_val v) s \ \ \ \ \; \; \ \ \ \ , AS_let2 x t (AS_val v) s \ \ \" (is "?A \ ?B") shows " \\'. \; \; \' \ \ \ , s[x::=v]\<^sub>s\<^sub>v \ \ \ \ \ \ \'" proof - have tt: "\ \ \ \ \" and fd: "(\fd\set \. check_fundef \ \ fd)" using assms by blast+ have "?A \ ?B" using assms by auto then have "\; \; {||}; GNil; \ \ s[x::=v]\<^sub>s\<^sub>v \ \" proof assume "\; \; \ \ \ \ , AS_let x (AE_val v) s \ \ \" hence * : " \; \; {||}; GNil; \ \ AS_let x (AE_val v) s \ \" by blast thus ?thesis using check_let_val by simp next assume "\; \; \ \ \ \ , AS_let2 x t (AS_val v) s \ \ \" hence * : "\; \; {||}; GNil; \ \ AS_let2 x t (AS_val v) s \ \" by blast thus ?thesis using check_let_val by simp qed thus ?thesis using tt config_typeI fd order_refl by metis qed lemma check_s_fst_snd: assumes "fst_snd = AE_fst \ v=v1 \ fst_snd = AE_snd \ v=v2" and "\; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \" shows "\; \; {||}; GNil; \ \ AS_let x ( AE_val v) s' \ \" proof - have 1: \ \; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \\ using assms by auto then obtain t1 where 2:"\; \; {||}; GNil; \ \ (fst_snd (V_pair v1 v2)) \ t1" using check_s_elims by auto show ?thesis using subtype_let 1 2 assms by (meson infer_e_fst_pair infer_e_snd_pair) qed lemma preservation_fst_snd: assumes "\; \; \ \ \ \ , LET x = (fst_snd (V_pair v1 v2)) IN s' \ \ \" and "fst_snd = AE_fst \ v=v1 \ fst_snd = AE_snd \ v=v2" shows "\\'. \; \; \ \ \ \ , LET x = (AE_val v) IN s' \ \ \ \ \ \ \'" proof - have tt: "\; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \ \ \ \ \ \ \" using assms by blast hence t2: "\; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \" by auto moreover have "\fd\set \. check_fundef \ \ fd" using assms config_type_elims by auto ultimately show ?thesis using config_typeI order_refl tt assms check_s_fst_snd by metis qed inductive_cases check_branch_s_elims2[elim!]: "\ ; \ ; \ ; \ ; \; tid ; cons ; const ; v \ cs \ \" lemmas freshers = freshers atom_dom.simps toSet.simps fresh_def x_not_in_b_set declare freshers [simp] lemma subtype_eq_if: fixes t::\ and va::v assumes "\ ; \ ; \ \\<^sub>w\<^sub>f \ z : b_of t | c_of t z \" and "\ ; \ ; \ \\<^sub>w\<^sub>f \ z : b_of t | c IMP c_of t z \ " shows "\ ; \ ; \ \ \ z : b_of t | c_of t z \ \ \ z : b_of t | c IMP c_of t z \ " proof - obtain x::x where xf:"atom x \ ((\, \, \, z, c_of t z, z, c IMP c_of t z ),c)" using obtain_fresh by metis moreover have "\ ; \ ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ (c IMP c_of t z )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v " unfolding subst_cv.simps proof(rule valid_eq_imp) have "\ ; \ ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f (c IMP (c_of t z))[z::=[ x ]\<^sup>v]\<^sub>v " apply(rule wfT_wfC_cons) apply(simp add: assms, simp add: assms, unfold fresh_prodN ) using xf fresh_prodN by metis+ thus "\ ; \ ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v IMP (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v " using subst_cv.simps subst_defs by auto qed ultimately show ?thesis using subtype_baseI assms fresh_Pair subst_defs by metis qed lemma subtype_eq_if_\: fixes t::\ and va::v assumes "\ ; \ ; \ \\<^sub>w\<^sub>f t" and "\ ; \ ; \ \\<^sub>w\<^sub>f \ z : b_of t | c IMP c_of t z \ " and "atom z \ t" shows "\ ; \ ; \ \ t \ \ z : b_of t | c IMP c_of t z \ " proof - have "t = \ z : b_of t | c_of t z \" using b_of_c_of_eq assms by auto thus ?thesis using subtype_eq_if assms c_of.simps b_of.simps by metis qed lemma valid_conj: assumes "\ ; \ ; \ \ c1" and "\ ; \ ; \ \ c2" shows "\ ; \ ; \ \ c1 AND c2" proof show \ \ ; \ ; \ \\<^sub>w\<^sub>f c1 AND c2 \ using valid.simps wfC_conjI assms by auto show \\i. \ ; \ \ i \ i \ \ \ i \ c1 AND c2 \ proof(rule+) fix i assume *:"\ ; \ \ i \ i \ \ " thus "i \ c1 \ ~ True" using assms valid.simps using is_satis.cases by blast show "i \ c2 \ ~ True" using assms valid.simps using is_satis.cases * by blast qed qed subsection \Other Statements\ lemma check_if: fixes s'::s and cs::branch_s and css::branch_list and v::v shows "\; \; B; G; \ \ s' \ \ \ s' = IF (V_lit ll) THEN s1 ELSE s2 \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ G = GNil \ B = {||} \ ll = L_true \ s = s1 \ ll = L_false \ s = s2 \ \; \; {||}; GNil; \ \ s \ \" and "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" proof(nominal_induct \ and \ and \ rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_ifI z \ \ \ \ \ v s1 s2 \) obtain z' where teq: "\ = \ z' : b_of \ | c_of \ z' \ \ atom z' \ (z,\)" using obtain_fresh_z_c_of by metis hence ceq: "(c_of \ z')[z'::=[ z ]\<^sup>v]\<^sub>c\<^sub>v = (c_of \ z)" using c_of_switch fresh_Pair by metis have zf: " atom z \ c_of \ z'" using c_of_fresh check_ifI teq fresh_Pair fresh_at_base by metis hence 1:"\; \; {||}; GNil; \ \ s \ \ z : b_of \ | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \ z \" using check_ifI by auto moreover have 2:"\ ; {||} ; GNil \ (\ z : b_of \ | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \ z \) \ \" proof - have "\ ; {||} ; GNil \\<^sub>w\<^sub>f (\ z : b_of \ | CE_val (V_lit ll ) == CE_val (V_lit ll) IMP c_of \ z \)" using check_ifI check_s_wf by auto moreover have "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" using check_s_wf check_ifI by auto ultimately show ?thesis using subtype_if_simp[of \ " {||}" z "b_of \" ll "c_of \ z'" z'] using teq ceq zf subst_defs by metis qed ultimately show ?case using check_s_supertype(1) check_ifI by metis qed(auto+) lemma preservation_if: assumes "\; \; \ \ \ \ , IF (V_lit ll) THEN s1 ELSE s2 \ \ \" and "ll = L_true \ s = s1 \ ll = L_false \ s = s2" shows "\; \; \ \ \\, s\ \ \ \ setD \ \ setD \" proof - have *: "\; \; {||}; GNil; \ \ AS_if (V_lit ll) s1 s2 \ \ \ (\fd\set \. check_fundef \ \ fd)" using assms config_type_elims by metis hence "\; \; {||}; GNil; \ \ s \ \" using check_s_wf check_if assms by metis hence "\; \; \ \ \\, s\ \ \ \ setD \ \ setD \" using config_typeI * using assms(1) by blast thus ?thesis by blast qed lemma wfT_conj: assumes "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c1 \" and "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c2 \" shows "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c1 AND c2\" proof show \atom z \ (\, \, GNil)\ apply(unfold fresh_prodN, intro conjI) using wfTh_fresh wfT_wf assms apply metis using fresh_GNil x_not_in_b_set fresh_def by metis+ show \ \ ; \ \\<^sub>w\<^sub>f b \ using wfT_elims assms by metis have "\ ; \ ; (z, b, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f c1" using wfT_wfC fresh_GNil assms by auto moreover have "\ ; \ ; (z, b, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f c2" using wfT_wfC fresh_GNil assms by auto ultimately show \ \ ; \ ; (z, b, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f c1 AND c2 \ using wfC_conjI by auto qed lemma subtype_conj: assumes "\ ; \ ; GNil \ t \ \ z : b | c1 \" and "\ ; \ ; GNil \ t \ \ z : b | c2 \" shows "\ ; \ ; GNil \ \ z : b | c_of t z \ \ \ z : b | c1 AND c2 \" proof - have beq: "b_of t = b" using subtype_eq_base2 b_of.simps assms by metis obtain x::x where x:\atom x \ (\, \, GNil, z, c_of t z, z, c1 AND c2 )\ using obtain_fresh by metis thus ?thesis proof have "atom z \ t" using subtype_wf wfT_supp fresh_def x_not_in_b_set atom_dom.simps toSet.simps assms dom.simps by fastforce hence t:"t = \ z : b_of t | c_of t z \" using b_of_c_of_eq by auto thus \ \ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c_of t z \ \ using subtype_wf beq assms by auto show \\ ; \ ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ (c1 AND c2 )[z::=[ x ]\<^sup>v]\<^sub>v \ proof - have \\ ; \ ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ c1[z::=[ x ]\<^sup>v]\<^sub>v \ proof(rule subtype_valid) show \\ ; \ ; GNil \ t \ \ z : b | c1 \\ using assms by auto show \atom x \ GNil\ using fresh_GNil by auto show \t = \ z : b | c_of t z \\ using t beq by auto show \\ z : b | c1 \ = \ z : b | c1 \\ by auto qed moreover have \\ ; \ ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ c2[z::=[ x ]\<^sup>v]\<^sub>v \ proof(rule subtype_valid) show \\ ; \ ; GNil \ t \ \ z : b | c2 \\ using assms by auto show \atom x \ GNil\ using fresh_GNil by auto show \t = \ z : b | c_of t z \\ using t beq by auto show \\ z : b | c2 \ = \ z : b | c2 \\ by auto qed ultimately show ?thesis unfolding subst_cv.simps subst_v_c_def using valid_conj by metis qed thus \ \ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c1 AND c2 \ \ using subtype_wf wfT_conj assms by auto qed qed lemma infer_v_conj: assumes "\ ; \ ; GNil \ v \ \ z : b | c1 \" and "\ ; \ ; GNil \ v \ \ z : b | c2 \" shows "\ ; \ ; GNil \ v \ \ z : b | c1 AND c2 \" proof - obtain t1 where t1: "\ ; \ ; GNil \ v \ t1 \ \ ; \ ; GNil \ t1 \ \ z : b | c1 \" using assms check_v_elims by metis obtain t2 where t2: "\ ; \ ; GNil \ v \ t2 \ \ ; \ ; GNil \ t2 \ \ z : b | c2 \" using assms check_v_elims by metis have teq: "t1 = \ z : b | c_of t1 z \" using subtype_eq_base2 b_of.simps by (metis (full_types) b_of_c_of_eq fresh_GNil infer_v_t_wf t1 wfT_x_fresh) have "t1 = t2" using infer_v_uniqueness t1 t2 by auto hence " \ ; \ ; GNil \ \ z : b | c_of t1 z \ \ \ z : b | c1 AND c2 \" using subtype_conj t1 t2 by simp hence " \ ; \ ; GNil \ t1 \ \ z : b | c1 AND c2 \" using teq by auto thus ?thesis using t1 using check_v.intros by auto qed lemma wfG_conj: fixes c1::c assumes "\ ; \ \\<^sub>w\<^sub>f (x, b, c1 AND c2) #\<^sub>\ \" shows "\ ; \ \\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\ \" proof(cases "c1 \ {TRUE, FALSE}") case True then show ?thesis using assms wfG_cons2I wfG_elims wfX_wfY by metis next case False then show ?thesis using assms wfG_cons1I wfG_elims wfX_wfY wfC_elims by (metis wfG_elim2) qed lemma check_match: fixes s'::s and s::s and css::branch_list and cs::branch_s shows "\ ; \ ; \ ; \ ; \ \ s \ \ \ True" and "\ ; \ ; B ; G ; \ ; tid ; dc ; const ; vcons \ cs \ \ \ vcons = V_cons tid dc v \ B = {||} \ G = GNil \ cs = (dc x' \ s') \ \ ; {||} ; GNil \ v \ const \ \; \; {||}; GNil; \ \ s'[x'::=v]\<^sub>s\<^sub>v \ \" and "\ ; \ ; B ; G ; \ ; tid ; dclist ; vcons \ css \ \ \ distinct (map fst dclist) \ vcons = V_cons tid dc v \ B = {||} \ (dc, const) \ set dclist \ G = GNil \ Some (AS_branch dc x' s') = lookup_branch dc css \ \ ; {||} ; GNil \ v \ const \ \; \; {||}; GNil; \ \ s'[x'::=v]\<^sub>s\<^sub>v \ \" proof(nominal_induct \ and \ and \ avoiding: x' v rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_branch_list_consI \ \ \ \ \ tid consa consta va cs \ dclist cssa) then obtain xa and sa where cseq:"cs = AS_branch consa xa sa" using check_branch_s_elims2[OF check_branch_list_consI(1)] by metis show ?case proof(cases "dc = consa") case True hence "cs = AS_branch consa x' s'" using check_branch_list_consI cseq by (metis lookup_branch.simps(2) option.inject) moreover have "const = consta" using check_branch_list_consI distinct.simps by (metis True dclist_distinct_unique list.set_intros(1)) moreover have "va = V_cons tid consa v" using check_branch_list_consI True by auto ultimately show ?thesis using check_branch_list_consI by auto next case False hence "Some (AS_branch dc x' s') = lookup_branch dc cssa" using lookup_branch.simps(2) check_branch_list_consI(10) cseq by auto moreover have "(dc, const) \ set dclist" using check_branch_list_consI False by simp ultimately show ?thesis using check_branch_list_consI by auto qed next case (check_branch_list_finalI \ \ \ \ \ tid cons const va cs \) hence " cs = AS_branch cons x' s'" using lookup.simps check_branch_list_finalI lookup_branch.simps option.inject by (metis map_of.simps(1) map_of_Cons_code(2) option.distinct(1) s_branch_s_branch_list.exhaust(2) weak_map_of_SomeI) then show ?case using check_branch_list_finalI by auto next case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons va s) text \Supporting facts here to make the main body of the proof concise\ have xf:"atom x \ \" proof - have "supp \ \ supp \" using wf_supp(4) check_branch_s_branchI atom_dom.simps toSet.simps dom.simps by fastforce thus ?thesis using fresh_def x_not_in_b_set by blast qed hence "\[x::=v]\<^sub>\\<^sub>v = \" using forget_subst_v subst_v_\_def by auto have "\[x::=v]\<^sub>\\<^sub>v = \" using forget_subst_dv wfD_x_fresh fresh_GNil check_branch_s_branchI by metis have "supp v = {}" using check_branch_s_branchI check_v_wf wfV_supp_nil by metis hence "supp va = {}" using \ va = V_cons tid cons v\ v.supp pure_supp by auto let ?G = "(x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\ \" obtain z::x where z: "const = \ z : b_of const | c_of const z \ \ atom z \ (x', v,x,const,va)" using obtain_fresh_z_c_of by metis have vt: \\ ; \ ; GNil \ v \ \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \\ proof(rule infer_v_conj) obtain t' where t: "\ ; \ ; GNil \ v \ t' \ \ ; \ ; GNil \ t' \ const" using check_v_elims check_branch_s_branchI by metis show "\ ; \ ; GNil \ v \ \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \" proof(rule check_v_top) show "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \ " proof(rule wfG_wfT) show \ \ ; \ \\<^sub>w\<^sub>f (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil \ proof - have 1: "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis moreover have 2: "\ ; \ \\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\ GNil " using check_branch_s_branchI(17)[THEN check_s_wf] \\ = GNil\ by auto moreover hence "\ ; \ \\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil" using wfG_conj by metis ultimately show ?thesis unfolding subst_cv.simps subst_cev.simps subst_vv.simps by auto qed show \atom x \ ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )\ unfolding c.fresh ce.fresh v.fresh apply(intro conjI ) using check_branch_s_branchI fresh_at_base z freshers apply simp using check_branch_s_branchI fresh_at_base z freshers apply simp using pure_supp apply force using z fresh_x_neq fresh_prod5 by metis qed show \[ va ]\<^sup>c\<^sup>e = [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e[z::=v]\<^sub>c\<^sub>e\<^sub>v\ using \ va = V_cons tid cons v\ subst_cev.simps subst_vv.simps by auto show \ \ ; \ ; GNil \ v \ const \ using check_branch_s_branchI by auto show "supp [ va ]\<^sup>c\<^sup>e \ supp \" using \supp va = {}\ ce.supp by simp qed show "\ ; \ ; GNil \ v \ \ z : b_of const | c_of const z \" using check_branch_s_branchI z by auto qed text \Main body of proof for this case\ have "\ ; \ ; \ ; (?G)[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" proof(rule subst_check_check_s) show \\ ; \ ; GNil \ v \ \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \\ using vt by auto show \atom z \ (x, v)\ using z fresh_prodN by auto show \ \ ; \ ; \ ; ?G ; \ \ s \ \\ using check_branch_s_branchI by auto show \ ?G = GNil @ (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil\ proof - have "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis moreover have "(c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c_of const x" using c_of_switch[of z const x] z fresh_prodN by metis ultimately show ?thesis unfolding subst_cv.simps subst_cev.simps subst_vv.simps append_g.simps using c_of_switch[of z const x] z fresh_prodN z fresh_prodN check_branch_s_branchI by argo qed qed moreover have "s[x::=v]\<^sub>s\<^sub>v = s'[x'::=v]\<^sub>s\<^sub>v" using check_branch_s_branchI subst_v_flip_eq_two subst_v_s_def s_branch_s_branch_list.eq_iff by metis ultimately show ?case using check_branch_s_branchI \\[x::=v]\<^sub>\\<^sub>v = \\ \\[x::=v]\<^sub>\\<^sub>v = \\ by auto qed(auto+) text \Lemmas for while reduction. Making these separate lemmas allows flexibility in wiring them into the main proof and robustness if we change it\ lemma check_unit: fixes \::\ and \::\ and \::\ and G::\ assumes "\ ; {||} ; GNil \ \ z : B_unit | TRUE \ \ \'" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" and "\ \\<^sub>w\<^sub>f \" and "\ ; {||} \\<^sub>w\<^sub>f G" shows \ \ ; \ ; {||} ; G ; \ \ [[ L_unit ]\<^sup>v]\<^sup>s \ \'\ proof - have *:"\ ; {||} ; GNil \ [L_unit]\<^sup>v \ \ z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \" using infer_l.intros(4) infer_v_litI fresh_GNil assms wfX_wfY by (meson subtype_g_wf) moreover have "\ ; {||} ; GNil \ \ z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \ \ \'" using subtype_top subtype_trans * infer_v_wf by (meson assms(1)) ultimately show ?thesis (*apply(rule check_valI, auto simp add: assms,rule * )*) using subtype_top subtype_trans fresh_GNil assms check_valI assms check_s_g_weakening assms toSet.simps by (metis bot.extremum infer_v_g_weakening subtype_weakening wfD_wf) qed lemma preservation_var: shows "\; \; {||}; GNil; \ \ VAR u : \' = v IN s \ \ \ \ \ \ \ \ \ atom u \ \ \ atom u \ \ \ \; \; {||}; GNil; (u,\')#\<^sub>\\ \ s \ \ \ \ \ (u,v)#\ \ (u,\')#\<^sub>\\" and "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" proof(nominal_induct "{||}::bv fset" GNil \ "VAR u : \' = v IN s" \ and \ and \ rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_varI u' \ \ \ \ s') hence "\; \; {||}; GNil; (u, \') #\<^sub>\ \ \ s \ \" using check_s_abs_u check_v_wf by metis moreover have "\ \ ((u,v)#\) \ ((u,\')#\<^sub>\\)" proof show \\ \ \ \ \ \ using check_varI by auto show \\ ; {||} ; GNil \ v \ \'\ using check_varI by auto show \u \ fst ` set \\ using check_varI fresh_d_fst_d by auto qed ultimately show ?case by simp qed(auto+) lemma check_while: shows "\; \; {||}; GNil; \ \ WHILE s1 DO { s2 } \ \ \ atom x \ (s1, s2) \ atom z' \ x \ \; \; {||}; GNil; \ \ LET x : (\ z' : B_bool | TRUE \) = s1 IN (IF (V_var x) THEN (s2 ;; (WHILE s1 DO {s2})) ELSE ([ V_lit L_unit]\<^sup>s)) \ \" and "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" proof(nominal_induct "{||}::bv fset" GNil \ "AS_while s1 s2" \ and \ and \ avoiding: s1 s2 x z' rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_whileI \ \ \ s1 z s2 \') have teq:"\ z' : B_bool | TRUE \ = \ z : B_bool | TRUE \" using \.eq_iff by auto show ?case proof have " atom x \ \' " using wfT_nil_supp fresh_def subtype_wfT check_whileI(5) by fast moreover have "atom x \ \ z' : B_bool | TRUE \ " using \.fresh c.fresh b.fresh by metis ultimately show \atom x \ (\, \, {||}, GNil, \, \ z' : B_bool | TRUE \, s1, \')\ apply(unfold fresh_prodN) using check_whileI wb_x_fresh check_s_wf wfD_x_fresh fresh_empty_fset fresh_GNil fresh_Pair \atom x \ \'\ by metis show \ \ ; \ ; {||} ; GNil ; \ \ s1 \ \ z' : B_bool | TRUE \\ using check_whileI teq by metis let ?G = "(x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil" have cof:"(c_of \ z' : B_bool | TRUE \ x) = C_true" using c_of.simps check_whileI subst_cv.simps by metis have wfg: "\ ; {||} \\<^sub>w\<^sub>f ?G" proof show "c_of \ z' : B_bool | TRUE \ x \ {TRUE, FALSE}" using subst_cv.simps cof by auto show "\ ; {||} \\<^sub>w\<^sub>f GNil " using wfG_nilI check_whileI wfX_wfY check_s_wf by metis show "atom x \ GNil" using fresh_GNil by auto show "\ ; {||} \\<^sub>w\<^sub>f b_of \ z' : B_bool | TRUE \ " using wfB_boolI wfX_wfY check_s_wf b_of.simps by (metis \\ ; {||} \\<^sub>w\<^sub>f GNil\) qed obtain zz::x where zf:\atom zz \ ((\, \, {||}::bv fset, ?G , \, [ x ]\<^sup>v, AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \'),x,?G)\ using obtain_fresh by blast show \ \ ; \ ; {||} ; ?G ; \ \ AS_if [ x ]\<^sup>v (AS_seq s2 (AS_while s1 s2)) (AS_val [ L_unit ]\<^sup>v) \ \'\ proof show "atom zz \ (\, \, {||}::bv fset, ?G , \, [ x ]\<^sup>v, AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \')" using zf by auto show \\ ; {||} ; ?G \ [ x ]\<^sup>v \ \ zz : B_bool | TRUE \\ proof have "atom zz \ x \ atom zz \ (\, {||}::bv fset, ?G)" using zf fresh_prodN by metis thus \ \ ; {||} ; ?G \ [ x ]\<^sup>v \\ zz : B_bool | [[zz]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \\ using infer_v_varI lookup.simps wfg b_of.simps by metis thus \\ ; {||} ; ?G \ \ zz : B_bool | [[ zz ]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \ \ \ zz : B_bool | TRUE \\ using subtype_top infer_v_wf by metis qed show \ \ ; \ ; {||} ; ?G ; \ \ AS_seq s2 (AS_while s1 s2) \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ proof have "\ zz : B_unit | TRUE \ = \ z : B_unit | TRUE \" using \.eq_iff by auto thus \ \ ; \ ; {||} ; ?G ; \ \ s2 \ \ zz : B_unit | TRUE \\ using check_s_g_weakening(1) [OF check_whileI(3) _ wfg] toSet.simps by (simp add: \\ zz : B_unit | TRUE \ = \ z : B_unit | TRUE \\) show \ \ ; \ ; {||} ; ?G ; \ \ AS_while s1 s2 \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ proof(rule check_s_supertype(1)) have \ \; \; {||}; GNil; \ \ AS_while s1 s2 \ \'\ using check_whileI by auto thus *:\ \ ; \ ; {||} ; ?G ; \ \ AS_while s1 s2 \ \' \ using check_s_g_weakening(1)[OF _ _ wfg] toSet.simps by auto show \\ ; {||} ; ?G \ \' \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ proof(rule subtype_eq_if_\) show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \' \ using * check_s_wf by auto show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \ \ apply(rule wfT_eq_imp, simp add: base_for_lit.simps) using subtype_wf check_whileI wfg zf fresh_prodN by metis+ show \atom zz \ \'\ using zf fresh_prodN by metis qed qed qed show \ \ ; \ ; {||} ; ?G ; \ \ AS_val [ L_unit ]\<^sup>v \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ proof(rule check_s_supertype(1)) show *:\ \ ; \ ; {||} ; ?G ; \ \ AS_val [ L_unit ]\<^sup>v \ \'\ using check_unit[OF check_whileI(5) _ _ wfg] using check_whileI wfg wfX_wfY check_s_wf by metis show \\ ; {||} ; ?G \ \' \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ proof(rule subtype_eq_if_\) show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \' \ using * check_s_wf by metis show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \ \ apply(rule wfT_eq_imp, simp add: base_for_lit.simps) using subtype_wf check_whileI wfg zf fresh_prodN by metis+ show \atom zz \ \'\ using zf fresh_prodN by metis qed qed qed qed qed(auto+) lemma check_s_narrow: fixes s::s and x::x assumes "atom x \ (\, \, \, \, \, c, \, s)" and "\ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \ s \ \" and "\ ; \ ; \ \ c " shows "\ ; \ ; \ ; \ ; \ \ s \ \" proof - let ?B = "({||}::bv fset)" let ?v = "V_lit L_true" obtain z::x where z: "atom z \ (x, [ L_true ]\<^sup>v,c)" using obtain_fresh by metis have "atom z \ c" using z fresh_prodN by auto hence c:" c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c" using subst_v_c_def by simp have "\ ; \ ; \ ; ((x,B_bool, c) #\<^sub>\ \)[x::=?v]\<^sub>\\<^sub>v ; \[x::=?v]\<^sub>\\<^sub>v \ s[x::=?v]\<^sub>s\<^sub>v \ \[x::=?v]\<^sub>\\<^sub>v" proof(rule subst_infer_check_s(1) ) show vt: \ \ ; \ ; \ \ [ L_true ]\<^sup>v \ \ z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \ \ using infer_v_litI check_s_wf wfG_elims(2) infer_trueI assms by metis show \\ ; \ ; \ \ \ z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \ \ \ z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \\ proof show \atom x \ (\, \, \, z, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e , z, c[x::=[ z ]\<^sup>v]\<^sub>v)\ apply(unfold fresh_prodN, intro conjI) prefer 5 using c.fresh ce.fresh v.fresh z fresh_prodN apply auto[1] prefer 6 using fresh_subst_v_if[of "atom x" c x] assms fresh_prodN apply simp using z assms fresh_prodN apply metis using z assms fresh_prodN apply metis using z assms fresh_prodN apply metis using z fresh_prodN assms fresh_at_base by metis+ show \ \ ; \ ; \ \\<^sub>w\<^sub>f \ z : B_bool | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ \ using vt infer_v_wf by simp show \ \ ; \ ; \ \\<^sub>w\<^sub>f \ z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \ \ proof(rule wfG_wfT) show \ \ ; \ \\<^sub>w\<^sub>f (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ using c check_s_wf assms by metis have " atom x \ [ z ]\<^sup>v" using v.fresh z fresh_at_base by auto thus \atom x \ c[x::=[ z ]\<^sup>v]\<^sub>v\ using fresh_subst_v_if[of "atom x" c ] by auto qed have wfg: "\ ; \ \\<^sub>w\<^sub>f(x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \" using wfT_wfG vt infer_v_wf fresh_prodN assms by simp show \\ ; \ ; (x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>v \ using c valid_weakening[OF assms(3) _ wfg ] toSet.simps using subst_v_c_def by auto qed show \atom z \ (x, [ L_true ]\<^sup>v)\ using z fresh_prodN by metis show \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \ s \ \\ using assms by auto thus \(x, B_bool, c) #\<^sub>\ \ = GNil @ (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\ using append_g.simps c by auto qed moreover have "((x,B_bool, c) #\<^sub>\ \)[x::=?v]\<^sub>\\<^sub>v = \ " using subst_gv.simps by auto ultimately show ?thesis using assms forget_subst_dv forget_subst_sv forget_subst_tv fresh_prodN by metis qed lemma check_assert_s: fixes s::s and x::x assumes "\; \; {||}; GNil; \ \ AS_assert c s \ \" shows "\; \; {||}; GNil; \ \ s \ \ \ \ ; {||} ; GNil \ c " proof - let ?B = "({||}::bv fset)" let ?v = "V_lit L_true" obtain x where x: "\ ; \ ; ?B ; (x,B_bool, c) #\<^sub>\ GNil ; \ \ s \ \ \ atom x \ (\, \, ?B, GNil, \, c, \, s ) \ \ ; ?B ; GNil \ c " using check_s_elims(10)[OF \\ ; \ ; ?B ; GNil ; \ \ AS_assert c s \ \\] valid.simps by metis show ?thesis using assms check_s_narrow x by metis qed lemma infer_v_pair2I: "atom z \ (v1, v2) \ atom z \ (\, \, \) \ \ ; \ ; \ \ v1 \ t1 \ \ ; \ ; \ \ v2 \ t2 \ b1 = b_of t1 \ b2 = b_of t2 \ \ ; \ ; \ \ [ v1 , v2 ]\<^sup>v \ \ z : [ b1 , b2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \" using infer_v_pairI by simp subsection \Main Lemma\ lemma preservation: assumes "\ \ \\, s\ \ \\', s'\" and "\; \; \ \ \\, s\ \ \" shows "\\'. \; \; \' \ \\', s'\ \ \ \ \ \ \'" using assms proof(induct arbitrary: \ rule: reduce_stmt.induct) case (reduce_let_plusI \ x n1 n2 s') then show ?case using preservation_plus by (metis order_refl) next case (reduce_let_leqI b n1 n2 \ x s) then show ?case using preservation_leq by (metis order_refl) next case (reduce_let_eqI b n1 n2 \ \ x s) then show ?case using preservation_eq[OF reduce_let_eqI(2)] order_refl by metis next case (reduce_let_appI f z b c \' s' \ \ x v s) hence tt: "\; \; {||}; GNil; \ \ AS_let x (AE_app f v) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims[OF reduce_let_appI(2)] by metis hence *:"\; \; {||}; GNil; \ \ AS_let x (AE_app f v) s \ \" by auto hence "\; \; {||}; GNil; \ \ AS_let2 x (\'[z::=v]\<^sub>\\<^sub>v) (s'[z::=v]\<^sub>s\<^sub>v) s \ \" using preservation_app reduce_let_appI tt by auto hence "\; \; \ \ \ \ , AS_let2 x (\'[z::=v]\<^sub>\\<^sub>v) s'[z::=v]\<^sub>s\<^sub>v s \ \ \" using config_typeI tt by auto then show ?case using tt order_refl reduce_let_appI by metis next case (reduce_let_appPI f bv z b c \' s' \ \ x b' v s) hence tt: "\; \; {||}; GNil; \ \ AS_let x (AE_appP f b' v) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims[OF reduce_let_appPI(2)] by metis hence *:"\; \; {||}; GNil; \ \ AS_let x (AE_appP f b' v) s \ \" by auto have "\; \; {||}; GNil; \ \ AS_let2 x (\'[bv::=b']\<^sub>\\<^sub>b[z::=v]\<^sub>\\<^sub>v) (s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v) s \ \" proof(rule preservation_poly_app) show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \' s'))) = lookup_fun \ f\ using reduce_let_appPI by metis show \\fd\set \. check_fundef \ \ fd\ using tt lookup_fun_member by metis show \ \ ; \ ;{||} ; GNil ; \ \ AS_let x (AE_appP f b' v) s \ \\ using * by auto show \ \ ; {||} \\<^sub>w\<^sub>f b' \ using check_s_elims infer_e_wf wfE_elims * by metis qed(auto+) hence "\; \; \ \ \ \ , AS_let2 x (\'[bv::=b']\<^sub>\\<^sub>b[z::=v]\<^sub>\\<^sub>v) s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v s \ \ \" using config_typeI tt by auto then show ?case using tt order_refl reduce_let_appI by metis next case (reduce_if_trueI \ s1 s2) then show ?case using preservation_if by metis next case (reduce_if_falseI uw \ s1 s2) then show ?case using preservation_if by metis next case (reduce_let_valI \ x v s) then show ?case using preservation_let_val by presburger next case (reduce_let_mvar u v \ \ x s) hence *:"\; \; {||}; GNil; \ \ AS_let x (AE_mvar u) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by blast hence **: "\; \; {||}; GNil; \ \ AS_let x (AE_mvar u) s \ \" by auto obtain xa::x and za::x and ca::c and ba::b and sa::s where sa1: "atom xa \ (\, \, {||}::bv fset, GNil, \, AE_mvar u, \) \ atom za \ (xa, \, \, {||}::bv fset, GNil, \, AE_mvar u, \, sa) \ \; \; {||}; GNil; \ \ AE_mvar u \ \ za : ba | ca \ \ \ ; \ ; {||} ; (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \ \ (\c. atom c \ (s, sa) \ atom c \ (x, xa, s, sa) \ (x \ c) \ s = (xa \ c) \ sa)" using check_s_elims(2)[OF **] subst_defs by metis have "\ ; {||} ; GNil \ v \ \ za : ba | ca \" proof - have " (u , \ za : ba | ca \) \ setD \" using infer_e_elims(11) sa1 by fast thus ?thesis using delta_sim_v reduce_let_mvar config_type_elims check_s_wf by metis qed then obtain \' where vst: "\ ; {||} ; GNil \ v \ \' \ \ ; {||} ; GNil \ \' \ \ za : ba | ca \" using check_v_elims by blast obtain za2 and ba2 and ca2 where zbc: "\' = (\ za2 : ba2 | ca2 \) \ atom za2 \ (xa, (xa, \, \, {||}::bv fset, GNil, \, AE_val v, \, sa))" using obtain_fresh_z by blast have beq: "ba=ba2" using subtype_eq_base vst zbc by blast moreover have xaf: "atom xa \ (za, za2)" apply(unfold fresh_prodN, intro conjI) using sa1 zbc fresh_prodN fresh_x_neq by metis+ have sat2: " \ ; \ ; {||} ; GNil@(xa, ba, ca2[za2::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \" proof(rule ctx_subtype_s) show "\ ; \ ; {||} ; GNil @ (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \" using sa1 by auto show "\ ; {||} ; GNil \ \ za2 : ba | ca2 \ \ \ za : ba | ca \" using beq zbc vst by fast show "atom xa \ (za, za2, ca, ca2)" proof - have *:"\ ; {||} ; GNil \\<^sub>w\<^sub>f (\ za2 : ba2 | ca2 \) " using zbc vst subtype_wf by auto hence "supp ca2 \ { atom za2 }" using wfT_supp_c[OF *] supp_GNil by simp moreover have "atom za2 \ xa" using zbc fresh_Pair fresh_x_neq by metis ultimately have "atom xa \ ca2" using zbc supp_at_base fresh_def by (metis empty_iff singleton_iff subset_singletonD) moreover have "atom xa \ ca" proof - have *:"\ ; {||} ; GNil \\<^sub>w\<^sub>f (\ za : ba | ca \) " using zbc vst subtype_wf by auto hence "supp ca \ { atom za }" using wfT_supp \.supp by force moreover have "xa \ za" using fresh_def fresh_x_neq xaf fresh_Pair by metis ultimately show ?thesis using fresh_def by auto qed ultimately show ?thesis using xaf sa1 fresh_prod4 fresh_Pair by metis qed qed hence dwf: "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" using sa1 infer_e_wf by meson have "\; \; {||}; GNil; \ \ AS_let xa (AE_val v) sa \ \" proof have "atom xa \ (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce thus "atom xa \ (\, \, {||}::bv fset, GNil, \, AE_val v, \)" using sa1 freshers by simp have "atom za2 \ (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce thus "atom za2 \ (xa, \, \, {||}::bv fset, GNil, \, AE_val v, \, sa)" using zbc freshers fresh_prodN by auto have " \ \\<^sub>w\<^sub>f \" using sa1 infer_e_wf by auto thus "\; \; {||}; GNil; \ \ AE_val v \ \ za2 : ba | ca2 \" using zbc vst beq dwf infer_e_valI by blast show "\ ; \ ; {||} ; (xa, ba, ca2[za2::=V_var xa]\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \" using sat2 append_g.simps subst_defs by metis qed moreover have "AS_let xa (AE_val v) sa = AS_let x (AE_val v) s" proof - have "[[atom x]]lst. s = [[atom xa]]lst. sa" using sa1 Abs1_eq_iff_all(3)[where z=" (s, sa)"] by metis thus ?thesis using s_branch_s_branch_list.eq_iff(2) by metis qed ultimately have "\; \; {||}; GNil; \ \ AS_let x (AE_val v) s \ \" by auto then show ?case using reduce_let_mvar * config_typeI by (meson order_refl) next case (reduce_let2I \ \ s1 \' s1' x t s2) hence **: "\; \; {||}; GNil; \ \ AS_let2 x t s1 s2 \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims[OF reduce_let2I(3)] by blast hence *:"\; \; {||}; GNil; \ \ AS_let2 x t s1 s2 \ \" by auto obtain xa::x and z::x and c and b and s2a::s where st: "atom xa \ (\, \, {||}::bv fset, GNil, \, t, s1, \) \ \; \; {||}; GNil; \ \ s1 \ t \ \ ; \ ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\ GNil ; \ \ s2a \ \ \ ([[atom x]]lst. s2 = [[atom xa]]lst. s2a) " using check_s_elims(4)[OF *] Abs1_eq_iff_all(3) by metis hence "\; \; \ \ \ \ , s1 \ \ t" using config_typeI ** by auto then obtain \' where s1r: "\; \; \' \ \ \' , s1' \ \ t \ \ \ \'" using reduce_let2I by presburger have "\; \; {||}; GNil; \' \ AS_let2 xa t s1' s2a \ \" proof(rule check_let2I) show *:"\; \; {||}; GNil; \' \ s1' \ t" using config_type_elims st s1r by metis show "atom xa \ (\, \, {||}::bv fset, GNil, \',t, s1', \)" proof - have "atom xa \ s1'" using check_s_x_fresh * by auto moreover have "atom xa \ \'" using check_s_x_fresh * by auto ultimately show ?thesis using st fresh_prodN by metis qed show "\ ; \ ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\ GNil ; \' \ s2a \ \" proof - have "\ ; {||} ; GNil \\<^sub>w\<^sub>f \'" using * check_s_wf by auto moreover have "\ ; {||} \\<^sub>w\<^sub>f ((xa, b_of t, c_of t xa) #\<^sub>\ GNil)" using st check_s_wf by auto ultimately have "\ ; {||} ; ((xa, b_of t , c_of t xa) #\<^sub>\ GNil) \\<^sub>w\<^sub>f \'" using wf_weakening by auto thus ?thesis using check_s_d_weakening check_s_wf st s1r by metis qed qed moreover have "AS_let2 xa t s1' s2a = AS_let2 x t s1' s2" using st s_branch_s_branch_list.eq_iff by metis ultimately have "\; \; {||}; GNil; \' \ AS_let2 x t s1' s2 \ \" using st by argo moreover have "\ \ \' \ \'" using config_type_elims s1r by fast ultimately show ?case using config_typeI ** by (meson s1r) next case (reduce_let2_valI vb \ x t v s) then show ?case using preservation_let_val by meson next case (reduce_varI u \ \ \' v s) hence ** : "\; \; {||}; GNil; \ \ AS_var u \' v s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by meson have uf: "atom u \ \" using reduce_varI delta_sim_fresh by force hence *: "\; \; {||}; GNil; \ \ AS_var u \' v s \ \" and " \ \ \ \ \" using ** by auto thus ?case using preservation_var reduce_varI config_typeI ** set_subset_Cons setD_ConsD subsetI by (metis delta_sim_fresh) next case (reduce_assignI \ \ u v ) hence *: "\; \; {||}; GNil; \ \ AS_assign u v \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by meson then obtain z and \' where zt: "\ ; {||} ; GNil \ (\ z : B_unit | TRUE \) \ \ \ (u,\') \ setD \ \ \ ; {||} ; GNil \ v \ \' \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \" using check_s_elims(8) by metis hence "\ \ update_d \ u v \ \" using update_d_sim * by metis moreover have "\; \; {||}; GNil; \ \ AS_val (V_lit L_unit ) \ \" using zt * check_s_v_unit check_s_wf by auto ultimately show ?case using config_typeI * by (meson order_refl) next case (reduce_seq1I \ \ s) hence "\ ; \ ; {||} ; GNil ; \ \ s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using check_s_elims config_type_elims by force then show ?case using config_typeI by blast next case (reduce_seq2I s1 v \ \ \' s1' s2) hence tt: "\; \; {||}; GNil; \ \ AS_seq s1 s2 \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by blast then obtain z where zz: "\ ; \ ; {||} ; GNil; \ \ s1 \ (\ z : B_unit | TRUE \) \ \; \; {||}; GNil; \ \ s2 \ \" using check_s_elims by blast hence "\; \; \ \ \ \ , s1 \ \ (\ z : B_unit | TRUE \)" using tt config_typeI tt by simp then obtain \' where *: "\; \; \' \ \ \' , s1' \ \ (\ z : B_unit | TRUE \) \ \ \ \'" using reduce_seq2I by meson moreover hence s't: "\; \; {||}; GNil; \' \ s1' \ (\ z : B_unit | TRUE \) \ \ \ \' \ \'" using config_type_elims by force moreover hence "\ ; {||} ; GNil \\<^sub>w\<^sub>f \'" using check_s_wf by meson moreover hence "\; \; {||}; GNil; \' \ s2 \ \" using calculation(1) zz check_s_d_weakening * by metis moreover hence "\; \; {||}; GNil; \' \ (AS_seq s1' s2) \ \" using check_seqI zz s't by meson ultimately have "\; \; \' \ \ \' , AS_seq s1' s2 \ \ \ \ \ \ \'" using zz config_typeI tt by meson then show ?case by meson next case (reduce_whileI x s1 s2 z' \ \ ) hence *: "\; \; {||}; GNil; \ \ AS_while s1 s2 \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by meson hence **:"\; \; {||}; GNil; \ \ AS_while s1 s2 \ \" by auto hence "\; \; {||}; GNil; \ \ AS_let2 x (\ z' : B_bool | TRUE \) s1 (AS_if (V_var x) (AS_seq s2 (AS_while s1 s2)) (AS_val (V_lit L_unit)) ) \ \" using check_while reduce_whileI by auto thus ?case using config_typeI * by (meson subset_refl) next case (reduce_caseI dc x' s' css \ \ tyid v) hence **: "\; \; {||}; GNil; \ \ AS_match (V_cons tyid dc v) css \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims[OF reduce_caseI(2)] by metis hence ***: "\; \; {||}; GNil; \ \ AS_match (V_cons tyid dc v) css \ \" by auto let ?vcons = "V_cons tyid dc v" obtain dclist tid and z::x where cv: "\; {||} ; GNil \ (V_cons tyid dc v) \ (\ z : B_id tid | TRUE \) \ \; \; {||}; GNil; \ ; tid ; dclist ; (V_cons tyid dc v) \ css \ \ \ AF_typedef tid dclist \ set \ \ \ ; {||} ; GNil \ V_cons tyid dc v \ \ z : B_id tid | TRUE \" using check_s_elims(9)[OF ***] by metis hence vi:" \ ; {||} ; GNil \ V_cons tyid dc v \ \ z : B_id tid | TRUE \" by auto obtain tcons where vi2:" \ ; {||} ; GNil \ V_cons tyid dc v \ tcons \ \ ; {||} ; GNil \ tcons \ \ z : B_id tid | TRUE \" using check_v_elims(1)[OF vi] by metis hence vi1: "\ ; {||} ; GNil \ V_cons tyid dc v \ tcons" by auto show ?case proof(rule infer_v_elims(4)[OF vi1],goal_cases) case (1 dclist2 tc tv z2) have "tyid = tid" using \.eq_iff using subtype_eq_base vi2 1 by fastforce hence deq:"dclist = dclist2" using check_v_wf wfX_wfY cv 1 wfTh_dclist_unique by metis have "\; \; {||}; GNil; \ \ s'[x'::=v]\<^sub>s\<^sub>v \ \" proof(rule check_match(3)) show \ \ ; \ ; {||} ; GNil ; \ ; tyid ; dclist ; ?vcons \ css \ \\ using \tyid = tid\ cv by auto show "distinct (map fst dclist)" using wfTh_dclist_distinct check_v_wf wfX_wfY cv by metis show \?vcons = V_cons tyid dc v\ by auto show \{||} = {||}\ by auto show \(dc, tc) \ set dclist\ using 1 deq by auto show \GNil = GNil\ by auto show \Some (AS_branch dc x' s') = lookup_branch dc css\ using reduce_caseI by auto show \\ ; {||} ; GNil \ v \ tc\ using 1 check_v.intros by auto qed thus ?case using config_typeI ** by blast qed next case (reduce_let_fstI \ \ x v1 v2 s) thus ?case using preservation_fst_snd order_refl by metis next case (reduce_let_sndI \ \ x v1 v2 s) thus ?case using preservation_fst_snd order_refl by metis next case (reduce_let_concatI \ \ x v1 v2 s) hence elim: "\; \; {||}; GNil; \ \ AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by metis obtain z::x where z: "atom z \ (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))))" using obtain_fresh by metis have "\ ; {||} \\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto have "\; \; {||}; GNil; \ \ AS_let x (AE_val (V_lit (L_bitvec (v1 @ v2)))) s \ \" proof(rule subtype_let) show \ \; \; {||}; GNil; \ \ AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \ \\ using elim by auto show \\; \; {||}; GNil; \ \ (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \ \ z : B_bitvec | CE_val (V_var z) == (CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e))\ \ (is "\; \; {||}; GNil; \ \ ?e1 \ ?t1") proof show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto show \ \ \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto show \ \ ; {||} ; GNil \ V_lit (L_bitvec v1) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v1)) \\ using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil by auto show \ \ ; {||} ; GNil \ V_lit (L_bitvec v2) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v2)) \\ using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil by auto show \atom z \ AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))\ using z fresh_Pair by metis show \atom z \ GNil\ using z fresh_Pair by auto qed show \\; \; {||}; GNil; \ \ AE_val (V_lit (L_bitvec (v1 @ v2))) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) \ \ (is "\; \; {||}; GNil; \ \ ?e2 \ ?t2") using infer_e_valI infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil check_s_wf elim by metis show \\ ; {||} ; GNil \ ?t2 \ ?t1\ using subtype_concat check_s_wf elim by auto qed thus ?case using config_typeI elim by (meson order_refl) next case (reduce_let_lenI \ \ x v s) hence elim: "\; \; {||}; GNil; \ \ AS_let x (AE_len (V_lit (L_bitvec v))) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using check_s_elims config_type_elims by metis then obtain t where t: "\; \; {||}; GNil; \ \ AE_len (V_lit (L_bitvec v)) \ t" using check_s_elims by meson moreover then obtain z::x where "t = \ z : B_int | CE_val (V_var z) == CE_len ([V_lit (L_bitvec v)]\<^sup>c\<^sup>e) \" using infer_e_elims by meson moreover obtain z'::x where "atom z' \ v" using obtain_fresh by metis moreover have "\; \; {||}; GNil; \ \ AE_val (V_lit (L_num (int (length v)))) \ \ z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \" using infer_e_valI infer_v_litI infer_l.intros(3) t check_s_wf elim by (metis infer_l_form2 type_for_lit.simps(3)) moreover have "\ ; {||} ; GNil \ \ z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \ \ \ z : B_int | CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e \" using subtype_len check_s_wf elim by auto ultimately have "\; \; {||}; GNil; \ \ AS_let x (AE_val (V_lit (L_num (int (length v))))) s \ \" using subtype_let by (meson elim) thus ?case using config_typeI elim by (meson order_refl) next case (reduce_let_splitI n v v1 v2 \ \ x s) hence elim: "\; \; {||}; GNil; \ \ AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by metis obtain z::x where z: "atom z \ (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))), ([ L_bitvec v1 ]\<^sup>v, [ L_bitvec v2 ]\<^sup>v), \, {||}::bv fset)" using obtain_fresh by metis have *:"\ ; {||} \\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto have "\; \; {||}; GNil; \ \ AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s \ \" proof(rule subtype_let) show \ \; \; {||}; GNil; \ \ AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \ \\ using elim by auto show \\; \; {||}; GNil; \ \ (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) \ \ z : B_pair B_bitvec B_bitvec | ((CE_val (V_lit (L_bitvec v))) == (CE_concat (CE_fst (CE_val (V_var z))) (CE_snd (CE_val (V_var z))))) AND (((CE_len (CE_fst (CE_val (V_var z))))) == (CE_val (V_lit (L_num n)))) \ \ (is "\; \; {||}; GNil; \ \ ?e1 \ ?t1") proof show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto show \ \ \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto show \ \ ; {||} ; GNil \ V_lit (L_bitvec v) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v)) \\ using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil by auto show "\ ; {||} ; GNil \ ([ L_num n ]\<^sup>v) \ \ z : B_int | (([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) == ([ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)) AND [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \" using split_n reduce_let_splitI check_v_num_leq * wfX_wfY by metis show \atom z \ AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\ using z fresh_Pair by auto show \atom z \ GNil\ using z fresh_Pair by auto show \atom z \ AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\ using z fresh_Pair by auto show \atom z \ GNil\ using z fresh_Pair by auto show \atom z \ AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\ using z fresh_Pair by auto show \atom z \ GNil\ using z fresh_Pair by auto qed show \\; \; {||}; GNil; \ \ AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \ \ z : B_pair B_bitvec B_bitvec | CE_val (V_var z) == CE_val ((V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) \ \ (is "\; \; {||}; GNil; \ \ ?e2 \ ?t2") apply(rule infer_e_valI) using check_s_wf elim apply metis using check_s_wf elim apply metis apply(rule infer_v_pair2I) using z fresh_prodN apply metis using z fresh_GNil fresh_prodN apply metis using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ b_of.simps apply blast+ using b_of.simps apply simp+ done show \\ ; {||} ; GNil \ ?t2 \ ?t1\ using subtype_split check_s_wf elim reduce_let_splitI by auto qed thus ?case using config_typeI elim by (meson order_refl) next case (reduce_assert1I \ \ c v) hence elim: "\; \; {||}; GNil; \ \ AS_assert c [v]\<^sup>s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims reduce_assert1I by metis hence *:"\; \; {||}; GNil; \ \ AS_assert c [v]\<^sup>s \ \" by auto have "\; \; {||}; GNil; \ \ [v]\<^sup>s \ \" using check_assert_s * by metis thus ?case using elim config_typeI by blast next case (reduce_assert2I \ \ s \' s' c) hence elim: "\; \; {||}; GNil; \ \ AS_assert c s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims by metis hence *:"\; \; {||}; GNil; \ \ AS_assert c s \ \" by auto have cv: "\; \; {||}; GNil; \ \ s \ \ \ \ ; {||} ; GNil \ c " using check_assert_s * by metis hence "\; \; \ \ \\, s\ \ \" using elim config_typeI by simp then obtain \' where D: "\; \; \' \ \ \' , s' \ \ \ \ \ \ \'" using reduce_assert2I by metis hence **:"\; \; {||}; GNil; \' \ s' \ \ \ \ \ \' \ \'" using config_type_elims by metis obtain x::x where x:"atom x \ (\, \, ({||}::bv fset), GNil, \', c, \, s')" using obtain_fresh by metis have *:"\; \; {||}; GNil; \' \ AS_assert c s' \ \" proof show "atom x \ (\, \, {||}, GNil, \', c, \, s')" using x by auto have "\ ; {||} ; GNil \\<^sub>w\<^sub>f c" using * check_s_wf by auto hence wfg:"\ ; {||} \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ GNil" using wfC_wfG wfB_boolI check_s_wf * fresh_GNil by auto moreover have cs: "\; \; {||}; GNil; \' \ s' \ \" using ** by auto ultimately show "\ ; \ ; {||} ; (x, B_bool, c) #\<^sub>\ GNil ; \' \ s' \ \" using check_s_g_weakening(1)[OF cs _ wfg] toSet.simps by simp show "\ ; {||} ; GNil \ c " using cv by auto show "\ ; {||} ; GNil \\<^sub>w\<^sub>f \' " using check_s_wf ** by auto qed thus ?case using elim config_typeI D ** by metis qed lemma preservation_many: assumes " \ \ \\, s\ \\<^sup>* \ \' , s' \" shows "\; \; \ \ \\, s\ \ \ \ \\'. \; \; \' \ \ \' , s' \ \ \ \ \ \ \'" using assms proof(induct arbitrary: \ rule: reduce_stmt_many.induct) case (reduce_stmt_many_oneI \ \ s \' s') then show ?case using preservation by simp next case (reduce_stmt_many_manyI \ \ s \' s' \'' s'') then show ?case using preservation subset_trans by metis qed section \Progress\ text \We prove that a well typed program is either a value or we can make a step\ lemma check_let_op_infer: assumes "\; \; {||}; \; \ \ LET x = (AE_op opp v1 v2) IN s \ \" and "supp ( LET x = (AE_op opp v1 v2) IN s) \ atom`fst`setD \" shows "\ z b c. \; \; {||}; \; \ \ (AE_op opp v1 v2) \ \z:b|c\" proof - have xx: "\; \; {||}; \; \ \ LET x = (AE_op opp v1 v2) IN s \ \" using assms by simp then show ?thesis using check_s_elims(2)[OF xx] by meson qed lemma infer_pair: assumes "\ ; B; \ \ v \ \ z : B_pair b1 b2 | c \" and "supp v = {}" obtains v1 and v2 where "v = V_pair v1 v2" using assms proof(nominal_induct v rule: v.strong_induct) case (V_lit x) then show ?case by auto next case (V_var x) then show ?case using v.supp supp_at_base by auto next case (V_pair x1a x2a) then show ?case by auto next case (V_cons x1a x2a x3) then show ?case by auto next case (V_consp x1a x2a x3 x4) then show ?case by auto qed lemma progress_fst: assumes "\; \; {||}; \; \ \ LET x = (AE_fst v) IN s \ \" and "\ \ \ \ \" and "supp (LET x = (AE_fst v) IN s) \ atom`fst`setD \" shows "\\' s'. \ \ \ \ , LET x = (AE_fst v) IN s \ \ \\', s'\" proof - have *:"supp v = {}" using assms s_branch_s_branch_list.supp by auto obtain z and b and c where "\; \; {||}; \; \ \ (AE_fst v ) \ \ z : b | c \" using check_s_elims(2) using assms by meson moreover obtain z' and b' and c' where "\ ; {||} ; \ \ v \ \ z' : B_pair b b' | c' \" using infer_e_elims(8) using calculation by auto moreover then obtain v1 and v2 where "V_pair v1 v2 = v" using * infer_pair by metis ultimately show ?thesis using reduce_let_fstI assms by metis qed lemma progress_let: assumes "\; \; {||}; \; \ \ LET x = e IN s \ \" and "\ \ \ \ \" and "supp (LET x = e IN s) \ atom ` fst ` setD \" and "sble \ \" shows "\\' s'. \ \ \ \ , LET x = e IN s\ \ \ \' , s'\" proof - obtain z b c where *: "\; \; {||}; \; \ \ e \ \ z : b | c \ " using check_s_elims(2)[OF assms(1)] by metis have **: "supp e \ atom ` fst ` setD \" using assms s_branch_s_branch_list.supp by auto from * ** assms show ?thesis proof(nominal_induct "\ z : b | c \" rule: infer_e.strong_induct) case (infer_e_valI \ \ \ \ \ v) then show ?case using reduce_stmt_elims reduce_let_valI by metis next case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) hence vf: "supp v1 = {} \ supp v2 = {}" by force then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \ v2 = (V_lit (L_num n2))" using infer_int infer_e_plusI by metis then show ?case using reduce_let_plusI * by metis next case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) hence vf: "supp v1 = {} \ supp v2 = {}" by force then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \ v2 = (V_lit (L_num n2))" using infer_int infer_e_leqI by metis then show ?case using reduce_let_leqI * by metis next case (infer_e_eqI \ \ \ \ \ v1 z1 bb c1 v2 z2 c2 z3) hence vf: "supp v1 = {} \ supp v2 = {}" by force then obtain n1 and n2 where *: "v1 = V_lit n1 \ v2 = (V_lit n2)" using infer_lit infer_e_eqI by metis then show ?case using reduce_let_eqI by blast next case (infer_e_appI \ \ \ \ \ f x b c \' s' v) then show ?case using reduce_let_appI by metis next case (infer_e_appPI \ \ \ \ \ b' f bv x b c \' s' v) then show ?case using reduce_let_appPI by metis next case (infer_e_fstI \ \ \ \ \ v z' b2 c z) hence "supp v = {}" by force then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_fstI infer_pair by metis then show ?case using reduce_let_fstI * by metis next case (infer_e_sndI \ \ \ \ \ v z' b1 c z) hence "supp v = {}" by force then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_sndI infer_pair by metis then show ?case using reduce_let_sndI * by metis next case (infer_e_lenI \ \ \ \ \ v z' c za) hence "supp v = {}" by force then obtain bvec where "v = V_lit (L_bitvec bvec)" using infer_e_lenI infer_bitvec by metis then show ?case using reduce_let_lenI * by metis next case (infer_e_mvarI \ \ \ \ \ u) hence "(u, \ z : b | c \) \ setD \" using infer_e_elims(10) by meson then obtain v where "(u,v) \ set \" using infer_e_mvarI delta_sim_delta_lookup by meson then show ?case using reduce_let_mvar by metis next case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) hence vf: "supp v1 = {} \ supp v2 = {}" by force then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \ v2 = (V_lit (L_bitvec n2))" using infer_bitvec infer_e_concatI by metis then show ?case using reduce_let_concatI * by metis next case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) hence vf: "supp v1 = {} \ supp v2 = {}" by force then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \ v2 = (V_lit (L_num n2))" using infer_bitvec infer_e_splitI check_int by metis have "0 \ n2 \ n2 \ int (length n1)" using check_v_range[OF _ * ] infer_e_splitI by simp then obtain bv1 and bv2 where "split n2 n1 (bv1 , bv2)" using obtain_split by metis then show ?case using reduce_let_splitI * by metis qed qed lemma check_css_lookup_branch_exist: fixes s::s and cs::branch_s and css::branch_list and v::v shows "\; \; B; G; \ \ s \ \ \ True" and "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and "\; \; \; \; \; tid; dclist; v \ css \ \ \ (dc, t) \ set dclist \ \x' s'. Some (AS_branch dc x' s') = lookup_branch dc css" proof(nominal_induct \ and \ and \ rule: check_s_check_branch_s_check_branch_list.strong_induct) case (check_branch_list_consI \ \ \ \ \ tid cons const v cs \ dclist css) then show ?case using lookup_branch.simps check_branch_list_finalI by force next case (check_branch_list_finalI \ \ \ \ \ tid cons const v cs \) then show ?case using lookup_branch.simps check_branch_list_finalI by force qed(auto+) lemma progress_aux: shows "\; \; \; \; \ \ s \ \ \ \ = {||} \ sble \ \ \ supp s \ atom ` fst ` setD \ \ \ \ \ \ \ \ (\v. s = [v]\<^sup>s) \ (\\' s'. \ \ \\, s\ \ \\', s'\)" and "\; \; {||}; \; \; tid; dc; const; v2 \ cs \ \ \ supp cs = {} \ True " "\; \; {||}; \; \; tid; dclist; v2 \ css \ \ \ supp css = {} \ True " proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) case (check_valI \ \ \ v \' \) then show ?case by auto next case (check_letI x \ \ \ \ \ e \ z s b c) hence "\; \; {||}; \; \ \ AS_let x e s \ \" using Typing.check_letI by meson then show ?case using progress_let check_letI by metis next case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) then show ?case by auto next case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) then show ?case by auto next case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) then show ?case by auto next case (check_ifI z \ \ \ \ \ v s1 s2 \) have "supp v = {}" using check_ifI s_branch_s_branch_list.supp by auto hence "v = V_lit L_true \ v = V_lit L_false" using check_bool_options check_ifI by auto then show ?case using reduce_if_falseI reduce_if_trueI check_ifI by meson next case (check_let2I x \ \ \ G \ t s1 \ s2 ) then consider " (\v. s1 = AS_val v)" | "(\\' a. \ \ \\, s1\ \ \\', a\)" by auto then show ?case proof(cases) case 1 then show ?thesis using reduce_let2_valI by fast next case 2 then show ?thesis using reduce_let2I check_let2I by meson qed next case (check_varI u \ \ \ \ \ \' v \ s) obtain uu::u where uf: "atom uu \ (u,\,s) " using obtain_fresh by blast obtain sa where " (uu \ u ) \ s = sa" by presburger moreover have "atom uu \ s" using uf fresh_prod3 by auto ultimately have "AS_var uu \' v sa = AS_var u \' v s" using s_branch_s_branch_list.eq_iff(7) Abs1_eq_iff(3)[of uu sa u s] by auto moreover have "atom uu \ \" using uf fresh_prod3 by auto ultimately have "\ \ \\, AS_var u \' v s\ \ \(uu, v) # \, sa\" using reduce_varI uf by metis then show ?case by auto next case (check_assignI \ u \ P G v z \') then show ?case using reduce_assignI by blast next case (check_whileI \ \ \ \ \ s1 z s2 \') obtain x::x where "atom x \ (s1,s2)" using obtain_fresh by metis moreover obtain z::x where "atom z \ x" using obtain_fresh by metis ultimately show ?case using reduce_whileI by fast next case (check_seqI P \ \ G \ s1 z s2 \) thus ?case proof(cases "\v. s1 = AS_val v") case True then obtain v where v: "s1 = AS_val v" by blast hence "supp v = {}" using check_seqI by auto have "\z1 c1. P; \; G \ v \ (\ z1 : B_unit | c1 \)" proof - obtain t where t:"P; \; G \ v \ t \ P; \ ; G \ t \ (\ z : B_unit | TRUE \)" using v check_seqI(1) check_s_elims(1) by blast obtain z1 and b1 and c1 where teq: "t = (\ z1 : b1 | c1 \)" using obtain_fresh_z by meson hence "b1 = B_unit" using subtype_eq_base t by meson thus ?thesis using t teq by fast qed then obtain z1 and c1 where "P ; \ ; G \ v \ (\ z1 : B_unit | c1 \)" by auto hence "v = V_lit L_unit" using infer_v_unit_form \supp v = {}\ by simp hence "s1 = AS_val (V_lit L_unit)" using v by auto then show ?thesis using check_seqI reduce_seq1I by meson next case False then show ?thesis using check_seqI reduce_seq2I by (metis Un_subset_iff s_branch_s_branch_list.supp(9)) qed next case (check_caseI \ \ \ \ \ tid dclist v cs \ z) hence "supp v = {}" by auto then obtain v' and dc and t::\ where v: "v = V_cons tid dc v' \ (dc, t) \ set dclist" using check_v_tid_form check_caseI by metis obtain z and b and c where teq: "t = (\ z : b | c \)" using obtain_fresh_z by meson moreover then obtain x' s' where "Some (AS_branch dc x' s') = lookup_branch dc cs" using v teq check_caseI check_css_lookup_branch_exist by metis ultimately show ?case using reduce_caseI v check_caseI dc_of.cases by metis next case (check_assertI x \ \ \ \ \ c \ s) hence sps: "supp s \ atom ` fst ` setD \" by auto have "atom x \ c " using check_assertI by auto have "atom x \ \" using check_assertI check_s_wf wfG_elims by metis have "sble \ ((x, B_bool, c) #\<^sub>\ \)" proof - obtain i' where i':" i' \ \ \ \; \ \ i'" using check_assertI sble_def by metis obtain i::valuation where i:"i = i' ( x \ SBool True)" by auto have "i \ (x, B_bool, c) #\<^sub>\ \" proof - have "i' \ c" using valid.simps i' check_assertI by metis hence "i \ c" using is_satis_weakening_x i \atom x \ c\ by auto moreover have "i \ \" using is_satis_g_weakening_x i' i check_assertI \atom x \ \\ by metis ultimately show ?thesis using is_satis_g.simps i by auto qed moreover have "\ ; ((x, B_bool, c) #\<^sub>\ \) \ i" proof(rule wfI_cons) show \ i' \ \ \ using i' by auto show \ \ ; \ \ i'\ using i' by auto show \i = i'(x \ SBool True)\ using i by auto show \ \ \ SBool True: B_bool\ using wfRCV_BBoolI by auto show \atom x \ \\ using check_assertI check_s_wf wfG_elims by auto qed ultimately show ?thesis using sble_def by auto qed then consider "(\v. s = [v]\<^sup>s)" | "(\\' a. \ \ \\, s\ \ \ \', a\)" using check_assertI sps by metis hence "(\\' a. \ \ \\, ASSERT c IN s\ \ \\', a\)" proof(cases) case 1 then show ?thesis using reduce_assert1I by metis next case 2 then show ?thesis using reduce_assert2I by metis qed thus ?case by auto qed lemma progress: assumes "\; \; \ \ \\, s\ \ \" shows "(\v. s = [v]\<^sup>s) \ (\\' s'. \ \ \\, s\ \ \\', s'\)" proof - have "\; \; {||}; GNil; \ \ s \ \" and "\ \ \ \ \" using config_type_elims[OF assms(1)] by auto+ moreover hence "supp s \ atom ` fst ` setD \" using check_s_wf wfS_supp by fastforce moreover have "sble \ GNil" using sble_def wfI_def is_satis_g.simps by simp ultimately show ?thesis using progress_aux by blast qed section \Safety\ lemma safety_stmt: assumes "\ \ \\, s\ \\<^sup>* \\', s'\" and "\; \; \ \ \\, s\ \ \" shows "(\v. s' = [v]\<^sup>s) \ (\\'' s''. \ \ \\', s'\ \ \\'', s''\)" using preservation_many progress assms by meson lemma safety: assumes "\ \PROG \ \ \ s\ \ \" and "\ \ \\_of \, s\ \\<^sup>* \\', s'\" shows "(\v. s' = [v]\<^sup>s) \ (\\'' s''. \ \ \\', s'\ \ \\'', s''\)" using assms config_type_prog_elims safety_stmt by metis end \ No newline at end of file diff --git a/thys/MiniSail/Syntax.thy b/thys/MiniSail/Syntax.thy --- a/thys/MiniSail/Syntax.thy +++ b/thys/MiniSail/Syntax.thy @@ -1,1174 +1,1174 @@ (*<*) theory Syntax imports "Nominal2.Nominal2" "Nominal-Utils" begin (*>*) chapter \Syntax\ text \Syntax of MiniSail programs and the contexts we use in judgements.\ section \Program Syntax\ subsection \AST Datatypes\ type_synonym num_nat = nat atom_decl x (* Immutable variables*) atom_decl u (* Mutable variables *) atom_decl bv (* Basic-type variables *) type_synonym f = string (* Function name *) type_synonym dc = string (* Data constructor *) type_synonym tyid = string text \Basic types. Types without refinement constraints\ nominal_datatype "b" = B_int | B_bool | B_id tyid | B_pair b b ("[ _ , _ ]\<^sup>b") | B_unit | B_bitvec | B_var bv | B_app tyid b nominal_datatype bit = BitOne | BitZero text \ Literals \ nominal_datatype "l" = L_num "int" | L_true | L_false | L_unit | L_bitvec "bit list" text \ Values. We include a type identifier, tyid, in the literal for constructors to make typing and well-formedness checking easier \ nominal_datatype "v" = V_lit "l" ( "[ _ ]\<^sup>v") | V_var "x" ( "[ _ ]\<^sup>v") | V_pair "v" "v" ( "[ _ , _ ]\<^sup>v") | V_cons tyid dc "v" | V_consp tyid dc b "v" text \ Binary Operations \ nominal_datatype "opp" = Plus ( "plus") | LEq ("leq") | Eq ("eq") text \ Expressions \ nominal_datatype "e" = AE_val "v" ( "[ _ ]\<^sup>e" ) | AE_app "f" "v" ( "[ _ ( _ ) ]\<^sup>e" ) | AE_appP "f" "b" "v" ( "[_ [ _ ] ( _ )]\<^sup>e" ) | AE_op "opp" "v" "v" ( "[ _ _ _ ]\<^sup>e" ) | AE_concat v v ( "[ _ @@ _ ]\<^sup>e" ) | AE_fst "v" ( "[#1_ ]\<^sup>e" ) | AE_snd "v" ( "[#2_ ]\<^sup>e" ) | AE_mvar "u" ( "[ _ ]\<^sup>e" ) | AE_len "v" ( "[| _ |]\<^sup>e" ) | AE_split "v" "v" ( "[ _ / _ ]\<^sup>e" ) text \ Expressions for constraints\ nominal_datatype "ce" = CE_val "v" ( "[ _ ]\<^sup>c\<^sup>e" ) | CE_op "opp" "ce" "ce" ( "[ _ _ _ ]\<^sup>c\<^sup>e" ) | CE_concat ce ce ( "[ _ @@ _ ]\<^sup>c\<^sup>e" ) | CE_fst "ce" ( "[#1_]\<^sup>c\<^sup>e" ) | CE_snd "ce" ( "[#2_]\<^sup>c\<^sup>e" ) | CE_len "ce" ( "[| _ |]\<^sup>c\<^sup>e" ) text \ Constraints \ nominal_datatype "c" = C_true ( "TRUE" [] 50 ) | C_false ( "FALSE" [] 50 ) | C_conj "c" "c" ("_ AND _ " [50, 50] 50) | C_disj "c" "c" ("_ OR _ " [50,50] 50) | C_not "c" ( "\ _ " [] 50 ) | C_imp "c" "c" ("_ IMP _ " [50, 50] 50) | C_eq "ce" "ce" ("_ == _ " [50, 50] 50) text \ Refined types \ nominal_datatype "\" = T_refined_type x::x b c::c binds x in c ("\ _ : _ | _ \" [50, 50] 1000) text \ Statements \ nominal_datatype s = AS_val v ( "[_]\<^sup>s") | AS_let x::x e s::s binds x in s ( "(LET _ = _ IN _)") | AS_let2 x::x \ s s::s binds x in s ( "(LET _ : _ = _ IN _)") | AS_if v s s ( "(IF _ THEN _ ELSE _)" [0, 61, 0] 61) | AS_var u::u \ v s::s binds u in s ( "(VAR _ : _ = _ IN _)") | AS_assign u v ( "(_ ::= _)") | AS_match v branch_list ( "(MATCH _ WITH { _ })") | AS_while s s ( "(WHILE _ DO { _ } )" [0, 0] 61) | AS_seq s s ( "( _ ;; _ )" [1000, 61] 61) | AS_assert c s ( "(ASSERT _ IN _ )" ) and branch_s = AS_branch dc x::x s::s binds x in s ( "( _ _ \ _ )") and branch_list = AS_final branch_s ( "{ _ }" ) | AS_cons branch_s branch_list ( "( _ | _ )") text \ Function and union type definitions \ nominal_datatype "fun_typ" = AF_fun_typ x::x "b" c::c \::\ s::s binds x in c \ s nominal_datatype "fun_typ_q" = AF_fun_typ_some bv::bv ft::fun_typ binds bv in ft | AF_fun_typ_none fun_typ nominal_datatype "fun_def" = AF_fundef f fun_typ_q nominal_datatype "type_def" = AF_typedef string "(string * \) list" | AF_typedef_poly string bv::bv dclist::"(string * \) list" binds bv in dclist lemma check_typedef_poly: "AF_typedef_poly ''option'' bv [ (''None'', \ zz : B_unit | TRUE \), (''Some'', \ zz : B_var bv | TRUE \) ] = AF_typedef_poly ''option'' bv2 [ (''None'', \ zz : B_unit | TRUE \), (''Some'', \ zz : B_var bv2 | TRUE \) ]" by auto nominal_datatype "var_def" = AV_def u \ v text \ Programs \ nominal_datatype "p" = AP_prog "type_def list" "fun_def list" "var_def list" "s" ("PROG _ _ _ _") declare l.supp [simp] v.supp [simp] e.supp [simp] s_branch_s_branch_list.supp [simp] \.supp [simp] c.supp [simp] b.supp[simp] subsection \Lemmas\ text \These lemmas deal primarily with freshness and alpha-equivalence\ subsubsection \Atoms\ lemma x_not_in_u_atoms[simp]: fixes u::u and x::x and us::"u set" shows "atom x \ atom`us" by (simp add: image_iff) lemma x_fresh_u[simp]: fixes u::u and x::x shows "atom x \ u" by auto lemma x_not_in_b_set[simp]: fixes x::x and bs::"bv fset" shows "atom x \ supp bs" by(induct bs,auto, simp add: supp_finsert supp_at_base) lemma x_fresh_b[simp]: fixes x::x and b::b shows "atom x \ b" apply (induct b rule: b.induct, auto simp: pure_supp) using pure_supp fresh_def by blast+ lemma x_fresh_bv[simp]: fixes x::x and bv::bv shows "atom x \ bv" using fresh_def supp_at_base by auto lemma u_not_in_x_atoms[simp]: fixes u::u and x::x and xs::"x set" shows "atom u \ atom`xs" by (simp add: image_iff) lemma bv_not_in_x_atoms[simp]: fixes bv::bv and x::x and xs::"x set" shows "atom bv \ atom`xs" by (simp add: image_iff) lemma u_not_in_b_atoms[simp]: fixes b :: b and u::u shows "atom u \ supp b" by (induct b rule: b.induct,auto simp: pure_supp supp_at_base) lemma u_not_in_b_set[simp]: fixes u::u and bs::"bv fset" shows "atom u \ supp bs" by(induct bs, auto simp add: supp_at_base supp_finsert) lemma u_fresh_b[simp]: fixes x::u and b::b shows "atom x \ b" by(induct b rule: b.induct, auto simp: pure_fresh ) lemma supp_b_v_disjoint: fixes x::x and bv::bv shows "supp (V_var x) \ supp (B_var bv) = {}" by (simp add: supp_at_base) lemma supp_b_u_disjoint[simp]: fixes b::b and u::u shows "supp u \ supp b = {}" by(nominal_induct b rule:b.strong_induct,(auto simp add: pure_supp b.supp supp_at_base)+) lemma u_fresh_bv[simp]: fixes u::u and b::bv shows "atom u \ b" using fresh_at_base by simp subsubsection \Basic Types\ nominal_function b_of :: "\ \ b" where "b_of \ z : b | c \ = b" apply(auto,simp add: eqvt_def b_of_graph_aux_def ) by (meson \.exhaust) nominal_termination (eqvt) by lexicographic_order lemma supp_b_empty[simp]: fixes b :: b and x::x shows "atom x \ supp b" by (induct b rule: b.induct, auto simp: pure_supp supp_at_base x_not_in_b_set) lemma flip_b_id[simp]: fixes x::x and b::b shows "(x \ x') \ b = b" by(rule flip_fresh_fresh, auto simp add: fresh_def) lemma flip_x_b_cancel[simp]: fixes x::x and y::x and b::b and bv::bv shows "(x \ y) \ b = b" and "(x \ y) \ bv = bv" using flip_b_id apply simp by (metis b.eq_iff(7) b.perm_simps(7) flip_b_id) lemma flip_bv_x_cancel[simp]: fixes bv::bv and z::bv and x::x shows "(bv \ z) \ x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto lemma flip_bv_u_cancel[simp]: fixes bv::bv and z::bv and x::u shows "(bv \ z) \ x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto subsubsection \Literals\ lemma supp_bitvec_empty: fixes bv::"bit list" shows "supp bv = {}" proof(induct bv) case Nil then show ?case using supp_Nil by auto next case (Cons a bv) then show ?case using supp_Cons bit.supp - by (metis (mono_tags, hide_lams) bit.strong_exhaust l.supp(5) sup_bot.right_neutral) + by (metis (mono_tags, opaque_lifting) bit.strong_exhaust l.supp(5) sup_bot.right_neutral) qed lemma bitvec_pure[simp]: fixes bv::"bit list" and x::x shows "atom x \ bv" using fresh_def supp_bitvec_empty by auto lemma supp_l_empty[simp]: fixes l:: l shows "supp (V_lit l) = {}" by(nominal_induct l rule: l.strong_induct, auto simp add: l.strong_exhaust pure_supp v.fv_defs supp_bitvec_empty) lemma type_l_nosupp[simp]: fixes x::x and l::l shows "atom x \ supp (\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \)" using supp_at_base supp_l_empty ce.supp(1) c.supp \.supp by force lemma flip_bitvec0: fixes x::"bit list" assumes "atom c \ (z, x, z')" shows "(z \ c) \ x = (z' \ c) \ x" proof - have "atom z \ x" and "atom z' \ x" using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+ moreover have "atom c \ x" using supp_bitvec_empty fresh_def by auto ultimately show ?thesis using assms flip_fresh_fresh by metis qed lemma flip_bitvec: assumes "atom c \ (z, L_bitvec x, z')" shows "(z \ c) \ x = (z' \ c) \ x" proof - have "atom z \ x" and "atom z' \ x" using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+ moreover have "atom c \ x" using supp_bitvec_empty fresh_def by auto ultimately show ?thesis using assms flip_fresh_fresh by metis qed lemma type_l_eq: shows "\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \ = (\ z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \)" by(auto,nominal_induct l rule: l.strong_induct,auto, metis permute_pure, auto simp add: flip_bitvec) lemma flip_l_eq: fixes x::l shows "(z \ c) \ x = (z' \ c) \ x" proof - have "atom z \ x" and "atom c \ x" and "atom z' \ x" using flip_fresh_fresh fresh_def supp_l_empty by fastforce+ thus ?thesis using flip_fresh_fresh by metis qed lemma flip_l_eq1: fixes x::l assumes "(z \ c) \ x = (z' \ c) \ x'" shows "x' = x" proof - have "atom z \ x" and "atom c \ x'" and "atom c \ x" and "atom z' \ x'" using flip_fresh_fresh fresh_def supp_l_empty by fastforce+ thus ?thesis using flip_fresh_fresh assms by metis qed subsubsection \Types\ lemma flip_base_eq: fixes b::b and x::x and y::x shows "(x \ y) \ b = b" using b.fresh by (simp add: flip_fresh_fresh fresh_def) text \ Obtain an alpha-equivalent type where the bound variable is fresh in some term t \ lemma has_fresh_z0: fixes t::"'b::fs" shows "\z. atom z \ (c',t) \ (\z' : b | c' \) = (\ z : b | (z \ z' ) \ c' \)" proof - obtain z::x where fr: "atom z \ (c',t)" using obtain_fresh by blast moreover hence "(\ z' : b | c' \) = (\ z : b | (z \ z') \ c' \)" using \.eq_iff Abs1_eq_iff by (metis flip_commute flip_fresh_fresh fresh_PairD(1)) ultimately show ?thesis by fastforce qed lemma has_fresh_z: fixes t::"'b::fs" shows "\z b c. atom z \ t \ \ = \ z : b | c \" proof - obtain z' and b and c' where teq: "\ = (\ z' : b | c' \)" using \.exhaust by blast obtain z::x where fr: "atom z \ (t,c')" using obtain_fresh by blast hence "(\ z' : b | c' \) = (\ z : b | (z \ z') \ c' \)" using \.eq_iff Abs1_eq_iff flip_commute flip_fresh_fresh fresh_PairD(1) by (metis fresh_PairD(2)) hence "atom z \ t \ \ = (\ z : b | (z \ z') \ c' \)" using fr teq by force thus ?thesis using teq fr by fast qed lemma obtain_fresh_z: fixes t::"'b::fs" obtains z and b and c where "atom z \ t \ \ = \ z : b | c \" using has_fresh_z by blast lemma has_fresh_z2: fixes t::"'b::fs" shows "\z c. atom z \ t \ \ = \ z : b_of \ | c \" proof - obtain z and b and c where "atom z \ t \ \ = \ z : b | c \" using obtain_fresh_z by metis moreover then have "b_of \ = b" using \.eq_iff by simp ultimately show ?thesis using obtain_fresh_z \.eq_iff by auto qed lemma obtain_fresh_z2: fixes t::"'b::fs" obtains z and c where "atom z \ t \ \ = \ z : b_of \ | c \" using has_fresh_z2 by blast subsubsection \Values\ lemma u_notin_supp_v[simp]: fixes u::u and v::v shows "atom u \ supp v" proof(nominal_induct v rule: v.strong_induct) case (V_lit l) then show ?case using supp_l_empty by auto next case (V_var x) then show ?case by (simp add: supp_at_base) next case (V_pair v1 v2) then show ?case by auto next case (V_cons tyid list v) then show ?case using pure_supp by auto next case (V_consp tyid list b v) then show ?case using pure_supp by auto qed lemma u_fresh_xv[simp]: fixes u::u and x::x and v::v shows "atom u \ (x,v)" proof - have "atom u \ x" using fresh_def by fastforce moreover have "atom u \ v" using fresh_def u_notin_supp_v by metis ultimately show ?thesis using fresh_prod2 by auto qed text \ Part of an effort to make the proofs across inductive cases more uniform by distilling the non-uniform parts into lemmas like this\ lemma v_flip_eq: fixes v::v and va::v and x::x and c::x assumes "atom c \ (v, va)" and "atom c \ (x, xa, v, va)" and "(x \ c) \ v = (xa \ c) \ va" shows "((v = V_lit l \ (\l'. va = V_lit l' \ (x \ c) \ l = (xa \ c) \ l'))) \ ((v = V_var y \ (\y'. va = V_var y' \ (x \ c) \ y = (xa \ c) \ y'))) \ ((v = V_pair vone vtwo \ (\v1' v2'. va = V_pair v1' v2' \ (x \ c) \ vone = (xa \ c) \ v1' \ (x \ c) \ vtwo = (xa \ c) \ v2'))) \ ((v = V_cons tyid dc vone \ (\v1'. va = V_cons tyid dc v1'\ (x \ c) \ vone = (xa \ c) \ v1'))) \ ((v = V_consp tyid dc b vone \ (\v1'. va = V_consp tyid dc b v1'\ (x \ c) \ vone = (xa \ c) \ v1')))" using assms proof(nominal_induct v rule:v.strong_induct) case (V_lit l) then show ?case using assms v.perm_simps empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh by (metis permute_swap_cancel2 v.distinct) next case (V_var x) then show ?case using assms v.perm_simps empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh by (metis permute_swap_cancel2 v.distinct) next case (V_pair v1 v2) have " (V_pair v1 v2 = V_pair vone vtwo \ (\v1' v2'. va = V_pair v1' v2' \ (x \ c) \ vone = (xa \ c) \ v1' \ (x \ c) \ vtwo = (xa \ c) \ v2'))" proof assume "V_pair v1 v2= V_pair vone vtwo" thus "(\v1' v2'. va = V_pair v1' v2' \ (x \ c) \ vone = (xa \ c) \ v1' \ (x \ c) \ vtwo = (xa \ c) \ v2')" using V_pair assms - by (metis (no_types, hide_lams) flip_def permute_swap_cancel v.perm_simps(3)) + by (metis (no_types, opaque_lifting) flip_def permute_swap_cancel v.perm_simps(3)) qed thus ?case using V_pair by auto next case (V_cons tyid dc v1) have " (V_cons tyid dc v1 = V_cons tyid dc vone \ (\ v1'. va = V_cons tyid dc v1' \ (x \ c) \ vone = (xa \ c) \ v1'))" proof assume as: "V_cons tyid dc v1 = V_cons tyid dc vone" hence "(x \ c) \ (V_cons tyid dc vone) = V_cons tyid dc ((x \ c) \ vone)" proof - have "(x \ c) \ dc = dc" using pure_permute_id by metis moreover have "(x \ c) \ tyid = tyid" using pure_permute_id by metis ultimately show ?thesis using v.perm_simps(4) by simp qed then obtain v1' where "(xa \ c) \ va = V_cons tyid dc v1' \ (x \ c) \ vone = v1'" using assms V_cons using as by fastforce hence " va = V_cons tyid dc ((xa \ c) \ v1') \ (x \ c) \ vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh by (metis pure_fresh v.perm_simps(4)) thus "(\v1'. va = V_cons tyid dc v1' \ (x \ c) \ vone = (xa \ c) \ v1')" using V_cons assms by simp qed thus ?case using V_cons by auto next case (V_consp tyid dc b v1) have " (V_consp tyid dc b v1 = V_consp tyid dc b vone \ (\ v1'. va = V_consp tyid dc b v1' \ (x \ c) \ vone = (xa \ c) \ v1'))" proof assume as: "V_consp tyid dc b v1 = V_consp tyid dc b vone" hence "(x \ c) \ (V_consp tyid dc b vone) = V_consp tyid dc b ((x \ c) \ vone)" proof - have "(x \ c) \ dc = dc" using pure_permute_id by metis moreover have "(x \ c) \ tyid = tyid" using pure_permute_id by metis ultimately show ?thesis using v.perm_simps(4) by simp qed then obtain v1' where "(xa \ c) \ va = V_consp tyid dc b v1' \ (x \ c) \ vone = v1'" using assms V_consp using as by fastforce hence " va = V_consp tyid dc b ((xa \ c) \ v1') \ (x \ c) \ vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh pure_fresh v.perm_simps - by (metis (mono_tags, hide_lams)) + by (metis (mono_tags, opaque_lifting)) thus "(\v1'. va = V_consp tyid dc b v1' \ (x \ c) \ vone = (xa \ c) \ v1')" using V_consp assms by simp qed thus ?case using V_consp by auto qed lemma flip_eq: fixes x::x and xa::x and s::"'a::fs" and sa::"'a::fs" assumes "(\c. atom c \ (s, sa) \ atom c \ (x, xa, s, sa) \ (x \ c) \ s = (xa \ c) \ sa)" and "x \ xa" shows "(x \ xa) \ s = sa" proof - have " ([[atom x]]lst. s = [[atom xa]]lst. sa)" using assms Abs1_eq_iff_all by simp hence "(xa = x \ sa = s \ xa \ x \ sa = (xa \ x) \ s \ atom xa \ s)" using assms Abs1_eq_iff[of xa sa x s] by simp thus ?thesis using assms by (metis flip_commute) qed lemma swap_v_supp: fixes v::v and d::x and z::x assumes "atom d \ v" shows "supp ((z \ d ) \ v) \ supp v - { atom z } \ { atom d}" using assms proof(nominal_induct v rule:v.strong_induct) case (V_lit l) then show ?case using l.supp by (metis supp_l_empty empty_subsetI l.strong_exhaust pure_supp supp_eqvt v.supp) next case (V_var x) hence "d \ x" using fresh_def by fastforce thus ?case apply(cases "z = x") using supp_at_base V_var \d\x\ by fastforce+ next case (V_cons tyid dc v) show ?case using v.supp(4) pure_supp using V_cons.hyps V_cons.prems fresh_def by auto next case (V_consp tyid dc b v) show ?case using v.supp(4) pure_supp using V_consp.hyps V_consp.prems fresh_def by auto qed(force+) subsubsection \Expressions\ lemma swap_e_supp: fixes e::e and d::x and z::x assumes "atom d \ e" shows "supp ((z \ d ) \ e) \ supp e - { atom z } \ { atom d}" using assms proof(nominal_induct e rule:e.strong_induct) case (AE_val v) then show ?case using swap_v_supp by simp next case (AE_app f v) then show ?case using swap_v_supp by (simp add: pure_supp) next case (AE_appP b f v) hence df: "atom d \ v" using fresh_def e.supp by force have "supp ((z \ d ) \ (AE_appP b f v)) = supp (AE_appP b f ((z \ d ) \ v))" using e.supp by (metis b.eq_iff(3) b.perm_simps(3) e.perm_simps(3) flip_b_id) also have "... = supp b \ supp f \ supp ((z \ d ) \ v)" using e.supp by auto also have "... \ supp b \ supp f \ supp v - { atom z } \ { atom d}" using swap_v_supp[OF df] pure_supp by auto finally show ?case using e.supp by auto next case (AE_op opp v1 v2) hence df: "atom d \ v1 \ atom d \ v2" using fresh_def e.supp by force have "((z \ d ) \ (AE_op opp v1 v2)) = AE_op opp ((z \ d ) \ v1) ((z \ d ) \ v2)" using e.perm_simps flip_commute opp.perm_simps AE_op opp.strong_exhaust pure_supp by (metis (full_types)) hence "supp ((z \ d) \ AE_op opp v1 v2) = supp (AE_op opp ((z \ d) \v1) ((z \ d) \v2))" by simp also have "... = supp ((z \ d) \v1) \ supp ((z \ d) \v2)" using e.supp - by (metis (mono_tags, hide_lams) opp.strong_exhaust opp.supp sup_bot.left_neutral) + by (metis (mono_tags, opaque_lifting) opp.strong_exhaust opp.supp sup_bot.left_neutral) also have "... \ (supp v1 - { atom z } \ { atom d}) \ (supp v2 - { atom z } \ { atom d})" using swap_v_supp AE_op df by blast finally show ?case using e.supp opp.supp by blast next case (AE_fst v) then show ?case using swap_v_supp by auto next case (AE_snd v) then show ?case using swap_v_supp by auto next case (AE_mvar u) then show ?case using Diff_empty Diff_insert0 Un_upper1 atom_x_sort flip_def flip_fresh_fresh fresh_def set_eq_subset supp_eqvt swap_set_in_eq by (metis sort_of_atom_eq) next case (AE_len v) then show ?case using swap_v_supp by auto next case (AE_concat v1 v2) then show ?case using swap_v_supp by auto next case (AE_split v1 v2) then show ?case using swap_v_supp by auto qed lemma swap_ce_supp: fixes e::ce and d::x and z::x assumes "atom d \ e" shows "supp ((z \ d ) \ e) \ supp e - { atom z } \ { atom d}" using assms proof(nominal_induct e rule:ce.strong_induct) case (CE_val v) then show ?case using swap_v_supp ce.fresh ce.supp by simp next case (CE_op opp v1 v2) hence df: "atom d \ v1 \ atom d \ v2" using fresh_def e.supp by force have "((z \ d ) \ (CE_op opp v1 v2)) = CE_op opp ((z \ d ) \ v1) ((z \ d ) \ v2)" using ce.perm_simps flip_commute opp.perm_simps CE_op opp.strong_exhaust x_fresh_b pure_supp by (metis (full_types)) hence "supp ((z \ d) \ CE_op opp v1 v2) = supp (CE_op opp ((z \ d) \v1) ((z \ d) \v2))" by simp also have "... = supp ((z \ d) \v1) \ supp ((z \ d) \v2)" using ce.supp - by (metis (mono_tags, hide_lams) opp.strong_exhaust opp.supp sup_bot.left_neutral) + by (metis (mono_tags, opaque_lifting) opp.strong_exhaust opp.supp sup_bot.left_neutral) also have "... \ (supp v1 - { atom z } \ { atom d}) \ (supp v2 - { atom z } \ { atom d})" using swap_v_supp CE_op df by blast finally show ?case using ce.supp opp.supp by blast next case (CE_fst v) then show ?case using ce.supp ce.fresh swap_v_supp by auto next case (CE_snd v) then show ?case using ce.supp ce.fresh swap_v_supp by auto next case (CE_len v) then show ?case using ce.supp ce.fresh swap_v_supp by auto next case (CE_concat v1 v2) then show ?case using ce.supp ce.fresh swap_v_supp ce.perm_simps proof - have "\x v xa. \ atom (x::x) \ (v::v) \ supp ((xa \ x) \ v) \ supp v - {atom xa} \ {atom x}" by (meson swap_v_supp) (* 0.0 ms *) then show ?thesis using CE_concat ce.supp by auto qed qed lemma swap_c_supp: fixes c::c and d::x and z::x assumes "atom d \ c" shows "supp ((z \ d ) \ c) \ supp c - { atom z } \ { atom d}" using assms proof(nominal_induct c rule:c.strong_induct) case (C_eq e1 e2) then show ?case using swap_ce_supp by auto qed(auto+) lemma type_e_eq: assumes "atom z \ e" and "atom z' \ e" shows "\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \ = (\ z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \)" by (auto,metis (full_types) assms(1) assms(2) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2)) lemma type_e_eq2: assumes "atom z \ e" and "atom z' \ e" and "b=b'" shows "\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \ = (\ z' : b' | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \)" using assms type_e_eq by fast lemma e_flip_eq: fixes e::e and ea::e assumes "atom c \ (e, ea)" and "atom c \ (x, xa, e, ea)" and "(x \ c) \ e = (xa \ c) \ ea" shows "(e = AE_val w \ (\w'. ea = AE_val w' \ (x \ c) \ w = (xa \ c) \ w')) \ (e = AE_op opp v1 v2 \ (\v1' v2'. ea = AS_op opp v1' v2' \ (x \ c) \ v1 = (xa \ c) \ v1') \ (x \ c) \ v2 = (xa \ c) \ v2') \ (e = AE_fst v \ (\v'. ea = AE_fst v' \ (x \ c) \ v = (xa \ c) \ v')) \ (e = AE_snd v \ (\v'. ea = AE_snd v' \ (x \ c) \ v = (xa \ c) \ v')) \ (e = AE_len v \ (\v'. ea = AE_len v' \ (x \ c) \ v = (xa \ c) \ v')) \ (e = AE_concat v1 v2 \ (\v1' v2'. ea = AS_concat v1' v2' \ (x \ c) \ v1 = (xa \ c) \ v1') \ (x \ c) \ v2 = (xa \ c) \ v2') \ (e = AE_app f v \ (\v'. ea = AE_app f v' \ (x \ c) \ v = (xa \ c) \ v'))" by (metis assms e.perm_simps permute_flip_cancel2) lemma fresh_opp_all: fixes opp::opp shows "z \ opp" using e.fresh opp.exhaust opp.fresh by metis lemma fresh_e_opp_all: shows "(z \ v1 \ z \ v2) = z \ AE_op opp v1 v2" using e.fresh opp.exhaust opp.fresh fresh_opp_all by simp lemma fresh_e_opp: fixes z::x assumes "atom z \ v1 \ atom z \ v2" shows "atom z \ AE_op opp v1 v2" using e.fresh opp.exhaust opp.fresh opp.supp by (metis assms) subsubsection \Statements\ lemma branch_s_flip_eq: fixes v::v and va::v assumes "atom c \ (v, va)" and "atom c \ (x, xa, v, va)" and "(x \ c) \ s = (xa \ c) \ sa" shows "(s = AS_val w \ (\w'. sa = AS_val w' \ (x \ c) \ w = (xa \ c) \ w')) \ (s = AS_seq s1 s2 \ (\s1' s2'. sa = AS_seq s1' s2' \ (x \ c) \ s1 = (xa \ c) \ s1') \ (x \ c) \ s2 = (xa \ c) \ s2') \ (s = AS_if v s1 s2 \ (\v' s1' s2'. sa = AS_if seq s1' s2' \ (x \ c) \ s1 = (xa \ c) \ s1') \ (x \ c) \ s2 = (xa \ c) \ s2' \ (x \ c) \ c = (xa \ c) \ v')" by (metis assms s_branch_s_branch_list.perm_simps permute_flip_cancel2) section \Context Syntax\ subsection \Datatypes\ text \Type and function/type definition contexts\ type_synonym \ = "fun_def list" type_synonym \ = "type_def list" type_synonym \ = "bv fset" datatype \ = GNil | GCons "x*b*c" \ (infixr "#\<^sub>\" 65) datatype \ = DNil ("[]\<^sub>\") | DCons "u*\" \ (infixr "#\<^sub>\" 65) subsection \Functions and Lemmas\ lemma \_induct [case_names GNil GCons] : "P GNil \ (\x b c \'. P \' \ P ((x,b,c) #\<^sub>\ \')) \ P \" proof(induct \ rule:\.induct) case GNil then show ?case by auto next case (GCons x1 x2) then obtain x and b and c where "x1=(x,b,c)" using prod_cases3 by blast then show ?case using GCons by presburger qed instantiation \ :: pt begin primrec permute_\ where DNil_eqvt: "p \ DNil = DNil" | DCons_eqvt: "p \ (x #\<^sub>\ xs) = p \ x #\<^sub>\ p \ (xs::\)" instance by standard (induct_tac [!] x, simp_all) end lemmas [eqvt] = permute_\.simps lemma \_induct [case_names DNil DCons] : "P DNil \ (\u t \'. P \' \ P ((u,t) #\<^sub>\ \')) \ P \" proof(induct \ rule: \.induct) case DNil then show ?case by auto next case (DCons x1 x2) then obtain u and t where "x1=(u,t)" by fastforce then show ?case using DCons by presburger qed lemma \_induct [case_names PNil PConsNone PConsSome] : "P [] \ (\f x b c \ s' \'. P \' \ P ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s'))) # \')) \ (\f bv x b c \ s' \'. P \' \ P ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s'))) # \')) \ P \" proof(induct \ rule: list.induct) case Nil then show ?case by auto next case (Cons x1 x2) then obtain f and t where ft: "x1 = (AF_fundef f t)" by (meson fun_def.exhaust) then show ?case proof(nominal_induct t rule:fun_typ_q.strong_induct) case (AF_fun_typ_some bv ft) then show ?case using Cons ft by (metis fun_typ.exhaust) next case (AF_fun_typ_none ft) then show ?case using Cons ft by (metis fun_typ.exhaust) qed qed lemma \_induct [case_names TNil AF_typedef AF_typedef_poly] : "P [] \ (\tid dclist \'. P \' \ P ((AF_typedef tid dclist) # \')) \ (\tid bv dclist \'. P \' \ P ((AF_typedef_poly tid bv dclist ) # \')) \ P \" proof(induct \ rule: list.induct) case Nil then show ?case by auto next case (Cons td T) show ?case by(cases td rule: type_def.exhaust, (simp add: Cons)+) qed instantiation \ :: pt begin primrec permute_\ where GNil_eqvt: "p \ GNil = GNil" | GCons_eqvt: "p \ (x #\<^sub>\ xs) = p \ x #\<^sub>\ p \ (xs::\)" instance by standard (induct_tac [!] x, simp_all) end lemmas [eqvt] = permute_\.simps lemma G_cons_eqvt[simp]: fixes \::\ shows "p \ ((x,b,c) #\<^sub>\ \) = ((p \ x, p \ b , p \ c) #\<^sub>\ (p \ \ ))" (is "?A = ?B" ) using Cons_eqvt triple_eqvt supp_b_empty by simp lemma G_cons_flip[simp]: fixes x::x and \::\ shows "(x\x') \ ((x'',b,c) #\<^sub>\ \) = (((x\x') \ x'', b , (x\x') \ c) #\<^sub>\ ((x\x') \ \ ))" using Cons_eqvt triple_eqvt supp_b_empty by auto lemma G_cons_flip_fresh[simp]: fixes x::x and \::\ assumes "atom x \ (c,\)" and "atom x' \ (c,\)" shows "(x\x') \ ((x',b,c) #\<^sub>\ \) = ((x, b , c) #\<^sub>\ \ )" using G_cons_flip flip_fresh_fresh assms by force lemma G_cons_flip_fresh2[simp]: fixes x::x and \::\ assumes "atom x \ (c,\)" and "atom x' \ (c,\)" shows "(x\x') \ ((x,b,c) #\<^sub>\ \) = ((x', b , c) #\<^sub>\ \ )" using G_cons_flip flip_fresh_fresh assms by force lemma G_cons_flip_fresh3[simp]: fixes x::x and \::\ assumes "atom x \ \" and "atom x' \ \" shows "(x\x') \ ((x',b,c) #\<^sub>\ \) = ((x, b , (x\x') \ c) #\<^sub>\ \ )" using G_cons_flip flip_fresh_fresh assms by force lemma neq_GNil_conv: "(xs \ GNil) = (\y ys. xs = y #\<^sub>\ ys)" by (induct xs) auto nominal_function toList :: "\ \ (x*b*c) list" where "toList GNil = []" | "toList (GCons xbc G) = xbc#(toList G)" apply (auto, simp add: eqvt_def toList_graph_aux_def ) using neq_GNil_conv surj_pair by metis nominal_termination (eqvt) by lexicographic_order nominal_function toSet :: "\ \ (x*b*c) set" where "toSet GNil = {}" | "toSet (GCons xbc G) = {xbc} \ (toSet G)" apply (auto,simp add: eqvt_def toSet_graph_aux_def ) using neq_GNil_conv surj_pair by metis nominal_termination (eqvt) by lexicographic_order nominal_function append_g :: "\ \ \ \ \" (infixr "@" 65) where "append_g GNil g = g" | "append_g (xbc #\<^sub>\ g1) g2 = (xbc #\<^sub>\ (g1@g2))" apply (auto,simp add: eqvt_def append_g_graph_aux_def ) using neq_GNil_conv surj_pair by metis nominal_termination (eqvt) by lexicographic_order nominal_function dom :: "\ \ x set" where "dom \ = (fst` (toSet \))" apply auto unfolding eqvt_def dom_graph_aux_def lfp_eqvt toSet.eqvt by simp nominal_termination (eqvt) by lexicographic_order text \ Use of this is sometimes mixed in with use of freshness and support for the context however it makes it clear that for immutable variables, the context is `self-supporting'\ nominal_function atom_dom :: "\ \ atom set" where "atom_dom \ = atom`(dom \)" apply auto unfolding eqvt_def atom_dom_graph_aux_def lfp_eqvt toSet.eqvt by simp nominal_termination (eqvt) by lexicographic_order subsection \Immutable Variable Context Lemmas\ lemma append_GNil[simp]: "GNil @ G = G" by simp lemma append_g_toSetU [simp]: "toSet (G1@G2) = toSet G1 \ toSet G2" by(induct G1, auto+) lemma supp_GNil: shows "supp GNil = {}" by (simp add: supp_def) lemma supp_GCons: fixes xs::\ shows "supp (x #\<^sub>\ xs) = supp x \ supp xs" by (simp add: supp_def Collect_imp_eq Collect_neg_eq) lemma atom_dom_eq[simp]: fixes G::\ shows "atom_dom ((x, b, c) #\<^sub>\ G) = atom_dom ((x, b, c') #\<^sub>\ G)" using atom_dom.simps toSet.simps by simp lemma dom_append[simp]: "atom_dom (\@\') = atom_dom \ \ atom_dom \'" using image_Un append_g_toSetU atom_dom.simps dom.simps by metis lemma dom_cons[simp]: "atom_dom ((x,b,c) #\<^sub>\ G) = { atom x } \ atom_dom G" using image_Un append_g_toSetU atom_dom.simps by auto lemma fresh_GNil[ms_fresh]: shows "a \ GNil" by (simp add: fresh_def supp_GNil) lemma fresh_GCons[ms_fresh]: fixes xs::\ shows "a \ (x #\<^sub>\ xs) \ a \ x \ a \ xs" by (simp add: fresh_def supp_GCons) lemma dom_supp_g[simp]: "atom_dom G \ supp G" apply(induct G rule: \_induct,simp) using supp_at_base supp_Pair atom_dom.simps supp_GCons by fastforce lemma fresh_append_g[ms_fresh]: fixes xs::\ shows "a \ (xs @ ys) \ a \ xs \ a \ ys" by (induct xs) (simp_all add: fresh_GNil fresh_GCons) lemma append_g_assoc: fixes xs::\ shows "(xs @ ys) @ zs = xs @ (ys @ zs)" by (induct xs) simp_all lemma append_g_inside: fixes xs::\ shows "xs @ (x #\<^sub>\ ys) = (xs @ (x #\<^sub>\ GNil)) @ ys" by(induct xs,auto+) lemma finite_\: "finite (toSet \)" by(induct \ rule: \_induct,auto) lemma supp_\: "supp \ = supp (toSet \)" proof(induct \ rule: \_induct) case GNil then show ?case using supp_GNil toSet.simps by (simp add: supp_set_empty) next case (GCons x b c \') then show ?case using supp_GCons toSet.simps finite_\ supp_of_finite_union using supp_of_finite_insert by fastforce qed lemma supp_of_subset: fixes G::"('a::fs set)" assumes "finite G" and "finite G'" and "G \ G'" shows "supp G \ supp G'" using supp_of_finite_sets assms by (metis subset_Un_eq supp_of_finite_union) lemma supp_weakening: assumes "toSet G \ toSet G'" shows "supp G \ supp G'" using supp_\ finite_\ by (simp add: supp_of_subset assms) lemma fresh_weakening[ms_fresh]: assumes "toSet G \ toSet G'" and "x \ G'" shows "x \ G" proof(rule ccontr) assume "\ x \ G" hence "x \ supp G" using fresh_def by auto hence "x \ supp G'" using supp_weakening assms by auto thus False using fresh_def assms by auto qed instance \ :: fs by (standard, induct_tac x, simp_all add: supp_GNil supp_GCons finite_supp) lemma fresh_gamma_elem: fixes \::\ assumes "a \ \" and "e \ toSet \" shows "a \ e" using assms by(induct \,auto simp add: fresh_GCons) lemma fresh_gamma_append: fixes xs::\ shows "a \ (xs @ ys) \ a \ xs \ a \ ys" by (induct xs, simp_all add: fresh_GNil fresh_GCons) lemma supp_triple[simp]: shows "supp (x, y, z) = supp x \ supp y \ supp z" proof - have "supp (x,y,z) = supp (x,(y,z))" by auto hence "supp (x,y,z) = supp x \ (supp y \ supp z)" using supp_Pair by metis thus ?thesis by auto qed lemma supp_append_g: fixes xs::\ shows "supp (xs @ ys) = supp xs \ supp ys" by(induct xs,auto simp add: supp_GNil supp_GCons ) lemma fresh_in_g[simp]: fixes \::\ and x'::x shows "atom x' \ \' @ (x, b0, c0) #\<^sub>\ \ = (atom x' \ supp \' \ supp x \ supp b0 \ supp c0 \ supp \)" proof - have "atom x' \ \' @ (x, b0, c0) #\<^sub>\ \ = (atom x' \ supp (\' @((x,b0,c0) #\<^sub>\ \)))" using fresh_def by auto also have "... = (atom x' \ supp \' \ supp ((x,b0,c0) #\<^sub>\ \))" using supp_append_g by fast also have "... = (atom x' \ supp \' \ supp x \ supp b0 \ supp c0 \ supp \)" using supp_GCons supp_append_g supp_triple by auto finally show ?thesis by fast qed lemma fresh_suffix[ms_fresh]: fixes \::\ assumes "atom x \ \'@\" shows "atom x \ \" using assms by(induct \' rule: \_induct, auto simp add: append_g.simps fresh_GCons) lemma not_GCons_self [simp]: fixes xs::\ shows "xs \ x #\<^sub>\ xs" by (induct xs) auto lemma not_GCons_self2 [simp]: fixes xs::\ shows "x #\<^sub>\ xs \ xs" by (rule not_GCons_self [symmetric]) lemma fresh_restrict: fixes y::x and \::\ assumes "atom y \ (\' @ (x, b, c) #\<^sub>\ \)" shows "atom y \ (\'@\)" using assms by(induct \' rule: \_induct, auto simp add:fresh_GCons fresh_GNil ) lemma fresh_dom_free: assumes "atom x \ \" shows "(x,b,c) \ toSet \" using assms proof(induct \ rule: \_induct) case GNil then show ?case by auto next case (GCons x' b' c' \') hence "x\x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by blast moreover have "atom x \ \'" using fresh_GCons GCons by auto ultimately show ?case using toSet.simps GCons by auto qed lemma \_set_intros: "x \ toSet ( x #\<^sub>\ xs)" and "y \ toSet xs \ y \ toSet (x #\<^sub>\ xs)" by simp+ lemma fresh_dom_free2: assumes "atom x \ atom_dom \" shows "(x,b,c) \ toSet \" using assms proof(induct \ rule: \_induct) case GNil then show ?case by auto next case (GCons x' b' c' \') hence "x\x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by auto moreover have "atom x \ atom_dom \'" using fresh_GCons GCons by auto ultimately show ?case using toSet.simps GCons by auto qed subsection \Mutable Variable Context Lemmas\ lemma supp_DNil: shows "supp DNil = {}" by (simp add: supp_def) lemma supp_DCons: fixes xs::\ shows "supp (x #\<^sub>\ xs) = supp x \ supp xs" by (simp add: supp_def Collect_imp_eq Collect_neg_eq) lemma fresh_DNil[ms_fresh]: shows "a \ DNil" by (simp add: fresh_def supp_DNil) lemma fresh_DCons[ms_fresh]: fixes xs::\ shows "a \ (x #\<^sub>\ xs) \ a \ x \ a \ xs" by (simp add: fresh_def supp_DCons) instance \ :: fs by (standard, induct_tac x, simp_all add: supp_DNil supp_DCons finite_supp) subsection \Lookup Functions\ nominal_function lookup :: "\ \ x \ (b*c) option" where "lookup GNil x = None" | "lookup ((x,b,c)#\<^sub>\G) y = (if x=y then Some (b,c) else lookup G y)" by (auto,simp add: eqvt_def lookup_graph_aux_def, metis neq_GNil_conv surj_pair) nominal_termination (eqvt) by lexicographic_order nominal_function replace_in_g :: "\ \ x \ c \ \" ("_[_\_]" [1000,0,0] 200) where "replace_in_g GNil _ _ = GNil" | "replace_in_g ((x,b,c)#\<^sub>\G) x' c' = (if x=x' then ((x,b,c')#\<^sub>\G) else (x,b,c)#\<^sub>\(replace_in_g G x' c'))" apply(auto,simp add: eqvt_def replace_in_g_graph_aux_def) using surj_pair \.exhaust by metis nominal_termination (eqvt) by lexicographic_order text \ Functions for looking up data-constructors in the Pi context \ nominal_function lookup_fun :: "\ \ f \ fun_def option" where "lookup_fun [] g = None" | "lookup_fun ((AF_fundef f ft)#\) g = (if (f=g) then Some (AF_fundef f ft) else lookup_fun \ g)" apply(auto,simp add: eqvt_def lookup_fun_graph_aux_def ) by (metis fun_def.exhaust neq_Nil_conv) nominal_termination (eqvt) by lexicographic_order nominal_function lookup_td :: "\ \ string \ type_def option" where "lookup_td [] g = None" | "lookup_td ((AF_typedef s lst ) # (\::\)) g = (if (s = g) then Some (AF_typedef s lst ) else lookup_td \ g)" | "lookup_td ((AF_typedef_poly s bv lst ) # (\::\)) g = (if (s = g) then Some (AF_typedef_poly s bv lst ) else lookup_td \ g)" apply(auto,simp add: eqvt_def lookup_td_graph_aux_def ) by (metis type_def.exhaust neq_Nil_conv) nominal_termination (eqvt) by lexicographic_order nominal_function name_of_type ::"type_def \ f " where "name_of_type (AF_typedef f _ ) = f" | "name_of_type (AF_typedef_poly f _ _) = f" apply(auto,simp add: eqvt_def name_of_type_graph_aux_def ) using type_def.exhaust by blast nominal_termination (eqvt) by lexicographic_order nominal_function name_of_fun ::"fun_def \ f " where "name_of_fun (AF_fundef f ft) = f" apply(auto,simp add: eqvt_def name_of_fun_graph_aux_def ) using fun_def.exhaust by blast nominal_termination (eqvt) by lexicographic_order nominal_function remove2 :: "'a::pt \ 'a list \ 'a list" where "remove2 x [] = []" | "remove2 x (y # xs) = (if x = y then xs else y # remove2 x xs)" by (simp add: eqvt_def remove2_graph_aux_def,auto+,meson list.exhaust) nominal_termination (eqvt) by lexicographic_order nominal_function base_for_lit :: "l \ b" where "base_for_lit (L_true) = B_bool " | "base_for_lit (L_false) = B_bool " | "base_for_lit (L_num n) = B_int " | "base_for_lit (L_unit) = B_unit " | "base_for_lit (L_bitvec v) = B_bitvec" apply (auto simp: eqvt_def base_for_lit_graph_aux_def ) using l.strong_exhaust by blast nominal_termination (eqvt) by lexicographic_order lemma neq_DNil_conv: "(xs \ DNil) = (\y ys. xs = y #\<^sub>\ ys)" by (induct xs) auto nominal_function setD :: "\ \ (u*\) set" where "setD DNil = {}" | "setD (DCons xbc G) = {xbc} \ (setD G)" apply (auto,simp add: eqvt_def setD_graph_aux_def ) using neq_DNil_conv surj_pair by metis nominal_termination (eqvt) by lexicographic_order lemma eqvt_triple: fixes y::"'a::at" and ya::"'a::at" and xa::"'c::at" and va::"'d::fs" and s::s and sa::s and f::"s*'c*'d \ s" assumes "atom y \ (xa, va)" and "atom ya \ (xa, va)" and "\c. atom c \ (s, sa) \ atom c \ (y, ya, s, sa) \ (y \ c) \ s = (ya \ c) \ sa" and "eqvt_at f (s,xa,va)" and "eqvt_at f (sa,xa,va)" and "atom c \ (s, va, xa, sa)" and "atom c \ (y, ya, f (s, xa, va), f (sa, xa, va))" shows "(y \ c) \ f (s, xa, va) = (ya \ c) \ f (sa, xa, va)" proof - have " (y \ c) \ f (s, xa, va) = f ( (y \ c) \ (s,xa,va))" using assms eqvt_at_def by metis also have "... = f ( (y \ c) \ s, (y \ c) \ xa ,(y \ c) \ va)" by auto also have "... = f ( (ya \ c) \ sa, (ya \ c) \ xa ,(ya \ c) \ va)" proof - have " (y \ c) \ s = (ya \ c) \ sa" using assms Abs1_eq_iff_all by auto moreover have "((y \ c) \ xa) = ((ya \ c) \ xa)" using assms flip_fresh_fresh fresh_prodN by metis moreover have "((y \ c) \ va) = ((ya \ c) \ va)" using assms flip_fresh_fresh fresh_prodN by metis ultimately show ?thesis by auto qed also have "... = f ( (ya \ c) \ (sa,xa,va))" by auto finally show ?thesis using assms eqvt_at_def by metis qed section \Functions for bit list/vectors\ inductive split :: "int \ bit list \ bit list * bit list \ bool" where "split 0 xs ([], xs)" | "split m xs (ys,zs) \ split (m+1) (x#xs) ((x # ys), zs)" equivariance split nominal_inductive split . lemma split_concat: assumes "split n v (v1,v2)" shows "v = append v1 v2" using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts) case 1 then show ?case by auto next case (2 m xs ys zs x) then show ?case by auto qed lemma split_n: assumes "split n v (v1,v2)" shows "0 \ n \ n \ int (length v)" using assms proof(induct rule: split.inducts) case (1 xs) then show ?case by auto next case (2 m xs ys zs x) then show ?case by auto qed lemma split_length: assumes "split n v (v1,v2)" shows "n = int (length v1)" using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts) case (1 xs) then show ?case by auto next case (2 m xs ys zs x) then show ?case by auto qed lemma obtain_split: assumes "0 \ n" and "n \ int (length bv)" shows "\ bv1 bv2. split n bv (bv1 , bv2)" using assms proof(induct bv arbitrary: n) case Nil then show ?case using split.intros by auto next case (Cons b bv) show ?case proof(cases "n = 0") case True then show ?thesis using split.intros by auto next case False then obtain m where m:"n=m+1" using Cons by (metis add.commute add_minus_cancel) moreover have "0 \ m" using False m Cons by linarith then obtain bv1 and bv2 where "split m bv (bv1 , bv2)" using Cons m by force hence "split n (b # bv) ((b#bv1), bv2)" using m split.intros by auto then show ?thesis by auto qed qed end \ No newline at end of file diff --git a/thys/MiniSail/WellformedL.thy b/thys/MiniSail/WellformedL.thy --- a/thys/MiniSail/WellformedL.thy +++ b/thys/MiniSail/WellformedL.thy @@ -1,4755 +1,4755 @@ (*<*) theory WellformedL imports Wellformed "SyntaxL" begin (*>*) chapter \Wellformedness Lemmas\ section \Prelude\ lemma b_of_subst_bb_commute: "(b_of (\[bv::=b]\<^sub>\\<^sub>b)) = (b_of \)[bv::=b]\<^sub>b\<^sub>b" proof - obtain z' and b' and c' where "\ = \ z' : b' | c' \ " using obtain_fresh_z by metis moreover hence "(b_of (\[bv::=b]\<^sub>\\<^sub>b)) = b_of \ z' : b'[bv::=b]\<^sub>b\<^sub>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 "\; \; \ \\<^sub>w\<^sub>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 \ \; \ \\<^sub>w\<^sub>f b \ atom bv \ (\, \, \, b, v) \ \; \; \ \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \ atom bv \ tm)" and "\; \; \ \\<^sub>w\<^sub>f c \ True" and "\; \ \\<^sub>w\<^sub>f \ \ True" and "\; \; \ \\<^sub>w\<^sub>f \ \ True" and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f \ \True" and "\; \ \\<^sub>w\<^sub>f b \ True " and "\; \; \ \\<^sub>w\<^sub>f ce : b' \ True" and "\ \\<^sub>w\<^sub>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" (" _ ; _ \\<^sub>w\<^sub>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; \ \\<^sub>w\<^sub>f (z,b,c) #\<^sub>\\" shows "P; \ \\<^sub>w\<^sub>f \ \ atom z \ \ \ wfB P \ b" using wfG_elims(2)[OF assms] by metis lemma wfG_cons2[ms_wb]: fixes \::\ assumes "P; \ \\<^sub>w\<^sub>f zbc #\<^sub>\\" shows "P; \ \\<^sub>w\<^sub>f \" proof - obtain z and b and c where zbc: "zbc=(z,b,c)" using prod_cases3 by blast hence "P; \ \\<^sub>w\<^sub>f (z,b,c) #\<^sub>\\" using assms by auto thus ?thesis using zbc wfG_cons assms by simp qed lemma wf_g_unique: fixes \::\ assumes "\; \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>f ((x'',b'',c'') #\<^sub>\\)" using GCons wfG_elims by blast hence "\; \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>f \" and "Some (b,c) = lookup \ x" shows "\; \ \\<^sub>w\<^sub>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) #\<^sub>\ \)" and "(x1,b1,c1) \ toSet ((x, b, c) #\<^sub>\ \)" and "x1=x" shows "b1 = b \ c1 = c" proof - have "(x, b, c) \ toSet ((x, b, c) #\<^sub>\ \)" by simp thus ?thesis using wf_g_unique assms by blast qed lemma wfG_unique_full: fixes \::\ assumes "wfG \ B (\'@(x, b, c) #\<^sub>\ \)" and "(x1,b1,c1) \ toSet (\'@(x, b, c) #\<^sub>\ \)" and "x1=x" shows "b1 = b \ c1 = c" proof - have "(x, b, c) \ toSet (\'@(x, b, c) #\<^sub>\ \)" 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: "\; \; \ \\<^sub>w\<^sub>f v : b \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and wfC_wf: "\; \; \ \\<^sub>w\<^sub>f c \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and wfG_wf :"\; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \" and wfT_wf: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f b_of \" and wfTs_wf:"\; \; \ \\<^sub>w\<^sub>f ts \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \" and "\\<^sub>w\<^sub>f \ \ True" and wfB_wf: "\; \ \\<^sub>w\<^sub>f b \ \\<^sub>w\<^sub>f \" and wfCE_wf: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and wfTD_wf: "\ \\<^sub>w\<^sub>f td \ \\<^sub>w\<^sub>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: "\; \; \; \; \ \\<^sub>w\<^sub>f e : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and wfS_wf: "\; \; \; \; \ \\<^sub>w\<^sub>f s : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and "\; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and wfPhi_wf: "\ \\<^sub>w\<^sub>f (\::\) \ \\<^sub>w\<^sub>f \" and wfD_wf: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and wfFTQ_wf: "\ ; \ \\<^sub>w\<^sub>f ftq \ \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \" and wfFT_wf: "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>f \" using wfX_wfY1 by auto moreover have "\; \; \ \\<^sub>w\<^sub>f \" using wfS_assertI by auto moreover have "\\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>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' #\<^sub>\ 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 "\; \; \ \\<^sub>w\<^sub>f \" shows "\(u,\) \ setD \. \; \; \ \\<^sub>w\<^sub>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]\<^sub>\\<^sub>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 "\; \ \\<^sub>w\<^sub>f \" and "atom x \ \" and "wfB \ \ b" and "c \ { TRUE, FALSE } \ \; \ \\<^sub>w\<^sub>f \ " and "c \ { TRUE, FALSE } \ \ ;\ ; (x,b,C_true) #\<^sub>\\ \\<^sub>w\<^sub>f c" shows "\; \ \\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\\)" using wfG_cons1I wfG_cons2I assms by metis lemma wfG_consI: fixes \::\ and \::\ and c::c assumes "\; \ \\<^sub>w\<^sub>f \" and "atom x \ \" and "wfB \ \ b" and "\ ; \ ; (x,b,C_true) #\<^sub>\\ \\<^sub>w\<^sub>f c" shows "\ ; \ \\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\\)" 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) #\<^sub>\\)" shows "P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c \ wfB P \ b" proof(cases "c \ {TRUE,FALSE}") case True have "P; \ \\<^sub>w\<^sub>f \ \ atom x \ \ \ wfB P \ b" using wfG_elims(2)[OF assms] by auto hence "P; \ \\<^sub>w\<^sub>f ((x,b,TRUE) #\<^sub>\\) \ 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 \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \" shows "\ ; B ; ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>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; \ \\<^sub>w\<^sub>f \" and "atom z \ \" and "P; \ \\<^sub>w\<^sub>f b" shows "P ; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \" using wfG_cons2I wfG_wfB assms by simp lemma wfG_cons_TRUE2: assumes "P; \ \\<^sub>w\<^sub>f (z,b,c) #\<^sub>\\" and "atom z \ \" shows "P; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \" 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; \ \\<^sub>w\<^sub>f \' @ \" using wfG_elims by auto then show ?case using GCons wfG_elims by auto qed lemma wfV_wfCE: fixes v::v assumes "\; \; \ \\<^sub>w\<^sub>f v : b" shows " \ ; \ ; \ \\<^sub>w\<^sub>f CE_val v : b" proof - have "\ \\<^sub>w\<^sub>f ([]::\) " using wfPhi_emptyI wfV_wf wfG_wf assms by metis moreover have "\; \; \ \\<^sub>w\<^sub>f []\<^sub>\" 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: "\; \; \ \\<^sub>w\<^sub>f v : b \ supp v \ atom_dom \ \ supp \" and wfC_supp: "\; \; \ \\<^sub>w\<^sub>f c \ supp c \ atom_dom \ \ supp \" and wfG_supp: "\; \ \\<^sub>w\<^sub>f \ \ atom_dom \ \ supp \" and wfT_supp: "\; \; \ \\<^sub>w\<^sub>f \ \ supp \ \ atom_dom \ \ supp \ " and wfTs_supp: "\; \; \ \\<^sub>w\<^sub>f ts \ supp ts \ atom_dom \ \ supp \" and wfTh_supp: "\\<^sub>w\<^sub>f \ \ supp \ = {}" and wfB_supp: "\; \ \\<^sub>w\<^sub>f b \ supp b \ supp \" and wfCE_supp: "\; \; \ \\<^sub>w\<^sub>f ce : b \ supp ce \ atom_dom \ \ supp \" and wfTD_supp: "\ \\<^sub>w\<^sub>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: "\; \; \; \; \ \\<^sub>w\<^sub>f e : b \ (supp e \ atom_dom \ \ supp \ \ atom ` fst ` setD \)" and (* \ ( \ = [] \ supp e \ supp \ = {})" and*) wfS_supp: "\; \; \; \; \ \\<^sub>w\<^sub>f s : b \ supp s \ atom_dom \ \ atom ` fst ` setD \ \ supp \" and "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ supp cs \ atom_dom \ \ atom ` fst ` setD \ \ supp \" and "\; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ supp css \ atom_dom \ \ atom ` fst ` setD \ \ supp \" and wfPhi_supp: "\ \\<^sub>w\<^sub>f (\::\) \ supp \ = {}" and wfD_supp: "\; \; \ \\<^sub>w\<^sub>f \ \ supp \ \ atom`fst`(setD \) \ atom_dom \ \ supp \ " and "\ ; \ \\<^sub>w\<^sub>f ftq \ supp ftq = {}" and "\ ; \ ; \ \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f v : ( b[bv::=b']\<^sub>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, \) #\<^sub>\ \) = 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, \) #\<^sub>\ \) \ 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 (* 0.0 ms *) have "supp b \ supp B'" using wfFTI.hyps(1) wf_supp1(7) by blast (* 0.0 ms *) 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 (* 234 ms *) 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) #\<^sub>\ \) \ atom ` fst ` setD \ \ supp \" using wfS_assertI * by blast have "z \ atom_dom ((x, B_bool, c) #\<^sub>\ \) \ z \ atom_dom \" using * ** by auto thus "z \ atom_dom \ \ atom ` fst ` setD \ \ supp \" using * ** using \z \ atom_dom ((x, B_bool, c) #\<^sub>\ \) \ 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 \\<^sub>w\<^sub>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; \ \\<^sub>w\<^sub>f b \ using assms by auto show \ P ;\ ; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>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) #\<^sub>\G \\<^sub>w\<^sub>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') #\<^sub>\ 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) #\<^sub>\G \\<^sub>w\<^sub>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') #\<^sub>\ 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 "\\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom z \ \" shows "\; \; (z,b,TRUE) #\<^sub>\\ \\<^sub>w\<^sub>f c" proof - obtain za ba ca where *:"\ z : b | c \ = \ za : ba | ca \ \ atom za \ (\,\,\) \ \; \; (za, ba, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>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 " \\<^sub>w\<^sub>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) #\<^sub>\ \ \\<^sub>w\<^sub>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) #\<^sub>\ \' \\<^sub>w\<^sub>f c" using wfG_cons_wfC GCons by simp hence "supp c \ atom_dom ((x, b, TRUE) #\<^sub>\ \') \ 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) #\<^sub>\\)" shows "wfC P \ ((x,b,TRUE) #\<^sub>\\) 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) #\<^sub>\\)" 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) #\<^sub>\\)" 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 "\; \; \ \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>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]\<^sub>c\<^sub>v) #\<^sub>\ G)" and "atom x \ c" shows "P; \ ; G \\<^sub>w\<^sub>f \ z : b | c \" proof - have " P; \ ; (x, b, TRUE) #\<^sub>\ G \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>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]\<^sub>c\<^sub>v \" using wfTI assms by metis moreover have "\ x : b | c[z::=V_var x]\<^sub>c\<^sub>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]\<^sub>c\<^sub>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) \) \\<^sub>w\<^sub>f (CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]\<^sub>c\<^sub>v)" using wfT_wfC assms fresh_Pair by auto hence "\; \; ((z2,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c[z::=V_var z2]\<^sub>c\<^sub>v" using wfC_elims by metis hence "wfT \ \ \ (\ z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>v\)" using assms fresh_Pair wfTI \wfB \ \ b\ * by auto moreover have "\ z : b | c \ = \ z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f \" and "\ \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>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 "\; \; \; \ ; \ \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>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 "\; \; \; \; \ \\<^sub>w\<^sub>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) #\<^sub>\\))" 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; \ \\<^sub>w\<^sub>f (\1 @ (x, b, c) #\<^sub>\ \)" using wfG_elims append_g.simps GCons by metis have "atom x1 \ (\1 @ (x, b, c) #\<^sub>\ \)" using GCons wfG_elims append_g.simps by metis hence "atom x1 \ atom_dom (\1 @ (x, b, c) #\<^sub>\ \)" 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]\<^sub>c\<^sub>v) #\<^sub>\ \)" 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]\<^sub>c\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and " \; \ \\<^sub>w\<^sub>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]\<^sub>c\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and " \; \ \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\ \') @ (x, b, c) #\<^sub>\ \" shows "\; \ \\<^sub>w\<^sub>f (\' @ (x, b, c) #\<^sub>\ \) \ atom x' \ (\' @ (x, b, c) #\<^sub>\ \) \ \; \ \\<^sub>w\<^sub>f b' \ x' \ x" proof - have "((x', b', c') #\<^sub>\ \') @ (x, b, c) #\<^sub>\ \ = (x', b', c') #\<^sub>\ (\' @ (x, b, c) #\<^sub>\ \)" using append_g.simps by auto hence *:"\; \ \\<^sub>w\<^sub>f (\' @ (x, b, c) #\<^sub>\ \) \ atom x' \ (\' @ (x, b, c) #\<^sub>\ \) \ \; \ \\<^sub>w\<^sub>f b'" using assms wfG_cons by metis moreover have "atom x' \ x" proof(rule wfG_inside_x_fresh[of "(\' @ (x, b, c) #\<^sub>\ \)"]) show "\' @ (x, b, c) #\<^sub>\ \ = \' @ (x, b, c[x::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \" by simp show " atom x' \ \' @ (x, b, c) #\<^sub>\ \" using * by auto show "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c) #\<^sub>\ \ " using * by auto qed ultimately show ?thesis by auto qed lemma flip_u_eq: fixes u::u and u'::u and \::\ and \::\ assumes "\; \; \ \\<^sub>w\<^sub>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]\<^sub>c\<^sub>v) #\<^sub>\\)" proof - have "\ x : b | c[z::=V_var x]\<^sub>c\<^sub>v \ = \ z : b | c \" using assms freshers type_eq_subst by metis hence *:"wfT P \ \ \ x : b | c[z::=V_var x]\<^sub>c\<^sub>v \" using assms by metis show ?thesis proof(rule wfG_consI) show \ P; \ \\<^sub>w\<^sub>f \ \ using assms wfT_wf by auto show \atom x \ \\ using assms by auto show \ P; \ \\<^sub>w\<^sub>f b \ using assms wfX_wfY b_of.simps by metis show \ P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>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) #\<^sub>\ \)" 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 "\; \; \ \\<^sub>w\<^sub>f v : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ v \ atom x \ \\<^sub>1 \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f v : b" and "\; \; \ \\<^sub>w\<^sub>f c \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ c\ atom x \ \\<^sub>1 \ \ ; \ ; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f c" and "\; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ \\<^sub>1 \ \; \ \\<^sub>w\<^sub>f \\<^sub>1@\\<^sub>2" and "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ \\ atom x \ \\<^sub>1 \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f \ \True" and "\; \ \\<^sub>w\<^sub>f b \ True" and "\; \; \ \\<^sub>w\<^sub>f ce : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ ce \ atom x \ \\<^sub>1 \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f ce : b" and "\ \\<^sub>w\<^sub>f td \ True" proof(induct arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>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 (\\<^sub>1@\\<^sub>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 "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v1 : b1" using wfV_pairI by auto show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>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 "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>f b" using wfV_conspI by auto show " \; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b" using wfV_conspI by auto show "atom bv \ (\, \, \\<^sub>1 @ \\<^sub>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 \ (\, \, \\<^sub>1 @ \\<^sub>2)\ using wfTI fresh_restrict[of z] using wfG_fresh_x wfX_wfY wfTI fresh_prodN by metis show \ \; \ \\<^sub>w\<^sub>f b \ using wfTI by auto have "\; \; ((z, b, TRUE) #\<^sub>\ \\<^sub>1) @ \\<^sub>2 \\<^sub>w\<^sub>f c " proof(rule wfTI(5)[of "(z, b, TRUE) #\<^sub>\ \\<^sub>1" ]) show \(z, b, TRUE) #\<^sub>\ \ = ((z, b, TRUE) #\<^sub>\ \\<^sub>1) @ (x, b', c') #\<^sub>\ \\<^sub>2\ using wfTI by auto show \atom x \ c\ using wfTI \.fresh \x \ z\ by auto show \atom x \ (z, b, TRUE) #\<^sub>\ \\<^sub>1\ using wfTI \x \ z\ fresh_GCons by simp qed thus \ \; \; (z, b, TRUE) #\<^sub>\ \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f c \ by auto qed next case (wfC_eqI \ \ \ e1 b e2) show ?case proof show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f e1 : b " using wfC_eqI c.fresh fresh_Nil by auto show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>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 "\\<^sub>1=GNil") case True then show ?thesis using wfG_cons1I wfG_consI by auto next case False then obtain G'::\ where *:"(x1, b1, c1) #\<^sub>\ G' = \\<^sub>1" using GCons_eq_append_conv wfG_cons1I by auto hence **:"G=G' @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons1I by auto have " \; \ \\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\ (G' @ \\<^sub>2)" proof(rule Wellformed.wfG_cons1I) show \c1 \ {TRUE, FALSE}\ using wfG_cons1I by auto show \atom x1 \ G' @ \\<^sub>2\ using wfG_cons1I(4) ** fresh_restrict by metis have " atom x \ G'" using wfG_cons1I * using fresh_GCons by blast thus \ \; \ \\<^sub>w\<^sub>f G' @ \\<^sub>2 \ using wfG_cons1I(3)[of G'] ** by auto have "atom x \ c1 \ atom x \ (x1, b1, TRUE) #\<^sub>\ G'" using fresh_GCons \atom x \ \\<^sub>1\ * by auto thus \ \; \; (x1, b1, TRUE) #\<^sub>\ G' @ \\<^sub>2 \\<^sub>w\<^sub>f c1 \ using wfG_cons1I(6)[of "(x1, b1, TRUE) #\<^sub>\ G'"] ** * wfG_cons1I by auto show \ \; \ \\<^sub>w\<^sub>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 "\\<^sub>1=GNil") case True then show ?thesis using wfG_cons2I wfG_consI by auto next case False then obtain G'::\ where *:"(x1, b1, c1) #\<^sub>\ G' = \\<^sub>1" using GCons_eq_append_conv wfG_cons2I by auto hence **:"G=G' @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons2I by auto have " \; \ \\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\ (G' @ \\<^sub>2)" proof(rule Wellformed.wfG_cons2I) show \c1 \ {TRUE, FALSE}\ using wfG_cons2I by auto show \atom x1 \ G' @ \\<^sub>2\ using wfG_cons2I ** fresh_restrict by metis have " atom x \ G'" using wfG_cons2I * using fresh_GCons by blast thus \ \; \ \\<^sub>w\<^sub>f G' @ \\<^sub>2 \ using wfG_cons2I ** by auto show \ \; \ \\<^sub>w\<^sub>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 "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ e \ atom x \ \\<^sub>1 \ atom x \ \ \ \; \; \; \\<^sub>1@\\<^sub>2 ; \ \\<^sub>w\<^sub>f e : b" and "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ True" and "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ True" and "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ True" and "\ \\<^sub>w\<^sub>f (\::\) \ True " and "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ \\<^sub>1 \ atom x \ \ \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" proof(induct arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>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 \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto have "atom bv \ \\<^sub>1 @ \\<^sub>2" using wfE_appPI fresh_prodN fresh_restrict by metis thus \atom bv \ (\, \, \, \\<^sub>1 @ \\<^sub>2, \, b', v, (b_of \)[bv::=b']\<^sub>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 \ \; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>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 "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f \" using wfD_cons fresh_DCons by metis show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>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) #\<^sub>\ \) \" and "atom x \ \" shows "\; \; \ \\<^sub>w\<^sub>f \" using wf_restrict1(4)[of \ \ "((x, b, c) #\<^sub>\ \)" \ GNil x "b" "c" \] assms fresh_GNil append_g.simps by auto lemma wfG_intros2: assumes "wfC P \ ((x,b,TRUE) #\<^sub>\\) c" shows "wfG P \ ((x,b,c) #\<^sub>\\)" proof - have "wfG P \ ((x,b,TRUE) #\<^sub>\\)" 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 "\; \; \ \\<^sub>w\<^sub>f v : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f v : b" and "\; \; \ \\<^sub>w\<^sub>f c \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f c" and "\; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f ts \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f ts" and "\\<^sub>w\<^sub>f P \ True " and "\; \ \\<^sub>w\<^sub>f b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ \\<^sub>w\<^sub>f b" and "\; \; \ \\<^sub>w\<^sub>f ce : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f ce : b" and "\ \\<^sub>w\<^sub>f td \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' \\<^sub>w\<^sub>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 \ \' ; \ ; \ \\<^sub>w\<^sub>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 \\' ; \ ; \ \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by auto show "\' ; \ \\<^sub>w\<^sub>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 \ \\<^sub>w\<^sub>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 \ \\<^sub>w\<^sub>f \' \ using wfB_appI by auto show \AF_typedef_poly s bv dclist \ set \'\ using wfB_appI by auto show "\' ; \ \\<^sub>w\<^sub>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 "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b" and "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f s : b" and "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and "\ \\<^sub>w\<^sub>f (\::\) \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' \\<^sub>w\<^sub>f (\::\)" and "\; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f ftq \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ \\<^sub>w\<^sub>f ftq" and "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>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 \ \' \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \' ; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI wf_theta_weakening1 by auto show \atom bv \ (\, \', \, \, \, b', v, (b_of \)[bv::=b']\<^sub>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 \ \' ; \ ; \ \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI wf_theta_weakening1 by auto qed next case (wfS_matchI \ \ \ v tid dclist \ \ cs b) show ?case proof show \ \' ; \ ; \ \\<^sub>w\<^sub>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 \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto show \ \' \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto show \\'; \; \; \; \; tid; dclist \\<^sub>w\<^sub>f cs : b \ using wfS_matchI by auto qed next case (wfS_varI \ \ \ \ v u \ \ b s) show ?case proof show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_varI wf_theta_weakening1 by auto show \ \' ; \ ; \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI wf_theta_weakening1 by auto show \atom u \ (\, \', \, \, \, \, v, b)\ using wfS_varI by auto show \ \' ; \ ; \ ; \ ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>f s : b \ using wfS_varI by auto qed next case (wfS_letI \ \ \ \ \ e b' x s b) show ?case proof show \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto show \ \' ; \ ; \ ; (x, b', TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_letI by auto show \ \' ; \ ; \ \\<^sub>w\<^sub>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 \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I by auto show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_let2I wf_theta_weakening1 by auto show \ \' ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>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) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_branchI by auto show \atom x \ (\, \', \, \, \, \, \)\ using wfS_branchI by auto show \ \' ; \ ; \ \\<^sub>w\<^sub>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 "\' ; \ \\<^sub>w\<^sub>f ft" using wfPhi_consI by auto show "\' \\<^sub>w\<^sub>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) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_assertI wf_theta_weakening1 by auto show \ \' ; \ ; \ \\<^sub>w\<^sub>f c \ using wfS_assertI wf_theta_weakening1 by auto show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_assertI wf_theta_weakening1 by auto have "atom x \ \'" using wf_supp(6)[OF \\\<^sub>w\<^sub>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 "\\<^sub>w\<^sub>f \" shows "\ \\<^sub>w\<^sub>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 "\' \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>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 "\\<^sub>w\<^sub>f \" shows "supp t = {}" proof - have "\ ; {||} ; GNil \\<^sub>w\<^sub>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 "\\<^sub>w\<^sub>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 "\\<^sub>w\<^sub>f P" shows "P ; {|bva|} ; GNil \\<^sub>w\<^sub>f dclist" proof - have *:"P \\<^sub>w\<^sub>f AF_typedef_poly tyid bva dclist" using lookup_wfTD assms by simp obtain bv lst where *:"P ; {|bv|} ; GNil \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>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 \\\<^sub>w\<^sub>f P\ theta_flip_eq) then show ?thesis using "*" "**" \(bv \ c) \ P ; (bv \ c) \ {|bv|} ; (bv \ c) \ GNil \\<^sub>w\<^sub>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 "\\<^sub>w\<^sub>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, hide_lams) Un_empty_left Un_empty_right pure_supp subset_Un_eq subset_singletonD supp_list_member) + 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 "\\<^sub>w\<^sub>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]\<^sub>c\<^sub>v) #\<^sub>\ \)" shows "(\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v = (\'[x::=v]\<^sub>\\<^sub>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]\<^sub>c\<^sub>v) #\<^sub>\ \) \ atom x' \ (G @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" using wfG_elims(2) using GCons.prems append_g.simps by metis hence "atom x \ atom_dom ((x', b', c') #\<^sub>\ 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]\<^sub>c\<^sub>v) \))[x::=v]\<^sub>\\<^sub>v = (GCons (x', b', c') (G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \)))[x::=v]\<^sub>\\<^sub>v" by auto also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \))[x::=v]\<^sub>\\<^sub>v)" using subst_gv.simps \x\x'\ by simp also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (G[x::=v]\<^sub>\\<^sub>v @ \)" using GCons wfg by blast also have "... = ((x', b', c') #\<^sub>\ G)[x::=v]\<^sub>\\<^sub>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 \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>f dclist" and "(dc,t) \ set dclist" shows "\; \; GNil \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>f t" proof - have "P \\<^sub>w\<^sub>f AF_typedef tid dclist" using lookup_wfTD assms by auto hence "P ; {||} ; GNil \\<^sub>w\<^sub>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) (* 31 ms *) 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 "\\<^sub>w\<^sub>f P" shows "b1[bva1::=b]\<^sub>b\<^sub>b = b2[bva2::=b]\<^sub>b\<^sub>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 " \; \; \ \\<^sub>w\<^sub>f \" and "\ ; \ ; \ ; (x1, b, c1) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s1 : b_of t1" shows "\ ; \ ; \ ; (x2, b, c2) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>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 "\\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>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) #\<^sub>\ \) ; \ \\<^sub>w\<^sub>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) #\<^sub>\ ((x2 \ x1) \ \)) ; \ \\<^sub>w\<^sub>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\) (* 125 ms *) then have "x2 \ x1 \ s2 = (x2 \ x1) \ s1 \ atom x2 \ s1 \ b_of t2 = (x2 \ x1) \ b_of t1" by (metis b_of.eqvt) (* 0.0 ms *) 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) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s1 : b_of t1\ \\ ; \ ; \ ; (x2 \ x1) \ ((x1, b, c1) #\<^sub>\ \) ; \ \\<^sub>w\<^sub>f (x2 \ x1) \ s1 : (x2 \ x1) \ b_of t1\ \atom x1 \ \\ \atom x2 \ \\) (* 593 ms *) qed qed section \Lookup\ lemma wf_not_in_prefix: assumes "\ ; B \\<^sub>w\<^sub>f (\'@(x,b1,c1) #\<^sub>\\)" 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 #\<^sub>\ \') @ (x, b1, c1) #\<^sub>\ \ = ((x',b',c') #\<^sub>\(\'@ ((x, b1, c1) #\<^sub>\ \)))" by simp hence "atom x' \ (\'@(x,b1,c1) #\<^sub>\\)" using wfG_elims(2) GCons by metis moreover have "\ ; B \\<^sub>w\<^sub>f (\' @ (x, b1, c1) #\<^sub>\ \)" using GCons wfG_elims * by metis ultimately have "atom x' \ atom_dom (\'@(x,b1,c1) #\<^sub>\\)" 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 \\<^sub>w\<^sub>f \' @ (x, b1, c1) #\<^sub>\ \\ atom_dom.simps by auto qed lemma lookup_inside_wf[simp]: assumes "\ ; B \\<^sub>w\<^sub>f (\'@(x,b1,c1) #\<^sub>\\)" shows "Some (b1,c1) = lookup (\'@(x,b1,c1) #\<^sub>\\) 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 "\; \ \\<^sub>w\<^sub>f \'" and "\; \ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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: "\ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f v : b \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f v : b" and "\; \; \ \\<^sub>w\<^sub>f c \\ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f c" and "\; \ \\<^sub>w\<^sub>f \ \\ |\| \' \ \; \' \\<^sub>w\<^sub>f \ " and "\; \; \ \\<^sub>w\<^sub>f \ \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f ts \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f ts" and "\\<^sub>w\<^sub>f P \ True " and "wfB \ \ b \ \ |\| \' \ wfB \ \' b" and "\; \; \ \\<^sub>w\<^sub>f ce : b \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f ce : b" and "\ \\<^sub>w\<^sub>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 \ \ ; \' \\<^sub>w\<^sub>f b \ using wfV_conspI by auto show \atom bv \ (\, \', \, b, v)\ using fresh_prodN wfV_conspI by auto thus \ \; \' ; \ \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by simp qed next case (wfTI z \ \ \ b c) show ?case proof show "atom z \ (\, \', \)" using wfTI by auto show "\; \' \\<^sub>w\<^sub>f b " using wfTI by auto show "\; \' ; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>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 "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f e : b" and "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f s : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and "\ \\<^sub>w\<^sub>f (\::\) \ True" and "\; \; \ \\<^sub>w\<^sub>f \ \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ |\| \' \ \ ; \ ; \' \\<^sub>w\<^sub>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 "\ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f AE_appP f1 b' v1 : (b_of \1)[bv1::=b']\<^sub>b" proof show "\ \\<^sub>w\<^sub>f \" using wfE_appPI by auto show "\; \' ; \ \\<^sub>w\<^sub>f \ " using wfE_appPI by auto show "\; \' \\<^sub>w\<^sub>f b' " using wfE_appPI wb_b_weakening1 by auto thus " atom bv1 \ (\, \, \', \, \, b', v1, (b_of \1)[bv1::=b']\<^sub>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 "\; \' ; \ \\<^sub>w\<^sub>f v1 : b1[bv1::=b']\<^sub>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 \ \ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto show \ \ ; \ ; \' ; (x, b', TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_letI by auto show \ \; \' ; \ \\<^sub>w\<^sub>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 "\; \' \\<^sub>w\<^sub>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) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ " using wfFTI wb_b_weakening1 by auto show "\ \\<^sub>w\<^sub>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) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wb_b_weakening1 wfS_assertI by simp show \ \; \' ; \ \\<^sub>w\<^sub>f c \ using wb_b_weakening1 wfS_assertI by simp show \ \; \' ; \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>f \" shows "\; \' \\<^sub>w\<^sub>f \ " using wb_b_weakening assms by auto lemma wfT_b_weakening: fixes \::\ and \::\ and \::\ assumes "\ |\| \'" and "\; \; \ \\<^sub>w\<^sub>f \" shows "\; \' ; \ \\<^sub>w\<^sub>f \ " using wb_b_weakening assms by auto lemma wfB_subst_wfB: fixes \::\ and b'::b and b::b assumes "\ ; {|bv|} \\<^sub>w\<^sub>f b" and "\; \ \\<^sub>w\<^sub>f b'" shows "\; \ \\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>b " using assms proof(nominal_induct b rule:b.strong_induct) case B_int hence "\ ; {||} \\<^sub>w\<^sub>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 "\ ; {||} \\<^sub>w\<^sub>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 " \; \ \\<^sub>w\<^sub>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 "\ ; {||} \\<^sub>w\<^sub>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 "\ ; {||} \\<^sub>w\<^sub>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 (* 781 ms *) 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|} \\<^sub>w\<^sub>f b" using wfB_elims by metis show ?case unfolding subst_b_simps proof show "\\<^sub>w\<^sub>f \ " using B_app wfX_wfY by metis show "\ ; \ \\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>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) #\<^sub>\ GNil \\<^sub>w\<^sub>f \" and "\; \ \\<^sub>w\<^sub>f b'" shows "\; \ \\<^sub>w\<^sub>f (b_of \)[bv::=b']\<^sub>b\<^sub>b " proof - obtain b where "\ ; {|bv|} \\<^sub>w\<^sub>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) #\<^sub>\\))" and "wfG \ \ (((x,b,c) #\<^sub>\\))" 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) #\<^sub>\\))" and "wfG \ \ (\'@((x,b,c) #\<^sub>\\))" 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) #\<^sub>\ \)" 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 "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b" and "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ \\<^sub>w\<^sub>f s : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and "\ \\<^sub>w\<^sub>f (\::\) \ True" and "\; \; \ \\<^sub>w\<^sub>f \ \ True" and "\ ; \ \\<^sub>w\<^sub>f ftq \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' \\<^sub>w\<^sub>f ftq" and "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ \\<^sub>w\<^sub>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 \ \ \\<^sub>w\<^sub>f \' \ using wfE_appPI by auto show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto show \atom bv \ (\', \, \, \, \, b', v, (b_of \)[bv::=b']\<^sub>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 \ \; \; \ \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>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 \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfS_varI by simp show \ \; \; \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI by simp show \atom u \ (\', \, \, \, \, \, v, b)\ using wfS_varI by simp show \ \ ; \' ; \ ; \ ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>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) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_assertI by auto next show \ \; \; \ \\<^sub>w\<^sub>f c \ using wfS_assertI by auto next show \ \; \; \ \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>f b \ using wfFTI by auto next show \supp c \ {atom x}\ using wfFTI by auto next show \ \ ; B ; (x, b, c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ \ using wfFTI by auto next show \ \ \\<^sub>w\<^sub>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) #\<^sub>\ GNil \\<^sub>w\<^sub>f \a'" and "(AF_fun_typ x b c \' s') = (AF_fun_typ xa b ca \a' sa')" shows "\; \; (x, b, c) #\<^sub>\ GNil \\<^sub>w\<^sub>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]\<^sub>c\<^sub>v = ca[xa::=V_var cb]\<^sub>c\<^sub>v \ \'[x::=V_var cb]\<^sub>\\<^sub>v = \a'[xa::=V_var cb]\<^sub>\\<^sub>v" using assms \.eq_iff Abs1_eq_iff_all by auto have **: "\; \; (xa \ cb ) \ ((xa, b, ca) #\<^sub>\ GNil) \\<^sub>w\<^sub>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) #\<^sub>\ GNil) \\<^sub>w\<^sub>f (x \ cb ) \ \'" proof - have "(xa \ cb ) \ xa = (x \ cb ) \ x" using xf by auto hence "(x \ cb ) \ ((x, b, c) #\<^sub>\ GNil) = (xa \ cb ) \ ((xa, b, ca) #\<^sub>\ 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 \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" shows "\ ; B ; (x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ supp s \ { atom x } \ supp B" proof - obtain xa and ca and sa and \' where *:"\ ; B \\<^sub>w\<^sub>f b \ (\ \\<^sub>w\<^sub>f \ ) \ supp sa \ {atom xa} \ supp B \ (\ ; B ; (xa, b, ca) #\<^sub>\ GNil \\<^sub>w\<^sub>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) #\<^sub>\GNil \\<^sub>w\<^sub>f \" using wfT_fun_return_t by metis moreover have " (\ \\<^sub>w\<^sub>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 "\ ; \ \\<^sub>w\<^sub>f (AF_fun_typ_none (AF_fun_typ x b c \ s))" shows "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ supp s \ { atom x } " proof - have *:"\ ; \ ; {||} \\<^sub>w\<^sub>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 "\ ; \ \\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))" shows "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ \ ; \ ; {|bv|} \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" proof - obtain bv1 ft1 where *:"\ ; \ ; {|bv1|} \\<^sub>w\<^sub>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, hide_lams) wfFT_wf_aux) + 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|} \\<^sub>w\<^sub>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|} \\<^sub>w\<^sub>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 "\ ; \ \\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))" shows "\ ; {| bv |} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" and "set \ \ set \'" and "\ \\<^sub>w\<^sub>f \'" shows "\ ; {||} ; (x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>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 " \ \\<^sub>w\<^sub>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 *:"\ ; \'' \\<^sub>w\<^sub>f AF_fun_typ_none (AF_fun_typ x b c \ s) " using wfPhi_elims(2)[OF PConsNone(3)] by metis hence "\ ; \'' ; {||} \\<^sub>w\<^sub>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 " \ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" shows "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" shows "supp b = {}" proof - have "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" shows "supp c \ { atom x }" proof - have "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" and "set \ \ set \'" and "\ \\<^sub>w\<^sub>f \'" shows "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \' \ \ ; \' ; {|bv|} \\<^sub>w\<^sub>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 " \ \\<^sub>w\<^sub>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 *:"\ ; \'' \\<^sub>w\<^sub>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 " \ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" shows "\ ; {| bv |} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>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 *:"\ ; \' \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" shows "supp b \ supp bv" proof - have "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>f \" shows "supp c \ { atom x, atom bv }" proof - have "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>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 "\ \\<^sub>w\<^sub>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|} \\<^sub>w\<^sub>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]\<^sub>\\<^sub>v = \2'[x2::=v]\<^sub>\\<^sub>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]\<^sub>s\<^sub>v = s2'[x2::=v]\<^sub>s\<^sub>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]\<^sub>\\<^sub>v = \2'[x2::=v]\<^sub>\\<^sub>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]\<^sub>b\<^sub>b | c1[bv1::=b]\<^sub>c\<^sub>b \ = \ x2 : b2[bv2::=b]\<^sub>b\<^sub>b | c2[bv2::=b]\<^sub>c\<^sub>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]\<^sub>\\<^sub>b = \ x2 : ((bv2 \ c) \ b2) | ((bv2 \ c) \ c2) \[c::=b]\<^sub>\\<^sub>b" by metis hence "\ x1 :((bv1 \ c) \ b1)[c::=b]\<^sub>b\<^sub>b | ((bv1 \ c) \ c1)[c::=b]\<^sub>c\<^sub>b \ = \ x2 : ((bv2 \ c) \ b2)[c::=b]\<^sub>b\<^sub>b | ((bv2 \ c) \ c2)[c::=b]\<^sub>c\<^sub>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]\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v = \2'[bv2::=b]\<^sub>\\<^sub>b[x2::=v]\<^sub>\\<^sub>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]\<^sub>\\<^sub>b = [[atom x2]]lst. \2'[bv2::=b]\<^sub>\\<^sub>b" apply(rule lst_snd[of _ "c1[bv1::=b]\<^sub>c\<^sub>b" _ _ "c2[bv2::=b]\<^sub>c\<^sub>b"]) apply(rule lst_fst[of _ _ "s1'[bv1::=b]\<^sub>s\<^sub>b" _ _ "s2'[bv2::=b]\<^sub>s\<^sub>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]\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v = s2'[bv2::=b]\<^sub>s\<^sub>b[x2::=v]\<^sub>s\<^sub>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]\<^sub>s\<^sub>b = [[atom x2]]lst. s2'[bv2::=b]\<^sub>s\<^sub>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]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v = \[x::=v]\<^sub>\\<^sub>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: "\; \; \ \\<^sub>w\<^sub>f v : b \ \; \ \\<^sub>w\<^sub>f b" and "\; \; \ \\<^sub>w\<^sub>f c \ True" and "\; \ \\<^sub>w\<^sub>f \ \ True" and wfT_wfB: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f b_of \ " and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f \ \ True" and "\; \ \\<^sub>w\<^sub>f b \ True" and wfCE_wfB: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \; \ \\<^sub>w\<^sub>f b" and "\ \\<^sub>w\<^sub>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: "\; \; \; \; \ \\<^sub>w\<^sub>f e : b \ \; \ \\<^sub>w\<^sub>f b" and wfS_wfB: "\; \; \; \; \ \\<^sub>w\<^sub>f s : b \ \; \ \\<^sub>w\<^sub>f b" and wfCS_wfB: "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \ \\<^sub>w\<^sub>f b" and wfCSS_wfB: "\; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \ \\<^sub>w\<^sub>f b" and "\ \\<^sub>w\<^sub>f \ \ True" and "\; \; \ \\<^sub>w\<^sub>f \ \ True" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ |\| \' \ \ ; \ ; \' \\<^sub>w\<^sub>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) #\<^sub>\ GNil \\<^sub>w\<^sub>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) #\<^sub>\ GNil \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>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: "\; \; \ \\<^sub>w\<^sub>f v : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f v : b" and wfC_weakening: "\; \; \ \\<^sub>w\<^sub>f c \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f c" and "\; \ \\<^sub>w\<^sub>f \ \ True" and wfT_weakening: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f P \ True " and wfB_weakening: "wfB \ \ b \ \ |\| \' \ wfB \ \ b" and wfCE_weakening: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f ce : b" and "\ \\<^sub>w\<^sub>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) (* This proof pattern is used elsewhere when proving weakening for typing predicates *) show ?case proof show \atom z \ (\, \, \')\ using wfTI by auto show \ \; \ \\<^sub>w\<^sub>f b \ using wfTI by auto have *:"toSet ((z, b, TRUE) #\<^sub>\ \) \ toSet ((z, b, TRUE) #\<^sub>\ \')" using toSet.simps wfTI by auto thus \ \; \; (z, b, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>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 \ \; \ \\<^sub>w\<^sub>f b \ using wfV_conspI by auto show \atom bv \ (\, \, \', b, v)\ using wfV_conspI by simp show \ \; \; \' \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>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: "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ \\<^sub>w\<^sub>f e : b" and wfS_weakening: "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ \\<^sub>w\<^sub>f s : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and "\ \\<^sub>w\<^sub>f (\::\) \ True" and wfD_weakning: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>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 \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto show \atom bv \ (\, \, \, \', \, b', v, (b_of \)[bv::=b']\<^sub>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 \ \; \; \' \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI wf_weakening1 by auto qed next case (wfS_letI \ \ \ \ \ e b' x s b) show ?case proof(rule) show \ \ ; \ ; \ ; \' ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto have "toSet ((x, b', TRUE) #\<^sub>\ \) \ toSet ((x, b', TRUE) #\<^sub>\ \')" using wfS_letI by auto thus \ \ ; \ ; \ ; (x, b', TRUE) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>f s : b \ using wfS_letI by (meson wfG_cons wfG_cons_TRUE wfS_wf) show \ \; \; \' \\<^sub>w\<^sub>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 \ \ ; \ ; \ ; \' ; \ \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I by auto show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wfS_let2I wf_weakening1 by auto have "toSet ((x, b_of \, TRUE) #\<^sub>\ \) \ toSet ((x, b_of \, TRUE) #\<^sub>\ \')" using wfS_let2I by auto thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>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 "\; \; \' \\<^sub>w\<^sub>f \ " using wfS_varI wf_weakening1 by auto show "\; \; \' \\<^sub>w\<^sub>f v : b_of \ " using wfS_varI wf_weakening1 by auto show "atom u \ (\, \, \, \', \, \, v, b)" using wfS_varI by auto show "\ ; \ ; \ ; \' ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>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) #\<^sub>\ \) \ toSet ((x, b_of \, TRUE) #\<^sub>\ \')" using wfS_branchI by auto thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>f s : b \ using wfS_branchI by (meson wfG_cons wfG_cons_TRUE wfS_wf) show \atom x \ (\, \, \, \', \, \', \)\ using wfS_branchI by auto show \ \; \; \' \\<^sub>w\<^sub>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 \ \; \; \' \\<^sub>w\<^sub>f c \ using wfS_assertI wf_weakening1 by auto have "\; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \'" proof(rule wfG_consI) show \ \; \ \\<^sub>w\<^sub>f \' \ using wfS_assertI by auto show \atom x \ \'\ using wfS_assertI by auto show \ \; \ \\<^sub>w\<^sub>f B_bool \ using wfS_assertI wfB_boolI wfX_wfY by metis have "\; \ \\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\ \'" proof show "(TRUE) \ {TRUE, FALSE}" by auto show \ \; \ \\<^sub>w\<^sub>f \' \ using wfS_assertI by auto show \atom x \ \'\ using wfS_assertI by auto show \ \; \ \\<^sub>w\<^sub>f B_bool \ using wfS_assertI wfB_boolI wfX_wfY by metis qed thus \ \; \; (x, B_bool, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>f c \ using wf_weakening1(2)[OF \ \; \; \' \\<^sub>w\<^sub>f c \ \ \; \ \\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\ \' \] by force qed thus \ \; \; \; (x, B_bool, c) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>f s : b \ using wfS_assertI by fastforce show \ \; \; \' \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f v : b" and "atom y \ \" and "\; \; ((y,b',TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f c" shows "\; \; (y,b',c) #\<^sub>\\ \\<^sub>w\<^sub>f v : b" proof - have "wfG \ \ ((y,b',c) #\<^sub>\\)" using wfG_intros2 assms by auto moreover have "toSet \ \ toSet ((y,b',c) #\<^sub>\\)" using toSet.simps by auto ultimately show ?thesis using wf_weakening using assms(1) by blast qed lemma wfG_cons_weakening: fixes \'::\ assumes "\; \ \\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\ \)" and "\; \ \\<^sub>w\<^sub>f \'" and "toSet \ \ toSet \'" and "atom x \ \'" shows "\; \ \\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\ \')" proof(cases "c \ {TRUE,FALSE}") case True then show ?thesis using wfG_wfB wfG_cons2I assms by auto next case False hence *:"\; \ \\<^sub>w\<^sub>f \ \ atom x \ \ \ \; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c" using wfG_elims(2)[OF assms(1)] by auto have a1:"\; \ \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \'" using wfG_wfB wfG_cons2I assms by simp moreover have a2:"toSet ((x, b, TRUE) #\<^sub>\ \ ) \ toSet ((x, b, TRUE) #\<^sub>\ \')" using toSet.simps assms by blast moreover have " \; \ \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \'" proof show "(TRUE) \ {TRUE, FALSE}" by auto show "\; \ \\<^sub>w\<^sub>f \'" using assms by auto show "atom x \ \'" using assms by auto show "\; \ \\<^sub>w\<^sub>f b" using assms wfG_elims by metis qed hence " \; \; (x, b, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "\; \ \\<^sub>w\<^sub>f \'" and "toSet \ \ toSet \'" and "atom z \ \'" shows "\; \; \' \\<^sub>w\<^sub>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 \ \; \ \\<^sub>w\<^sub>f b \ using assms wfT_elims by metis show \ \; \; (z, b, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>f c \ proof - have *:"\; \; (z,b,TRUE) #\<^sub>\\ \\<^sub>w\<^sub>f c" using wfT_wfC fresh_weakening assms by auto moreover have a1:"\; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \'" using wfG_cons2I assms \\; \ \\<^sub>w\<^sub>f b\ by simp moreover have a2:"toSet ((z, b, TRUE) #\<^sub>\ \ ) \ toSet ((z, b, TRUE) #\<^sub>\ \')" using toSet.simps assms by blast moreover have " \; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \' " proof show "(TRUE) \ {TRUE, FALSE}" by auto show "\; \ \\<^sub>w\<^sub>f \'" using assms by auto show "atom z \ \'" using assms by auto show "\; \ \\<^sub>w\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f \" and "\; \' \\<^sub>w\<^sub>f \'" and "toSet \ \ toSet \'" and "\ |\| \'" shows "\; \' ; \' \\<^sub>w\<^sub>f \" using wb_b_weakening assms wfT_weakening by metis lemma wfT_weakening_nil: fixes \::\ and \'::\ and \::\ assumes "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" and "\; \' \\<^sub>w\<^sub>f \'" shows "\; \' ; \' \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>f G" shows "supp \ = {}" and "\[x::=v]\<^sub>\\<^sub>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]\<^sub>\\<^sub>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 "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' \\<^sub>w\<^sub>f e : b" and "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' \\<^sub>w\<^sub>f s : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' ; tid ; dclist \\<^sub>w\<^sub>f css : b" and "\ \\<^sub>w\<^sub>f (\::\) \ True" and "\; \; \ \\<^sub>w\<^sub>f \ \ True" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>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']\<^sub>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 \ \; \; \ \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI by auto qed next case (wfE_mvarI \ \ \ \ \ u \) show ?case proof show \ \ \\<^sub>w\<^sub>f \ \ using wfE_mvarI by auto show \ \; \; \ \\<^sub>w\<^sub>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 \ \; \; \; \; \' \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto have "\; \ \\<^sub>w\<^sub>f (x, b', TRUE) #\<^sub>\ \" using wfG_cons2I wfX_wfY wfS_letI by metis hence "\; \; (x, b', TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) wfS_letI by force thus \ \ ; \ ; \ ; (x, b', TRUE) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>f s : b \ using wfS_letI by metis show \ \; \; \ \\<^sub>w\<^sub>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) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" proof(rule wf_weakening2(6)) show \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfS_assertI by auto next show \ \; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \ \ using wfS_assertI wfX_wfY by metis next show \toSet \ \ toSet ((x, B_bool, c) #\<^sub>\ \)\ using wfS_assertI by auto qed thus \ \; \; \; (x, B_bool, c) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>f s : b \ using wfS_assertI wfX_wfY by metis next show \ \; \; \ \\<^sub>w\<^sub>f c \ using wfS_assertI by auto next show \ \; \; \ \\<^sub>w\<^sub>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 \ \; \; \; \; \' \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I by auto show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfS_let2I by auto have "\; \ \\<^sub>w\<^sub>f (x, b_of \, TRUE) #\<^sub>\ \" using wfG_cons2I wfX_wfY wfS_let2I by metis hence "\; \; (x, b_of \, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) wfS_let2I by force thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>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 \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfS_varI by auto show \ \; \; \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI by auto show \atom u \ (\, \, \, \, \', \, v, b)\ using wfS_varI setD.simps by auto have "\; \; \ \\<^sub>w\<^sub>f (u, \) #\<^sub>\ \'" using wfS_varI wfD_cons setD.simps u_fresh_d by metis thus \ \ ; \ ; \ ; \ ; (u, \) #\<^sub>\ \' \\<^sub>w\<^sub>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 \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfS_assignI by auto show \ \ \\<^sub>w\<^sub>f \ \ using wfS_assignI by auto show \ \; \; \ \\<^sub>w\<^sub>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 "\; \ \\<^sub>w\<^sub>f (x, b_of \, TRUE) #\<^sub>\ \" using wfG_cons2I wfX_wfY wfS_branchI by metis hence "\; \; (x, b_of \, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) wfS_branchI by force thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>f s : b \ using wfS_branchI by simp show \ atom x \ (\, \, \, \, \', \, \)\ using wfS_branchI by auto show \ \; \; \ \\<^sub>w\<^sub>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 "\ ; \ ; \ \\<^sub>w\<^sub>f ce : b" and "atom x \ \ " shows "\ ; \ ; ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f (CE_val (V_var x) == ce )" proof - have "\; \ \\<^sub>w\<^sub>f b" using assms wfX_wfB by auto hence wbg: "\; \ \\<^sub>w\<^sub>f \" using wfX_wfY assms by auto show ?thesis proof show *:"\ ; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var x) : b" proof(rule) show "\ ; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f V_var x : b " proof show "\ ; \ \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \ " using wfG_cons2I wfX_wfY assms \\; \ \\<^sub>w\<^sub>f b\ by auto show "Some (b, TRUE) = lookup ((x, b, TRUE) #\<^sub>\ \) x" using lookup.simps by auto qed qed show "\ ; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f ce : b" using assms wf_weakening1(8)[OF assms(1), of "(x, b, TRUE) #\<^sub>\ \ "] * toSet.simps wfX_wfY by (metis Un_subset_iff equalityE) qed qed lemma wfC_e_eq2: fixes e1::ce and e2::ce assumes "\ ; \ ; \ \\<^sub>w\<^sub>f e1 : b" and "\ ; \ ; \ \\<^sub>w\<^sub>f e2 : b" and " \\<^sub>w\<^sub>f \" and "atom x \ \" shows "\; \; (x, b, (CE_val (V_var x)) == e1) #\<^sub>\ \ \\<^sub>w\<^sub>f (CE_val (V_var x)) == e2 " proof(rule wfC_eqI) have *: "\; \ \\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \" proof(rule wfG_cons1I ) show "(CE_val (V_var x) == e1 ) \ {TRUE, FALSE}" by auto show "\; \ \\<^sub>w\<^sub>f \" using assms wfX_wfY by metis show *:"atom x \ \" using assms by auto show "\; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var x) == e1" using wfC_e_eq assms * by auto show "\; \ \\<^sub>w\<^sub>f b" using assms wfX_wfB by auto qed show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var x) : b" using assms * wfCE_valI wfV_varI by auto show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \ \\<^sub>w\<^sub>f e2 : b" proof(rule wf_weakening1(8)) show "\; \; \ \\<^sub>w\<^sub>f e2 : b " using assms by auto show "\; \ \\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \" using * by auto show "toSet \ \ toSet ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \)" 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; \ \\<^sub>w\<^sub>f b_of t \ using wfX_wfY assms by meson have wfg: " P; \ \\<^sub>w\<^sub>f (z1, b_of t, TRUE) #\<^sub>\ \" using assms wfV_wf wfG_cons2I wfX_wfY by (meson wfG_cons_TRUE) show \ P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t z1 \ proof show *: \ P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e \ proof(rule wfC_eqI[where b="base_for_lit l"]) show "P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e : base_for_lit l" using assms wf_intros wf_weakening wfg by (meson wfV_weakening_cons) show "P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ [ l ]\<^sup>v ]\<^sup>c\<^sup>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) #\<^sub>\ \ \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>f \'" and "\ ; {||} \\<^sub>w\<^sub>f (x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil" and "atom zz \ x" shows "\ ; {||} ; (x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ ll ]\<^sup>v ]\<^sup>c\<^sup>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) #\<^sub>\ GNil \\<^sub>w\<^sub>f [ x ]\<^sup>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) #\<^sub>\ GNil \\<^sub>w\<^sub>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) #\<^sub>\ 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 "\ ; \ ; \ \\<^sub>w\<^sub>f v : b" and "atom x \ \ " shows "\ ; \ ; ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>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 "\ ; \ ; \ \\<^sub>w\<^sub>f ce : b" and "atom z \ \" shows "\; \; \ \\<^sub>w\<^sub>f \ z : b | CE_val (V_var z) == ce \" proof show "\; \ \\<^sub>w\<^sub>f b" using wfX_wfB assms by auto show " atom z \ (\, \, \)" using assms wfG_fresh_x wfX_wfY by metis show "\ ; \ ; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>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 ; \ \\<^sub>w\<^sub>f c" and "\ ; B \\<^sub>w\<^sub>f b" and "atom x \ \" shows "\ ; B \\<^sub>w\<^sub>f (x,b,c)#\<^sub>\ \" proof - have " \ ; B \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \" using wfG_cons2I assms wfX_wfY by fast hence " \ ; B ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>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') #\<^sub>\ \' @ (x, b, c) #\<^sub>\ \))" shows "x'\x" proof - have "atom x' \ (\' @ (x, b, c) #\<^sub>\ \)" using assms wfG_elims(2) by blast thus ?thesis using fresh_gamma_append[of "atom x'" \' "(x, b, c) #\<^sub>\ \"] 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') #\<^sub>\\))" shows "replace_in_g (\'@((x,b0,c0') #\<^sub>\\)) x c0 = (\'@((x,b0,c0) #\<^sub>\\))" 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; \ \\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\ (\''@ (x, b0, c0') #\<^sub>\ \ ))" 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) #\<^sub>\ \)" and "wfG P \ (\'' @ (x, b0, c0') #\<^sub>\ \)" shows "supp (\'' @ (x, b0, c0') #\<^sub>\ \) \ supp \ = supp (\'' @ (x, b0, c0) #\<^sub>\ \) \ supp \" using assms proof(induct \'') case GNil have "supp (GNil @ (x, b0, c0') #\<^sub>\ \) \ supp \ = supp ((x, b0, c0') #\<^sub>\ \) \ 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) #\<^sub>\ \)) \ supp \ " using supp_GCons by auto finally have "supp (GNil @ (x, b0, c0') #\<^sub>\ \) \ supp \ = supp (GNil @ (x, b0, c0) #\<^sub>\ \) \ 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 #\<^sub>\ \1) @ (x, b0, c0) #\<^sub>\ \ = (xbc #\<^sub>\ (\1 @ (x, b0, c0) #\<^sub>\ \))" by simp moreover have " (xbc #\<^sub>\ \1) @ (x, b0, c0') #\<^sub>\ \ = (xbc #\<^sub>\ (\1 @ (x, b0, c0') #\<^sub>\ \))" by simp ultimately have "(P; \ \\<^sub>w\<^sub>f \1 @ ((x, b0, c0) #\<^sub>\ \)) \ P; \ \\<^sub>w\<^sub>f \1 @ ((x, b0, c0') #\<^sub>\ \)" 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) #\<^sub>\ \)" and "wfG P \ (\'' @ (x, b, c') #\<^sub>\ \)" shows "atom y \ (\'' @ (x, b, c) #\<^sub>\ \) = atom y \ (\'' @ (x, b, c') #\<^sub>\ \)" 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 \\<^sub>w\<^sub>f v : b' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f v : b'" and wfC_replace_inside: "\; \; G \\<^sub>w\<^sub>f c'' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f c''" and wfG_replace_inside: "\; \ \\<^sub>w\<^sub>f G \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \; \ \\<^sub>w\<^sub>f (\' @ (x, b, c) #\<^sub>\ \) " and wfT_replace_inside: "\; \; G \\<^sub>w\<^sub>f \ \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f P \ True" and "\; \ \\<^sub>w\<^sub>f b \ True" and wfCE_replace_inside: "\ ; \ ; G \\<^sub>w\<^sub>f ce : b' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f ce : b'" and "\ \\<^sub>w\<^sub>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 \ \ ; \ \\<^sub>w\<^sub>f b1 \ using wfV_conspI by auto show *: \ \; \; \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b \ using wfV_conspI by auto moreover have "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c') #\<^sub>\ \" using wfV_wf wfV_conspI by simp ultimately have "atom bv \ \' @ (x, b, c) #\<^sub>\ \" unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfV_conspI by (metis Un_iff fresh_def) thus \atom bv \ (\, \, \' @ (x, b, c) #\<^sub>\ \, b1, v)\ unfolding fresh_prodN using fresh_prodN wfV_conspI by metis qed next case (wfTI z \ \ G b1 c1) show ?case proof show \ \; \ \\<^sub>w\<^sub>f b1 \ using wfTI by auto have "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \" using wfG_consI wfTI wfG_cons wfX_wfY by metis moreover hence *:"wfG \ \ (\' @ (x, b, c) #\<^sub>\ \)" 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) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b c \ c' z,OF *] wfTI wfX_wfY wfG_elims by metis thus \atom z \ (\, \, \' @ (x, b, c) #\<^sub>\ \)\ using wfG_fresh_x[OF *] by auto have "(z, b1, TRUE) #\<^sub>\ G = ((z, b1, TRUE) #\<^sub>\ \') @ (x, b, c') #\<^sub>\ \" using wfTI append_g.simps by metis thus \ \; \; (z, b1, TRUE) #\<^sub>\ \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f c1 \ using wfTI(9)[OF _ wfTI(11)] by fastforce qed next case (wfG_nilI \) hence "GNil = (x, b, c') #\<^sub>\ \" 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) #\<^sub>\ G' = \'" using wfG_cons1I wfG_cons1I(7) GCons_eq_append_conv by auto hence **:" G = G' @ (x, b, c') #\<^sub>\ \" using wfG_cons1I by auto hence " \; \ \\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\ \" using wfG_cons1I by auto have "\; \ \\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\ G' @ (x, b, c) #\<^sub>\ \" proof(rule Wellformed.wfG_cons1I) show "c1 \ {TRUE, FALSE}" using wfG_cons1I by auto show "\; \ \\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\ \ " using wfG_cons1I(3)[of G',OF **] wfG_cons1I by auto show "atom x1 \ G' @ (x, b, c) #\<^sub>\ \" using wfG_cons1I * ** fresh_replace_inside by metis show "\; \; (x1, b1, TRUE) #\<^sub>\ G' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f c1" using wfG_cons1I(6)[of " (x1, b1, TRUE) #\<^sub>\ G'"] wfG_cons1I ** by auto show "\; \ \\<^sub>w\<^sub>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) #\<^sub>\ G' = \'" using wfG_cons2I GCons_eq_append_conv by auto hence **:" G = G' @ (x, b, c') #\<^sub>\ \" using wfG_cons2I by auto moreover have " \; \ \\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\ \" using wfG_cons2I * ** by auto moreover hence "atom x1 \ G' @ (x, b, c) #\<^sub>\ \" using wfG_cons2I * ** fresh_replace_inside by metis ultimately show ?thesis using Wellformed.wfG_cons2I[OF wfG_cons2I(1), of \ \ "G'@ (x, b, c) #\<^sub>\ \" 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 \\<^sub>w\<^sub>f e : b' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \); D \\<^sub>w\<^sub>f e : b'" and "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ True" and "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ True" and "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ True" and "\ \\<^sub>w\<^sub>f \ \ True" and "\; \; G \\<^sub>w\<^sub>f \ \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>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 \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \; \; \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto show *:\ \; \; \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \ using wfE_appPI wf_replace_inside1 by auto moreover have "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c') #\<^sub>\ \" using wfV_wf wfE_appPI by metis ultimately have "atom bv \ \' @ (x, b, c) #\<^sub>\ \" unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfE_appPI Un_iff fresh_def by metis thus \atom bv \ (\, \, \, \' @ (x, b, c) #\<^sub>\ \, \, b', v, (b_of \)[bv::=b']\<^sub>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) #\<^sub>\\)" and "wfC P \ ((x,b,TRUE) #\<^sub>\\) c2" shows "wfC P \ ((x,b,c1) #\<^sub>\\) c2" proof - have "wfC P \ (GNil@((x,b,c1) #\<^sub>\\)) c2" proof(rule wf_replace_inside1(2)) show " P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c2 " using wfG_elim2 assms by auto show \(x, b, TRUE) #\<^sub>\ \ = GNil @ (x, b, TRUE) #\<^sub>\ \\ using append_g.simps by auto show \P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>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') #\<^sub>\\)" shows "wfC \ \ ((x, b', c') #\<^sub>\\) 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 "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom y \ (c,\)" shows "\; \ \\<^sub>w\<^sub>f (y, b, c[z::=V_var y]\<^sub>c\<^sub>v) #\<^sub>\ \" 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]\<^sub>c\<^sub>v" using assms subst_v_c_def by auto ultimately have "\ z : b | c \ = \ y : b | c[z::=V_var y]\<^sub>c\<^sub>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]\<^sub>v) #\<^sub>\\) (c2[z2::=V_var x]\<^sub>v)" (is "wfC P \ ?G ?c") proof - have eq: "\ z2 : b | c2 \ = \ x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \" using type_eq_subst assms fresh_prod3 by simp have eq2: "\ z1 : b | c1 \ = \ x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \" using type_eq_subst assms fresh_prod3 by simp moreover have "wfT P \ \ \ x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \" using assms eq2 by auto moreover hence "wfG P \ ((x,b,c1[z1::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\\)" using wfT_wf_cons fresh_prod3 assms by auto moreover have "wfT P \ \ \ x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \" using assms eq by auto moreover hence "wfC P \ ((x,b,TRUE) #\<^sub>\\) (c2[z2::=V_var x]\<^sub>c\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom x \ \" shows "\; \; (x,b,TRUE)#\<^sub>\\ \\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>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 ]\<^sup>v]\<^sub>v \ = \ z : b | c \" using \.eq_iff Abs1_eq_iff(3)[of x "c[z::=[ x ]\<^sup>v]\<^sub>v" z c] by (metis flip_subst_v type_eq_flip) hence " \; \; \ \\<^sub>w\<^sub>f \ x : b | c[z::=[ x ]\<^sup>v]\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom x \ \" shows "\; \ \\<^sub>w\<^sub>f (x,b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \" proof - have "\; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>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') #\<^sub>\ \)" and "wfG P \ ((x,b,c) #\<^sub>\\)" shows "wfG P \ (\' @ (x, b, c) #\<^sub>\ \)" proof - have "wfC P \ ((x,b,TRUE) #\<^sub>\\) 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') #\<^sub>\ \)" and "wfG P \ (\'@((x,b,c) #\<^sub>\\))" shows "wfG P \ (\' @ (x, b, c) #\<^sub>\ \)" proof - have "wfG P \ ((x,b,c) #\<^sub>\\)" 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') #\<^sub>\ \) t" and "wfG \ \ (\'@((x,b,c) #\<^sub>\\))" shows "wfT \ \ (\' @ (x, b, c) #\<^sub>\ \) t" proof - have "wfG \ \ (((x,b,c) #\<^sub>\\))" using wfG_suffix assms by auto hence "wfC \ \ ((x,b,TRUE) #\<^sub>\\) 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') #\<^sub>\ 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 \\; \ \\<^sub>w\<^sub>f G[x'\c'']\ and "atom x \ G" and \\; \ \\<^sub>w\<^sub>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) #\<^sub>\ \ \\<^sub>w\<^sub>f c'" shows "\; \; (x, b, c') #\<^sub>\ \ \\<^sub>w\<^sub>f c" proof - have *:"\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \" using wfC_wf assms by auto hence "atom x \ \ \ wfG \ \ \ \ \; \ \\<^sub>w\<^sub>f b" using wfG_cons by auto hence " \; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f TRUE " using wfC_trueI wfG_cons2I by simp hence "\; \;(x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c'" using wf_replace_inside1(2)[of \ \ "(x, b, c) #\<^sub>\ \" c' GNil x b c \ TRUE] assms by auto hence "wfG \ \ ((x,b,c') #\<^sub>\\)" using wf_replace_inside1(3)[OF *, of GNil x b c \ c'] by auto moreover have "wfC \ \ ((x,b,TRUE) #\<^sub>\\) c" proof(cases "c \ { TRUE, FALSE }") case True have "\; \ \\<^sub>w\<^sub>f \ \ atom x \ \ \ \; \ \\<^sub>w\<^sub>f b" using wfG_elims(2)[OF *] by auto hence "\; \ \\<^sub>w\<^sub>f (x,b,TRUE) #\<^sub>\ \" 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 \\<^sub>1::\ and \\<^sub>2::\ assumes "wfG P \ (\\<^sub>1@((x,b,c) #\<^sub>\\\<^sub>2))" shows "(\\<^sub>1@((x,b,c) #\<^sub>\\\<^sub>2))[x::=v]\<^sub>\\<^sub>v = \\<^sub>1[x::=v]\<^sub>\\<^sub>v@\\<^sub>2" using assms proof(induct \\<^sub>1 rule: \_induct) case GNil then show ?case using subst_gv.simps by simp next case (GCons x' b' c' G) hence *:"P; \ \\<^sub>w\<^sub>f (x', b', c') #\<^sub>\ (G @ (x, b, c) #\<^sub>\ \\<^sub>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) \\<^sub>2))[x::=v]\<^sub>\\<^sub>v = (GCons (x', b', c') (G @ (GCons (x, b, c) \\<^sub>2)))[x::=v]\<^sub>\\<^sub>v" by auto also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c) \\<^sub>2))[x::=v]\<^sub>\\<^sub>v)" using subst_gv.simps \x\x'\ by simp also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (G[x::=v]\<^sub>\\<^sub>v @ \\<^sub>2)" using GCons * wfG_elims by metis also have "... = ((x', b', c') #\<^sub>\ G)[x::=v]\<^sub>\\<^sub>v @ \\<^sub>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']\<^sub>c\<^sub>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']\<^sub>\\<^sub>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) #\<^sub>\ \1)[x'::=v']\<^sub>\\<^sub>v = ((x1, b1, c1[x'::=v']\<^sub>c\<^sub>v) #\<^sub>\ \1[x'::=v']\<^sub>\\<^sub>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\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)) x" and "x \ x'" and "\; \ \\<^sub>w\<^sub>f (\'@((x',b\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\))" shows "\c'. Some (b,c') = lookup (\'[x'::=v']\<^sub>\\<^sub>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: "\; \; \ \\<^sub>w\<^sub>f v : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \;\\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b" and wfC_subst: "\; \; \ \\<^sub>w\<^sub>f c \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v" and wfG_subst: "\; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" and wfT_subst: "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f \ \True" and "\; \ \\<^sub>w\<^sub>f b \ True " and wfCE_subst: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f ce[x::=v']\<^sub>c\<^sub>e\<^sub>v : b" and "\ \\<^sub>w\<^sub>f td \ True" proof(nominal_induct b and c and \ and \ and ts and \ and b and b and td avoiding: x v' arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>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']\<^sub>v\<^sub>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']\<^sub>\\<^sub>v \\<^sub>w\<^sub>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']\<^sub>v\<^sub>v = (V_var x1) " using subst_vv.simps by auto moreover have "\; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" using wfV_varI by simp moreover obtain c1' where "Some (b1, c1') = lookup \[x::=v']\<^sub>\\<^sub>v x1" using wfV_varI False lookup_subst by metis ultimately show ?thesis using Wellformed.wfV_varI[of \ \ "\[x::=v']\<^sub>\\<^sub>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 \ \ ;\ \\<^sub>w\<^sub>f b \ using wfV_conspI by auto have "atom bv \ \[x::=v']\<^sub>\\<^sub>v" using fresh_subst_gv_if wfV_conspI by metis moreover have "atom bv \ va[x::=v']\<^sub>v\<^sub>v" using wfV_conspI fresh_subst_if by simp ultimately show \atom bv \ (\, \, \[x::=v']\<^sub>\\<^sub>v, b, va[x::=v']\<^sub>v\<^sub>v)\ unfolding fresh_prodN using wfV_conspI by auto show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f va[x::=v']\<^sub>v\<^sub>v : b'[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by auto qed next case (wfTI z \ \ \ b c) have " \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \ z : b | c[x::=v']\<^sub>c\<^sub>v \" proof have \\; \; ((z, b, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ proof(rule wfTI(9)) show \(z, b, TRUE) #\<^sub>\ \ = ((z, b, TRUE) #\<^sub>\ \\<^sub>1) @ (x, b', c') #\<^sub>\ \\<^sub>2\ using wfTI append_g.simps by simp show \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ using wfTI by auto qed thus *:\\; \; (z, b, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ using subst_gv.simps subst_cv.simps wfTI fresh_x_neq by auto have "atom z \ \[x::=v']\<^sub>\\<^sub>v" using fresh_subst_gv_if wfTI by metis moreover have "\; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" using wfTI wfX_wfY wfG_elims subst_gv.simps * by metis ultimately show \atom z \ (\, \, \[x::=v']\<^sub>\\<^sub>v)\ using wfG_fresh_x by metis show \ \; \ \\<^sub>w\<^sub>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']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e1[x::=v']\<^sub>c\<^sub>e\<^sub>v : b " using wfC_eqI subst_dv.simps by auto show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e2[x::=v']\<^sub>c\<^sub>e\<^sub>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) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v = \" using subst_gv.simps by auto moreover have "\; \ \\<^sub>w\<^sub>f \" using wfG_cons1I by auto ultimately show ?thesis by auto next case False have "\\<^sub>1 \ GNil" using wfG_cons1I False by auto then obtain G where "\\<^sub>1 = (y, b, c) #\<^sub>\ G" using GCons_eq_append_conv wfG_cons1I by auto hence *:"\ = G @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons1I by auto hence "((y, b, c) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" using subst_gv.simps False by auto moreover have "\; \ \\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" proof(rule Wellformed.wfG_cons1I) show \c[x::=v']\<^sub>c\<^sub>v \ {TRUE, FALSE}\ using wfG_cons1I subst_c_TRUE_FALSE by auto show \ \; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfG_cons1I * by auto have "\ = (G @ ((x, b', c') #\<^sub>\GNil)) @ \\<^sub>2" using * append_g_assoc by auto hence "atom y \ \\<^sub>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']\<^sub>\\<^sub>v\ using fresh_subst_gv wfG_cons1I by auto have "((y, b, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v = (y, b, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v" using subst_gv.simps subst_cv.simps False by auto thus \ \; \; (y, b, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ using wfG_cons1I(6)[of "(y,b,TRUE) #\<^sub>\G"] * subst_gv.simps wfG_cons1I by fastforce show "\; \ \\<^sub>w\<^sub>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) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v = \" using subst_gv.simps by auto moreover have "\; \ \\<^sub>w\<^sub>f \" using wfG_cons2I by auto ultimately show ?thesis by auto next case False have "\\<^sub>1 \ GNil" using wfG_cons2I False by auto then obtain G where "\\<^sub>1 = (y, b, c) #\<^sub>\ G" using GCons_eq_append_conv wfG_cons2I by auto hence *:"\ = G @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons2I by auto hence "((y, b, c) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" using subst_gv.simps False by auto moreover have "\; \ \\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" proof(rule Wellformed.wfG_cons2I) show \c[x::=v']\<^sub>c\<^sub>v \ {TRUE, FALSE}\ using subst_cv.simps wfG_cons2I by auto show \ \; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfG_cons2I * by auto have "\ = (G @ ((x, b', c') #\<^sub>\GNil)) @ \\<^sub>2" using * append_g_assoc by auto hence "atom y \ \\<^sub>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']\<^sub>\\<^sub>v\ using fresh_subst_gv wfG_cons2I by auto show "\; \ \\<^sub>w\<^sub>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 "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b" and "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \ ;\ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b" and "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dc ; t \\<^sub>w\<^sub>f subst_branchv cs x v' : b" and "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dclist \\<^sub>w\<^sub>f subst_branchlv css x v' : b" and "\ \\<^sub>w\<^sub>f (\::\) \ True " and "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" proof(nominal_induct b and b and b and b and \ and \ and ftq and ft avoiding: x v' arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>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) (* 31 ms *) 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)) (* 78 ms *) 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 "\ \\<^sub>w\<^sub>f \" using wfE_appPI wfX_wfY by metis show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>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']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v1[x::=v']\<^sub>v\<^sub>v : b1[bv1::=b']\<^sub>b " using wfE_appPI wf_subst1 by auto show "\; \ \\<^sub>w\<^sub>f b' " using wfE_appPI by auto have "atom bv1 \ \[x::=v']\<^sub>\\<^sub>v" using fresh_subst_gv_if wfE_appPI by metis moreover have "atom bv1 \ v1[x::=v']\<^sub>v\<^sub>v" using wfE_appPI fresh_subst_if by simp moreover have "atom bv1 \ \[x::=v']\<^sub>\\<^sub>v" using wfE_appPI fresh_subst_dv_if by simp ultimately show "atom bv1 \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, b', v1[x::=v']\<^sub>v\<^sub>v, (b_of \1)[bv1::=b']\<^sub>b)" using wfE_appPI fresh_prodN by metis qed next case (wfE_mvarI \ \ \ \ \ u \) have " \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f (AE_mvar u) : b_of \[x::=v']\<^sub>\\<^sub>v" proof show "\ \\<^sub>w\<^sub>f \ " using wfE_mvarI by auto show "\; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v " using wfE_mvarI by auto show "(u, \[x::=v']\<^sub>\\<^sub>v) \ setD \[x::=v']\<^sub>\\<^sub>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']\<^sub>\\<^sub>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']\<^sub>c\<^sub>v) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \ using wfS_assertI(4)[of "(x2, B_bool, c) #\<^sub>\ \\<^sub>1" x ] wfS_assertI by auto show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ using wfS_assertI wf_subst1 by auto show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_assertI wf_subst1 by auto show \atom x2 \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, c[x::=v']\<^sub>c\<^sub>v, b, s[x::=v']\<^sub>s\<^sub>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']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f LET y = (e[x::=v']\<^sub>e\<^sub>v) IN (s[x::=v']\<^sub>s\<^sub>v) : b2" proof show \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b1 \ using wfS_letI by auto have \ \ ; \ ; \ ; ((y, b1, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \ using wfS_letI(6) wfS_letI append_g.simps by metis thus \ \ ; \ ; \ ; (y, b1, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \ using wfS_letI subst_gv.simps by auto show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_letI by auto show \atom y \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, e[x::=v']\<^sub>e\<^sub>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']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f LET y : \[x::=v']\<^sub>\\<^sub>v = (s1[x::=v']\<^sub>s\<^sub>v) IN (s2[x::=v']\<^sub>s\<^sub>v) : b" proof show \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s1[x::=v']\<^sub>s\<^sub>v : b_of (\[x::=v']\<^sub>\\<^sub>v) \ using wfS_let2I b_of_subst by simp have \ \ ; \ ; \ ; ((y, b_of \, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \ using wfS_let2I append_g.simps by metis thus \ \ ; \ ; \ ; (y, b_of \[x::=v']\<^sub>\\<^sub>v, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \ using wfS_let2I subst_gv.simps append_g.simps using b_of_subst by simp show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_let2I wf_subst1 by metis show \atom y \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, s1[x::=v']\<^sub>s\<^sub>v, b, \[x::=v']\<^sub>\\<^sub>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']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_varI wf_subst1 by auto have "b_of (\[x::=v']\<^sub>\\<^sub>v) = b_of \" using b_of_subst by auto thus \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \[x::=v']\<^sub>\\<^sub>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']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, v[x::=v']\<^sub>v\<^sub>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']\<^sub>\\<^sub>v ; (u, \[x::=v']\<^sub>\\<^sub>v) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>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']\<^sub>\\<^sub>v) \ setD \[x::=v']\<^sub>\\<^sub>v\ using subst_dv_iff wfS_assignI using subst_dv_fst_eq using subst_dv_member by auto show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_assignI by auto show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \[x::=v']\<^sub>\\<^sub>v \ using wfS_assignI b_of_subst wf_subst1 by auto show "\ \\<^sub>w\<^sub>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']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>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']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dclist \\<^sub>w\<^sub>f subst_branchlv cs x v' : b \ using wfS_matchI by simp show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v " using wfS_matchI by auto show "\ \\<^sub>w\<^sub>f \ " using wfS_matchI by auto qed next case (wfS_branchI \ \ \ y \ \ \ s b tid dc) have " \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dc ; \ \\<^sub>w\<^sub>f dc y \ (s[x::=v']\<^sub>s\<^sub>v) : b" proof have \ \ ; \ ; \ ; ((y, b_of \, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \ using wfS_branchI append_g.simps by metis thus \ \ ; \ ; \ ; (y, b_of \, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \ using subst_gv.simps b_of_subst wfS_branchI by simp show \atom y \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>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']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>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 "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c0[z0::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \" and "wfV \ \ \ v b" shows "\; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>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) #\<^sub>\\))" and "x \ x1" shows "\c1'. (x1,b1,c1') \ toSet ((\'[x::=v]\<^sub>\\<^sub>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]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using append_g_toSetU by auto then show ?thesis by auto next case rhs hence "(x1,b1,c1) \ toSet (\'[x::=v]\<^sub>\\<^sub>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) #\<^sub>\\))" and "wfG \ \ (\'@((x,b,c) #\<^sub>\\))" and "x \ x1" shows "\c1'. (x1,b1,c1') \ toSet ((\'[x::=v]\<^sub>\\<^sub>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]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using append_g_toSetU by auto then show ?thesis by auto next case rhs hence "(x1,b1,c1) \ toSet (\'[x::=v]\<^sub>\\<^sub>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') #\<^sub>\\)) c" and "\; \; \ \\<^sub>w\<^sub>f v : b" shows "\; \; ((\'[x::=v]\<^sub>\\<^sub>v)@\) \\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" proof - have "(\'@((x,b,c') #\<^sub>\\))[x::=v]\<^sub>\\<^sub>v = ((\'[x::=v]\<^sub>\\<^sub>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) #\<^sub>\\)" shows "atom x \ \" proof - have "wfG P B ((x,b,c) #\<^sub>\\)" 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' ; \ \\<^sub>w\<^sub>f v : b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and "\ ; B' ; \ \\<^sub>w\<^sub>f c \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b" and "\ ; B' \\<^sub>w\<^sub>f \ \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" and "\ ; B' ; \ \\<^sub>w\<^sub>f \ \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f \ \True" and "\ ; B' \\<^sub>w\<^sub>f b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " and "\ ; B' ; \ \\<^sub>w\<^sub>f ce : b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and "\ \\<^sub>w\<^sub>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 \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wfV_varI by auto show "Some (b1[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \[bv::=b]\<^sub>\\<^sub>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]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv::=b]\<^sub>b\<^sub>b" using wfV_pairI by simp show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>b : b2[bv::=b]\<^sub>b\<^sub>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]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>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]\<^sub>b\<^sub>b = b'" using forget_subst subst_bb_def fresh_def by (metis empty_iff subst_b_b_def) ultimately show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>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 \\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by metis have "atom bva \ \[bv::=b]\<^sub>\\<^sub>b" using fresh_subst_if subst_b_\_def wfV_conspI by metis moreover have "atom bva \ ba[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfV_conspI by metis moreover have "atom bva \ v[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfV_conspI by metis ultimately show \atom bva \ (\, B, \[bv::=b]\<^sub>\\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b)\ unfolding fresh_prodN using wfV_conspI fresh_def supp_fset by auto show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bva::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>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]\<^sub>\\<^sub>b)" using wfTI subst_g_b_x_fresh by simp show "\ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfTI by auto show "\ ; B ; (z, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\ \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>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]\<^sub>c\<^sub>b \ {TRUE, FALSE}" using wfG_cons1I(1) by(nominal_induct c' rule: c.strong_induct,auto+) show "\ ; B \\<^sub>w\<^sub>f \'[bv::=b]\<^sub>\\<^sub>b " using wfG_cons1I by auto show "atom x \ \'[bv::=b]\<^sub>\\<^sub>b" using wfG_cons1I subst_g_b_x_fresh by auto show "\ ; B ; (x, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\ \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c'[bv::=b]\<^sub>c\<^sub>b" using wfG_cons1I by auto show "\ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>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]\<^sub>c\<^sub>b \ {TRUE, FALSE}" using wfG_cons2I by auto show "\ ; B \\<^sub>w\<^sub>f \'[bv::=b]\<^sub>\\<^sub>b " using wfG_cons2I by auto show "atom x \ \'[bv::=b]\<^sub>\\<^sub>b" using wfG_cons2I subst_g_b_x_fresh by auto show "\ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>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 (* 46 ms *) 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' ; \ ; \ \\<^sub>w\<^sub>f e : b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f e[bv::=b]\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f s : b \ True" and "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ True" and "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ True" and "\ \\<^sub>w\<^sub>f (\::\) \ True " and "\ ; B' ; \ \\<^sub>w\<^sub>f \ \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" and "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and "\ ; \ ; \ \\<^sub>w\<^sub>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]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=ba]\<^sub>v\<^sub>b : b[bv::=ba]\<^sub>b\<^sub>b) \ \ ts ; \ ; g \\<^sub>w\<^sub>f v : b ) \ \ ts ; f \\<^sub>w\<^sub>f ba" using wfE_plusI.prems(1) wf_b_subst1(1) by force (* 0.0 ms *) then show "\ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f [ plus v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_int[bv::=b]\<^sub>b\<^sub>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 \\<^sub>w\<^sub>f b) \ \ (ts ; {|ba|} ; g \\<^sub>w\<^sub>f v : B_int) \ (ts ; f ; g[ba::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[ba::=b]\<^sub>v\<^sub>b : B_int)" by (metis wf_b_subst1(1) wf_b_subst_lemmas(86)) (* 46 ms *) then show "\ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f [ leq v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_bool[bv::=b]\<^sub>b\<^sub>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)) (* 46 ms *) qed next case (wfE_eqI \ \ \ \ \ v1 bb v2) show ?case unfolding subst_eb.simps subst_bb.simps proof show \ \ \\<^sub>w\<^sub>f \ \ using wfX_wfY wfE_eqI by metis show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b \ using wfX_wfY wfE_eqI by metis show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : bb \ using subst_bb.simps wfE_eqI - by (metis (no_types, hide_lams) empty_iff insert_iff wf_b_subst1(1)) + by (metis (no_types, opaque_lifting) empty_iff insert_iff wf_b_subst1(1)) show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>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]\<^sub>b\<^sub>b = b'" using subst_bb.simps wf_b_subst_lemmas by metis have "\ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f (AE_app f (v[bv::=b]\<^sub>v\<^sub>b)) : (b_of (\[bv::=b]\<^sub>\\<^sub>b))" proof show "\ \\<^sub>w\<^sub>f \" using wfE_appI by auto show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>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]\<^sub>\\<^sub>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]\<^sub>\\<^sub>b s))) = lookup_fun \ f" using wfE_appI by simp show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>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]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f AE_appP f b'[bv::=b]\<^sub>b\<^sub>b (v1[bv::=b]\<^sub>v\<^sub>b) : (b_of \1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b" proof show \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b \ using wfE_appPI by auto show \ \ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b \ using wfE_appPI wf_b_subst1 by auto have "atom bv1 \ \[bv::=b]\<^sub>\\<^sub>b" using fresh_subst_if subst_b_\_def wfE_appPI by metis moreover have "atom bv1 \ b'[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis moreover have "atom bv1 \ v1[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfE_appPI by metis moreover have "atom bv1 \ \[bv::=b]\<^sub>\\<^sub>b" using fresh_subst_if subst_b_\_def wfE_appPI by metis moreover have "atom bv1 \ (b_of \1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis ultimately show "atom bv1 \ (\, \, B, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v1[bv::=b]\<^sub>v\<^sub>b, (b_of \1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>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]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b \ using wfE_appPI subst_b_b_def * wf_b_subst1 by metis thus \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>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 \\<^sub>w\<^sub>f (AE_mvar u)[bv::=b]\<^sub>e\<^sub>b : (b_of (\[bv::=b]\<^sub>\\<^sub>b))" proof(subst subst_eb.simps,rule Wellformed.wfE_mvarI) show "\ \\<^sub>w\<^sub>f \ " using wfE_mvarI by simp show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" using wfE_mvarI by metis show "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>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]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wfD_cons by auto show "\ ; B ; \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wfD_cons wf_b_subst1 by auto show "u \ fst ` setD \[bv::=b]\<^sub>\\<^sub>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) #\<^sub>\GNil \\<^sub>w\<^sub>f \" and "\ ; B \\<^sub>w\<^sub>f b'" shows "\ ; B ; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\GNil \\<^sub>w\<^sub>f (\[bv::=b']\<^sub>\\<^sub>b)" proof - have "\ ; B ; ((x,b,c) #\<^sub>\GNil)[bv::=b']\<^sub>\\<^sub>b \\<^sub>w\<^sub>f (\[bv::=b']\<^sub>\\<^sub>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 "\; \; \ \\<^sub>w\<^sub>f v : b' \ \ = (x, b, c2) #\<^sub>\ G \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c2 \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f v : b'" and "\; \; \ \\<^sub>w\<^sub>f c \ \ = (x, b, c2) #\<^sub>\ G \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c2 \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c" and "\; \ \\<^sub>w\<^sub>f \ \ True" and "\; \; \ \\<^sub>w\<^sub>f \ \ True" and "\; \; \ \\<^sub>w\<^sub>f ts \ True" and "\\<^sub>w\<^sub>f \ \True" and "\; \ \\<^sub>w\<^sub>f b \ True " and "\; \; \ \\<^sub>w\<^sub>f ce : b' \ \ = (x, b, c2) #\<^sub>\ G \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c2 \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f ce : b' " and "\ \\<^sub>w\<^sub>f td \ True" proof(nominal_induct b' and c and \ and \ and ts and \ and b and b' and td avoiding: c1 arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) case (wfV_varI \ \ \ b' c' x') have wbg: "\; \ \\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\ G " using wfC_wf wfV_varI by simp show ?case proof(cases "x=x'") case True have "Some (b', c1) = lookup ((x, b, c1) #\<^sub>\ 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) #\<^sub>\ 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 \\; \ \\<^sub>w\<^sub>f b1 \ using wfV_conspI by auto show \atom bv \ (\, \, (x, b, c1) #\<^sub>\ G, b1, v)\ unfolding fresh_prodN fresh_GCons using wfV_conspI fresh_prodN fresh_GCons by simp show \\; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b\ using wfV_conspI by auto qed qed( (auto | metis wfC_wf wf_intros) +) end \ No newline at end of file