diff --git a/thys/MiniSail/BTVSubst.thy b/thys/MiniSail/BTVSubst.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/BTVSubst.thy @@ -0,0 +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)) + 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/BTVSubstTypingL.thy b/thys/MiniSail/BTVSubstTypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/BTVSubstTypingL.thy @@ -0,0 +1,635 @@ +(*<*) +theory BTVSubstTypingL + imports "HOL-Eisbach.Eisbach_Tools" ContextSubtypingL SubstMethods +begin + (*>*) + +chapter \Basic Type Variable Substitution Lemmas\ +text \Lemmas that show that types are preserved, in some way, +under basic type variable substitution\ + +lemma subst_vv_subst_bb_commute: + fixes bv::bv and b::b and x::x and v::v + assumes "atom bv \ v" + shows "(v'[x::=v]\<^sub>v\<^sub>v)[bv::=b]\<^sub>v\<^sub>b = (v'[bv::=b]\<^sub>v\<^sub>b)[x::=v]\<^sub>v\<^sub>v" + using assms proof(nominal_induct v' rule: v.strong_induct) + case (V_lit x) + then show ?case using subst_vb.simps subst_vv.simps by simp +next + case (V_var y) + hence "v[bv::=b]\<^sub>v\<^sub>b=v" using forget_subst subst_b_v_def by metis + then show ?case unfolding subst_vb.simps(2) subst_vv.simps(2) using V_var by auto +next + case (V_pair x1a x2a) + then show ?case using subst_vb.simps subst_vv.simps by simp +next + case (V_cons x1a x2a x3) + then show ?case using subst_vb.simps subst_vv.simps by simp +next + case (V_consp x1a x2a x3 x4) + then show ?case using subst_vb.simps subst_vv.simps by simp +qed + +lemma subst_cev_subst_bb_commute: + fixes bv::bv and b::b and x::x and v::v + assumes "atom bv \ v" + shows "(ce[x::=v]\<^sub>v)[bv::=b]\<^sub>c\<^sub>e\<^sub>b = (ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b)[x::=v]\<^sub>v" + using assms apply (nominal_induct ce rule: ce.strong_induct, (simp add: subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps)) + using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps + by (simp add: subst_v_ce_def)+ + +lemma subst_cv_subst_bb_commute: + fixes bv::bv and b::b and x::x and v::v + assumes "atom bv \ v" + shows "c[x::=v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[x::=v]\<^sub>c\<^sub>v" + using assms apply (nominal_induct c rule: c.strong_induct ) + using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps + subst_v_c_def subst_b_c_def apply auto + using subst_cev_subst_bb_commute subst_v_ce_def by auto+ + +lemma subst_b_c_of: + "(c_of \ z)[bv::=b]\<^sub>c\<^sub>b = c_of (\[bv::=b]\<^sub>\\<^sub>b) z" +proof(nominal_induct \ avoiding: z rule:\.strong_induct) + case (T_refined_type z' b' c') + moreover have "atom bv \ [ z ]\<^sup>v " using fresh_at_base v.fresh by auto + ultimately show ?case using subst_cv_subst_bb_commute[of bv "V_var z" c' z' b] c_of.simps subst_tb.simps by metis +qed + +lemma subst_b_b_of: + "(b_of \)[bv::=b]\<^sub>b\<^sub>b = b_of (\[bv::=b]\<^sub>\\<^sub>b)" + by(nominal_induct \ rule:\.strong_induct, simp add: b_of.simps subst_tb.simps ) + +lemma subst_b_if: + "\ z : b_of \[bv::=b]\<^sub>\\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit ll) IMP c_of \[bv::=b]\<^sub>\\<^sub>b z \ = \ z : b_of \ | CE_val (v) == CE_val (V_lit ll) IMP c_of \ z \[bv::=b]\<^sub>\\<^sub>b " + unfolding subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps using subst_b_b_of subst_b_c_of by auto + +lemma subst_b_top_eq: + "\ z : B_unit | TRUE \[bv::=b]\<^sub>\\<^sub>b = \ z : B_unit | TRUE \" and "\ z : B_bool | TRUE \[bv::=b]\<^sub>\\<^sub>b = \ z : B_bool | TRUE \" and "\ z : B_id tid | TRUE \ = \ z : B_id tid | TRUE \[bv::=b]\<^sub>\\<^sub>b" + unfolding subst_tb.simps subst_bb.simps subst_cb.simps by auto + +lemmas subst_b_eq = subst_b_c_of subst_b_b_of subst_b_if subst_b_top_eq + +lemma subst_cx_subst_bb_commute[simp]: + fixes bv::bv and b::b and x::x and v'::c + shows "(v'[x::=V_var y]\<^sub>c\<^sub>v)[bv::=b]\<^sub>c\<^sub>b = (v'[bv::=b]\<^sub>c\<^sub>b)[x::=V_var y]\<^sub>c\<^sub>v" + using subst_cv_subst_bb_commute fresh_at_base v.fresh by auto + +lemma subst_b_infer_b: + fixes l::l and b::b + assumes " \ l \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ l \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms infer_l_form3 infer_l_form4 wf_b_subst infer_l_wf subst_tb.simps base_for_lit.simps subst_tb.simps + subst_b_base_for_lit subst_cb.simps(6) subst_ceb.simps(1) subst_vb.simps(1) subst_vb.simps(2) type_l_eq + by metis + +lemma subst_b_subtype: + fixes s::s and b'::b + assumes "\ ; {|bv|} ; \ \ \1 \ \2" and "\ ; {||} \\<^sub>w\<^sub>f b'" and "B = {|bv|}" + shows "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \ \1[bv::=b']\<^sub>\\<^sub>b \ \2[bv::=b']\<^sub>\\<^sub>b" + using assms proof(nominal_induct "{|bv|}" \ \1 \2 rule:subtype.strong_induct) + case (subtype_baseI x \ \ z c z' c' b) + + hence **: "\ ; {|bv|} ; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \ c'[z'::=V_var x]\<^sub>c\<^sub>v " using validI subst_defs by metis + + have "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \ \ z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ \ \ z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \" proof(rule Typing.subtype_baseI) + show "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \ z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ " + using subtype_baseI assms wf_b_subst(4) subst_tb.simps subst_defs by metis + show "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \ z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \ " + using subtype_baseI assms wf_b_subst(4) subst_tb.simps by metis + show "atom x \(\, {||}::bv fset, \[bv::=b']\<^sub>\\<^sub>b, z , c[bv::=b']\<^sub>c\<^sub>b , z' , c'[bv::=b']\<^sub>c\<^sub>b )" + apply(unfold fresh_prodN,auto simp add: * fresh_prodN fresh_empty_fset) + using subst_b_fresh_x * fresh_prodN \atom x \ c\ \atom x \ c'\ subst_defs subtype_baseI by metis+ + have "\ ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, c[z::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \[bv::=b']\<^sub>\\<^sub>b \ c'[z'::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b" + using ** subst_b_valid subst_gb.simps assms subtype_baseI by metis + thus "\ ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, (c[bv::=b']\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\ \[bv::=b']\<^sub>\\<^sub>b \ (c'[bv::=b']\<^sub>c\<^sub>b)[z'::=V_var x]\<^sub>v" + using subst_defs subst_cv_subst_bb_commute by (metis subst_cx_subst_bb_commute) + qed + thus ?case using subtype_baseI subst_tb.simps subst_defs by metis +qed + +lemma b_of_subst_bv: + "(b_of \)[x::=v]\<^sub>b\<^sub>b = b_of (\[x::=v]\<^sub>\\<^sub>b)" +proof - + obtain z b c where *:"\ = \ z : b | c \ \ atom z \ (x,v)" using obtain_fresh_z by metis + thus ?thesis using subst_tv.simps * by auto +qed + +lemma subst_b_infer_v: + fixes v::v and b::b + assumes "\ ; B ; G \ v \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms proof(nominal_induct avoiding: b bv rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ b' c x z) + show ?case unfolding subst_b_simps proof + show "\ ; {||} \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using infer_v_varI wf_b_subst by metis + show "Some (b'[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \[bv::=b]\<^sub>\\<^sub>b x" using subst_b_lookup infer_v_varI by metis + show "atom z \ x" using infer_v_varI by auto + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b) " by(fresh_mth add: infer_v_varI subst_b_fresh_x subst_b_\_def fresh_prodN fresh_empty_fset ) + qed +next + case (infer_v_litI \ \ \ l \) + then show ?case using Typing.infer_v_litI subst_b_infer_b + using wf_b_subst1(3) by auto +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + show ?case unfolding subst_b_simps b_of_subst_bv proof + show "atom z \ (v1[bv::=b]\<^sub>v\<^sub>b, v2[bv::=b]\<^sub>v\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x) + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x subst_b_\_def fresh_empty_fset) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ t1[bv::=b]\<^sub>\\<^sub>b" using infer_v_pairI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ t2[bv::=b]\<^sub>\\<^sub>b" using infer_v_pairI by auto + qed +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + show ?case unfolding subst_b_simps b_of_subst_bv proof + show "AF_typedef s dclist \ set \" using infer_v_consI by auto + show "(dc, tc) \ set dclist" using infer_v_consI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b" using infer_v_consI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc" proof - + have "atom bv \ tc" using wfTh_lookup_supp_empty fresh_def infer_v_consI infer_v_wf by fast + moreover have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc[bv::=b]\<^sub>\\<^sub>b" + using subst_b_subtype infer_v_consI by simp + ultimately show ?thesis using forget_subst subst_b_\_def by metis + qed + show "atom z \ v[bv::=b]\<^sub>v\<^sub>b" using infer_v_consI using subst_b_fresh_x subst_b_v_def by metis + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b)" by(fresh_mth add: infer_v_consI subst_b_fresh_x subst_b_\_def fresh_empty_fset) + qed +next + case (infer_v_conspI s bv2 dclist2 \ dc tc \ \ v tv ba z) + + have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) \ \ z : B_app s (ba[bv::=b]\<^sub>b\<^sub>b) | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) ]\<^sup>c\<^sup>e \" + proof(rule Typing.infer_v_conspI) + show "AF_typedef_poly s bv2 dclist2 \ set \" using infer_v_conspI by auto + show "(dc, tc) \ set dclist2" using infer_v_conspI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b" + using infer_v_conspI subst_tb.simps by metis + + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b" proof - + have "supp tc \ { atom bv2 }" using infer_v_conspI wfTh_poly_lookup_supp wfX_wfY by metis + moreover have "bv2 \ bv" using \atom bv2 \ \\ \\ = {|bv|} \ fresh_at_base fresh_def + using fresh_finsert by fastforce + ultimately have "atom bv \ tc" unfolding fresh_def by auto + hence "tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b = tc[bv2::=ba]\<^sub>\\<^sub>b[bv::=b]\<^sub>\\<^sub>b" + using subst_tb_commute by metis + moreover have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc[bv2::=ba]\<^sub>\\<^sub>b[bv::=b]\<^sub>\\<^sub>b" + using infer_v_conspI(7) subst_b_subtype infer_v_conspI by metis + ultimately show ?thesis by auto + qed + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)" + apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset) + using \atom z \ \\ fresh_subst_if subst_b_\_def x_fresh_b apply metis + using \atom z \ v\ fresh_subst_if subst_b_v_def x_fresh_b by metis + show "atom bv2 \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)" + apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset) + using \atom bv2 \ b\ \atom bv2 \ \\ fresh_subst_if subst_b_\_def apply metis + using \atom bv2 \ b\ \atom bv2 \ v\ fresh_subst_if subst_b_v_def apply metis + using \atom bv2 \ b\ \atom bv2 \ ba\ fresh_subst_if subst_b_b_def by metis + show "\ ; {||} \\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b " + using infer_v_conspI wf_b_subst by metis + qed + thus ?case using subst_vb.simps subst_tb.simps subst_bb.simps by simp + +qed + +lemma subst_b_check_v: + fixes v::v and b::b + assumes "\ ; B ; G \ v \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ (\[bv::=b]\<^sub>\\<^sub>b)" +proof - + obtain \' where "\ ; B ; G \ v \ \' \ \ ; B ; G \ \' \ \" using check_v_elims[OF assms(1)] by metis + thus ?thesis using subst_b_subtype subst_b_infer_v assms + by (metis (no_types) check_v_subtypeI subst_b_infer_v subst_b_subtype) +qed + +lemma subst_vv_subst_vb_switch: + shows "(v'[bv::=b']\<^sub>v\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>v\<^sub>v = v'[x::=v]\<^sub>v\<^sub>v[bv::=b']\<^sub>v\<^sub>b" +proof(nominal_induct v' rule:v.strong_induct) + case (V_lit x) + then show ?case using subst_vv.simps subst_vb.simps by auto +next + case (V_var x) + then show ?case using subst_vv.simps subst_vb.simps by auto +next + case (V_pair x1a x2a) + then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto +next + case (V_cons x1a x2a x3) + then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto +next + case (V_consp x1a x2a x3 x4) + then show ?case using subst_vv.simps subst_vb.simps v.fresh pure_fresh + by (metis forget_subst subst_b_b_def) +qed + +lemma subst_cev_subst_vb_switch: + shows "(ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>e\<^sub>v = (ce[x::=v]\<^sub>c\<^sub>e\<^sub>v)[bv::=b']\<^sub>c\<^sub>e\<^sub>b" + by(nominal_induct ce rule:ce.strong_induct, auto simp add: subst_vv_subst_vb_switch ce.fresh) + +lemma subst_cv_subst_vb_switch: + shows "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b" + by(nominal_induct c rule:c.strong_induct, auto simp add: subst_cev_subst_vb_switch c.fresh) + +lemma subst_tv_subst_vb_switch: + shows "(\[bv::=b']\<^sub>\\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \[x::=v]\<^sub>\\<^sub>v[bv::=b']\<^sub>\\<^sub>b" +proof(nominal_induct \ avoiding: x v rule:\.strong_induct) + case (T_refined_type z b c ) + hence ceq: "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b" using subst_cv_subst_vb_switch by auto + + moreover have "atom z \ v[bv::=b']\<^sub>v\<^sub>b" using x_fresh_b fresh_subst_if subst_b_v_def T_refined_type by metis + + hence "\ z : b | c \[bv::=b']\<^sub>\\<^sub>b[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \ z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \" + using subst_tv.simps subst_tb.simps T_refined_type fresh_Pair by metis + + moreover have "\ z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \ = \ z : b | c[x::=v]\<^sub>c\<^sub>v \[bv::=b']\<^sub>\\<^sub>b" + using subst_tv.simps subst_tb.simps ceq \.fresh forget_subst[of "bv" b b'] subst_b_b_def T_refined_type by metis + + ultimately show ?case using subst_tv.simps subst_tb.simps ceq T_refined_type by auto +qed + +lemma subst_tb_triple: + assumes "atom bv \ \'" + shows "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \'[bv'::=b']\<^sub>\\<^sub>b[x'::=v']\<^sub>\\<^sub>v[bv::=b]\<^sub>\\<^sub>b" +proof - + have "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \'[bv'::=b']\<^sub>\\<^sub>b[bv::=b]\<^sub>\\<^sub>b [x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\\<^sub>v" + using subst_tb_commute \atom bv \ \'\ by auto + also have "... = \'[bv'::=b']\<^sub>\\<^sub>b [x'::=v']\<^sub>\\<^sub>v[bv::=b]\<^sub>\\<^sub>b" + using subst_tv_subst_vb_switch by auto + finally show ?thesis using fresh_subst_if forget_subst by auto +qed + +lemma subst_b_infer_e: + fixes s::s and b::b + assumes "\ ; \ ; B ; G; D \ e \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b \ (e[bv::=b]\<^sub>e\<^sub>b) \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms proof(nominal_induct avoiding: b rule: infer_e.strong_induct) + case (infer_e_valI \ \ \ \ \ v \) + thus ?case using subst_eb.simps infer_e.intros wf_b_subst subst_db.simps wf_b_subst infer_v_wf subst_b_infer_v + by (metis forget_subst ms_fresh_all(1) wfV_b_fresh) +next + case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + + show ?case unfolding subst_b_simps subst_eb.simps proof(rule Typing.infer_e_plusI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_plusI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_plusI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_plusI subst_b_simps by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_plusI subst_b_simps by force + show "atom z3 \ AE_op Plus (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_plusI subst_b_fresh_x subst_b_e_def by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_plusI by auto + qed +next + case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_leqI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_leqI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_leqI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_leqI subst_b_simps by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_leqI subst_b_simps by force + show "atom z3 \ AE_op LEq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_leqI subst_b_fresh_x subst_b_e_def by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_leqI by auto + qed +next + case (infer_e_eqI \ \ \ \ \ v1 z1 bb c1 v2 z2 c2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_eqI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_eqI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_eqI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : bb[bv::=b]\<^sub>b\<^sub>b | c1[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_eqI subst_b_simps by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : bb[bv::=b]\<^sub>b\<^sub>b | c2[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_eqI subst_b_simps by force + show "atom z3 \ AE_op Eq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_eqI subst_b_fresh_x subst_b_e_def by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_eqI by auto + show "bb[bv::=b]\<^sub>b\<^sub>b \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed +next + case (infer_e_appI \ \ \ \ \ f x b' c \' s' v \) + show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" using wf_b_subst(10) subst_db.simps infer_e_appI wfX_wfY by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_appI by auto + show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c \' s'))) = lookup_fun \ f" using infer_e_appI by auto + (*show "\ ; {||} ; (x, b', c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \'" using infer_e_appI sory*) + have "atom bv \ b'" using \\ \\<^sub>w\<^sub>f \ \ infer_e_appI wfPhi_f_supp fresh_def[of "atom bv" b'] by simp + hence "b' = b'[bv::=b]\<^sub>b\<^sub>b" using subst_b_simps + using has_subst_b_class.forget_subst subst_b_b_def by force + moreover have ceq:"c = c[bv::=b]\<^sub>c\<^sub>b" using subst_b_simps proof - + have "supp c \ {atom x}" using infer_e_appI wfPhi_f_simple_supp_c[OF _ \\ \\<^sub>w\<^sub>f \ \] by simp + hence "atom bv \ c" using + fresh_def[of "atom bv" c] + using fresh_def fresh_finsert insert_absorb + insert_subset ms_fresh_all supp_at_base x_not_in_b_set fresh_prodN by metis + thus ?thesis + using forget_subst subst_b_c_def fresh_def[of "atom bv" c] by metis + qed + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ x : b' | c \" + using subst_b_check_v subst_tb.simps subst_vb.simps infer_e_appI + proof - + have "\ ; {|bv|} ; \ \ v \ \ x : b' | c \" + by (metis \\ = {|bv|}\ \\ ; \ ; \ \ v \ \ x : b' | c \\) (* 0.0 ms *) + then show ?thesis + by (metis (no_types) \\ ; {||} \\<^sub>w\<^sub>f b\ \b' = b'[bv::=b]\<^sub>b\<^sub>b\ subst_b_check_v subst_tb.simps ceq) + qed + show "atom x \ (\, \, {||}::bv fset, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)" + apply (fresh_mth add: fresh_prodN subst_g_b_x_fresh infer_e_appI ) + using subst_b_fresh_x infer_e_appI apply metis+ + done + have "supp \' \ { atom x }" using wfPhi_f_simple_supp_t infer_e_appI by auto + hence "atom bv \ \'" using fresh_def fresh_at_base by force + then show "\'[x::=v[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \[bv::=b]\<^sub>\\<^sub>b" using infer_e_appI + forget_subst subst_b_\_def subst_tv_subst_vb_switch subst_defs by metis + qed +next + case (infer_e_appPI \' \ \' \ \' b' f' bv' x' b1 c \' s' v' \1) + + have beq: "b1[bv'::=b']\<^sub>b\<^sub>b[bv::=b]\<^sub>b\<^sub>b = b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b" + proof - + have "supp b1 \ { atom bv' }" using wfPhi_f_poly_supp_b infer_e_appPI + using supp_at_base by blast + moreover have "bv \ bv'" using infer_e_appPI fresh_def supp_at_base + by (simp add: fresh_def supp_at_base) + ultimately have "atom bv \ b1" using fresh_def fresh_at_base by force + thus ?thesis by simp + qed + + have ceq: "(c[bv'::=b']\<^sub>c\<^sub>b)[bv::=b]\<^sub>c\<^sub>b = c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>c\<^sub>b" proof - + have "supp c \ { atom bv', atom x' }" using wfPhi_f_poly_supp_c infer_e_appPI + using supp_at_base by blast + moreover have "bv \ bv'" using infer_e_appPI fresh_def supp_at_base + by (simp add: fresh_def supp_at_base) + moreover have "atom x' \ atom bv" by auto + ultimately have "atom bv \ c" using fresh_def[of "atom bv" c] fresh_at_base by auto + thus ?thesis by simp + qed + + show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appPI) + show "\' ; {||} ; \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst subst_db.simps infer_e_appPI wfX_wfY by metis + show "\' \\<^sub>w\<^sub>f \' " using infer_e_appPI by auto + show "Some (AF_fundef f' (AF_fun_typ_some bv' (AF_fun_typ x' b1 c \' s'))) = lookup_fun \' f'" using infer_e_appPI by auto + thus "\' ; {||} ; \'[bv::=b]\<^sub>\\<^sub>b \ v'[bv::=b]\<^sub>v\<^sub>b \ \ x' : b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b | c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b \" + using subst_b_check_v subst_tb.simps subst_b_simps infer_e_appPI + proof - + have "\' ; {||} ; \'[bv::=b]\<^sub>\\<^sub>b \ v'[bv::=b]\<^sub>v\<^sub>b \ \ x' : b1[bv'::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b | (c[bv'::=b']\<^sub>b)[bv::=b]\<^sub>c\<^sub>b \" + using infer_e_appPI subst_b_check_v subst_tb.simps by metis + thus ?thesis using beq ceq subst_defs by metis + qed + show "atom x' \ \'[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_appPI by auto + show "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \1[bv::=b]\<^sub>\\<^sub>b" proof - (* \'[bv'::=b']\<^sub>\\<^sub>b[x'::=v']\<^sub>\\<^sub>v = \1 *) + + have "supp \' \ { atom x', atom bv' }" using wfPhi_f_poly_supp_t infer_e_appPI by auto + moreover hence "bv \ bv'" using infer_e_appPI fresh_def supp_at_base + by (simp add: fresh_def supp_at_base) + ultimately have "atom bv \ \'" using fresh_def by force + hence "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \'[bv'::=b']\<^sub>b[x'::=v']\<^sub>v[bv::=b]\<^sub>\\<^sub>b" using subst_tb_triple subst_defs by auto + thus ?thesis using infer_e_appPI by metis + qed + show "atom bv' \ (\', \', {||}, \'[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v'[bv::=b]\<^sub>v\<^sub>b, \1[bv::=b]\<^sub>\\<^sub>b)" + unfolding fresh_prodN apply( auto simp add: infer_e_appPI fresh_empty_fset) + using fresh_subst_if subst_b_\_def subst_b_\_def subst_b_b_def subst_b_v_def subst_b_\_def infer_e_appPI by metis+ + show "\' ; {||} \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using infer_e_appPI wf_b_subst by simp + qed +next + case (infer_e_fstI \ \ \ \ \ v z' b1 b2 c z) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_fstI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_fstI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_fstI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_fstI by force + show "atom z \ AE_fst (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_fstI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_fstI by auto + qed +next + case (infer_e_sndI \ \ \ \ \ v z' b1 b2 c z) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_sndI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_sndI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_sndI by force + show "atom z \ AE_snd (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_sndI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_sndI by auto + qed +next + case (infer_e_lenI \ \ \ \ \ v z' c z) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_lenI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_lenI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_lenI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z' : B_bitvec | c[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_lenI by force + show "atom z \ AE_len (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_lenI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_lenI by auto + qed +next + case (infer_e_mvarI \ \ \ \ \ u \) + show ?case proof(subst subst subst_eb.simps, rule Typing.infer_e_mvarI) + show "\ ; {||} \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using infer_e_mvarI wf_b_subst by auto + show "\ \\<^sub>w\<^sub>f \ " using infer_e_mvarI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using infer_e_mvarI using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY + by (metis wf_b_subst(15)) + show "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>b" using infer_e_mvarI subst_db.simps set_insert + subst_d_b_member by simp + qed +next + case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_concatI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_concatI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_concatI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_bitvec | c2[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force + show "atom z3 \ AE_concat (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using infer_e_concatI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_concatI by auto + qed +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_splitI) + show \ \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b \ using wf_b_subst(10) subst_db.simps infer_e_splitI wfX_wfY + by (metis wf_b_subst(15)) + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_splitI by auto + show \ \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \\ + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_splitI by force + show \\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND + [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[bv::=b]\<^sub>v\<^sub>b ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \\ + using subst_b_check_v subst_tb.simps subst_b_simps infer_e_splitI + proof - + have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \[bv::=b]\<^sub>\\<^sub>b" + using infer_e_splitI.hyps(7) infer_e_splitI.prems(1) infer_e_splitI.prems(2) subst_b_check_v by presburger (* 0.0 ms *) + then show ?thesis + by simp (* 0.0 ms *) + qed + show \atom z1 \ AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\ using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis + show \atom z1 \ \[bv::=b]\<^sub>\\<^sub>b\ using subst_g_b_x_fresh infer_e_splitI by auto + + show \atom z2 \ AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\ using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis + + show \atom z2 \ \[bv::=b]\<^sub>\\<^sub>b\ using subst_g_b_x_fresh infer_e_splitI by auto + show \atom z3 \ AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\ using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis + show \atom z3 \ \[bv::=b]\<^sub>\\<^sub>b\ using subst_g_b_x_fresh infer_e_splitI by auto + qed +qed + +text \This is needed for preservation. When we apply a function "f [b] v" we need to +substitute into the body of the function f replacing type-variable with b\ + +lemma subst_b_c_of_forget: + assumes "atom bv \ const" + shows "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x" + using assms proof(nominal_induct const avoiding: x rule:\.strong_induct) + case (T_refined_type x' b' c') + hence "c_of \ x' : b' | c' \ x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by metis + moreover have "atom bv \ c'[x'::=V_var x]\<^sub>c\<^sub>v" proof - + have "atom bv \ c'" using T_refined_type \.fresh by simp + moreover have "atom bv \ V_var x" using v.fresh by simp + ultimately show ?thesis + using T_refined_type \.fresh subst_b_c_def fresh_subst_if + \_fresh_c fresh_subst_cv_if has_subst_b_class.subst_b_fresh_x ms_fresh_all(37) ms_fresh_all assms by metis + qed + ultimately show ?case using forget_subst subst_b_c_def by metis +qed + +lemma subst_b_check_s: + fixes s::s and b::b and cs::branch_s and css::branch_list and v::v and \::\ + assumes "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" (*and "D = []"*) + shows "\ ; \ ; B ; G; D \ s \ \ \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b \ (s[bv::=b]\<^sub>s\<^sub>b) \ (\[bv::=b]\<^sub>\\<^sub>b)" and + "\ ; \ ; B ; G; D ; tid ; cons ; const ; v \ cs \ \ \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b ; tid ; cons ; const ; v[bv::=b]\<^sub>v\<^sub>b \ (subst_branchb cs bv b) \ (\[bv::=b]\<^sub>\\<^sub>b)" and + "\ ; \ ; B ; G; D ; tid ; dclist ; v \ css \ \ \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \ (subst_branchlb css bv b ) \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) + note facts = wfD_emptyI wfX_wfY wf_b_subst subst_b_subtype subst_b_infer_v + case (check_valI \ \ \ \ \ v \' \) + show ?case + apply(subst subst_sb.simps, rule Typing.check_valI) + using facts check_valI apply metis + using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast + using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast + using check_valI subst_b_infer_v wf_b_subst subst_b_subtype by metis +next + case (check_letI x \ \ \ \ \ e \ z s b' c) + show ?case proof(subst subst_sb.simps, rule Typing.check_letI) + + show "atom x \(\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)" (*using check_letI subst_b_\_def subst_b_\_def subst_b_fresh_x fresh_prod4 subst_b_c_def sory*) + apply(unfold fresh_prodN,auto) + apply(simp add: check_letI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done + show "atom z \ (x, \, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)" + apply(unfold fresh_prodN,auto) + apply(simp add: check_letI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done + show "\ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ e[bv::=b]\<^sub>e\<^sub>b \ \ z : b'[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \" + using check_letI subst_b_infer_e subst_tb.simps by metis + have "c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>c\<^sub>v" + using subst_cv_subst_bb_commute[of bv "V_var x" c z b] fresh_at_base by simp + thus "\ ; \ ; {||} ; ((x, b'[bv::=b]\<^sub>b\<^sub>b , (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\ \[bv::=b]\<^sub>\\<^sub>b) ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" + using check_letI subst_gb.simps subst_defs by metis + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof(subst subst_sb.simps, rule Typing.check_assertI) + show "atom x \ (\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)" + apply(unfold fresh_prodN,auto) + apply(simp add: check_assertI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_assertI fresh_prodN)+ done + + have "\ ; \ ; {||} ; ((x, B_bool, c) #\<^sub>\ \)[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" using check_assertI + by metis + thus "\ ; \ ; {||} ; (x, B_bool, c[bv::=b]\<^sub>c\<^sub>b) #\<^sub>\ \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" using subst_gb.simps by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ c[bv::=b]\<^sub>c\<^sub>b " using subst_b_valid check_assertI by simp + show " \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst2(6) check_assertI by simp + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_consI by simp +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_finalI by simp +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + + show ?case unfolding subst_b_simps proof(rule Typing.check_branch_s_branchI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using check_branch_s_branchI wf_b_subst subst_db.simps by metis + show "\\<^sub>w\<^sub>f \ " using check_branch_s_branchI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using check_branch_s_branchI wf_b_subst by metis + + show "atom x \(\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, tid, cons , const, v[bv::=b]\<^sub>v\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)" + apply(unfold fresh_prodN,auto) + apply(simp add: check_branch_s_branchI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_branch_s_branchI fresh_prodN)+ + done + show wft:"\ ; {||} ; GNil \\<^sub>w\<^sub>f const" using check_branch_s_branchI by auto + hence "(b_of const) = (b_of const)[bv::=b]\<^sub>b\<^sub>b" + using wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_b_def \.supp + bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset + by (metis b_of_supp) + moreover have "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x" + using wft wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_c_def \.supp + bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset subst_b_c_of_forget by metis + ultimately show "\ ; \ ; {||} ; (x, b_of const, CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" + using check_branch_s_branchI subst_gb.simps by auto + qed +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case unfolding subst_b_simps proof(rule Typing.check_ifI) + show \atom z \ (\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, s2[bv::=b]\<^sub>s\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)\ + by(unfold fresh_prodN,auto, auto simp add: check_ifI fresh_empty_fset subst_b_fresh_x ) + have "\ z : B_bool | TRUE \[bv::=b]\<^sub>\\<^sub>b = \ z : B_bool | TRUE \" by auto + thus \\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z : B_bool | TRUE \\ using check_ifI subst_b_check_v by metis + show \ \ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s1[bv::=b]\<^sub>s\<^sub>b \ \ z : b_of \[bv::=b]\<^sub>\\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_true) IMP c_of \[bv::=b]\<^sub>\\<^sub>b z \\ + using subst_b_if check_ifI by metis + show \ \ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s2[bv::=b]\<^sub>s\<^sub>b \ \ z : b_of \[bv::=b]\<^sub>\\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_false) IMP c_of \[bv::=b]\<^sub>\\<^sub>b z \\ + using subst_b_if check_ifI by metis + qed +next + case (check_let2I x \ \ \ G \ t s1 \ s2 ) + show ?case unfolding subst_b_simps proof (rule Typing.check_let2I) + have "atom x \ b" using x_fresh_b by auto + show \atom x \ (\, \, {||}, G[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, t[bv::=b]\<^sub>\\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)\ + apply(unfold fresh_prodN, auto, auto simp add: check_let2I fresh_prodN fresh_empty_fset) + apply(metis subst_b_fresh_x check_let2I fresh_prodN)+ + done + + show \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s1[bv::=b]\<^sub>s\<^sub>b \ t[bv::=b]\<^sub>\\<^sub>b \ using check_let2I subst_tb.simps by auto + show \ \ ; \ ; {||} ; (x, b_of t[bv::=b]\<^sub>\\<^sub>b, c_of t[bv::=b]\<^sub>\\<^sub>b x) #\<^sub>\ G[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s2[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b\ + using check_let2I subst_tb.simps subst_gb.simps b_of.simps subst_b_c_of subst_b_b_of by auto + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case unfolding subst_b_simps proof(rule Typing.check_varI) + show "atom u \ (\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, \'[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b) " + by(unfold fresh_prodN,auto simp add: check_varI fresh_empty_fset subst_b_fresh_u ) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \'[bv::=b]\<^sub>\\<^sub>b" using check_varI subst_b_check_v by auto + show "\ ; \ ; {||} ; (subst_gb \ bv b) ; (u, (\'[bv::=b]\<^sub>\\<^sub>b)) #\<^sub>\ (subst_db \ bv b) \ (s[bv::=b]\<^sub>s\<^sub>b) \ (\[bv::=b]\<^sub>\\<^sub>b)" using check_varI by auto + qed +next + case (check_assignI \ \ \ \ \ u \ v z \') + show ?case unfolding subst_b_simps proof( rule Typing.check_assignI) + show "\ \\<^sub>w\<^sub>f \ " using check_assignI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst check_assignI by auto + show "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>b" using check_assignI subst_d_b_member by simp + show " \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" using check_assignI subst_b_check_v by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ \ z : B_unit | TRUE \ \ \'[bv::=b]\<^sub>\\<^sub>b" using check_assignI subst_b_subtype subst_b_simps subst_tb.simps by fastforce + qed +next + case (check_whileI \ \ \ \ \ s1 z s2 \') + show ?case unfolding subst_b_simps proof(rule Typing.check_whileI) + show "\ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s1[bv::=b]\<^sub>s\<^sub>b \ \ z : B_bool | TRUE \" using check_whileI by auto + show "\ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s2[bv::=b]\<^sub>s\<^sub>b \ \ z : B_unit | TRUE \" using check_whileI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ \ z : B_unit | TRUE \ \ \'[bv::=b]\<^sub>\\<^sub>b" using subst_b_subtype check_whileI by fastforce + qed +next + case (check_seqI \ \ \ \ \ s1 z s2 \) + then show ?case unfolding subst_sb.simps using check_seqI Typing.check_seqI subst_b_eq by metis +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + show ?case unfolding subst_b_simps proof(rule Typing.check_caseI) + show \ \ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \ subst_branchlb cs bv b \ \[bv::=b]\<^sub>\\<^sub>b\ using check_caseI by auto + show \AF_typedef tid dclist \ set \\ using check_caseI by auto + show \\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z : B_id tid | TRUE \\ using check_caseI subst_b_check_v subst_b_simps subst_tb.simps subst_b_simps + proof - + have "\ z : B_id tid | TRUE \ = \ z : B_id tid | TRUE \[bv::=b]\<^sub>\\<^sub>b" using subst_b_eq by auto + then show ?thesis + by (metis (no_types) check_caseI.hyps(4) check_caseI.prems(1) check_caseI.prems(2) subst_b_check_v) (* 31 ms *) + qed + show \ \\<^sub>w\<^sub>f \ \ using check_caseI by auto + qed +qed + +end diff --git a/thys/MiniSail/ContextSubtypingL.thy b/thys/MiniSail/ContextSubtypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/ContextSubtypingL.thy @@ -0,0 +1,1068 @@ +(*<*) +theory ContextSubtypingL + imports TypingL "HOL-Eisbach.Eisbach_Tools" SubstMethods +begin + (*>*) + +declare freshers[simp del] + +chapter \Context Subtyping Lemmas\ + +text \Lemmas allowing us to replace the type of a variable in the context with a subtype +and have the judgement remain valid. Also known as narrowing.\ + +section \Replace or exchange type of variable in a context\ + +text \ Because the G-context is extended by the statements like let, we will need a generalised +substitution lemma for statements. +For this we setup a function that replaces in G (rig) for a particular x the constraint for it. +We also define a well-formedness relation for RIGs that ensures that each new constraint +implies the old one\ + +nominal_function replace_in_g_many :: "\ \ (x*c) list \ \" where + "replace_in_g_many G xcs = List.foldr (\(x,c) G. G[x \ c]) xcs G" + by(auto,simp add: eqvt_def replace_in_g_many_graph_aux_def) +nominal_termination (eqvt) by lexicographic_order + +inductive replace_in_g_subtyped :: "\ \ \ \ \ \ (x*c) list \ \ \ bool" (" _ ; _ \ _ \ _ \ \ _" [100,50,50] 50) where + replace_in_g_subtyped_nilI: "\; \ \ G \ [] \ \ G" +| replace_in_g_subtyped_consI: "\ + Some (b,c') = lookup G x ; + \; \; G \\<^sub>w\<^sub>f c ; + \; \; G[x\c] \ c' ; + \; \ \ G[x\c] \ xcs \ \ G'; x \ fst ` set xcs \ \ + \; \ \ G \ (x,c)#xcs \ \ G'" +equivariance replace_in_g_subtyped +nominal_inductive replace_in_g_subtyped . + +inductive_cases replace_in_g_subtyped_elims[elim!]: + "\; \ \ G \ [] \ \ G'" + "\; \ \ ((x,b,c)#\<^sub>\\ G) \ acs \ \ ((x,b,c)#\<^sub>\G')" + "\; \ \ G' \ (x,c)# acs \ \ G" + +lemma rigs_atom_dom_eq: + assumes "\; \ \ G \ xcs \ \ G'" + shows "atom_dom G = atom_dom G'" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI G) + then show ?case by simp +next + case (replace_in_g_subtyped_consI b c' G x \ \ c xcs G') + then show ?case using rig_dom_eq atom_dom.simps dom.simps by simp +qed + +lemma replace_in_g_wfG: + assumes "\; \ \ G \ xcs \ \ G'" and "wfG \ \ G " + shows "wfG \ \ G'" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI \ G) + then show ?case by auto +next + case (replace_in_g_subtyped_consI b c' G x \ c xcs G') + then show ?case using valid_g_wf by auto +qed + +lemma wfD_rig_single: + fixes \::\ and x::x and c::c and G::\ + assumes "\; \; G \\<^sub>w\<^sub>f \ " and "wfG \ \ (G[x\c])" + shows "\; \; G[x\c] \\<^sub>w\<^sub>f \" +proof(cases "atom x \ atom_dom G") + case False + hence "(G[x\c]) = G" using assms replace_in_g_forget wfX_wfY by metis + then show ?thesis using assms by auto +next + case True + then obtain G1 G2 b c' where *: "G=G1@(x,b,c')#\<^sub>\G2" using split_G by fastforce + hence **: "(G[x\c]) = G1@(x,b,c)#\<^sub>\G2" using replace_in_g_inside wfD_wf assms wfD_wf by metis + + hence "wfG \ \ ((x,b,c)#\<^sub>\G2)" using wfG_suffix assms by auto + hence "\; \; (x, b, TRUE) #\<^sub>\ G2 \\<^sub>w\<^sub>f c" using wfG_elim2 by auto + + thus ?thesis using wf_replace_inside1 assms * ** + by (simp add: wf_replace_inside2(6)) +qed + +lemma wfD_rig: + assumes "\; \ \ G \ xcs \ \ G'" and "wfD \ \ G \" + shows "wfD \ \ G' \" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI \ G) + then show ?case by auto +next + case (replace_in_g_subtyped_consI b c' G x \ c xcs G') + then show ?case using wfD_rig_single valid.simps wfC_wf by auto +qed + +lemma replace_in_g_fresh: + fixes x::x + assumes "\; \ \ \ \ xcs \ \ \'" and "wfG \ \ \" and "wfG \ \ \'" and "atom x \ \" + shows "atom x \ \'" + using wfG_dom_supp assms fresh_def rigs_atom_dom_eq by metis + +lemma replace_in_g_fresh1: + fixes x::x + assumes "\; \ \ \ \ xcs \ \ \'" and "wfG \ \ \" and "atom x \ \" + shows "atom x \ \'" +proof - + have "wfG \ \ \'" using replace_in_g_wfG assms by auto + thus ?thesis using assms replace_in_g_fresh by metis +qed + +text \ Wellscoping for an eXchange list\ + +inductive wsX:: "\ \ (x*c) list \ bool" where + wsX_NilI: "wsX G []" +| wsX_ConsI: "\ wsX G xcs ; atom x \ atom_dom G ; x \ fst ` set xcs \ \ wsX G ((x,c)#xcs)" +equivariance wsX +nominal_inductive wsX . + +lemma wsX_if1: + assumes "wsX G xcs" + shows "(( atom ` fst ` set xcs) \ atom_dom G) \ List.distinct (List.map fst xcs)" + using assms by(induct rule: wsX.induct,force+ ) + +lemma wsX_if2: + assumes "(( atom ` fst ` set xcs) \ atom_dom G) \ List.distinct (List.map fst xcs)" + shows "wsX G xcs" + using assms proof(induct xcs) + case Nil + then show ?case using wsX_NilI by fast +next + case (Cons a xcs) + then obtain x and c where xc: "a=(x,c)" by force + have " wsX G xcs" proof - + have "distinct (map fst xcs)" using Cons by force + moreover have " atom ` fst ` set xcs \ atom_dom G" using Cons by simp + ultimately show ?thesis using Cons by fast + qed + moreover have "atom x \ atom_dom G" using Cons xc + by simp + moreover have "x \ fst ` set xcs" using Cons xc + by simp + ultimately show ?case using wsX_ConsI xc by blast +qed + +lemma wsX_iff: + "wsX G xcs = ((( atom ` fst ` set xcs) \ atom_dom G) \ List.distinct (List.map fst xcs))" + using wsX_if1 wsX_if2 by meson + +inductive_cases wsX_elims[elim!]: + "wsX G []" + "wsX G ((x,c)#xcs)" + +lemma wsX_cons: + assumes "wsX \ xcs" and "x \ fst ` set xcs" + shows "wsX ((x, b, c1) #\<^sub>\ \) ((x, c2) # xcs)" + using assms proof(induct \) + case GNil + then show ?case using atom_dom.simps wsX_iff by auto +next + case (GCons xbc \) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + then have "atom ` fst ` set xcs \ atom_dom (xbc #\<^sub>\ \) \ distinct (map fst xcs)" + using GCons.prems(1) wsX_iff by blast + then have "wsX ((x, b, c1) #\<^sub>\ xbc #\<^sub>\ \) xcs" + by (simp add: Un_commute subset_Un_eq wsX_if2) + then show ?case by (simp add: GCons.prems(2) wsX_ConsI) +qed + +lemma wsX_cons2: + assumes "wsX \ xcs" and "x \ fst ` set xcs" + shows "wsX ((x, b, c1) #\<^sub>\ \) xcs" + using assms proof(induct \) + case GNil + then show ?case using atom_dom.simps wsX_iff by auto +next + case (GCons xbc \) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + then have "atom ` fst ` set xcs \ atom_dom (xbc #\<^sub>\ \) \ distinct (map fst xcs)" + using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2) +qed + +lemma wsX_cons3: + assumes "wsX \ xcs" + shows "wsX ((x, b, c1) #\<^sub>\ \) xcs" + using assms proof(induct \) + case GNil + then show ?case using atom_dom.simps wsX_iff by auto +next + case (GCons xbc \) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + then have "atom ` fst ` set xcs \ atom_dom (xbc #\<^sub>\ \) \ distinct (map fst xcs)" + using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2) +qed + +lemma wsX_fresh: + assumes "wsX G xcs" and "atom x \ G" and "wfG \ \ G " + shows "x \ fst ` set xcs" +proof - + have "atom x \ atom_dom G" using assms + using fresh_def wfG_dom_supp by auto + thus ?thesis using wsX_iff assms by blast +qed + +lemma replace_in_g_dist: + assumes "x' \ x" + shows "replace_in_g ((x, b,c) #\<^sub>\ G) x' c'' = ((x, b,c) #\<^sub>\ (replace_in_g G x' c''))" using replace_in_g.simps assms by presburger + +lemma wfG_replace_inside_rig: + fixes c''::c + assumes \\; \ \\<^sub>w\<^sub>f G[x'\c'']\ \\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G \ + shows "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G[x'\c'']" +proof(rule wfG_consI) + + have "wfG \ \ G " using wfG_cons assms by auto + + show *:"\; \ \\<^sub>w\<^sub>f G[x'\c'']" using assms by auto + show "atom x \ G[x'\c'']" using replace_in_g_fresh_single[OF *] assms wfG_elims assms by metis + show **:"\; \ \\<^sub>w\<^sub>f b " using wfG_elim2 assms by auto + show "\; \; (x, b, TRUE) #\<^sub>\ G[x'\c''] \\<^sub>w\<^sub>f c " + proof(cases "atom x' \ atom_dom G") + case True + hence "G = G[x'\c'']" using replace_in_g_forget \wfG \ \ G\ by auto + thus ?thesis using assms wfG_wfC by auto + next + case False + then obtain G1 G2 b' c' where **:"G=G1@(x',b',c')#\<^sub>\G2" + using split_G by fastforce + hence ***: "(G[x'\c'']) = G1@(x',b',c'')#\<^sub>\G2" + using replace_in_g_inside \wfG \ \ G \ by metis + hence "\; \; (x, b, TRUE) #\<^sub>\ G1@(x',b',c')#\<^sub>\G2 \\<^sub>w\<^sub>f c" using * ** assms wfG_wfC by auto + hence "\; \; (x, b, TRUE) #\<^sub>\ G1@(x',b',c'')#\<^sub>\G2 \\<^sub>w\<^sub>f c" using * *** wf_replace_inside assms + by (metis "**" append_g.simps(2) wfG_elim2 wfG_suffix) + thus ?thesis using ** * *** by auto + qed +qed + +lemma replace_in_g_valid_weakening: + assumes "\; \; \[x'\c''] \ c'" and "x' \ x" and "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \[x'\c'']" + shows "\; \; ((x, b,c) #\<^sub>\ \)[x'\ c''] \ c'" + apply(subst replace_in_g_dist,simp add: assms,rule valid_weakening) + using assms by auto+ + +lemma replace_in_g_subtyped_cons: + assumes "replace_in_g_subtyped \ \ G xcs G'" and "wfG \ \ ((x,b,c)#\<^sub>\G)" + shows "x \ fst ` set xcs \ replace_in_g_subtyped \ \ ((x,b,c)#\<^sub>\G) xcs ((x,b,c)#\<^sub>\G')" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI G) + then show ?case + by (simp add: replace_in_g_subtyped.replace_in_g_subtyped_nilI) +next + case (replace_in_g_subtyped_consI b' c' G x' \ \ c'' xcs' G') + hence "\; \ \\<^sub>w\<^sub>f G[x'\c'']" using valid.simps wfC_wf by auto + + show ?case proof(rule replace_in_g_subtyped.replace_in_g_subtyped_consI) + show " Some (b', c') = lookup ((x, b,c) #\<^sub>\ G) x'" using lookup.simps + fst_conv image_iff \_set_intros surj_pair replace_in_g_subtyped_consI by force + show wbc: " \; \; (x, b, c) #\<^sub>\ G \\<^sub>w\<^sub>f c'' " using wf_weakening \ \; \; G \\<^sub>w\<^sub>f c''\ \\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G \ by fastforce + have "x' \ x" using replace_in_g_subtyped_consI by auto + have wbc1: "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G[x'\c'']" proof - + have "(x, b, c) #\<^sub>\ G[x'\c''] = ((x, b, c) #\<^sub>\ G)[x'\c'']" using \x' \ x\ using replace_in_g.simps by auto + thus ?thesis using wfG_replace_inside_rig \\; \ \\<^sub>w\<^sub>f G[x'\c'']\ \\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G \ by fastforce + qed + show *: "\; \; replace_in_g ((x, b,c) #\<^sub>\ G) x' c'' \ c'" + proof - + have "\; \; G[x'\c''] \ c'" using replace_in_g_subtyped_consI by auto + thus ?thesis using replace_in_g_valid_weakening wbc1 \x'\x\ by auto + qed + + show "replace_in_g_subtyped \ \ (replace_in_g ((x, b,c) #\<^sub>\ G) x' c'') xcs' ((x, b,c) #\<^sub>\ G')" + using replace_in_g_subtyped_consI wbc1 by auto + show "x' \ fst ` set xcs'" + using replace_in_g_subtyped_consI by linarith + qed +qed + +lemma replace_in_g_split: + fixes G::\ + assumes "\ = replace_in_g \' x c" and "\' = G'@(x,b,c')#\<^sub>\G" and "wfG \ \ \'" + shows "\ = G'@(x,b,c)#\<^sub>\G" + using assms proof(induct G' arbitrary: G \ \' rule: \_induct) + case GNil + then show ?case by simp +next + case (GCons x1 b1 c1 \1) + hence "x1 \ x" + using wfG_cons_fresh2[of \ \ x1 b1 c1 \1 x b ] + using GCons.prems(2) GCons.prems(3) append_g.simps(2) by auto + moreover hence *: "\; \ \\<^sub>w\<^sub>f (\1 @ (x, b, c') #\<^sub>\ G)" using GCons append_g.simps wfG_elims by metis + moreover hence "replace_in_g (\1 @ (x, b, c') #\<^sub>\ G) x c = \1 @ (x, b, c) #\<^sub>\ G" using GCons replace_in_g_inside[OF *, of c] by auto + + ultimately show ?case using replace_in_g.simps(2)[of x1 b1 c1 " \1 @ (x, b, c') #\<^sub>\ G" x c] GCons + by (simp add: GCons.prems(1) GCons.prems(2)) +qed + +lemma replace_in_g_subtyped_split0: + fixes G::\ + assumes "replace_in_g_subtyped \ \ \'[(x,c)] \" and "\' = G'@(x,b,c')#\<^sub>\G" and "wfG \ \ \'" + shows "\ = G'@(x,b,c)#\<^sub>\G" +proof - + have "\ = replace_in_g \' x c" using assms replace_in_g_subtyped.simps + by (metis Pair_inject list.distinct(1) list.inject) + thus ?thesis using assms replace_in_g_split by blast +qed + +lemma replace_in_g_subtyped_split: + assumes "Some (b, c') = lookup G x" and "\; \; replace_in_g G x c \ c'" and "wfG \ \ G " + shows "\ \ \'. G = \'@(x,b,c')#\<^sub>\\ \ \; \; \'@(x,b,c)#\<^sub>\\ \ c'" +proof - + obtain \ and \' where "G = \'@(x,b,c')#\<^sub>\\" using assms lookup_split by blast + moreover hence "replace_in_g G x c = \'@(x,b,c)#\<^sub>\\" using replace_in_g_split assms by blast + ultimately show ?thesis by (metis assms(2)) +qed + +section \Validity and Subtyping\ + +lemma wfC_replace_in_g: + fixes c::c and c0::c + assumes "\; \; \'@(x,b,c0')#\<^sub>\\ \\<^sub>w\<^sub>f c" and "\; \; (x,b,TRUE)#\<^sub>\\ \\<^sub>w\<^sub>f c0" + shows "\; \; \' @ (x, b, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f c" + using wf_replace_inside1(2) assms by auto + + +lemma ctx_subtype_valid: + assumes "\; \; \'@(x,b,c0')#\<^sub>\\ \ c" and + "\; \; \'@(x,b,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@(x,b,c0)#\<^sub>\\ \ c" +proof(rule validI) + show "\; \; \' @ (x, b, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f c" proof - + have "\; \; \'@(x,b,c0')#\<^sub>\\ \\<^sub>w\<^sub>f c" using valid.simps assms by auto + moreover have "\; \; (x,b,TRUE)#\<^sub>\\ \\<^sub>w\<^sub>f c0" proof - + have "wfG \ \ (\'@(x,b,c0)#\<^sub>\\)" using assms valid.simps wfC_wf by auto + hence "wfG \ \ ((x,b,c0)#\<^sub>\\)" using wfG_suffix by auto + thus ?thesis using wfG_wfC by auto + qed + ultimately show ?thesis using assms wfC_replace_in_g by auto + qed + + show "\i. wfI \ (\' @ (x, b, c0) #\<^sub>\ \) i \ is_satis_g i (\' @ (x, b, c0) #\<^sub>\ \) \ is_satis i c" proof(rule,rule) + fix i + assume * : "wfI \ (\' @ (x, b, c0) #\<^sub>\ \) i \ is_satis_g i (\' @ (x, b, c0) #\<^sub>\ \) " + + hence "is_satis_g i (\'@(x, b, c0) #\<^sub>\ \) \ wfI \ (\'@(x, b, c0) #\<^sub>\ \) i" using is_satis_g_append wfI_suffix by metis + moreover hence "is_satis i c0'" using valid.simps assms by presburger + + moreover have "is_satis_g i \'" using is_satis_g_append * by simp + ultimately have "is_satis_g i (\' @ (x, b, c0') #\<^sub>\ \) " using is_satis_g_append by simp + moreover have "wfI \ (\' @ (x, b, c0') #\<^sub>\ \) i" using wfI_def wfI_suffix * wfI_def wfI_replace_inside by metis + ultimately show "is_satis i c" using assms valid.simps by metis + qed +qed + +lemma ctx_subtype_subtype: + fixes \::\ + shows "\; \; G \ t1 \ t2 \ G = \'@(x,b0,c0')#\<^sub>\\ \ \; \; \'@(x,b0,c0)#\<^sub>\\ \ c0' \ \; \; \'@(x,b0,c0)#\<^sub>\\ \ t1 \ t2" +proof(nominal_induct avoiding: c0 rule: subtype.strong_induct) + + case (subtype_baseI x' \ \ \'' z c z' c' b) + let ?\c0 = "\'@(x,b0,c0)#\<^sub>\\" + have wb1: "wfG \ \ ?\c0" using valid.simps wfC_wf subtype_baseI by metis + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ z : b | c \ \ using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ z' : b | c' \ \ using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis + have "atom x' \ \' @ (x, b0, c0) #\<^sub>\ \" using fresh_prodN subtype_baseI fresh_replace_inside wb1 subtype_wf wfX_wfY by metis + thus \atom x' \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \, z, c , z' , c' )\ using subtype_baseI fresh_prodN by metis + have "\; \; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \ \ c'[z'::=V_var x']\<^sub>v " proof(rule ctx_subtype_valid) + show 1: \\; \; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0') #\<^sub>\ \ \ c'[z'::=V_var x']\<^sub>v \ + using subtype_baseI append_g.simps subst_defs by metis + have *:"\; \ \\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \ " proof(rule wfG_replace_inside2) + show "\; \ \\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0') #\<^sub>\ \" + using * valid_wf_all wfC_wf 1 append_g.simps by metis + show "\; \ \\<^sub>w\<^sub>f (x, b0, c0) #\<^sub>\ \" using wfG_suffix wb1 by auto + qed + moreover have "toSet (\' @ (x, b0, c0) #\<^sub>\ \) \ toSet (((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \)" using toSet.simps append_g.simps by auto + ultimately show \\; \; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \ \ c0' \ using valid_weakening subtype_baseI * by blast + qed + thus \\; \; (x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \' @ (x, b0, c0) #\<^sub>\ \ \ c'[z'::=V_var x']\<^sub>v \ using append_g.simps subst_defs by simp + qed +qed + +lemma ctx_subtype_subtype_rig: + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and "\; \; \' \ t1 \ t2" + shows "\; \; \ \ t1 \ t2" +proof - + have wf: "wfG \ \ \'" using subtype_g_wf assms by auto + obtain b and c0' where "Some (b, c0') = lookup \' x \ (\; \; replace_in_g \' x c0 \ c0')" using + replace_in_g_subtyped.simps[of \ \ \' "[(x, c0)]" \] assms(1) + + by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair) + moreover then obtain G and G' where *: "\' = G'@(x,b,c0')#\<^sub>\G \ \; \; G'@(x,b,c0)#\<^sub>\G \ c0'" + using replace_in_g_subtyped_split[of b c0' \' x \ \ c0] wf by metis + + ultimately show ?thesis using ctx_subtype_subtype + assms(1) assms(2) replace_in_g_subtyped_split0 subtype_g_wf + by (metis (no_types, lifting) local.wf replace_in_g_split) +qed + +text \ We now prove versions of the @{text "ctx_subtype"} lemmas above using @{text "replace_in_g"}. First we do case where +the replace is just for a single variable (indicated by suffix rig) and then the general case for +multiple replacements (indicated by suffix rigs)\ + +lemma ctx_subtype_subtype_rigs: + assumes "replace_in_g_subtyped \ \ \' xcs \" and "\; \; \' \ t1 \ t2" + shows "\; \; \ \ t1 \ t2" + using assms proof(induct xcs arbitrary: \ \' ) + case Nil + moreover have "\' = \" using replace_in_g_subtyped_nilI + using calculation(1) by blast + ultimately show ?case by auto +next + case (Cons a xcs) + + then obtain x and c where "a=(x,c)" by fastforce + then obtain b and c' where bc: "Some (b, c') = lookup \' x \ + replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \ \ \; \; \' \\<^sub>w\<^sub>f c \ + x \ fst ` set xcs \ \; \; (replace_in_g \' x c) \ c' " using replace_in_g_subtyped_elims(3)[of \ \ \' x c xcs \] Cons + by (metis valid.simps) + + hence *: "replace_in_g_subtyped \ \ \' [(x,c)] (replace_in_g \' x c)" using replace_in_g_subtyped_consI + by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI) + + hence "\; \; (replace_in_g \' x c) \ t1 \ t2" + using ctx_subtype_subtype_rig * assms Cons.prems(2) by auto + + moreover have "replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \" using Cons + using bc by blast + + ultimately show ?case using Cons by blast +qed + +lemma replace_in_g_inside_valid: + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and "wfG \ \ \'" + shows "\b c0' G G'. \' = G' @ (x,b,c0')#\<^sub>\G \ \ = G' @ (x,b,c0)#\<^sub>\G \ \; \; G'@ (x,b,c0)#\<^sub>\G \ c0'" +proof - + obtain b and c0' where bc: "Some (b, c0') = lookup \' x \ \; \; replace_in_g \' x c0 \ c0'" using + replace_in_g_subtyped.simps[of \ \ \' "[(x, c0)]" \] assms(1) + by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair) + then obtain G and G' where *: "\' = G'@(x,b,c0')#\<^sub>\G \ \; \; G'@(x,b,c0)#\<^sub>\G \ c0'" using replace_in_g_subtyped_split[of b c0' \' x \ \ c0] assms + by metis + thus ?thesis using replace_in_g_inside bc + using assms(1) assms(2) by blast +qed + +lemma replace_in_g_valid: + assumes "\; \ \ G \ xcs \ \ G'" and "\; \; G \ c " + shows \\; \; G' \ c \ + using assms proof(induct rule: replace_in_g_subtyped.inducts) + case (replace_in_g_subtyped_nilI \ \ G) + then show ?case by auto +next + case (replace_in_g_subtyped_consI b c1 G x \ \ c2 xcs G') + hence "\; \; G[x\c2] \ c" + by (metis ctx_subtype_valid replace_in_g_split replace_in_g_subtyped_split valid_g_wf) + then show ?case using replace_in_g_subtyped_consI by auto +qed + +section \Literals\ + +section \Values\ + +lemma lookup_inside_unique_b[simp]: + assumes "\ ; B \\<^sub>w\<^sub>f (\'@(x,b0,c0)#\<^sub>\\)" and "\ ; B \\<^sub>w\<^sub>f (\'@(x,b0,c0')#\<^sub>\\)" + and "Some (b, c) = lookup (\' @ (x, b0, c0') #\<^sub>\ \) y" and "Some (b0,c0) = lookup (\'@((x,b0,c0))#\<^sub>\\) x" and "x=y" + shows "b = b0" + by (metis assms(2) assms(3) assms(5) lookup_inside_wf old.prod.exhaust option.inject prod.inject) + +lemma ctx_subtype_v_aux: + fixes v::v + assumes "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1" + using assms proof(nominal_induct "\'@((x,b0,c0')#\<^sub>\\)" v t1 avoiding: c0 rule: infer_v.strong_induct) + case (infer_v_varI \ \ b c xa z) + have wf:\ \; \ \\<^sub>w\<^sub>f \' @ (x, b0, c0) #\<^sub>\ \ \ using wfG_inside_valid2 infer_v_varI by metis + have xf1:\atom z \ xa\ using infer_v_varI by metis + have xf2: \atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \)\ apply( fresh_mth add: infer_v_varI ) + using fresh_def infer_v_varI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show ?case proof (cases "x=xa") + case True + moreover have "b = b0" using infer_v_varI True by simp + moreover hence \Some (b, c0) = lookup (\' @ (x, b0, c0) #\<^sub>\ \) xa\ using lookup_inside_wf[OF wf] infer_v_varI True by auto + ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by metis + next + case False + moreover hence \Some (b, c) = lookup (\' @ (x, b0, c0) #\<^sub>\ \) xa\ using lookup_inside2 infer_v_varI by metis + ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by simp + qed +next + case (infer_v_litI \ \ l \) + thus ?case using Typing.infer_v_litI wfG_inside_valid2 by simp +next + case (infer_v_pairI z v1 v2 \ \ t1' t2' c0) + show ?case proof + show "atom z \ (v1, v2)" using infer_v_pairI fresh_Pair by simp + show "atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \)" apply( fresh_mth add: infer_v_pairI ) + using fresh_def infer_v_pairI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show "\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ t1'" using infer_v_pairI by simp + show "\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ t2'" using infer_v_pairI by simp + qed +next + case (infer_v_consI s dclist \ dc tc \ v tv z) + show ?case proof + show \AF_typedef s dclist \ set \\ using infer_v_consI by auto + show \(dc, tc) \ set dclist\ using infer_v_consI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ tv\ using infer_v_consI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ tv \ tc\ using infer_v_consI ctx_subtype_subtype by auto + show \atom z \ v\ using infer_v_consI by auto + show \atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \)\ apply( fresh_mth add: infer_v_consI ) + using fresh_def infer_v_consI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + qed +next + case (infer_v_conspI s bv dclist \ dc tc \ v tv b z) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show \(dc, tc) \ set dclist\ using infer_v_conspI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ tv\ using infer_v_conspI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ tv \ tc[bv::=b]\<^sub>\\<^sub>b\ using infer_v_conspI ctx_subtype_subtype by auto + show \atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \, v, b)\ apply( fresh_mth add: infer_v_conspI ) + using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show \atom bv \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \, v, b)\ apply( fresh_mth add: infer_v_conspI ) + using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show \ \; \ \\<^sub>w\<^sub>f b \ using infer_v_conspI by auto + qed +qed + +lemma ctx_subtype_v: + fixes v::v + assumes "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\t2. \; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t2 \ \; \; \'@((x,b0,c0)#\<^sub>\\) \ t2 \ t1" +proof - + + have "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1 " using ctx_subtype_v_aux assms by auto + moreover hence "\; \; \'@((x,b0,c0)#\<^sub>\\) \ t1 \ t1" using subtype_reflI2 infer_v_wf by simp + ultimately show ?thesis by auto +qed + +lemma ctx_subtype_v_eq: + fixes v::v + assumes + "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and + " \; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1" +proof - + obtain t1' where "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1'" using ctx_subtype_v assms by metis + moreover have "replace_in_g (\'@((x,b0,c0')#\<^sub>\\)) x c0 = \'@((x,b0,c0)#\<^sub>\\)" using replace_in_g_inside infer_v_wf assms by metis + ultimately show ?thesis using infer_v_uniqueness_rig assms by metis +qed + +lemma ctx_subtype_check_v_eq: + assumes "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and " \; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1" +proof - + obtain t2 where t2: "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t2 \ \; \; \'@((x,b0,c0')#\<^sub>\\) \ t2 \ t1" + using check_v_elims assms by blast + hence t3: "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t2" + using assms ctx_subtype_v_eq by blast + + have "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t2" using t3 by auto + moreover have " \; \; \'@((x,b0,c0)#\<^sub>\\) \ t2 \ t1" proof - + + have " \; \; \'@((x,b0,c0')#\<^sub>\\) \ t2 \ t1" using t2 by auto + thus ?thesis using subtype_trans + using assms(2) ctx_subtype_subtype by blast + qed + ultimately show ?thesis using check_v.intros by presburger +qed + +text \ Basically the same as @{text "ctx_subtype_v_eq"} but in a different form\ +lemma ctx_subtype_v_rig_eq: + fixes v::v + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and + "\; \; \' \ v \ t1" + shows "\; \; \ \ v \ t1" +proof - + obtain b and c0' and G and G' where "\' = G' @ (x,b,c0')#\<^sub>\G \ \ = G' @ (x,b,c0)#\<^sub>\G \ \; \; G'@ (x,b,c0)#\<^sub>\G \ c0'" + using assms replace_in_g_inside_valid infer_v_wf by metis + thus ?thesis using ctx_subtype_v_eq[of \ \ G' x b c0' G v t1 c0] assms by simp +qed + +lemma ctx_subtype_v_rigs_eq: + fixes v::v + assumes "replace_in_g_subtyped \ \ \' xcs \" and + "\; \; \' \ v \ t1" + shows "\; \; \ \ v \ t1" + using assms proof(induct xcs arbitrary: \ \' t1 ) + case Nil + then show ?case by auto +next + case (Cons a xcs) + then obtain x and c where "a=(x,c)" by fastforce + + then obtain b and c' where bc: "Some (b, c') = lookup \' x \ + replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \ \ \; \; \' \\<^sub>w\<^sub>f c \ + x \ fst ` set xcs \ \; \; (replace_in_g \' x c) \ c' " + using replace_in_g_subtyped_elims(3)[of \ \ \' x c xcs \] Cons by (metis valid.simps) + + hence *: "replace_in_g_subtyped \ \ \' [(x,c)] (replace_in_g \' x c)" using replace_in_g_subtyped_consI + by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI) + hence t2: "\; \; (replace_in_g \' x c) \ v \ t1 " using ctx_subtype_v_rig_eq[OF * Cons(3)] by blast + moreover have **: "replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \" using bc by auto + ultimately have t2': "\; \; \ \ v \ t1" using Cons by blast + thus ?case by blast +qed + +lemma ctx_subtype_check_v_rigs_eq: + assumes "replace_in_g_subtyped \ \ \' xcs \" and + "\; \; \' \ v \ t1" + shows "\; \; \ \ v \ t1" +proof - + obtain t2 where "\; \; \' \ v \ t2 \ \; \; \' \ t2 \ t1" using check_v_elims assms by fast + hence "\; \; \ \ v \ t2 \ \; \; \ \ t2 \ t1" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs + using assms(1) by blast + thus ?thesis + using check_v_subtypeI by blast +qed + +section \Expressions\ + +lemma valid_wfC: + fixes c0::c + assumes "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; (x, b0, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c0" + using wfG_elim2 valid.simps wfG_suffix + using assms valid_g_wf by metis + +lemma ctx_subtype_e_eq: + fixes G::\ + assumes + "\ ; \ ; \ ; G ; \ \ e \ t1" and "G = \'@((x,b0,c0')#\<^sub>\\)" + "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\ ; \ ; \ ; \'@((x,b0,c0)#\<^sub>\\) ; \ \ e \ t1" + using assms proof(nominal_induct t1 avoiding: c0 rule: infer_e.strong_induct) + case (infer_e_valI \ \ \'' \ \ v \) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_valI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_valI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \\ using infer_e_valI ctx_subtype_v_eq by auto + qed +next + case (infer_e_plusI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_plusI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_plusI by auto + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_int | c1 \\ using infer_e_plusI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : B_int | c2 \\ using infer_e_plusI ctx_subtype_v_eq by auto + show \atom z3 \ AE_op Plus v1 v2\ using infer_e_plusI by auto + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using * infer_e_plusI fresh_replace_inside infer_v_wf by metis + qed +next + case (infer_e_leqI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_leqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_leqI by auto + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_int | c1 \\ using infer_e_leqI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : B_int | c2 \\ using infer_e_leqI ctx_subtype_v_eq by auto + show \atom z3 \ AE_op LEq v1 v2\ using infer_e_leqI by auto + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using * infer_e_leqI fresh_replace_inside infer_v_wf by metis + qed +next + case (infer_e_eqI \ \ \'' \ \ v1 z1 bb c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_eqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_eqI by auto + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : bb | c1 \\ using infer_e_eqI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : bb | c2 \\ using infer_e_eqI ctx_subtype_v_eq by auto + show \atom z3 \ AE_op Eq v1 v2\ using infer_e_eqI by auto + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using * infer_e_eqI fresh_replace_inside infer_v_wf by metis + show "bb \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed +next + case (infer_e_appI \ \ \'' \ \ f x' b c \' s' v \) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_appI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appI by auto + show \Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \' s'))) = lookup_fun \ f\ using infer_e_appI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ x' : b | c \\ using infer_e_appI ctx_subtype_check_v_eq by auto + thus \atom x' \ (\, \, \, \' @ (x, b0, c0) #\<^sub>\ \, \, v, \)\ + using infer_e_appI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 x'] infer_v_wf by auto + show \\'[x'::=v]\<^sub>v = \\ using infer_e_appI by auto + qed +next + case (infer_e_appPI \ \ \1 \ \ b' f bv x1 b c \' s' v \) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_appPI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appPI by auto + show \ \; \ \\<^sub>w\<^sub>f b' \ using infer_e_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b c \' s'))) = lookup_fun \ f\ using infer_e_appPI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ x1 : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \\ using infer_e_appPI ctx_subtype_check_v_eq subst_defs by auto + thus \atom x1 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 x1 ] infer_v_wf infer_e_appPI by auto + show \\'[bv::=b']\<^sub>b[x1::=v]\<^sub>v = \\ using infer_e_appPI by auto + have "atom bv \ \' @ (x, b0, c0') #\<^sub>\ \" using infer_e_appPI by metis + hence "atom bv \ \' @ (x, b0, c0) #\<^sub>\ \" + unfolding fresh_append_g fresh_GCons fresh_prod3 using \atom bv \ c0 \ fresh_append_g by metis + thus \atom bv \ (\, \, \, \' @ (x, b0, c0) #\<^sub>\ \, \, b', v, \)\ using infer_e_appPI by auto + qed +next + case (infer_e_fstI \ \ \'' \ \ v z' b1 b2 c z) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_fstI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_fstI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ z' : B_pair b1 b2 | c \\ using infer_e_fstI ctx_subtype_v_eq by auto + thus \atom z \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_fstI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z] infer_v_wf by auto + show \atom z \ AE_fst v\ using infer_e_fstI by auto + qed +next + case (infer_e_sndI \ \ \'' \ \ v z' b1 b2 c z) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_sndI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_sndI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ z' : B_pair b1 b2 | c \\ using infer_e_sndI ctx_subtype_v_eq by auto + thus \atom z \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_sndI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z] infer_v_wf by auto + show \atom z \ AE_snd v\ using infer_e_sndI by auto + qed +next + case (infer_e_lenI \ \ \'' \ \ v z' c z) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_lenI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_lenI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ z' : B_bitvec | c \\ using infer_e_lenI ctx_subtype_v_eq by auto + thus \atom z \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_lenI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z] infer_v_wf by auto + show \atom z \ AE_len v\ using infer_e_lenI by auto + qed +next + case (infer_e_mvarI \ \ \'' \ \ u \) + show ?case proof + show "\; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \" using wf_replace_inside2(6) valid_wfC infer_e_mvarI by auto + thus "\; \ \\<^sub>w\<^sub>f \' @ (x, b0, c0) #\<^sub>\ \" using infer_e_mvarI fresh_replace_inside wfD_wf by blast + show "\ \\<^sub>w\<^sub>f \ " using infer_e_mvarI by auto + show "(u, \) \ setD \" using infer_e_mvarI by auto + qed +next + case (infer_e_concatI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_concatI by auto + thus \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_concatI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z3] infer_v_wf wfX_wfY by metis + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_concatI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_bitvec | c1 \\ using infer_e_concatI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : B_bitvec | c2 \\ using infer_e_concatI ctx_subtype_v_eq by auto + show \atom z3 \ AE_concat v1 v2\ using infer_e_concatI by auto + qed +next + case (infer_e_splitI \ \ \'' \ \ v1 z1 c1 v2 z2 z3) + show ?case proof + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_splitI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_splitI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_bitvec | c1 \\ using infer_e_splitI ctx_subtype_v_eq by auto + show \\; \; \' @ + (x, b0, c0) #\<^sub>\ + \ \ v2 \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND + [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \\ + using infer_e_splitI ctx_subtype_check_v_eq by auto + + show \atom z1 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z1] infer_e_splitI infer_v_wf wfX_wfY * by metis + show \atom z2 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis + show \atom z1 \ AE_split v1 v2\ using infer_e_splitI by auto + show \atom z2 \ AE_split v1 v2\ using infer_e_splitI by auto + show \atom z3 \ AE_split v1 v2\ using infer_e_splitI by auto + qed +qed + +lemma ctx_subtype_e_rig_eq: + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and + "\ ; \ ; \ ; \' ; \ \ e \ t1" + shows "\ ; \ ; \ ; \ ; \ \ e \ t1" +proof - + obtain b and c0' and G and G' where "\' = G' @ (x,b,c0')#\<^sub>\G \ \ = G' @ (x,b,c0)#\<^sub>\G \ \; \; G'@ (x,b,c0)#\<^sub>\G \ c0'" + using assms replace_in_g_inside_valid infer_e_wf by meson + thus ?thesis + using assms ctx_subtype_e_eq by presburger +qed + +lemma ctx_subtype_e_rigs_eq: + assumes "replace_in_g_subtyped \ \ \' xcs \" and + "\ ; \ ; \ ; \'; \ \ e \ t1" + shows "\ ; \ ; \ ; \ ; \ \ e \ t1" + using assms proof(induct xcs arbitrary: \ \' t1 ) + case Nil + moreover have "\' = \" using replace_in_g_subtyped_nilI + using calculation(1) by blast + moreover have "\;\;\ \ t1 \ t1" using subtype_reflI2 Nil infer_e_t_wf by blast + ultimately show ?case by blast +next + case (Cons a xcs) + + then obtain x and c where "a=(x,c)" by fastforce + then obtain b and c' where bc: "Some (b, c') = lookup \' x \ + replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \ \ \; \; \' \\<^sub>w\<^sub>f c \ + x \ fst ` set xcs \ \; \; (replace_in_g \' x c) \ c' " using replace_in_g_subtyped_elims(3)[of \ \ \' x c xcs \] Cons + by (metis valid.simps) + + hence *: "replace_in_g_subtyped \ \ \' [(x,c)] (replace_in_g \' x c)" using replace_in_g_subtyped_consI + by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI) + hence t2: "\ ; \ ; \ ; (replace_in_g \' x c) ; \ \ e \ t1 " using ctx_subtype_e_rig_eq[OF * Cons(3)] by blast + moreover have **: "replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \" using bc by auto + ultimately have t2': "\ ; \ ; \ ; \ ; \ \ e \ t1" using Cons by blast + thus ?case by blast +qed + +section \Statements\ + +lemma ctx_subtype_s_rigs: + fixes c0::c and s::s and G'::\ and xcs :: "(x*c) list" and css::branch_list + shows + "check_s \ \ \ G \ s t1 \ wsX G xcs \ replace_in_g_subtyped \ \ G xcs G' \ check_s \ \ \ G' \ s t1" and + "check_branch_s \ \ \ G \ tid cons const v cs t1 \ wsX G xcs \ replace_in_g_subtyped \ \ G xcs G' \ check_branch_s \ \ \ G' \ tid cons const v cs t1" + "check_branch_list \ \ \ G \ tid dclist v css t1 \ wsX G xcs \ replace_in_g_subtyped \ \ G xcs G' \ check_branch_list \ \ \ G' \ tid dclist v css t1" +proof(induction arbitrary: xcs G' and xcs G' and xcs G' rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ \ \ \ \ v \' \) + hence *:"\; \; G' \ v \ \' \ \; \; G' \ \' \ \" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs + by (meson check_v.simps) + show ?case proof + show \\; \; G' \\<^sub>w\<^sub>f \\ using check_valI wfD_rig by auto + show \\ \\<^sub>w\<^sub>f \ \ using check_valI by auto + show \\; \; G' \ v \ \'\ using * by auto + show \\; \; G' \ \' \ \\ using * by auto + qed +next + case (check_letI x \ \ \ \ \ e \ z' s b' c') + + show ?case proof + have wfG: "\; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f G'" using infer_e_wf check_letI replace_in_g_wfG using infer_e_wf(2) by (auto simp add: freshers) + hence "atom x \ G'" using check_letI replace_in_g_fresh replace_in_g_wfG by auto + thus "atom x \ (\, \, \, G', \, e, \)" using check_letI by auto + have "atom z' \ G'" apply(rule replace_in_g_fresh[OF check_letI(7)]) + using replace_in_g_wfG check_letI fresh_prodN infer_e_wf by metis+ + thus "atom z' \ (x, \, \, \, G', \, e, \, s)" using check_letI fresh_prodN by metis + + show "\ ; \ ; \ ; G' ; \ \ e \ \ z' : b' | c' \" + using check_letI ctx_subtype_e_rigs_eq by blast + show "\ ; \ ; \ ; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G' ; \ \ s \ \" + proof(rule check_letI(5)) + have vld: "\;\;((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) \ c'[z'::=V_var x]\<^sub>c\<^sub>v" proof - + have "wfG \ \ ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \)" using check_letI check_s_wf by metis + hence "wfC \ \ ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) (c'[z'::=V_var x]\<^sub>c\<^sub>v)" using wfC_refl subst_defs by auto + thus ?thesis using valid_reflI[of \ \ x b' "c'[z'::=V_var x]\<^sub>v" \ " c'[z'::=V_var x]\<^sub>v"] subst_defs by auto + qed + have xf: "x \ fst ` set xcs" proof - + have "atom ` fst ` set xcs \ atom_dom \" using check_letI wsX_iff by meson + moreover have "wfG \ \ \" using infer_e_wf check_letI by metis + ultimately show ?thesis using fresh_def check_letI wfG_dom_supp + using wsX_fresh by auto + qed + show "replace_in_g_subtyped \ \ ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) ((x, c'[z'::=V_var x]\<^sub>v) # xcs) ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G')" proof - + + have "Some (b', c'[z'::=V_var x]\<^sub>v) = lookup ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x" by auto + + moreover have "\; \; replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v) \ c'[z'::=V_var x]\<^sub>v" proof - + have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v) = ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \)" + using replace_in_g.simps by presburger + thus ?thesis using vld subst_defs by auto + qed + + moreover have " replace_in_g_subtyped \ \ (replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G'))" proof - + have "wfG \ \ ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \))" using check_letI check_s_wf by metis + hence "replace_in_g_subtyped \ \ ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G'))" + using check_letI replace_in_g_subtyped_cons xf by meson + moreover have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v) = ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \))" + using replace_in_g.simps by presburger + ultimately show ?thesis by argo + qed + moreover have "\; \; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>v " using vld subst_defs by auto + ultimately show ?thesis using replace_in_g_subtyped_consI xf replace_in_g.simps(2) by metis + qed + + show " wsX ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) ((x, c'[z'::=V_var x]\<^sub>v) # xcs)" + using check_letI xf subst_defs by (simp add: wsX_cons) + qed + qed + +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case using Typing.check_branch_list_consI by auto +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case using Typing.check_branch_list_finalI by auto +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + + have wfcons: "wfG \ \ ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \)" using check_s_wf check_branch_s_branchI + by meson + hence wf: "wfG \ \ \" using wfG_cons by metis + + moreover have "atom x \ (const, G',v)" proof - + have "atom x \ G'" using check_branch_s_branchI wf replace_in_g_fresh + wfG_dom_supp replace_in_g_wfG by simp + thus ?thesis using check_branch_s_branchI fresh_prodN by simp + qed + + moreover have st: "\ ; \ ; \ ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ G' ; \ \ s \ \ " proof - + have " wsX ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) xcs" using check_branch_s_branchI wsX_cons2 wsX_fresh wf by force + moreover have "replace_in_g_subtyped \ \ ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) xcs ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ G' )" + using replace_in_g_subtyped_cons wsX_fresh wf check_branch_s_branchI wfcons by auto + thus ?thesis using check_branch_s_branchI calculation by meson + qed + moreover have wft: " wfT \ \ G' \ " using + check_branch_s_branchI ctx_subtype_subtype_rigs subtype_reflI2 subtype_wf by metis + moreover have "wfD \ \ G' \" using check_branch_s_branchI wfD_rig by presburger + ultimately show ?case using + Typing.check_branch_s_branchI + using check_branch_s_branchI.hyps by simp + +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + hence wf:"wfG \ \ \" using check_s_wf by presburger + show ?case proof(rule check_s_check_branch_s_check_branch_list.check_ifI) + show \atom z \ (\, \, \, G', \, v, s1, s2, \)\ using fresh_prodN replace_in_g_fresh1 wf check_ifI by auto + show \\; \; G' \ v \ \ z : B_bool | TRUE \\ using ctx_subtype_check_v_rigs_eq check_ifI by presburger + show \ \ ; \ ; \ ; G' ; \ \ s1 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_true) IMP c_of \ z \\ using check_ifI by auto + show \ \ ; \ ; \ ; G' ; \ \ s2 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \\ using check_ifI by auto + qed +next + + case (check_let2I x P \ \ G \ t s1 \ s2 ) + show ?case proof + have "wfG P \ G" using check_let2I check_s_wf by metis + show *: "P ; \ ; \ ; G' ; \ \ s1 \t" using check_let2I by blast + show "atom x \ (P, \, \, G', \, t, s1, \)" proof - + have "wfG P \ G'" using check_s_wf * by blast + hence "atom_dom G = atom_dom G'" using check_let2I rigs_atom_dom_eq by presburger + moreover have "atom x \ G" using check_let2I by auto + moreover have "wfG P \ G" using check_s_wf * replace_in_g_wfG check_let2I by simp + ultimately have "atom x \ G'" using wfG_dom_supp fresh_def \wfG P \ G'\ by metis + thus ?thesis using check_let2I by auto + qed + show "P ; \ ; \ ;(x, b_of t, c_of t x) #\<^sub>\ G' ; \ \ s2 \ \ " proof - + have "wsX ((x, b_of t, c_of t x) #\<^sub>\ G) xcs" using check_let2I wsX_cons2 wsX_fresh \wfG P \ G\ by simp + moreover have "replace_in_g_subtyped P \ ((x, b_of t, c_of t x) #\<^sub>\ G) xcs ((x, b_of t, c_of t x) #\<^sub>\ G')" proof(rule replace_in_g_subtyped_cons ) + show "replace_in_g_subtyped P \ G xcs G'" using check_let2I by auto + have "atom x \ G" using check_let2I by auto + moreover have "wfT P \ G t" using check_let2I check_s_wf by metis + + moreover have "atom x \ t" using check_let2I check_s_wf wfT_supp by auto + ultimately show "wfG P \ ((x, b_of t, c_of t x) #\<^sub>\ G)" using wfT_wf_cons b_of_c_of_eq[of x t] by auto + show "x \ fst ` set xcs" using check_let2I wsX_fresh \wfG P \ G\ by simp + qed + ultimately show ?thesis using check_let2I by presburger + qed + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof + have "atom u \ G'" unfolding fresh_def + apply(rule u_not_in_g , rule replace_in_g_wfG) + using check_v_wf check_varI by simp+ + thus \atom u \ (\, \, \, G', \, \', v, \)\ unfolding fresh_prodN using check_varI by simp + show \\; \; G' \ v \ \'\ using ctx_subtype_check_v_rigs_eq check_varI by auto + show \ \ ; \ ; \ ; G' ; (u, \') #\<^sub>\ \ \ s \ \\ using check_varI by auto + qed +next + case (check_assignI P \ \ G \ u \ v z \') + show ?case proof + show \P \\<^sub>w\<^sub>f \ \ using check_assignI by auto + show \P ; \ ; G' \\<^sub>w\<^sub>f \ \ using check_assignI wfD_rig by auto + show \(u, \) \ setD \\ using check_assignI by auto + show \P ; \ ; G' \ v \ \\ using ctx_subtype_check_v_rigs_eq check_assignI by auto + show \P ; \ ; G' \ \ z : B_unit | TRUE \ \ \'\ using ctx_subtype_subtype_rigs check_assignI by auto + qed +next + case (check_whileI \ G P s1 z s2 \') + then show ?case using Typing.check_whileI + by (meson ctx_subtype_subtype_rigs) +next + case (check_seqI \ G P s1 z s2 \) + then show ?case + using check_s_check_branch_s_check_branch_list.check_seqI by blast +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + show ?case proof + show " \ ; \ ; \ ; G' ; \ ; tid ; dclist ; v \ cs \ \" using check_caseI ctx_subtype_check_v_rigs_eq by auto + show "AF_typedef tid dclist \ set \" using check_caseI by auto + show "\; \; G' \ v \ \ z : B_id tid | TRUE \" using check_caseI ctx_subtype_check_v_rigs_eq by auto + show "\\<^sub>w\<^sub>f \ " using check_caseI by auto + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + have wfG: "\; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f G'" using check_s_wf check_assertI replace_in_g_wfG wfX_wfY by metis + hence "atom x \ G'" using check_assertI replace_in_g_fresh replace_in_g_wfG by auto + thus \atom x \ (\, \, \, G', \, c, \, s)\ using check_assertI fresh_prodN by auto + show \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ G' ; \ \ s \ \\ proof(rule check_assertI(5) ) + show "wsX ((x, B_bool, c) #\<^sub>\ \) xcs" using check_assertI wsX_cons3 by simp + show "\; \ \ (x, B_bool, c) #\<^sub>\ \ \ xcs \ \ (x, B_bool, c) #\<^sub>\ G'" proof(rule replace_in_g_subtyped_cons) + show \ \; \ \ \ \ xcs \ \ G'\ using check_assertI by auto + show \ \; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \ \ using check_assertI check_s_wf by metis + thus \x \ fst ` set xcs\ using check_assertI wsX_fresh wfG_elims wfX_wfY by metis + qed + qed + show \\; \; G' \ c \ using check_assertI replace_in_g_valid by auto + show \ \; \; G' \\<^sub>w\<^sub>f \ \ using check_assertI wfD_rig by auto + qed +qed + +lemma replace_in_g_subtyped_empty: + assumes "wfG \ \ (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + shows "replace_in_g_subtyped \ \ (replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" +proof - + have "replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) = (\' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + using assms proof(induct \' rule: \_induct) + case GNil + then show ?case using replace_in_g.simps by auto + next + case (GCons x1 b1 c1 \1) + have "x \ fst ` toSet ((x1,b1,c1)#\<^sub>\\1)" using GCons wfG_inside_fresh atom_dom.simps dom.simps toSet.simps append_g.simps by fast + hence "x1 \ x" using assms wfG_inside_fresh GCons by force + hence "((x1,b1,c1) #\<^sub>\ (\1 @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \))[x\c'[z'::=V_var x]\<^sub>c\<^sub>v] = (x1,b1,c1) #\<^sub>\ (\1 @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + using replace_in_g.simps GCons wfG_elims append_g.simps by metis + thus ?case using append_g.simps by simp + qed + thus ?thesis using replace_in_g_subtyped_nilI by presburger +qed + +lemma ctx_subtype_s: + fixes s::s + assumes "\ ; \ ; \ ; \'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\) ; \ \ s \ \" and + "\; \; \ \ \ z' : b | c' \ \ \ z : b | c \" and + "atom x \ (z,z',c,c')" + shows "\ ; \ ; \ ; \'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\ ; \ \ s \ \" +proof - + + have wf: "wfG \ \ (\'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\))" using check_s_wf assms by meson + hence *:"x \ fst ` toSet \'" using wfG_inside_fresh by force + have "wfG \ \ ((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)" using wf wfG_suffix by metis + hence xfg: "atom x \ \" using wfG_elims by metis + have "x \ z'" using assms fresh_at_base fresh_prod4 by metis + hence a2: "atom x \ c'" using assms fresh_prod4 by metis + + have "atom x \ (z', c', z, c, \)" proof - + have "x \ z" using assms using assms fresh_at_base fresh_prod4 by metis + hence a1 : "atom x \ c" using assms subtype_wf subtype_wf assms wfT_fresh_c xfg by meson + thus ?thesis using a1 a2 \atom x \ (z,z',c,c')\ fresh_prod4 fresh_Pair xfg by simp + qed + hence wc1:" \; \; (x, b, c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \ \ c[z::=V_var x]\<^sub>v" + using subtype_valid assms fresh_prodN by metis + + have vld: "\;\ ; (\'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) \ c[z::=V_var x]\<^sub>c\<^sub>v" proof - + + have "toSet ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) \ toSet (\'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" by auto + moreover have "wfG \ \ (\'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" proof - + have *:"wfT \ \ \ (\ z' : b | c' \)" using subtype_wf assms by meson + moreover have "atom x \ (c',\)" using xfg a2 by simp + ultimately have "wfG \ \ ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" using wfT_wf_cons_flip freshers by blast + thus ?thesis using wfG_replace_inside2 check_s_wf assms by metis + qed + ultimately show ?thesis using wc1 valid_weakening subst_defs by metis + qed + hence wbc: "\; \; \' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using valid.simps by auto + have wbc1: "\; \; (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using wc1 valid.simps subst_defs by auto + have "wsX (\'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)]" proof + show "wsX (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) []" using wsX_NilI by auto + show "atom x \ atom_dom (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" by simp + show "x \ fst ` set []" by auto + qed + moreover have "replace_in_g_subtyped \ \ (\'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)] (\'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)" proof + show "Some (b, c[z::=V_var x]\<^sub>c\<^sub>v) = lookup (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x" using lookup_inside* by auto + show "\; \; replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) \ c[z::=V_var x]\<^sub>c\<^sub>v" using vld replace_in_g_split wf by metis + show "replace_in_g_subtyped \ \ (replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + using replace_in_g_subtyped_empty wf by presburger + show "x \ fst ` set []" by auto + show "\; \; \' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>c\<^sub>v" + proof(rule wf_weakening) + show \\; \; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c'[z'::=[ x ]\<^sup>v]\<^sub>c\<^sub>v \ using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis + show \\; \ \\<^sub>w\<^sub>f \' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis + show \toSet ((x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) \ toSet (\' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)\ using append_g.simps toSet.simps by auto + qed + qed + ultimately show ?thesis using ctx_subtype_s_rigs(1)[OF assms(1)] by presburger +qed + +end \ No newline at end of file diff --git a/thys/MiniSail/Examples.thy b/thys/MiniSail/Examples.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Examples.thy @@ -0,0 +1,162 @@ +(*<*) +theory Examples +imports SubstMethods fuzzyrule.fuzzyrule +begin +(*>*) + +subsection \Examples from thesis\ + +text \ +let x = false in +let y = 0b11 in +let z = 42 in +let w = (true,(z,y)) in () +\ + +method fresh_mth_aux uses add = ( + (match conclusion in "atom z \ GNil" for z \ \ simp add: fresh_GNil[of "atom z"] add\) + | (match conclusion in "atom z \ DNil" for z \ \ simp add: fresh_DNil[of "atom z" ] add\) + | (match conclusion in "atom z \ Nil" for z \ \ simp add: fresh_Nil[of "atom z" ] add\) + | (match conclusion in "atom z \ {||}" for z \ \ simp add: fresh_empty_fset[of "atom z" ] add\) + | (match conclusion in "atom (z::x) \ ((x::x,b::b,c::c)#\<^sub>\G)" for z x b c G \ \simp add: fresh_GCons[of "atom z" "(x,b,c)" G] add\) + | (match conclusion in "atom z \ (b::b)" for z b \ \ simp add: b.fresh[of "atom z" ] add\) + | (match conclusion in "atom z \ (c::c)" for z c \ \ simp add: c.fresh[of "atom z" ] add\) + | (match conclusion in "atom z \ (i::int)" for z i \ \ simp add: pure_fresh add\) +(* tbc delta and types + | (auto simp add: add x_fresh_b pure_fresh) (* Cases where there is no subst and so can most likely get what we want from induction premises *)*) +) + +method fresh_mth_basic uses add = ( + (unfold fresh_prodN, intro conjI)?, + (fresh_mth_aux add: add)+) + +method wf_ctx uses add = ( + (rule wfPhi_emptyI | rule wfTh_emptyI| rule wfD_emptyI| rule wfG_nilI | rule wfG_cons2I )+ +) + +method wfb uses add = ( + (rule wfB_boolI | rule wfB_intI | rule wfB_unitI | rule wfB_pairI )+ +) + +method wfc uses add = ( + (rule wfC_trueI | rule wfC_falseI | rule wfC_notI | rule wfC_eqI )+ +) + +method wfv uses add = ( + (rule wfV_varI | fuzzy_rule wfV_litI | rule wfV_pairI )+ +) + +method wfce uses add = ( + (rule wfCE_valI | rule wfCE_fstI | rule wfCE_sndI | rule wfCE_eqI )+ +) + +method wfe uses add = ( + (rule wfE_valI | rule wfE_fstI | rule wfE_sndI )+ +) + +method wfs uses add = ( + (rule wfS_letI | rule wfS_valI )+ +) + +method wf_all uses add = ( + (wfb | wfv | wfc | wfce | wfe | wfs | auto simp add: add | wf_ctx | fresh_mth_basic add: add )+ +) + +method wf_all2 uses add = ( + (wfb | wfv | wfc | wfce | wfe | wfs | wf_ctx )+ +) + +method wf_let uses add = ( + (rule wfS_letI, rule wfE_valI, wf_ctx add: add ), + (rule wfV_litI,rule wfG_nilI, rule wfTh_emptyI)) + +(* +rule wfS_letI, rule wfE_valI, rule wfPhi_emptyI , rule wfTh_emptyI, rule wfD_emptyI, rule wfG_nilI, rule wfTh_emptyI), + (rule wfV_litI,rule wfG_nilI, rule wfTh_emptyI) +*) + +lemma + shows "[] ; [] ; {||} ; GNil ; []\<^sub>\ \\<^sub>w\<^sub>f LET x = [ [ L_false ]\<^sup>v ]\<^sup>e IN [[ L_unit ]\<^sup>v]\<^sup>s : B_unit" + apply wf_let + apply(rule wfS_valI,wf_ctx) + by(fuzzy_rule wfV_litI, wf_all) + +lemma + assumes " atom z \ x" + shows "[] ; [] ; {||} ; GNil ; []\<^sub>\ \\<^sub>w\<^sub>f LET x = [ [ L_false ]\<^sup>v ]\<^sup>e IN + LET z = [ [ L_num 42]\<^sup>v ]\<^sup>e IN [[ L_unit ]\<^sup>v]\<^sup>s : B_unit" + + + apply wf_let + apply(rule wfS_letI) + + apply(rule wfE_valI, wf_ctx,auto,wf_ctx) + by(wf_all add: assms) + + +text \ +val not_bool : (x : bool [ T ]) -> { z : bool | ~ (z = x) } +function not_bool(x) = if x then F else T + +val eq_int : ( x : int x int [ T ]) -> { z : bool | z = (fst x = snd x ) } +function eq_int ( x ) = + let x1 = fst x in + let x2 = snd x in + let x3 = x1 = x2 in x3 + +val neq_int : ( ( x : int x int [ T ]) -> { z : bool | ~ ( z = (fst x = snd x )) } +function neq_int (x) = not_bool(eq_int(x)) +\ + + +lemma + assumes "atom z \ x" + shows "wfFT [] [] {||} (AF_fun_typ (x::x) B_bool C_true (\ z : B_bool | C_not (C_eq (CE_val (V_var z)) (CE_val (V_var x))) \) + (AS_if (V_var x) [[L_false]\<^sup>v]\<^sup>s [[L_true]\<^sup>v]\<^sup>s))" + apply(rule wfFTI,wf_all) + apply (simp add: supp_at_base)+ + apply(rule wfTI) + by(wf_all add: assms) + +lemma + assumes "atom z \ x" + shows "wfFT [] [] {||} (AF_fun_typ (x::x) (B_pair B_int B_int) C_true (\ z : B_bool | (C_eq (CE_val (V_var z)) (CE_op Eq (CE_fst (CE_val (V_var x))) (CE_fst (CE_val (V_var x))))) \) + (AS_let x1 (AE_fst ((V_var x))) ( + AS_let x2 (AE_snd ((V_var x))) ( + AS_let x3 (AE_op Eq ((V_var x1)) ((V_var x2))) (AS_val (V_var x3))))))" + + apply(rule wfFTI,wf_all2) + defer 1 + + + apply simp + apply(rule wfTI) + apply(fresh_mth_basic ) + defer 1 + apply(wf_all ) + using assms apply simp + apply(wf_all ) + using assms apply simp + apply(wf_all ) + using assms apply simp + apply(wf_all ) + defer 1 + apply(wf_all ) + using assms apply simp +(*,wf_ctx,rule wfC_eqI,rule wfCE_valI,rule wfV_varI, (wf_ctx|auto|wf1)+)*) + apply(fresh_mth_basic ) + apply((wf_ctx|auto|wf1)+) + apply(rule wfCE_eqI,rule wfCE_fstI, rule wfCE_valI,rule wfV_varI) + apply((wf_ctx|auto|wf1)+) + apply(fresh_mth_basic add: assms)+ + apply((wf_ctx|auto|wf1)+) + using assms apply simp + apply(rule wfCE_fstI ,rule wfCE_valI,rule wfV_varI) + apply((wf_ctx|auto|wf1)+) + apply(fresh_mth_basic add: assms)+ + apply((wf_ctx|auto|wf1)+) + using assms apply simp + by((wf_ctx|auto|wf1)+) + + +end \ No newline at end of file diff --git a/thys/MiniSail/IVSubst.thy b/thys/MiniSail/IVSubst.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/IVSubst.thy @@ -0,0 +1,1446 @@ +(*<*) +theory IVSubst + imports Syntax +begin + (*>*) + +chapter \Immutable Variable Substitution\ + +text \Substitution involving immutable variables. We define a class and instances for all +of the term forms\ + +section \Class\ + +class has_subst_v = fs + + fixes subst_v :: "'a::fs \ x \ v \ 'a::fs" ("_[_::=_]\<^sub>v" [1000,50,50] 1000) + assumes fresh_subst_v_if: "y \ (subst_v a x v) \ (atom x \ a \ y \ a) \ (y \ v \ (y \ a \ y = atom x))" + and forget_subst_v[simp]: "atom x \ a \ subst_v a x v = a" + and subst_v_id[simp]: "subst_v a x (V_var x) = a" + and eqvt[simp,eqvt]: "(p::perm) \ (subst_v a x v) = (subst_v (p \ a) (p \x) (p \v))" + and flip_subst_v[simp]: "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + and subst_v_simple_commute[simp]: "atom x \ c \(c[z::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = c[z::=b]\<^sub>v" +begin + +lemma subst_v_flip_eq_one: + fixes z1::x and z2::x and x1::x and x2::x + assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + and "atom x1 \ (z1,z2,c1,c2)" + shows "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (c2[z2::=[x1]\<^sup>v]\<^sub>v)" +proof - + have "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (x1 \ z1) \ c1" using assms flip_subst_v by auto + moreover have "(c2[z2::=[x1]\<^sup>v]\<^sub>v) = (x1 \ z2) \ c2" using assms flip_subst_v 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_v_flip_eq_two: + fixes z1::x and z2::x and x1::x and x2::x + assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + shows "(c1[z1::=b]\<^sub>v) = (c2[z2::=b]\<^sub>v)" +proof - + obtain x::x where *:"atom x \ (z1,z2,c1,c2)" using obtain_fresh by metis + hence "(c1[z1::=[x]\<^sup>v]\<^sub>v) = (c2[z2::=[x]\<^sup>v]\<^sub>v)" using subst_v_flip_eq_one[OF assms, of x] by metis + hence "(c1[z1::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = (c2[z2::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v" by auto + thus ?thesis using subst_v_simple_commute * fresh_prod4 by metis +qed + +lemma subst_v_flip_eq_three: + assumes "[[atom z1]]lst. c1 = [[atom z1']]lst. c1'" and "atom x \ c1" and "atom x' \ (x,z1,z1', c1, c1')" + shows "(x \ x') \ (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1'[z1'::=[x']\<^sup>v]\<^sub>v" +proof - + have "atom x' \ c1[z1::=[x]\<^sup>v]\<^sub>v" using assms fresh_subst_v_if by simp + hence "(x \ x') \ (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1[z1::=[x]\<^sup>v]\<^sub>v[x::=[x']\<^sup>v]\<^sub>v" using flip_subst_v[of x' "c1[z1::=[x]\<^sup>v]\<^sub>v" x] flip_commute by metis + also have "... = c1[z1::=[x']\<^sup>v]\<^sub>v" using subst_v_simple_commute fresh_prod4 assms by auto + also have "... = c1'[z1'::=[x']\<^sup>v]\<^sub>v" using subst_v_flip_eq_one[of z1 c1 z1' c1' x'] using assms by auto + finally show ?thesis by auto +qed + +end + +section \Values\ + +nominal_function + subst_vv :: "v \ x \ v \ v" where + "subst_vv (V_lit l) x v = V_lit l" +| "subst_vv (V_var y) x v = (if x = y then v else V_var y)" +| "subst_vv (V_cons tyid c v') x v = V_cons tyid c (subst_vv v' x v)" +| "subst_vv (V_consp tyid c b v') x v = V_consp tyid c b (subst_vv v' x v)" +| "subst_vv (V_pair v1 v2) x v = V_pair (subst_vv v1 x v ) (subst_vv v2 x v )" + by(auto simp: eqvt_def subst_vv_graph_aux_def, metis v.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_vv_abbrev :: "v \ x \ v \ v" ("_[_::=_]\<^sub>v\<^sub>v" [1000,50,50] 1000) + where + "v[x::=v']\<^sub>v\<^sub>v \ subst_vv v x v'" + +lemma fresh_subst_vv_if [simp]: + "j \ t[i::=x]\<^sub>v\<^sub>v = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using supp_l_empty apply (induct t rule: v.induct,auto simp add: subst_vv.simps fresh_def, auto) + by (simp add: supp_at_base |metis b.supp supp_b_empty )+ + +lemma forget_subst_vv [simp]: "atom a \ tm \ tm[a::=x]\<^sub>v\<^sub>v = tm" + by (induct tm rule: v.induct) (simp_all add: fresh_at_base) + +lemma subst_vv_id [simp]: "tm[a::=V_var a]\<^sub>v\<^sub>v = tm" + by (induct tm rule: v.induct) simp_all + +lemma subst_vv_commute [simp]: + "atom j \ tm \ tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>v\<^sub>v " + by (induct tm rule: v.induct) (auto simp: fresh_Pair) + +lemma subst_vv_commute_full [simp]: + "atom j \ t \ atom i \ u \ i \ j \ tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[j::=u]\<^sub>v\<^sub>v[i::=t]\<^sub>v\<^sub>v" + by (induct tm rule: v.induct) auto + +lemma subst_vv_var_flip[simp]: + fixes v::v + assumes "atom y \ v" + shows "(y \ x) \ v = v [x::=V_var y]\<^sub>v\<^sub>v" + using assms apply(induct v rule:v.induct) + apply auto + using l.fresh l.perm_simps l.strong_exhaust supp_l_empty permute_pure permute_list.simps fresh_def flip_fresh_fresh apply fastforce + using permute_pure apply blast+ + done + +instantiation v :: has_subst_v +begin + +definition + "subst_v = subst_vv" + +instance proof + fix j::atom and i::x and x::v and t::v + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_vv_if[of j t i x] subst_v_v_def by metis + + fix a::x and tm::v and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_vv subst_v_v_def by simp + + fix a::x and tm::v + show "subst_v tm a (V_var a) = tm" using subst_vv_id subst_v_v_def by simp + + fix p::perm and x1::x and v::v and t1::v + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_v_v_def by simp + + fix x::x and c::v and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_v_v_def by simp + + fix x::x and c::v and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_v_v_def by simp +qed + +end + +section \Expressions\ + +nominal_function subst_ev :: "e \ x \ v \ e" where + "subst_ev ( (AE_val v') ) x v = ( (AE_val (subst_vv v' x v)) )" +| "subst_ev ( (AE_app f v') ) x v = ( (AE_app f (subst_vv v' x v )) )" +| "subst_ev ( (AE_appP f b v') ) x v = ( (AE_appP f b (subst_vv v' x v )) )" +| "subst_ev ( (AE_op opp v1 v2) ) x v = ( (AE_op opp (subst_vv v1 x v ) (subst_vv v2 x v )) )" +| "subst_ev [#1 v']\<^sup>e x v = [#1 (subst_vv v' x v )]\<^sup>e" +| "subst_ev [#2 v']\<^sup>e x v = [#2 (subst_vv v' x v )]\<^sup>e" +| "subst_ev ( (AE_mvar u)) x v = AE_mvar u" +| "subst_ev [| v' |]\<^sup>e x v = [| (subst_vv v' x v ) |]\<^sup>e" +| "subst_ev ( AE_concat v1 v2) x v = AE_concat (subst_vv v1 x v ) (subst_vv v2 x v )" +| "subst_ev ( AE_split v1 v2) x v = AE_split (subst_vv v1 x v ) (subst_vv v2 x v )" + by(simp add: eqvt_def subst_ev_graph_aux_def,auto)(meson e.strong_exhaust) + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_ev_abbrev :: "e \ x \ v \ e" ("_[_::=_]\<^sub>e\<^sub>v" [1000,50,50] 500) + where + "e[x::=v']\<^sub>e\<^sub>v \ subst_ev e x v' " + +lemma size_subst_ev [simp]: "size ( subst_ev A i x) = size A" + apply (nominal_induct A avoiding: i x rule: e.strong_induct) + by auto + +lemma forget_subst_ev [simp]: "atom a \ A \ subst_ev A a x = A" + apply (nominal_induct A avoiding: a x rule: e.strong_induct) + by (auto simp: fresh_at_base) + +lemma subst_ev_id [simp]: "subst_ev A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: e.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_ev_if [simp]: + "j \ (subst_ev A i x ) = ((atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i)))" + apply (induct A rule: e.induct) + unfolding subst_ev.simps fresh_subst_vv_if apply auto+ + using pure_fresh fresh_opp_all apply metis+ + done + +lemma subst_ev_commute [simp]: + "atom j \ A \ (A[i::=t]\<^sub>e\<^sub>v)[j::=u]\<^sub>e\<^sub>v = A[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>e\<^sub>v" + by (nominal_induct A avoiding: i j t u rule: e.strong_induct) (auto simp: fresh_at_base) + +lemma subst_ev_var_flip[simp]: + fixes e::e and y::x and x::x + assumes "atom y \ e" + shows "(y \ x) \ e = e [x::=V_var y]\<^sub>e\<^sub>v" + using assms apply(nominal_induct e rule:e.strong_induct) + apply (simp add: subst_v_v_def) + apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip) + apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip) + subgoal for x + apply (rule_tac y=x in opp.strong_exhaust) + using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+ + using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+ + +lemma subst_ev_flip: + fixes e::e and ea::e and c::x + assumes "atom c \ (e, ea)" and "atom c \ (x, xa, e, ea)" and "(x \ c) \ e = (xa \ c) \ ea" + shows "e[x::=v']\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>e\<^sub>v" +proof - + have "e[x::=v']\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>e\<^sub>v)[c::=v']\<^sub>e\<^sub>v" using subst_ev_commute assms by simp + also have "... = ((c \ x) \ e)[c::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + also have "... = ((c \ xa) \ ea)[c::=v']\<^sub>e\<^sub>v" using assms flip_commute by metis + also have "... = ea[xa::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + finally show ?thesis by auto +qed + +lemma subst_ev_var[simp]: + "(AE_val (V_var x))[x::=[z]\<^sup>v]\<^sub>e\<^sub>v = AE_val (V_var z)" + by auto + +instantiation e :: has_subst_v +begin + +definition + "subst_v = subst_ev" + +instance proof + fix j::atom and i::x and x::v and t::e + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_ev_if[of j t i x] subst_v_e_def by metis + + fix a::x and tm::e and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_ev subst_v_e_def by simp + + fix a::x and tm::e + show "subst_v tm a (V_var a) = tm" using subst_ev_id subst_v_e_def by simp + + fix p::perm and x1::x and v::v and t1::e + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_ev_commute subst_v_e_def by simp + + fix x::x and c::e and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_v_e_def by simp + + fix x::x and c::e and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_v_e_def by simp +qed +end + +lemma subst_ev_commute_full: + fixes e::e and w::v and v::v + assumes "atom z \ v" and "atom x \ w" and "x \ z" + shows "subst_ev (e[z::=w]\<^sub>e\<^sub>v) x v = subst_ev (e[x::=v]\<^sub>e\<^sub>v) z w" + using assms by(nominal_induct e rule: e.strong_induct,simp+) + +lemma subst_ev_v_flip1[simp]: + fixes e::e + assumes "atom z1 \ (z,e)" and "atom z1' \ (z,e)" + shows"(z1 \ z1') \ e[z::=v]\<^sub>e\<^sub>v = e[z::= ((z1 \ z1') \ v)]\<^sub>e\<^sub>v" + using assms proof(nominal_induct e rule:e.strong_induct) +qed (simp add: flip_def fresh_Pair swap_fresh_fresh)+ + +section \Expressions in Constraints\ + +nominal_function subst_cev :: "ce \ x \ v \ ce" where + "subst_cev ( (CE_val v') ) x v = ( (CE_val (subst_vv v' x v )) )" +| "subst_cev ( (CE_op opp v1 v2) ) x v = ( (CE_op opp (subst_cev v1 x v ) (subst_cev v2 x v )) )" +| "subst_cev ( (CE_fst v')) x v = CE_fst (subst_cev v' x v )" +| "subst_cev ( (CE_snd v')) x v = CE_snd (subst_cev v' x v )" +| "subst_cev ( (CE_len v')) x v = CE_len (subst_cev v' x v )" +| "subst_cev ( CE_concat v1 v2) x v = CE_concat (subst_cev v1 x v ) (subst_cev v2 x v )" + apply (simp add: eqvt_def subst_cev_graph_aux_def,auto) + by (meson ce.strong_exhaust) + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_cev_abbrev :: "ce \ x \ v \ ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>v" [1000,50,50] 500) + where + "e[x::=v']\<^sub>c\<^sub>e\<^sub>v \ subst_cev e x v'" + +lemma size_subst_cev [simp]: "size ( subst_cev A i x ) = size A" + by (nominal_induct A avoiding: i x rule: ce.strong_induct,auto) + +lemma forget_subst_cev [simp]: "atom a \ A \ subst_cev A a x = A" + by (nominal_induct A avoiding: a x rule: ce.strong_induct, auto simp: fresh_at_base) + +lemma subst_cev_id [simp]: "subst_cev A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: ce.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_cev_if [simp]: + "j \ (subst_cev A i x ) = ((atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i)))" +proof(nominal_induct A avoiding: i x rule: ce.strong_induct) + case (CE_op opp v1 v2) + then show ?case using fresh_subst_vv_if subst_ev.simps e.supp pure_fresh opp.fresh + fresh_e_opp + using fresh_opp_all by auto +qed(auto)+ + +lemma subst_cev_commute [simp]: + "atom j \ A \ (subst_cev (subst_cev A i t ) j u) = subst_cev A i (subst_vv t j u )" + by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base) + +lemma subst_cev_var_flip[simp]: + fixes e::ce and y::x and x::x + assumes "atom y \ e" + shows "(y \ x) \ e = e [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v" + using assms proof(nominal_induct e rule:ce.strong_induct) + case (CE_val v) + then show ?case using subst_vv_var_flip by auto +next + case (CE_op opp v1 v2) + hence yf: "atom y \ v1 \ atom y \ v2" using ce.fresh by blast + have " (y \ x) \ (CE_op opp v1 v2 ) = CE_op ((y \ x) \ opp) ( (y \ x) \ v1 ) ( (y \ x) \ v2)" + using opp.perm_simps ce.perm_simps permute_pure ce.fresh opp.strong_exhaust by presburger + also have "... = CE_op ((y \ x) \ opp) (v1[x::=V_var y]\<^sub>c\<^sub>e\<^sub>v) (v2 [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v)" using yf + by (simp add: CE_op.hyps(1) CE_op.hyps(2)) + finally show ?case using subst_cev.simps opp.perm_simps opp.strong_exhaust + by (metis (full_types)) +qed( (auto simp add: permute_pure subst_vv_var_flip)+) + +lemma subst_cev_flip: + fixes e::ce and ea::ce and c::x + assumes "atom c \ (e, ea)" and "atom c \ (x, xa, e, ea)" and "(x \ c) \ e = (xa \ c) \ ea" + shows "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v" +proof - + have "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>c\<^sub>e\<^sub>v)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_commute assms by simp + also have "... = ((c \ x) \ e)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + also have "... = ((c \ xa) \ ea)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using assms flip_commute by metis + also have "... = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + finally show ?thesis by auto +qed + +lemma subst_cev_var[simp]: + fixes z::x and x::x + shows "[[x]\<^sup>v]\<^sup>c\<^sup>e [x::=[z]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v = [[z]\<^sup>v]\<^sup>c\<^sup>e" + by auto + +instantiation ce :: has_subst_v +begin + +definition + "subst_v = subst_cev" + +instance proof + fix j::atom and i::x and x::v and t::ce + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_cev_if[of j t i x] subst_v_ce_def by metis + + fix a::x and tm::ce and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_cev subst_v_ce_def by simp + + fix a::x and tm::ce + show "subst_v tm a (V_var a) = tm" using subst_cev_id subst_v_ce_def by simp + + fix p::perm and x1::x and v::v and t1::ce + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_cev_commute subst_v_ce_def by simp + + fix x::x and c::ce and z::x + show "atom x \ c \ ((x \ z) \ c) = c [z::=V_var x]\<^sub>v" + using subst_v_ce_def by simp + + fix x::x and c::ce and z::x + show "atom x \ c \ c [z::=V_var x]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_v_ce_def by simp +qed + +end + +lemma subst_cev_commute_full: + fixes e::ce and w::v and v::v + assumes "atom z \ v" and "atom x \ w" and "x \ z" + shows "subst_cev (e[z::=w]\<^sub>c\<^sub>e\<^sub>v) x v = subst_cev (e[x::=v]\<^sub>c\<^sub>e\<^sub>v) z w " + using assms by(nominal_induct e rule: ce.strong_induct,simp+) + + +lemma subst_cev_v_flip1[simp]: + fixes e::ce + assumes "atom z1 \ (z,e)" and "atom z1' \ (z,e)" + shows"(z1 \ z1') \ e[z::=v]\<^sub>c\<^sub>e\<^sub>v = e[z::= ((z1 \ z1') \ v)]\<^sub>c\<^sub>e\<^sub>v" + using assms apply(nominal_induct e rule:ce.strong_induct) + by (simp add: flip_def fresh_Pair swap_fresh_fresh)+ + +section \Constraints\ + +nominal_function subst_cv :: "c \ x \ v \ c" where + "subst_cv (C_true) x v = C_true" +| "subst_cv (C_false) x v = C_false" +| "subst_cv (C_conj c1 c2) x v = C_conj (subst_cv c1 x v ) (subst_cv c2 x v )" +| "subst_cv (C_disj c1 c2) x v = C_disj (subst_cv c1 x v ) (subst_cv c2 x v )" +| "subst_cv (C_imp c1 c2) x v = C_imp (subst_cv c1 x v ) (subst_cv c2 x v )" +| "subst_cv (e1 == e2) x v = ((subst_cev e1 x v ) == (subst_cev e2 x v ))" +| "subst_cv (C_not c) x v = C_not (subst_cv c x v )" + apply (simp add: eqvt_def subst_cv_graph_aux_def,auto) + using c.strong_exhaust by metis +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_cv_abbrev :: "c \ x \ v \ c" ("_[_::=_]\<^sub>c\<^sub>v" [1000,50,50] 1000) + where + "c[x::=v']\<^sub>c\<^sub>v \ subst_cv c x v'" + +lemma size_subst_cv [simp]: "size ( subst_cv A i x ) = size A" + by (nominal_induct A avoiding: i x rule: c.strong_induct,auto) + +lemma forget_subst_cv [simp]: "atom a \ A \ subst_cv A a x = A" + by (nominal_induct A avoiding: a x rule: c.strong_induct, auto simp: fresh_at_base) + +lemma subst_cv_id [simp]: "subst_cv A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: c.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_cv_if [simp]: + "j \ (subst_cv A i x ) \ (atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i))" + by (nominal_induct A avoiding: i x rule: c.strong_induct, (auto simp add: pure_fresh)+) + +lemma subst_cv_commute [simp]: + "atom j \ A \ (subst_cv (subst_cv A i t ) j u ) = subst_cv A i (subst_vv t j u )" + by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base) + +lemma let_s_size [simp]: "size s \ size (AS_let x e s)" + apply (nominal_induct s avoiding: e x rule: s_branch_s_branch_list.strong_induct(1)) + apply auto + done + +lemma subst_cv_var_flip[simp]: + fixes c::c + assumes "atom y \ c" + shows "(y \ x) \ c = c[x::=V_var y]\<^sub>c\<^sub>v" + using assms by(nominal_induct c rule:c.strong_induct,(simp add: flip_subst_v subst_v_ce_def)+) + +instantiation c :: has_subst_v +begin + +definition + "subst_v = subst_cv" + +instance proof + fix j::atom and i::x and x::v and t::c + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_cv_if[of j t i x] subst_v_c_def by metis + + fix a::x and tm::c and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_cv subst_v_c_def by simp + + fix a::x and tm::c + show "subst_v tm a (V_var a) = tm" using subst_cv_id subst_v_c_def by simp + + fix p::perm and x1::x and v::v and t1::c + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_cv_commute subst_v_c_def by simp + + fix x::x and c::c and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_cv_var_flip subst_v_c_def by simp + + fix x::x and c::c and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_cv_var_flip subst_v_c_def by simp +qed + +end + +lemma subst_cv_var_flip1[simp]: + fixes c::c + assumes "atom y \ c" + shows "(x \ y) \ c = c[x::=V_var y]\<^sub>c\<^sub>v" + using subst_cv_var_flip flip_commute + by (metis assms) + +lemma subst_cv_v_flip3[simp]: + fixes c::c + assumes "atom z1 \ c" and "atom z1' \ c" + shows"(z1 \ z1') \ c[z::=[z1]\<^sup>v]\<^sub>c\<^sub>v = c[z::=[z1']\<^sup>v]\<^sub>c\<^sub>v" +proof - + consider "z1' = z" | "z1 = z" | "atom z1 \ z \ atom z1' \ z" by force + then show ?thesis proof(cases) + case 1 + then show ?thesis using 1 assms by auto + next + case 2 + then show ?thesis using 2 assms by auto + next + case 3 + then show ?thesis using assms by auto + qed +qed + +lemma subst_cv_v_flip[simp]: + fixes c::c + assumes "atom x \ c" + shows "((x \ z) \ c)[x::=v]\<^sub>c\<^sub>v = c [z::=v]\<^sub>c\<^sub>v" + using assms subst_v_c_def by auto + +lemma subst_cv_commute_full: + fixes c::c + assumes "atom z \ v" and "atom x \ w" and "x\z" + shows "(c[z::=w]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v = (c[x::=v]\<^sub>c\<^sub>v)[z::=w]\<^sub>c\<^sub>v" + using assms proof(nominal_induct c rule: c.strong_induct) + case (C_eq e1 e2) + then show ?case using subst_cev_commute_full by simp +qed(force+) + +lemma subst_cv_eq[simp]: + assumes "atom z1 \ e1" + shows "(CE_val (V_var z1) == e1 )[z1::=[x]\<^sup>v]\<^sub>c\<^sub>v = (CE_val (V_var x) == e1 )" (is "?A = ?B") +proof - + have "?A = (((CE_val (V_var z1))[z1::=[x]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v) == e1)" using subst_cv.simps assms by simp + thus ?thesis by simp +qed + +section \Variable Context\ + +text \The idea of this substitution is to remove x from the context. We really want to add the condition +that x is fresh in v but this causes problems with proofs.\ + +nominal_function subst_gv :: "\ \ x \ v \ \" where + "subst_gv GNil x v = GNil" +| "subst_gv ((y,b,c) #\<^sub>\ \) x v = (if x = y then \ else ((y,b,c[x::=v]\<^sub>c\<^sub>v)#\<^sub>\ (subst_gv \ x v )))" +proof(goal_cases) + case 1 + then show ?case by(simp add: eqvt_def subst_gv_graph_aux_def ) +next + case (3 P x) + then show ?case by (metis neq_GNil_conv prod_cases3) +qed(fast+) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_gv_abbrev :: "\ \ x \ v \ \" ("_[_::=_]\<^sub>\\<^sub>v" [1000,50,50] 1000) + where + "g[x::=v]\<^sub>\\<^sub>v \ subst_gv g x v" + +lemma size_subst_gv [simp]: "size ( subst_gv G i x ) \ size G" + by (induct G,auto) + +lemma forget_subst_gv [simp]: "atom a \ G \ subst_gv G a x = G" + apply (induct G ,auto) + using fresh_GCons fresh_PairD(1) not_self_fresh apply blast + apply (simp add: fresh_GCons)+ + done + +lemma fresh_subst_gv: "atom a \ G \ atom a \ v \ atom a \ subst_gv G x v" +proof(induct G) + case GNil + then show ?case by auto +next + case (GCons xbc G) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + show ?case proof(cases "x=x'") + case True + have "atom a \ G" using GCons fresh_GCons by blast + thus ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc True by presburger + next + case False + then show ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc False fresh_GCons by simp + qed +qed + +lemma subst_gv_flip: + fixes x::x and xa::x and z::x and c::c and b::b and \::\ + assumes "atom xa \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" and "atom xa \ \" and "atom x \ \" and "atom x \ (z, c)" and "atom xa \ (z, c)" + shows "(x \ xa) \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) = (xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \" +proof - + have "(x \ xa) \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) = (( (x \ xa) \ x, b, (x \ xa) \ c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ ((x \ xa) \ \))" + using subst Cons_eqvt flip_fresh_fresh using G_cons_flip by simp + also have "... = ((xa, b, (x \ xa) \ c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ ((x \ xa) \ \))" using assms by fastforce + also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ ((x \ xa) \ \))" using assms subst_cv_var_flip by fastforce + also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \)" using assms flip_fresh_fresh by blast + finally show ?thesis by simp +qed + +section \Types\ + +nominal_function subst_tv :: "\ \ x \ v \ \" where + "atom z \ (x,v) \ subst_tv \ z : b | c \ x v = \ z : b | c[x::=v]\<^sub>c\<^sub>v \" + apply (simp add: eqvt_def subst_tv_graph_aux_def ) + apply auto + subgoal for P a aa b + apply(rule_tac y=a and c="(aa,b)" in \.strong_exhaust) + by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) + apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) +proof - + fix z :: x and c :: c and za :: x and xa :: x and va :: v and ca :: c and cb :: x + assume a1: "atom za \ va" and a2: "atom z \ va" and a3: "\cb. atom cb \ c \ atom cb \ ca \ cb \ z \ cb \ za \ c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v" + assume a4: "atom cb \ c" and a5: "atom cb \ ca" and a6: "cb \ z" and a7: "cb \ za" and "atom cb \ va" and a8: "za \ xa" and a9: "z \ xa" + assume a10:"cb \ xa" + note assms = a10 a9 a8 a7 a6 a5 a4 a3 a2 a1 + + have "c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v" using assms by auto + hence "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by simp + moreover have "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v" using subst_cv_commute_full[of z va xa "V_var cb" ] assms fresh_def v.supp by fastforce + moreover have "ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v" + using subst_cv_commute_full[of za va xa "V_var cb" ] assms fresh_def v.supp by fastforce + + ultimately show "c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v" by simp +qed + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_tv_abbrev :: "\ \ x \ v \ \" ("_[_::=_]\<^sub>\\<^sub>v" [1000,50,50] 1000) + where + "t[x::=v]\<^sub>\\<^sub>v \ subst_tv t x v" + +lemma size_subst_tv [simp]: "size ( subst_tv A i x ) = size A" +proof (nominal_induct A avoiding: i x rule: \.strong_induct) + case (T_refined_type x' b' c') + then show ?case by auto +qed + +lemma forget_subst_tv [simp]: "atom a \ A \ subst_tv A a x = A" + apply (nominal_induct A avoiding: a x rule: \.strong_induct) + apply(auto simp: fresh_at_base) + done + +lemma subst_tv_id [simp]: "subst_tv A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: \.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_tv_if [simp]: + "j \ (subst_tv A i x ) \ (atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i))" + apply (nominal_induct A avoiding: i x rule: \.strong_induct) + using fresh_def supp_b_empty x_fresh_b by auto + +lemma subst_tv_commute [simp]: + "atom y \ \ \ (\[x::= t]\<^sub>\\<^sub>v)[y::=v]\<^sub>\\<^sub>v = \[x::= t[y::=v]\<^sub>v\<^sub>v]\<^sub>\\<^sub>v " + by (nominal_induct \ avoiding: x y t v rule: \.strong_induct) (auto simp: fresh_at_base) + +lemma subst_tv_var_flip [simp]: + fixes x::x and xa::x and \::\ + assumes "atom xa \ \" + shows "(x \ xa) \ \ = \[x::=V_var xa]\<^sub>\\<^sub>v" +proof - + obtain z::x and b and c where zbc: "atom z \ (x,xa, V_var xa) \ \ = \ z : b | c \" + using obtain_fresh_z by (metis prod.inject subst_tv.cases) + hence "atom xa \ supp c - { atom z }" using \.supp[of z b c] fresh_def supp_b_empty assms + by auto + moreover have "xa \ z" using zbc fresh_prod3 by force + ultimately have xaf: "atom xa \ c" using fresh_def by auto + have "(x \ xa) \ \ = \ z : b | (x \ xa) \ c \" + by (metis \.perm_simps empty_iff flip_at_base_simps(3) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2) fresh_def not_self_fresh supp_b_empty v.fresh(2) zbc) + also have "... = \ z : b | c[x::=V_var xa]\<^sub>c\<^sub>v \" using subst_cv_v_flip xaf + by (metis permute_flip_cancel permute_flip_cancel2 subst_cv_var_flip) + finally show ?thesis using subst_tv.simps zbc + using fresh_PairD(1) not_self_fresh by force +qed + +instantiation \ :: has_subst_v +begin + +definition + "subst_v = subst_tv" + +instance proof + fix j::atom and i::x and x::v and t::\ + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + + proof(nominal_induct t avoiding: i x rule:\.strong_induct) + case (T_refined_type z b c) + hence " j \ \ z : b | c \[i::=x]\<^sub>v = j \ \ z : b | c[i::=x]\<^sub>c\<^sub>v \" using subst_tv.simps subst_v_\_def fresh_Pair by simp + also have "... = (atom i \ \ z : b | c \ \ j \ \ z : b | c \ \ j \ x \ (j \ \ z : b | c \ \ j = atom i))" + unfolding \.fresh using subst_v_c_def fresh_subst_v_if + using T_refined_type.hyps(1) T_refined_type.hyps(2) x_fresh_b by auto + finally show ?case by auto + qed + + fix a::x and tm::\ and x::v + show "atom a \ tm \ subst_v tm a x = tm" + apply(nominal_induct tm avoiding: a x rule:\.strong_induct) + using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\_def fresh_Pair by simp + + fix a::x and tm::\ + show "subst_v tm a (V_var a) = tm" + apply(nominal_induct tm avoiding: a rule:\.strong_induct) + using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\_def fresh_Pair by simp + + fix p::perm and x1::x and v::v and t1::\ + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + apply(nominal_induct tm avoiding: a x rule:\.strong_induct) + using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\_def fresh_Pair by simp + + fix x::x and c::\ and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + apply(nominal_induct c avoiding: z x rule:\.strong_induct) + using subst_v_c_def flip_subst_v subst_tv.simps subst_v_\_def fresh_Pair by auto + + fix x::x and c::\ and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + apply(nominal_induct c avoiding: x v z rule:\.strong_induct) + using subst_v_c_def subst_tv.simps subst_v_\_def fresh_Pair + by (metis flip_commute subst_tv_commute subst_tv_var_flip subst_v_\_def subst_vv.simps(2)) +qed + +end + +lemma subst_tv_commute_full: + fixes c::\ + assumes "atom z \ v" and "atom x \ w" and "x\z" + shows "(c[z::=w]\<^sub>\\<^sub>v)[x::=v]\<^sub>\\<^sub>v = (c[x::=v]\<^sub>\\<^sub>v)[z::=w]\<^sub>\\<^sub>v" + using assms proof(nominal_induct c avoiding: x v z w rule: \.strong_induct) + case (T_refined_type x1a x2a x3a) + then show ?case using subst_cv_commute_full by simp +qed + +lemma type_eq_subst_eq: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = \ z2 : b2 | c2 \" + shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" + using subst_v_flip_eq_two[of z1 c1 z2 c2 v] \.eq_iff assms subst_v_c_def by simp + +text \Extract constraint from a type. We cannot just project out the constraint as this would +mean alpha-equivalent types give different answers \ + +nominal_function c_of :: "\ \ x \ c" where + "atom z \ x \ c_of (T_refined_type z b c) x = c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" +proof(goal_cases) + case 1 + then show ?case using eqvt_def c_of_graph_aux_def by force +next + case (2 x y) + then show ?case using eqvt_def c_of_graph_aux_def by force +next + case (3 P x) + then obtain x1::\ and x2::x where *:"x = (x1,x2)" by force + obtain z' and b' and c' where "x1 = \ z' : b' | c' \ \ atom z' \ x2" using obtain_fresh_z by metis + then show ?case using 3 * by auto +next + case (4 z1 x1 b1 c1 z2 x2 b2 c2) + then show ?case using subst_v_flip_eq_two \.eq_iff by (metis prod.inject type_eq_subst_eq) +qed + +nominal_termination (eqvt) by lexicographic_order + +lemma c_of_eq: + shows "c_of \ x : b | c \ x = c" +proof(nominal_induct "\ x : b | c \" avoiding: x rule: \.strong_induct) + case (T_refined_type x' c') + moreover hence "c_of \ x' : b | c' \ x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by auto + moreover have "\ x' : b | c' \ = \ x : b | c \" using T_refined_type \.eq_iff by metis + moreover have "c'[x'::=V_var x]\<^sub>c\<^sub>v = c" using T_refined_type Abs1_eq_iff flip_subst_v subst_v_c_def + by (metis subst_cv_id) + ultimately show ?case by auto +qed + +lemma obtain_fresh_z_c_of: + fixes t::"'b::fs" + obtains z where "atom z \ t \ \ = \ z : b_of \ | c_of \ z \" +proof - + obtain z and c where "atom z \ t \ \ = \ z : b_of \ | c \" using obtain_fresh_z2 by metis + moreover hence "c = c_of \ z" using c_of.simps using c_of_eq by metis + ultimately show ?thesis + using that by auto +qed + +lemma c_of_fresh: + fixes x::x + assumes "atom x \ (t,z)" + shows "atom x \ c_of t z" +proof - + obtain z' and c' where z:"t = \ z' : b_of t | c' \ \ atom z' \ (x,z)" using obtain_fresh_z_c_of by metis + hence *:"c_of t z = c'[z'::=V_var z]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair by metis + have "(atom x \ c' \ atom x \ set [atom z']) \ atom x \ b_of t" using \.fresh assms z fresh_Pair by metis + hence "atom x \ c'" using fresh_Pair z fresh_at_base(2) by fastforce + moreover have "atom x \ V_var z" using assms fresh_Pair v.fresh by metis + ultimately show ?thesis using assms fresh_subst_v_if[of "atom x" c' z' "V_var z"] subst_v_c_def * by metis +qed + +lemma c_of_switch: + fixes z::x + assumes "atom z \ t" + shows "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c_of t x" +proof - + obtain z' and c' where z:"t = \ z' : b_of t | c' \ \ atom z' \ (x,z)" using obtain_fresh_z_c_of by metis + hence "(atom z \ c' \ atom z \ set [atom z']) \ atom z \ b_of t" using \.fresh[of "atom z" z' "b_of t" c'] assms by metis + moreover have " atom z \ set [atom z']" using z fresh_Pair by force + ultimately have **:"atom z \ c'" using fresh_Pair z fresh_at_base(2) by metis + + have "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c'[z'::=V_var z]\<^sub>c\<^sub>v[z::=V_var x]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair z by metis + also have "... = c'[z'::=V_var x]\<^sub>c\<^sub>v" using subst_v_simple_commute subst_v_c_def assms c_of.simps z ** by metis + finally show ?thesis using c_of.simps[of z' x "b_of t" c'] fresh_Pair z by metis +qed + +lemma type_eq_subst_eq1: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = (\ z2 : b2 | c2 \)" and "atom z1 \ c2" + shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and " c1 = (z1 \ z2) \ c2" +proof - + show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast + show "b1=b2" using \.eq_iff assms by blast + have "z1 = z2 \ c1 = c2 \ z1 \ z2 \ c1 = (z1 \ z2) \ c2 \ atom z1 \ c2" + using \.eq_iff Abs1_eq_iff[of z1 c1 z2 c2] assms by blast + thus "c1 = (z1 \ z2) \ c2" by auto +qed + +lemma type_eq_subst_eq2: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = (\ z2 : b2 | c2 \)" + shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" +proof - + show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast + show "b1=b2" using \.eq_iff assms by blast + show "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + using \.eq_iff assms by auto +qed + +lemma type_eq_subst_eq3: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = (\ z2 : b2 | c2 \)" and "atom z1 \ c2" + shows "c1 = c2[z2::=V_var z1]\<^sub>c\<^sub>v" and "b1=b2" + using type_eq_subst_eq1 assms subst_v_c_def + by (metis subst_cv_var_flip)+ + +lemma type_eq_flip: + assumes "atom x \ c" + shows "\ z : b | c \ = \ x : b | (x \ z ) \ c \" + using \.eq_iff Abs1_eq_iff assms + by (metis (no_types, lifting) flip_fresh_fresh) + +lemma c_of_true: + "c_of \ z' : B_bool | TRUE \ x = C_true" +proof(nominal_induct "\ z' : B_bool | TRUE \" avoiding: x rule:\.strong_induct) + case (T_refined_type x1a x3a) + hence "\ z' : B_bool | TRUE \ = \ x1a : B_bool | x3a \" using \.eq_iff by metis + then show ?case using subst_cv.simps c_of.simps T_refined_type + type_eq_subst_eq3 + by (metis type_eq_subst_eq) +qed + +lemma type_eq_subst: + assumes "atom x \ c" + shows "\ z : b | c \ = \ x : b | c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v \" + using \.eq_iff Abs1_eq_iff assms + using subst_cv_var_flip type_eq_flip by auto + +lemma type_e_subst_fresh: + fixes x::x and z::x + assumes "atom z \ (x,v)" and "atom x \ e" + shows "\ z : b | CE_val (V_var z) == e \[x::=v]\<^sub>\\<^sub>v = \ z : b | CE_val (V_var z) == e \" + using assms subst_tv.simps subst_cv.simps forget_subst_cev by simp + +lemma type_v_subst_fresh: + fixes x::x and z::x + assumes "atom z \ (x,v)" and "atom x \ v'" + shows "\ z : b | CE_val (V_var z) == CE_val v' \[x::=v]\<^sub>\\<^sub>v = \ z : b | CE_val (V_var z) == CE_val v' \" + using assms subst_tv.simps subst_cv.simps by simp + +lemma subst_tbase_eq: + "b_of \ = b_of \[x::=v]\<^sub>\\<^sub>v" +proof - + obtain z and b and c where zbc: "\ = \ z:b|c\ \ atom z \ (x,v)" using \.exhaust + by (metis prod.inject subst_tv.cases) + hence "b_of \ z:b|c\ = b_of \ z:b|c\[x::=v]\<^sub>\\<^sub>v" using subst_tv.simps by simp + thus ?thesis using zbc by blast +qed + +lemma subst_tv_if: + assumes "atom z1 \ (x,v)" and "atom z' \ (x,v)" + shows "\ z1 : b | CE_val (v'[x::=v]\<^sub>v\<^sub>v) == CE_val (V_lit l) IMP (c'[x::=v]\<^sub>c\<^sub>v)[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \ = + \ z1 : b | CE_val v' == CE_val (V_lit l) IMP c'[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \[x::=v]\<^sub>\\<^sub>v" + using subst_cv_commute_full[of z' v x "V_var z1" c'] subst_tv.simps subst_vv.simps(1) subst_ev.simps subst_cv.simps assms + by simp + +lemma subst_tv_tid: + assumes "atom za \ (x,v)" + shows "\ za : B_id tid | TRUE \ = \ za : B_id tid | TRUE \[x::=v]\<^sub>\\<^sub>v" + using assms subst_tv.simps subst_cv.simps by presburger + + +lemma b_of_subst: + "b_of (\[x::=v]\<^sub>\\<^sub>v) = b_of \" +proof - + obtain z b c where *:"\ = \ z : b | c \ \ atom z \ (x,v)" using obtain_fresh_z by metis + thus ?thesis using subst_tv.simps * by auto +qed + +lemma subst_tv_flip: + assumes "\'[x::=v]\<^sub>\\<^sub>v = \" and "atom x \ (v,\)" and "atom x' \ (v,\)" + shows "((x' \ x) \ \')[x'::=v]\<^sub>\\<^sub>v = \" +proof - + have "(x' \ x) \ v = v \ (x' \ x) \ \ = \" using assms flip_fresh_fresh by auto + thus ?thesis using subst_tv.eqvt[of "(x' \ x)" \' x v ] assms by auto +qed + +lemma subst_cv_true: + "\ z : B_id tid | TRUE \ = \ z : B_id tid | TRUE \[x::=v]\<^sub>\\<^sub>v" +proof - + obtain za::x where "atom za \ (x,v)" using obtain_fresh by auto + hence "\ z : B_id tid | TRUE \ = \ za: B_id tid | TRUE \" using \.eq_iff Abs1_eq_iff by fastforce + moreover have "\ za : B_id tid | TRUE \ = \ za : B_id tid | TRUE \[x::=v]\<^sub>\\<^sub>v" + using subst_cv.simps subst_tv.simps by (simp add: \atom za \ (x, v)\) + ultimately show ?thesis by argo +qed + +lemma t_eq_supp: + assumes "(\ z : b | c \) = (\ z1 : b1 | c1 \)" + shows "supp c - { atom z } = supp c1 - { atom z1 }" +proof - + have "supp c - { atom z } \ supp b = supp c1 - { atom z1 } \ supp b1" using \.supp assms + by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty) + moreover have "supp b = supp b1" using assms \.eq_iff by simp + moreover have "atom z1 \ supp b1 \ atom z \ supp b" using supp_b_empty by simp + ultimately show ?thesis + by (metis \.eq_iff \.supp assms b.supp(1) list.set(1) list.set(2) sup_bot.right_neutral) +qed + +lemma fresh_t_eq: + fixes x::x + assumes "(\ z : b | c \) = (\ zz : b | cc \)" and "atom x \ c" and "x \ zz" + shows "atom x \ cc" +proof - + have "supp c - { atom z } \ supp b = supp cc - { atom zz } \ supp b" using \.supp assms + by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty) + moreover have "atom x \ supp c" using assms fresh_def by blast + ultimately have "atom x \ supp cc - { atom zz } \ supp b" by force + hence "atom x \ supp cc" using assms by simp + thus ?thesis using fresh_def by auto +qed + +section \Mutable Variable Context\ + +nominal_function subst_dv :: "\ \ x \ v \ \" where + "subst_dv DNil x v = DNil" +| "subst_dv ((u,t) #\<^sub>\ \) x v = ((u,t[x::=v]\<^sub>\\<^sub>v) #\<^sub>\ (subst_dv \ x v ))" + apply (simp add: eqvt_def subst_dv_graph_aux_def,auto ) + using delete_aux.elims by (metis \.exhaust surj_pair) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_dv_abbrev :: "\ \ x \ v \ \" ("_[_::=_]\<^sub>\\<^sub>v" [1000,50,50] 1000) + where + "\[x::=v]\<^sub>\\<^sub>v \ subst_dv \ x v " + +nominal_function dmap :: "(u*\ \ u*\) \ \ \ \" where + "dmap f DNil = DNil" +| "dmap f ((u,t)#\<^sub>\\) = (f (u,t) #\<^sub>\ (dmap f \ ))" + apply (simp add: eqvt_def dmap_graph_aux_def,auto ) + using delete_aux.elims by (metis \.exhaust surj_pair) +nominal_termination (eqvt) by lexicographic_order + +lemma subst_dv_iff: + "\[x::=v]\<^sub>\\<^sub>v = dmap (\(u,t). (u, t[x::=v]\<^sub>\\<^sub>v)) \" + by(induct \, auto) + +lemma size_subst_dv [simp]: "size ( subst_dv G i x) \ size G" + by (induct G,auto) + +lemma forget_subst_dv [simp]: "atom a \ G \ subst_dv G a x = G" + apply (induct G ,auto) + using fresh_DCons fresh_PairD(1) not_self_fresh apply fastforce + apply (simp add: fresh_DCons)+ + done + +lemma subst_dv_member: + assumes "(u,\) \ setD \" + shows "(u, \[x::=v]\<^sub>\\<^sub>v) \ setD (\[x::=v]\<^sub>\\<^sub>v)" + using assms by(induct \ rule: \_induct,auto) + +lemma fresh_subst_dv: + fixes x::x + assumes "atom xa \ \" and "atom xa \ v" + shows "atom xa \\[x::=v]\<^sub>\\<^sub>v" + using assms proof(induct \ rule:\_induct) + case DNil + then show ?case by auto +next + case (DCons u t \) + then show ?case using subst_dv.simps subst_v_\_def fresh_DCons fresh_Pair by simp +qed + +lemma fresh_subst_dv_if: + fixes j::atom and i::x and x::v and t::\ + assumes "j \ t \ j \ x" + shows "(j \ subst_dv t i x)" + using assms proof(induct t rule: \_induct) + case DNil + then show ?case using subst_gv.simps fresh_GNil by auto +next + case (DCons u' t' D') + then show ?case unfolding subst_dv.simps using fresh_DCons fresh_subst_tv_if fresh_Pair by metis +qed + +section \Statements\ + +text \ Using ideas from proofs at top of AFP/Launchbury/Substitution.thy. + Subproofs borrowed from there; hence the apply style proofs. \ + +nominal_function (default "case_sum (\x. Inl undefined) (case_sum (\x. Inl undefined) (\x. Inr undefined))") + subst_sv :: "s \ x \ v \ s" + and subst_branchv :: "branch_s \ x \ v \ branch_s" + and subst_branchlv :: "branch_list \ x \ v \ branch_list" where + "subst_sv ( (AS_val v') ) x v = (AS_val (subst_vv v' x v ))" +| "atom y \ (x,v) \ subst_sv (AS_let y e s) x v = (AS_let y (e[x::=v]\<^sub>e\<^sub>v) (subst_sv s x v ))" +| "atom y \ (x,v) \ subst_sv (AS_let2 y t s1 s2) x v = (AS_let2 y (t[x::=v]\<^sub>\\<^sub>v) (subst_sv s1 x v ) (subst_sv s2 x v ))" +| " subst_sv (AS_match v' cs) x v = AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v )" +| "subst_sv (AS_assign y v') x v = AS_assign y (subst_vv v' x v )" +| "subst_sv ( (AS_if v' s1 s2) ) x v = (AS_if (subst_vv v' x v ) (subst_sv s1 x v ) (subst_sv s2 x v ) )" +| "atom u \ (x,v) \ subst_sv (AS_var u \ v' s) x v = AS_var u (subst_tv \ x v ) (subst_vv v' x v ) (subst_sv s x v ) " +| "subst_sv (AS_while s1 s2) x v = AS_while (subst_sv s1 x v ) (subst_sv s2 x v )" +| "subst_sv (AS_seq s1 s2) x v = AS_seq (subst_sv s1 x v ) (subst_sv s2 x v )" +| "subst_sv (AS_assert c s) x v = AS_assert (subst_cv c x v) (subst_sv s x v)" +| "atom x1 \ (x,v) \ subst_branchv (AS_branch dc x1 s1 ) x v = AS_branch dc x1 (subst_sv s1 x v )" + +| "subst_branchlv (AS_final cs) x v = AS_final (subst_branchv cs x v )" +| "subst_branchlv (AS_cons cs css) x v = AS_cons (subst_branchv cs x v ) (subst_branchlv css x v )" + apply (auto,simp add: eqvt_def subst_sv_subst_branchv_subst_branchlv_graph_aux_def ) +proof(goal_cases) + + have eqvt_at_proj: "\ s xa va . eqvt_at subst_sv_subst_branchv_subst_branchlv_sumC (Inl (s, xa, va)) \ + eqvt_at (\a. projl (subst_sv_subst_branchv_subst_branchlv_sumC (Inl a))) (s, xa, va)" + apply(simp add: eqvt_at_def) + apply(rule) + apply(subst Projl_permute) + apply(thin_tac _)+ + apply (simp add: subst_sv_subst_branchv_subst_branchlv_sumC_def) + apply (simp add: THE_default_def) + apply (case_tac "Ex1 (subst_sv_subst_branchv_subst_branchlv_graph (Inl (s,xa,va)))") + apply simp + apply(auto)[1] + apply (erule_tac x="x" in allE) + apply simp + apply(cases rule: subst_sv_subst_branchv_subst_branchlv_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 P x') + then show ?case proof(cases x') + case (Inl a) thus P + proof(cases a) + case (fields aa bb cc) + thus P using Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis + qed + next + case (Inr b) thus P + proof(cases b) + case (Inl a) thus P proof(cases a) + case (fields aa bb cc) + then show ?thesis using Inr Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis + qed + next + case Inr2: (Inr b) thus P proof(cases b) + case (fields aa bb cc) + then show ?thesis using Inr Inr2 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis + qed + qed + qed + next + case (2 y s ya xa va sa c) + thus ?case using eqvt_triple eqvt_at_proj by blast + next + case (3 y s2 ya xa va s1a s2a c) + thus ?case using eqvt_triple eqvt_at_proj by blast + next + case (4 u xa va s ua sa c) + moreover have "atom u \ (xa, va) \ atom ua \ (xa, va)" + using fresh_Pair u_fresh_xv by auto + ultimately show ?case using eqvt_triple[of u xa va ua s sa] subst_sv_def eqvt_at_proj by metis + next + case (5 x1 s1 x1a xa va s1a c) + thus ?case using eqvt_triple eqvt_at_proj by blast + } +qed +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_sv_abbrev :: "s \ x \ v \ s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000) + where + "s[x::=v]\<^sub>s\<^sub>v \ subst_sv s x v" + +abbreviation + subst_branchv_abbrev :: "branch_s \ x \ v \ branch_s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000) + where + "s[x::=v]\<^sub>s\<^sub>v \ subst_branchv s x v" + +lemma size_subst_sv [simp]: "size (subst_sv A i x ) = size A" and "size (subst_branchv B i x ) = size B" and "size (subst_branchlv C i x ) = size C" + by(nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct,auto) + +lemma forget_subst_sv [simp]: shows "atom a \ A \ subst_sv A a x = A" and "atom a \ B \ subst_branchv B a x = B" and "atom a \ C \ subst_branchlv C a x = C" + by (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct,auto simp: fresh_at_base) + +lemma subst_sv_id [simp]: "subst_sv A a (V_var a) = A" and "subst_branchv B a (V_var a) = B" and "subst_branchlv C a (V_var a) = C" +proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct) + case (AS_let x option e s) + then show ?case + by (metis (no_types, lifting) fresh_Pair not_None_eq subst_ev_id subst_sv.simps(2) subst_sv.simps(3) subst_tv_id v.fresh(2)) +next + case (AS_match v branch_s) + then show ?case using fresh_Pair not_None_eq subst_ev_id subst_sv.simps subst_sv.simps subst_tv_id v.fresh subst_vv_id + by metis +qed(auto)+ + +lemma fresh_subst_sv_if_rl: + shows + "(atom x \ s \ j \ s) \ (j \ v \ (j \ s \ j = atom x)) \ j \ (subst_sv s x v )" and + "(atom x \ cs \ j \ cs) \ (j \ v \ (j \ cs \ j = atom x)) \ j \ (subst_branchv cs x v)" and + "(atom x \ css \ j \ css) \ (j \ v \ (j \ css \ j = atom x)) \ j \ (subst_branchlv css x v )" + apply(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct) + using pure_fresh by force+ + +lemma fresh_subst_sv_if_lr: + shows "j \ (subst_sv s x v) \ (atom x \ s \ j \ s) \ (j \ v \ (j \ s \ j = atom x))" and + "j \ (subst_branchv cs x v) \ (atom x \ cs \ j \ cs) \ (j \ v \ (j \ cs \ j = atom x))" and + "j \ (subst_branchlv css x v ) \ (atom x \ css \ j \ css) \ (j \ v \ (j \ css \ j = atom x))" +proof(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct) + case (AS_branch list x s ) + then show ?case using s_branch_s_branch_list.fresh fresh_Pair list.distinct(1) list.set_cases pure_fresh set_ConsD subst_branchv.simps by metis +next + case (AS_let y e s') + thus ?case proof(cases "atom x \ (AS_let y e s')") + case True + hence "subst_sv (AS_let y e s') x v = (AS_let y e s')" using forget_subst_sv by simp + hence "j \ (AS_let y e s')" using AS_let by argo + then show ?thesis using True by blast + next + case False + have "subst_sv (AS_let y e s') x v = AS_let y (e[x::=v]\<^sub>e\<^sub>v) (s'[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps(2) AS_let by force + hence "((j \ s'[x::=v]\<^sub>s\<^sub>v \ j \ set [atom y]) \ j \ None \ j \ e[x::=v]\<^sub>e\<^sub>v)" using s_branch_s_branch_list.fresh AS_let + by (simp add: fresh_None) + then show ?thesis using AS_let fresh_None fresh_subst_ev_if list.discI list.set_cases s_branch_s_branch_list.fresh set_ConsD + by metis + qed +next + case (AS_let2 y \ s1 s2) + thus ?case proof(cases "atom x \ (AS_let2 y \ s1 s2)") + case True + hence "subst_sv (AS_let2 y \ s1 s2) x v = (AS_let2 y \ s1 s2)" using forget_subst_sv by simp + hence "j \ (AS_let2 y \ s1 s2)" using AS_let2 by argo + then show ?thesis using True by blast + next + case False + have "subst_sv (AS_let2 y \ s1 s2) x v = AS_let2 y (\[x::=v]\<^sub>\\<^sub>v) (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps AS_let2 by force + then show ?thesis using AS_let2 + fresh_subst_tv_if list.discI list.set_cases s_branch_s_branch_list.fresh(4) set_ConsD by auto + qed +qed(auto)+ + +lemma fresh_subst_sv_if[simp]: + fixes x::x and v::v + shows "j \ (subst_sv s x v) \ (atom x \ s \ j \ s) \ (j \ v \ (j \ s \ j = atom x))" and + "j \ (subst_branchv cs x v) \ (atom x \ cs \ j \ cs) \ (j \ v \ (j \ cs \ j = atom x))" + using fresh_subst_sv_if_lr fresh_subst_sv_if_rl by metis+ + +lemma subst_sv_commute [simp]: + fixes A::s and t::v and j::x and i::x + shows "atom j \ A \ (subst_sv (subst_sv A i t) j u ) = subst_sv A i (subst_vv t j u )" and + "atom j \ B \ (subst_branchv (subst_branchv B i t ) j u ) = subst_branchv B i (subst_vv t j u )" and + "atom j \ C \ (subst_branchlv (subst_branchlv C i t) j u ) = subst_branchlv C i (subst_vv t j u ) " + apply(nominal_induct A and B and C avoiding: i j t u rule: s_branch_s_branch_list.strong_induct) + by(auto simp: fresh_at_base) + +lemma c_eq_perm: + assumes "( (atom z) \ (atom z') ) \ c = c'" and "atom z' \ c" + shows "\ z : b | c \ = \ z' : b | c' \" + using \.eq_iff Abs1_eq_iff(3) + by (metis Nominal2_Base.swap_commute assms(1) assms(2) flip_def swap_fresh_fresh) + +lemma subst_sv_flip: + fixes s::s and sa::s and v'::v + assumes "atom c \ (s, sa)" and "atom c \ (v',x, xa, s, sa)" "atom x \ v'" and "atom xa \ v'" and "(x \ c) \ s = (xa \ c) \ sa" + shows "s[x::=v']\<^sub>s\<^sub>v = sa[xa::=v']\<^sub>s\<^sub>v" +proof - + have "atom x \ (s[x::=v']\<^sub>s\<^sub>v)" and xafr: "atom xa \ (sa[xa::=v']\<^sub>s\<^sub>v)" + and "atom c \ ( s[x::=v']\<^sub>s\<^sub>v, sa[xa::=v']\<^sub>s\<^sub>v)" using assms using fresh_subst_sv_if assms by( blast+ ,force) + + hence "s[x::=v']\<^sub>s\<^sub>v = (x \ c) \ (s[x::=v']\<^sub>s\<^sub>v)" by (simp add: flip_fresh_fresh fresh_Pair) + also have " ... = ((x \ c) \ s)[ ((x \ c) \ x) ::= ((x \ c) \ v') ]\<^sub>s\<^sub>v" using subst_sv_subst_branchv_subst_branchlv.eqvt by blast + also have "... = ((xa \ c) \ sa)[ ((x \ c) \ x) ::= ((x \ c) \ v') ]\<^sub>s\<^sub>v" using assms by presburger + also have "... = ((xa \ c) \ sa)[ ((xa \ c) \ xa) ::= ((xa \ c) \ v') ]\<^sub>s\<^sub>v" using assms + by (metis flip_at_simps(1) flip_fresh_fresh fresh_PairD(1)) + also have "... = (xa \ c) \ (sa[xa::=v']\<^sub>s\<^sub>v)" using subst_sv_subst_branchv_subst_branchlv.eqvt by presburger + also have "... = sa[xa::=v']\<^sub>s\<^sub>v" using xafr assms by (simp add: flip_fresh_fresh fresh_Pair) + finally show ?thesis by simp +qed + +lemma if_type_eq: + fixes \::\ and v::v and z1::x + assumes "atom z1' \ (v, ca, (x, b, c) #\<^sub>\ \, (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" and "atom z1 \ v" + and "atom z1 \ (za,ca)" and "atom z1' \ (za,ca)" + shows "(\ z1' : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v \) = \ z1 : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v \" +proof - + have "atom z1' \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )" using assms fresh_prod4 by blast + moreover hence "(CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v) = (z1' \ z1) \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )" + proof - + have "(z1' \ z1) \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ) = ( (z1' \ z1) \ (CE_val v == CE_val (V_lit ll)) IMP ((z1' \ z1) \ ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" + by auto + also have "... = ((CE_val v == CE_val (V_lit ll)) IMP ((z1' \ z1) \ ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" + using \atom z1 \ v\ assms + by (metis (mono_tags) \atom z1' \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )\ c.fresh(6) c.fresh(7) ce.fresh(1) flip_at_simps(2) flip_fresh_fresh fresh_at_base_permute_iff fresh_def supp_l_empty v.fresh(1)) + also have "... = ((CE_val v == CE_val (V_lit ll)) IMP (ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v ))" + using assms by fastforce + finally show ?thesis by auto + qed + ultimately show ?thesis + using \.eq_iff Abs1_eq_iff(3)[of z1' "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v" + z1 "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v"] by blast +qed + +lemma subst_sv_var_flip: + fixes x::x and s::s and z::x + shows "atom x \ s \ ((x \ z) \ s) = s[z::=[x]\<^sup>v]\<^sub>s\<^sub>v" and + "atom x \ cs \ ((x \ z) \ cs) = subst_branchv cs z [x]\<^sup>v" and + "atom x \ css \ ((x \ z) \ css) = subst_branchlv css z [x]\<^sup>v" + apply(nominal_induct s and cs and css avoiding: z x rule: s_branch_s_branch_list.strong_induct) + using [[simproc del: alpha_lst]] + apply (auto ) (* This unpacks subst, perm *) + using subst_tv_var_flip flip_fresh_fresh v.fresh s_branch_s_branch_list.fresh + subst_v_\_def subst_v_v_def subst_vv_var_flip subst_v_e_def subst_ev_var_flip pure_fresh apply auto + defer 1 (* Sometimes defering hard goals to the end makes it easier to finish *) + using x_fresh_u apply blast (* Next two involve u and flipping with x *) + defer 1 + using x_fresh_u apply blast + defer 1 + using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh + apply (simp add: subst_v_c_def) + using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh + by (simp add: flip_fresh_fresh) + +instantiation s :: has_subst_v +begin + +definition + "subst_v = subst_sv" + +instance proof + fix j::atom and i::x and x::v and t::s + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_sv_if subst_v_s_def by auto + + fix a::x and tm::s and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_sv subst_v_s_def by simp + + fix a::x and tm::s + show "subst_v tm a (V_var a) = tm" using subst_sv_id subst_v_s_def by simp + + fix p::perm and x1::x and v::v and t1::s + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_sv_commute subst_v_s_def by simp + + fix x::x and c::s and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_sv_var_flip subst_v_s_def by simp + + fix x::x and c::s and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_sv_var_flip subst_v_s_def by simp +qed +end + +section \Type Definition\ + +nominal_function subst_ft_v :: "fun_typ \ x \ v \ fun_typ" where + "atom z \ (x,v) \ subst_ft_v ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z b c[x::=v]\<^sub>c\<^sub>v t[x::=v]\<^sub>\\<^sub>v s[x::=v]\<^sub>s\<^sub>v" + apply(simp add: eqvt_def subst_ft_v_graph_aux_def ) + apply(simp add:fun_typ.strong_exhaust ) + apply(auto) + apply(rule_tac y=a and c="(aa,b)" in fun_typ.strong_exhaust) + apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) + +proof(goal_cases) + case (1 z xa va c t s za ca ta sa cb) + hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v" + by (metis flip_commute subst_cv_var_flip) + hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by auto + then show ?case using subst_cv_commute atom_eq_iff fresh_atom fresh_atom_at_base subst_cv_commute_full v.fresh + using 1 subst_cv_var_flip flip_commute by metis +next + case (2 z xa va c t s za ca ta sa cb) + hence "t[z::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v" by metis + hence "t[z::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v[xa::=va]\<^sub>\\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v[xa::=va]\<^sub>\\<^sub>v" by auto + then show ?case using subst_tv_commute_full 2 + by (metis atom_eq_iff fresh_atom fresh_atom_at_base v.fresh(2)) +qed + +nominal_termination (eqvt) by lexicographic_order + +nominal_function subst_ftq_v :: "fun_typ_q \ x \ v \ fun_typ_q" where + "atom bv \ (x,v) \ subst_ftq_v (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_v ft x v))" +| "subst_ftq_v (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_v ft x v))" + apply(simp add: eqvt_def subst_ftq_v_graph_aux_def ) + apply(simp add:fun_typ_q.strong_exhaust ) + apply(auto) + apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust) + apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) +proof(goal_cases) + case (1 bv ft bva fta xa va c) + then show ?case using subst_ft_v.simps by (simp add: flip_fresh_fresh) +qed +nominal_termination (eqvt) by lexicographic_order + +lemma size_subst_ft[simp]: "size (subst_ft_v A x v) = size A" + by(nominal_induct A avoiding: x v rule: fun_typ.strong_induct,auto) + +lemma forget_subst_ft [simp]: shows "atom x \ A \ subst_ft_v A x a = A" + by (nominal_induct A avoiding: a x rule: fun_typ.strong_induct,auto simp: fresh_at_base) + +lemma subst_ft_id [simp]: "subst_ft_v A a (V_var a) = A" + by(nominal_induct A avoiding: a rule: fun_typ.strong_induct,auto) + +instantiation fun_typ :: has_subst_v +begin + +definition + "subst_v = subst_ft_v" + +instance proof + + fix j::atom and i::x and x::v and t::fun_typ + show "(j \ subst_v 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(simp only: subst_v_fun_typ_def subst_ft_v.simps ) + using fun_typ.fresh fresh_subst_v_if apply simp + by auto + + fix a::x and tm::fun_typ and x::v + show "atom a \ tm \ subst_v tm a x = tm" + proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct) + case (AF_fun_typ x1a x2a x3a x4a x5a) + then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\_def by fastforce + qed + + fix a::x and tm::fun_typ + show "subst_v tm a (V_var a) = tm" + proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct) + case (AF_fun_typ x1a x2a x3a x4a x5a) + then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\_def by fastforce + qed + + fix p::perm and x1::x and v::v and t1::fun_typ + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + proof(nominal_induct t1 avoiding: x1 v rule:fun_typ.strong_induct) + case (AF_fun_typ x1a x2a x3a x4a x5a) + then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\_def by fastforce + qed + + fix x::x and c::fun_typ and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + apply(nominal_induct c avoiding: x z rule:fun_typ.strong_induct) + by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\_def subst_v_fun_typ_def) + + fix x::x and c::fun_typ and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + apply(nominal_induct c avoiding: z x v rule:fun_typ.strong_induct) + apply auto + by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\_def subst_v_fun_typ_def ) +qed +end + +instantiation fun_typ_q :: has_subst_v +begin + +definition + "subst_v = subst_ftq_v" + +instance proof + fix j::atom and i::x and x::v and t::fun_typ_q + show "(j \ subst_v 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_q.strong_induct,auto) + apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + by (metis (no_types) fresh_subst_v_if subst_v_fun_typ_def)+ + + fix i::x and t::fun_typ_q and x::v + show "atom i \ t \ subst_v t i x = t" + apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix i::x and t::fun_typ_q + show "subst_v t i (V_var i) = t" using subst_cv_id subst_v_fun_typ_def + apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix p::perm and x1::x and v::v and t1::fun_typ_q + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + apply(nominal_induct t1 avoiding: v x1 rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix x::x and c::fun_typ_q and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + apply(nominal_induct c avoiding: x z rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix x::x and c::fun_typ_q and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + apply(nominal_induct c avoiding: z x v rule:fun_typ_q.strong_induct,auto) + apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + by (metis subst_v_fun_typ_def flip_bv_x_cancel subst_ft_v.eqvt subst_v_simple_commute v.perm_simps )+ +qed + +end + +section \Variable Context\ + +lemma subst_dv_fst_eq: + "fst ` setD (\[x::=v]\<^sub>\\<^sub>v) = fst ` setD \" + by(induct \ rule: \_induct,simp,force) + +lemma subst_gv_member_iff: + fixes x'::x and x::x and v::v and c'::c + assumes "(x',b',c') \ toSet \" and "atom x \ atom_dom \" + shows "(x',b',c'[x::=v]\<^sub>c\<^sub>v) \ toSet \[x::=v]\<^sub>\\<^sub>v" +proof - + have "x' \ x" using assms fresh_dom_free2 by metis + then show ?thesis using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto + next + case (GCons x1 b1 c1 \') + show ?case proof(cases "(x',b',c') = (x1,b1,c1)") + case True + hence "((x1, b1, c1) #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v = ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v))" using subst_gv.simps \x'\x\ by auto + then show ?thesis using True by auto + next + case False + have "x1\x" using fresh_def fresh_GCons fresh_Pair supp_at_base GCons fresh_dom_free2 by auto + hence "(x', b', c') \ toSet \'" using GCons False toSet.simps by auto + moreover have "atom x \ atom_dom \'" using fresh_GCons GCons dom.simps toSet.simps by simp + ultimately have "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \ toSet \'[x::=v]\<^sub>\\<^sub>v" using GCons by auto + hence "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \ toSet ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v))" by auto + then show ?thesis using subst_gv.simps \x1\x\ by auto + qed + qed +qed + +lemma fresh_subst_gv_if: + fixes j::atom and i::x and x::v and t::\ + assumes "j \ t \ j \ x" + shows "(j \ subst_gv t i x)" + using assms proof(induct t rule: \_induct) + case GNil + then show ?case using subst_gv.simps fresh_GNil by auto +next + case (GCons x' b' c' \') + then show ?case unfolding subst_gv.simps using fresh_GCons fresh_subst_cv_if by auto +qed + +section \Lookup\ + +lemma set_GConsD: "y \ toSet (x #\<^sub>\ xs) \ y=x \ y \ toSet xs" + by auto + +lemma subst_g_assoc_cons: + assumes "x \ x'" + shows "(((x', b', c') #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v @ G) = ((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ ((\'[x::=v]\<^sub>\\<^sub>v) @ G))" + using subst_gv.simps append_g.simps assms by auto + +end \ No newline at end of file diff --git a/thys/MiniSail/IVSubstTypingL.thy b/thys/MiniSail/IVSubstTypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/IVSubstTypingL.thy @@ -0,0 +1,1534 @@ +(*<*) +theory IVSubstTypingL + imports SubstMethods ContextSubtypingL +begin + (*>*) + +chapter \Immutable Variable Substitution Lemmas\ +text \Lemmas that show that types are preserved, in some way, +under immutable variable substitution\ + +section \Proof Methods\ + +method subst_mth = (metis subst_g_inside infer_e_wf infer_v_wf infer_v_wf) + +method subst_tuple_mth uses add = ( + (unfold fresh_prodN), (simp add: add )+, + (rule,metis fresh_z_subst_g add fresh_Pair ), + (metis fresh_subst_dv add fresh_Pair ) ) + +section \Prelude\ + +lemma subst_top_eq: + "\ z : b | TRUE \ = \ z : b | TRUE \[x::=v]\<^sub>\\<^sub>v" +proof - + obtain z'::x and c' where zeq: "\ z : b | TRUE \ = \ z' : b | c' \ \ atom z' \ (x,v)" using obtain_fresh_z2 b_of.simps by metis + hence "\ z' : b | TRUE \[x::=v]\<^sub>\\<^sub>v = \ z' : b | TRUE \" using subst_tv.simps subst_cv.simps by metis + moreover have "c' = C_true" using \.eq_iff Abs1_eq_iff(3) c.fresh flip_fresh_fresh by (metis zeq) + ultimately show ?thesis using zeq by metis +qed + +lemma wfD_subst: + fixes \\<^sub>1::\ and v::v and \::\ and \::\ and \::\ + assumes "\ ; \ ; \ \ v \ \\<^sub>1" and "wfD \ \ (\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)) \" and "b_of \\<^sub>1=b\<^sub>1" + shows " \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v" +proof - + have "\ ; \ ; \\\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_v_wf assms by auto + moreover have "(\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\))[x::=v]\<^sub>\\<^sub>v = \'[x::=v]\<^sub>\\<^sub>v @ \" using subst_g_inside wfD_wf assms by metis + ultimately show ?thesis using wf_subst assms by metis +qed + +lemma subst_v_c_of: + assumes "atom xa \ (v,x)" + shows "c_of t[x::=v]\<^sub>\\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v" + using assms proof(nominal_induct t avoiding: x v xa rule:\.strong_induct) + case (T_refined_type z' b' c') + then have " c_of \ z' : b' | c' \[x::=v]\<^sub>\\<^sub>v xa = c_of \ z' : b' | c'[x::=v]\<^sub>c\<^sub>v \ xa" + using subst_tv.simps fresh_Pair by metis + also have "... = c'[x::=v]\<^sub>c\<^sub>v [z'::=V_var xa]\<^sub>c\<^sub>v" using c_of.simps T_refined_type by metis + also have "... = c' [z'::=V_var xa]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v" + using subst_cv_commute_full[of z' v x "V_var xa" c'] subst_v_c_def T_refined_type fresh_Pair fresh_at_base v.fresh fresh_x_neq by metis + finally show ?case using c_of.simps T_refined_type by metis +qed + +section \Context\ + +lemma subst_lookup: + assumes "Some (b,c) = lookup (\'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\\)) y" and "x \ y" and "wfG \ \ (\'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\\))" + shows "\d. Some (b,d) = lookup ((\'[x::=v]\<^sub>\\<^sub>v)@\) y" + using assms proof(induct \' rule: \_induct) + case GNil + hence "Some (b,c) = lookup \ y" by (simp add: assms(1)) + then show ?case using subst_gv.simps by auto +next + case (GCons x1 b1 c1 \1) + show ?case proof(cases "x1 = x") + case True + hence "atom x \ (\1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\ \)" using GCons wfG_elims(2) + append_g.simps by metis + moreover have "atom x \ atom_dom (\1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\ \)" by simp + ultimately show ?thesis + using forget_subst_gv not_GCons_self2 subst_gv.simps append_g.simps + by (metis GCons.prems(3) True wfG_cons_fresh2 ) + 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=y") + case True + then show ?thesis using GCons using lookup.simps + by (metis \((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\ append_g.simps fst_conv option.inject) + next + case False + then show ?thesis using GCons using lookup.simps + using \((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\ append_g.simps \.distinct \.inject wfG.simps wfG_elims by metis + qed + qed +qed + +section \Validity\ + +lemma subst_self_valid: + fixes v::v + assumes "\ ; \ ; G \ v \ \ z : b | c \" and "atom z \ v" + shows " \ ; \ ; G \ c[z::=v]\<^sub>c\<^sub>v" +proof - + have "c= (CE_val (V_var z) == CE_val v )" using infer_v_form2 assms by presburger + hence "c[z::=v]\<^sub>c\<^sub>v = (CE_val (V_var z) == CE_val v )[z::=v]\<^sub>c\<^sub>v" by auto + also have "... = (((CE_val (V_var z))[z::=v]\<^sub>c\<^sub>e\<^sub>v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" by fastforce + also have "... = ((CE_val v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" using subst_cev.simps subst_vv.simps by presburger + also have "... = (CE_val v == CE_val v )" using infer_v_form subst_cev.simps assms forget_subst_vv by presburger + finally have *:"c[z::=v]\<^sub>c\<^sub>v = (CE_val v == CE_val v )" by auto + + have **:"\ ; \ ; G\\<^sub>w\<^sub>f CE_val v : b" using wfCE_valI assms infer_v_v_wf b_of.simps by metis + + show ?thesis proof(rule validI) + show "\ ; \ ; G\\<^sub>w\<^sub>f c[z::=v]\<^sub>c\<^sub>v" proof - + have "\ ; \ ; G\\<^sub>w\<^sub>f v : b" using infer_v_v_wf assms b_of.simps by metis + moreover have "\ \\<^sub>w\<^sub>f ([]::\) \ \ ; \ ; G\\<^sub>w\<^sub>f []\<^sub>\" using wfD_emptyI wfPhi_emptyI infer_v_wf assms by auto + ultimately show ?thesis using * wfCE_valI wfC_eqI by metis + qed + show "\i. wfI \ G i \ is_satis_g i G \ is_satis i c[z::=v]\<^sub>c\<^sub>v" proof(rule,rule) + fix i + assume \wfI \ G i \ is_satis_g i G\ + thus \is_satis i c[z::=v]\<^sub>c\<^sub>v\ using * ** is_satis_eq by auto + qed + qed +qed + +lemma subst_valid_simple: + fixes v::v + assumes "\ ; \ ; G \ v \ \ z0 : b | c0 \" and + "atom z0 \ c" and "atom z0 \ v" + "\; \ ; (z0,b,c0)#\<^sub>\G \ c[z::=V_var z0]\<^sub>c\<^sub>v" + shows " \ ; \ ; G \ c[z::=v]\<^sub>c\<^sub>v" +proof - + have " \ ; \ ; G \ c0[z0::=v]\<^sub>c\<^sub>v" using subst_self_valid assms by metis + moreover have "atom z0 \ G" using assms valid_wf_all by meson + moreover have "wfV \ \ G v b" using infer_v_v_wf assms b_of.simps by metis + moreover have "(c[z::=V_var z0]\<^sub>c\<^sub>v)[z0::=v]\<^sub>c\<^sub>v = c[z::=v]\<^sub>c\<^sub>v" using subst_v_simple_commute assms subst_v_c_def by metis + ultimately show ?thesis using valid_trans assms subst_defs by metis +qed + +lemma wfI_subst1: + assumes " wfI \ (G'[x::=v]\<^sub>\\<^sub>v @ G) i" and "wfG \ \ (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ G)" and "eval_v i v sv" and "wfRCV \ sv b" + shows "wfI \ (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ G) ( i( x \ sv))" +proof - + { + fix xa::x and ba::b and ca::c + assume as: "(xa,ba,ca) \ toSet ((G' @ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ G)))" + then have " \s. Some s = (i(x \ sv)) xa \ wfRCV \ s ba" + proof(cases "x=xa") + case True + have "Some sv = (i(x \ sv)) x \ wfRCV \ sv b" using as assms wfI_def by auto + moreover have "b=ba" using assms as True wfG_member_unique by metis + ultimately show ?thesis using True by auto + next + case False + + then obtain ca' where "(xa, ba, ca') \ toSet (G'[x::=v]\<^sub>\\<^sub>v @ G)" using wfG_member_subst2 assms as by metis + then obtain s where " Some s = i xa \ wfRCV \ s ba" using wfI_def assms False by blast + thus ?thesis using False by auto + qed + } + from this show ?thesis using wfI_def allI by blast +qed + +lemma subst_valid: + fixes v::v and c'::c and \ ::\ + assumes "\ ; \ ; \ \ c[z::=v]\<^sub>c\<^sub>v" and "\ ; \ ; \\\<^sub>w\<^sub>f v : b" and + "\ ; \\\<^sub>w\<^sub>f \" and "atom x \ c" and "atom x \ \" and + "\ ; \\\<^sub>w\<^sub>f (\'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ \)" and + "\ ; \ ; \'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ \ \ c'" (is " \ ; \; ?G \ c'") + shows "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v@\ \ c'[x::=v]\<^sub>c\<^sub>v" +proof - + have *:"wfC \ \ (\'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ \) c'" using valid.simps assms by metis + hence "wfC \ \ (\'[x::=v]\<^sub>\\<^sub>v @ \) (c'[x::=v]\<^sub>c\<^sub>v)" using wf_subst(2)[OF *] b_of.simps assms subst_g_inside wfC_wf by metis + moreover have "\i. wfI \ (\'[x::=v]\<^sub>\\<^sub>v @ \) i \ is_satis_g i (\'[x::=v]\<^sub>\\<^sub>v @ \) \ is_satis i (c'[x::=v]\<^sub>c\<^sub>v)" + + proof(rule,rule) + fix i + assume as: " wfI \ (\'[x::=v]\<^sub>\\<^sub>v @ \) i \ is_satis_g i (\'[x::=v]\<^sub>\\<^sub>v @ \)" + + hence wfig: "wfI \ \ i" using wfI_suffix infer_v_wf assms by metis + then obtain s where s:"eval_v i v s" and b:"wfRCV \ s b" using eval_v_exist infer_v_v_wf b_of.simps assms by metis + + have is1: "is_satis_g ( i( x \ s)) (\' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" proof(rule is_satis_g_i_upd2) + show "is_satis (i(x \ s)) (c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)" proof - + have "is_satis i (c[z::=v]\<^sub>c\<^sub>v)" + using subst_valid_simple assms as valid.simps infer_v_wf assms + is_satis_g_suffix wfI_suffix by metis + hence "is_satis i ((c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v)" using assms subst_v_simple_commute[of x c z v] subst_v_c_def by metis + moreover have "\ ; \ ; (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" using wfC_refl wfG_suffix assms by metis + moreover have "\ ; \ ; \\\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf b_of.simps by metis + ultimately show ?thesis using subst_c_satis[OF s , of \ \ x b "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" \ "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v"] wfig by auto + qed + show "atom x \ \" using assms by metis + show "wfG \ \ (\' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" using valid_wf_all assms by metis + show "\ ; \ ; \\\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf by force + show "i \ v \ ~ s " using s by auto + show "\ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ i" using as by auto + show "i \ \'[x::=v]\<^sub>\\<^sub>v @ \ " using as by auto + qed + hence "is_satis ( i( x \ s)) c'" proof - + have "wfI \ (\' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) ( i( x \ s))" + using wfI_subst1[of \ \' x v \ i \ b c z s] as b s assms by metis + thus ?thesis using is1 valid.simps assms by presburger + qed + + thus "is_satis i (c'[x::=v]\<^sub>c\<^sub>v)" using subst_c_satis_full[OF s] valid.simps as infer_v_v_wf b_of.simps assms by metis + + qed + ultimately show ?thesis using valid.simps by auto +qed + +lemma subst_valid_infer_v: + fixes v::v and c'::c + assumes "\ ; \ ; G \ v \ \ z0 : b | c0 \" and "atom x \ c" and "atom x \ G" and "wfG \ \ (G'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ G)" and "atom z0 \ v" + " \;\;(z0,b,c0)#\<^sub>\G \ c[z::=V_var z0]\<^sub>c\<^sub>v" and "atom z0 \ c" and + " \;\;G'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ G \ c'" (is " \ ; \; ?G \ c'") + shows " \;\;G'[x::=v]\<^sub>\\<^sub>v@G \ c'[x::=v]\<^sub>c\<^sub>v" +proof - + have "\ ; \ ; G \ c[z::=v]\<^sub>c\<^sub>v" + using infer_v_wf subst_valid_simple valid.simps assms using subst_valid_simple assms valid.simps infer_v_wf assms + is_satis_g_suffix wfI_suffix by metis + moreover have "wfV \ \ G v b" and "wfG \ \ G" + using assms infer_v_wf b_of.simps apply metis using assms infer_v_wf by metis + ultimately show ?thesis using assms subst_valid by metis +qed + +section \Subtyping\ + +lemma subst_subtype: + fixes v::v + assumes "\ ; \ ; \ \ v \ (\z0:b|c0\)" and + " \;\;\ \ (\z0:b|c0\) \ (\ z : b | c \)" and + " \;\;\'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ (\ z1 : b1 | c1 \) \ (\ z2 : b1 | c2 \)" (is " \ ; \; ?G1 \ ?t1 \ ?t2" ) and + "atom z \ (x,v) \ atom z0 \ (c,x,v,z,\) \ atom z1 \ (x,v) \ atom z2 \ (x,v)" and "wsV \ \ \ v" + shows " \;\;\'[x::=v]\<^sub>\\<^sub>v@\ \ \ z1 : b1 | c1 \[x::=v]\<^sub>\\<^sub>v \ \ z2 : b1 | c2 \[x::=v]\<^sub>\\<^sub>v" +proof - + have z2: "atom z2 \ (x,v) " using assms by auto + hence "x \ z2" by auto + + obtain xx::x where xxf: "atom xx \ (x,z1, c1, z2, c2, \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \, c1[x::=v]\<^sub>c\<^sub>v, c2[x::=v]\<^sub>c\<^sub>v, \'[x::=v]\<^sub>\\<^sub>v @ \, + (\ , \ , \'[x::=v]\<^sub>\\<^sub>v@\, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 , c2[x::=v]\<^sub>c\<^sub>v ))" (is "atom xx \ ?tup") + using obtain_fresh by blast + hence xxf2: "atom xx \ (z1, c1, z2, c2, \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" using fresh_prod9 fresh_prod5 by fast + + have vd1: " \;\;((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v @ \ \ (c2[z2::=V_var xx]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v" + proof(rule subst_valid_infer_v[of \ _ _ _ z0 b c0 _ c, where z=z]) + show "\ ; \ ; \ \ v \ \ z0 : b | c0 \" using assms by auto + + show xf: "atom x \ \" using subtype_g_wf wfG_inside_fresh_suffix assms by metis + + show "atom x \ c" proof - + have "wfT \ \ \ (\ z : b | c \)" using subtype_wf[OF assms(2)] by auto + moreover have "x \ z" using assms(4) + using fresh_Pair not_self_fresh by blast + ultimately show ?thesis using xf wfT_fresh_c assms by presburger + qed + + show " \ ; \\\<^sub>w\<^sub>f ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ \') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ " + proof(subst append_g.simps,rule wfG_consI) + show *: \ \ ; \\\<^sub>w\<^sub>f \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ using subtype_g_wf assms by metis + show \atom xx \ \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\ using xxf fresh_prod9 by metis + show \ \ ; \\\<^sub>w\<^sub>f b1 \ using subtype_elims[OF assms(3)] wfT_wfC wfC_wf wfG_cons by metis + show "\ ; \ ; (xx, b1, TRUE) #\<^sub>\ \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c1[z1::=V_var xx]\<^sub>c\<^sub>v " proof(rule wfT_wfC) + have "\ z1 : b1 | c1 \ = \ xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \ " using xxf fresh_prod9 type_eq_subst xxf2 fresh_prodN by metis + thus "\ ; \ ; \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f \ xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \ " using subtype_wfT[OF assms(3)] by metis + show "atom xx \ \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \" using xxf fresh_prod9 by metis + qed + qed + + show "atom z0 \ v" using assms fresh_prod5 by auto + have "\ ; \ ; (z0, b, c0) #\<^sub>\ \ \ c[z::=V_var z0]\<^sub>v " + apply(rule obtain_fresh[of "(z0,c0, \, c, z)"],rule subtype_valid[OF assms(2), THEN valid_flip], + (fastforce simp add: assms fresh_prodN)+) done + thus "\ ; \ ; (z0, b, c0) #\<^sub>\ \ \ c[z::=V_var z0]\<^sub>c\<^sub>v " using subst_defs by auto + + show "atom z0 \ c" using assms fresh_prod5 by auto + show "\ ; \ ; ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ \') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ c2[z2::=V_var xx]\<^sub>c\<^sub>v " + using subtype_valid assms(3) xxf xxf2 fresh_prodN append_g.simps subst_defs by metis + qed + + have xfw1: "atom z1 \ v \ atom x \ [ xx ]\<^sup>v \ x \ z1" + apply(intro conjI) + apply(simp add: assms xxf fresh_at_base fresh_prodN freshers fresh_x_neq)+ + using fresh_x_neq fresh_prodN xxf apply blast + using fresh_x_neq fresh_prodN assms by blast + + have xfw2: "atom z2 \ v \ atom x \ [ xx ]\<^sup>v \ x \ z2" + apply(auto simp add: assms xxf fresh_at_base fresh_prodN freshers) + by(insert xxf fresh_at_base fresh_prodN assms, fast+) + + have wf1: "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \)" proof - + have "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z1 : b1 | c1 \)[x::=v]\<^sub>\\<^sub>v" + using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis + moreover have "atom z1 \ (x,v)" using assms by auto + ultimately show ?thesis using subst_tv.simps by auto + qed + moreover have wf2: "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \)" proof - + have "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z2 : b1 | c2 \)[x::=v]\<^sub>\\<^sub>v" using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis + moreover have "atom z2 \ (x,v)" using assms by auto + ultimately show ?thesis using subst_tv.simps by auto + qed + moreover have "\ ; \ ; (xx, b1, c1[x::=v]\<^sub>c\<^sub>v[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v @ \ ) \ (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" proof - + have "xx \ x" using xxf fresh_Pair fresh_at_base by fast + hence "((xx, b1, subst_cv c1 z1 (V_var xx) ) #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v = (xx, b1, (subst_cv c1 z1 (V_var xx) )[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v)" + using subst_gv.simps by auto + moreover have "(c1[z1::=V_var xx]\<^sub>c\<^sub>v )[x::=v]\<^sub>c\<^sub>v = (c1[x::=v]\<^sub>c\<^sub>v) [z1::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw1 by metis + moreover have "c2[z2::=[ xx ]\<^sup>v]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v = (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw2 by metis + ultimately show ?thesis using vd1 append_g.simps by metis + qed + moreover have "atom xx \ (\ , \ , \'[x::=v]\<^sub>\\<^sub>v@\, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 ,c2[x::=v]\<^sub>c\<^sub>v )" + using xxf fresh_prodN by metis + ultimately have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v@\ \ \ z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \ \ \ z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \" + using subtype_baseI subst_defs by metis + thus ?thesis using subst_tv.simps assms by presburger +qed + +lemma subst_subtype_tau: + fixes v::v + assumes "\ ; \ ; \ \ v \ \" and + "\ ; \ ; \ \ \ \ (\ z : b | c \)" + "\ ; \ ; \'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ \1 \ \2" and + "atom z \ (x,v)" + shows "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v@\ \ \1[x::=v]\<^sub>\\<^sub>v \ \2[x::=v]\<^sub>\\<^sub>v" +proof - + obtain z0 and b0 and c0 where zbc0: "\=(\ z0 : b0 | c0 \) \ atom z0 \ (c,x,v,z,\)" + using obtain_fresh_z by metis + obtain z1 and b1 and c1 where zbc1: "\1=(\ z1 : b1 | c1 \) \ atom z1 \ (x,v)" + using obtain_fresh_z by metis + obtain z2 and b2 and c2 where zbc2: "\2=(\ z2 : b2 | c2 \) \ atom z2 \ (x,v)" + using obtain_fresh_z by metis + + have "b0=b" using subtype_eq_base zbc0 assms by blast + + hence vinf: "\ ; \ ; \ \ v \ \ z0 : b | c0 \" using assms zbc0 by blast + have vsub: "\ ; \ ; \ \\ z0 : b | c0 \ \ \ z : b | c \" using assms zbc0 \b0=b\ by blast + have beq:"b1=b2" using subtype_eq_base + using zbc1 zbc2 assms by blast + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ \ z1 : b1 | c1 \[x::=v]\<^sub>\\<^sub>v \ \ z2 : b1 | c2 \[x::=v]\<^sub>\\<^sub>v" + proof(rule subst_subtype[OF vinf vsub] ) + show "\ ; \ ; \'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ \ z1 : b1 | c1 \ \ \ z2 : b1 | c2 \" + using beq assms zbc1 zbc2 by auto + show "atom z \ (x, v) \ atom z0 \ (c, x, v, z, \) \ atom z1 \ (x, v) \ atom z2 \ (x, v)" + using zbc0 zbc1 zbc2 assms by blast + show "wfV \ \ \ v (b_of \)" using infer_v_wf assms by simp + qed + + thus ?thesis using zbc1 zbc2 \b1=b2\ assms by blast +qed + +lemma subtype_if1: + fixes v::v + assumes "P ; \ ; \ \ t1 \ t2" and "wfV P \ \ v (base_for_lit l)" and + "atom z1 \ v" and "atom z2 \ v" and "atom z1 \ t1" and "atom z2 \ t2" and "atom z1 \ \" and "atom z2 \ \" + shows "P ; \ ; \ \ \ z1 : b_of t1 | CE_val v == CE_val (V_lit l) IMP (c_of t1 z1) \ \ \ z2 : b_of t2 | CE_val v == CE_val (V_lit l) IMP (c_of t2 z2) \" +proof - + obtain z1' where t1: "t1 = \ z1' : b_of t1 | c_of t1 z1'\ \ atom z1' \ (z1,\,t1)" using obtain_fresh_z_c_of by metis + obtain z2' where t2: "t2 = \ z2' : b_of t2 | c_of t2 z2'\ \ atom z2' \ (z2,t2) " using obtain_fresh_z_c_of by metis + have beq:"b_of t1 = b_of t2" using subtype_eq_base2 assms by auto + + have c1: "(c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t1 z1" using c_of_switch t1 assms by simp + have c2: "(c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t2 z2" using c_of_switch t2 assms by simp + + have "P ; \ ; \ \ \ z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[z1]\<^sup>v]\<^sub>v \ \ \ z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[z2]\<^sup>v]\<^sub>v \" + proof(rule subtype_if) + show \P ; \ ; \ \ \ z1' : b_of t1 | c_of t1 z1' \ \ \ z2' : b_of t1 | c_of t2 z2' \\ using t1 t2 assms beq by auto + show \ P ; \ ; \ \\<^sub>w\<^sub>f \ z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>v \ \ using wfT_wfT_if_rev assms subtype_wfT c1 subst_defs by metis + show \ P ; \ ; \ \\<^sub>w\<^sub>f \ z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>v \ \ using wfT_wfT_if_rev assms subtype_wfT c2 subst_defs beq by metis + show \atom z1 \ v\ using assms by auto + show \atom z1' \ \\ using t1 by auto + show \atom z1 \ c_of t1 z1'\ using t1 assms c_of_fresh by force + show \atom z2 \ c_of t2 z2'\ using t2 assms c_of_fresh by force + show \atom z2 \ v\ using assms by auto + qed + then show ?thesis using t1 t2 assms c1 c2 beq subst_defs by metis +qed + +section \Values\ + +lemma subst_infer_aux: + fixes \\<^sub>1::\ and v'::v + assumes "\ ; \ ; \ \ v'[x::=v]\<^sub>v\<^sub>v \ \\<^sub>1" and "\ ; \ ; \' \ v' \ \\<^sub>2" and "b_of \\<^sub>1 = b_of \\<^sub>2" + shows "\\<^sub>1 = (\\<^sub>2[x::=v]\<^sub>\\<^sub>v)" +proof - + obtain z1 and b1 where zb1: "\\<^sub>1 = (\ z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \) \ atom z1 \ ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v),v'[x::=v]\<^sub>v\<^sub>v)" + using infer_v_form_fresh[OF assms(1)] by fastforce + obtain z2 and b2 where zb2: "\\<^sub>2 = (\ z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \) \ atom z2 \ ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v,x,v),v')" + using infer_v_form_fresh [OF assms(2)] by fastforce + have beq: "b1 = b2" using assms zb1 zb2 by simp + + hence "(\ z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \)[x::=v]\<^sub>\\<^sub>v = (\ z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \)" + using subst_tv.simps subst_cv.simps subst_ev.simps forget_subst_vv[of x "V_var z2"] zb2 by force + also have "... = (\ z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \)" + using type_e_eq[of z2 "CE_val (v'[x::=v]\<^sub>v\<^sub>v)"z1 b1 ] zb1 zb2 fresh_PairD(1) assms beq by metis + finally show ?thesis using zb1 zb2 by argo +qed + +lemma subst_t_b_eq: + fixes x::x and v::v + shows "b_of (\[x::=v]\<^sub>\\<^sub>v) = b_of \" +proof - + obtain z and b and c where "\ = \ z : b | c \ \ atom z \ (x,v)" + using has_fresh_z by blast + thus ?thesis using subst_tv.simps by simp +qed + +lemma fresh_g_fresh_v: + fixes x::x + assumes "atom x \ \" and "wfV \ \ \ v b" + shows "atom x \ v" + using assms wfV_supp wfX_wfY wfG_atoms_supp_eq fresh_def + by (metis wfV_x_fresh) + +lemma infer_v_fresh_g_fresh_v: + fixes x::x and \::\ and v::v + assumes "atom x \ \'@\" and "\ ; \ ; \ \ v \ \" + shows "atom x \ v" +proof - + have "atom x \ \" using fresh_suffix assms by auto + moreover have "wfV \ \ \ v (b_of \)" using infer_v_wf assms by auto + ultimately show ?thesis using fresh_g_fresh_v by metis +qed + +lemma infer_v_fresh_g_fresh_xv: + fixes xa::x and v::v and \::\ + assumes "atom xa \ \'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)" and "\ ; \ ; \ \ v \ \" + shows "atom xa \ (x,v)" +proof - + have "atom xa \ x" using assms fresh_in_g fresh_def by blast + moreover have "\'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) = ((\'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\GNil)@\)" using append_g.simps append_g_assoc by simp + moreover hence "atom xa \ v" using infer_v_fresh_g_fresh_v assms by metis + ultimately show ?thesis by auto +qed + +lemma wfG_subst_infer_v: + fixes v::v + assumes "\ ; \ \\<^sub>w\<^sub>f \' @ (x, b, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \" and "\ ; \ ; \ \ v \ \" and "b_of \ = b" + shows "\ ; \\\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \ " + using wfG_subst_wfV infer_v_v_wf assms by auto + +lemma fresh_subst_gv_inside: + fixes \::\ + assumes "atom z \ \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \" and "atom z \ v" + shows "atom z \ \'[x::=v]\<^sub>\\<^sub>v@\" + unfolding fresh_append_g using fresh_append_g assms fresh_subst_gv fresh_GCons by metis + +lemma subst_t: + fixes x::x and b::b and xa::x + assumes "atom z \ x" and "atom z \ v" + shows "(\ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v'[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \) = (\ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v)" + using assms subst_vv.simps subst_tv.simps subst_cv.simps subst_cev.simps by auto + +lemma infer_l_fresh: + assumes "\ l \ \" + shows "atom x \ \" +proof - + have "[] ; {||} \\<^sub>w\<^sub>f GNil" using wfG_nilI wfTh_emptyI by auto + hence "[] ; {||} ; GNil \\<^sub>w\<^sub>f \" using assms infer_l_wf by auto + thus ?thesis using fresh_def wfT_supp by force +qed + +lemma subst_infer_v: + fixes v::v and v'::v + assumes "\ ; \ ; \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ v' \ \\<^sub>2" and + "\ ; \ ; \ \ v \ \\<^sub>1" and + "\ ; \ ; \ \ \\<^sub>1 \ (\ z0 : b\<^sub>1 | c0 \)" and "atom z0 \ (x,v)" + shows "\ ; \ ; (\'[x::=v]\<^sub>\\<^sub>v)@\ \ v'[x::=v]\<^sub>v\<^sub>v \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct "\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)" v' \\<^sub>2 avoiding: x v rule: infer_v.strong_induct) + case (infer_v_varI \ \ b c xa z) + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof(cases "x=xa") + case True + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \" + proof(rule infer_v_g_weakening) + show *:\ \ ; \ ; \ \ v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \\ + using infer_v_form infer_v_varI + by (metis True lookup_inside_unique_b lookup_inside_wf ms_fresh_all(32) subtype_eq_base type_e_eq) + show \toSet \ \ toSet (\'[x::=v]\<^sub>\\<^sub>v @ \)\ by simp + have "\ ; \ ; \ \\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis + thus \ \ ; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \ \ + using wfG_subst[OF infer_v_varI(3), of \' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \ v] subst_g_inside infer_v_varI by metis + qed + + thus ?thesis using subst_vv.simps True by simp + next + case False + then obtain c' where c: "Some (b, c') = lookup (\'[x::=v]\<^sub>\\<^sub>v @ \) xa" using lookup_subst2 infer_v_varI by metis + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ [ xa ]\<^sup>v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v ]\<^sup>c\<^sup>e \" + proof + have "\ ; \ ; \ \\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis + thus "\ ; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \" using infer_v_varI + using wfG_subst[OF infer_v_varI(3), of \' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \ v] subst_g_inside infer_v_varI by metis + show "atom z \ xa" using infer_v_varI by auto + show "Some (b, c') = lookup (\'[x::=v]\<^sub>\\<^sub>v @ \) xa" using c by auto + show "atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \)" by (fresh_mth add: infer_v_varI fresh_subst_gv_inside) + qed + then show ?thesis using subst_vv.simps False by simp + qed + thus ?case using subst_t fresh_prodN infer_v_varI by metis +next + case (infer_v_litI \ \ l \) + show ?case unfolding subst_vv.simps proof + show "\ ; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \" using wfG_subst_infer_v infer_v_litI subtype_eq_base2 b_of.simps by metis + have "atom x \ \" using infer_v_litI infer_l_fresh by metis + thus "\ l \ \[x::=v]\<^sub>\\<^sub>v" using infer_v_litI type_v_subst_fresh by simp + qed +next + case (infer_v_pairI z v1 v2 \ \ t1 t2) + have " \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ + \ \ [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v \ \ z : [ b_of t1[x::=v]\<^sub>\\<^sub>v , b_of + t2[x::=v]\<^sub>\\<^sub>v ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v ]\<^sup>c\<^sup>e \" + proof + show \atom z \ (v1[x::=v]\<^sub>v\<^sub>v, v2[x::=v]\<^sub>v\<^sub>v)\ by (fresh_mth add: infer_v_pairI) + show \atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \)\ by (fresh_mth add: infer_v_pairI fresh_subst_gv_inside) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ t1[x::=v]\<^sub>\\<^sub>v\ using infer_v_pairI by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ t2[x::=v]\<^sub>\\<^sub>v\ using infer_v_pairI by metis + qed + then show ?case using subst_vv.simps subst_tv.simps infer_v_pairI b_of_subst by simp +next + case (infer_v_consI s dclist \ dc tc \ va tv z) + + have " \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ (V_cons s dc va[x::=v]\<^sub>v\<^sub>v) \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof + show td:\AF_typedef s dclist \ set \\ using infer_v_consI by auto + show dc:\(dc, tc) \ set dclist\ using infer_v_consI by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ va[x::=v]\<^sub>v\<^sub>v \ tv[x::=v]\<^sub>\\<^sub>v\ using infer_v_consI by auto + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc[x::=v]\<^sub>\\<^sub>v\ + using subst_subtype_tau infer_v_consI by metis + moreover have "atom x \ tc" using wfTh_lookup_supp_empty[OF td dc] infer_v_wf infer_v_consI fresh_def by fast + ultimately show \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc\ by simp + show \atom z \ va[x::=v]\<^sub>v\<^sub>v\ using infer_v_consI by auto + show \atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \)\ by (fresh_mth add: infer_v_consI fresh_subst_gv_inside) + qed + thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_consI by metis + +next + case (infer_v_conspI s bv dclist \ dc tc \ va tv b z) + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ (V_consp s dc b va[x::=v]\<^sub>v\<^sub>v) \ \ z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof + show td:\AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show dc:\(dc, tc) \ set dclist\ using infer_v_conspI by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ va[x::=v]\<^sub>v\<^sub>v \ tv[x::=v]\<^sub>\\<^sub>v\ using infer_v_conspI by metis + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc[bv::=b]\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v\ + using subst_subtype_tau infer_v_conspI by metis + moreover have "atom x \ tc[bv::=b]\<^sub>\\<^sub>b" proof - + have "supp tc \ { atom bv }" using wfTh_poly_lookup_supp infer_v_conspI wfX_wfY by metis + hence "atom x \ tc" using x_not_in_b_set + using fresh_def by fastforce + moreover have "atom x \ b" using x_fresh_b by auto + ultimately show ?thesis using fresh_subst_if subst_b_\_def by metis + qed + ultimately show \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc[bv::=b]\<^sub>\\<^sub>b\ by simp + show \atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \, va[x::=v]\<^sub>v\<^sub>v, b)\ proof - + have "atom z \ va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis + moreover have "atom z \ \'[x::=v]\<^sub>\\<^sub>v @ \" using fresh_subst_gv_inside infer_v_conspI by metis + ultimately show ?thesis using fresh_prodN infer_v_conspI by metis + qed + show \atom bv \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \, va[x::=v]\<^sub>v\<^sub>v, b)\ proof - + have "atom bv \ va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis + moreover have "atom bv \ \'[x::=v]\<^sub>\\<^sub>v @ \" using fresh_subst_gv_inside infer_v_conspI by metis + ultimately show ?thesis using fresh_prodN infer_v_conspI by metis + qed + show "\ ; \ \\<^sub>w\<^sub>f b" using infer_v_conspI by auto + qed + thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_conspI by metis + +qed + +lemma subst_infer_check_v: + fixes v::v and v'::v + assumes "\ ; \ ; \ \ v \ \\<^sub>1" and + "check_v \ \ (\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)) v' \\<^sub>2" and + "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" and "atom z0 \ (x,v)" + shows "check_v \ \ ((\'[x::=v]\<^sub>\\<^sub>v)@\) (v'[x::=v]\<^sub>v\<^sub>v) (\\<^sub>2[x::=v]\<^sub>\\<^sub>v)" +proof - + obtain \\<^sub>2' where t2: "infer_v \ \ (\' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) v' \\<^sub>2' \ \ ; \ ; (\' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) \ \\<^sub>2' \ \\<^sub>2" + using check_v_elims assms by blast + hence "infer_v \ \ ((\'[x::=v]\<^sub>\\<^sub>v)@\) (v'[x::=v]\<^sub>v\<^sub>v) (\\<^sub>2'[x::=v]\<^sub>\\<^sub>v)" + using subst_infer_v[OF _ assms(1) assms(3) assms(4)] by blast + moreover hence "\; \ ; ((\'[x::=v]\<^sub>\\<^sub>v)@\) \\\<^sub>2'[x::=v]\<^sub>\\<^sub>v \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using subst_subtype assms t2 by (meson subst_subtype_tau subtype_trans) + ultimately show ?thesis using check_v.intros by blast +qed + +lemma type_veq_subst[simp]: + assumes "atom z \ (x,v)" + shows "\ z : b | CE_val (V_var z) == CE_val v' \[x::=v]\<^sub>\\<^sub>v = \ z : b | CE_val (V_var z) == CE_val v'[x::=v]\<^sub>v\<^sub>v \" + using assms by auto + +lemma subst_infer_v_form: + fixes v::v and v'::v and \::\ + assumes "\ ; \ ; \ \ v \ \\<^sub>1" and + "\ ; \ ; \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ v' \ \\<^sub>2" and "b= b_of \\<^sub>2" + "\ ; \ ; \ \ \\<^sub>1 \ (\ z0 : b\<^sub>1 | c0 \)" and "atom z0 \ (x,v)" and "atom z3' \ (x,v,v', \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) )" + shows \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : b | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \\ +proof - + have "\ ; \ ; \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ v' \ \ z3' : b_of \\<^sub>2 | C_eq (CE_val (V_var z3')) (CE_val v') \" + proof(rule infer_v_form4) + show \ \ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \\<^sub>2\ using assms by metis + show \atom z3' \ (v', \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)\ using assms fresh_prodN by metis + show \b_of \\<^sub>2 = b_of \\<^sub>2\ by auto + qed + hence \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : b_of \\<^sub>2 | CE_val (V_var z3') == CE_val v' \[x::=v]\<^sub>\\<^sub>v\ + using subst_infer_v assms by metis + thus ?thesis using type_veq_subst fresh_prodN assms by metis +qed + +section \Expressions\ + +text \ + For operator, fst and snd cases, we use elimination to get one or more values, apply the substitution lemma for values. The types always have +the same form and are equal under substitution. + For function application, the subst value is a subtype of the value which is a subtype of the argument. The return of the function is the same + under substitution. +\ + +text \ Observe a similar pattern for each case \ + +lemma subst_infer_e: + fixes v::v and e::e and \'::\ + assumes + "\ ; \ ; \ ; G ; \ \ e \ \\<^sub>2" and "G = (\'@((x,b\<^sub>1,subst_cv c0 z0 (V_var x))#\<^sub>\\))" + "\ ; \ ; \ \ v \ \\<^sub>1" and + "\; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" and "atom z0 \ (x,v)" + shows "\ ; \ ; \ ; ((\'[x::=v]\<^sub>\\<^sub>v)@\) ; (\[x::=v]\<^sub>\\<^sub>v) \ (subst_ev e x v ) \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct avoiding: x v rule: infer_e.strong_induct) + case (infer_e_valI \ \ \'' \ \ v' \) + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_val (v'[x::=v]\<^sub>v\<^sub>v)) \ \[x::=v]\<^sub>\\<^sub>v" + proof + show "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v" using wfD_subst infer_e_valI subtype_eq_base2 + by (metis b_of.simps infer_v_v_wf subst_g_inside_simple wfD_wf wf_subst(11)) + show "\\\<^sub>w\<^sub>f \" using infer_e_valI by auto + show "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" using subst_infer_v infer_e_valI using wfD_subst infer_e_valI subtype_eq_base2 + by metis + qed + thus ?case using subst_ev.simps by simp +next + case (infer_e_plusI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + + hence z3f: "atom z3 \ CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_op Plus v1 v2, CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v , CE_op Plus [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e,\'[x::=v]\<^sub>\\<^sub>v @ \ )" + using obtain_fresh by metis + hence **:"(\ z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_plusI fresh_Pair z3f by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z1 : B_int | c1 \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + obtain z2' b2' c2' where z2:"atom z2' \ (x,v) \ \ z2 : B_int | c2 \ = \ z2' : b2' | c2' \" using obtain_fresh_z by metis + + have bb:"b1' = B_int \ b2' = B_int" using z1 z2 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_op Plus (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ + using infer_e_plusI wfD_subst subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_plusI by blast + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_plusI z1 bb by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_plusI z2 bb by metis + show \atom z3' \ AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_prod6 * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + by(subst subst_tv.simps,auto simp add: * ) + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_leqI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + + hence z3f: "atom z3 \ CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_op LEq v1 v2, CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op LEq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\'[x::=v]\<^sub>\\<^sub>v @ \ )" + using obtain_fresh by metis + hence **:"(\ z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_leqI fresh_Pair z3f by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z1 : B_int | c1 \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + obtain z2' b2' c2' where z2:"atom z2' \ (x,v) \ \ z2 : B_int | c2 \ = \ z2' : b2' | c2' \" using obtain_fresh_z by metis + + have bb:"b1' = B_int \ b2' = B_int" using z1 z2 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_op LEq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_leqI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_leqI(2) by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_leqI z1 bb by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_leqI z2 bb by metis + show \atom z3' \ AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_eqI \ \ \'' \ \ v1 z1 bb c1 v2 z2 c2 z3) + + hence z3f: "atom z3 \ CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_op Eq v1 v2, CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op Eq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\'[x::=v]\<^sub>\\<^sub>v @ \ )" + using obtain_fresh by metis + hence **:"(\ z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_eqI fresh_Pair z3f by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z1 : bb | c1 \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + obtain z2' b2' c2' where z2:"atom z2' \ (x,v) \ \ z2 : bb | c2 \ = \ z2' : b2' | c2' \" using obtain_fresh_z by metis + + have bb:"b1' = bb \ b2' = bb" using z1 z2 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_op Eq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_eqI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_eqI(2) by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1' : bb | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_eqI z1 bb by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2' : bb | c2'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_eqI z2 bb by metis + show \atom z3' \ AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + show "bb \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed + moreover have "\ z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_appI \ \ \'' \ \ f x' b c \' s' v' \) + + hence "x \ x'" using \atom x' \ \''\ using wfG_inside_x_neq wfX_wfY by metis + + show ?case proof(subst subst_ev.simps,rule) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using infer_e_appI wfD_subst subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_appI by metis + show \Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \' s'))) = lookup_fun \ f\ using infer_e_appI by metis + + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b | c \[x::=v]\<^sub>\\<^sub>v\ proof(rule subst_infer_check_v ) + show "\ ; \ ; \ \ v \ \\<^sub>1" using infer_e_appI by metis + show "\ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \ x' : b | c \" using infer_e_appI by metis + show "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" using infer_e_appI by metis + show "atom z0 \ (x, v)" using infer_e_appI by metis + qed + moreover have "atom x \ c" using wfPhi_f_simple_supp_c infer_e_appI fresh_def \x\x'\ + atom_eq_iff empty_iff infer_e_appI.hyps insert_iff subset_singletonD by metis + + moreover hence "atom x \ \ x' : b | c \" using \.fresh supp_b_empty fresh_def by blast + ultimately show \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b | c \\ using forget_subst_tv by metis + + have *: "atom x' \ (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appI check_v_wf by blast + + show \atom x' \ (\, \, \, \'[x::=v]\<^sub>\\<^sub>v @ \, \[x::=v]\<^sub>\\<^sub>v, v'[x::=v]\<^sub>v\<^sub>v, \[x::=v]\<^sub>\\<^sub>v)\ + apply(unfold fresh_prodN, intro conjI) + apply (fresh_subst_mth_aux add: infer_e_appI fresh_subst_gv wfD_wf subst_g_inside) + using infer_e_appI fresh_subst_gv wfD_wf subst_g_inside apply metis + using infer_e_appI fresh_subst_dv_if apply metis + done + + have "supp \' \ { atom x' } \ supp \" using infer_e_appI wfT_supp wfPhi_f_simple_wfT + by (meson infer_e_appI.hyps(2) le_supI1 wfPhi_f_simple_supp_t) + hence "atom x \ \'" using \x\x'\ fresh_def supp_at_base x_not_in_b_set by fastforce + thus \\'[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \[x::=v]\<^sub>\\<^sub>v\ using subst_tv_commute infer_e_appI subst_defs by metis + qed +next + case (infer_e_appPI \ \ \'' \ \ b' f bv x' b c \' s' v' \) + + hence "x \ x'" using \atom x' \ \''\ using wfG_inside_x_neq wfX_wfY by metis + + show ?case proof(subst subst_ev.simps,rule) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using infer_e_appPI wfD_subst subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_appPI(4) by auto + show "\ ; \ \\<^sub>w\<^sub>f b'" using infer_e_appPI(5) by auto + show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b c \' s'))) = lookup_fun \ f" using infer_e_appPI(6) by auto + + show "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \" proof - + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \[x::=v]\<^sub>\\<^sub>v\ proof(rule subst_infer_check_v ) + show "\ ; \ ; \ \ v \ \\<^sub>1" using infer_e_appPI by metis + show "\ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \ x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" using infer_e_appPI subst_defs by metis + show "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" using infer_e_appPI by metis + show "atom z0 \ (x, v)" using infer_e_appPI by metis + qed + moreover have "atom x \ c" proof - + have "supp c \ {atom x', atom bv}" using wfPhi_f_poly_supp_c[OF infer_e_appPI(6)] infer_e_appPI by metis + thus ?thesis unfolding fresh_def using \x\x'\ atom_eq_iff by auto + qed + moreover hence "atom x \ \ x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" using \.fresh supp_b_empty fresh_def subst_b_fresh_x + by (metis subst_b_c_def) + ultimately show ?thesis using forget_subst_tv subst_defs by metis + qed + have "supp \' \ { atom x', atom bv }" using wfPhi_f_poly_supp_t infer_e_appPI by metis + hence "atom x \ \'" using fresh_def \x\x'\ by force + hence *:"atom x \ \'[bv::=b']\<^sub>\\<^sub>b" using subst_b_fresh_x subst_b_\_def by metis + have "atom x' \ (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appPI check_v_wf by blast + thus "atom x' \ \'[x::=v]\<^sub>\\<^sub>v @ \" using infer_e_appPI fresh_subst_gv wfD_wf subst_g_inside fresh_Pair by metis + show "\'[bv::=b']\<^sub>b[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \[x::=v]\<^sub>\\<^sub>v" using infer_e_appPI subst_tv_commute[OF * ] subst_defs by metis + show "atom bv \ (\, \, \, \'[x::=v]\<^sub>\\<^sub>v @ \, \[x::=v]\<^sub>\\<^sub>v, b', v'[x::=v]\<^sub>v\<^sub>v, \[x::=v]\<^sub>\\<^sub>v)" + by (fresh_mth add: infer_e_appPI fresh_subst_gv_inside) + qed +next + case (infer_e_fstI \ \ \'' \ \ v' z' b1 b2 c z) + + hence zf: "atom z \ CE_fst [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_fst v', CE_fst [v']\<^sup>c\<^sup>e , AE_fst v'[x::=v]\<^sub>v\<^sub>v ,\'[x::=v]\<^sub>\\<^sub>v @ \ )" using obtain_fresh by auto + hence **:"(\ z : b1 | CE_val (V_var z) == CE_fst [v']\<^sup>c\<^sup>e \) = \ z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \" + using type_e_eq infer_e_fstI(4) fresh_Pair zf by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z' : B_pair b1 b2 | c \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + + have bb:"b1' = B_pair b1 b2" using z1 \.eq_iff by metis + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_fst v'[x::=v]\<^sub>v\<^sub>v) \ \ z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_fstI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_fstI by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_fstI z1 bb by metis + + show \atom z3' \ AE_fst v'[x::=v]\<^sub>v\<^sub>v \ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \ = \ z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_sndI \ \ \'' \ \ v' z' b1 b2 c z) + hence zf: "atom z \ CE_snd [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_snd v', CE_snd [v']\<^sup>c\<^sup>e , AE_snd v'[x::=v]\<^sub>v\<^sub>v ,\'[x::=v]\<^sub>\\<^sub>v @ \ ,v', \'')" using obtain_fresh by auto + hence **:"(\ z : b2 | CE_val (V_var z) == CE_snd [v']\<^sup>c\<^sup>e \) = \ z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \" + using type_e_eq infer_e_sndI(4) fresh_Pair zf by metis + + obtain z1' b2' c1' where z1:"atom z1' \ (x,v) \ \ z' : B_pair b1 b2 | c \ = \ z1' : b2' | c1' \" using obtain_fresh_z by metis + + have bb:"b2' = B_pair b1 b2" using z1 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_snd (v'[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_sndI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_sndI by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_sndI z1 bb by metis + + show \atom z3' \ AE_snd v'[x::=v]\<^sub>v\<^sub>v \ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + by(subst subst_tv.simps, auto simp add: fresh_prodN *) + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_lenI \ \ \'' \ \ v' z' c z) + hence zf: "atom z \ CE_len [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + obtain z3'::x where *:"atom z3' \ (x,v,AE_len v', CE_len [v']\<^sup>c\<^sup>e , AE_len v'[x::=v]\<^sub>v\<^sub>v ,\'[x::=v]\<^sub>\\<^sub>v @ \ , \'',v')" using obtain_fresh by auto + hence **:"(\ z : B_int | CE_val (V_var z) == CE_len [v']\<^sup>c\<^sup>e \) = \ z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \" + using type_e_eq infer_e_lenI fresh_Pair zf by metis + + have ***: "\ ; \ ; \'' \ v' \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \" + using infer_e_lenI infer_v_form3[OF infer_e_lenI(3), of z3'] b_of.simps * fresh_Pair by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_len (v'[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_lenI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_lenI by metis + + have \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \[x::=v]\<^sub>\\<^sub>v\ + proof(rule subst_infer_v) + show \ \ ; \ ; \ \ v \ \\<^sub>1\ using infer_e_lenI by metis + show \ \ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \ z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v' ]\<^sup>c\<^sup>e \\ using *** infer_e_lenI by metis + show "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" using infer_e_lenI by metis + show "atom z0 \ (x, v)" using infer_e_lenI by metis + qed + + thus \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \\ + using subst_tv.simps subst_ev.simps fresh_Pair * fresh_prodN subst_vv.simps by auto + + show \atom z3' \ AE_len v'[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_Pair * by metis + qed + + moreover have "\ z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_mvarI \ \ \'' \ \ u \) + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_mvar u) \ \[x::=v]\<^sub>\\<^sub>v" + proof + show \ \ ; \\\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \\ proof - + have "wfV \ \ \ v (b_of \\<^sub>1)" using infer_v_wf infer_e_mvarI by auto + moreover have "b_of \\<^sub>1 = b\<^sub>1" using subtype_eq_base2 infer_e_mvarI b_of.simps by simp + ultimately show ?thesis using wf_subst(3)[OF infer_e_mvarI(1), of \' x b\<^sub>1 "c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v" \ v] infer_e_mvarI subst_g_inside by metis + qed + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_mvarI by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_mvarI subtype_eq_base2 b_of.simps by metis + show \(u, \[x::=v]\<^sub>\\<^sub>v) \ setD \[x::=v]\<^sub>\\<^sub>v\ using infer_e_mvarI subst_dv_member by metis + qed + moreover have " (AE_mvar u) = (AE_mvar u)[x::=v]\<^sub>e\<^sub>v" using subst_ev.simps by auto + ultimately show ?case by auto + +next + case (infer_e_concatI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + + hence zf: "atom z3 \ CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,v1,v2,AE_concat v1 v2, CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v) ,\'[x::=v]\<^sub>\\<^sub>v @ \ , \'',v1 , v2)" using obtain_fresh by auto + + hence **:"(\ z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_concatI fresh_Pair zf by metis + have zfx: "atom x \ z3'" using fresh_at_base fresh_prodN * by auto + + have v1: "\ ; \ ; \'' \ v1 \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v1 \" + using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis + have v2: "\ ; \ ; \'' \ v2 \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v2 \" + using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_concatI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ by(simp add: infer_e_concatI) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val (v1[x::=v]\<^sub>v\<^sub>v) \\ + using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val (v2[x::=v]\<^sub>v\<^sub>v) \\ + using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis + show \atom z3' \ AE_concat v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_Pair * by metis + qed + + moreover have "\ z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + + ultimately show ?case using subst_ev.simps ** * by metis +next + case (infer_e_splitI \ \ \'' \ \ v1 z1 c1 v2 z2 z3) + hence *:"atom z3 \ (x,v)" using fresh_Pair by auto + have \x \ z3 \ using infer_e_splitI by force + have "\ ; \ ; \ ; (\'[x::=v]\<^sub>\\<^sub>v @ \) ; \[x::=v]\<^sub>\\<^sub>v \ (AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v) \ + \ z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND + [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_splitI subtype_eq_base2 b_of.simps by metis + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_splitI by auto + have \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1 : B_bitvec | c1 \[x::=v]\<^sub>\\<^sub>v\ + using subst_infer_v infer_e_splitI by metis + thus \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1 : B_bitvec | c1[x::=v]\<^sub>c\<^sub>v \\ + using infer_e_splitI subst_tv.simps fresh_Pair by metis + have \x \ z2 \ using infer_e_splitI by force + have "(\ z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \) = + (\ z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) + AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \[x::=v]\<^sub>\\<^sub>v)" + unfolding subst_cv.simps subst_cev.simps subst_vv.simps using \x \ z2\ infer_e_splitI fresh_Pair by simp + thus \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ + \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e + AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \\ + using infer_e_splitI subst_infer_check_v fresh_Pair by metis + + show \atom z1 \ AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using infer_e_splitI fresh_subst_vv_if by auto + show \atom z2 \ AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using infer_e_splitI fresh_subst_vv_if by auto + show \atom z3 \ AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using infer_e_splitI fresh_subst_vv_if by auto + + show \atom z3 \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_subst_gv_inside infer_e_splitI by metis + show \atom z2 \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_subst_gv_inside infer_e_splitI by metis + show \atom z1 \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_subst_gv_inside infer_e_splitI by metis + qed + thus ?case apply (subst subst_tv.simps) + using infer_e_splitI fresh_Pair apply metis + unfolding subst_tv.simps subst_ev.simps subst_cv.simps subst_cev.simps subst_vv.simps * + using \x \ z3\ by simp +qed + +lemma infer_e_uniqueness: + assumes "\ ; \ ; \ ; GNil ; \ \ e\<^sub>1 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ e\<^sub>1 \ \\<^sub>2" + shows "\\<^sub>1 = \\<^sub>2" + using assms proof(nominal_induct rule:e.strong_induct) + case (AE_val x) + then show ?case using infer_e_elims(7)[OF AE_val(1)] infer_e_elims(7)[OF AE_val(2)] infer_v_uniqueness by metis +next + case (AE_app f v) + + obtain x1 b1 c1 s1' \1' where t1: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f \ \\<^sub>1 = \1'[x1::=v]\<^sub>\\<^sub>v" using infer_e_app2E[OF AE_app(1)] by metis + moreover obtain x2 b2 c2 s2' \2' where t2: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \2' s2'))) = lookup_fun \ f \ \\<^sub>2 = \2'[x2::=v]\<^sub>\\<^sub>v" using infer_e_app2E[OF AE_app(2)] by metis + + have "\1'[x1::=v]\<^sub>\\<^sub>v = \2'[x2::=v]\<^sub>\\<^sub>v" using t1 and t2 fun_ret_unique by metis + thus ?thesis using t1 t2 by auto +next + case (AE_appP f b v) + obtain bv1 x1 b1 c1 s1' \1' where t1: "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f \ \\<^sub>1 = \1'[bv1::=b]\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v" using infer_e_appP2E[OF AE_appP(1)] by metis + moreover obtain bv2 x2 b2 c2 s2' \2' where t2: "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2'))) = lookup_fun \ f \ \\<^sub>2 = \2'[bv2::=b]\<^sub>\\<^sub>b[x2::=v]\<^sub>\\<^sub>v" using infer_e_appP2E[OF AE_appP(2)] by metis + + have "\1'[bv1::=b]\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v = \2'[bv2::=b]\<^sub>\\<^sub>b[x2::=v]\<^sub>\\<^sub>v" using t1 and t2 fun_poly_ret_unique by metis + thus ?thesis using t1 t2 by auto +next + case (AE_op opp v1 v2) + show ?case proof(rule opp.exhaust) + assume "opp = plus" + hence "\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>2" using AE_op by auto + thus ?thesis using infer_e_elims(11)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>1\ ] infer_e_elims(11)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>2\ ] + by force + next + assume "opp = leq" + hence "opp = LEq" using opp.strong_exhaust by auto + hence "\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>2" using AE_op by auto + thus ?thesis using infer_e_elims(12)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>1\ ] infer_e_elims(12)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>2\ ] + by force + next + assume "opp = eq" + hence "opp = Eq" using opp.strong_exhaust by auto + hence "\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>2" using AE_op by auto + thus ?thesis using infer_e_elims(25)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>1\ ] infer_e_elims(25)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>2\ ] + by force + qed +next + case (AE_concat v1 v2) + + obtain z3::x where t1:"\\<^sub>1 = \ z3 : B_bitvec | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ \ atom z3 \ v1 \ atom z3 \ v2 " using infer_e_elims(18)[OF AE_concat(1)] by metis + obtain z3'::x where t2:"\\<^sub>2 = \ z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ \ atom z3' \ v1 \ atom z3' \ v2" using infer_e_elims(18)[OF AE_concat(2)] by metis + + thus ?case using t1 t2 type_e_eq ce.fresh by metis + +next + case (AE_fst v) + + obtain z1 and b1 where "\\<^sub>1 = \ z1 : b1 | CE_val (V_var z1) == (CE_fst [v]\<^sup>c\<^sup>e) \" using infer_v_form AE_fst by auto + + obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where + f1: "\\<^sub>2 = \ xx : bb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxa : B_pair bb bba | cc \ \ atom xx \ v" + using infer_e_elims(8)[OF AE_fst(2)] by metis + obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where + f2: "\\<^sub>1 = \ xxb : bbb | CE_val (V_var xxb) == CE_fst [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxc : B_pair bbb bbc | cca \ \ atom xxb \ v" + using infer_e_elims(8)[OF AE_fst(1)] by metis + then have "B_pair bb bba = B_pair bbb bbc" + using f1 by (metis (no_types) b_of.simps infer_v_uniqueness) + then have "\ xx : bbb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \ = \\<^sub>2" + using f1 by auto + then show ?thesis + using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple) +next + case (AE_snd v) + obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where + f1: "\\<^sub>2 = \ xx : bba | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxa : B_pair bb bba | cc \ \ atom xx \ v" + using infer_e_elims(9)[OF AE_snd(2)] by metis + obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where + f2: "\\<^sub>1 = \ xxb : bbc | CE_val (V_var xxb) == CE_snd [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxc : B_pair bbb bbc | cca \ \ atom xxb \ v" + using infer_e_elims(9)[OF AE_snd(1)] by metis + then have "B_pair bb bba = B_pair bbb bbc" + using f1 by (metis (no_types) b_of.simps infer_v_uniqueness) + then have "\ xx : bbc | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \ = \\<^sub>2" + using f1 by auto + then show ?thesis + using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple) +next + case (AE_mvar x) + then show ?case using infer_e_elims(10)[OF AE_mvar(1)] infer_e_elims(10)[OF AE_mvar(2)] wfD_unique by metis +next + case (AE_len x) + then show ?case using infer_e_elims(16)[OF AE_len(1)] infer_e_elims(16)[OF AE_len(2)] by force +next + case (AE_split x1a x2) + then show ?case using infer_e_elims(22)[OF AE_split(1)] infer_e_elims(22)[OF AE_split(2)] by force +qed + +section \Statements\ + +lemma subst_infer_check_v1: + fixes v::v and v'::v and \::\ + assumes "\ = \\<^sub>1@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>2)" and + "\ ; \ ; \\<^sub>2 \ v \ \\<^sub>1" and + "\ ; \ ; \ \ v' \ \\<^sub>2" and + "\ ; \ ; \\<^sub>2 \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" and "atom z0 \ (x,v)" + shows "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using subst_g_inside check_v_wf assms subst_infer_check_v by metis + +lemma infer_v_c_valid: + assumes " \ ; \ ; \ \ v \ \" and "\ ; \ ; \ \ \ \ \ z : b | c \" + shows \\ ; \ ; \ \ c[z::=v]\<^sub>c\<^sub>v \ +proof - + obtain z1 and b1 and c1 where *:"\ = \ z1 : b1 | c1 \ \ atom z1 \ (c,v,\)" using obtain_fresh_z by metis + then have "b1 = b" using assms subtype_eq_base by metis + moreover then have "\ ; \ ; \ \ v \ \ z1 : b | c1 \" using assms * by auto + + moreover have "\ ; \ ; (z1, b, c1) #\<^sub>\ \ \ c[z::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v " proof - + have "\ ; \ ; (z1, b, c1[z1::=[ z1 ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c[z::=[ z1 ]\<^sup>v]\<^sub>v " + using subtype_valid[OF assms(2), of z1 z1 b c1 z c ] * fresh_prodN \b1 = b\ by metis + moreover have "c1[z1::=[ z1 ]\<^sup>v]\<^sub>v = c1" using subst_v_v_def by simp + ultimately show ?thesis using subst_v_c_def by metis + qed + ultimately show ?thesis using * fresh_prodN subst_valid_simple by metis +qed + +text \ Substitution Lemma for Statements \ + +lemma subst_infer_check_s: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and + \\<^sub>1::\ and \\<^sub>2::\ and css::branch_list + assumes "\ ; \ ; \\<^sub>1 \ v \ \" and "\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" and + "atom z \ (x, v)" + shows "\ ; \ ; \ ; \; \ \ s \ \' \ + \ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1)) \ + \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" + and + "\ ; \ ; \ ; \; \; tid ; cons ; const ; v' \ cs \ \' \ + \ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1)) \ + \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v; + tid ; cons ; const ; v'[x::=v]\<^sub>v\<^sub>v \ cs[x::=v]\<^sub>s\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" + and + "\ ; \ ; \ ; \; \; tid ; dclist ; v' \ css \ \' \ + \ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1)) \ + \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v; tid ; dclist ; v'[x::=v]\<^sub>v\<^sub>v \ + subst_branchlv css x v \ \'[x::=v]\<^sub>\\<^sub>v" + + using assms proof(nominal_induct \' and \' and \' avoiding: x v arbitrary: \\<^sub>2 and \\<^sub>2 and \\<^sub>2 + rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_valI \ \ \ \ \ v' \' \'') + + have sg: " \[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using check_valI by subst_mth + have "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ (AS_val (v'[x::=v]\<^sub>v\<^sub>v)) \ \''[x::=v]\<^sub>\\<^sub>v" proof + have "\ ; \ ; \\<^sub>1\\<^sub>w\<^sub>f v : b " using infer_v_v_wf subtype_eq_base2 b_of.simps check_valI by metis + thus \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v\\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v\ using wf_subst(15) check_valI by auto + show \ \\\<^sub>w\<^sub>f \ \ using check_valI by auto + show \ \ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v\ proof(subst sg, rule subst_infer_v) + show "\ ; \ ; \\<^sub>1 \ v \ \" using check_valI by auto + show "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ v' \ \'" using check_valI by metis + show "\ ; \ ; \\<^sub>1 \ \ \ \ z: b | c \" using check_valI by auto + show "atom z \ (x, v)" using check_valI by auto + qed + show \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v \ \''[x::=v]\<^sub>\\<^sub>v\ using subst_subtype_tau check_valI sg by metis + qed + + thus ?case using Typing.check_valI subst_sv.simps sg by auto +next + case (check_letI xa \ \ \ \ \ ea \a za sa ba ca) + have *:"(AS_let xa ea sa)[x::=v]\<^sub>s\<^sub>v=(AS_let xa (ea[x::=v]\<^sub>e\<^sub>v) sa[x::=v]\<^sub>s\<^sub>v)" + using subst_sv.simps \ atom xa \ x\ \ atom xa \ v\ by auto + show ?case unfolding * proof + + show "atom xa \ (\,\,\,\[x::=v]\<^sub>\\<^sub>v,\[x::=v]\<^sub>\\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v,\a[x::=v]\<^sub>\\<^sub>v)" + by(subst_tuple_mth add: check_letI) + + show "atom za \ (xa,\,\,\,\[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v, + \a[x::=v]\<^sub>\\<^sub>v,sa[x::=v]\<^sub>s\<^sub>v)" + by(subst_tuple_mth add: check_letI) + + show "\; \; \; \[x::=v]\<^sub>\\<^sub>v; \[x::=v]\<^sub>\\<^sub>v \ + ea[x::=v]\<^sub>e\<^sub>v \ \ za : ba | ca[x::=v]\<^sub>c\<^sub>v \" + proof - + have "\; \; \; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1; \[x::=v]\<^sub>\\<^sub>v \ + ea[x::=v]\<^sub>e\<^sub>v \ \ za : ba | ca \[x::=v]\<^sub>\\<^sub>v" + using check_letI subst_infer_e by metis + thus ?thesis using check_letI subst_tv.simps + by (metis fresh_prod2I infer_e_wf subst_g_inside_simple) + qed + + show "\; \; \; (xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v; + \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + proof - + have "\; \; \; ((xa, ba, ca[za::=V_var xa]\<^sub>v) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v ; + \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + apply(rule check_letI(23)[of "(xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>2"]) + by(metis check_letI append_g.simps subst_defs)+ + + moreover have "(xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v = + ((xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v" + using subst_cv_commute subst_gv.simps check_letI + by (metis ms_fresh_all(39) ms_fresh_all(49) subst_cv_commute_full) + ultimately show ?thesis + using subst_defs by auto + qed + qed +next + case (check_assertI xa \ \ \ \ \ ca \ s) + show ?case unfolding subst_sv.simps proof + show \atom xa \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, ca[x::=v]\<^sub>c\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, s[x::=v]\<^sub>s\<^sub>v)\ + by(subst_tuple_mth add: check_assertI) + have "xa \ x" using check_assertI by fastforce + thus \ \ ; \ ; \ ; (xa, B_bool, ca[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v\ + using check_assertI(12)[of "(xa, B_bool, c) #\<^sub>\ \\<^sub>2" x v] check_assertI subst_gv.simps append_g.simps by metis + have \\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ ca[x::=v]\<^sub>c\<^sub>v \ proof(rule subst_valid ) + show \\ ; \ ; \\<^sub>1 \ c[z::=v]\<^sub>c\<^sub>v \ using infer_v_c_valid check_assertI by metis + show \ \ ; \ ; \\<^sub>1 \\<^sub>w\<^sub>f v : b \ using check_assertI infer_v_wf b_of.simps subtype_eq_base + by (metis subtype_eq_base2) + show \ \ ; \ \\<^sub>w\<^sub>f \\<^sub>1 \ using check_assertI infer_v_wf by metis + have " \ ; \ \\<^sub>w\<^sub>f \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1" using check_assertI wfX_wfY by metis + thus \atom x \ \\<^sub>1\ using check_assertI wfG_suffix wfG_elims by metis + + moreover have "\ ; \ ; \\<^sub>1 \\<^sub>w\<^sub>f \ z : b | c \" using subtype_wfT check_assertI by metis + moreover have "x \ z" using fresh_Pair check_assertI fresh_x_neq by metis + ultimately show \atom x \ c\ using check_assertI wfT_fresh_c by metis + + show \ \ ; \ \\<^sub>w\<^sub>f \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ using check_assertI wfX_wfY by metis + show \\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ ca \ using check_assertI by auto + qed + thus \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ ca[x::=v]\<^sub>c\<^sub>v \ using check_assertI + proof - + show ?thesis + by (metis (no_types) \\ = \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1\ \\ ; \ ; \ \ ca\ \\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ ca[x::=v]\<^sub>c\<^sub>v\ subst_g_inside valid_g_wf) (* 93 ms *) + qed + + have "\ ; \ ; \\<^sub>1 \\<^sub>w\<^sub>f v : b" using infer_v_wf b_of.simps check_assertI + by (metis subtype_eq_base2) + thus \ \ ; \ ; \[x::=v]\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wf_subst2(6) check_assertI by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist vv cs \ css) + show ?case unfolding * using subst_sv.simps check_branch_list_consI by simp +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + show ?case unfolding * using subst_sv.simps check_branch_list_finalI by simp +next + case (check_branch_s_branchI \ \ \ \ \ const xa \ tid cons va sa) + hence *:"(AS_branch cons xa sa)[x::=v]\<^sub>s\<^sub>v = (AS_branch cons xa sa[x::=v]\<^sub>s\<^sub>v)" using subst_branchv.simps fresh_Pair by metis + show ?case unfolding * proof + + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v\\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v " + using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis + + show "\\<^sub>w\<^sub>f \ " using check_branch_s_branchI by metis + + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v " + using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis + + show wft:"\ ; {||} ; GNil\\<^sub>w\<^sub>f const " using check_branch_s_branchI by metis + + show "atom xa \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, tid, cons, const,va[x::=v]\<^sub>v\<^sub>v, \[x::=v]\<^sub>\\<^sub>v)" + apply(unfold fresh_prodN, (simp add: check_branch_s_branchI )+) + apply(rule,metis fresh_z_subst_g check_branch_s_branchI fresh_Pair ) + by(metis fresh_subst_dv check_branch_s_branchI fresh_Pair ) + + have "\ ; \ ; \ ; ((xa, b_of const, CE_val va == CE_val(V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" + using check_branch_s_branchI by (metis append_g.simps(2)) + + moreover have "(xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of (const) xa) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v = + ((xa, b_of const , CE_val va == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v" + proof - + have *:"xa \ x" using check_branch_s_branchI fresh_at_base by metis + have "atom x \ const" using wfT_nil_supp[OF wft] fresh_def by auto + hence "atom x \ (const,xa)" using fresh_at_base wfT_nil_supp[OF wft] fresh_Pair fresh_def * by auto + moreover hence "(c_of (const) xa)[x::=v]\<^sub>c\<^sub>v = c_of (const) xa" + using c_of_fresh[of x const xa] forget_subst_cv wfT_nil_supp wft by metis + moreover hence "(V_cons tid cons (V_var xa))[x::=v]\<^sub>v\<^sub>v = (V_cons tid cons (V_var xa))" using check_branch_s_branchI subst_vv.simps * by metis + ultimately show ?thesis using subst_gv.simps check_branch_s_branchI subst_cv.simps subst_cev.simps * by presburger + qed + + ultimately show "\ ; \ ; \ ; (xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" + by metis + qed + +next + case (check_let2I xa \ \ \ G \ t s1 \a s2 ) + hence *:"(AS_let2 xa t s1 s2)[x::=v]\<^sub>s\<^sub>v = (AS_let2 xa t[x::=v]\<^sub>\\<^sub>v (s1[x::=v]\<^sub>s\<^sub>v) s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps fresh_Pair by metis + have "xa \ x" using check_let2I fresh_at_base by metis + show ?case unfolding * proof + show "atom xa \ (\, \, \, G[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, t[x::=v]\<^sub>\\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, \a[x::=v]\<^sub>\\<^sub>v)" + by(subst_tuple_mth add: check_let2I) + show "\ ; \ ; \ ; G[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s1[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v" using check_let2I by metis + + have "\ ; \ ; \ ; ((xa, b_of t, c_of t xa) #\<^sub>\ G)[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + proof(rule check_let2I(14)) + show \(xa, b_of t, c_of t xa) #\<^sub>\ G = (((xa, b_of t, c_of t xa)#\<^sub>\ \\<^sub>2)) @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1\ + using check_let2I append_g.simps by metis + show \ \ ; \ ; \\<^sub>1 \ v \ \\ using check_let2I by metis + show \\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \\ using check_let2I by metis + show \atom z \ (x, v)\ using check_let2I by metis + qed + moreover have "c_of t[x::=v]\<^sub>\\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v" using subst_v_c_of fresh_Pair check_let2I by metis + moreover have "b_of t[x::=v]\<^sub>\\<^sub>v = b_of t" using b_of.simps subst_tv.simps b_of_subst by metis + ultimately show " \ ; \ ; \ ; (xa, b_of t[x::=v]\<^sub>\\<^sub>v, c_of t[x::=v]\<^sub>\\<^sub>v xa) #\<^sub>\ G[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + using check_let2I(14) subst_gv.simps \xa \ x\ b_of.simps by metis + qed + +next + + case (check_varI u \ \ \ \ \ \' va \'' s) + have **: "\[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside check_s_wf check_varI by meson + + have "\ ; \ ;\ ; subst_gv \ x v ; \[x::=v]\<^sub>\\<^sub>v \ AS_var u \'[x::=v]\<^sub>\\<^sub>v (va[x::=v]\<^sub>v\<^sub>v) (subst_sv s x v) \ \''[x::=v]\<^sub>\\<^sub>v" + proof(rule Typing.check_varI) + show "atom u \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, \'[x::=v]\<^sub>\\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, \''[x::=v]\<^sub>\\<^sub>v)" + by(subst_tuple_mth add: check_varI) + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ va[x::=v]\<^sub>v\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" using check_varI subst_infer_check_v ** by metis + show "\ ; \ ; \ ; subst_gv \ x v ; (u, \'[x::=v]\<^sub>\\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \''[x::=v]\<^sub>\\<^sub>v" proof - + have "wfD \ \ (\\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1) ((u,\')#\<^sub>\ \)" using check_varI check_s_wf by meson + moreover have "wfV \ \ \\<^sub>1 v (b_of \)" using infer_v_wf check_varI(6) check_varI by metis + have "wfD \ \ (\[x::=v]\<^sub>\\<^sub>v) ((u, \'[x::=v]\<^sub>\\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v)" proof(subst subst_dv.simps(2)[symmetric], subst **, rule wfD_subst) + show "\ ; \ ; \\<^sub>1 \ v \ \" using check_varI by auto + show "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1\\<^sub>w\<^sub>f (u, \') #\<^sub>\ \" using check_varI check_s_wf by simp + show "b_of \ = b" using check_varI subtype_eq_base2 b_of.simps by auto + qed + thus ?thesis using check_varI by auto + qed + qed + moreover have "atom u \ (x,v)" using u_fresh_xv by auto + ultimately show ?case using subst_sv.simps(7) by auto + +next + case (check_assignI P \ \ \ \ u \1 v' z1 \') (* may need to revisit subst in \ as well *) + + have "wfG P \ \" using check_v_wf check_assignI by simp + hence gs: "\\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 = \[x::=v]\<^sub>\\<^sub>v" using subst_g_inside check_assignI by simp + + have "P ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_assign u (v'[x::=v]\<^sub>v\<^sub>v) \ \'[x::=v]\<^sub>\\<^sub>v" + proof(rule Typing.check_assignI) + show "P \\<^sub>w\<^sub>f \ " using check_assignI by auto + show "wfD P \ (\[x::=v]\<^sub>\\<^sub>v) \[x::=v]\<^sub>\\<^sub>v" using wf_subst(15)[OF check_assignI(2)] gs infer_v_v_wf check_assignI b_of.simps subtype_eq_base2 by metis + thus "(u, \1[x::=v]\<^sub>\\<^sub>v) \ setD \[x::=v]\<^sub>\\<^sub>v" using check_assignI subst_dv_member by metis + thus "P ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \1[x::=v]\<^sub>\\<^sub>v" using subst_infer_check_v check_assignI gs by metis + + have "P ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ \ z : B_unit | TRUE \[x::=v]\<^sub>\\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" proof(rule subst_subtype_tau) + show "P ; \ ; \\<^sub>1 \ v \ \" using check_assignI by auto + show "P ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" using check_assignI by meson + show "P ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ \ z : B_unit | TRUE \ \ \'" using check_assignI + by (metis Abs1_eq_iff(3) \.eq_iff c.fresh(1) c.perm_simps(1)) + show "atom z \ (x, v)" using check_assignI by auto + qed + moreover have "\ z : B_unit | TRUE \[x::=v]\<^sub>\\<^sub>v = \ z : B_unit | TRUE \" using subst_tv.simps subst_cv.simps check_assignI by presburger + ultimately show "P ; \ ; \[x::=v]\<^sub>\\<^sub>v \ \ z : B_unit | TRUE \ \ \'[x::=v]\<^sub>\\<^sub>v" using gs by auto + qed + thus ?case using subst_sv.simps(5) by auto + +next + case (check_whileI \ \ \ \ \ s1 z' s2 \') + have " wfG \ \ (\\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1)" using check_whileI check_s_wf by meson + hence **: " \[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside wf check_whileI by auto + have teq: "(\ z : B_unit | TRUE \)[x::=v]\<^sub>\\<^sub>v = (\ z : B_unit | TRUE \)" by(auto simp add: subst_sv.simps check_whileI) + moreover have "(\ z : B_unit | TRUE \) = (\ z' : B_unit | TRUE \)" using type_eq_flip c.fresh flip_fresh_fresh by metis + ultimately have teq2:"(\ z' : B_unit | TRUE \)[x::=v]\<^sub>\\<^sub>v = (\ z' : B_unit | TRUE \)" by metis + + hence "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s1[x::=v]\<^sub>s\<^sub>v \ \ z' : B_bool | TRUE \" using check_whileI subst_sv.simps subst_top_eq by metis + moreover have "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \ z' : B_unit | TRUE \" using check_whileI subst_top_eq by metis + moreover have "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ \ z' : B_unit | TRUE \ \ \'[x::=v]\<^sub>\\<^sub>v" proof - + have "\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ \ z' : B_unit | TRUE \[x::=v]\<^sub>\\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" proof(rule subst_subtype_tau) + show "\ ; \ ; \\<^sub>1 \ v \ \" by(auto simp add: check_whileI) + show "\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" by(auto simp add: check_whileI) + show "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ \ z' : B_unit | TRUE \ \ \'" using check_whileI by metis + show "atom z \ (x, v)" by(auto simp add: check_whileI) + qed + thus ?thesis using teq2 ** by auto + qed + + ultimately have " \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_while s1[x::=v]\<^sub>s\<^sub>v s2[x::=v]\<^sub>s\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" + using Typing.check_whileI by metis + then show ?case using subst_sv.simps by metis +next + case (check_seqI P \ \ \ \ s1 z s2 \ ) + hence "P ; \; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_seq (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v) \ \[x::=v]\<^sub>\\<^sub>v" using Typing.check_seqI subst_top_eq check_seqI by metis + then show ?case using subst_sv.simps by metis +next + case (check_caseI \ \ \ \ \ tid dclist v' cs \ za) + + have wf: "wfG \ \ \" using check_caseI check_v_wf by simp + have **: "\[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside wf check_caseI by auto + + have "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v) \ \[x::=v]\<^sub>\\<^sub>v" proof(rule Typing.check_caseI ) + show "check_branch_list \ \ \ (\[x::=v]\<^sub>\\<^sub>v) \[x::=v]\<^sub>\\<^sub>v tid dclist v'[x::=v]\<^sub>v\<^sub>v (subst_branchlv cs x v ) (\[x::=v]\<^sub>\\<^sub>v)" using check_caseI by auto + show "AF_typedef tid dclist \ set \" using check_caseI by auto + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \ za : B_id tid | TRUE \" proof - + have "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ v' \ \ za : B_id tid | TRUE \" + using check_caseI by argo + hence "\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ v'[x::=v]\<^sub>v\<^sub>v \ (\ za : B_id tid | TRUE \)[x::=v]\<^sub>\\<^sub>v" + using check_caseI subst_infer_check_v[OF check_caseI(7) _ check_caseI(8) check_caseI(9)] by meson + moreover have "(\ za : B_id tid | TRUE \) = ((\ za : B_id tid | TRUE \)[x::=v]\<^sub>\\<^sub>v)" + using subst_cv.simps subst_tv.simps subst_cv_true by fast + ultimately show ?thesis using check_caseI ** by argo + qed + show "wfTh \" using check_caseI by auto + qed + thus ?case using subst_branchlv.simps subst_sv.simps(4) by metis +next + case (check_ifI z' \ \ \ \ \ va s1 s2 \') + show ?case unfolding subst_sv.simps proof + show \atom z' \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, s2[x::=v]\<^sub>s\<^sub>v, \'[x::=v]\<^sub>\\<^sub>v)\ + by(subst_tuple_mth add: check_ifI) + have *:"\ z' : B_bool | TRUE \[x::=v]\<^sub>\\<^sub>v = \ z' : B_bool | TRUE \" using subst_tv.simps check_ifI + by (metis freshers(19) subst_cv.simps(1) type_eq_subst) + have **: "\[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside wf check_ifI check_v_wf by metis + show \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ va[x::=v]\<^sub>v\<^sub>v \ \ z' : B_bool | TRUE \\ + proof(subst *[symmetric], rule subst_infer_check_v1[where \\<^sub>1=\\<^sub>2 and \\<^sub>2=\\<^sub>1]) + show "\ = \\<^sub>2 @ ((x, b ,c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1)" using check_ifI by metis + show \ \ ; \ ; \\<^sub>1 \ v \ \\ using check_ifI by metis + show \\ ; \ ; \ \ va \ \ z' : B_bool | TRUE \\ using check_ifI by metis + show \\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \\ using check_ifI by metis + show \atom z \ (x, v)\ using check_ifI by metis + qed + + have " \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \ = \ z' : b_of \' | [ va ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' z' \[x::=v]\<^sub>\\<^sub>v" + by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of) + + thus \ \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s1[x::=v]\<^sub>s\<^sub>v \ \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \\ + using check_ifI by metis + have " \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \ = \ z' : b_of \' | [ va ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' z' \[x::=v]\<^sub>\\<^sub>v" + by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of) + thus \ \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \\ + using check_ifI by metis + qed +qed + +lemma subst_check_check_s: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \\<^sub>1::\ and \\<^sub>2::\ + assumes "\ ; \ ; \\<^sub>1 \ v \ \ z : b | c \" and "atom z \ (x, v)" + and "check_s \ \ \ \ \ s \'" and "\ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1))" + shows "check_s \ \ \ (subst_gv \ x v) \[x::=v]\<^sub>\\<^sub>v (s[x::=v]\<^sub>s\<^sub>v) (subst_tv \' x v )" +proof - + obtain \ where "\ ; \ ; \\<^sub>1 \ v \ \ \ \ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" using check_v_elims assms by auto + thus ?thesis using subst_infer_check_s assms by metis +qed + +text \ If a statement checks against a type @{text "\"} then it checks against a supertype of @{text "\"} \ +lemma check_s_supertype: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \::\ and \'::\ and css::branch_list + shows "check_s \ \ \ G \ s t1 \ \ ; \ ; G \ t1 \ t2 \ check_s \ \ \ G \ s t2" and + "check_branch_s \ \ \ G \ tid cons const v cs t1 \ \ ; \ ; G \ t1 \ t2 \ check_branch_s \ \ \ G \ tid cons const v cs t2" and + "check_branch_list \ \ \ G \ tid dclist v css t1 \ \ ; \ ; G \ t1 \ t2 \ check_branch_list \ \ \ G \ tid dclist v css t2" +proof(induct arbitrary: t2 and t2 and t2 rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ \ \ \ \ v \' \ ) + hence " \ ; \ ; \ \ \' \ t2" using subtype_trans by meson + then show ?case using subtype_trans Typing.check_valI check_valI by metis + +next + case (check_letI x \ \ \ \ \ e \ z s b c) + show ?case proof(rule Typing.check_letI) + show "atom x \(\, \, \, \, \, e, t2)" using check_letI subtype_fresh_tau fresh_prodN by metis + show "atom z \ (x, \, \, \, \, \, e, t2, s)" using check_letI(2) subtype_fresh_tau[of z \ \ _ _ t2] fresh_prodN check_letI(6) by auto + show "\ ; \ ; \ ; \ ; \ \ e \ \ z : b | c \" using check_letI by meson + + have "wfG \ \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \)" using check_letI check_s_wf subst_defs by metis + moreover have "toSet \ \ toSet ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \)" by auto + ultimately have " \ ; \ ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ \ \ t2" using subtype_weakening[OF check_letI(6)] by auto + thus "\ ; \ ; \ ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \ ; \ \ s \ t2" using check_letI subst_defs by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case using Typing.check_branch_list_consI by auto +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case using Typing.check_branch_list_finalI by auto +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + show ?case proof + have "atom x \ t2" using subtype_fresh_tau[of x \ ] check_branch_s_branchI(5,8) fresh_prodN by metis + thus "atom x \ (\, \, \, \, \, tid, cons, const, v, t2)" using check_branch_s_branchI fresh_prodN by metis + show "wfT \ \ \ t2" using subtype_wf check_branch_s_branchI by meson + show "\ ; \ ; \ ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \ ; \ \ s \ t2" proof - + have "wfG \ \ ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \)" using check_s_wf check_branch_s_branchI by metis + moreover have "toSet \ \ toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \)" by auto + hence "\ ; \ ; ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) \ \ \ t2" + using check_branch_s_branchI subtype_weakening + using calculation by presburger + thus ?thesis using check_branch_s_branchI by presburger + qed + qed(auto simp add: check_branch_s_branchI) +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case proof(rule Typing.check_ifI) + have *:"atom z \ t2" using subtype_fresh_tau[of z \ \ ] check_ifI fresh_prodN by auto + thus \atom z \ (\, \, \, \, \, v, s1, s2, t2)\ using check_ifI fresh_prodN by auto + show \\ ; \ ; \ \ v \ \ z : B_bool | TRUE \\ using check_ifI by auto + show \ \ ; \ ; \ ; \ ; \ \ s1 \ \ z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \\ + using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis + + show \ \ ; \ ; \ ; \ ; \ \ s2 \ \ z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \\ + using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + + show ?case proof + have "atom x \ t2" using subtype_fresh_tau[OF _ _ \\ ; \ ; \ \ \ \ t2\] check_assertI fresh_prodN by simp + thus "atom x \ (\, \, \, \, \, c, t2, s)" using subtype_fresh_tau check_assertI fresh_prodN by simp + have "\ ; \ ; (x, B_bool, c) #\<^sub>\ \ \ \ \ t2" apply(rule subtype_weakening) + using check_assertI apply simp + using toSet.simps apply blast + using check_assertI check_s_wf by simp + thus "\ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \ s \ t2" using check_assertI by simp + show "\ ; \ ; \ \ c " using check_assertI by auto + show "\ ; \ ; \ \\<^sub>w\<^sub>f \ " using check_assertI by auto + qed +next + case (check_let2I x P \ \ G \ t s1 \ s2 ) + have "wfG P \ ((x, b_of t, c_of t x) #\<^sub>\ G)" + using check_let2I check_s_wf by metis + moreover have "toSet G \ toSet ((x, b_of t, c_of t x) #\<^sub>\ G)" by auto + ultimately have *:"P ; \ ; (x, b_of t, c_of t x) #\<^sub>\ G \ \ \ t2" using check_let2I subtype_weakening by metis + show ?case proof(rule Typing.check_let2I) + have "atom x \ t2" using subtype_fresh_tau[of x \ ] check_let2I fresh_prodN by metis + thus "atom x \ (P, \, \, G, \, t, s1, t2)" using check_let2I fresh_prodN by metis + show "P ; \ ; \ ; G ; \ \ s1 \ t" using check_let2I by blast + show "P ; \ ; \ ;(x, b_of t, c_of t x ) #\<^sub>\ G ; \ \ s2 \ t2" using check_let2I * by blast + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof(rule Typing.check_varI) + have "atom u \ t2" using u_fresh_t by auto + thus \atom u \ (\, \, \, \, \, \', v, t2)\ using check_varI fresh_prodN by auto + show \\ ; \ ; \ \ v \ \'\ using check_varI by auto + show \ \ ; \ ; \ ; \ ; (u, \') #\<^sub>\ \ \ s \ t2\ using check_varI by auto + qed +next + case (check_assignI \ u \ P G v z \') + then show ?case using Typing.check_assignI by (meson subtype_trans) +next + case (check_whileI \ G P s1 z s2 \') + then show ?case using Typing.check_whileI by (meson subtype_trans) +next + case (check_seqI \ G P s1 z s2 \) + then show ?case using Typing.check_seqI by blast +next + case (check_caseI \ \ \ tid cs \ v z) + then show ?case using Typing.check_caseI subtype_trans by meson + +qed + +lemma subtype_let: + fixes s'::s and cs::branch_s and css::branch_list and v::v + shows "\ ; \ ; \ ; GNil ; \ \ AS_let x e\<^sub>1 s \ \ \ \ ; \ ; \ ; GNil ; \ \ e\<^sub>1 \ \\<^sub>1 \ + \ ; \ ; \ ; GNil ; \ \ e\<^sub>2 \ \\<^sub>2 \ \ ; \ ; GNil \ \\<^sub>2 \ \\<^sub>1 \ \ ; \ ; \ ; GNil ; \ \ AS_let x e\<^sub>2 s \ \" and + "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" +proof(nominal_induct GNil \ "AS_let x e\<^sub>1 s" \ and \ and \ avoiding: e\<^sub>2 \\<^sub>1 \\<^sub>2 + rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_letI x1 \ \ \ \ \1 z1 s1 b1 c1) + obtain z2 and b2 and c2 where t2:"\\<^sub>2 = \ z2 : b2 | c2 \ \ atom z2 \ (x1, \, \, \, GNil, \, e\<^sub>2, \1, s1) " + using obtain_fresh_z by metis + + obtain z1a and b1a and c1a where t1:"\\<^sub>1 = \ z1a : b1a | c1a \ \ atom z1a \ x1" using infer_e_uniqueness check_letI by metis + hence t3: " \ z1a : b1a | c1a \ = \ z1 : b1 | c1 \ " using infer_e_uniqueness check_letI by metis + + have beq: "b1a = b2 \ b2 = b1" using check_letI subtype_eq_base t1 t2 t3 by metis + + have " \ ; \ ; \ ; GNil ; \ \ AS_let x1 e\<^sub>2 s1 \ \1" proof + show \atom x1 \ (\, \, \, GNil, \, e\<^sub>2, \1)\ using check_letI t2 fresh_prodN by metis + show \atom z2 \ (x1, \, \, \, GNil, \, e\<^sub>2, \1, s1)\ using check_letI t2 using check_letI t2 fresh_prodN by metis + show \\ ; \ ; \ ; GNil ; \ \ e\<^sub>2 \ \ z2 : b2 | c2 \\ using check_letI t2 by metis + + have \ \ ; \ ; \ ; GNil@(x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ s1 \ \1\ + proof(rule ctx_subtype_s) + have "c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v = c1[z1::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v" using subst_v_flip_eq_two subst_v_c_def t3 \.eq_iff by metis + thus \ \ ; \ ; \ ; GNil @ (x1, b2, c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ s1 \ \1\ using check_letI beq append_g.simps subst_defs by metis + show \\ ; \ ; GNil \ \ z2 : b2 | c2 \ \ \ z1a : b2 | c1a \\ using check_letI beq t1 t2 by metis + have "atom x1 \ c2" using t2 check_letI \_fresh_c fresh_prodN by blast + moreover have "atom x1 \ c1a" using t1 check_letI \_fresh_c fresh_prodN by blast + ultimately show \atom x1 \ (z1a, z2, c1a, c2)\ using t1 t2 fresh_prodN fresh_x_neq by metis + qed + + thus \ \ ; \ ; \ ; (x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil ; \ \ s1 \ \1\ using append_g.simps subst_defs by metis + qed + + moreover have "AS_let x1 e\<^sub>2 s1 = AS_let x e\<^sub>2 s" using check_letI s_branch_s_branch_list.eq_iff by metis + + ultimately show ?case by metis + +qed(auto+) + +end \ No newline at end of file diff --git a/thys/MiniSail/MiniSail.thy b/thys/MiniSail/MiniSail.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/MiniSail.thy @@ -0,0 +1,10 @@ +(*<*) +theory MiniSail + imports + "Safety" +begin +(*>*) + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MiniSail/Nominal-Utils.thy b/thys/MiniSail/Nominal-Utils.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Nominal-Utils.thy @@ -0,0 +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)) + +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) + +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)) + +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/Operational.thy b/thys/MiniSail/Operational.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Operational.thy @@ -0,0 +1,171 @@ +(*<*) +theory Operational + imports Typing +begin + (*>*) + +chapter \Operational Semantics\ + +text \ Here we define the operational semantics in terms of a small-step reduction relation. \ + +section \Reduction Rules\ + +text \ The store for mutable variables \ +type_synonym \ = "(u*v) list" + +nominal_function update_d :: "\ \ u \ v \ \" where + "update_d [] _ _ = []" +| "update_d ((u',v')#\) u v = (if u = u' then ((u,v)#\) else ((u',v')# (update_d \ u v)))" + by(auto,simp add: eqvt_def update_d_graph_aux_def ,metis neq_Nil_conv old.prod.exhaust) +nominal_termination (eqvt) by lexicographic_order + +text \ Relates constructor to the branch in the case and binding variable and statement \ +inductive find_branch :: "dc \ branch_list \ branch_s \ bool" where + find_branch_finalI: "dc' = dc \ find_branch dc' (AS_final (AS_branch dc x s )) (AS_branch dc x s)" +| find_branch_branch_eqI: "dc' = dc \ find_branch dc' (AS_cons (AS_branch dc x s) css) (AS_branch dc x s)" +| find_branch_branch_neqI: "\ dc \ dc'; find_branch dc' css cs \ \ find_branch dc' (AS_cons (AS_branch dc x s) css) cs" +equivariance find_branch +nominal_inductive find_branch . + +inductive_cases find_branch_elims[elim!]: + "find_branch dc (AS_final cs') cs" + "find_branch dc (AS_cons cs' css) cs" + +nominal_function lookup_branch :: "dc \ branch_list \ branch_s option" where + "lookup_branch dc (AS_final (AS_branch dc' x s)) = (if dc = dc' then (Some (AS_branch dc' x s)) else None)" +| "lookup_branch dc (AS_cons (AS_branch dc' x s) css) = (if dc = dc' then (Some (AS_branch dc' x s)) else lookup_branch dc css)" + apply(auto,simp add: eqvt_def lookup_branch_graph_aux_def) + by(metis neq_Nil_conv old.prod.exhaust s_branch_s_branch_list.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +text \ Reduction rules \ +inductive reduce_stmt :: "\ \ \ \ s \ \ \ s \ bool" (" _ \ \ _ , _\ \ \ _ , _\" [50, 50, 50] 50) where + reduce_if_trueI: " \ \ \\, AS_if [L_true]\<^sup>v s1 s2\ \ \\, s1\ " +| reduce_if_falseI: " \ \ \\, AS_if [L_false]\<^sup>v s1 s2\ \ \\, s2\ " +| reduce_let_valI: " \ \ \\, AS_let x (AE_val v) s\ \ \\, s[x::=v]\<^sub>s\<^sub>v\" +| reduce_let_plusI: " \ \ \\, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\ \ + \\, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \ " +| reduce_let_leqI: "b = (if (n1 \ n2) then L_true else L_false) \ + \ \ \\, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\ \ + \\, AS_let x (AE_val (V_lit b)) s\" +| reduce_let_eqI: "b = (if (n1 = n2) then L_true else L_false) \ + \ \ \\, AS_let x ((AE_op Eq (V_lit n1) (V_lit n2))) s\ \ + \\, AS_let x (AE_val (V_lit b)) s\" +| reduce_let_appI: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ z b c \ s'))) = lookup_fun \ f \ + \ \ \\, AS_let x ((AE_app f v)) s\ \ \\, AS_let2 x \[z::=v]\<^sub>\\<^sub>v s'[z::=v]\<^sub>s\<^sub>v s\ " +| reduce_let_appPI: "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \ s'))) = lookup_fun \ f \ + \ \ \\, AS_let x ((AE_appP f b' v)) s\ \ \\, 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\ " +| reduce_let_fstI: "\ \ \\, AS_let x (AE_fst (V_pair v1 v2)) s\ \ \\, AS_let x (AE_val v1) s\" +| reduce_let_sndI: "\ \ \\, AS_let x (AE_snd (V_pair v1 v2)) s\ \ \\, AS_let x (AE_val v2) s\" +| reduce_let_concatI: "\ \ \\, AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s\ \ + \\, AS_let x (AE_val (V_lit (L_bitvec (v1@v2)))) s\" +| reduce_let_splitI: " split n v (v1 , v2 ) \ \ \ \\, AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s\ \ + \\, AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s\" +| reduce_let_lenI: "\ \ \\, AS_let x (AE_len (V_lit (L_bitvec v))) s\ \ + \\, AS_let x (AE_val (V_lit (L_num (int (List.length v))))) s\" +| reduce_let_mvar: "(u,v) \ set \ \ \ \ \\, AS_let x (AE_mvar u) s\ \ \\, AS_let x (AE_val v) s \" +| reduce_assert1I: "\ \ \\, AS_assert c (AS_val v)\ \ \\, AS_val v\" +| reduce_assert2I: " \ \ \\, s \ \ \ \', s'\ \ \ \ \\, AS_assert c s\ \ \ \', AS_assert c s'\" +| reduce_varI: "atom u \ \ \ \ \ \\, AS_var u \ v s \ \ \ ((u,v)#\) , s\" +| reduce_assignI: " \ \ \\, AS_assign u v \ \ \ update_d \ u v , AS_val (V_lit L_unit)\" +| reduce_seq1I: "\ \ \\, AS_seq (AS_val (V_lit L_unit )) s \ \ \\, s\" +| reduce_seq2I: "\s1 \ AS_val v ; \ \ \\, s1\ \ \ \', s1'\ \ \ + \ \ \\, AS_seq s1 s2\ \ \ \', AS_seq s1' s2\" +| reduce_let2_valI: "\ \ \\, AS_let2 x t (AS_val v) s\ \ \\, s[x::=v]\<^sub>s\<^sub>v\" +| reduce_let2I: " \ \ \\, s1 \ \ \ \', s1'\ \ \ \ \\, AS_let2 x t s1 s2\ \ \ \', AS_let2 x t s1' s2\" + +| reduce_caseI: "\ Some (AS_branch dc x' s') = lookup_branch dc cs \ \ \ \ \\, AS_match (V_cons tyid dc v) cs\ \ \\, s'[x'::=v]\<^sub>s\<^sub>v\ " +| reduce_whileI: "\ atom x \ (s1,s2); atom z \ x \ \ + \ \ \\, AS_while s1 s2 \ \ + \\, 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))) \" + +equivariance reduce_stmt +nominal_inductive reduce_stmt . + +inductive_cases reduce_stmt_elims[elim!]: + "\ \ \\, AS_if (V_lit L_true) s1 s2\ \ \\ , s1\" + "\ \ \\, AS_if (V_lit L_false) s1 s2\ \ \\,s2\" + "\ \ \\, AS_let x (AE_val v) s\ \ \\,s'\" + "\ \ \\, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\ \ + \\, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \" + "\ \ \\, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\ \ \\, AS_let x (AE_val (V_lit b)) s\" + "\ \ \\, AS_let x ((AE_app f v)) s\ \ \\, AS_let2 x \ (subst_sv s' x v ) s\ " + "\ \ \\, AS_let x ((AE_len v)) s\ \ \\, AS_let x v' s\ " + "\ \ \\, AS_let x ((AE_concat v1 v2)) s\ \ \\, AS_let x v' s\ " + "\ \ \\, AS_seq s1 s2\ \ \ \' ,s'\" + "\ \ \\, AS_let x ((AE_appP f b v)) s\ \ \\, AS_let2 x \ (subst_sv s' z v) s\ " + "\ \ \\, AS_let x ((AE_split v1 v2)) s\ \ \\, AS_let x v' s\ " + "\ \ \\, AS_assert c s \ \ \\, s'\ " + "\ \ \\, AS_let x ((AE_op Eq (V_lit (n1)) (V_lit (n2)))) s\ \ \\, AS_let x (AE_val (V_lit b)) s\" + +inductive reduce_stmt_many :: "\ \ \ \ s \ \ \ s \ bool" ("_ \ \ _ , _\ \\<^sup>* \ _ , _\" [50, 50, 50] 50) where + reduce_stmt_many_oneI: "\ \ \\, s\ \ \\', s'\ \ \ \ \\ , s\ \\<^sup>* \\', s'\ " +| reduce_stmt_many_manyI: "\ \ \ \\, s\ \ \\', s'\ ; \ \ \\', s'\ \\<^sup>* \\'', s''\ \ \ \ \ \\, s\ \\<^sup>* \\'', s''\" + +nominal_function convert_fds :: "fun_def list \ (f*fun_def) list" where + "convert_fds [] = []" +| "convert_fds ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)))#fs) = ((f,AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)))#convert_fds fs)" +| "convert_fds ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s)))#fs) = ((f,AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s)))#convert_fds fs)" + apply(auto) + apply (simp add: eqvt_def convert_fds_graph_aux_def ) + using fun_def.exhaust fun_typ.exhaust fun_typ_q.exhaust neq_Nil_conv + by metis +nominal_termination (eqvt) by lexicographic_order + +nominal_function convert_tds :: "type_def list \ (f*type_def) list" where + "convert_tds [] = []" +| "convert_tds ((AF_typedef s dclist)#fs) = ((s,AF_typedef s dclist)#convert_tds fs)" +| "convert_tds ((AF_typedef_poly s bv dclist)#fs) = ((s,AF_typedef_poly s bv dclist)#convert_tds fs)" + apply(auto) + apply (simp add: eqvt_def convert_tds_graph_aux_def ) + by (metis type_def.exhaust neq_Nil_conv) +nominal_termination (eqvt) by lexicographic_order + +inductive reduce_prog :: "p \ v \ bool" where + "\ reduce_stmt_many \ [] s \ (AS_val v) \ \ reduce_prog (AP_prog \ \ [] s) v" + +section \Reduction Typing\ + +text \ Checks that the store is consistent with @{typ \} \ +inductive delta_sim :: "\ \ \ \ \ \ bool" ( "_ \ _ \ _ " [50,50] 50 ) where + delta_sim_nilI: "\ \ [] \ []\<^sub>\ " +| delta_sim_consI: "\ \ \ \ \ \ ; \ ; {||} ; GNil \ v \ \ ; u \ fst ` set \ \ \ \ \ ((u,v)#\) \ ((u,\)#\<^sub>\\)" + +equivariance delta_sim +nominal_inductive delta_sim . + +inductive_cases delta_sim_elims[elim!]: + "\ \ [] \ []\<^sub>\" + "\ \ ((u,v)#ds) \ (u,\) #\<^sub>\ D" + "\ \ ((u,v)#ds) \ D" + +text \A typing judgement that combines typing of the statement, the store and the condition that definitions are well-typed\ +inductive config_type :: "\ \ \ \ \ \ \ \ s \ \ \ bool" ("_ ; _ ; _ \ \ _ , _\ \ _ " [50, 50, 50] 50) where + config_typeI: "\ \ ; \ ; {||} ; GNil ; \ \ s \ \; + (\ fd \ set \. \ ; \ \ fd) ; + \ \ \ \ \ \ + \ \ ; \ ; \ \ \ \ , s\ \ \" +equivariance config_type +nominal_inductive config_type . + +inductive_cases config_type_elims [elim!]: + " \ ; \ ; \ \ \ \ , s\ \ \" + +nominal_function \_of :: "var_def list \ \" where + "\_of [] = []" +| "\_of ((AV_def u t v)#vs) = (u,v) # (\_of vs)" + apply auto + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust + by (metis var_def.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +inductive config_type_prog :: "p \ \ \ bool" (" \ \ _\ \ _") where + "\ + \ ; \ ; \_of \ \ \ \_of \ , s\ \ \ +\ \ \ \ AP_prog \ \ \ s\ \ \" + +inductive_cases config_type_prog_elims [elim!]: + "\ \ AP_prog \ \ \ s\ \ \" + +end diff --git a/thys/MiniSail/RCLogic.thy b/thys/MiniSail/RCLogic.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/RCLogic.thy @@ -0,0 +1,222 @@ +(*<*) +theory RCLogic + imports Wellformed +begin + (*>*) + +hide_const Syntax.dom + +chapter \Refinement Constraint Logic\ + +text \ Semantics for the logic we use in the refinement constraints. It is a multi-sorted, quantifier free +logic with polymorphic datatypes and linear arithmetic. We could have modelled by using one of the +encodings to FOL however we wanted to explore using a more direct model. \ + +section \Evaluation and Satisfiability\ + +subsection \Valuation\ + +text \ Refinement constraint logic values. SUt is a value for the uninterpreted + sort that corresponds to basic type variables. For now we only need one of these universes. + We wrap an smt\_val inside it during a process we call 'boxing' + which is introduced in the RCLogicL theory \ +nominal_datatype rcl_val = SBitvec "bit list" | SNum int | SBool bool | SPair rcl_val rcl_val | + SCons tyid string rcl_val | SConsp tyid string b rcl_val | + SUnit | SUt rcl_val + +text \ RCL sorts. Represent our domains. The universe is the union of all of the these. + S\_Ut is the single uninterpreted sort. These map almost directly to basic type + but we have them to distinguish syntax (basic types) and semantics (RCL sorts) \ +nominal_datatype rcl_sort = S_bool | S_int | S_unit | S_pair rcl_sort rcl_sort | S_id tyid | S_app tyid rcl_sort | S_bitvec | S_ut + +type_synonym valuation = "(x,rcl_val) map" + +type_synonym type_valuation = "(bv,rcl_sort) map" + +text \Well-sortedness for RCL values\ +inductive wfRCV:: "\ \ rcl_val \ b \ bool" ( " _ \ _ : _" [50,50] 50) where + wfRCV_BBitvecI: "P \ (SBitvec bv) : B_bitvec" +| wfRCV_BIntI: "P \ (SNum n) : B_int" +| wfRCV_BBoolI: "P \ (SBool b) : B_bool" +| wfRCV_BPairI: "\ P \ s1 : b1 ; P \ s2 : b2 \ \ P \ (SPair s1 s2) : (B_pair b1 b2)" +| wfRCV_BConsI: "\ AF_typedef s dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + \ \ s1 : b \ \ \ \(SCons s dc s1) : (B_id s)" +| wfRCV_BConsPI:"\ AF_typedef_poly s bv dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + atom bv \ (\, SConsp s dc b' s1, B_app s b'); + \ \ s1 : b[bv::=b']\<^sub>b\<^sub>b \ \ \ \(SConsp s dc b' s1) : (B_app s b')" +| wfRCV_BUnitI: "P \ SUnit : B_unit" +| wfRCV_BVarI: "P \ (SUt n) : (B_var bv)" +equivariance wfRCV +nominal_inductive wfRCV + avoids wfRCV_BConsPI: bv +proof(goal_cases) + case (1 s bv dclist \ dc x b c b' s1) + then show ?case using fresh_star_def by auto +next + case (2 s bv dclist \ dc x b c s1 b') + then show ?case by auto +qed + +inductive_cases wfRCV_elims : + "wfRCV P s B_bitvec" + "wfRCV P s (B_pair b1 b2)" + "wfRCV P s (B_int)" + "wfRCV P s (B_bool)" + "wfRCV P s (B_id ss)" + "wfRCV P s (B_var bv)" + "wfRCV P s (B_unit)" + "wfRCV P s (B_app tyid b)" + "wfRCV P (SBitvec bv) b" + "wfRCV P (SNum n) b" + "wfRCV P (SBool n) b" + "wfRCV P (SPair s1 s2) b" + "wfRCV P (SCons s dc s1) b" + "wfRCV P (SConsp s dc b' s1) b" + "wfRCV P SUnit b" + "wfRCV P (SUt s1) b" + +text \ Sometimes we want to assert @{text "P \ s ~ b[bv=b']"} and we want to know what b is +however substitution is not injective so we can't write this in terms of @{text "wfRCV"}. +So we define a relation that makes the components of the substitution explicit. \ + +inductive wfRCV_subst:: "\ \ rcl_val \ b \ (bv*b) option \ bool" where + wfRCV_subst_BBitvecI: "wfRCV_subst P (SBitvec bv) B_bitvec sub " +| wfRCV_subst_BIntI: "wfRCV_subst P (SNum n) B_int sub " +| wfRCV_subst_BBoolI: "wfRCV_subst P (SBool b) B_bool sub " +| wfRCV_subst_BPairI: "\ wfRCV_subst P s1 b1 sub ; wfRCV_subst P s2 b2 sub \ \ wfRCV_subst P (SPair s1 s2) (B_pair b1 b2) sub" +| wfRCV_subst_BConsI: "\ AF_typedef s dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + wfRCV_subst \ s1 b None \ \ wfRCV_subst \ (SCons s dc s1) (B_id s) sub" +| wfRCV_subst_BConspI: "\ AF_typedef_poly s bv dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + wfRCV_subst \ s1 (b[bv::=b']\<^sub>b\<^sub>b) sub \ \ wfRCV_subst \ (SConsp s dc b' s1) (B_app s b') sub" +| wfRCV_subst_BUnitI: "wfRCV_subst P SUnit B_unit sub " +| wfRCV_subst_BVar1I: "bvar \ bv \ wfRCV_subst P (SUt n) (B_var bv) (Some (bvar,bin))" +| wfRCV_subst_BVar2I: "\ bvar = bv; wfRCV_subst P s bin None \ \ wfRCV_subst P s (B_var bv) (Some (bvar,bin))" +| wfRCV_subst_BVar3I: "wfRCV_subst P (SUt n) (B_var bv) None" +equivariance wfRCV_subst +nominal_inductive wfRCV_subst . + +subsection \Evaluation base-types\ + +inductive eval_b :: "type_valuation \ b \ rcl_sort \ bool" ( "_ \ _ \ ~ _ " ) where + "v \ B_bool \ ~ S_bool" +| "v \ B_int \ ~ S_int" +| "Some s = v bv \ v \ B_var bv \ ~ s" +equivariance eval_b +nominal_inductive eval_b . + +subsection \Wellformed vvaluations\ + +definition wfI :: "\ \ \ \ valuation \ bool" ( " _ ; _ \ _" ) where + "\ ; \ \ i = (\ (x,b,c) \ toSet \. \s. Some s = i x \ \ \ s : b)" + +subsection \Evaluating Terms\ + +nominal_function eval_l :: "l \ rcl_val" ( "\ _ \ " ) where + "\ L_true \ = SBool True" +| "\ L_false \ = SBool False" +| "\ L_num n \ = SNum n" +| "\ L_unit \ = SUnit" +| "\ L_bitvec n \ = SBitvec n" + apply(auto simp: eqvt_def eval_l_graph_aux_def) + by (metis l.exhaust) +nominal_termination (eqvt) by lexicographic_order + +inductive eval_v :: "valuation \ v \ rcl_val \ bool" ( "_ \ _ \ ~ _ " ) where + eval_v_litI: "i \ V_lit l \ ~ \ l \ " +| eval_v_varI: "Some sv = i x \ i \ V_var x \ ~ sv" +| eval_v_pairI: "\ i \ v1 \ ~ s1 ; i \ v2 \ ~ s2 \ \ i \ V_pair v1 v2 \ ~ SPair s1 s2" +| eval_v_consI: "i \ v \ ~ s \ i \ V_cons tyid dc v \ ~ SCons tyid dc s" +| eval_v_conspI: "i \ v \ ~ s \ i \ V_consp tyid dc b v \ ~ SConsp tyid dc b s" +equivariance eval_v +nominal_inductive eval_v . + +inductive_cases eval_v_elims: + "i \ V_lit l \ ~ s" + "i \ V_var x \ ~ s" + "i \ V_pair v1 v2 \ ~ s" + "i \ V_cons tyid dc v \ ~ s" + "i \ V_consp tyid dc b v \ ~ s" + +inductive eval_e :: "valuation \ ce \ rcl_val \ bool" ( "_ \ _ \ ~ _ " ) where + eval_e_valI: "i \ v \ ~ sv \ i \ CE_val v \ ~ sv" +| eval_e_plusI: "\ i \ v1 \ ~ SNum n1; i \ v2 \ ~ SNum n2 \ \ i \ (CE_op Plus v1 v2) \ ~ (SNum (n1+n2))" +| eval_e_leqI: "\ i \ v1 \ ~ (SNum n1); i \ v2 \ ~ (SNum n2) \ \ i \ (CE_op LEq v1 v2) \ ~ (SBool (n1 \ n2))" +| eval_e_eqI: "\ i \ v1 \ ~ s1; i \ v2 \ ~ s2 \ \ i \ (CE_op Eq v1 v2) \ ~ (SBool (s1 = s2))" +| eval_e_fstI: "\ i \ v \ ~ SPair v1 v2 \ \ i \ (CE_fst v) \ ~ v1" +| eval_e_sndI: "\ i \ v \ ~ SPair v1 v2 \ \ i \ (CE_snd v) \ ~ v2" +| eval_e_concatI:"\ i \ v1 \ ~ (SBitvec bv1); i \ v2 \ ~ (SBitvec bv2) \ \ i \ (CE_concat v1 v2) \ ~ (SBitvec (bv1@bv2))" +| eval_e_lenI:"\ i \ v \ ~ (SBitvec bv) \ \ i \ (CE_len v) \ ~ (SNum (int (List.length bv)))" +equivariance eval_e +nominal_inductive eval_e . + +inductive_cases eval_e_elims: + "i \ (CE_val v) \ ~ s" + "i \ (CE_op Plus v1 v2) \ ~ s" + "i \ (CE_op LEq v1 v2) \ ~ s" + "i \ (CE_op Eq v1 v2) \ ~ s" + "i \ (CE_fst v) \ ~ s" + "i \ (CE_snd v) \ ~ s" + "i \ (CE_concat v1 v2) \ ~ s" + "i \ (CE_len v) \ ~ s" + +inductive eval_c :: "valuation \ c \ bool \ bool" ( " _ \ _ \ ~ _ ") where + eval_c_trueI: "i \ C_true \ ~ True" +| eval_c_falseI:"i \ C_false \ ~ False" +| eval_c_conjI: "\ i \ c1 \ ~ b1 ; i \ c2 \ ~ b2 \ \ i \ (C_conj c1 c2) \ ~ (b1 \ b2)" +| eval_c_disjI: "\ i \ c1 \ ~ b1 ; i \ c2 \ ~ b2 \ \ i \ (C_disj c1 c2) \ ~ (b1 \ b2)" +| eval_c_impI:"\ i \ c1 \ ~ b1 ; i \ c2 \ ~ b2 \ \ i \ (C_imp c1 c2) \ ~ (b1 \ b2)" +| eval_c_notI:"\ i \ c \ ~ b \ \ i \ (C_not c) \ ~ (\ b)" +| eval_c_eqI:"\ i \ e1 \ ~ sv1; i \ e2 \ ~ sv2 \ \ i \ (C_eq e1 e2) \ ~ (sv1=sv2)" +equivariance eval_c +nominal_inductive eval_c . + +inductive_cases eval_c_elims: + "i \ C_true \ ~ True" + "i \ C_false \ ~ False" + "i \ (C_conj c1 c2)\ ~ s" + "i \ (C_disj c1 c2)\ ~ s" + "i \ (C_imp c1 c2)\ ~ s" + "i \ (C_not c) \ ~ s" + "i \ (C_eq e1 e2)\ ~ s" + "i \ C_true \ ~ s" + "i \ C_false \ ~ s" + +subsection \Satisfiability\ + +inductive is_satis :: "valuation \ c \ bool" ( " _ \ _ " ) where + "i \ c \ ~ True \ i \ c" +equivariance is_satis +nominal_inductive is_satis . + +nominal_function is_satis_g :: "valuation \ \ \ bool" ( " _ \ _ " ) where + "i \ GNil = True" +| "i \ ((x,b,c) #\<^sub>\ G) = ( i \ c \ i \ G)" + apply(auto simp: eqvt_def is_satis_g_graph_aux_def) + by (metis \.exhaust old.prod.exhaust) +nominal_termination (eqvt) by lexicographic_order + +section \Validity\ + +nominal_function valid :: "\ \ \ \ \ \ c \ bool" ("_ ; _ ; _ \ _ " [50, 50] 50) where + "P ; B ; G \ c = ( (P ; B ; G \\<^sub>w\<^sub>f c) \ (\i. (P ; G \ i) \ i \ G \ i \ c))" + by (auto simp: eqvt_def wfI_def valid_graph_aux_def) +nominal_termination (eqvt) by lexicographic_order + +section \Lemmas\ +text \Lemmas needed for Examples\ + +lemma valid_trueI [intro]: + fixes G::\ + assumes "P ; B \\<^sub>w\<^sub>f G" + shows "P ; B ; G \ C_true" +proof - + have "\i. i \ C_true" using is_satis.simps eval_c_trueI by simp + moreover have "P ; B ; G \\<^sub>w\<^sub>f C_true" using wfC_trueI assms by simp + ultimately show ?thesis using valid.simps by simp +qed + +end \ No newline at end of file diff --git a/thys/MiniSail/RCLogicL.thy b/thys/MiniSail/RCLogicL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/RCLogicL.thy @@ -0,0 +1,2729 @@ +(*<*) +theory RCLogicL + imports RCLogic WellformedL +begin + (*>*) + +hide_const Syntax.dom + +chapter \Refinement Constraint Logic Lemmas\ + +section \Lemmas\ + +lemma wfI_domi: + assumes "\ ; \ \ i" + shows "fst ` toSet \ \ dom i" + using wfI_def toSet.simps assms by fastforce + +lemma wfI_lookup: + fixes G::\ and b::b + assumes "Some (b,c) = lookup G x" and "P ; G \ i" and "Some s = i x" and "P ; B \\<^sub>w\<^sub>f G" + shows "P \ s : b" +proof - + have "(x,b,c) \ toSet G" using lookup.simps assms + using lookup_in_g by blast + then obtain s' where *:"Some s' = i x \ wfRCV P s' b" using wfI_def assms by auto + hence "s' = s" using assms by (metis option.inject) + thus ?thesis using * by auto +qed + +lemma wfI_restrict_weakening: + assumes "wfI \ \' i'" and "i = restrict_map i' (fst `toSet \)" and "toSet \ \ toSet \'" + shows "\ ; \ \ i" +proof - + { fix x + assume "x \ toSet \" + have "case x of (x, b, c) \ \s. Some s = i x \ \ \ s : b" using assms wfI_def + proof - + have "case x of (x, b, c) \ \s. Some s = i' x \ \ \ s : b" + using \x \ toSet \\ assms wfI_def by auto + then have "\s. Some s = i (fst x) \ wfRCV \ s (fst (snd x))" + by (simp add: \x \ toSet \\ assms(2) case_prod_unfold) + then show ?thesis + by (simp add: case_prod_unfold) + qed + } + thus ?thesis using wfI_def assms by auto +qed + +lemma wfI_suffix: + fixes G::\ + assumes "wfI P (G'@G) i" and "P ; B \\<^sub>w\<^sub>f G" + shows "P ; G \ i" + using wfI_def append_g.simps assms toSet.simps by simp + +lemma wfI_replace_inside: + assumes "wfI \ (\' @ (x, b, c) #\<^sub>\ \) i" + shows "wfI \ (\' @ (x, b, c') #\<^sub>\ \) i" + using wfI_def toSet_splitP assms by simp + +section \Existence of evaluation\ + +lemma eval_l_base: + "\ \ \ l \ : (base_for_lit l)" + apply(nominal_induct l rule:l.strong_induct) + using wfRCV.intros eval_l.simps base_for_lit.simps by auto+ + +lemma obtain_fresh_bv_dclist: + fixes tm::"'a::fs" + assumes "(dc, \ x : b | c \) \ set dclist" + obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 + \ (dc, \ x1 : b1 | c1 \) \ set dclist1 \ atom bv1 \ tm" +proof - + obtain bv1 dclist1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \ atom bv1 \ tm" + using obtain_fresh_bv by metis + moreover hence "[[atom bv]]lst. dclist = [[atom bv1]]lst. dclist1" using type_def.eq_iff by metis + moreover then obtain x1 b1 c1 where "(dc, \ x1 : b1 | c1 \) \ set dclist1" using td_lookup_eq_iff assms by metis + ultimately show ?thesis using that by blast +qed + +lemma obtain_fresh_bv_dclist_b_iff: + fixes tm::"'a::fs" + assumes "(dc, \ x : b | c \) \ set dclist" and "AF_typedef_poly tyid bv dclist \ set P" and "\\<^sub>w\<^sub>f P" + obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 + \ (dc, \ x1 : b1 | c1 \) \ set dclist1 \ atom bv1 \ tm \ b[bv::=b']\<^sub>b\<^sub>b=b1[bv1::=b']\<^sub>b\<^sub>b" +proof - + obtain bv1 dclist1 x1 b1 c1 where *:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \ atom bv1 \ tm + \ (dc, \ x1 : b1 | c1 \) \ set dclist1" using obtain_fresh_bv_dclist assms by metis + + hence "AF_typedef_poly tyid bv1 dclist1 \ set P" using assms by metis + + hence "b[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b" + using wfTh_typedef_poly_b_eq_iff[OF assms(2) assms(1) _ _ assms(3),of bv1 dclist1 x1 b1 c1 b'] * by metis + + from this that show ?thesis using * by metis +qed + +lemma eval_v_exist: + fixes \::\ and v::v and b::b + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "\s. i \ v \ ~ s \ P \ s : b " + using assms proof(nominal_induct v arbitrary: b rule:v.strong_induct) + case (V_lit x) + then show ?case using eval_l_base eval_v.intros eval_l.simps wfV_elims rcl_val.supp pure_supp by metis +next + case (V_var x) + then obtain c where *:"Some (b,c) = lookup \ x" using wfV_elims by metis + hence "x \ fst ` toSet \" using lookup_x by blast + hence "x \ dom i" using wfI_domi using assms by blast + then obtain s where "i x = Some s" by auto + moreover hence "P \ s : b" using wfRCV.intros wfI_lookup * V_var + by (metis wfV_wf) + + ultimately show ?case using eval_v.intros rcl_val.supp b.supp by metis +next + case (V_pair v1 v2) + then obtain b1 and b2 where *:"P ; B ; \ \\<^sub>w\<^sub>f v1 : b1 \ P ; B ; \ \\<^sub>w\<^sub>f v2 : b2 \ b = B_pair b1 b2" using wfV_elims by metis + then obtain s1 and s2 where "eval_v i v1 s1 \ wfRCV P s1 b1 \ eval_v i v2 s2 \ wfRCV P s2 b2" using V_pair by metis + thus ?case using eval_v.intros wfRCV.intros * by metis +next + case (V_cons tyid dc v) + then obtain s and b'::b and dclist and x::x and c::c where "(wfV P B \ v b') \ i \ v \ ~ s \ P \ s : b' \ b = B_id tyid \ + AF_typedef tyid dclist \ set P \ (dc, \ x : b' | c \) \ set dclist" using wfV_elims(4) by metis + then show ?case using eval_v.intros(4) wfRCV.intros(5) V_cons by metis +next + case (V_consp tyid dc b' v) + + obtain b'a::b and bv and dclist and x::x and c::c where *:"(wfV P B \ v b'a[bv::=b']\<^sub>b\<^sub>b) \ b = B_app tyid b' \ + AF_typedef_poly tyid bv dclist \ set P \ (dc, \ x : b'a | c \) \ set dclist \ + atom bv \ (P, B_app tyid b',B) " using wf_strong_elim(1)[OF V_consp(3)] by metis + + then obtain s where **:"i \ v \ ~ s \ P \ s : b'a[bv::=b']\<^sub>b\<^sub>b" using V_consp by auto + + have " \\<^sub>w\<^sub>f P" using wfX_wfY V_consp by metis + then obtain bv1::bv and dclist1 x1 b1 c1 where 3:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 + \ (dc, \ x1 : b1 | c1 \) \ set dclist1 \ atom bv1 \ (P, SConsp tyid dc b' s, B_app tyid b') + \ b'a[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b" + using * obtain_fresh_bv_dclist_b_iff by blast + + have " i \ V_consp tyid dc b' v \ ~ SConsp tyid dc b' s" using eval_v.intros by (simp add: "**") + + moreover have " P \ SConsp tyid dc b' s : B_app tyid b'" proof + show \AF_typedef_poly tyid bv1 dclist1 \ set P\ using 3 * by metis + show \(dc, \ x1 : b1 | c1 \) \ set dclist1\ using 3 by auto + thus \atom bv1 \ (P, SConsp tyid dc b' s, B_app tyid b')\ using * 3 fresh_prodN by metis + show \ P \ s : b1[bv1::=b']\<^sub>b\<^sub>b\ using 3 ** by auto + qed + + ultimately show ?case using eval_v.intros wfRCV.intros V_consp * by auto +qed + +lemma eval_v_uniqueness: + fixes v::v + assumes "i \ v \ ~ s" and "i \ v \ ~ s'" + shows "s=s'" + using assms proof(nominal_induct v arbitrary: s s' rule:v.strong_induct) + case (V_lit x) + then show ?case using eval_v_elims eval_l.simps by metis +next + case (V_var x) + then show ?case using eval_v_elims by (metis option.inject) +next + case (V_pair v1 v2) + obtain s1 and s2 where s:"i \ v1 \ ~ s1 \ i \ v2 \ ~ s2 \ s = SPair s1 s2" using eval_v_elims V_pair by metis + obtain s1' and s2' where s':"i \ v1 \ ~ s1' \ i \ v2 \ ~ s2' \ s' = SPair s1' s2'" using eval_v_elims V_pair by metis + then show ?case using eval_v_elims using V_pair s s' by auto +next + case (V_cons tyid dc v1) + obtain sv1 where 1:"i \ v1 \ ~ sv1 \ s = SCons tyid dc sv1" using eval_v_elims V_cons by metis + moreover obtain sv2 where 2:"i \ v1 \ ~ sv2 \ s' = SCons tyid dc sv2" using eval_v_elims V_cons by metis + ultimately have "sv1 = sv2" using V_cons by auto + then show ?case using 1 2 by auto +next + case (V_consp tyid dc b v1) + then show ?case using eval_v_elims by metis + +qed + +lemma eval_v_base: + fixes \::\ and v::v and b::b + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f v : b" and "i \ v \ ~ s" + shows "P \ s : b" + using eval_v_exist eval_v_uniqueness assms by metis + +lemma eval_e_uniqueness: + fixes e::ce + assumes "i \ e \ ~ s" and "i \ e \ ~ s'" + shows "s=s'" + using assms proof(nominal_induct e arbitrary: s s' rule:ce.strong_induct) + case (CE_val x) + then show ?case using eval_v_uniqueness eval_e_elims by metis +next + case (CE_op opp x1 x2) + consider "opp = Plus" | "opp = LEq" | "opp = Eq" using opp.exhaust by metis + thus ?case proof(cases) + case 1 + hence a1:"eval_e i (CE_op Plus x1 x2) s" and a2:"eval_e i (CE_op Plus x1 x2) s'" using CE_op by auto + then show ?thesis using eval_e_elims(2)[OF a1] eval_e_elims(2)[OF a2] + CE_op eval_e_plusI + by (metis rcl_val.eq_iff(2)) + next + case 2 + hence a1:"eval_e i (CE_op LEq x1 x2) s" and a2:"eval_e i (CE_op LEq x1 x2) s'" using CE_op by auto + then show ?thesis using eval_v_uniqueness eval_e_elims(3)[OF a1] eval_e_elims(3)[OF a2] + CE_op eval_e_plusI + by (metis rcl_val.eq_iff(2)) + next + case 3 + hence a1:"eval_e i (CE_op Eq x1 x2) s" and a2:"eval_e i (CE_op Eq x1 x2) s'" using CE_op by auto + then show ?thesis using eval_v_uniqueness eval_e_elims(4)[OF a1] eval_e_elims(4)[OF a2] + CE_op eval_e_plusI + by (metis rcl_val.eq_iff(2)) + qed +next + case (CE_concat x1 x2) + hence a1:"eval_e i (CE_concat x1 x2) s" and a2:"eval_e i (CE_concat x1 x2) s'" using CE_concat by auto + show ?case using eval_e_elims(7)[OF a1] eval_e_elims(7)[OF a2] CE_concat eval_e_concatI rcl_val.eq_iff + proof - + assume "\P. (\bv1 bv2. \s' = SBitvec (bv1 @ bv2); i \ x1 \ ~ SBitvec bv1 ; i \ x2 \ ~ SBitvec bv2 \ \ P) \ P" + obtain bbs :: "bit list" and bbsa :: "bit list" where + "i \ x2 \ ~ SBitvec bbs \ i \ x1 \ ~ SBitvec bbsa \ SBitvec (bbsa @ bbs) = s'" + by (metis \\P. (\bv1 bv2. \s' = SBitvec (bv1 @ bv2); i \ x1 \ ~ SBitvec bv1 ; i \ x2 \ ~ SBitvec bv2 \ \ P) \ P\) (* 93 ms *) + then have "s' = s" + by (metis (no_types) \\P. (\bv1 bv2. \s = SBitvec (bv1 @ bv2); i \ x1 \ ~ SBitvec bv1 ; i \ x2 \ ~ SBitvec bv2 \ \ P) \ P\ \\s' s. \i \ x1 \ ~ s ; i \ x1 \ ~ s' \ \ s = s'\ \\s' s. \i \ x2 \ ~ s ; i \ x2 \ ~ s' \ \ s = s'\ rcl_val.eq_iff(1)) (* 125 ms *) + then show ?thesis + by metis (* 0.0 ms *) + qed +next + case (CE_fst x) + then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff) +next + case (CE_snd x) + then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff) +next + case (CE_len x) + then show ?case using eval_e_elims rcl_val.eq_iff + proof - + obtain bbs :: "rcl_val \ ce \ (x \ rcl_val option) \ bit list" where + "\x0 x1 x2. (\v3. x0 = SNum (int (length v3)) \ x2 \ x1 \ ~ SBitvec v3 ) = (x0 = SNum (int (length (bbs x0 x1 x2))) \ x2 \ x1 \ ~ SBitvec (bbs x0 x1 x2) )" + by moura (* 0.0 ms *) + then have "\f c r. \ f \ [| c |]\<^sup>c\<^sup>e \ ~ r \ r = SNum (int (length (bbs r c f))) \ f \ c \ ~ SBitvec (bbs r c f)" + by (meson eval_e_elims(8)) (* 46 ms *) + then show ?thesis + by (metis (no_types) CE_len.hyps CE_len.prems(1) CE_len.prems(2) rcl_val.eq_iff(1)) (* 31 ms *) + qed + +qed + +lemma wfV_eval_bitvec: + fixes v::v + assumes "P ; B ; \ \\<^sub>w\<^sub>f v : B_bitvec" and "P ; \ \ i" + shows "\bv. eval_v i v (SBitvec bv)" +proof - + obtain s where "i \ v \ ~ s \ wfRCV P s B_bitvec" using eval_v_exist assms by metis + moreover then obtain bv where "s = SBitvec bv" using wfRCV_elims(1)[of P s] by metis + ultimately show ?thesis by metis +qed + +lemma wfV_eval_pair: + fixes v::v + assumes "P ; B ; \ \\<^sub>w\<^sub>f v : B_pair b1 b2" and "P ; \ \ i" + shows "\s1 s2. eval_v i v (SPair s1 s2)" +proof - + obtain s where "i \ v \ ~ s \ wfRCV P s (B_pair b1 b2)" using eval_v_exist assms by metis + moreover then obtain s1 and s2 where "s = SPair s1 s2" using wfRCV_elims(2)[of P s] by metis + ultimately show ?thesis by metis +qed + +lemma wfV_eval_int: + fixes v::v + assumes "P ; B ; \ \\<^sub>w\<^sub>f v : B_int" and "P ; \ \ i" + shows "\n. eval_v i v (SNum n)" +proof - + obtain s where "i \ v \ ~ s \ wfRCV P s (B_int)" using eval_v_exist assms by metis + moreover then obtain n where "s = SNum n" using wfRCV_elims(3)[of P s] by metis + ultimately show ?thesis by metis +qed + +text \Well sorted value with a well sorted valuation evaluates\ +lemma wfI_wfV_eval_v: + fixes v::v and b::b + assumes "\ ; B ; \ \\<^sub>w\<^sub>f v : b" and "wfI \ \ i" + shows "\s. i \ v \ ~ s \ \ \ s : b" + using eval_v_exist assms by auto + +lemma wfI_wfCE_eval_e: + fixes e::ce and b::b + assumes "wfCE P B G e b" and "P ; G \ i" + shows "\s. i \ e \ ~ s \ P \ s : b" + using assms proof(nominal_induct e arbitrary: b rule: ce.strong_induct) + case (CE_val v) + obtain s where "i \ v \ ~ s \ P \ s : b" using wfI_wfV_eval_v[of P B G v b i] assms wfCE_elims(1) [of P B G v b] CE_val by auto + then show ?case using CE_val eval_e.intros(1)[of i v s ] by auto +next + case (CE_op opp v1 v2) + + consider "opp =Plus" | "opp=LEq" | "opp=Eq" using opp.exhaust by auto + + thus ?case proof(cases) + case 1 + hence "wfCE P B G v1 B_int \ wfCE P B G v2 B_int" using wfCE_elims(2) CE_op + + by blast + then obtain s1 and s2 where *: "eval_e i v1 s1 \ wfRCV P s1 B_int \ eval_e i v2 s2 \ wfRCV P s2 B_int" + using wfI_wfV_eval_v CE_op by metis + then obtain n1 and n2 where **:"s2=SNum n2 \ s1 = SNum n1" using wfRCV_elims by meson + hence "eval_e i (CE_op Plus v1 v2) (SNum (n1+n2))" using eval_e_plusI * ** by simp + moreover have "wfRCV P (SNum (n1+n2)) B_int" using wfRCV.intros by auto + ultimately show ?thesis using 1 + using CE_op.prems(1) wfCE_elims(2) by blast + next + case 2 + hence "wfCE P B G v1 B_int \ wfCE P B G v2 B_int" using wfCE_elims(3) CE_op + by blast + then obtain s1 and s2 where *: "eval_e i v1 s1 \ wfRCV P s1 B_int \ eval_e i v2 s2 \ wfRCV P s2 B_int" + using wfI_wfV_eval_v CE_op by metis + then obtain n1 and n2 where **:"s2=SNum n2 \ s1 = SNum n1" using wfRCV_elims by meson + hence "eval_e i (CE_op LEq v1 v2) (SBool (n1 \ n2))" using eval_e_leqI * ** by simp + moreover have "wfRCV P (SBool (n1\n2)) B_bool" using wfRCV.intros by auto + ultimately show ?thesis using 2 + using CE_op.prems wfCE_elims by metis + next + case 3 + then obtain b2 where "wfCE P B G v1 b2 \ wfCE P B G v2 b2" using wfCE_elims(9) CE_op + by blast + then obtain s1 and s2 where *: "eval_e i v1 s1 \ wfRCV P s1 b2 \ eval_e i v2 s2 \ wfRCV P s2 b2" + using wfI_wfV_eval_v CE_op by metis + hence "eval_e i (CE_op Eq v1 v2) (SBool (s1 = s2))" using eval_e_leqI * + by (simp add: eval_e_eqI) + moreover have "wfRCV P (SBool (s1 = s2)) B_bool" using wfRCV.intros by auto + ultimately show ?thesis using 3 + using CE_op.prems wfCE_elims by metis + qed +next + case (CE_concat v1 v2) + then obtain s1 and s2 where *:"b = B_bitvec \ eval_e i v1 s1 \ eval_e i v2 s2 \ + wfRCV P s1 B_bitvec \ wfRCV P s2 B_bitvec" using + CE_concat + by (meson wfCE_elims(6)) + thus ?case using eval_e_concatI wfRCV.intros(1) wfRCV_elims + proof - + obtain bbs :: "type_def list \ rcl_val \ bit list" where + "\ts s. \ ts \ s : B_bitvec \ s = SBitvec (bbs ts s)" + using wfRCV_elims(1) by moura (* 156 ms *) + then show ?thesis + by (metis (no_types) "local.*" wfRCV_BBitvecI eval_e_concatI) (* 78 ms *) + qed +next + case (CE_fst v1) + thus ?case using eval_e_fstI wfRCV.intros wfCE_elims wfI_wfV_eval_v + by (metis wfRCV_elims(2) rcl_val.eq_iff(4)) +next + case (CE_snd v1) + thus ?case using eval_e_sndI wfRCV.intros wfCE_elims wfI_wfV_eval_v + by (metis wfRCV_elims(2) rcl_val.eq_iff(4)) +next + case (CE_len x) + thus ?case using eval_e_lenI wfRCV.intros wfCE_elims wfI_wfV_eval_v wfV_eval_bitvec + by (metis wfRCV_elims(1)) +qed + +lemma eval_e_exist: + fixes \::\ and e::ce + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f e : b" + shows "\s. i \ e \ ~ s" + using assms proof(nominal_induct e arbitrary: b rule:ce.strong_induct) + case (CE_val v) + then show ?case using eval_v_exist wfCE_elims eval_e.intros by metis +next + case (CE_op op v1 v2) + + show ?case proof(rule opp.exhaust) + assume \op = Plus\ + hence "P ; B ; \ \\<^sub>w\<^sub>f v1 : B_int \ P ; B ; \ \\<^sub>w\<^sub>f v2 : B_int \ b = B_int" using wfCE_elims CE_op by metis + then obtain n1 and n2 where "eval_e i v1 (SNum n1) \ eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int + by (metis wfI_wfCE_eval_e wfRCV_elims(3)) + then show \\a. eval_e i (CE_op op v1 v2) a\ using eval_e_plusI[of i v1 _ v2] \op=Plus\ by auto + next + assume \op = LEq\ + hence "P ; B ; \ \\<^sub>w\<^sub>f v1 : B_int \ P ; B ; \ \\<^sub>w\<^sub>f v2 : B_int \ b = B_bool" using wfCE_elims CE_op by metis + then obtain n1 and n2 where "eval_e i v1 (SNum n1) \ eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int + by (metis wfI_wfCE_eval_e wfRCV_elims(3)) + then show \\a. eval_e i (CE_op op v1 v2) a\ using eval_e_leqI[of i v1 _ v2] eval_v_exist \op=LEq\ CE_op by auto + next + assume \op = Eq\ + then obtain b1 where "P ; B ; \ \\<^sub>w\<^sub>f v1 : b1 \ P ; B ; \ \\<^sub>w\<^sub>f v2 : b1 \ b = B_bool" using wfCE_elims CE_op by metis + then obtain s1 and s2 where "eval_e i v1 s1 \ eval_e i v2 s2" using CE_op eval_v_exist wfV_eval_int + by (metis wfI_wfCE_eval_e wfRCV_elims(3)) + then show \\a. eval_e i (CE_op op v1 v2) a\ using eval_e_eqI[of i v1 _ v2] eval_v_exist \op=Eq\ CE_op by auto + qed +next + case (CE_concat v1 v2) + then obtain bv1 and bv2 where "eval_e i v1 (SBitvec bv1) \ eval_e i v2 (SBitvec bv2)" + using wfV_eval_bitvec wfCE_elims(6) + by (meson eval_e_elims(7) wfI_wfCE_eval_e) + then show ?case using eval_e.intros by metis +next + case (CE_fst ce) + then obtain b2 where b:"P ; B ; \ \\<^sub>w\<^sub>f ce : B_pair b b2" using wfCE_elims by metis + then obtain s where s:"i \ ce \ ~ s" using CE_fst by auto + then obtain s1 and s2 where "s = (SPair s1 s2)" using eval_e_elims(4) CE_fst wfI_wfCE_eval_e[of P B \ ce "B_pair b b2" i,OF b] wfRCV_elims(2)[of P s b b2] + by (metis eval_e_uniqueness) + then show ?case using s eval_e.intros by metis +next + case (CE_snd ce) + then obtain b1 where b:"P ; B ; \ \\<^sub>w\<^sub>f ce : B_pair b1 b" using wfCE_elims by metis + then obtain s where s:"i \ ce \ ~ s" using CE_snd by auto + then obtain s1 and s2 where "s = (SPair s1 s2)" + using eval_e_elims(5) CE_snd wfI_wfCE_eval_e[of P B \ ce "B_pair b1 b" i,OF b] wfRCV_elims(2)[of P s b b1] + eval_e_uniqueness + by (metis wfRCV_elims(2)) + then show ?case using s eval_e.intros by metis +next + case (CE_len v1) + then obtain bv1 where "eval_e i v1 (SBitvec bv1)" + using wfV_eval_bitvec CE_len wfCE_elims eval_e_uniqueness + by (metis eval_e_elims(7) wfCE_concatI wfI_wfCE_eval_e) + then show ?case using eval_e.intros by metis +qed + +lemma eval_c_exist: + fixes \::\ and c::c + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f c" + shows "\s. i \ c \ ~ s" + using assms proof(nominal_induct c rule: c.strong_induct) + case C_true + then show ?case using eval_c.intros wfC_elims by metis +next + case C_false + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_conj c1 c2) + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_disj x1 x2) + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_not x) + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_imp x1 x2) + then show ?case using eval_c.intros eval_e_exist wfC_elims by metis +next + case (C_eq x1 x2) + then show ?case using eval_c.intros eval_e_exist wfC_elims by metis +qed + +lemma eval_c_uniqueness: + fixes c::c + assumes "i \ c \ ~ s" and "i \ c \ ~ s'" + shows "s=s'" + using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct) + case C_true + then show ?case using eval_c_elims by metis +next + case C_false + then show ?case using eval_c_elims by metis +next + case (C_conj x1 x2) + then show ?case using eval_c_elims(3) by (metis (full_types)) +next + case (C_disj x1 x2) + then show ?case using eval_c_elims(4) by (metis (full_types)) +next + case (C_not x) + then show ?case using eval_c_elims(6) by (metis (full_types)) +next + case (C_imp x1 x2) + then show ?case using eval_c_elims(5) by (metis (full_types)) +next + case (C_eq x1 x2) + then show ?case using eval_e_uniqueness eval_c_elims(7) by metis +qed + +lemma wfI_wfC_eval_c: + fixes c::c + assumes "wfC P B G c" and "P ; G \ i" + shows "\s. i \ c \ ~ s" + using assms proof(nominal_induct c rule: c.strong_induct) +qed(metis wfC_elims wfI_wfCE_eval_e eval_c.intros)+ + +section \Satisfiability\ + +lemma satis_reflI: + fixes c::c + assumes "i \ ((x, b, c) #\<^sub>\ G)" + shows "i \ c" + using assms by auto + +lemma is_satis_mp: + fixes c1::c and c2::c + assumes "i \ (c1 IMP c2)" and "i \ c1" + shows "i \ c2" + using assms proof - + have "eval_c i (c1 IMP c2) True" using is_satis.simps using assms by blast + then obtain b1 and b2 where "True = (b1 \ b2) \ eval_c i c1 b1 \ eval_c i c2 b2" + using eval_c_elims(5) by metis + moreover have "eval_c i c1 True" using is_satis.simps using assms by blast + moreover have "b1 = True" using calculation eval_c_uniqueness by blast + ultimately have "eval_c i c2 True" by auto + thus ?thesis using is_satis.intros by auto +qed + +lemma is_satis_imp: + fixes c1::c and c2::c + assumes "i \ c1 \ i \ c2" and "i \ c1 \ ~ b1" and "i \ c2 \ ~ b2" + shows "i \ (c1 IMP c2)" +proof(cases b1) + case True + hence "i \ c2" using assms is_satis.simps by simp + hence "b2 = True" using is_satis.simps assms + using eval_c_uniqueness by blast + then show ?thesis using eval_c_impI is_satis.simps assms by force +next + case False + then show ?thesis using assms eval_c_impI is_satis.simps by metis +qed + +lemma is_satis_iff: + "i \ G = (\x b c. (x,b,c) \ toSet G \ i \ c)" + by(induct G,auto) + +lemma is_satis_g_append: + "i \ (G1@G2) = (i \ G1 \ i \ G2)" + using is_satis_g.simps is_satis_iff by auto + +section \Substitution for Evaluation\ + +lemma eval_v_i_upd: + fixes v::v + assumes "atom x \ v" and "i \ v \ ~ s'" + shows "eval_v ((i ( x \s))) v s' " + using assms proof(nominal_induct v arbitrary: s' rule:v.strong_induct) + case (V_lit x) + then show ?case by (metis eval_v_elims(1) eval_v_litI) +next + case (V_var y) + then obtain s where *: "Some s = i y \ s=s'" using eval_v_elims by metis + moreover have "x \ y" using \atom x \ V_var y\ v.supp by simp + ultimately have "(i ( x \s)) y = Some s" + by (simp add: \Some s = i y \ s = s'\) + then show ?case using eval_v_varI * \x \ y\ + by (simp add: eval_v.eval_v_varI) +next + case (V_pair v1 v2) + hence "atom x \ v1 \ atom x \ v2" using v.supp by simp + moreover obtain s1 and s2 where *: "eval_v i v1 s1 \ eval_v i v2 s2 \ s' = SPair s1 s2" using eval_v_elims V_pair by metis + ultimately have "eval_v ((i ( x \s))) v1 s1 \ eval_v ((i ( x \s))) v2 s2" using V_pair by blast + thus ?case using eval_v_pairI * by meson +next + case (V_cons tyid dc v1) + hence "atom x \ v1" using v.supp by simp + moreover obtain s1 where *: "eval_v i v1 s1 \ s' = SCons tyid dc s1" using eval_v_elims V_cons by metis + ultimately have "eval_v ((i ( x \s))) v1 s1" using V_cons by blast + thus ?case using eval_v_consI * by meson +next + case (V_consp tyid dc b1 v1) + + hence "atom x \ v1" using v.supp by simp + moreover obtain s1 where *: "eval_v i v1 s1 \ s' = SConsp tyid dc b1 s1" using eval_v_elims V_consp by metis + ultimately have "eval_v ((i ( x \s))) v1 s1" using V_consp by blast + thus ?case using eval_v_conspI * by meson +qed + +lemma eval_e_i_upd: + fixes e::ce + assumes "i \ e \ ~ s'" and "atom x \ e" + shows " (i ( x \s)) \ e \ ~ s'" + using assms apply(induct rule: eval_e.induct) using eval_v_i_upd eval_e_elims + by (meson ce.fresh eval_e.intros)+ + +lemma eval_c_i_upd: + fixes c::c + assumes "i \ c \ ~ s'" and "atom x \ c" + shows "((i ( x \s))) \ c \ ~ s' " + using assms proof(induct rule:eval_c.induct) + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using RCLogic.eval_c_eqI eval_e_i_upd c.fresh by metis +qed(simp add: eval_c.intros)+ + +lemma subst_v_eval_v[simp]: + fixes v::v and v'::v + assumes "i \ v \ ~ s" and "i \ (v'[x::=v]\<^sub>v\<^sub>v) \ ~ s'" + shows "(i ( x \ s )) \ v' \ ~ s'" + using assms proof(nominal_induct v' arbitrary: s' rule:v.strong_induct) + case (V_lit x) + then show ?case using subst_vv.simps + by (metis eval_v_elims(1) eval_v_litI) +next + case (V_var x') + then show ?case proof(cases "x=x'") + case True + hence "(V_var x')[x::=v]\<^sub>v\<^sub>v = v" using subst_vv.simps by auto + then show ?thesis using V_var eval_v_elims eval_v_varI eval_v_uniqueness True + by (simp add: eval_v.eval_v_varI) + next + case False + hence "atom x \ (V_var x')" by simp + then show ?thesis using eval_v_i_upd False V_var by fastforce + qed +next + case (V_pair v1 v2) + then obtain s1 and s2 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \ eval_v i (v2[x::=v]\<^sub>v\<^sub>v) s2 \ s' = SPair s1 s2" using V_pair eval_v_elims subst_vv.simps by metis + hence "(i ( x \ s )) \ v1 \ ~ s1 \ (i ( x \ s )) \ v2 \ ~ s2" using V_pair by metis + thus ?case using eval_v_pairI subst_vv.simps * V_pair by metis +next + case (V_cons tyid dc v1) + then obtain s1 where "eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1" using eval_v_elims subst_vv.simps by metis + thus ?case using eval_v_consI V_cons + by (metis eval_v_elims subst_vv.simps) +next + case (V_consp tyid dc b1 v1) + + then obtain s1 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \ s' = SConsp tyid dc b1 s1" using eval_v_elims subst_vv.simps by metis + hence "i ( x \ s ) \ v1 \ ~ s1" using V_consp by metis + thus ?case using * eval_v_conspI by metis +qed + +lemma subst_e_eval_v[simp]: + fixes y::x and e::ce and v::v and e'::ce + assumes "i \ e' \ ~ s'" and "e'=(e[y::=v]\<^sub>c\<^sub>e\<^sub>v)" and "i \ v \ ~ s" + shows "(i ( y \ s )) \ e \ ~ s'" + using assms proof(induct arbitrary: e rule: eval_e.induct) + case (eval_e_valI i v1 sv) + then obtain v1' where *:"e = CE_val v1' \ v1 = v1'[y::=v]\<^sub>v\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_v i (v1'[y::=v]\<^sub>v\<^sub>v) sv" using eval_e_valI by simp + hence "eval_v (i ( y \ s )) v1' sv" using subst_v_eval_v eval_e_valI by simp + then show ?case using RCLogic.eval_e_valI * by meson +next + case (eval_e_plusI i v1 n1 v2 n2) + then obtain v1' and v2' where *:"e = CE_op Plus v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_plusI by simp + hence "eval_e (i ( y \ s )) v1' (SNum n1) \ eval_e (i ( y \ s )) v2' (SNum n2)" using subst_v_eval_v eval_e_plusI + using "local.*" by blast + then show ?case using RCLogic.eval_e_plusI * by meson +next + case (eval_e_leqI i v1 n1 v2 n2) + then obtain v1' and v2' where *:"e = CE_op LEq v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_leqI by simp + hence "eval_e (i ( y \ s )) v1' (SNum n1) \ eval_e (i ( y \ s )) v2' (SNum n2)" using subst_v_eval_v eval_e_leqI + using * by blast + then show ?case using RCLogic.eval_e_leqI * by meson +next + case (eval_e_eqI i v1 n1 v2 n2) + then obtain v1' and v2' where *:"e = CE_op Eq v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n1 \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n2" using eval_e_eqI by simp + hence "eval_e (i ( y \ s )) v1' n1 \ eval_e (i ( y \ s )) v2' n2" using subst_v_eval_v eval_e_eqI + using * by blast + then show ?case using RCLogic.eval_e_eqI * by meson +next + case (eval_e_fstI i v1 s1 s2) + then obtain v1' and v2' where *:"e = CE_fst v1' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_fstI by simp + hence "eval_e (i ( y \ s )) v1' (SPair s1 s2)" using eval_e_fstI * by metis + then show ?case using RCLogic.eval_e_fstI * by meson +next + case (eval_e_sndI i v1 s1 s2) + then obtain v1' and v2' where *:"e = CE_snd v1' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_sndI by simp + hence "eval_e (i ( y \ s )) v1' (SPair s1 s2)" using subst_v_eval_v eval_e_sndI * by blast + then show ?case using RCLogic.eval_e_sndI * by meson +next + case (eval_e_concatI i v1 bv1 v2 bv2) + then obtain v1' and v2' where *:"e = CE_concat v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv1) \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv2)" using eval_e_concatI by simp + moreover hence "eval_e (i ( y \ s )) v1' (SBitvec bv1) \ eval_e (i ( y \ s )) v2' (SBitvec bv2)" + using subst_v_eval_v eval_e_concatI * by blast + ultimately show ?case using RCLogic.eval_e_concatI * eval_v_uniqueness by (metis eval_e_concatI.hyps(1)) +next + case (eval_e_lenI i v1 bv) + then obtain v1' where *:"e = CE_len v1' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv)" using eval_e_lenI by simp + hence "eval_e (i ( y \ s )) v1' (SBitvec bv)" using subst_v_eval_v eval_e_lenI * by blast + then show ?case using RCLogic.eval_e_lenI * by meson +qed + +lemma subst_c_eval_v[simp]: + fixes v::v and c :: c + assumes "i \ v \ ~ s" and "i \ c[x::=v]\<^sub>c\<^sub>v \ ~ s1" and + "(i ( x \ s)) \ c \ ~ s2" + shows "s1 = s2" + using assms proof(nominal_induct c arbitrary: s1 s2 rule: c.strong_induct) + case C_true + hence "s1 = True \ s2 = True" using eval_c_elims subst_cv.simps by auto + then show ?case by auto +next + case C_false + hence "s1 = False \ s2 = False" using eval_c_elims subst_cv.simps by metis + then show ?case by auto +next + case (C_conj c1 c2) + hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v AND c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 \ s12)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \ eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using + eval_c_elims(3) by metis + moreover obtain s21 and s22 where "eval_c (i ( x \ s)) c1 s21 \ eval_c (i ( x \ s)) c2 s22 \ (s2 = (s21 \ s22))" using + eval_c_elims(3) C_conj by metis + ultimately show ?case using C_conj by (meson eval_c_elims) +next + case (C_disj c1 c2) + hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v OR c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 \ s12)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \ eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using + eval_c_elims(4) by metis + moreover obtain s21 and s22 where "eval_c (i ( x \ s)) c1 s21 \ eval_c (i ( x \ s)) c2 s22 \ (s2 = (s21 \ s22))" using + eval_c_elims(4) C_disj by metis + ultimately show ?case using C_disj by (meson eval_c_elims) +next + case (C_not c1) + then obtain s11 where "(s1 = (\ s11)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11" using + eval_c_elims(6) by (metis subst_cv.simps(7)) + moreover obtain s21 where "eval_c (i ( x \ s)) c1 s21 \ (s2 = (\s21))" using + eval_c_elims(6) C_not by metis + ultimately show ?case using C_not by (meson eval_c_elims) +next + case (C_imp c1 c2) + hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v IMP c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 \ s12)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \ eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using + eval_c_elims(5) by metis + moreover obtain s21 and s22 where "eval_c (i ( x \ s)) c1 s21 \ eval_c (i ( x \ s)) c2 s22 \ (s2 = (s21 \ s22))" using + eval_c_elims(5) C_imp by metis + ultimately show ?case using C_imp by (meson eval_c_elims) +next + case (C_eq e1 e2) + hence *:"eval_c i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v == e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 = s12)) \ eval_e i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v) s11 \ eval_e i (e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s12" using + eval_c_elims(7) by metis + moreover obtain s21 and s22 where "eval_e (i ( x \ s)) e1 s21 \ eval_e (i ( x \ s)) e2 s22 \ (s2 = (s21 = s22))" using + eval_c_elims(7) C_eq by metis + ultimately show ?case using C_eq subst_e_eval_v by (metis eval_e_uniqueness) +qed + +lemma wfI_upd: + assumes "wfI \ \ i" and "wfRCV \ s b" and "wfG \ B ((x, b, c) #\<^sub>\ \)" + shows "wfI \ ((x, b, c) #\<^sub>\ \) (i(x \ s))" +proof(subst wfI_def,rule) + fix xa + assume as:"xa \ toSet ((x, b, c) #\<^sub>\ \)" + + then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps + using prod_cases3 by blast + + have "\sa. Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" proof(cases "x=x1") + case True + hence "b=b1" using as xa wfG_unique assms by metis + hence "Some s = (i(x \ s)) x1 \ wfRCV \ s b1" using assms True by simp + then show ?thesis by auto + next + case False + hence "(x1,b1,c1) \ toSet \" using xa as by auto + then obtain sa where "Some sa = i x1 \ wfRCV \ sa b1" using assms wfI_def as xa by auto + hence "Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" using False by auto + then show ?thesis by auto + qed + + thus "case xa of (xa, ba, ca) \ \sa. Some sa = (i(x \ s)) xa \ wfRCV \ sa ba" using xa by auto +qed + +lemma wfI_upd_full: + fixes v::v + assumes "wfI \ G i" and "G = ((\'[x::=v]\<^sub>\\<^sub>v)@\)" and "wfRCV \ s b" and "wfG \ B (\'@((x,b,c)#\<^sub>\\))" and "\ ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "wfI \ (\'@((x,b,c)#\<^sub>\\)) (i(x \ s))" +proof(subst wfI_def,rule) + fix xa + assume as:"xa \ toSet (\'@((x,b,c)#\<^sub>\\))" + + then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps + using prod_cases3 by blast + + have "\sa. Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" + proof(cases "x=x1") + case True + hence "b=b1" using as xa wfG_unique_full assms by metis + hence "Some s = (i(x \ s)) x1 \ wfRCV \ s b1" using assms True by simp + then show ?thesis by auto + next + case False + hence "(x1,b1,c1) \ toSet (\'@\)" using as xa by auto + then obtain c1' where "(x1,b1,c1') \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using xa as wfG_member_subst assms False by metis + then obtain sa where "Some sa = i x1 \ wfRCV \ sa b1" using assms wfI_def as xa by blast + hence "Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" using False by auto + then show ?thesis by auto + qed + + thus "case xa of (xa, ba, ca) \ \sa. Some sa = (i(x \ s)) xa \ wfRCV \ sa ba" using xa by auto +qed + +lemma subst_c_satis[simp]: + fixes v::v + assumes "i \ v \ ~ s" and "wfC \ B ((x,b,c')#\<^sub>\\) c" and "wfI \ \ i" and "\ ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "i \ (c[x::=v]\<^sub>c\<^sub>v) \ (i ( x \ s)) \ c" +proof - + have "wfI \ ((x, b, c') #\<^sub>\ \) (i(x \ s))" using wfI_upd assms wfC_wf eval_v_base by blast + then obtain s1 where s1:"eval_c (i(x \ s)) c s1" using eval_c_exist[of \ "((x,b,c')#\<^sub>\\)" "(i ( x \ s))" B c ] assms by auto + + have "\ ; B ; \ \\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wf_subst1(2)[OF assms(2) _ assms(4) , of GNil x ] subst_gv.simps by simp + then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \ "\" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto + + show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases + using eval_c_uniqueness is_satis.simps by auto +qed + +text \ Key theorem telling us we can replace a substitution with an update to the valuation \ +lemma subst_c_satis_full: + fixes v::v and \'::\ + assumes "i \ v \ ~ s" and "wfC \ B (\'@((x,b,c')#\<^sub>\\)) c" and "wfI \ ((\'[x::=v]\<^sub>\\<^sub>v)@\) i" and "\ ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "i \ (c[x::=v]\<^sub>c\<^sub>v) \ (i ( x \ s)) \ c" +proof - + have "wfI \ (\'@((x, b, c')) #\<^sub>\ \) (i(x \ s))" using wfI_upd_full assms wfC_wf eval_v_base wfI_suffix wfI_def wfV_wf by fast + then obtain s1 where s1:"eval_c (i(x \ s)) c s1" using eval_c_exist[of \ "(\'@(x,b,c')#\<^sub>\\)" "(i ( x \ s))" B c ] assms by auto + + have "\ ; B ; ((\'[x::=v]\<^sub>\\<^sub>v)@\) \\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wbc_subst assms by auto + + then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \ "((\'[x::=v]\<^sub>\\<^sub>v)@\)" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto + + show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases + using eval_c_uniqueness is_satis.simps by auto +qed + +section \Validity\ + +lemma validI[intro]: + fixes c::c + assumes "wfC P B G c" and "\i. P ; G \ i \ i \ G \ i \ c" + shows "P ; B ; G \ c" + using assms valid.simps by presburger + +lemma valid_g_wf: + fixes c::c and G::\ + assumes "P ; B ; G \ c" + shows "P ; B \\<^sub>w\<^sub>f G" + using assms wfC_wf valid.simps by blast + +lemma valid_reflI [intro]: + fixes b::b + assumes "P ; B ; ((x,b,c1)#\<^sub>\G) \\<^sub>w\<^sub>f c1" and "c1 = c2" + shows "P ; B ; ((x,b,c1)#\<^sub>\G) \ c2" + using satis_reflI assms by simp + +subsection \Weakening and Strengthening\ + +text \ Adding to the domain of a valuation doesn't change the result \ + +lemma eval_v_weakening: + fixes c::v and B::"bv fset" + assumes "i = i'|` d" and "supp c \ atom ` d \ supp B " and "i \ c \ ~ s" + shows "i' \ c \ ~ s" + using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct) + case (V_lit x) + then show ?case using eval_v_elims eval_v_litI by metis +next + case (V_var x) + have "atom x \ atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base + proof - + show ?thesis + by (metis UnE V_var.prems(2) \atom x \ supp B\ singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *) + qed + moreover have "Some s = i x" using assms eval_v_elims(2) + using V_var.prems(3) by blast + hence "Some s= i' x" using assms insert_subset restrict_in + proof - + show ?thesis + by (metis (no_types) \i = i' |` d\ \Some s = i x\ atom_eq_iff calculation imageE restrict_in) (* 31 ms *) + qed + thus ?case using eval_v.eval_v_varI by simp + +next + case (V_pair v1 v2) + then show ?case using eval_v_elims(3) eval_v_pairI v.supp + by (metis assms le_sup_iff) +next + case (V_cons dc v1) + then show ?case using eval_v_elims(4) eval_v_consI v.supp + by (metis assms le_sup_iff) +next + case (V_consp tyid dc b1 v1) + + then obtain sv1 where *:"i \ v1 \ ~ sv1 \ s = SConsp tyid dc b1 sv1" using eval_v_elims by metis + hence "i' \ v1 \ ~ sv1" using V_consp by auto + then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis +qed + +lemma eval_v_restrict: + fixes c::v and B::"bv fset" + assumes "i = i' |` d" and "supp c \ atom ` d \ supp B " and "i' \ c \ ~ s" + shows "i \ c \ ~ s" + using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct) + case (V_lit x) + then show ?case using eval_v_elims eval_v_litI by metis +next + case (V_var x) + have "atom x \ atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base + proof - + show ?thesis + by (metis UnE V_var.prems(2) \atom x \ supp B\ singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *) + qed + moreover have "Some s = i' x" using assms eval_v_elims(2) + using V_var.prems(3) by blast + hence "Some s= i x" using assms insert_subset restrict_in + proof - + show ?thesis + by (metis (no_types) \i = i' |` d\ \Some s = i' x\ atom_eq_iff calculation imageE restrict_in) (* 31 ms *) + qed + thus ?case using eval_v.eval_v_varI by simp +next + case (V_pair v1 v2) + then show ?case using eval_v_elims(3) eval_v_pairI v.supp + by (metis assms le_sup_iff) +next + case (V_cons dc v1) + then show ?case using eval_v_elims(4) eval_v_consI v.supp + by (metis assms le_sup_iff) +next + case (V_consp tyid dc b1 v1) + then obtain sv1 where *:"i' \ v1 \ ~ sv1 \ s = SConsp tyid dc b1 sv1" using eval_v_elims by metis + hence "i \ v1 \ ~ sv1" using V_consp by auto + then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis +qed + +lemma eval_e_weakening: + fixes e::ce and B::"bv fset" + assumes "i \ e \ ~ s" and "i = i' |` d" and "supp e \ atom ` d \ supp B " + shows "i' \ e \ ~ s" + using assms proof(induct rule: eval_e.induct) + case (eval_e_valI i v sv) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_plusI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_leqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_eqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_fstI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by metis +next + case (eval_e_sndI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by metis +next + case (eval_e_concatI i v1 bv2 v2 bv1) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_lenI i v bv) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +qed + +lemma eval_e_restrict : + fixes e::ce and B::"bv fset" + assumes "i' \ e \ ~ s" and "i = i' |` d" and "supp e \ atom ` d \ supp B " + shows "i \ e \ ~ s" + using assms proof(induct rule: eval_e.induct) + case (eval_e_valI i v sv) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_plusI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_leqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_eqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_fstI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by metis +next + case (eval_e_sndI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by metis +next + case (eval_e_concatI i v1 bv2 v2 bv1) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_lenI i v bv) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +qed + +lemma eval_c_i_weakening: + fixes c::c and B::"bv fset" + assumes "i \ c \ ~ s" and "i = i' |` d" and "supp c \ atom ` d \ supp B" + shows "i' \ c \ ~ s" + using assms proof(induct rule:eval_c.induct) + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using eval_c.intros eval_e_weakening by auto +qed(auto simp add: eval_c.intros)+ + +lemma eval_c_i_restrict: + fixes c::c and B::"bv fset" + assumes "i' \ c \ ~ s" and "i = i' |` d" and "supp c \ atom ` d \ supp B" + shows "i \ c \ ~ s" + using assms proof(induct rule:eval_c.induct) + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using eval_c.intros eval_e_restrict by auto +qed(auto simp add: eval_c.intros)+ + +lemma is_satis_i_weakening: + fixes c::c and B::"bv fset" + assumes "i = i' |` d" and "supp c \ atom ` d \ supp B " and "i \ c" + shows "i' \ c" + using is_satis.simps eval_c_i_weakening[OF _ assms(1) assms(2) ] + using assms(3) by auto + +lemma is_satis_i_restrict: + fixes c::c and B::"bv fset" + assumes "i = i' |` d" and "supp c \ atom ` d \ supp B" and "i' \ c" + shows "i \ c" + using is_satis.simps eval_c_i_restrict[OF _ assms(1) assms(2) ] + using assms(3) by auto + +lemma is_satis_g_restrict1: + fixes \'::\ and \::\ + assumes "toSet \ \ toSet \'" and "i \ \'" + shows "i \ \" + using assms proof(induct \ rule: \.induct) + case GNil + then show ?case by auto +next + case (GCons xbc G) + obtain x and b and c::c where xbc: "xbc=(x,b,c)" + using prod_cases3 by blast + hence "i \ G" using GCons by auto + moreover have "i \ c" using GCons + is_satis_iff toSet.simps subset_iff + using xbc by blast + ultimately show ?case using is_satis_g.simps GCons + by (simp add: xbc) +qed + +lemma is_satis_g_restrict2: + fixes \'::\ and \::\ + assumes "i \ \" and "i' = i |` d" and "atom_dom \ \ atom ` d" and "\ ; B \\<^sub>w\<^sub>f \" + shows "i' \ \" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x b c G) + + hence "i' \ G" proof - + have "i \ G" using GCons by simp + moreover have "atom_dom G \ atom ` d" using GCons by simp + ultimately show ?thesis using GCons wfG_cons2 by blast + qed + + moreover have "i' \ c" proof - + have "i \ c" using GCons by auto + moreover have "\ ; B ; (x, b, TRUE) #\<^sub>\ G \\<^sub>w\<^sub>f c" using wfG_wfC GCons by simp + moreover hence "supp c \ atom ` d \ supp B" using wfC_supp GCons atom_dom_eq by blast + ultimately show ?thesis using is_satis_i_restrict[of i' i d c] GCons by simp + qed + + ultimately show ?case by auto +qed + +lemma is_satis_g_restrict: + fixes \'::\ and \::\ + assumes "toSet \ \ toSet \'" and "i' \ \'" and "i = i' |` (fst ` toSet \)" and "\ ; B \\<^sub>w\<^sub>f \" + shows "i \ \" + using assms is_satis_g_restrict1[OF assms(1) assms(2)] is_satis_g_restrict2[OF _ assms(3)] by simp + +subsection \Updating valuation\ + +lemma is_satis_c_i_upd: + fixes c::c + assumes "atom x \ c" and "i \ c" + shows "((i ( x \s))) \ c" + using assms eval_c_i_upd is_satis.simps by simp + +lemma is_satis_g_i_upd: + fixes G::\ + assumes "atom x \ G" and "i \ G" + shows "((i ( x \s))) \ G" + using assms proof(induct G rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' G') + + hence *:"atom x \ G' \ atom x \ c'" + using fresh_def fresh_GCons GCons by force + moreover hence "is_satis ((i ( x \s))) c'" + using is_satis_c_i_upd GCons is_satis_g.simps by auto + moreover have " is_satis_g (i(x \ s)) G'" using GCons * by fastforce + ultimately show ?case + using GCons is_satis_g.simps(2) by metis +qed + +lemma valid_weakening: + assumes "\ ; B ; \ \ c" and "toSet \ \ toSet \'" and "wfG \ B \'" + shows "\ ; B ; \' \ c" +proof - + have "wfC \ B \ c" using assms valid.simps by auto + hence sp: "supp c \ atom `(fst `toSet \) \ supp B" using wfX_wfY wfG_elims + using atom_dom.simps dom.simps wf_supp(2) by metis + have wfg: "wfG \ B \" using assms valid.simps wfC_wf by auto + + moreover have a1: "(\i. wfI \ \' i \ i \ \' \ i \ c)" proof(rule allI, rule impI) + fix i + assume as: "wfI \ \' i \ i \ \'" + hence as1: "fst ` toSet \ \ dom i" using assms wfI_domi by blast + obtain i' where idash: "i' = restrict_map i (fst `toSet \)" by blast + hence as2: "dom i' = (fst `toSet \)" using dom_restrict as1 by auto + + have id2: "\ ; \ \ i' \ i' \ \" proof - + have "wfI \ \ i'" using as2 wfI_restrict_weakening[of \ \' i i' \] as assms + using idash by blast + moreover have "i' \ \" using is_satis_g_restrict[OF assms(2)] wfg as idash by auto + ultimately show ?thesis using idash by auto + qed + hence "i' \ c" using assms valid.simps by auto + thus "i \ c" using assms valid.simps is_satis_i_weakening idash sp by blast + qed + moreover have "wfC \ B \' c" using wf_weakening assms valid.simps + by (meson wfg) + ultimately show ?thesis using assms valid.simps by auto +qed + +lemma is_satis_g_suffix: + fixes G::\ + assumes "i \ (G'@G)" + shows "i \ G" + using assms proof(induct G' rule:\.induct) + case GNil + then show ?case by auto +next + case (GCons xbc x2) + obtain x and b and c::c where xbc: "xbc=(x,b,c)" + using prod_cases3 by blast + hence " i \ (xbc #\<^sub>\ (x2 @ G))" using append_g.simps GCons by fastforce + then show ?case using is_satis_g.simps GCons xbc by blast +qed + +lemma wfG_inside_valid2: + fixes x::x and \::\ and c0::c and c0'::c + assumes "wfG \ B (\'@((x,b0,c0')#\<^sub>\\))" and + "\ ; B ; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "wfG \ B (\'@((x,b0,c0)#\<^sub>\\))" +proof - + have "wfG \ B (\'@(x,b0,c0)#\<^sub>\\)" using valid.simps wfC_wf assms by auto + thus ?thesis using wfG_replace_inside_full assms by auto +qed + +lemma valid_trans: + assumes " \ ; \ ; \ \ c0[z::=v]\<^sub>v" and " \ ; \ ; (z,b,c0)#\<^sub>\\ \ c1" and "atom z \ \" and "wfV \ \ \ v b" + shows " \ ; \ ; \ \ c1[z::=v]\<^sub>v" +proof - + have *:"wfC \ \ ((z,b,c0)#\<^sub>\\) c1" using valid.simps assms by auto + hence "wfC \ \ \ (c1[z::=v]\<^sub>v)" using wf_subst1(2)[OF * , of GNil ] assms subst_gv.simps subst_v_c_def by force + + moreover have "\i. wfI \ \ i \ is_satis_g i \ \ is_satis i (c1[z::=v]\<^sub>v)" + proof(rule,rule) + fix i + assume as: "wfI \ \ i \ is_satis_g i \" + then obtain sv where sv: "eval_v i v sv \ wfRCV \ sv b" using eval_v_exist assms by metis + hence "is_satis i (c0[z::=v]\<^sub>v)" using assms valid.simps as by metis + hence "is_satis (i(z \ sv)) c0" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis + moreover have "is_satis_g (i(z \ sv)) \" + using is_satis_g_i_upd assms by (simp add: as) + ultimately have "is_satis_g (i(z \ sv)) ((z,b,c0)#\<^sub>\\)" + using is_satis_g.simps by simp + moreover have "wfI \ ((z,b,c0)#\<^sub>\\) (i(z \ sv))" using as wfI_upd sv assms valid.simps wfC_wf by metis + ultimately have "is_satis (i(z \ sv)) c1" using assms valid.simps by auto + thus "is_satis i (c1[z::=v]\<^sub>v)" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis + qed + + ultimately show ?thesis using valid.simps by auto +qed + +lemma valid_trans_full: + assumes "\ ; \ ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \) \ c2[z2::=V_var x]\<^sub>v" and + "\ ; \ ; ((x, b, c2[z2::=V_var x]\<^sub>v) #\<^sub>\ \) \ c3[z3::=V_var x]\<^sub>v" + shows "\ ; \ ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \) \ c3[z3::=V_var x]\<^sub>v" + unfolding valid.simps proof + show "\ ; \ ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c3[z3::=V_var x]\<^sub>v" using wf_trans valid.simps assms by metis + + show "\i. ( wfI \ ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \) i \ (is_satis_g i ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \)) \ (is_satis i (c3[z3::=V_var x]\<^sub>v)) ) " + proof(rule,rule) + fix i + assume as: "\ ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \ i \ i \ (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \" + have "i \ c2[z2::=V_var x]\<^sub>v" using is_satis_g.simps as assms by simp + moreover have "i \ \" using is_satis_g.simps as by simp + ultimately show "i \ c3[z3::=V_var x]\<^sub>v " using assms is_satis_g.simps valid.simps + by (metis append_g.simps(1) as wfI_replace_inside) + qed +qed + +lemma eval_v_weakening_x: + fixes c::v + assumes "i' \ c \ ~ s" and "atom x \ c" and "i = i' (x \ s')" + shows "i \ c \ ~ s" + using assms proof(induct rule: eval_v.induct) + case (eval_v_litI i l) + then show ?case using eval_v.intros by auto +next + case (eval_v_varI sv i1 x1) + hence "x \ x1" using v.fresh fresh_at_base by auto + hence "i x1 = Some sv" using eval_v_varI by simp + then show ?case using eval_v.intros by auto +next + case (eval_v_pairI i v1 s1 v2 s2) + then show ?case using eval_v.intros by auto +next + case (eval_v_consI i v s tyid dc) + then show ?case using eval_v.intros by auto +next + case (eval_v_conspI i v s tyid dc b) + then show ?case using eval_v.intros by auto +qed + +lemma eval_e_weakening_x: + fixes c::ce + assumes "i' \ c \ ~ s" and "atom x \ c" and "i = i' (x \ s')" + shows "i \ c \ ~ s" + using assms proof(induct rule: eval_e.induct) + case (eval_e_valI i v sv) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_plusI i v1 n1 v2 n2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_leqI i v1 n1 v2 n2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_eqI i v1 n1 v2 n2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_fstI i v v1 v2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_sndI i v v1 v2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_concatI i v1 bv1 v2 bv2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_lenI i v bv) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +qed + +lemma eval_c_weakening_x: + fixes c::c + assumes "i' \ c \ ~ s" and "atom x \ c" and "i = i' (x \ s')" + shows "i \ c \ ~ s" + using assms proof(induct rule: eval_c.induct) + case (eval_c_trueI i) + then show ?case using eval_c.intros by auto +next + case (eval_c_falseI i) + then show ?case using eval_c.intros by auto +next + case (eval_c_conjI i c1 b1 c2 b2) + then show ?case using eval_c.intros by auto +next + case (eval_c_disjI i c1 b1 c2 b2) + then show ?case using eval_c.intros by auto +next + case (eval_c_impI i c1 b1 c2 b2) + then show ?case using eval_c.intros by auto +next + case (eval_c_notI i c b) + then show ?case using eval_c.intros by auto +next + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using eval_e_weakening_x c.fresh eval_c.intros by metis +qed + +lemma is_satis_weakening_x: + fixes c::c + assumes "i' \ c" and "atom x \ c" and "i = i' (x \ s)" + shows "i \ c" + using eval_c_weakening_x assms is_satis.simps by simp + +lemma is_satis_g_weakening_x: + fixes G::\ + assumes "i' \ G" and "atom x \ G" and "i = i' (x \ s)" + shows "i \ G" + using assms proof(induct G rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + hence "atom x \ c'" using fresh_GCons fresh_prodN by simp + moreover hence "i \ c'" using is_satis_weakening_x is_satis_g.simps(2) GCons by metis + then show ?case using is_satis_g.simps(2)[of i x' b' c' \'] GCons fresh_GCons by simp +qed + +section \Base Type Substitution\ + +text \The idea of boxing is to take an smt val and its base type and at nodes in the smt val that correspond to type variables we +wrap them in an SUt smt val node. Another way of looking at it is that s' where the node for the base type variable is an 'any node'. +It is needed to prove subst\_b\_valid - the base-type variable substitution lemma for validity. + +The first @{text "rcl_val"} is the expanded form (has type with base-variables replaced with base-type terms) + ; the second is its corresponding form + +We only have one variable so we need to ensure that in all of the @{text "bs_boxed_BVarI"} cases, the s has the same + base type. + +For example is an SMT value is (SPair (SInt 1) (SBool true)) and it has sort (BPair (BVar x) BBool)[x::=BInt] then +the boxed version is SPair (SUt (SInt 1)) (SBool true) and is has sort (BPair (BVar x) BBool). We need to do this +so that we can obtain from a valuation i, that gives values like the first smt value, to a valuation i' that gives values like +the second. +\ + +inductive boxed_b :: "\ \ rcl_val \ b \ bv \ b \ rcl_val \ bool" ( " _ \ _ ~ _ [ _ ::= _ ] \ _ " [50,50] 50) where + boxed_b_BVar1I: "\ bv = bv' ; wfRCV P s b \ \ boxed_b P s (B_var bv') bv b (SUt s)" +| boxed_b_BVar2I: "\ bv \ bv'; wfRCV P s (B_var bv') \ \ boxed_b P s (B_var bv') bv b s" +| boxed_b_BIntI:"wfRCV P s B_int \ boxed_b P s B_int _ _ s" +| boxed_b_BBoolI:"wfRCV P s B_bool \ boxed_b P s B_bool _ _ s " +| boxed_b_BUnitI:"wfRCV P s B_unit \ boxed_b P s B_unit _ _ s" +| boxed_b_BPairI:"\ boxed_b P s1 b1 bv b s1' ; boxed_b P s2 b2 bv b s2' \ \ boxed_b P (SPair s1 s2) (B_pair b1 b2) bv b (SPair s1' s2')" + +| boxed_b_BConsI:"\ + AF_typedef tyid dclist \ set P; + (dc, \ x : b | c \) \ set dclist ; + boxed_b P s1 b bv b' s1' + \ \ + boxed_b P (SCons tyid dc s1) (B_id tyid) bv b' (SCons tyid dc s1')" + +| boxed_b_BConspI:"\ AF_typedef_poly tyid bva dclist \ set P; + atom bva \ (b1,bv,b',s1,s1'); + (dc, \ x : b | c \) \ set dclist ; + boxed_b P s1 (b[bva::=b1]\<^sub>b\<^sub>b) bv b' s1' + \ \ + boxed_b P (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1) (B_app tyid b1) bv b' (SConsp tyid dc b1 s1')" + +| boxed_b_Bbitvec: "wfRCV P s B_bitvec \ boxed_b P s B_bitvec bv b s" + +equivariance boxed_b +nominal_inductive boxed_b . + +inductive_cases boxed_b_elims: + "boxed_b P s (B_var bv) bv' b s'" + "boxed_b P s B_int bv b s'" + "boxed_b P s B_bool bv b s'" + "boxed_b P s B_unit bv b s'" + "boxed_b P s (B_pair b1 b2) bv b s'" + "boxed_b P s (B_id dc) bv b s'" + "boxed_b P s B_bitvec bv b s'" + "boxed_b P s (B_app dc b') bv b s'" + +lemma boxed_b_wfRCV: + assumes "boxed_b P s b bv b' s'" (*and "supp s = {}"*) and "\\<^sub>w\<^sub>f P" + shows "wfRCV P s b[bv::=b']\<^sub>b\<^sub>b \ wfRCV P s' b" + using assms proof(induct rule: boxed_b.inducts) + case (boxed_b_BVar1I bv bv' P s b ) + then show ?case using wfRCV.intros by auto +next + case (boxed_b_BVar2I bv bv' P s ) + then show ?case using wfRCV.intros by auto +next + case (boxed_b_BPairI P s1 b1 bv b s1' s2 b2 s2') + then show ?case using wfRCV.intros rcl_val.supp by simp +next + case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1') + hence "supp b = {}" using wfTh_supp_b by metis + hence "b [ bv ::= b' ]\<^sub>b\<^sub>b = b" using fresh_def subst_b_b_def forget_subst[of "bv" b b'] by auto + hence " P \ SCons tyid dc s1 : (B_id tyid)" using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis + moreover have "P \ SCons tyid dc s1' : B_id tyid" using boxed_b_BConsI + using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis + ultimately show ?case using subst_bb.simps by metis +next + case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c) + + obtain bva2 and dclist2 where *:"AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \ + atom bva2 \ (bv,(P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b))" + using obtain_fresh_bv by metis + + then obtain x2 and b2 and c2 where **:\(dc, \ x2 : b2 | c2 \) \ set dclist2\ + using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff by metis + + have "P \ SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1 : (B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" proof + show 1: \AF_typedef_poly tyid bva2 dclist2 \ set P\ using boxed_b_BConspI * by auto + show 2: \(dc, \ x2 : b2 | c2 \) \ set dclist2\ using boxed_b_BConspI using ** by simp + + hence "atom bv \ b2" proof - + have "supp b2 \ { atom bva2 }" using wfTh_poly_supp_b 1 2 boxed_b_BConspI by auto + moreover have "bv \ bva2" using * fresh_Pair fresh_at_base by metis + ultimately show ?thesis using fresh_def by force + qed + moreover have "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff * 2 boxed_b_BConspI by metis + ultimately show \ P \ s1 : b2[bva2::=b1[bv::=b']\<^sub>b\<^sub>b]\<^sub>b\<^sub>b\ using boxed_b_BConspI subst_b_b_def subst_bb_commute by auto + show "atom bva2 \ (P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" using * fresh_Pair by metis + qed + + moreover have "P \ SConsp tyid dc b1 s1' : B_app tyid b1" proof + show \AF_typedef_poly tyid bva dclist \ set P\ using boxed_b_BConspI by auto + show \(dc, \ x : b | c \) \ set dclist\ using boxed_b_BConspI by auto + show \ P \ s1' : b[bva::=b1]\<^sub>b\<^sub>b\ using boxed_b_BConspI by auto + have "atom bva \ P" using boxed_b_BConspI wfTh_fresh by metis + thus "atom bva \ (P, SConsp tyid dc b1 s1', B_app tyid b1)" using boxed_b_BConspI rcl_val.fresh b.fresh pure_fresh fresh_prodN by metis + qed + + ultimately show ?case using subst_bb.simps by simp +qed(auto)+ + +lemma subst_b_var: + assumes "B_var bv2 = b[bv::=b']\<^sub>b\<^sub>b" + shows "(b = B_var bv \ b' = B_var bv2) \ (b=B_var bv2 \ bv \ bv2)" + using assms by(nominal_induct b rule: b.strong_induct,auto+) + +text \Here the valuation i' is the conv wrap version of i. For every x in G, i' x is the conv wrap version of i x. +The next lemma for a clearer explanation of what this is. i produces values of sort b[bv::=b'] and i' produces values of sort b\ + +inductive boxed_i :: "\ \ \ \ b \ bv \ valuation \ valuation \ bool" ( " _ ; _ ; _ , _ \ _ \ _" [50,50] 50) where + boxed_i_GNilI: "\ ; GNil ; b , bv \ i \ i" +| boxed_i_GConsI: "\ Some s = i x; boxed_b \ s b bv b' s' ; \ ; \ ; b' , bv \ i \ i' \ \ \ ; ((x,b,c)#\<^sub>\\) ; b' , bv \ i \ (i'(x \ s'))" +equivariance boxed_i +nominal_inductive boxed_i . + +inductive_cases boxed_i_elims: + "\ ;GNil ; b , bv \ i \ i'" + "\ ; ((x,b,c)#\<^sub>\\) ; b' , bv \ i \ i'" + +lemma wfRCV_poly_elims: + fixes tm::"'a::fs" and b::b + assumes "T \ SConsp typid dc bdc s : b" + obtains bva dclist x1 b1 c1 where "b = B_app typid bdc \ + AF_typedef_poly typid bva dclist \ set T \ (dc, \ x1 : b1 | c1 \) \ set dclist \ T \ s : b1[bva::=bdc]\<^sub>b\<^sub>b \ atom bva \ tm" + using assms proof(nominal_induct "SConsp typid dc bdc s" b avoiding: tm rule:wfRCV.strong_induct) + case (wfRCV_BConsPI bv dclist \ x b c) + then show ?case by simp +qed + +lemma boxed_b_ex: + assumes "wfRCV T s b[bv::=b']\<^sub>b\<^sub>b" and "wfTh T" + shows "\s'. boxed_b T s b bv b' s'" + using assms proof(nominal_induct s arbitrary: b rule: rcl_val.strong_induct) + case (SBitvec x) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bitvec" using wfRCV_elims(9)[OF SBitvec(1)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SBitvec x : B_bitvec" using wfRCV.intros by simp + moreover hence "b' = B_bitvec" using True SBitvec subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + hence "b = B_bitvec" using subst_bb_inject * by metis + then show ?thesis using * SBitvec boxed_b.intros by metis + qed +next + case (SNum x) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_int" using wfRCV_elims(10)[OF SNum(1)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SNum x : B_int" using wfRCV.intros by simp + moreover hence "b' = B_int" using True SNum subst_bb.simps(1) * by simp + ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis + next + case False + hence "b = B_int" using subst_bb_inject(1) * by metis + then show ?thesis using * SNum boxed_b_BIntI by metis + qed +next + case (SBool x) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bool" using wfRCV_elims(11)[OF SBool(1)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SBool x : B_bool" using wfRCV.intros by simp + moreover hence "b' = B_bool" using True SBool subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + hence "b = B_bool" using subst_bb_inject * by metis + then show ?thesis using * SBool boxed_b.intros by metis + qed +next + case (SPair s1 s2) + then obtain b1 and b2 where *:"b[bv::=b']\<^sub>b\<^sub>b = B_pair b1 b2 \ wfRCV T s1 b1 \ wfRCV T s2 b2" using wfRCV_elims(12) by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SPair s1 s2 : B_pair b1 b2" using wfRCV.intros * by simp + moreover hence "b' = B_pair b1 b2" using True SPair subst_bb.simps(1) * by simp + ultimately show ?thesis using boxed_b_BVar1I by metis + next + case False + then obtain b1' and b2' where "b = B_pair b1' b2' \ b1=b1'[bv::=b']\<^sub>b\<^sub>b \ b2=b2'[bv::=b']\<^sub>b\<^sub>b" using subst_bb_inject(5)[OF _ False ] * by metis + then show ?thesis using * SPair boxed_b_BPairI by blast + qed +next + case (SCons tyid dc s1) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_id tyid" using wfRCV_elims(13)[OF SCons(2)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SCons tyid dc s1 : B_id tyid" using wfRCV.intros + using "local.*" SCons.prems by auto + moreover hence "b' = B_id tyid" using True SCons subst_bb.simps(1) * by simp + ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis + next + case False + then obtain b1' where beq: "b = B_id tyid" using subst_bb_inject * by metis + then obtain b2 dclist x c where **:"AF_typedef tyid dclist \ set T \ (dc, \ x : b2 | c \) \ set dclist \ wfRCV T s1 b2" using wfRCV_elims(13) * SCons by metis + then have "atom bv \ b2" using \wfTh T\ wfTh_lookup_supp_empty[of tyid dclist T dc "\ x : b2 | c \"] \.fresh fresh_def by auto + then have "b2 = b2[ bv ::= b']\<^sub>b\<^sub>b" using forget_subst subst_b_b_def by metis + then obtain s1' where s1:"T \ s1 ~ b2 [ bv ::= b' ] \ s1'" using SCons ** by metis + + have "T \ SCons tyid dc s1 ~ (B_id tyid) [ bv ::= b' ] \ SCons tyid dc s1'" proof(rule boxed_b_BConsI) + show "AF_typedef tyid dclist \ set T" using ** by auto + show "(dc, \ x : b2 | c \) \ set dclist" using ** by auto + show "T \ s1 ~ b2 [ bv ::= b' ] \ s1' " using s1 ** by auto + + qed + thus ?thesis using beq by metis + qed +next + case (SConsp typid dc bdc s) + + obtain bva dclist x1 b1 c1 where **:"b[bv::=b']\<^sub>b\<^sub>b = B_app typid bdc \ + AF_typedef_poly typid bva dclist \ set T \ (dc, \ x1 : b1 | c1 \) \ set dclist \ T \ s : b1[bva::=bdc]\<^sub>b\<^sub>b \ atom bva \ bv" + using wfRCV_poly_elims[OF SConsp(2)] by metis + + then have *:"B_app typid bdc = b[bv::=b']\<^sub>b\<^sub>b" using wfRCV_elims(14)[OF SConsp(2)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SConsp typid dc bdc s : B_app typid bdc" using wfRCV.intros + using "local.*" SConsp.prems(1) by auto + moreover hence "b' = B_app typid bdc" using True SConsp subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + then obtain bdc' where bdc: "b = B_app typid bdc' \ bdc = bdc'[bv::=b']\<^sub>b\<^sub>b" using * subst_bb_inject(8)[OF *] by metis + (*hence beq:"b = B_app typid bdc" using subst_bb_inject * sory (* going to be like the above as bdc is closed *)*) + have "atom bv \ b1" proof - + have "supp b1 \ { atom bva }" using wfTh_poly_supp_b ** SConsp by metis + moreover have "bv \ bva" using ** by auto + ultimately show ?thesis using fresh_def by force + qed + have "T \ s : b1[bva::=bdc]\<^sub>b\<^sub>b" using ** by auto + moreover have "b1[bva::=bdc']\<^sub>b\<^sub>b[bv::=b']\<^sub>b\<^sub>b = b1[bva::=bdc]\<^sub>b\<^sub>b" using bdc subst_bb_commute \atom bv \ b1\ by auto + ultimately obtain s' where s':"T \ s ~ b1[bva::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \ s'" + using SConsp(1)[of "b1[bva::=bdc']\<^sub>b\<^sub>b"] bdc SConsp by metis + have "T \ SConsp typid dc bdc'[bv::=b']\<^sub>b\<^sub>b s ~ (B_app typid bdc') [ bv ::= b' ] \ SConsp typid dc bdc' s'" + proof - + + obtain bva3 and dclist3 where 3:"AF_typedef_poly typid bva3 dclist3 = AF_typedef_poly typid bva dclist \ + atom bva3 \ (bdc', bv, b', s, s')" using obtain_fresh_bv by metis + then obtain x3 b3 c3 where 4:"(dc, \ x3 : b3 | c3 \) \ set dclist3" + using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff + by (metis "**") + + show ?thesis proof + show \AF_typedef_poly typid bva3 dclist3 \ set T\ using 3 ** by metis + show \atom bva3 \ (bdc', bv, b', s, s')\ using 3 by metis + show 4:\(dc, \ x3 : b3 | c3 \) \ set dclist3\ using 4 by auto + have "b3[bva3::=bdc']\<^sub>b\<^sub>b = b1[bva::=bdc']\<^sub>b\<^sub>b" proof(rule wfTh_typedef_poly_b_eq_iff) + show \AF_typedef_poly typid bva3 dclist3 \ set T\ using 3 ** by metis + show \(dc, \ x3 : b3 | c3 \) \ set dclist3\ using 4 by auto + show \AF_typedef_poly typid bva dclist \ set T\ using ** by auto + show \(dc, \ x1 : b1 | c1 \) \ set dclist\ using ** by auto + qed(simp add: ** SConsp) + thus \ T \ s ~ b3[bva3::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \ s' \ using s' by auto + qed + qed + then show ?thesis using bdc by auto + + qed +next + case SUnit + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_unit" using wfRCV_elims SUnit by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SUnit : B_unit" using wfRCV.intros by simp + moreover hence "b' = B_unit" using True SUnit subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + hence "b = B_unit" using subst_bb_inject * by metis + then show ?thesis using * SUnit boxed_b.intros by metis + qed +next + case (SUt x) + then obtain bv' where *:"b[bv::=b']\<^sub>b\<^sub>b= B_var bv'" using wfRCV_elims by metis + show ?case proof (cases "b = B_var bv") + case True + then show ?thesis using boxed_b_BVar1I + using "local.*" wfRCV_BVarI by fastforce + next + case False + then show ?thesis using boxed_b_BVar1I boxed_b_BVar2I + using "local.*" wfRCV_BVarI by (metis subst_b_var) + qed +qed + +lemma boxed_i_ex: + assumes "wfI T \[bv::=b]\<^sub>\\<^sub>b i" and "wfTh T" + shows "\i'. T ; \ ; b , bv \ i \ i'" + using assms proof(induct \ arbitrary: i rule:\_induct) + case GNil + then show ?case using boxed_i_GNilI by metis +next + case (GCons x' b' c' \') + then obtain s where 1:"Some s = i x' \ wfRCV T s b'[bv::=b]\<^sub>b\<^sub>b" using wfI_def subst_gb.simps by auto + then obtain s' where 2: "boxed_b T s b' bv b s'" using boxed_b_ex GCons by metis + then obtain i' where 3: "boxed_i T \' b bv i i'" using GCons wfI_def subst_gb.simps by force + have "boxed_i T ((x', b', c') #\<^sub>\ \') b bv i (i'(x' \ s'))" proof + show "Some s = i x'" using 1 by auto + show "boxed_b T s b' bv b s'" using 2 by auto + show "T ; \' ; b , bv \ i \ i'" using "3" by auto + qed + thus ?case by auto +qed + +lemma boxed_b_eq: + assumes "boxed_b \ s1 b bv b' s1'" and "\\<^sub>w\<^sub>f \" + shows "wfTh \ \ boxed_b \ s2 b bv b' s2' \ ( s1 = s2 ) = ( s1' = s2' )" + using assms proof(induct arbitrary: s2 s2' rule: boxed_b.inducts ) + case (boxed_b_BVar1I bv bv' P s b ) + then show ?case + using boxed_b_elims(1) rcl_val.eq_iff by metis +next + case (boxed_b_BVar2I bv bv' P s b) + then show ?case using boxed_b_elims(1) by metis +next + case (boxed_b_BIntI P s uu uv) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BBoolI P s uw ux) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BUnitI P s uy uz) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BPairI P s1 b1 bv b s1' s2a b2 s2a') + then show ?case + by (metis boxed_b_elims(5) rcl_val.eq_iff(4)) +next + case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1') + obtain s22 and s22' dclist2 dc2 x2 b2 c2 where *:"s2 = SCons tyid dc2 s22 \ s2' = SCons tyid dc2 s22' \ boxed_b P s22 b2 bv b' s22' + \ AF_typedef tyid dclist2 \ set P \ (dc2, \ x2 : b2 | c2 \) \ set dclist2" using boxed_b_elims(6)[OF boxed_b_BConsI(6)] by metis + show ?case proof(cases "dc = dc2") + case True + hence "b = b2" using wfTh_ctor_unique \.eq_iff wfTh_dclist_unique wf boxed_b_BConsI * by metis + then show ?thesis using boxed_b_BConsI True * by auto + next + case False + then show ?thesis using * boxed_b_BConsI by simp + qed +next + case (boxed_b_Bbitvec P s bv b) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c) + obtain bva2 s22 s22' dclist2 dc2 x2 b2 c2 where *:" + s2 = SConsp tyid dc2 b1[bv::=b']\<^sub>b\<^sub>b s22 \ + s2' = SConsp tyid dc2 b1 s22' \ + boxed_b P s22 b2[bva2::=b1]\<^sub>b\<^sub>b bv b' s22' \ + AF_typedef_poly tyid bva2 dclist2 \ set P \ (dc2, \ x2 : b2 | c2 \) \ set dclist2" using boxed_b_elims(8)[OF boxed_b_BConspI(7)] by metis + show ?case proof(cases "dc = dc2") + case True + hence "AF_typedef_poly tyid bva2 dclist2 \ set P \ (dc, \ x2 : b2 | c2 \) \ set dclist2" using * by auto + hence "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff[OF boxed_b_BConspI(1) boxed_b_BConspI(3)] * boxed_b_BConspI by metis + then show ?thesis using boxed_b_BConspI True * by auto + next + case False + then show ?thesis using * boxed_b_BConspI by simp + qed +qed + +lemma bs_boxed_var: + assumes "boxed_i \ \ b' bv i i'" + shows "Some (b,c) = lookup \ x \ Some s = i x \ Some s' = i' x \ boxed_b \ s b bv b' s'" + using assms proof(induct rule: boxed_i.inducts) + case (boxed_i_GNilI T i) + then show ?case using lookup.simps by auto +next + case (boxed_i_GConsI s i x1 \ b1 bv b' s' \ i' c) + show ?case proof (cases "x=x1") + case True + then show ?thesis using boxed_i_GConsI + fun_upd_same lookup.simps(2) option.inject prod.inject by metis + next + case False + then show ?thesis using boxed_i_GConsI + fun_upd_same lookup.simps option.inject prod.inject by auto + qed +qed + +lemma eval_l_boxed_b: + assumes "\ l \ = s" + shows "boxed_b \ s (base_for_lit l) bv b' s" + using assms proof(nominal_induct l arbitrary: s rule:l.strong_induct) +qed(auto simp add: boxed_b.intros wfRCV.intros )+ + +lemma boxed_i_eval_v_boxed_b: + fixes v::v + assumes "boxed_i \ \ b' bv i i'" and "i \ v[bv::=b']\<^sub>v\<^sub>b \ ~ s" and "i' \ v \ ~ s'" and "wfV \ B \ v b" and "wfI \ \ i'" + shows "boxed_b \ s b bv b' s'" + using assms proof(nominal_induct v arbitrary: s s' b rule:v.strong_induct) + case (V_lit l) + hence "\ l \ = s \ \ l \ = s'" using eval_v_elims by auto + moreover have "b = base_for_lit l" using wfV_elims(2) V_lit by metis + ultimately show ?case using V_lit using eval_l_boxed_b subst_b_base_for_lit by metis +next + case (V_var x) (* look at defn of bs_boxed *) + hence "Some s = i x \ Some s' = i' x" using eval_v_elims subst_vb.simps by metis + moreover obtain c1 where bc:"Some (b,c1) = lookup \ x" using wfV_elims V_var by metis + ultimately show ?case using bs_boxed_var V_var by metis + +next + case (V_pair v1 v2) + then obtain b1 and b2 where b:"b=B_pair b1 b2" using wfV_elims subst_vb.simps by metis + obtain s1 and s2 where s: "eval_v i (v1[bv::=b']\<^sub>v\<^sub>b) s1 \ eval_v i (v2[bv::=b']\<^sub>v\<^sub>b) s2 \ s = SPair s1 s2" using eval_v_elims V_pair subst_vb.simps by metis + obtain s1' and s2' where s': "eval_v i' v1 s1' \ eval_v i' v2 s2' \ s' = SPair s1' s2'" using eval_v_elims V_pair by metis + have "boxed_b \ (SPair s1 s2) (B_pair b1 b2) bv b' (SPair s1' s2') " proof(rule boxed_b_BPairI) + show "boxed_b \ s1 b1 bv b' s1'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis + show "boxed_b \ s2 b2 bv b' s2'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis + qed + then show ?case using s s' b by auto +next + case (V_cons tyid dc v1) + + obtain dclist x b1 c where *: "b = B_id tyid \ AF_typedef tyid dclist \ set \ \ (dc, \ x : b1 | c \) \ set dclist \ \ ; B ; \ \\<^sub>w\<^sub>f v1 : b1" + using wfV_elims(4)[OF V_cons(5)] V_cons by metis + obtain s2 where s2: "s = SCons tyid dc s2 \ i \ (v1[bv::=b']\<^sub>v\<^sub>b) \ ~ s2" using eval_v_elims V_cons subst_vb.simps by metis + obtain s2' where s2': "s' = SCons tyid dc s2' \ i' \ v1 \ ~ s2'" using eval_v_elims V_cons by metis + have sp: "supp \ x : b1 | c \ = {}" using wfTh_lookup_supp_empty * wfX_wfY by metis + + have "boxed_b \ (SCons tyid dc s2) (B_id tyid) bv b' (SCons tyid dc s2')" + proof(rule boxed_b_BConsI) + show 1:"AF_typedef tyid dclist \ set \" using * by auto + show 2:"(dc, \ x : b1 | c \) \ set dclist" using * by auto + have bvf:"atom bv \ b1" using sp \.fresh fresh_def by auto + show "\ \ s2 ~ b1 [ bv ::= b' ] \ s2'" using V_cons s2 s2' * by metis + qed + then show ?case using * s2 s2' by simp +next + case (V_consp tyid dc b1 v1) + + obtain bv2 dclist x2 b2 c2 where *: "b = B_app tyid b1 \ AF_typedef_poly tyid bv2 dclist \ set \ \ + (dc, \ x2 : b2 | c2 \) \ set dclist \ \ ; B ; \ \\<^sub>w\<^sub>f v1 : b2[bv2::=b1]\<^sub>b\<^sub>b" + using wf_strong_elim(1)[OF V_consp (5)] by metis + + obtain s2 where s2: "s = SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2 \ i \ (v1[bv::=b']\<^sub>v\<^sub>b) \ ~ s2" + using eval_v_elims V_consp subst_vb.simps by metis + + obtain s2' where s2': "s' = SConsp tyid dc b1 s2' \ i' \ v1 \ ~ s2'" + using eval_v_elims V_consp by metis + + have "\\<^sub>w\<^sub>f \" using V_consp wfX_wfY by metis + then obtain bv3::bv and dclist3 x3 b3 c3 where **:" AF_typedef_poly tyid bv2 dclist = AF_typedef_poly tyid bv3 dclist3 \ + (dc, \ x3 : b3 | c3 \) \ set dclist3 \ atom bv3 \ (b1, bv, b', s2, s2') \ b2[bv2::=b1]\<^sub>b\<^sub>b = b3[bv3::=b1]\<^sub>b\<^sub>b" + using * obtain_fresh_bv_dclist_b_iff[where tm="(b1, bv, b', s2, s2')"] by metis + + have "boxed_b \ (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2) (B_app tyid b1) bv b' (SConsp tyid dc b1 s2')" + proof(rule boxed_b_BConspI[of tyid bv3 dclist3 \, where x=x3 and b=b3 and c=c3]) + show 1:"AF_typedef_poly tyid bv3 dclist3 \ set \" using * ** by auto + show 2:"(dc, \ x3 : b3 | c3 \) \ set dclist3" using ** by auto + show "atom bv3 \ (b1, bv, b', s2, s2')" using ** by auto + show " \ \ s2 ~ b3[bv3::=b1]\<^sub>b\<^sub>b [ bv ::= b' ] \ s2'" using V_consp s2 s2' * ** by metis + qed + then show ?case using * s2 s2' by simp +qed + +lemma boxed_b_eq_eq: + assumes "boxed_b \ n1 b1 bv b' n1'" and "boxed_b \ n2 b1 bv b' n2'" and "s = SBool (n1 = n2)" and "\\<^sub>w\<^sub>f \" + "s' = SBool (n1' = n2')" + shows "s=s'" + using boxed_b_eq assms by auto + +lemma boxed_i_eval_ce_boxed_b: + fixes e::ce + assumes "i' \ e \ ~ s'" and "i \ e[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ s" and "wfCE \ B \ e b" and "boxed_i \ \ b' bv i i'" and "wfI \ \ i'" + shows "boxed_b \ s b bv b' s'" + using assms proof(nominal_induct e arbitrary: s s' b b' rule: ce.strong_induct) + case (CE_val x) + then show ?case using boxed_i_eval_v_boxed_b eval_e_elims wfCE_elims subst_ceb.simps by metis +next + case (CE_op opp v1 v2) + + show ?case proof(rule opp.exhaust) + assume \opp = Plus\ + have 1:"wfCE \ B \ v1 (B_int)" using wfCE_elims CE_op \opp = Plus\ by metis + have 2:"wfCE \ B \ v2 (B_int)" using wfCE_elims CE_op \opp = Plus\ by metis + have *:"b = B_int" using CE_op wfCE_elims + by (metis \opp = plus\) + + obtain n1 and n2 where n:"s = SNum (n1 + n2) \ i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n1 \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n2" using eval_e_elims CE_op subst_ceb.simps \opp = plus\ by metis + obtain n1' and n2' where n':"s' = SNum (n1' + n2') \ i' \ v1 \ ~ SNum n1' \ i' \ v2 \ ~ SNum n2'" using eval_e_elims Plus CE_op \opp = plus\ by metis + + have "boxed_b \ (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op \opp = plus\ by metis + moreover have "boxed_b \ (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis + ultimately have "s=s'" using n' n boxed_b_elims(2) + by (metis rcl_val.eq_iff(2)) + thus ?thesis using * n n' boxed_b_BIntI CE_op wfRCV.intros Plus by simp + next + assume \opp = LEq\ + have 1:"wfCE \ B \ v1 (B_int)" using wfCE_elims CE_op \opp = LEq\ by metis + have 2:"wfCE \ B \ v2 (B_int)" using wfCE_elims CE_op \opp = LEq\ by metis + hence *:"b = B_bool" using CE_op wfCE_elims \opp = LEq\ by metis + obtain n1 and n2 where n:"s = SBool (n1 \ n2) \ i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n1 \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n2" using eval_e_elims subst_ceb.simps CE_op \opp = LEq\ by metis + obtain n1' and n2' where n':"s' = SBool (n1' \ n2') \ i' \ v1 \ ~ SNum n1' \ i' \ v2 \ ~ SNum n2'" using eval_e_elims CE_op \opp = LEq\ by metis + + have "boxed_b \ (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis + moreover have "boxed_b \ (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis + ultimately have "s=s'" using n' n boxed_b_elims(2) + by (metis rcl_val.eq_iff(2)) + thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \opp = LEq\ by simp + next + assume \opp = Eq\ + obtain b1 where b1:"wfCE \ B \ v1 b1 \ wfCE \ B \ v2 b1" using wfCE_elims CE_op \opp = Eq\ by metis + + hence *:"b = B_bool" using CE_op wfCE_elims \opp = Eq\ by metis + obtain n1 and n2 where n:"s = SBool (n1 = n2) \ i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ n1 \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ n2" using eval_e_elims subst_ceb.simps CE_op \opp = Eq\ by metis + obtain n1' and n2' where n':"s' = SBool (n1' = n2') \ i' \ v1 \ ~ n1' \ i' \ v2 \ ~ n2'" using eval_e_elims CE_op \opp = Eq\ by metis + + have "boxed_b \ n1 b1 bv b' n1'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis + moreover have "boxed_b \ n2 b1 bv b' n2'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis + moreover have "\\<^sub>w\<^sub>f \" using b1 wfX_wfY by metis + ultimately have "s=s'" using n' n boxed_b_elims + boxed_b_eq_eq by metis + thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \opp = Eq\ by simp + qed + +next + case (CE_concat v1 v2) + + obtain bv1 and bv2 where s : "s = SBitvec (bv1 @ bv2) \ (i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SBitvec bv1) \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SBitvec bv2" + using eval_e_elims(7) subst_ceb.simps CE_concat.prems(2) eval_e_elims(6) subst_ceb.simps(6) by metis + obtain bv1' and bv2' where s' : "s' = SBitvec (bv1' @ bv2') \ i' \ v1 \ ~ SBitvec bv1' \ i' \ v2 \ ~ SBitvec bv2'" + using eval_e_elims(7) CE_concat by metis + + then show ?case using boxed_i_eval_v_boxed_b wfCE_elims s s' CE_concat + by (metis CE_concat.prems(3) assms assms(5) wfRCV_BBitvecI boxed_b_Bbitvec boxed_b_elims(7) eval_e_concatI eval_e_uniqueness) +next + case (CE_fst ce) + obtain s2 where 1:"i \ ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SPair s s2" using CE_fst eval_e_elims subst_ceb.simps by metis + obtain s2' where 2:"i' \ ce \ ~ SPair s' s2'" using CE_fst eval_e_elims by metis + obtain b2 where 3:"wfCE \ B \ ce (B_pair b b2)" using wfCE_elims(4) CE_fst by metis + + have "boxed_b \ (SPair s s2) (B_pair b b2) bv b' (SPair s' s2')" + using 1 2 3 CE_fst boxed_i_eval_v_boxed_b boxed_b_BPairI by auto + thus ?case using boxed_b_elims(5) by force +next + case (CE_snd v) + obtain s1 where 1:"i \ v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SPair s1 s" using CE_snd eval_e_elims subst_ceb.simps by metis + obtain s1' where 2:"i' \ v \ ~ SPair s1' s'" using CE_snd eval_e_elims by metis + obtain b1 where 3:"wfCE \ B \ v (B_pair b1 b )" using wfCE_elims(5) CE_snd by metis + + have "boxed_b \ (SPair s1 s ) (B_pair b1 b ) bv b' (SPair s1' s')" using 1 2 3 CE_snd boxed_i_eval_v_boxed_b by simp + thus ?case using boxed_b_elims(5) by force +next + case (CE_len v) + obtain s1 where s: "i \ v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SBitvec s1" using CE_len eval_e_elims subst_ceb.simps by metis + obtain s1' where s': "i' \ v \ ~ SBitvec s1'" using CE_len eval_e_elims by metis + + have "\ ; B ; \ \\<^sub>w\<^sub>f v : B_bitvec \ b = B_int" using wfCE_elims CE_len by metis + then show ?case using boxed_i_eval_v_boxed_b s s' CE_len + by (metis boxed_b_BIntI boxed_b_elims(7) eval_e_lenI eval_e_uniqueness subst_ceb.simps(5) wfI_wfCE_eval_e) +qed + +lemma eval_c_eq_bs_boxed: + fixes c::c + assumes "i \ c[bv::=b]\<^sub>c\<^sub>b \ ~ s" and "i' \ c \ ~ s'" and "wfC \ B \ c" and "wfI \ \ i'" and "\ ; \[bv::=b]\<^sub>\\<^sub>b \ i " + and "boxed_i \ \ b bv i i'" + shows "s = s'" + using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct) + case C_true + then show ?case using eval_c_elims subst_cb.simps by metis +next + case C_false + then show ?case using eval_c_elims subst_cb.simps by metis +next + case (C_conj c1 c2) + obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \ eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \ s = (s1\s2)" using C_conj eval_c_elims(3) subst_cb.simps(3) by metis + obtain s1' and s2' where 2:"eval_c i' c1 s1' \ eval_c i' c2 s2' \ s' = (s1'\s2')" using C_conj eval_c_elims(3) by metis + then show ?case using 1 2 wfC_elims C_conj by metis +next + case (C_disj c1 c2) + + obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \ eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \ s = (s1\s2)" using C_disj eval_c_elims(4) subst_cb.simps(4) by metis + obtain s1' and s2' where 2:"eval_c i' c1 s1' \ eval_c i' c2 s2' \ s' = (s1'\s2')" using C_disj eval_c_elims(4) by metis + then show ?case using 1 2 wfC_elims C_disj by metis +next + case (C_not c) + obtain s1::bool where 1: "(i \ c[bv::=b]\<^sub>c\<^sub>b \ ~ s1) \ (s = (\ s1))" using C_not eval_c_elims(6) subst_cb.simps(7) by metis + obtain s1'::bool where 2: "(i' \ c \ ~ s1') \ (s' = (\ s1'))" using C_not eval_c_elims(6) by metis + then show ?case using 1 2 wfC_elims C_not by metis +next + case (C_imp c1 c2) + obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \ eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \ s = (s1 \ s2)" using C_imp eval_c_elims(5) subst_cb.simps(5) by metis + obtain s1' and s2' where 2:"eval_c i' c1 s1' \ eval_c i' c2 s2' \ s' = (s1' \ s2')" using C_imp eval_c_elims(5) by metis + then show ?case using 1 2 wfC_elims C_imp by metis +next + case (C_eq e1 e2) + obtain be where be: "wfCE \ B \ e1 be \ wfCE \ B \ e2 be" using C_eq wfC_elims by metis + obtain s1 and s2 where 1: "eval_e i (e1[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s1 \ eval_e i (e2[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s2 \ s = (s1 = s2)" using C_eq eval_c_elims(7) subst_cb.simps(6) by metis + obtain s1' and s2' where 2:"eval_e i' e1 s1' \ eval_e i' e2 s2' \ s' = (s1' = s2' )" using C_eq eval_c_elims(7) by metis + have "\\<^sub>w\<^sub>f \" using C_eq wfX_wfY by metis + moreover have "\ ; \[bv::=b]\<^sub>\\<^sub>b \ i " using C_eq by auto + ultimately show ?case using boxed_b_eq[of \ s1 be bv b s1' s2 s2'] 1 2 boxed_i_eval_ce_boxed_b C_eq wfC_elims subst_cb.simps 1 2 be by auto +qed + +lemma is_satis_bs_boxed: + fixes c::c + assumes "boxed_i \ \ b bv i i'" and "wfC \ B \ c" and "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" and "\ ; \ \ i'" + and "(i \ c[bv::=b]\<^sub>c\<^sub>b)" + shows "(i' \ c)" +proof - + have "eval_c i (c[bv::=b]\<^sub>c\<^sub>b) True" using is_satis.simps assms by auto + moreover obtain s where "i' \ c \ ~ s" using eval_c_exist assms by metis + ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis +qed + +lemma is_satis_bs_boxed_rev: + fixes c::c + assumes "boxed_i \ \ b bv i i'" and "wfC \ B \ c" and "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" and "\ ; \ \ i'" and "wfC \ {||} \[bv::=b]\<^sub>\\<^sub>b (c[bv::=b]\<^sub>c\<^sub>b)" + and "(i' \ c)" + shows "(i \ c[bv::=b]\<^sub>c\<^sub>b)" +proof - + have "eval_c i' c True" using is_satis.simps assms by auto + moreover obtain s where "i \ c[bv::=b]\<^sub>c\<^sub>b \ ~ s" using eval_c_exist assms by metis + ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis +qed + +lemma bs_boxed_wfi_aux: + fixes b::b and bv::bv and \::\ and B::\ + assumes "boxed_i \ \ b bv i i'" and "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" and "\\<^sub>w\<^sub>f \" and "wfG \ B \" + shows "\ ; \ \ i'" + using assms proof(induct rule: boxed_i.inducts) + case (boxed_i_GNilI T i) + then show ?case using wfI_def by auto +next + case (boxed_i_GConsI s i x1 T b1 bv b s' G i' c1) + { + fix x2 b2 c2 + assume as : "(x2,b2,c2) \ toSet ((x1, b1, c1) #\<^sub>\ G)" + + then consider (hd) "(x2,b2,c2) = (x1, b1, c1)" | (tail) "(x2,b2,c2) \ toSet G" using toSet.simps by auto + hence "\s. Some s = (i'(x1 \ s')) x2 \ wfRCV T s b2" proof(cases) + case hd + hence "b1=b2" by auto + moreover have "(x2,b2[bv::=b]\<^sub>b\<^sub>b,c2[bv::=b]\<^sub>c\<^sub>b) \ toSet ((x1, b1, c1) #\<^sub>\ G)[bv::=b]\<^sub>\\<^sub>b" using hd subst_gb.simps by simp + moreover hence "wfRCV T s b2[bv::=b]\<^sub>b\<^sub>b" using wfI_def boxed_i_GConsI hd + proof - + obtain ss :: "b \ x \ (x \ rcl_val option) \ type_def list \ rcl_val" where + "\x1a x2a x3 x4. (\v5. Some v5 = x3 x2a \ wfRCV x4 v5 x1a) = (Some (ss x1a x2a x3 x4) = x3 x2a \ wfRCV x4 (ss x1a x2a x3 x4) x1a)" + by moura (* 0.0 ms *) + then have f1: "Some (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) = i x1 \ wfRCV T (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) b2[bv::=b]\<^sub>b\<^sub>b" + using boxed_i_GConsI.prems(1) hd wfI_def by auto (* 31 ms *) + then have "ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T = s" + by (metis (no_types) boxed_i_GConsI.hyps(1) option.inject) (* 0.0 ms *) + then show ?thesis + using f1 by blast (* 0.0 ms *) + qed + ultimately have "wfRCV T s' b2" using boxed_i_GConsI boxed_b_wfRCV by metis + + then show ?thesis using hd by simp + next + case tail + hence "wfI T G i'" using boxed_i_GConsI wfI_suffix wfG_suffix subst_gb.simps + by (metis (no_types, lifting) Un_iff toSet.simps(2) wfG_cons2 wfI_def) + then show ?thesis using wfI_def[of T G i'] tail + using boxed_i_GConsI.prems(3) split_G wfG_cons_fresh2 by fastforce + qed + } + thus ?case using wfI_def by fast + +qed + +lemma is_satis_g_bs_boxed_aux: + fixes G::\ + assumes "boxed_i \ G1 b bv i i'" and "wfI \ G1[bv::=b]\<^sub>\\<^sub>b i" and "wfI \ G1 i'" and "G1 = (G2@G)" and "wfG \ B G1" + and "(i \ G[bv::=b]\<^sub>\\<^sub>b) " + shows "(i' \ G)" + using assms proof(induct G arbitrary: G2 rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \' G2) + show ?case proof(subst is_satis_g.simps,rule) + have *:"wfC \ B G1 c'" using GCons wfG_wfC_inside by force + show "i' \ c'" using is_satis_bs_boxed[OF assms(1) * ] GCons by auto + obtain G3 where "G1 = G3 @ \'" using GCons append_g.simps + by (metis append_g_assoc) + then show "i' \ \'" using GCons append_g.simps by simp + qed +qed + +lemma is_satis_g_bs_boxed: + fixes G::\ + assumes "boxed_i \ G b bv i i'" and "wfI \ G[bv::=b]\<^sub>\\<^sub>b i" and "wfI \ G i'" and "wfG \ B G" + and "(i \ G[bv::=b]\<^sub>\\<^sub>b) " + shows "(i' \ G)" + using is_satis_g_bs_boxed_aux assms + by (metis (full_types) append_g.simps(1)) + +lemma subst_b_valid: + fixes s::s and b::b + assumes "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" and "\ ; {|bv|} ;\ \ c" + shows "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ c[bv::=b]\<^sub>c\<^sub>b " +proof(rule validI) + + show **:"\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using assms valid.simps wf_b_subst subst_gb.simps by metis + show "\i. (wfI \ \[bv::=b]\<^sub>\\<^sub>b i \ i \ \[bv::=b]\<^sub>\\<^sub>b) \ i \ c[bv::=b]\<^sub>c\<^sub>b " + proof(rule,rule) + fix i + assume *:"wfI \ \[bv::=b]\<^sub>\\<^sub>b i \ i \ \[bv::=b]\<^sub>\\<^sub>b" + + obtain i' where idash: "boxed_i \ \ b bv i i'" using boxed_i_ex wfX_wfY assms * by fastforce + + have wfc: "\ ; {|bv|} ; \ \\<^sub>w\<^sub>f c" using valid.simps assms by simp + have wfg: "\ ; {|bv|} \\<^sub>w\<^sub>f \" using valid.simps wfX_wfY assms by metis + hence wfi: "wfI \ \ i'" using idash * bs_boxed_wfi_aux subst_gb.simps wfX_wfY by metis + moreover have "i' \ \" proof (rule is_satis_g_bs_boxed[OF idash ] wfX_wfY(2)[OF wfc]) + show "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" using subst_gb.simps * by simp + show "wfI \ \ i'" using wfi by auto + show "\ ; B \\<^sub>w\<^sub>f \ " using wfg assms by auto + show "i \ \[bv::=b]\<^sub>\\<^sub>b" using subst_gb.simps * by simp + qed + ultimately have ic:"i' \ c" using assms valid_def using valid.simps by blast + + show "i \ c[bv::=b]\<^sub>c\<^sub>b" proof(rule is_satis_bs_boxed_rev) + show "\ ; \ ; b , bv \ i \ i'" using idash by auto + show "\ ; B ; \ \\<^sub>w\<^sub>f c " using wfc assms by auto + show "\ ; \[bv::=b]\<^sub>\\<^sub>b \ i" using subst_gb.simps * by simp + show "\ ; \ \ i'" using wfi by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using ** by auto + show "i' \ c" using ic by auto + qed + + qed +qed + +section \Expression Operator Lemmas\ + +lemma is_satis_len_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (L_num (int (length v)))) )" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit (L_num (int (length v))))) (SNum (int (length v)))" + using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SNum (int (length v)))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e) (SNum (int (length v)))" + using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI) + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma is_satis_plus_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (L_num (n1+n2))))" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e))" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit (L_num (n1+n2)))) (SNum (n1+n2))" + using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SNum (n1+n2))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)) (SNum (n1+n2))" + using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI) + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma is_satis_leq_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (if (n1 \ n2) then L_true else L_false)))" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit ((if (n1 \ n2) then L_true else L_false)))) (SBool (n1\n2))" + using eval_e_elims(1) eval_v_elims eval_l.simps + by (metis (full_types) eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SBool (n1\n2))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2) )]\<^sup>c\<^sup>e) (SBool (n1\n2))" + using eval_e_elims(3) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI) + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma eval_lit_inj: + fixes n1::l and n2::l + assumes "\ n1 \ = s" and "\ n2 \ = s" + shows "n1=n2" + using assms proof(nominal_induct s rule: rcl_val.strong_induct) + case (SBitvec x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SNum x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SBool x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SPair x1a x2a) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SCons x1a x2a x3a) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SConsp x1a x2a x3a x4) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case SUnit + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SUt x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +qed + +lemma eval_e_lit_inj: + fixes n1::l and n2::l + assumes "i \ [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s" and "i \ [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s" + shows "n1=n2" + using eval_lit_inj assms eval_e_elims eval_v_elims by metis + +lemma is_satis_eq_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (if (n1 = n2) then L_true else L_false)))" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2))]\<^sup>c\<^sup>e)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit ((if (n1=n2) then L_true else L_false)))) (SBool (n1=n2))" + using eval_e_elims(1) eval_v_elims eval_l.simps + by (metis (full_types) eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SBool (n1=n2))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2) )]\<^sup>c\<^sup>e) (SBool (n1=n2))" + proof - + obtain s1 and s2 where *:"i \ [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s1 \ i \ [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s2" using eval_l.simps eval_e.intros eval_v_litI by metis + moreover have " SBool (n1 = n2) = SBool (s1 = s2)" proof(cases "n1=n2") + case True + then show ?thesis using * + by (simp add: calculation eval_e_uniqueness) + next + case False + then show ?thesis using * eval_e_lit_inj by auto + qed + ultimately show ?thesis using eval_e_eqI[of i "[(V_lit (n1))]\<^sup>c\<^sup>e" s1 "[(V_lit (n2))]\<^sup>c\<^sup>e" s2 ] by auto + qed + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma valid_eq_e: + assumes "\i s1 s2. wfG P \ GNil \ wfI P GNil i \ eval_e i e1 s1 \ eval_e i e2 s2 \ s1 = s2" + and "wfCE P \ GNil e1 b" and "wfCE P \ GNil e2 b" + shows "P ; \ ; (x, b , CE_val (V_var x) == e1 )#\<^sub>\ GNil \ CE_val (V_var x) == e2" + unfolding valid.simps +proof(intro conjI) + show \ P ; \ ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \ + using assms wf_intros wfX_wfY b.eq_iff fresh_GNil wfC_e_eq2 wfV_elims by meson + show \\i. ((P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil \ i) \ (i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil) \ + (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2)) \ proof(rule+) + fix i + assume as:"P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil \ i \ i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil" + + have *: "P ; GNil \ i " using wfI_def by auto + + then obtain s1 where s1:"eval_e i e1 s1" using assms eval_e_exist by metis + obtain s2 where s2:"eval_e i e2 s2" using assms eval_e_exist * by metis + moreover have "i x = Some s1" proof - + have "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1" using as is_satis_g.simps by auto + thus ?thesis using s1 + by (metis eval_c_elims(7) eval_e_elims(1) eval_e_uniqueness eval_v_elims(2) is_satis.cases) + qed + moreover have "s1 = s2" using s1 s2 * assms wfG_nilI wfX_wfY by metis + + ultimately show "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \ ~ True" + using eval_c.intros eval_e.intros eval_v.intros + proof - + have "i \ e2 \ ~ s1" + by (metis \s1 = s2\ s2) (* 0.0 ms *) + then show ?thesis + by (metis (full_types) \i x = Some s1\ eval_c_eqI eval_e_valI eval_v_varI) (* 31 ms *) + qed + qed +qed + +lemma valid_len: + assumes " \\<^sub>w\<^sub>f \" + shows "\ ; \ ; (x, B_int, [[x]\<^sup>v]\<^sup>c\<^sup>e == [[ L_num (int (length v)) ]\<^sup>v]\<^sup>c\<^sup>e) #\<^sub>\ GNil \ [[x]\<^sup>v]\<^sup>c\<^sup>e == CE_len [[ L_bitvec v ]\<^sup>v]\<^sup>c\<^sup>e" (is "\ ; \ ; ?G \ ?c" ) +proof - + have *:"\ \\<^sub>w\<^sub>f ([]::\) \ \ ; \ ; GNil \\<^sub>w\<^sub>f []\<^sub>\ " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto + + moreover hence "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int" + using wfCE_valI * wfV_litI base_for_lit.simps + by (metis wfE_valI wfX_wfY) + + moreover have "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int" + using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY wfCE_valI + by (metis wfCE_lenI) + moreover have "atom x \ GNil" by auto + ultimately have "\ ; \ ; ?G \\<^sub>w\<^sub>f ?c" using wfC_e_eq2 assms by simp + moreover have "(\i. wfI \ ?G i \ is_satis_g i ?G \ is_satis i ?c)" using is_satis_len_imp by auto + ultimately show ?thesis using valid.simps by auto +qed + +lemma valid_arith_bop: + assumes "wfG \ \ \" and "opp = Plus \ ll = (L_num (n1+n2)) \ (opp = LEq \ ll = ( if n1\n2 then L_true else L_false))" + and "(opp = Plus \ b = B_int) \ (opp = LEq \ b = B_bool)" and + "atom x \ \" + shows "\; \ ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e ))" (is "\ ; \ ; ?G \ ?c") +proof - + have "wfC \ \ ?G ?c" proof(rule wfC_e_eq2) + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI wfV_litI assms base_for_lit.simps by metis + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) : b " + using wfCE_plusI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis + show "\\<^sub>w\<^sub>f \" using assms wfX_wfY by auto + show "atom x \ \" using assms by auto + qed + + moreover have "\i. wfI \ ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI , rule impI) + fix i + assume "wfI \ ?G i \ is_satis_g i ?G" + + hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (ll)) ))" by auto + thus "is_satis i ((CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)))" + using is_satis_plus_imp assms opp.exhaust is_satis_leq_imp by auto + qed + ultimately show ?thesis using valid.simps by metis +qed + +lemma valid_eq_bop: + assumes "wfG \ \ \" and "atom x \ \" and "base_for_lit l1 = base_for_lit l2" + shows "\; \ ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e ))" (is "\ ; \ ; ?G \ ?c") +proof - + let ?ll = "(if l1 = l2 then L_true else L_false)" + have "wfC \ \ ?G ?c" proof(rule wfC_e_eq2) + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI wfV_litI assms base_for_lit.simps by metis + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e) : B_bool " + using wfCE_eqI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis + show "\\<^sub>w\<^sub>f \" using assms wfX_wfY by auto + show "atom x \ \" using assms by auto + qed + + moreover have "\i. wfI \ ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI , rule impI) + fix i + assume "wfI \ ?G i \ is_satis_g i ?G" + + hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (?ll)) ))" by auto + thus "is_satis i ((CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e)))" + using is_satis_eq_imp assms by auto + qed + ultimately show ?thesis using valid.simps by metis +qed + +lemma valid_fst: + fixes x::x and v\<^sub>1::v and v\<^sub>2::v + assumes "wfTh \" and "wfV \ \ GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)" + shows "\ ; \ ; (x, b\<^sub>1, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e) #\<^sub>\ GNil \ [[x]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e" +proof(rule valid_eq_e) + show \\i s1 s2. (\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ (i \ [ v\<^sub>1 ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2) \ s1 = s2\ + proof(rule+) + fix i s1 s2 + assume as:"\ ; \ \\<^sub>w\<^sub>f GNil \ \ ; GNil \ i \ (i \ [ v\<^sub>1 ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2)" + then obtain s2' where *:"i \ [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \ ~ SPair s2 s2'" + using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims + by meson + then have " i \ v\<^sub>1 \ ~ s2" using eval_v_elims(3)[OF *] by auto + then show "s1 = s2" using eval_v_uniqueness as + using eval_e_uniqueness eval_e_valI by blast + qed + + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [ v\<^sub>1 ]\<^sup>c\<^sup>e : b\<^sub>1 \ using assms + by (metis b.eq_iff(4) wfV_elims(3) wfV_wfCE) + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>1 \ using assms using wfCE_fstI + using wfCE_valI by blast +qed + +lemma valid_snd: + fixes x::x and v\<^sub>1::v and v\<^sub>2::v + assumes "wfTh \" and "wfV \ \ GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)" + shows "\ ; \ ; (x, b\<^sub>2, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>2]\<^sup>c\<^sup>e) #\<^sub>\ GNil \ [[x]\<^sup>v]\<^sup>c\<^sup>e == [#2[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e" +proof(rule valid_eq_e) + show \\i s1 s2. (\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ (i \ [ v\<^sub>2 ]\<^sup>c\<^sup>e \ ~ s1) \ +(i \ [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2) \ s1 = s2\ + proof(rule+) + fix i s1 s2 + assume as:"\ ; \ \\<^sub>w\<^sub>f GNil \ \ ; GNil \ i \ (i \ [ v\<^sub>2 ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2)" + then obtain s2' where *:"i \ [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \ ~ SPair s2' s2" + using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims + by meson + then have " i \ v\<^sub>2 \ ~ s2" using eval_v_elims(3)[OF *] by auto + then show "s1 = s2" using eval_v_uniqueness as + using eval_e_uniqueness eval_e_valI by blast + qed + + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [ v\<^sub>2 ]\<^sup>c\<^sup>e : b\<^sub>2 \ using assms + by (metis b.eq_iff wfV_elims wfV_wfCE) + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>2 \ using assms using wfCE_sndI wfCE_valI by blast +qed + +lemma valid_concat: + fixes v1::"bit list" and v2::"bit list" + assumes " \\<^sub>w\<^sub>f \" + shows "\ ; \ ; (x, B_bitvec, (CE_val (V_var x) == CE_val (V_lit (L_bitvec (v1@ v2))))) #\<^sub>\ GNil \ + (CE_val (V_var x) == CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e ) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e) )" +proof(rule valid_eq_e) + show \\i s1 s2. ((\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ + (i \ [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ s2) \ + s1 = s2)\ + proof(rule+) + fix i s1 s2 + assume as: "(\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ (i \ [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s1) \ + (i \ [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2) " + + hence *: "i \ [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2" by auto + obtain bv1 bv2 where s2:"s2 = SBitvec (bv1 @ bv2) \ i \ [ L_bitvec v1 ]\<^sup>v \ ~ SBitvec bv1 \ (i \ [ L_bitvec v2 ]\<^sup>v \ ~ SBitvec bv2)" + using eval_e_elims(7)[OF *] eval_e_elims(1) by metis + hence "v1 = bv1 \ v2 = bv2" using eval_v_elims(1) eval_l.simps(5) by force + moreover then have "s1 = SBitvec (bv1 @ bv2)" using s2 using eval_v_elims(1) eval_l.simps(5) + by (metis as eval_e_elims(1)) + + then show "s1 = s2" using s2 by auto + qed + + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e : B_bitvec \ + by (metis assms base_for_lit.simps(5) wfG_nilI wfV_litI wfV_wfCE) + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : B_bitvec \ + by (metis assms base_for_lit.simps(5) wfCE_concatI wfG_nilI wfV_litI wfCE_valI) +qed + +lemma valid_ce_eq: + fixes ce::ce + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f ce : b" + shows \\ ; \ ; \ \ ce == ce \ + unfolding valid.simps proof + show \ \ ; \ ; \ \\<^sub>w\<^sub>f ce == ce \ using assms wfC_eqI by auto + show \\i. \ ; \ \ i \ i \ \ \ i \ ce == ce \ proof(rule+) + fix i + assume "\ ; \ \ i \ i \ \" + then obtain s where "i\ ce \ ~ s" using assms eval_e_exist by metis + then show "i \ ce == ce \ ~ True " using eval_c_eqI by metis + qed +qed + +lemma valid_eq_imp: + fixes c1::c and c2::c + assumes " \ ; \ ; (x, b, c2) #\<^sub>\ \ \\<^sub>w\<^sub>f c1 IMP c2" + shows " \ ; \ ; (x, b, c2) #\<^sub>\ \ \ c1 IMP c2 " +proof - + have "\i. (\ ; (x, b, c2) #\<^sub>\ \ \ i \ i \ (x, b, c2) #\<^sub>\ \) \ i \ ( c1 IMP c2 )" + proof(rule,rule) + fix i + assume as:"\ ; (x, b, c2) #\<^sub>\ \ \ i \ i \ (x, b, c2) #\<^sub>\ \" + + have "\ ; \ ; (x, b, c2) #\<^sub>\ \ \\<^sub>w\<^sub>f c1" using wfC_elims assms by metis + + then obtain sc where "i \ c1 \ ~ sc" using eval_c_exist assms as by metis + moreover have "i \ c2 \ ~ True" using as is_satis_g.simps is_satis.simps by auto + + ultimately have "i \ c1 IMP c2 \ ~ True" using eval_c_impI by metis + + thus "i \ c1 IMP c2" using is_satis.simps by auto + qed + thus ?thesis using assms by auto +qed + +lemma valid_range: + assumes "0 \ n \ n \ m" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ GNil \ + (C_eq (CE_op LEq (CE_val (V_var x)) (CE_val (V_lit (L_num m)))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)" + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") +proof(rule validI) + have wfg: " \ ; {||} \\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil " + using assms base_for_lit.simps wfG_nilI wfV_litI fresh_GNil wfB_intI wfC_v_eq wfG_cons1I wfG_cons2I by metis + + show "\ ; {||} ; ?G \\<^sub>w\<^sub>f ?c1 AND ?c2" + using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI + by metis + + show "\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" proof(rule,rule) + fix i + assume a:"\ ; ?G \ i \ i \ ?G" + hence *:"i \ V_var x \ ~ SNum n" + proof - + obtain sv where sv: "i x = Some sv \ \ \ sv : B_int" using a wfI_def by force + have "i \ (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \ ~ True" + using a is_satis_g.simps + using is_satis.cases by blast + hence "i x = Some(SNum n)" using sv + by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2)) + thus ?thesis using eval_v_varI by auto + qed + + show "i \ ?c1 AND ?c2" + proof - + have "i \ ?c1 \ ~ True" + proof - + have "i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num m ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + + moreover have "i \ ?c2 \ ~ True" + proof - + have "i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + ultimately show ?thesis using eval_c_conjI is_satis.simps by metis + qed + qed +qed + +lemma valid_range_length: + fixes \::\ + assumes "0 \ n \ n \ int (length v)" and " \ ; {||} \\<^sub>w\<^sub>f \" and "atom x \ \" + shows "\ ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ \ \ + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + " + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") +proof(rule validI) + have wfg: " \ ; {||} \\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ \ " apply(rule wfG_cons1I) + apply simp + using assms apply simp+ + using assms base_for_lit.simps wfG_nilI wfV_litI wfB_intI wfC_v_eq wfB_intI wfX_wfY assms by metis+ + + show "\ ; {||} ; ?G \\<^sub>w\<^sub>f ?c1 AND ?c2" + using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI + by (metis (full_types) wfCE_lenI) + + show "\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" proof(rule,rule) + fix i + assume a:"\ ; ?G \ i \ i \ ?G" + hence *:"i \ V_var x \ ~ SNum n" + proof - + obtain sv where sv: "i x = Some sv \ \ \ sv : B_int" using a wfI_def by force + have "i \ (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \ ~ True" + using a is_satis_g.simps + using is_satis.cases by blast + hence "i x = Some(SNum n)" using sv + by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2)) + thus ?thesis using eval_v_varI by auto + qed + + show "i \ ?c1 AND ?c2" + proof - + have "i \ ?c2 \ ~ True" + proof - + have "i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_lenI eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + + moreover have "i \ ?c1 \ ~ True" + proof - + have "i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + ultimately show ?thesis using eval_c_conjI is_satis.simps by metis + qed + qed +qed + +lemma valid_range_length_inv_gnil: + fixes \::\ + assumes "\\<^sub>w\<^sub>f \ " + and "\ ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ GNil \ + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + " + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") + shows "0 \ n \ n \ int (length v)" +proof - + have *:"\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" using assms valid.simps by simp + + obtain i where i: "i x = Some (SNum n)" by auto + have "\ ; ?G \ i \ i \ ?G" proof + show "\ ; ?G \ i" unfolding wfI_def using wfRCV_BIntI i * by auto + have "i \ ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \ ~ True" + using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps + by (metis (full_types) i) + thus "i \ ?G" unfolding is_satis_g.simps is_satis.simps by auto + qed + hence **:"i \ ?c1 AND ?c2" using * by auto + + hence 1: "i \ ?c1 \ ~ True" using eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where "(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where "SBool True = SBool (n1 \ n2) \ (i \ [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"] + using \(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2\ \sv1 = SBool True\ by fastforce + moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i + apply (metis eval_l.simps(3) rcl_val.eq_iff(2)) + using eval_e_elims eval_v_elims i + by (metis calculation option.inject rcl_val.eq_iff(2)) + ultimately have le1: "0 \ n " by simp + + hence 2: "i \ ?c2 \ ~ True" using ** eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \ i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where ***:"SBool True = SBool (n1 \ n2) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3) + using sv \sv1 = SBool True\ by metis + moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto + moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i + by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2)) + ultimately have le2: "n \ int (length v) " by simp + + show ?thesis using le1 le2 by auto +qed + +lemma wfI_cons: + fixes i::valuation and \::\ + assumes "i' \ \" and "\ ; \ \ i'" and "i = i' ( x \ s)" and "\ \ s : b" and "atom x \ \" + shows "\ ; (x,b,c) #\<^sub>\ \ \ i" + unfolding wfI_def proof - + { + fix x' b' c' + assume "(x',b',c') \ toSet ((x, b, c) #\<^sub>\ \)" + then consider "(x',b',c') = (x,b,c)" | "(x',b',c') \ toSet \" using toSet.simps by auto + then have "\s. Some s = i x' \ \ \ s : b'" proof(cases) + case 1 + then show ?thesis using assms by auto + next + case 2 + then obtain s where s:"Some s = i' x' \ \ \ s : b'" using assms wfI_def by auto + moreover have "x' \ x" using assms 2 fresh_dom_free by auto + ultimately have "Some s = i x'" using assms by auto + then show ?thesis using s wfI_def by auto + qed + } + thus "\(x, b, c)\toSet ((x, b, c) #\<^sub>\ \). \s. Some s = i x \ \ \ s : b" by auto +qed + +lemma valid_range_length_inv: + fixes \::\ + assumes "\ ; B \\<^sub>w\<^sub>f \ " and "atom x \ \" and "\i. i \ \ \ \ ; \ \ i" + and "\ ; B ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ \ \ + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + " + (is "\ ; ?B ; ?G \ ?c1 AND ?c2") + shows "0 \ n \ n \ int (length v)" +proof - + have *:"\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" using assms valid.simps by simp + + obtain i' where idash: "is_satis_g i' \ \ \ ; \ \ i'" using assms by auto + obtain i where i: "i = i' ( x \ SNum n)" by auto + hence ix: "i x = Some (SNum n)" by auto + have "\ ; ?G \ i \ i \ ?G" proof + show "\ ; ?G \ i" using wfI_cons i idash ix wfRCV_BIntI assms by simp + + have **:"i \ ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \ ~ True" + using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps i + by (metis (full_types) ix) + + show "i \ ?G" unfolding is_satis_g.simps proof + show \ i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ using ** is_satis.simps by auto + show \ i \ \ \ using idash i assms is_satis_g_i_upd by metis + qed + qed + hence **:"i \ ?c1 AND ?c2" using * by auto + + hence 1: "i \ ?c1 \ ~ True" using eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where "(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where "SBool True = SBool (n1 \ n2) \ (i \ [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"] + using \(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2\ \sv1 = SBool True\ by fastforce + moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i + apply (metis eval_l.simps(3) rcl_val.eq_iff(2)) + using eval_e_elims eval_v_elims i + calculation option.inject rcl_val.eq_iff(2) + by (metis ix) + ultimately have le1: "0 \ n " by simp + + hence 2: "i \ ?c2 \ ~ True" using ** eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \ i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where ***:"SBool True = SBool (n1 \ n2) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3) + using sv \sv1 = SBool True\ by metis + moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto + moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i + by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2)) + ultimately have le2: "n \ int (length v) " by simp + + show ?thesis using le1 le2 by auto +qed + +lemma eval_c_conj2I[intro]: + assumes "i \ c1 \ ~ True" and "i \ c2 \ ~ True" + shows "i \ (C_conj c1 c2) \ ~ True" + using assms eval_c_conjI by metis + +lemma valid_split: + assumes "split n v (v1,v2)" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; (z , [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\ GNil +\ ([ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) AND ([| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e)" + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") + unfolding valid.simps proof + + have wfg: " \ ; {||} \\<^sub>w\<^sub>f (z, [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\ GNil" + using wf_intros assms base_for_lit.simps fresh_GNil wfC_v_eq wfG_intros2 by metis + + show "\ ; {||} ; ?G \\<^sub>w\<^sub>f ?c1 AND ?c2" + apply(rule wfC_conjI) + apply(rule wfC_eqI) + apply(rule wfCE_valI) + apply(rule wfV_litI) + using wf_intros wfg lookup.simps base_for_lit.simps wfC_v_eq + apply (metis )+ + done + + have len:"int (length v1) = n" using assms split_length by auto + + show "\i. \ ; ?G \ i \ i \ ?G \ i \ (?c1 AND ?c2)" + proof(rule,rule) + fix i + assume a:"\ ; ?G \ i \ i \ ?G" + hence "i \ [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ True" + using is_satis_g.simps is_satis.simps by simp + then obtain sv where "i \ [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv \ i \ [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv" + using eval_c_elims by metis + hence "i \ [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ (SPair (SBitvec v1) (SBitvec v2))" using eval_c_eqI eval_v.intros eval_l.simps + by (metis eval_e_elims(1) eval_v_uniqueness) + hence b:"i z = Some (SPair (SBitvec v1) (SBitvec v2))" using a eval_e_elims eval_v_elims by metis + + have v1: "i \ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ SBitvec v1 " + using eval_e_fstI eval_e_valI eval_v_varI b by metis + have v2: "i \ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ SBitvec v2" + using eval_e_sndI eval_e_valI eval_v_varI b by metis + + have "i \ [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBitvec v" using eval_e.intros eval_v.intros eval_l.simps by metis + moreover have "i \ [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ SBitvec v" + using assms split_concat v1 v2 eval_e_concatI by metis + moreover have "i \ [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \ ~ SNum (int (length v1))" + using v1 eval_e_lenI by auto + moreover have "i \ [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n" using eval_e.intros eval_v.intros eval_l.simps by metis + ultimately show "i \ ?c1 AND ?c2" using is_satis.intros eval_c_conj2I eval_c_eqI len by metis + qed +qed + + +lemma is_satis_eq: + assumes "wfI \ G i" and "wfCE \ \ G e b" + shows "is_satis i (e == e)" +proof(rule) + obtain s where "eval_e i e s" using eval_e_exist assms by metis + thus "eval_c i (e == e ) True" using eval_c_eqI by metis +qed + +lemma is_satis_g_i_upd2: + assumes "eval_v i v s" and "is_satis ((i ( x \ s))) c0" and "atom x \ G" and "wfG \ \ (G3@((x,b,c0)#\<^sub>\G))" and "wfV \ \ G v b" and "wfI \ (G3[x::=v]\<^sub>\\<^sub>v@G) i" + and "is_satis_g i (G3[x::=v]\<^sub>\\<^sub>v@G)" + shows "is_satis_g (i ( x \ s)) (G3@((x,b,c0)#\<^sub>\G))" + using assms proof(induct G3 rule: \_induct) + case GNil + hence "is_satis_g (i(x \ s)) G" using is_satis_g_i_upd by auto + then show ?case using GNil using is_satis_g.simps append_g.simps by metis +next + case (GCons x' b' c' \') + hence "x\x'" using wfG_cons_append by metis + hence "is_satis_g i (((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v) @ G))" using subst_gv.simps GCons by auto + hence *:"is_satis i c'[x::=v]\<^sub>c\<^sub>v \ is_satis_g i ((\'[x::=v]\<^sub>\\<^sub>v) @ G)" using subst_gv.simps by auto + + have "is_satis_g (i(x \ s)) ((x', b', c') #\<^sub>\ (\'@ (x, b, c0) #\<^sub>\ G))" proof(subst is_satis_g.simps,rule) + show "is_satis (i(x \ s)) c'" proof(subst subst_c_satis_full[symmetric]) + show \eval_v i v s\ using GCons by auto + show \ \ ; \ ; ((x', b', c') #\<^sub>\\')@(x, b, c0) #\<^sub>\ G \\<^sub>w\<^sub>f c' \ using GCons wfC_refl by auto + show \wfI \ ((((x', b', c') #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v) @ G) i\ using GCons by auto + show \ \ ; \ ; G \\<^sub>w\<^sub>f v : b \ using GCons by auto + show \is_satis i c'[x::=v]\<^sub>c\<^sub>v\ using * by auto + qed + show "is_satis_g (i(x \ s)) (\' @ (x, b, c0) #\<^sub>\ G)" proof(rule GCons(1)) + show \eval_v i v s\ using GCons by auto + show \is_satis (i(x \ s)) c0\ using GCons by metis + show \atom x \ G\ using GCons by auto + show \ \ ; \\\<^sub>w\<^sub>f \' @ (x, b, c0) #\<^sub>\ G \ using GCons wfG_elims append_g.simps by metis + show \is_satis_g i (\'[x::=v]\<^sub>\\<^sub>v @ G)\ using * by auto + show "wfI \ (\'[x::=v]\<^sub>\\<^sub>v @ G) i" using GCons wfI_def subst_g_assoc_cons \x\x'\ by auto + show "\ ; \ ; G \\<^sub>w\<^sub>f v : b " using GCons by auto + qed + qed + moreover have "((x', b', c') #\<^sub>\ \' @ (x, b, c0) #\<^sub>\ G) = (((x', b', c') #\<^sub>\ \') @ (x, b, c0) #\<^sub>\ G)" by auto + ultimately show ?case using GCons by metis +qed + + +end \ No newline at end of file diff --git a/thys/MiniSail/ROOT b/thys/MiniSail/ROOT new file mode 100644 --- /dev/null +++ b/thys/MiniSail/ROOT @@ -0,0 +1,15 @@ +chapter AFP + +session "MiniSail" (AFP) = "HOL-Library" + + description \Formalisation of MiniSail\ + options [timeout = 1200] + sessions + "Nominal2" + "HOL-Eisbach" + "Native_Word" + "Show" + theories + "MiniSail" + document_files + "root.tex" + "root.bib" diff --git a/thys/MiniSail/Safety.thy b/thys/MiniSail/Safety.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Safety.thy @@ -0,0 +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) +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/SubstMethods.thy b/thys/MiniSail/SubstMethods.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/SubstMethods.thy @@ -0,0 +1,71 @@ +theory SubstMethods + (* Its seems that it's best to load the Eisbach tools last *) + imports IVSubst WellformedL "HOL-Eisbach.Eisbach_Tools" +begin + +text \ + See Eisbach/Examples.thy as well as Eisbach User Manual. + +Freshness for various substitution situations. It seems that if undirected and we throw all the +facts at them to try to solve in one shot, the automatic methods are *sometimes* unable +to handle the different variants, so some guidance is needed. +First we split into subgoals using fresh\_prodN and intro conjI + +The 'add', for example, will be induction premises that will contain freshness facts or freshness conditions from +prior obtains + +Use different arguments for different things or just lump into one bucket\ + +method fresh_subst_mth_aux uses add = ( + (match conclusion in "atom z \ (\::\)[x::=v]\<^sub>\\<^sub>v" for z x v \ \ \auto simp add: fresh_subst_gv_if[of "atom z" \ v x] add\) + | (match conclusion in "atom z \ (v'::v)[x::=v]\<^sub>v\<^sub>v" for z x v v' \ \auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_v_def add\ ) + | (match conclusion in "atom z \ (ce::ce)[x::=v]\<^sub>c\<^sub>e\<^sub>v" for z x v ce \ \auto simp add: fresh_subst_v_if subst_v_ce_def add\ ) + | (match conclusion in "atom z \ (\::\)[x::=v]\<^sub>\\<^sub>v" for z x v \ \ \auto simp add: fresh_subst_v_if fresh_subst_dv_if add\ ) + | (match conclusion in "atom z \ \'[x::=v]\<^sub>\\<^sub>v @ \" for z x v \' \ \ \metis add \ ) + | (match conclusion in "atom z \ (\::\)[x::=v]\<^sub>\\<^sub>v" for z x v \ \ \auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_\_def add\ ) + | (match conclusion in "atom z \ ({||} :: bv fset)" for z \ \auto simp add: fresh_empty_fset\) + (* tbc delta and types *) + | (auto simp add: add x_fresh_b pure_fresh) (* Cases where there is no subst and so can most likely get what we want from induction premises *) + ) + +method fresh_mth uses add = ( + (unfold fresh_prodN, intro conjI)?, + (fresh_subst_mth_aux add: add)+) + + +notepad +begin + fix \::\ and z::x and x::x and v::v and \::\ and v'::v and w::x and tyid::string and dc::string and b::b and ce::ce and bv::bv + + assume as:"atom z \ (\,v',\, v,w,ce) \ atom bv \ (\,v',\, v,w,ce,b) " + + have "atom z \ \[x::=v]\<^sub>\\<^sub>v" + by (fresh_mth add: as) + + hence "atom z \ v'[x::=v]\<^sub>v\<^sub>v" + by (fresh_mth add: as) + + hence "atom z \ \" + by (fresh_mth add: as) + + hence "atom z \ \" + by (fresh_mth add: as) + + hence "atom z \ (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v" + using as by auto + + hence "atom bv \ (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v" + using as by auto + + have "atom z \ (\,\[x::=v]\<^sub>\\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v, (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v) " + by (fresh_mth add: as) + + have "atom bv \ (\,\[x::=v]\<^sub>\\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v) " + by (fresh_mth add: as) + +end + + + + +end \ No newline at end of file diff --git a/thys/MiniSail/Syntax.thy b/thys/MiniSail/Syntax.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Syntax.thy @@ -0,0 +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) +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)) + 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)) + 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) + 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) + 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/SyntaxL.thy b/thys/MiniSail/SyntaxL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/SyntaxL.thy @@ -0,0 +1,383 @@ +(*<*) +theory SyntaxL + imports Syntax IVSubst +begin + (*>*) + +chapter \Syntax Lemmas\ + +section \Support, lookup and contexts\ + +lemma supp_v_tau [simp]: + assumes "atom z \ v" + shows "supp (\ z : b | CE_val (V_var z) == CE_val v \) = supp v \ supp b" + using assms \.supp c.supp ce.supp + by (simp add: fresh_def supp_at_base) + +lemma supp_v_var_tau [simp]: + assumes "z \ x" + shows "supp (\ z : b | CE_val (V_var z) == CE_val (V_var x) \) = { atom x } \ supp b" + using supp_v_tau assms + using supp_at_base by fastforce + +text \ Sometimes we need to work with a version of a binder where the variable is fresh +in something else, such as a bigger context. I think these could be generated automatically \ + +lemma obtain_fresh_fun_def: + fixes t::"'b::fs" + shows "\y::x. atom y \ (s,c,\,t) \ (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \ x) \ c ) ((y \ x) \ \) ((y \ x) \ s))))" +proof - + obtain y::x where y: "atom y \ (s,c,\,t)" using obtain_fresh by blast + moreover have "AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \ x) \ c ) ((y \ x) \ \) ((y \ x) \ s))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)))" + proof(cases "x=y") + case True + then show ?thesis using fun_def.eq_iff Abs1_eq_iff(3) flip_commute flip_fresh_fresh fresh_PairD by auto + next + case False + + have "(AF_fun_typ y b ((y \ x) \ c) ((y \ x) \ \) ((y \ x) \ s)) =(AF_fun_typ x b c \ s)" proof(subst fun_typ.eq_iff, subst Abs1_eq_iff(3)) + show \(y = x \ (((y \ x) \ c, (y \ x) \ \), (y \ x) \ s) = ((c, \), s) \ + y \ x \ (((y \ x) \ c, (y \ x) \ \), (y \ x) \ s) = (y \ x) \ ((c, \), s) \ atom y \ ((c, \), s)) \ + b = b\ using False flip_commute flip_fresh_fresh fresh_PairD y by auto + qed + thus ?thesis by metis + qed + ultimately show ?thesis using y fresh_Pair by metis +qed + +lemma lookup_fun_member: + assumes "Some (AF_fundef f ft) = lookup_fun \ f" + shows "AF_fundef f ft \ set \" + using assms proof (induct \) + case Nil + then show ?case by auto +next + case (Cons a \) + then show ?case using lookup_fun.simps + by (metis fun_def.exhaust insert_iff list.simps(15) option.inject) +qed + +lemma rig_dom_eq: + "dom (G[x \ c]) = dom G" +proof(induct G rule: \.induct) + case GNil + then show ?case using replace_in_g.simps by presburger +next + case (GCons xbc \') + obtain x' and b' and c' where xbc: "xbc=(x',b',c')" using prod_cases3 by blast + then show ?case using replace_in_g.simps GCons by simp +qed + +lemma lookup_in_rig_eq: + assumes "Some (b,c) = lookup \ x" + shows "Some (b,c') = lookup (\[x\c']) x" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x b c \') + then show ?case using replace_in_g.simps lookup.simps by auto +qed + +lemma lookup_in_rig_neq: + assumes "Some (b,c) = lookup \ y" and "x\y" + shows "Some (b,c) = lookup (\[x\c']) y" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + then show ?case using replace_in_g.simps lookup.simps by auto +qed + +lemma lookup_in_rig: + assumes "Some (b,c) = lookup \ y" + shows "\c''. Some (b,c'') = lookup (\[x\c']) y" +proof(cases "x=y") + case True + then show ?thesis using lookup_in_rig_eq using assms by blast +next + case False + then show ?thesis using lookup_in_rig_neq using assms by blast +qed + +lemma lookup_inside[simp]: + assumes "x \ fst ` toSet \'" + shows "Some (b1,c1) = lookup (\'@(x,b1,c1)#\<^sub>\\) x" + using assms by(induct \',auto) + +lemma lookup_inside2: + assumes "Some (b1,c1) = lookup (\'@((x,b0,c0)#\<^sub>\\)) y" and "x\y" + shows "Some (b1,c1) = lookup (\'@((x,b0,c0')#\<^sub>\\)) y" + using assms by(induct \' rule: \.induct,auto+) + +fun tail:: "'a list \ 'a list" where + "tail [] = []" +| "tail (x#xs) = xs" + +lemma lookup_options: + assumes "Some (b,c) = lookup (xt#\<^sub>\G) x" + shows "((x,b,c) = xt) \ (Some (b,c) = lookup G x)" + by (metis assms lookup.simps(2) option.inject surj_pair) + +lemma lookup_x: + assumes "Some (b,c) = lookup G x" + shows "x \ fst ` toSet G" + using assms + by(induct "G" rule: \.induct ,auto+) + +lemma GCons_eq_appendI: + fixes xs1::\ + shows "[| x #\<^sub>\ xs1 = ys; xs = xs1 @ zs |] ==> x #\<^sub>\ xs = ys @ zs" + by (drule sym) simp + +lemma split_G: "x : toSet xs \ \ys zs. xs = ys @ x #\<^sub>\ zs" +proof (induct xs) + case GNil thus ?case by simp +next + case GCons thus ?case using GCons_eq_appendI + by (metis Un_iff append_g.simps(1) singletonD toSet.simps(2)) +qed + +lemma lookup_not_empty: + assumes "Some \ = lookup G x" + shows "G \ GNil" + using assms by auto + +lemma lookup_in_g: + assumes "Some (b,c) = lookup \ x" + shows "(x,b,c) \ toSet \" + using assms apply(induct \, simp) + using lookup_options by fastforce + +lemma lookup_split: + fixes \::\ + assumes "Some (b,c) = lookup \ x" + shows "\G G' . \ = G'@(x,b,c)#\<^sub>\G" + by (meson assms(1) lookup_in_g split_G) + +lemma toSet_splitU[simp]: + "(x',b',c') \ toSet (\' @ (x, b, c) #\<^sub>\ \) \ (x',b',c') \ (toSet \' \ {(x, b, c)} \ toSet \)" + using append_g_toSetU toSet.simps by auto + +lemma toSet_splitP[simp]: + "(\(x', b', c')\toSet (\' @ (x, b, c) #\<^sub>\ \). P x' b' c') \ (\ (x', b', c')\toSet \'. P x' b' c') \ P x b c \ (\ (x', b', c')\toSet \. P x' b' c')" (is "?A \ ?B") + using toSet_splitU by force + +lemma lookup_restrict: + assumes "Some (b',c') = lookup (\'@(x,b,c)#\<^sub>\\) y" and "x \ y" + shows "Some (b',c') = lookup (\'@\) y" + using assms proof(induct \' rule:\_induct) + case GNil + then show ?case by auto +next + case (GCons x1 b1 c1 \') + then show ?case by auto +qed + +lemma supp_list_member: + fixes x::"'a::fs" and l::"'a list" + assumes "x \ set l" + shows "supp x \ supp l" + using assms apply(induct l, auto) + using supp_Cons by auto + +lemma GNil_append: + assumes "GNil = G1@G2" + shows "G1 = GNil \ G2 = GNil" +proof(rule ccontr) + assume "\ (G1 = GNil \ G2 = GNil)" + hence "G1@G2 \ GNil" using append_g.simps by (metis \.distinct(1) \.exhaust) + thus False using assms by auto +qed + +lemma GCons_eq_append_conv: + fixes xs::\ + shows "x#\<^sub>\xs = ys@zs = (ys = GNil \ x#\<^sub>\xs = zs \ (\ys'. x#\<^sub>\ys' = ys \ xs = ys'@zs))" + by(cases ys) auto + +lemma dclist_distinct_unique: + assumes "(dc, const) \ set dclist2" and "(cons, const1) \ set dclist2" and "dc=cons" and "distinct (List.map fst dclist2)" + shows "(const) = const1" +proof - + have "(cons, const) = (dc, const1)" + using assms by (metis (no_types, lifting) assms(3) assms(4) distinct.simps(1) distinct.simps(2) empty_iff insert_iff list.set(1) list.simps(15) list.simps(8) list.simps(9) map_of_eq_Some_iff) + thus ?thesis by auto +qed + +lemma fresh_d_fst_d: + assumes "atom u \ \" + shows "u \ fst ` set \" + using assms proof(induct \) + case Nil + then show ?case by auto +next + case (Cons ut \') + obtain u' and t' where *:"ut = (u',t') " by fastforce + hence "atom u \ ut \ atom u \ \'" using fresh_Cons Cons by auto + moreover hence "atom u \ fst ut" using * fresh_Pair[of "atom u" u' t'] Cons by auto + ultimately show ?case using Cons by auto +qed + +lemma bv_not_in_bset_supp: + fixes bv::bv + assumes "bv |\| B" + shows "atom bv \ supp B" +proof - + have *:"supp B = fset (fimage atom B)" + by (metis fimage.rep_eq finite_fset supp_finite_set_at_base supp_fset) + thus ?thesis using assms + using notin_fset by fastforce +qed + +lemma u_fresh_d: + assumes "atom u \ D" + shows "u \ fst ` setD D" + using assms proof(induct D rule: \_induct) + case DNil + then show ?case by auto +next + case (DCons u' t' \') + then show ?case unfolding setD.simps + using fresh_DCons fresh_Pair by (simp add: fresh_Pair fresh_at_base(2)) +qed + +section \Type Definitions\ + +lemma exist_fresh_bv: + fixes tm::"'a::fs" + shows "\bva2 dclist2. AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \ + atom bva2 \ tm" +proof - + obtain bva2::bv where *:"atom bva2 \ (bva, dclist,tyid,tm)" using obtain_fresh by metis + moreover hence "bva2 \ bva" using fresh_at_base by auto + moreover have " dclist = (bva \ bva2) \ (bva2 \ bva) \ dclist" by simp + moreover have "atom bva \ (bva2 \ bva) \ dclist" proof - + have "atom bva2 \ dclist" using * fresh_prodN by auto + hence "atom ((bva2 \ bva) \ bva2) \ (bva2 \ bva) \ dclist" using fresh_eqvt True_eqvt + proof - + have "(bva2 \ bva) \ atom bva2 \ (bva2 \ bva) \ dclist" + by (metis True_eqvt \atom bva2 \ dclist\ fresh_eqvt) (* 62 ms *) + then show ?thesis + by simp (* 125 ms *) + qed + thus ?thesis by auto + qed + ultimately have "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 ((bva2 \ bva ) \ dclist)" + unfolding type_def.eq_iff Abs1_eq_iff by metis + thus ?thesis using * fresh_prodN by metis +qed + +lemma obtain_fresh_bv: + fixes tm::"'a::fs" + obtains bva2::bv and dclist2 where "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \ + atom bva2 \ tm" + using exist_fresh_bv by metis + +section \Function Definitions\ + +lemma fun_typ_flip: + fixes bv1::bv and c::bv + shows "(bv1 \ c) \ AF_fun_typ x1 b1 c1 \1 s1 = AF_fun_typ x1 ((bv1 \ c) \ b1) ((bv1 \ c) \ c1) ((bv1 \ c) \ \1) ((bv1 \ c) \ s1)" + using fun_typ.perm_simps flip_fresh_fresh supp_at_base fresh_def + flip_fresh_fresh fresh_def supp_at_base + by (simp add: flip_fresh_fresh) + +lemma fun_def_eq: + assumes "AF_fundef fa (AF_fun_typ_none (AF_fun_typ xa ba ca \a sa)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))" + shows "f = fa" and "b = ba" and "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \a = [[atom x]]lst. \" and + " [[atom xa]]lst. ca = [[atom x]]lst. c" + using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis + using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis +proof - + have "([[atom xa]]lst. ((ca, \a), sa) = [[atom x]]lst. ((c, \), s))" using assms fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff by auto + thus "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \a = [[atom x]]lst. \" and + " [[atom xa]]lst. ca = [[atom x]]lst. c" using lst_snd lst_fst by metis+ +qed + +lemma fun_arg_unique_aux: + assumes "AF_fun_typ x1 b1 c1 \1' s1' = AF_fun_typ x2 b2 c2 \2' s2'" + shows "\ x1 : b1 | c1 \ = \ x2 : b2 | c2\" +proof - + have " ([[atom x1]]lst. c1 = [[atom x2]]lst. c2)" using fun_def_eq assms by metis + moreover have "b1 = b2" using fun_typ.eq_iff assms by metis + ultimately show ?thesis using \.eq_iff by fast +qed + +lemma fresh_x_neq: + fixes x::x and y::x + shows "atom x \ y = (x \ y)" + using fresh_at_base fresh_def by auto + +lemma obtain_fresh_z3: + fixes tm::"'b::fs" + obtains z::x where "\ x : b | c \ = \ z : b | c[x::=V_var z]\<^sub>c\<^sub>v \ \ atom z \ tm \ atom z \ (x,c)" +proof - + obtain z::x and c'::c where z:"\ x : b | c \ = \ z : b | c' \ \ atom z \ (tm,x,c)" using obtain_fresh_z2 b_of.simps by metis + hence "c' = c[x::=V_var z]\<^sub>c\<^sub>v" proof - + have "([[atom z]]lst. c' = [[atom x]]lst. c)" using z \.eq_iff by metis + hence "c' = (z \ x) \ c" using Abs1_eq_iff[of z c' x c] fresh_x_neq fresh_prodN by fastforce + also have "... = c[x::=V_var z]\<^sub>c\<^sub>v" + using subst_v_c_def flip_subst_v[of z c x] z fresh_prod3 by metis + finally show ?thesis by auto + qed + thus ?thesis using z fresh_prodN that by metis +qed + +lemma u_fresh_v: + fixes u::u and t::v + shows "atom u \ t" + by(nominal_induct t rule:v.strong_induct,auto) + +lemma u_fresh_ce: + fixes u::u and t::ce + shows "atom u \ t" + apply(nominal_induct t rule:ce.strong_induct) + using u_fresh_v pure_fresh + apply (auto simp add: opp.fresh ce.fresh opp.fresh opp.exhaust) + unfolding ce.fresh opp.fresh opp.exhaust by (simp add: fresh_opp_all) + +lemma u_fresh_c: + fixes u::u and t::c + shows "atom u \ t" + by(nominal_induct t rule:c.strong_induct,auto simp add: c.fresh u_fresh_ce) + +lemma u_fresh_g: + fixes u::u and t::\ + shows "atom u \ t" + by(induct t rule:\_induct, auto simp add: u_fresh_b u_fresh_c fresh_GCons fresh_GNil) + +lemma u_fresh_t: + fixes u::u and t::\ + shows "atom u \ t" + by(nominal_induct t rule:\.strong_induct,auto simp add: \.fresh u_fresh_c u_fresh_b) + +lemma b_of_c_of_eq: + assumes "atom z \ \" + shows "\ z : b_of \ | c_of \ z \ = \" + using assms proof(nominal_induct \ avoiding: z rule: \.strong_induct) + case (T_refined_type x1a x2a x3a) + + hence " \ z : b_of \ x1a : x2a | x3a \ | c_of \ x1a : x2a | x3a \ z \ = \ z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \" + using b_of.simps c_of.simps c_of_eq by auto + moreover have "\ z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \ = \ x1a : x2a | x3a \" using T_refined_type \.fresh by auto + ultimately show ?case by auto +qed + +lemma fresh_d_not_in: + assumes "atom u2 \ \'" + shows "u2 \ fst ` setD \'" + using assms proof(induct \' rule: \_induct) + case DNil + then show ?case by simp +next + case (DCons u t \') + hence *: "atom u2 \ \' \ atom u2 \ (u,t)" + by (simp add: fresh_def supp_DCons) + hence "u2 \ fst ` setD \'" using DCons by auto + moreover have "u2 \ u" using * fresh_Pair + by (metis eq_fst_iff not_self_fresh) + ultimately show ?case by simp +qed + +end \ No newline at end of file diff --git a/thys/MiniSail/Typing.thy b/thys/MiniSail/Typing.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Typing.thy @@ -0,0 +1,586 @@ +(*<*) +theory Typing + imports RCLogic WellformedL +begin + (*>*) + +chapter \Type System\ + +text \The MiniSail type system. We define subtyping judgement first and then typing judgement +for the term forms\ + +section \Subtyping\ + +text \ Subtyping is defined on top of refinement constraint logic (RCL). +A subtyping check is converted into an RCL validity check. \ + +inductive subtype :: "\ \ \ \ \ \ \ \ \ \ bool" ("_ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + subtype_baseI: "\ + atom x \ (\, \, \, z,c,z',c') ; + \; \; \ \\<^sub>w\<^sub>f \ z : b | c \; + \; \; \ \\<^sub>w\<^sub>f \ z' : b | c' \; + \; \ ; (x,b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c'[z'::=[x]\<^sup>v]\<^sub>v +\ \ + \; \; \ \ \ z : b | c \ \ \ z' : b | c' \" + +equivariance subtype +nominal_inductive subtype + avoids subtype_baseI: x +proof(goal_cases) + case (1 \ \ \ z b c z' c' x) + then show ?case using fresh_star_def 1 by force +next + case (2 \ \ \ z b c z' c' x) + then show ?case by auto +qed + +inductive_cases subtype_elims: + "\; \; \ \ \ z : b | c \ \ \ z' : b | c' \" + "\; \; \ \ \\<^sub>1 \ \\<^sub>2" + +section \Literals\ + +text \The type synthesised has the constraint that z equates to the literal\ + +inductive infer_l :: "l \ \ \ bool" (" \ _ \ _" [50, 50] 50) where + infer_trueI: " \ L_true \ \ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_falseI: " \ L_false \ \ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_natI: " \ L_num n \ \ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_num n]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_unitI: " \ L_unit \ \ z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_unit]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_bitvecI: " \ L_bitvec bv \ \ z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_bitvec bv]\<^sup>v]\<^sup>c\<^sup>e \" + +nominal_inductive infer_l . +equivariance infer_l + +inductive_cases infer_l_elims[elim!]: + "\ L_true \ \" + "\ L_false \ \" + "\ L_num n \ \" + "\ L_unit \ \" + "\ L_bitvec x \ \" + "\ l \ \" + +lemma infer_l_form2[simp]: + shows "\z. \ l \ (\ z : base_for_lit l | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \)" +proof (nominal_induct l rule: l.strong_induct) + case (L_num x) + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case L_true + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case L_false + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case L_unit + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case (L_bitvec x) + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +qed + +section \Values\ + +inductive infer_v :: "\ \ \ \ \ \ v \ \ \ bool" (" _ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + +infer_v_varI: "\ + \; \ \\<^sub>w\<^sub>f \ ; + Some (b,c) = lookup \ x; + atom z \ x ; atom z \ (\, \, \) +\ \ + \; \; \ \ [x]\<^sup>v \ \ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[x]\<^sup>v]\<^sup>c\<^sup>e \" + +| infer_v_litI: "\ + \; \ \\<^sub>w\<^sub>f \ ; + \ l \ \ +\ \ + \; \; \ \ [l]\<^sup>v \ \" + +| infer_v_pairI: "\ + atom z \ (v1, v2); atom z \ (\, \, \) ; + \; \; \ \ (v1::v) \ t1 ; + \; \ ; \ \ (v2::v) \ t2 +\ \ + \; \; \ \ V_pair v1 v2 \ (\ z : B_pair (b_of t1) (b_of t2) | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \) " + +| infer_v_consI: "\ + AF_typedef s dclist \ set \; + (dc, tc) \ set dclist ; + \; \; \ \ v \ tv ; + \; \; \ \ tv \ tc ; + atom z \ v ; atom z \ (\, \, \) +\ \ + \; \; \ \ V_cons s dc v \ (\ z : B_id s | [[z]\<^sup>v]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \)" + +| infer_v_conspI: "\ + AF_typedef_poly s bv dclist \ set \; + (dc, tc) \ set dclist ; + \; \; \ \ v \ tv; + \; \; \ \ tv \ tc[bv::=b]\<^sub>\\<^sub>b ; + atom z \ (\, \, \, v, b); + atom bv \ (\, \, \, v, b); + \; \ \\<^sub>w\<^sub>f b +\ \ + \; \; \ \ V_consp s dc b v \ (\ z : B_app s b | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_val (V_consp s dc b v)) \)" + +equivariance infer_v +nominal_inductive infer_v + avoids infer_v_conspI: bv and z | infer_v_varI: z | infer_v_pairI: z | infer_v_consI: z +proof(goal_cases) + case (1 \ \ \ b c x z) + hence "atom z \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \" using \.fresh by simp + then show ?case unfolding fresh_star_def using 1 by simp +next + case (2 \ \ \ b c x z) + then show ?case by auto +next + case (3 z v1 v2 \ \ \ t1 t2) + hence "atom z \ \ z : [ b_of t1 , b_of t2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \" using \.fresh by simp + then show ?case unfolding fresh_star_def using 3 by simp +next + case (4 z v1 v2 \ \ \ t1 t2) + then show ?case by auto +next + case (5 s dclist \ dc tc \ \ v tv z) + hence "atom z \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \" using \.fresh b.fresh pure_fresh by auto + moreover have "atom z \ V_cons s dc v" using v.fresh 5 using v.fresh fresh_prodN pure_fresh by metis + then show ?case unfolding fresh_star_def using 5 by simp +next + case (6 s dclist \ dc tc \ \ v tv z) + then show ?case by auto +next + case (7 s bv dclist \ dc tc \ \ v tv b z) + hence "atom bv \ V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis + moreover then have "atom bv \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \" + using \.fresh ce.fresh v.fresh by auto + moreover have "atom z \ V_consp s dc b v" using v.fresh fresh_prodN pure_fresh 7 by metis + moreover then have "atom z \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \" + using \.fresh ce.fresh v.fresh by auto + ultimately show ?case using fresh_star_def 7 by force +next + case (8 s bv dclist \ dc tc \ \ v tv b z) + then show ?case by auto +qed + +inductive_cases infer_v_elims[elim!]: + "\; \; \ \ V_var x \ \" + "\; \; \ \ V_lit l \ \" + "\; \; \ \ V_pair v1 v2 \ \" + "\; \; \ \ V_cons s dc v \ \" + "\; \; \ \ V_pair v1 v2 \ (\ z : b | c \) " + "\; \; \ \ V_pair v1 v2 \ (\ z : [ b1 , b2 ]\<^sup>b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \) " + "\; \; \ \ V_consp s dc b v \ \ " + +inductive check_v :: "\ \ \ \ \ \ v \ \ \ bool" ("_ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + check_v_subtypeI: "\ \; \; \ \ \1 \ \2; \; \; \ \ v \ \1 \ \ \; \ ; \ \ v \ \2" +equivariance check_v +nominal_inductive check_v . + +inductive_cases check_v_elims[elim!]: + "\; \ ; \ \ v \ \" + +section \Expressions\ + +text \ Type synthesis for expressions \ + +inductive infer_e :: "\ \ \ \ \ \ \ \ \ \ e \ \ \ bool" ("_ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50] 50) where + +infer_e_valI: "\ + (\; \; \ \\<^sub>w\<^sub>f \) ; + (\ \\<^sub>w\<^sub>f (\::\)) ; + (\; \; \ \ v \ \) \ \ + \; \; \; \; \ \ (AE_val v) \ \" + +| infer_e_plusI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : B_int | c1 \ ; + \; \; \ \ v2 \ \ z2 : B_int | c2 \; + atom z3 \ (AE_op Plus v1 v2); atom z3 \ \ \ \ + \; \; \; \; \ \ AE_op Plus v1 v2 \ \ z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_leqI: "\ + \; \; \ \\<^sub>w\<^sub>f \; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : B_int | c1 \ ; + \; \; \ \ v2 \ \ z2 : B_int | c2 \; + atom z3 \ (AE_op LEq v1 v2); atom z3 \ \ +\ \ + \; \; \; \; \ \ AE_op LEq v1 v2 \ \ z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_eqI: "\ + \; \; \ \\<^sub>w\<^sub>f \; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : b | c1 \ ; + \; \; \ \ v2 \ \ z2 : b | c2 \; + atom z3 \ (AE_op Eq v1 v2); atom z3 \ \ ; + b \ { B_bool, B_int, B_unit } +\ \ + \; \; \; \; \ \ AE_op Eq v1 v2 \ \ z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_appI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f; + \; \; \ \ v \ \ x : b | c \; + atom x \ (\, \, \, \, \,v , \); + \'[x::=v]\<^sub>v = \ +\ \ + \; \; \; \; \ \ AE_app f v \ \" + +| infer_e_appPI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \ \\<^sub>w\<^sub>f b' ; + Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \' s'))) = lookup_fun \ f; + \; \; \ \ v \ \ x : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \; atom x \ \; + (\'[bv::=b']\<^sub>b[x::=v]\<^sub>v) = \ ; + atom bv \ (\, \, \, \, \, b', v, \) +\ \ + \; \; \; \; \ \ AE_appP f b' v \ \" + +| infer_e_fstI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v \ \ z' : [b1,b2]\<^sup>b | c \; + atom z \ AE_fst v ; atom z \ \ \ \ + \; \; \; \; \ \ AE_fst v \ \ z : b1 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_fst [v]\<^sup>c\<^sup>e)) \" + +| infer_e_sndI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v \ \ z' : [ b1, b2]\<^sup>b | c \; + atom z \ AE_snd v ; atom z \ \ \ \ + \; \; \; \; \ \ AE_snd v \ \ z : b2 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_snd [v]\<^sup>c\<^sup>e)) \" + +| infer_e_lenI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v \ \ z' : B_bitvec | c \; + atom z \ AE_len v ; atom z \ \\ \ + \; \; \; \; \ \ AE_len v \ \ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e)) \" + +| infer_e_mvarI: "\ + \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \\<^sub>w\<^sub>f \; + (u,\) \ setD \ \ \ + \; \; \; \; \ \ AE_mvar u \ \" + +| infer_e_concatI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : B_bitvec | c1 \ ; + \; \; \ \ v2 \ \ z2 : B_bitvec | c2 \; + atom z3 \ (AE_concat v1 v2); atom z3 \ \ \ \ + \; \; \; \; \ \ AE_concat v1 v2 \ \ z3 : B_bitvec | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_splitI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\); + \ ; \ ; \ \ v1 \ \ z1 : B_bitvec | c1 \ ; + \ ; \ ; \ \ v2 \ \ z2 : B_int | (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var z2))) == (CE_val (V_lit L_true)) AND + (CE_op LEq (CE_val (V_var z2)) (CE_len (CE_val (v1)))) == (CE_val (V_lit L_true)) \; + atom z1 \ (AE_split v1 v2); atom z1 \ \; + atom z2 \ (AE_split v1 v2); atom z2 \ \; + atom z3 \ (AE_split v1 v2); atom z3 \ \ +\ \ + \; \; \; \; \ \ (AE_split v1 v2) \ \ z3 : B_pair B_bitvec B_bitvec | + ((CE_val v1) == (CE_concat (CE_fst (CE_val (V_var z3))) (CE_snd (CE_val (V_var z3))))) + AND (((CE_len (CE_fst (CE_val (V_var z3))))) == (CE_val ( v2))) \" + +equivariance infer_e +nominal_inductive infer_e + avoids infer_e_appI: x |infer_e_appPI: bv | infer_e_splitI: z3 and z1 and z2 +proof(goal_cases) + case (1 \ \ \ \ \ f x b c \' s' v \) + moreover hence "atom x \ [ f v ]\<^sup>e" using fresh_prodN pure_fresh e.fresh by force + ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh by simp +next + case (2 \ \ \ \ \ f x b c \' s' v \) + then show ?case by auto +next + case (3 \ \ \ \ \ b' f bv x b c \' s' v \) + moreover hence "atom bv \ AE_appP f b' v" using fresh_prodN pure_fresh e.fresh by force + ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh fresh_Pair by auto +next + case (4 \ \ \ \ \ b' f bv x b c \' s' v \) + then show ?case by auto +next + case (5 \ \ \ \ \ v1 z1 c1 v2 z2 z3) + have "atom z3 \ \ z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \" + using \.fresh by simp + then show ?case unfolding fresh_star_def fresh_prod7 using wfG_fresh_x2 5 by auto +next + case (6 \ \ \ \ \ v1 z1 c1 v2 z2 z3) + then show ?case by auto +qed + +inductive_cases infer_e_elims[elim!]: + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \ z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \ z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \ z3 : B_int | c \" + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \ z3 : b | c \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \ z3 : b | c \" + "\; \; \; \; \ \ (AE_app f v ) \ \" + "\; \; \; \; \ \ (AE_val v) \ \" + "\; \; \; \; \ \ (AE_fst v) \ \" + "\; \; \; \; \ \ (AE_snd v) \ \" + "\; \; \; \; \ \ (AE_mvar u) \ \" + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \ z3 : B_bool | c \" + "\; \; \; \; \ \ (AE_app f v ) \ \[x::=v]\<^sub>\\<^sub>v" + "\; \; \; \; \ \ (AE_op opp v1 v2) \ \" + "\; \; \; \; \ \ (AE_len v) \ \" + "\; \; \; \; \ \ (AE_len v) \ \ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e))\ " + "\; \; \; \; \ \ AE_concat v1 v2 \ \" + "\; \; \; \; \ \ AE_concat v1 v2 \ (\ z : b | c \) " + "\; \; \; \; \ \ AE_concat v1 v2 \ (\ z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v1]\<^sup>c\<^sup>e) \) " + "\; \; \; \; \ \ (AE_appP f b v ) \ \" + "\; \; \; \; \ \ AE_split v1 v2 \ \" + "\; \; \; \; \ \ (AE_op Eq v1 v2) \ \ z3 : b | c \" + "\; \; \; \; \ \ (AE_op Eq v1 v2) \ \ z3 : B_bool | c \" + "\; \; \; \; \ \ (AE_op Eq v1 v2) \ \" +nominal_termination (eqvt) by lexicographic_order + +section \Statements\ + +inductive check_s :: "\ \ \ \ \ \ \ \ \ \ s \ \ \ bool" (" _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50,50] 50) and + check_branch_s :: "\ \ \ \ \ \ \ \ \ \ tyid \ string \ \ \ v \ branch_s \ \ \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50,50] 50) and + check_branch_list :: "\ \ \ \ \ \ \ \ \ \ tyid \ (string * \) list \ v \ branch_list \ \ \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50,50] 50) where + check_valI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \ v \ \'; + \; \; \ \ \' \ \ \ \ + \; \; \; \; \ \ (AS_val v) \ \" + +| check_letI: "\ + atom x \ (\, \, \, \, \, e, \); + atom z \ (x, \, \, \, \, \, e, \, s); + \; \; \; \; \ \ e \ \ z : b | c \ ; + \; \ ; \ ; ((x,b,c[z::=V_var x]\<^sub>v)#\<^sub>\\) ; \ \ s \ \ +\ \ + \; \; \; \; \ \ (AS_let x e s) \ \" + +| check_assertI: "\ + atom x \ (\, \, \, \, \, c, \, s); + \; \ ; \ ; ((x,B_bool,c)#\<^sub>\\) ; \ \ s \ \ ; + \; \; \ \ c; + \; \; \ \\<^sub>w\<^sub>f \ +\ \ + \; \; \; \; \ \ (AS_assert c s) \ \" + +|check_branch_s_branchI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \ ; + \ ; {||} ; GNil \\<^sub>w\<^sub>f const; + atom x \ (\, \, \, \, \, tid, cons , const, v, \); + \; \; \; ((x,b_of const, ([v]\<^sup>c\<^sup>e == [ V_cons tid cons [x]\<^sup>v]\<^sup>c\<^sup>e ) AND (c_of const x))#\<^sub>\\) ; \ \ s \ \ +\ \ + \; \; \; \; \ ; tid ; cons ; const ; v \ (AS_branch cons x s) \ \" + +|check_branch_list_consI: "\ + \; \; \; \; \; tid; cons; const; v \ cs \ \ ; + \; \; \; \; \; tid; dclist; v \ css \ \ +\ \ + \; \; \; \; \ ; tid ; (cons,const)#dclist ; v \ AS_cons cs css \ \" + +|check_branch_list_finalI: "\ + \; \;\; \; \; tid ; cons ; const ; v \ cs \ \ +\ \ + \; \; \; \; \ ; tid ; [(cons,const)] ; v \ AS_final cs \ \" + +| check_ifI: "\ + atom z \ (\, \, \, \, \, v , s1 , s2 , \ ); + (\; \; \ \ v \ (\ z : B_bool | TRUE \)) ; + \; \; \; \; \ \ s1 \ (\ z : b_of \ | ([v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \ z) \) ; + \; \; \; \; \ \ s2 \ (\ z : b_of \ | ([v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \ z) \) +\ \ + \; \; \; \; \ \ IF v THEN s1 ELSE s2 \ \" + +| check_let2I: "\ + atom x \ (\, \, \, G, \, t, s1, \) ; + \; \ ; \ ; G; \ \ s1 \ t; + \; \ ; \ ; ((x,b_of t,c_of t x)#\<^sub>\G) ; \ \ s2 \ \ +\ \ + \; \ ; \ ; G ; \ \ (LET x : t = s1 IN s2) \ \" + +| check_varI: "\ + atom u \ (\, \, \, \, \, \', v, \) ; + \; \; \ \ v \ \'; + \; \; \; \ ; ((u,\') #\<^sub>\ \) \ s \ \ +\ \ + \; \; \; \; \ \ (VAR u : \' = v IN s) \ \" + +| check_assignI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \ ; + (u,\) \ setD \ ; + \; \; \ \ v \ \; + \; \; \ \ (\ z : B_unit | TRUE \) \ \' +\ \ + \; \; \; \; \ \ (u ::= v) \ \'" + +| check_whileI: "\ + \; \; \; \; \ \ s1 \ \ z : B_bool | TRUE \; + \; \; \; \; \ \ s2 \ \ z : B_unit | TRUE \; + \; \; \ \ (\ z : B_unit | TRUE \) \ \' +\ \ + \; \; \; \; \ \ WHILE s1 DO { s2 } \ \'" + +| check_seqI: "\ + \; \; \; \; \ \ s1 \ \ z : B_unit | TRUE \; + \; \; \; \; \ \ s2 \ \ +\ \ + \; \; \; \; \ \ s1 ;; s2 \ \" + +| check_caseI: "\ + \; \; \; \; \; tid ; dclist ; v \ cs \ \ ; + (AF_typedef tid dclist ) \ set \ ; + \; \; \ \ v \ \ z : B_id tid | TRUE \; + \\<^sub>w\<^sub>f \ +\ \ + \; \; \; \; \ \ AS_match v cs \ \" + +equivariance check_s + +text \We only need avoidance for cases where a variable is added to a context\ +nominal_inductive check_s + avoids check_letI: x and z | check_branch_s_branchI: x | check_let2I: x | check_varI: u | check_ifI: z | check_assertI: x +proof(goal_cases) + case (1 x \ \ \ \ \ e \ z s b c) + hence "atom x \ AS_let x e s" using s_branch_s_branch_list.fresh(2) by auto + moreover have "atom z \ AS_let x e s" using s_branch_s_branch_list.fresh(2) 1 fresh_prod8 by auto + then show ?case using fresh_star_def 1 by force +next + case (3 x \ \ \ \ \ c \ s) + hence "atom x \ AS_assert c s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto + then show ?case using fresh_star_def 3 by force +next + case (5 \ \ \ \ \ const x \ tid cons v s) + hence "atom x \ AS_branch cons x s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto + then show ?case using fresh_star_def 5 by force +next + case (7 z \ \ \ \ \ v s1 s2 \) + hence "atom z \ AS_if v s1 s2 " using s_branch_s_branch_list.fresh by auto + then show ?case using 7 fresh_prodN fresh_star_def by fastforce +next + case (9 x \ \ \ G \ t s1 \ s2) + hence "atom x \ AS_let2 x t s1 s2" using s_branch_s_branch_list.fresh by auto + thus ?case using fresh_star_def 9 by force +next + case (11 u \ \ \ \ \ \' v \ s) + hence "atom u \ AS_var u \' v s" using s_branch_s_branch_list.fresh by auto + then show ?case using fresh_star_def 11 by force + +qed(auto+) + +inductive_cases check_s_elims[elim!]: + "\; \; \; \; \ \ AS_val v \ \" + "\; \; \; \; \ \ AS_let x e s \ \" + "\; \; \; \; \ \ AS_if v s1 s2 \ \" + "\; \; \; \; \ \ AS_let2 x t s1 s2 \ \" + "\; \; \; \; \ \ AS_while s1 s2 \ \" + "\; \; \; \; \ \ AS_var u t v s \ \" + "\; \; \; \; \ \ AS_seq s1 s2 \ \" + "\; \; \; \; \ \ AS_assign u v \ \" + "\; \; \; \; \ \ AS_match v cs \ \" + "\; \; \; \; \ \ AS_assert c s \ \" + +inductive_cases check_branch_s_elims[elim!]: + "\; \; \; \; \; tid ; dclist ; v \ (AS_final cs) \ \" + "\; \; \; \; \; tid ; dclist ; v \ (AS_cons cs css) \ \" + "\; \; \; \; \; tid ; cons ; const ; v \ (AS_branch dc x s ) \ \" + +section \Programs\ + +text \Type check function bodies\ + +inductive check_funtyp :: "\ \ \ \ \ \ fun_typ \ bool" ( " _ ; _ ; _ \ _ " ) where + check_funtypI: "\ + atom x \ (\, \, B , b ); + \; \ ; B ; ((x,b,c) #\<^sub>\ GNil) ; []\<^sub>\ \ s \ \ +\ \ + \; \ ; B \ (AF_fun_typ x b c \ s)" + +equivariance check_funtyp +nominal_inductive check_funtyp + avoids check_funtypI: x +proof(goal_cases) + case (1 x \ \ B b c s \ ) + hence "atom x \ (AF_fun_typ x b c \ s)" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp + then show ?case using fresh_star_def 1 fresh_prodN by fastforce +next + case (2 \ \ x b c s \ f) + then show ?case by auto +qed + +inductive check_funtypq :: "\ \ \ \ fun_typ_q \ bool" ( " _ ; _ \ _ " ) where + check_fundefq_simpleI: "\ + \; \ ; {||} \ (AF_fun_typ x b c t s) +\ \ + \; \ \ ((AF_fun_typ_none (AF_fun_typ x b c t s)))" + +|check_funtypq_polyI: "\ + atom bv \ (\, \, (AF_fun_typ x b c t s)); + \; \; {|bv|} \ (AF_fun_typ x b c t s) +\ \ + \; \ \ (AF_fun_typ_some bv (AF_fun_typ x b c t s))" + +equivariance check_funtypq +nominal_inductive check_funtypq + avoids check_funtypq_polyI: bv +proof(goal_cases) + case (1 bv \ \ x b c t s ) + hence "atom bv \ (AF_fun_typ_some bv (AF_fun_typ x b c t s))" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp + thus ?case using fresh_star_def 1 fresh_prodN by fastforce +next + case (2 bv \ \ ft ) + then show ?case by auto +qed + +inductive check_fundef :: "\ \ \ \ fun_def \ bool" ( " _ ; _ \ _ " ) where + check_fundefI: "\ + \; \ \ ft +\ \ + \; \ \ (AF_fundef f ft)" + +equivariance check_fundef +nominal_inductive check_fundef . + +text \Temporarily remove this simproc as it produces untidy eliminations\ +declare[[ simproc del: alpha_lst]] + +inductive_cases check_funtyp_elims[elim!]: + "check_funtyp \ \ B ft" + +inductive_cases check_funtypq_elims[elim!]: + "check_funtypq \ \ (AF_fun_typ_none (AF_fun_typ x b c \ s))" + "check_funtypq \ \ (AF_fun_typ_some bv (AF_fun_typ x b c \ s))" + +inductive_cases check_fundef_elims[elim!]: + "check_fundef \ \ (AF_fundef f ftq)" + +declare[[ simproc add: alpha_lst]] + +nominal_function \_of :: "var_def list \ \" where + "\_of [] = DNil" +| "\_of ((AV_def u t v)#vs) = (u,t) #\<^sub>\ (\_of vs)" + apply auto + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust + by (metis var_def.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +inductive check_prog :: "p \ \ \ bool" ( "\ _ \ _") where + "\ + \; \; {||}; GNil ; \_of \ \ s \ \ +\ \ \ (AP_prog \ \ \ s) \ \" + +inductive_cases check_prog_elims[elim!]: + "\ (AP_prog \ \ \ s) \ \" + +end \ No newline at end of file diff --git a/thys/MiniSail/TypingL.thy b/thys/MiniSail/TypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/TypingL.thy @@ -0,0 +1,2800 @@ +(*<*) +theory TypingL + imports Typing RCLogicL "HOL-Eisbach.Eisbach" +begin + (*>*) + +chapter \Typing Lemmas\ + +section \Prelude\ + +text \Needed as the typing elimination rules give us facts for an alpha-equivalent version of a term + and so need to be able to 'jump back' to a typing judgement for the orginal term\ + +lemma \_fresh_c[simp]: + assumes "atom x \ \ z : b | c \" and "atom z \ x" + shows "atom x \ c" + using \.fresh assms fresh_at_base + by (simp add: fresh_at_base(2)) + +lemmas subst_defs = subst_b_b_def subst_b_c_def subst_b_\_def subst_v_v_def subst_v_c_def subst_v_\_def + +lemma wfT_wfT_if1: + assumes "wfT \ \ \ (\ z : b_of t | CE_val v == CE_val (V_lit L_false) IMP c_of t z \)" and "atom z \ (\,t)" + shows "wfT \ \ \ t" + using assms proof(nominal_induct t avoiding: \ z rule: \.strong_induct) + case (T_refined_type z' b' c') + show ?case proof(rule wfT_wfT_if) + show \ \; \; \ \\<^sub>w\<^sub>f \ z : b' | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c'[z'::=[ z]\<^sup>v]\<^sub>c\<^sub>v \ \ + using T_refined_type b_of.simps c_of.simps subst_defs by metis + show \atom z \ (c', \)\ using T_refined_type fresh_prodN \_fresh_c by metis + qed +qed + +lemma fresh_u_replace_true: + fixes bv::bv and \::\ + assumes "atom bv \ \' @ (x, b, c) #\<^sub>\ \" + shows "atom bv \ \' @ (x, b, TRUE) #\<^sub>\ \" + using fresh_append_g fresh_GCons assms fresh_Pair c.fresh(1) by auto + +lemma wf_replace_true1: + 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 \\<^sub>w\<^sub>f v : b' \ G = \' @ (x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f v : b'" and + "\; \; G \\<^sub>w\<^sub>f c'' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f c''" and + "\ ; \ \\<^sub>w\<^sub>f G \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ \\<^sub>w\<^sub>f \' @ ((x, b, TRUE) #\<^sub>\ \) " and + "\; \; G \\<^sub>w\<^sub>f \ \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^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 + "\ ; \ ; G \\<^sub>w\<^sub>f ce : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^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 + arbitrary: \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfB_intI \ \) + then show ?case using wf_intros by metis +next + case (wfB_boolI \ \) + then show ?case using wf_intros by metis +next + case (wfB_unitI \ \) + then show ?case using wf_intros by metis +next + case (wfB_bitvecI \ \) + then show ?case using wf_intros by metis +next + case (wfB_pairI \ \ b1 b2) + then show ?case using wf_intros by metis +next + case (wfB_consI \ s dclist \) + then show ?case using wf_intros by metis +next + case (wfB_appI \ b s bv dclist \) + then show ?case using wf_intros by metis +next + case (wfV_varI \ \ \'' b' c x') + hence wfg: \ \ ; \ \\<^sub>w\<^sub>f \' @ (x, b, TRUE) #\<^sub>\ \ \ by auto + show ?case proof(cases "x=x'") + case True + hence "Some (b, TRUE) = lookup (\' @ (x, b, TRUE) #\<^sub>\ \) x'" using lookup.simps lookup_inside_wf wfg by simp + thus ?thesis using Wellformed.wfV_varI[OF wfg] + by (metis True lookup_inside_wf old.prod.inject option.inject wfV_varI.hyps(1) wfV_varI.hyps(3) wfV_varI.prems) + next + case False + hence "Some (b', c) = lookup (\' @ (x, b, TRUE) #\<^sub>\ \) x'" using lookup_inside2 wfV_varI by metis + then show ?thesis using Wellformed.wfV_varI[OF wfg] + by (metis wfG_elim2 wfG_suffix wfV_varI.hyps(1) wfV_varI.hyps(2) wfV_varI.hyps(3) + wfV_varI.prems Wellformed.wfV_varI wf_replace_inside(1)) + qed +next + case (wfV_litI \ \ \ l) + then show ?case using wf_intros using wf_intros by metis +next + case (wfV_pairI \ \ \ v1 b1 v2 b2) + then show ?case using wf_intros by metis +next + case (wfV_consI s dclist \ dc x b' c \ \ v) + then show ?case using wf_intros by metis +next + case (wfV_conspI s bv dclist \ dc xc bc cc \ b' \'' v) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using wfV_conspI by metis + show \(dc, \ xc : bc | cc \) \ set dclist\ using wfV_conspI by metis + show \ \ ;\ \\<^sub>w\<^sub>f b' \ using wfV_conspI by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : bc[bv::=b']\<^sub>b\<^sub>b \ using wfV_conspI by metis + have "atom bv \ \' @ (x, b, TRUE) #\<^sub>\ \" using fresh_u_replace_true wfV_conspI by metis + thus \atom bv \ (\, \, \' @ (x, b, TRUE) #\<^sub>\ \, b', v)\ using wfV_conspI fresh_prodN by metis + qed +next + case (wfCE_valI \ \ \ v b) + then show ?case using wf_intros by metis +next + case (wfCE_plusI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_leqI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_eqI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_fstI \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfCE_sndI \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfCE_concatI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_lenI \ \ \ v1) + then show ?case using wf_intros by metis +next + case (wfTI z \ \ \'' b' c') + show ?case proof + show \atom z \ (\, \, \' @ (x, b, TRUE) #\<^sub>\ \)\ using wfTI fresh_append_g fresh_GCons fresh_prodN by auto + show \ \ ; \ \\<^sub>w\<^sub>f b' \ using wfTI by metis + show \ \; \; (z, b', TRUE) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c' \ using wfTI append_g.simps by metis + qed +next + case (wfC_eqI \ \ \ e1 b e2) + then show ?case using wf_intros by metis +next + case (wfC_trueI \ \ \) + then show ?case using wf_intros by metis +next + case (wfC_falseI \ \ \) + then show ?case using wf_intros by metis +next + case (wfC_conjI \ \ \ c1 c2) + then show ?case using wf_intros by metis +next + case (wfC_disjI \ \ \ c1 c2) + then show ?case using wf_intros by metis +next + case (wfC_notI \ \ \ c1) + then show ?case using wf_intros by metis +next + case (wfC_impI \ \ \ c1 c2) + then show ?case using wf_intros by metis +next + case (wfG_nilI \ \) + then show ?case using GNil_append by blast +next + case (wfG_cons1I c \ \ \'' x b) + then show ?case using wf_intros wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix + by (metis (no_types, lifting)) +next + case (wfG_cons2I c \ \ \'' x' b) + then show ?case using wf_intros + by (metis wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix) +next + case wfTh_emptyI + then show ?case using wf_intros by metis +next + case (wfTh_consI tdef \) + then show ?case using wf_intros by metis +next + case (wfTD_simpleI \ lst s) + then show ?case using wf_intros by metis +next + case (wfTD_poly \ bv lst s) + then show ?case using wf_intros by metis +next + case (wfTs_nil \ \ \) + then show ?case using wf_intros by metis +next + case (wfTs_cons \ \ \ \ dc ts) + then show ?case using wf_intros by metis +qed + +lemma wf_replace_true2: + 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>\ \); D \\<^sub>w\<^sub>f e : b'" and + "\ ; \ ; \ ; G ; \ \\<^sub>w\<^sub>f s : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) ; \ \\<^sub>w\<^sub>f s : b'" and + "\ ; \ ; \ ; G ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b'" and + "\ ; \ ; \ ; G ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b'" 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 \" 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 + arbitrary: \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + + case (wfE_valI \ \ \ \ \ v b) + then show ?case using wf_intros using wf_intros wf_replace_true1 by metis +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_leqI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_appI \ \ \ \ \ f x b c \ s v) + then show ?case using wf_intros wf_replace_true1 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 wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by metis + show \ \ ; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by metis + have "atom bv \ \' @ (x, b, TRUE) #\<^sub>\ \" using fresh_u_replace_true wfE_appPI fresh_prodN by metis + thus \atom bv \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, 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 x1 b1 c1 \ s))) = lookup_fun \ f\ using wfE_appPI by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \ using wfE_appPI wf_replace_true1 by metis + qed +next + case (wfE_mvarI \ \ \ \ \ u \) + then show ?case using wf_intros wf_replace_true1 by metis +next + + case (wfS_valI \ \ \ \ v b \) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfS_letI \ \ \ \'' \ e b' x1 s b1) + show ?case proof + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI wf_replace_true1 by metis + have \ \ ; \ ; \ ; ((x1, b', TRUE) #\<^sub>\ \') @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b1 \ apply(rule wfS_letI(4)) + using wfS_letI append_g.simps by simp + thus \ \ ; \ ; \ ; (x1, b', TRUE) #\<^sub>\ \'@ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b1 \ using append_g.simps by auto + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_letI by metis + show "atom x1 \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, e, b1)" using fresh_append_g fresh_GCons fresh_prodN wfS_letI by auto + qed +next + case (wfS_assertI \ \ \ x' c \'' \ s b') + show ?case proof + show \ \ ; \ ; \ ; (x', B_bool, c) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b' \ + using wfS_assertI (2)[of "(x', B_bool, c) #\<^sub>\ \'" \] wfS_assertI by simp + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c \ using wfS_assertI wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_assertI by metis + show \atom x' \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, c, b', s)\ using wfS_assertI fresh_prodN by simp + qed +next + case (wfS_let2I \ \ \ \'' \ s1 \ x' s2 ba') + show ?case proof + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_let2I wf_replace_true1 by metis + have \ \ ; \ ; \ ; ((x', b_of \, TRUE) #\<^sub>\ \') @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : ba' \ + apply(rule wfS_let2I(5)) + using wfS_let2I append_g.simps by auto + thus \ \ ; \ ; \ ; (x', b_of \, TRUE) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : ba' \ using wfS_let2I append_g.simps by auto + show \atom x' \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, s1, ba', \)\ using fresh_append_g fresh_GCons fresh_prodN wfS_let2I by auto + qed +next + case (wfS_ifI \ \ \ v \ \ s1 b s2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfS_varI \ \ \'' \ v u \ \ b' s) + show ?case proof + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_varI wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI wf_replace_true1 by metis + show \atom u \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, \, v, b')\ using wfS_varI u_fresh_g fresh_prodN by auto + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>f s : b' \ using wfS_varI by metis + qed + +next + case (wfS_assignI u \ \ \ \ \ \ v) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfS_whileI \ \ \ \ \ s1 s2 b) + then show ?case using wf_intros wf_replace_true1 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') + show ?case proof + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : B_id tid \ using wfS_matchI wf_replace_true1 by auto + show \AF_typedef tid dclist \ set \\ using wfS_matchI by auto + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f cs : b' \ using wfS_matchI by auto + qed +next + case (wfS_branchI \ \ \ x' \ \'' \ s b' tid dc) + show ?case proof + have \ \ ; \ ; \ ; ((x', b_of \, TRUE) #\<^sub>\ \') @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b' \ using wfS_branchI append_g.simps by metis + thus \ \ ; \ ; \ ; (x', b_of \, TRUE) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b' \ using wfS_branchI append_g.simps append_g.simps by metis + show \atom x' \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \)\ using wfS_branchI by auto + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_branchI by auto + qed +next + case (wfS_finalI \ \ \ \ \ tid dc t cs b) + then show ?case using wf_intros by metis +next + case (wfS_cons \ \ \ \ \ tid dc t cs b dclist css) + then show ?case using wf_intros by metis +next + case (wfD_emptyI \ \ \) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfD_cons \ \ \ \ \ u) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfPhi_emptyI \) + then show ?case using wf_intros by metis +next + case (wfPhi_consI f \ \ ft) + then show ?case using wf_intros by metis +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 \) + then show ?case using wf_intros by metis +qed + +lemmas wf_replace_true = wf_replace_true1 wf_replace_true2 + +section \Subtyping\ + +lemma subtype_reflI2: + fixes \::\ + assumes "\; \; \ \\<^sub>w\<^sub>f \" + shows "\; \; \ \ \ \ \" +proof - + obtain z b c where *:"\ = \ z : b | c \ \ atom z \ (\,\,\) \ \; \; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c" + using wfT_elims(1)[OF assms] by metis + obtain x::x where **: "atom x \ (\, \, \, c, z ,c, z , c )" using obtain_fresh by metis + have "\; \; \ \ \ z : b | c \ \ \ z : b | c \" proof + show "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" using * assms by auto + show "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" using * assms by auto + show "atom x \ (\, \, \, z , c , z , c )" using fresh_prod6 fresh_prod5 ** by metis + thus "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ \ c[z::=V_var x]\<^sub>v " using wfT_wfC_cons assms * ** subst_v_c_def by simp + qed + thus ?thesis using * by auto +qed + +lemma subtype_reflI: + assumes "\ z1 : b | c1 \ = \ z2 : b | c2 \" and wf1: "\; \;\ \\<^sub>w\<^sub>f (\ z1 : b | c1 \)" + shows "\; \; \ \ (\ z1 : b | c1 \) \ (\ z2 : b | c2 \)" + using assms subtype_reflI2 by metis + +nominal_function base_eq :: "\ \ \ \ \ \ bool" where + "base_eq _ \ z1 : b1| c1 \ \ z2 : b2 | c2 \ = (b1 = b2)" + apply(auto,simp add: eqvt_def base_eq_graph_aux_def ) + by (meson \.exhaust) +nominal_termination (eqvt) by lexicographic_order + +lemma subtype_wfT: + fixes t1::\ and t2::\ + assumes "\; \; \ \ t1 \ t2" + shows "\; \; \ \\<^sub>w\<^sub>f t1 \ \; \; \ \\<^sub>w\<^sub>f t2" + using assms subtype_elims by metis + +lemma subtype_eq_base: + assumes "\; \; \ \ (\ z1 : b1| c1 \) \ (\ z2 : b2 | c2 \)" + shows "b1=b2" + using subtype.simps assms by auto + +lemma subtype_eq_base2: + assumes "\; \; \ \ t1 \ t2" + shows "b_of t1 = b_of t2" + using assms proof(rule subtype.induct[of \ \ \ t1 t2],goal_cases) + case (1 \ \ z1 b c1 z2 c2 x) + then show ?case using subtype_eq_base by auto +qed + +lemma subtype_wf: + fixes \1::\ and \2::\ + assumes "\; \; \ \ \1 \ \2" + shows "\; \; \ \\<^sub>w\<^sub>f \1 \\; \; \ \\<^sub>w\<^sub>f \2" + using assms +proof(rule subtype.induct[of \ \ \ \1 \2],goal_cases) + case (1 \ \G z1 b c1 z2 c2 x) + then show ?case by blast +qed + +lemma subtype_g_wf: + fixes \1::\ and \2::\ and \::\ + assumes "\; \; \ \ \1 \ \2" + shows "\ ; \\\<^sub>w\<^sub>f \" + using assms +proof(rule subtype.induct[of \ \ \ \1 \2],goal_cases) + case (1 \ \ \ z1 b c1 z2 c2 x) + then show ?case using wfX_wfY by auto +qed + +text \ For when we have a particular y that satisfies the freshness conditions that we want the validity check to use \ + +lemma valid_flip_simple: + assumes "\; \; (z, b, c) #\<^sub>\ \ \ c'" and "atom z \ \" and "atom x \ (z, c, z, c', \)" + shows "\; \; (x, b, (z \ x ) \ c) #\<^sub>\ \ \ (z \ x ) \ c'" +proof - + have "(z \ x ) \ \; \; (z \ x ) \ ((z, b, c) #\<^sub>\ \) \ (z \ x ) \ c'" + using True_eqvt valid.eqvt assms beta_flip_eq wfX_wfY by metis + moreover have " \\<^sub>w\<^sub>f \" using valid.simps wfC_wf wfG_wf assms by metis + ultimately show ?thesis + using theta_flip_eq G_cons_flip_fresh3[of x \ z b c] assms fresh_Pair flip_commute by metis +qed + +lemma valid_wf_all: + assumes " \; \; (z0,b,c0)#\<^sub>\G \ c" + shows "wfG \ \ G" and "wfC \ \ ((z0,b,c0)#\<^sub>\G) c" and "atom z0 \ G" + using valid.simps wfC_wf wfG_cons assms by metis+ + +lemma valid_wfT: + fixes z::x + assumes " \; \; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\G \ c[z::=V_var z0]\<^sub>v" and "atom z0 \ (\, \, G,c,c0)" + shows "\; \; G \\<^sub>w\<^sub>f \ z : b | c0 \" and "\; \; G \\<^sub>w\<^sub>f \ z : b | c \" +proof - + have "atom z0 \ c0" using assms fresh_Pair by auto + moreover have *:" \ ; \ \\<^sub>w\<^sub>f (z0,b,c0[z::=V_var z0]\<^sub>c\<^sub>v)#\<^sub>\G" using valid_wf_all wfX_wfY assms subst_v_c_def by metis + ultimately show wft: "\; \; G \\<^sub>w\<^sub>f \ z : b | c0 \" using wfG_wfT[OF *] by auto + + have "atom z0 \ c" using assms fresh_Pair by auto + moreover have wfc: "\; \; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\G \\<^sub>w\<^sub>f c[z::=V_var z0]\<^sub>v" using valid_wf_all assms by metis + have "\; \; G \\<^sub>w\<^sub>f \ z0 : b | c[z::=V_var z0]\<^sub>v \" proof + show \atom z0 \ (\, \, G)\ using assms fresh_prodN by simp + show \ \ ; \ \\<^sub>w\<^sub>f b \ using wft wfT_wfB by force + show \ \; \; (z0, b, TRUE) #\<^sub>\ G \\<^sub>w\<^sub>f c[z::=[ z0 ]\<^sup>v]\<^sub>v \ using wfc wfC_replace_inside[OF wfc, of GNil z0 b "c0[z::=[ z0 ]\<^sup>v]\<^sub>v" G C_true] wfC_trueI + append_g.simps + by (metis "local.*" wfG_elim2 wf_trans(2)) + qed + moreover have "\ z0 : b | c[z::=V_var z0]\<^sub>v \ = \ z : b | c \" using \atom z0 \ c0\ \.eq_iff Abs1_eq_iff(3) + using calculation(1) subst_v_c_def by auto + ultimately show "\; \; G \\<^sub>w\<^sub>f \ z : b | c \" by auto +qed + +lemma valid_flip: + fixes c::c and z::x and z0::x and xx2::x + assumes " \; \; (xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\ \ \ c[z::=V_var xx2]\<^sub>v" and + "atom xx2 \ (c0,\,c,z) " and "atom z0 \ (\,c,z)" + shows " \; \; (z0, b, c0) #\<^sub>\ \ \ c[z::=V_var z0]\<^sub>v" +proof - + + have " \\<^sub>w\<^sub>f \" using assms valid_wf_all wfX_wfY by metis + hence " \ ; \ ; (xx2 \ z0 ) \ ((xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\ \) \ ((xx2 \ z0 ) \ c[z::=V_var xx2]\<^sub>v)" + using valid.eqvt True_eqvt assms beta_flip_eq theta_flip_eq by metis + hence " \; \; (((xx2 \ z0 ) \ xx2, b, (xx2 \ z0 ) \ c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\ (xx2 \ z0 ) \ \) \ ((xx2 \ z0 ) \(c[z::=V_var xx2]\<^sub>v))" + using G_cons_flip[of xx2 z0 xx2 b "c0[z0::=V_var xx2]\<^sub>v" \] by auto + moreover have "(xx2 \ z0 ) \ xx2 = z0" by simp + moreover have "(xx2 \ z0 ) \ c0[z0::=V_var xx2]\<^sub>v = c0" + using assms subst_cv_v_flip[of xx2 c0 z0 "V_var z0"] assms fresh_prod4 by auto + moreover have "(xx2 \ z0 ) \ \ = \" proof - + have "atom xx2 \ \" using assms by auto + moreover have "atom z0 \ \" using assms by auto + ultimately show ?thesis using flip_fresh_fresh by auto + qed + moreover have "(xx2 \ z0 ) \ (c[z::=V_var xx2]\<^sub>v) =c[z::=V_var z0]\<^sub>v" + using subst_cv_v_flip3 assms by simp + ultimately show ?thesis by auto +qed + +lemma subtype_valid: + assumes "\; \; \ \ t1 \ t2" and "atom y \ \" and "t1 = \ z1 : b | c1 \" and "t2 = \ z2 : b | c2 \" + shows "\; \; ((y, b, c1[z1::=V_var y]\<^sub>v) #\<^sub>\ \) \ c2[z2::=V_var y]\<^sub>v" + using assms proof(nominal_induct t2 avoiding: y rule: subtype.strong_induct) + case (subtype_baseI x \ \ \ z c z' c' ba) + + hence "(x \ y) \ \ ; (x \ y) \ \ ; (x \ y) \ ((x, ba, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \) \ (x \ y) \ c'[z'::=[ x ]\<^sup>v]\<^sub>v" using valid.eqvt + using permute_boolI by blast + moreover have " \\<^sub>w\<^sub>f \" using valid.simps wfC_wf wfG_wf subtype_baseI by metis + ultimately have "\; \; ((y, ba, (x \ y) \ c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \) \ (x \ y) \ c'[z'::=[ x ]\<^sup>v]\<^sub>v" + using subtype_baseI theta_flip_eq beta_flip_eq \.eq_iff G_cons_flip_fresh3[of y \ x ba] by (metis flip_commute) + moreover have "(x \ y) \ c[z::=[ x ]\<^sup>v]\<^sub>v = c1[z1::=[ y ]\<^sup>v]\<^sub>v" + by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def) + moreover have "(x \ y) \ c'[z'::=[ x ]\<^sup>v]\<^sub>v = c2[z2::=[ y ]\<^sup>v]\<^sub>v" + by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def) + + ultimately show ?case using subtype_baseI \.eq_iff by metis +qed + +lemma subtype_valid_simple: + assumes "\; \; \ \ t1 \ t2" and "atom z \ \" and "t1 = \ z : b | c1 \" and "t2 = \ z : b | c2 \" + shows "\; \; ((z, b, c1) #\<^sub>\ \) \ c2" + using subst_v_c_def subst_v_id assms subtype_valid[OF assms] by simp + +lemma obtain_for_t_with_fresh: + assumes "atom x \ t" + shows "\b c. t = \ x : b | c \" +proof - + obtain z1 b1 c1 where *:" t = \ z1 : b1 | c1 \ \ atom z1 \ t" using obtain_fresh_z by metis + then have "t = (x \ z1) \ t" using flip_fresh_fresh assms by metis + also have "... = \ (x \ z1) \ z1 : (x \ z1) \ b1 | (x \ z1) \ c1 \" using * assms by simp + also have "... = \ x : b1 | (x \ z1) \ c1 \" using * assms by auto + finally show ?thesis by auto +qed + +lemma subtype_trans: + assumes "\; \; \ \ \1 \ \2" and "\; \; \ \ \2 \ \3" + shows "\; \; \ \ \1 \ \3" + using assms proof(nominal_induct avoiding: \3 rule: subtype.strong_induct) + case (subtype_baseI x \ \ \ z c z' c' b) + hence "b_of \3 = b" using subtype_eq_base2 b_of.simps by metis + then obtain z'' c'' where t3: "\3 = \ z'' : b | c''\ \ atom z'' \ x" + using obtain_fresh_z2 by metis + hence xf: "atom x \ (z'', c'')" using fresh_prodN subtype_baseI \.fresh by auto + have "\; \; \ \ \ z : b | c \ \ \ z'' : b | c''\" + proof(rule Typing.subtype_baseI) + show \atom x \ (\, \, \, z, c, z'', c'')\ using t3 fresh_prodN subtype_baseI xf by simp + show \ \; \; \ \\<^sub>w\<^sub>f \ z : b | c \ \ using subtype_baseI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ z'' : b | c'' \ \ using subtype_baseI t3 subtype_elims by metis + have " \; \; (x, b, c'[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c''[z''::=[ x ]\<^sup>v]\<^sub>v " + using subtype_valid[OF \\; \; \ \ \ z' : b | c' \ \ \3\ , of x z' b c' z'' c''] subtype_baseI + t3 by simp + thus \\; \; (x, b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c''[z''::=[ x ]\<^sup>v]\<^sub>v \ + using valid_trans_full[of \ \ x b c z \ c' z' c'' z'' ] subtype_baseI t3 by simp + qed + thus ?case using t3 by simp +qed + +lemma subtype_eq_e: + assumes "\i s1 s2 G. wfG P \ G \ wfI P G i \ eval_e i e1 s1 \ eval_e i e2 s2 \ s1 = s2" and "atom z1 \ e1" and "atom z2 \ e2" and "atom z1 \ \" and "atom z2 \ \" + and "wfCE P \ \ e1 b" and "wfCE P \ \ e2 b" + shows "P; \; \ \ \ z1 : b | CE_val (V_var z1) == e1 \ \ (\ z2 : b | CE_val (V_var z2) == e2 \)" +proof - + + have "wfCE P \ \ e1 b" and "wfCE P \ \ e2 b" using assms by auto + + have wst1: "wfT P \ \ (\ z1 : b | CE_val (V_var z1) == e1 \)" + using wfC_e_eq wfTI assms wfX_wfB wfG_fresh_x + by (simp add: wfT_e_eq) + + moreover have wst2:"wfT P \ \ (\ z2 : b | CE_val (V_var z2) == e2 \)" + using wfC_e_eq wfX_wfB wfTI assms wfG_fresh_x + by (simp add: wfT_e_eq) + + moreover obtain x::x where xf: "atom x \ (P, \ , z1, CE_val (V_var z1) == e1 , z2, CE_val (V_var z2) == e2 , \)" using obtain_fresh by blast + moreover have vld: "P; \ ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \ (CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v " (is "P; \ ; ?G \ ?c") + proof - + + have wbg: "P; \ \\<^sub>w\<^sub>f ?G \ P ; \ \\<^sub>w\<^sub>f \ \ toSet \ \ toSet ?G" proof - + have "P; \ \\<^sub>w\<^sub>f ?G" proof(rule wfG_consI) + show "P; \ \\<^sub>w\<^sub>f \" using assms wfX_wfY by metis + show "atom x \ \" using xf by auto + show "P; \ \\<^sub>w\<^sub>f b " using assms(6) wfX_wfB by auto + show "P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v " + using wfC_e_eq[OF assms(6)] wf_subst(2) + by (simp add: \atom x \ \\ assms(2) subst_v_c_def) + qed + moreover hence "P; \ \\<^sub>w\<^sub>f \" using wfG_elims by metis + ultimately show ?thesis using toSet.simps by auto + qed + + have wsc: "wfC P \ ?G ?c" proof - + have "wfCE P \ ?G (CE_val (V_var x)) b" proof + show \ P; \ ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f V_var x : b \ using wfV_varI lookup.simps wbg by auto + qed + moreover have "wfCE P \ ?G e2 b" using wf_weakening assms wbg by metis + ultimately have "wfC P \ ?G (CE_val (V_var x) == e2 )" using wfC_eqI by simp + thus ?thesis using subst_cv.simps(6) \atom z2 \ e2\ subst_v_c_def by simp + qed + + moreover have "\i. wfI P ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI , rule impI) + fix i + assume as: "wfI P ?G i \ is_satis_g i ?G" + hence "is_satis i ((CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v)" + by (simp add: is_satis_g.simps(2)) + hence "is_satis i (CE_val (V_var x) == e1 )" using subst_cv.simps assms subst_v_c_def by auto + then obtain s1 and s2 where *:"eval_e i (CE_val (V_var x)) s1 \ eval_e i e1 s2 \ s1=s2" using is_satis.simps eval_c_elims by metis + moreover hence "eval_e i e2 s1" proof - + have **:"wfI P ?G i" using as by auto + moreover have "wfCE P \ ?G e1 b \ wfCE P \ ?G e2 b" using assms xf wf_weakening wbg by metis + moreover then obtain s2' where "eval_e i e2 s2'" using assms wfI_wfCE_eval_e ** by metis + ultimately show ?thesis using * assms(1) wfX_wfY by metis + qed + ultimately have "is_satis i (CE_val (V_var x) == e2 )" using is_satis.simps eval_c_eqI by force + thus "is_satis i ((CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v)" using is_satis.simps eval_c_eqI assms subst_cv.simps subst_v_c_def by auto + qed + ultimately show ?thesis using valid.simps by simp + qed + moreover have "atom x \ (P, \, \, z1 , CE_val (V_var z1) == e1, z2, CE_val (V_var z2) == e2 ) " + unfolding fresh_prodN using xf fresh_prod7 \.fresh by fast + ultimately show ?thesis using subtype_baseI[OF _ wst1 wst2 vld] xf by simp +qed + +lemma subtype_eq_e_nil: + assumes "\i s1 s2 G. wfG P \ G \ wfI P G i \ eval_e i e1 s1 \ eval_e i e2 s2 \ s1 = s2" and "supp e1 = {}" and "supp e2 = {}" and "wfTh P" + and "wfCE P \ GNil e1 b" and "wfCE P \ GNil e2 b" and "atom z1 \ GNil" and "atom z2 \ GNil" + shows "P; \ ; GNil \ \ z1 : b | CE_val (V_var z1) == e1 \ \ (\ z2 : b | CE_val (V_var z2) == e2 \)" + apply(rule subtype_eq_e,auto simp add: assms e.fresh) + using assms fresh_def e.fresh supp_GNil by metis+ + +lemma subtype_gnil_fst_aux: + assumes "supp v\<^sub>1 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \ GNil (CE_val v\<^sub>1) b" and "wfCE P \ GNil (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) b" and + "wfCE P \ GNil (CE_val v\<^sub>2) b2" and "atom z1 \ GNil" and "atom z2 \ GNil" + shows "P; \ ; GNil \ (\ z1 : b | CE_val (V_var z1) == CE_val v\<^sub>1 \) \ (\ z2 : b | CE_val (V_var z2) == CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e \)" +proof - + have "\i s1 s2 G . wfG P \ G \ wfI P G i \ eval_e i ( CE_val v\<^sub>1) s1 \ eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2 \ s1 = s2" proof(rule+) + fix i s1 s2 G + assume as: "wfG P \ G \ wfI P G i \ eval_e i (CE_val v\<^sub>1) s1 \ eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2" + hence "wfCE P \ G (CE_val v\<^sub>2) b2" using assms wf_weakening + by (metis empty_subsetI toSet.simps(1)) + then obtain s3 where "eval_e i (CE_val v\<^sub>2) s3" using wfI_wfCE_eval_e as by metis + hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s1 s3)" + by (meson as eval_e_elims(1) eval_v_pairI) + hence "eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1" using eval_e_fstI eval_e_valI by metis + show "s1 = s2" using as eval_e_uniqueness + using \eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\ by auto + qed + thus ?thesis using subtype_eq_e_nil ce.supp assms by auto +qed + +lemma subtype_gnil_snd_aux: + assumes "supp v\<^sub>2 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \ GNil (CE_val v\<^sub>2) b" and + "wfCE P \ GNil (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) b" and + "wfCE P \ GNil (CE_val v\<^sub>1) b2" and "atom z1 \ GNil" and "atom z2 \ GNil" + shows "P; \ ; GNil \ (\ z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \) \ (\ z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \)" +proof - + have "\i s1 s2 G. wfG P \ G \ wfI P G i \ eval_e i ( CE_val v\<^sub>2) s1 \ eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2 \ s1 = s2" proof(rule+) + fix i s1 s2 G + assume as: " wfG P \ G \ wfI P G i \ eval_e i (CE_val v\<^sub>2) s1 \ eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2" + hence "wfCE P \ G (CE_val v\<^sub>1) b2" using assms wf_weakening + by (metis empty_subsetI toSet.simps(1)) + then obtain s3 where "eval_e i (CE_val v\<^sub>1) s3" using wfI_wfCE_eval_e as by metis + hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s3 s1)" + by (meson as eval_e_elims eval_v_pairI) + hence "eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s1" using eval_e_sndI eval_e_valI by metis + show "s1 = s2" using as eval_e_uniqueness + using \eval_e i (CE_snd [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\ by auto + qed + thus ?thesis using assms subtype_eq_e_nil by (simp add: ce.supp ce.supp) +qed + +lemma subtype_gnil_fst: + assumes "\ ; {||} ; GNil \\<^sub>w\<^sub>f [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b" + shows "\ ; {||} ; GNil \ (\ z\<^sub>1 : b | [[z\<^sub>1]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e \) \ + (\ z\<^sub>2 : b | [[z\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1, v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \)" +proof - + obtain b2 where **:" \ ; {||} ; GNil \\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b b2" using wfCE_elims(4)[OF assms ] wfCE_elims by metis + obtain b1' b2' where *:"B_pair b b2 = B_pair b1' b2' \ \ ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>1 : b1' \ \ ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>2 : b2'" + using wfV_elims(3)[OF **] by metis + + show ?thesis proof(rule subtype_gnil_fst_aux) + show \supp v\<^sub>1 = {}\ using * wfV_supp_nil by auto + show \supp (V_pair v\<^sub>1 v\<^sub>2) = {}\ using ** wfV_supp_nil e.supp by metis + show \\\<^sub>w\<^sub>f \\ using assms wfX_wfY by metis + show \\; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>1 : b \ using wfCE_valI "*" by auto + show \\; {||}; GNil \\<^sub>w\<^sub>f CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e : b \ using assms by auto + show \\; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>2 : b2 \using wfCE_valI "*" by auto + show \atom z\<^sub>1 \ GNil\ using fresh_GNil by metis + show \atom z\<^sub>2 \ GNil\ using fresh_GNil by metis + qed +qed + +lemma subtype_gnil_snd: + assumes "wfCE P {||} GNil (CE_snd ([V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e)) b" + shows "P ; {||} ; GNil \ (\ z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \) \ (\ z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \)" +proof - + obtain b1 where **:" P ; {||} ; GNil \\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b1 b " using wfCE_elims assms by metis + obtain b1' b2' where *:"B_pair b1 b = B_pair b1' b2' \ P ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>1 : b1' \ P ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>2 : b2'" using wfV_elims(3)[OF **] by metis + + show ?thesis proof(rule subtype_gnil_snd_aux) + show \supp v\<^sub>2 = {}\ using * wfV_supp_nil by auto + show \supp (V_pair v\<^sub>1 v\<^sub>2) = {}\ using ** wfV_supp_nil e.supp by metis + show \\\<^sub>w\<^sub>f P\ using assms wfX_wfY by metis + show \P; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>1 : b1 \ using wfCE_valI "*" by simp + show \P; {||}; GNil \\<^sub>w\<^sub>f CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e : b \ using assms by auto + show \P; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>2 : b \using wfCE_valI "*" by simp + show \atom z1 \ GNil\ using fresh_GNil by metis + show \atom z2 \ GNil\ using fresh_GNil by metis + qed +qed + +lemma subtype_fresh_tau: + fixes x::x + assumes "atom x \ t1" and "atom x \ \" and "P; \; \ \ t1 \ t2" + shows "atom x \ t2" +proof - + have wfg: "P; \ \\<^sub>w\<^sub>f \" using subtype_wf wfX_wfY assms by metis + have wft: "wfT P \ \ t2" using subtype_wf wfX_wfY assms by blast + hence "supp t2 \ atom_dom \ \ supp \" using wf_supp + using atom_dom.simps by auto + moreover have "atom x \ atom_dom \" using \atom x \ \\ wfG_atoms_supp_eq wfg fresh_def by blast + ultimately show "atom x \ t2" using fresh_def + by (metis Un_iff contra_subsetD x_not_in_b_set) +qed + +lemma subtype_if_simp: + assumes "wfT P \ GNil (\ z1 : b | CE_val (V_lit l ) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \)" and + "wfT P \ GNil (\ z : b | c \)" and "atom z1 \ c" + shows "P; \ ; GNil \ (\ z1 : b | CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \) \ (\ z : b | c \)" +proof - + obtain x::x where xx: "atom x \ ( P , \ , z1, CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c, GNil)" using obtain_fresh_z by blast + hence xx2: " atom x \ (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c, GNil)" using fresh_prod7 fresh_prod3 by fast + have *:"P; \ ; (x, b, (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\ GNil \ c[z::=V_var x]\<^sub>v " (is "P; \ ; ?G \ ?c" ) proof - + have "wfC P \ ?G ?c" using wfT_wfC_cons[OF assms(1) assms(2),of x] xx fresh_prod5 fresh_prod3 subst_v_c_def by metis + moreover have "(\i. wfI P ?G i \ is_satis_g i ?G \ is_satis i ?c)" proof(rule allI, rule impI) + fix i + assume as1: "wfI P ?G i \ is_satis_g i ?G" + have "((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))" + using assms subst_v_c_def by auto + hence "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))" using is_satis_g.simps as1 by presburger + moreover have "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l)))" using is_satis.simps eval_c_eqI[of i "(CE_val (V_lit l))" "eval_l l"] eval_e_uniqueness + eval_e_valI eval_v_litI by metis + ultimately show "is_satis i ?c" using is_satis_mp[of i] by metis + qed + ultimately show ?thesis using valid.simps by simp + qed + moreover have "atom x \ (P, \, GNil, z1 , CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c )" + unfolding fresh_prod5 \.fresh using xx fresh_prodN x_fresh_b by metis + ultimately show ?thesis using subtype_baseI assms xx xx2 by metis +qed + +lemma subtype_if: + assumes "P; \; \ \ \ z : b | c \ \ \ z' : b | c' \" and + "wfT P \ \ (\ z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \)" and + "wfT P \ \ (\ z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \)" and + "atom z1 \ v" and "atom z \ \" and "atom z1 \ c" and "atom z2 \ c'" and "atom z2 \ v" + shows "P; \ ; \ \ \ z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \ \ \ z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \" +proof - + obtain x::x where xx: "atom x \ (P,\,z,c,z', c', z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \)" + using obtain_fresh_z by blast + hence xf: "atom x \ (z, c, z', c', \)" by simp + have xf2: "atom x \ (z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \)" + using xx fresh_prod4 fresh_prodN by metis + + moreover have "P; \ ; (x, b, (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \ (CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v" + (is "P; \ ; ?G \ ?c" ) + proof - + have wbc: "wfC P \ ?G ?c" using assms xx fresh_prod4 fresh_prod2 wfT_wfC_cons assms subst_v_c_def by metis + + moreover have "\i. wfI P ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI, rule impI) + fix i + assume a1: "wfI P ?G i \ is_satis_g i ?G" + + have *: " is_satis i ((CE_val v == CE_val (V_lit l))) \ is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" + proof + assume a2: "is_satis i ((CE_val v == CE_val (V_lit l)))" + + have "is_satis i ((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v)" + using a1 is_satis_g.simps by simp + moreover have "((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v) = (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))" + using assms subst_v_c_def by simp + ultimately have "is_satis i (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))" by argo + + hence "is_satis i ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v)" using a2 is_satis_mp by auto + moreover have "((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((c[z::=V_var x]\<^sub>v ))" using assms by auto + ultimately have "is_satis i ((c[z::=V_var x]\<^sub>v ))" using a2 is_satis.simps by auto + + hence "is_satis_g i ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\ \)" using a1 is_satis_g.simps by meson + moreover have "wfI P ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\ \) i" proof - + obtain s where "Some s = i x \ wfRCV P s b \ wfI P \ i" using wfI_def a1 by auto + thus ?thesis using wfI_def by auto + qed + ultimately have "is_satis i ((c'[z'::=V_var x]\<^sub>v))" using subtype_valid assms(1) xf valid.simps by simp + + moreover have "(c'[z'::=V_var x]\<^sub>v) = ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" using assms by auto + ultimately show "is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" by auto + qed + + moreover have "?c = ((CE_val v == CE_val (V_lit l)) IMP ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v))" + using assms subst_v_c_def by simp + + moreover have "\b1 b2. eval_c i (CE_val v == CE_val (V_lit l) ) b1 \ + eval_c i c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v b2" proof - + + have "wfC P \ ?G (CE_val v == CE_val (V_lit l))" using wbc wfC_elims(7) assms subst_cv.simps subst_v_c_def by fastforce + + moreover have "wfC P \ ?G (c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v)" proof(rule wfT_wfC_cons) + show \ P; \; \ \\<^sub>w\<^sub>f \ z1 : b | CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v) \ \ using assms subst_v_c_def by auto + have " \ z2 : b | c'[z'::=V_var z2]\<^sub>v \ = \ z' : b | c' \" using assms subst_v_c_def by auto + thus \ P; \; \ \\<^sub>w\<^sub>f \ z2 : b | c'[z'::=V_var z2]\<^sub>v \ \ using assms subtype_elims by metis + show \atom x \ (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c'[z'::=V_var z2]\<^sub>v, \)\ using xx fresh_Pair c.fresh by metis + qed + + ultimately show ?thesis using wfI_wfC_eval_c a1 subst_v_c_def by simp + qed + + ultimately show "is_satis i ?c" using is_satis_imp[OF *] by auto + qed + ultimately show ?thesis using valid.simps by simp + qed + moreover have "atom x \ (P, \, \, z1 , CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2 , CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )" + unfolding fresh_prod5 \.fresh using xx xf2 fresh_prodN x_fresh_b by metis + ultimately show ?thesis using subtype_baseI assms xf2 by metis +qed + +lemma eval_e_concat_eq: + assumes "wfI \ \ i" + shows "\s. eval_e i (CE_val (V_lit (L_bitvec (v1 @ v2))) ) s \ eval_e i (CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e) s" + using eval_e_valI eval_e_concatI eval_v_litI eval_l.simps by metis + +lemma is_satis_eval_e_eq_imp: + assumes "wfI \ \ i" and "eval_e i e1 s" and "eval_e i e2 s" + and "is_satis i (CE_val (V_var x) == e1)" (is "is_satis i ?c1") + shows "is_satis i (CE_val (V_var x) == e2)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + hence "eval_e i (CE_val (V_var x)) s" using assms is_satis.simps eval_c_elims + by (metis (full_types) eval_e_uniqueness) + thus ?thesis using is_satis.simps eval_c.intros assms by fastforce +qed + +lemma valid_eval_e_eq: + fixes e1::ce and e2::ce + assumes "\\ i. wfI \ \ i \ (\s. eval_e i e1 s \ eval_e i e2 s)" and "\; \; GNil \\<^sub>w\<^sub>f e1 : b " and "\; \; GNil \\<^sub>w\<^sub>f e2 : b" + shows "\; \; (x, b, (CE_val (V_var x) == e1 )) #\<^sub>\ GNil \ (CE_val (V_var x) == e2) " +proof(rule validI) + show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f CE_val (V_var x) == e2" + proof + have " \ ; \ ; (x, b, TRUE )#\<^sub>\GNil \\<^sub>w\<^sub>f CE_val (V_var x) == e1" using assms wfC_eqI wfE_valI wfV_varI wfX_wfY + by (simp add: fresh_GNil wfC_e_eq) + hence "\ ; \ \\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil" using wfG_consI fresh_GNil wfX_wfY assms wfX_wfB by metis + thus "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f CE_val (V_var x) : b " using wfCE_valI wfV_varI wfX_wfY + lookup.simps assms wfX_wfY by simp + show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f e2 : b " using assms wf_weakening wfX_wfY + by (metis (full_types) \\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f CE_val (V_var x) : b\ empty_iff subsetI toSet.simps(1)) + qed + show " \i. wfI \ ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil) i \ is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil) \ is_satis i (CE_val (V_var x) == e2 )" + proof(rule,rule) + fix i + assume "wfI \ ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil) i \ is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil)" + moreover then obtain s where "eval_e i e1 s \ eval_e i e2 s" using assms by auto + ultimately show "is_satis i (CE_val (V_var x) == e2 )" using assms is_satis_eval_e_eq_imp is_satis_g.simps by meson + qed +qed + +lemma subtype_concat: + assumes " \\<^sub>w\<^sub>f \" + shows "\; \; GNil \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ 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 \ ?t1 \ ?t2") +proof - + obtain x::x where x: "atom x \ (\, \, GNil, z , CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))), + z , 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 "?xfree" ) + using obtain_fresh by auto + + have wb1: "\; \; GNil \\<^sub>w\<^sub>f CE_val (V_lit (L_bitvec (v1 @ v2))) : B_bitvec" using wfX_wfY wfCE_valI wfV_litI assms base_for_lit.simps wfG_nilI by metis + hence wb2: "\; \; GNil \\<^sub>w\<^sub>f CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e : B_bitvec" + using wfCE_concatI wfX_wfY wfV_litI base_for_lit.simps wfCE_valI by metis + + show ?thesis proof + show " \; \; GNil \\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil wb1 wb2 by metis + show " \; \; GNil \\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil wb1 wb2 by metis + show ?xfree using x by auto + show "\; \; (x, B_bitvec, (CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) )[z::=V_var x]\<^sub>v) #\<^sub>\ + GNil \ (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 )[z::=V_var x]\<^sub>v " + using valid_eval_e_eq eval_e_concat_eq wb1 wb2 subst_v_c_def by fastforce + qed +qed + +lemma subtype_len: + assumes " \\<^sub>w\<^sub>f \" + shows "\; \; 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 \" (is "\; \; GNil \ ?t1 \ ?t2") +proof - + + have *: "\ \\<^sub>w\<^sub>f [] \ \; \; GNil \\<^sub>w\<^sub>f []\<^sub>\ " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto + obtain x::x where x: "atom x \ (\, \, GNil, z' , CE_val (V_var z') == + CE_val (V_lit (L_num (int (length v)))) , z, CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )" + (is "atom x \ ?F") + using obtain_fresh by metis + then show ?thesis proof + have "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int" + using wfCE_valI * wfV_litI base_for_lit.simps + by (metis wfE_valI wfX_wfY) + + thus "\; \; GNil \\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil by auto + + have "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int" + using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY + by (metis wfCE_lenI wfCE_valI) + + thus "\; \; GNil \\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil by auto + + show "\; \; (x, B_int, (CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) )[z'::=V_var x]\<^sub>v) #\<^sub>\ GNil \ (CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )[z::=V_var x]\<^sub>v" + (is "\; \; ?G \ ?c" ) using valid_len assms subst_v_c_def by auto + qed +qed + +lemma subtype_base_fresh: + assumes "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "\; \; \ \\<^sub>w\<^sub>f \ z : b | c' \ " and + "atom z \ \" and "\; \; (z, b, c) #\<^sub>\ \ \ c'" + shows "\; \; \ \ \ z : b | c \ \ \ z : b | c' \" +proof - + obtain x::x where *:"atom x \ ((\ , \ , z, c, z, c', \) , (\, \, \, \ z : b | c \, \ z : b | c' \))" using obtain_fresh by metis + moreover hence "atom x \ \" using fresh_Pair by auto + moreover hence "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ \ c'[z::=V_var x]\<^sub>v" using assms valid_flip_simple * subst_v_c_def by auto + ultimately show ?thesis using subtype_baseI assms \.fresh fresh_Pair by metis +qed + +lemma subtype_bop_arith: + assumes "wfG \ \ \" and "(opp = Plus \ ll = (L_num (n1+n2))) \ (opp = LEq \ ll = ( if n1\n2 then L_true else L_false))" + and "(opp = Plus \ b = B_int) \ (opp = LEq \ b = B_bool)" + shows "\; \; \ \ (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit (ll))) \) \ + \ z : b | C_eq (CE_val (V_var z)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?T1 \ ?T2") +proof - + obtain x::x where xf: "atom x \ (z, CE_val (V_var z) == CE_val (V_lit (ll)) , z, CE_val (V_var z) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e , \)" + using obtain_fresh by blast + + have "\; \; \ \ (\ x : b | C_eq (CE_val (V_var x)) (CE_val (V_lit (ll))) \) \ + \ x : b | C_eq (CE_val (V_var x)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?S1 \ ?S2") + proof(rule subtype_base_fresh) + + show "atom x \ \" using xf fresh_Pair by auto + + show "wfT \ \ \ (\ x : b | CE_val (V_var x) == CE_val (V_lit ll) \)" (is "wfT \ \ ?A ?B") + proof(rule wfT_e_eq) + have " \; \; \ \\<^sub>w\<^sub>f (V_lit ll) : b" using wfV_litI base_for_lit.simps assms by metis + thus " \; \; \ \\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI by auto + show "atom x \ \" using xf fresh_Pair by auto + qed + + consider "opp = Plus" | "opp = LEq" using opp.exhaust assms by blast + then show "wfT \ \ \ (\ x : b | CE_val (V_var x) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e \)" (is "wfT \ \ ?A ?C") + proof(cases) + case 1 + then show "\ ; \ ; \ \\<^sub>w\<^sub>f \ x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \" + using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms + by (metis \atom x \ \\ wfT_e_eq) + next + case 2 + then show "\ ; \ ; \ \\<^sub>w\<^sub>f \ x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ " + using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms + + by (metis \atom x \ \\ wfCE_leqI wfT_e_eq) + qed + + show "\; \ ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op opp [V_lit (L_num n1)]\<^sup>c\<^sup>e [V_lit (L_num n2)]\<^sup>c\<^sup>e)" (is "\; \; ?G \ ?c") + using valid_arith_bop assms xf by simp + + qed + moreover have "?S1 = ?T1 " using type_l_eq by auto + moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp + by (metis ms_fresh_all(4)) + ultimately show ?thesis by auto + +qed + +lemma subtype_bop_eq: + assumes "wfG \ \ \" and "base_for_lit l1 = base_for_lit l2" + shows "\; \; \ \ (\ z : B_bool | C_eq (CE_val (V_var z)) (CE_val (V_lit (if l1 = l2 then L_true else L_false))) \) \ + \ z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?T1 \ ?T2") +proof - + let ?ll = "if l1 = l2 then L_true else L_false" + obtain x::x where xf: "atom x \ (z, CE_val (V_var z) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) , z, CE_val (V_var z) == CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e , \, (\, \, \))" + using obtain_fresh by blast + + have "\; \; \ \ (\ x : B_bool | C_eq (CE_val (V_var x)) (CE_val (V_lit (?ll))) \) \ + \ x : B_bool | C_eq (CE_val (V_var x)) (CE_op Eq [(V_lit (l1))]\<^sup>c\<^sup>e [(V_lit (l2))]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?S1 \ ?S2") + proof(rule subtype_base_fresh) + + show "atom x \ \" using xf fresh_Pair by auto + + show "wfT \ \ \ (\ x : B_bool | CE_val (V_var x) == CE_val (V_lit ?ll) \)" (is "wfT \ \ ?A ?B") + proof(rule wfT_e_eq) + have " \; \; \ \\<^sub>w\<^sub>f (V_lit ?ll) : B_bool" using wfV_litI base_for_lit.simps assms by metis + thus " \; \; \ \\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI by auto + show "atom x \ \" using xf fresh_Pair by auto + qed + + show " \ ; \ ; \ \\<^sub>w\<^sub>f \ x : B_bool | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ " + proof(rule wfT_e_eq) + show "\ ; \ ; \ \\<^sub>w\<^sub>f [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e : B_bool" + apply(rule wfCE_eqI, rule wfCE_valI) + apply(rule wfV_litI, simp add: assms) + using wfV_litI assms wfCE_valI by auto + show "atom x \ \" using xf fresh_Pair by auto + qed + + show "\; \ ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (?ll)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op Eq [V_lit (l1)]\<^sup>c\<^sup>e [V_lit (l2)]\<^sup>c\<^sup>e)" (is "\; \; ?G \ ?c") + using valid_eq_bop assms xf by auto + + qed + moreover have "?S1 = ?T1 " using type_l_eq by auto + moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp + by (metis ms_fresh_all(4)) + ultimately show ?thesis by auto + +qed + +lemma subtype_top: + assumes "wfT \ \ G (\ z : b | c \)" + shows "\; \; G \ (\ z : b | c \) \ (\ z : b | TRUE \)" +proof - + obtain x::x where *: "atom x \ (\, \, G, z , c, z , TRUE)" using obtain_fresh by blast + then show ?thesis proof(rule subtype_baseI) + show \ \; \; G \\<^sub>w\<^sub>f \ z : b | c \ \ using assms by auto + show \ \; \;G \\<^sub>w\<^sub>f \ z : b | TRUE \ \ using wfT_TRUE assms wfX_wfY b_of.simps wfT_wf + by (metis wfX_wfB(8)) + hence "\ ; \ \\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ G" using wfT_wf_cons3 assms fresh_Pair * subst_v_c_def by auto + thus \\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ G \ (TRUE)[z::=V_var x]\<^sub>v \ using valid_trueI subst_cv.simps subst_v_c_def by metis + qed +qed + +lemma if_simp: + "(if x = x then e1 else e2) = e1" + by auto + +lemma subtype_split: + assumes "split n v (v1,v2)" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; GNil \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec + v1 ]\<^sup>v , [ L_bitvec + v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \ \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ L_bitvec + v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num + n ]\<^sup>v ]\<^sup>c\<^sup>e \" + (is "\ ;?B ; GNil \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c1 \ \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c2 \") +proof - + obtain x::x where xf:"atom x \ (\, ?B, GNil, z, ?c1,z, ?c2 )" using obtain_fresh by auto + then show ?thesis proof(rule subtype_baseI) + show *: \\ ; ?B ; (x, [ B_bitvec , B_bitvec ]\<^sup>b, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ + GNil \ (?c2)[z::=[ x ]\<^sup>v]\<^sub>v \ + unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp + using valid_split[OF assms, of x] by simp + show \ \ ; ?B ; GNil \\<^sub>w\<^sub>f \ z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c1 \ \ using valid_wfT[OF *] xf fresh_prodN by metis + show \ \ ; ?B ; GNil \\<^sub>w\<^sub>f \ z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c2 \ \ using valid_wfT[OF *] xf fresh_prodN by metis + qed +qed + +lemma subtype_range: + fixes n::int and \::\ + assumes "0 \ n \ n \ int (length v)" and "\ ; {||} \\<^sub>w\<^sub>f \" + shows "\ ; {||} ; \ \ \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ \ + \ 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) \" + (is "\ ; ?B ; \ \ \ z : B_int | ?c1 \ \ \ z : B_int | ?c2 AND ?c3 \") +proof - + obtain x::x where *:\atom x \ (\, ?B, \, z, ?c1 , z, ?c2 AND ?c3)\ using obtain_fresh by auto + moreover have **:\\ ; ?B ; (x, B_int, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ (?c2 AND ?c3)[z::=[ x ]\<^sup>v]\<^sub>v\ + unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp using valid_range_length[OF assms(1)] assms fresh_prodN * by simp + + moreover hence \ \ ; ?B ; \ \\<^sub>w\<^sub>f \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ \ using + valid_wfT * fresh_prodN by metis + moreover have \ \ ; ?B ; \ \\<^sub>w\<^sub>f \ z : B_int | ?c2 AND ?c3 \ \ + using valid_wfT[OF **] * fresh_prodN by metis + ultimately show ?thesis using subtype_baseI by auto +qed + +lemma check_num_range: + assumes "0 \ n \ n \ int (length v)" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; 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 assms subtype_range check_v.intros infer_v_litI wfG_nilI + by (meson infer_natI) + +section \Literals\ + +nominal_function type_for_lit :: "l \ \" where + "type_for_lit (L_true) = (\ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_true]\<^sup>c\<^sup>e \)" +| "type_for_lit (L_false) = (\ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_false]\<^sup>c\<^sup>e \)" +| "type_for_lit (L_num n) = (\ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_num n)]\<^sup>c\<^sup>e \)" +| "type_for_lit (L_unit) = (\ z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_unit )]\<^sup>c\<^sup>e \)" (* could have done { z : unit | True } but want the uniformity *) +| "type_for_lit (L_bitvec v) = (\ z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_bitvec v)]\<^sup>c\<^sup>e \)" + by (auto simp: eqvt_def type_for_lit_graph_aux_def, metis l.strong_exhaust,(simp add: permute_int_def flip_bitvec0)+ ) +nominal_termination (eqvt) by lexicographic_order + + +nominal_function type_for_var :: "\ \ \ \ x \ \" where + "type_for_var G \ x = (case lookup G x of + None \ \ + | Some (b,c) \ (\ x : b | c \)) " + apply auto unfolding eqvt_def apply(rule allI) unfolding type_for_var_graph_aux_def eqvt_def by simp +nominal_termination (eqvt) by lexicographic_order + +lemma infer_l_form: + fixes l::l and tm::"'a::fs" + assumes "\ l \ \" + shows "\z b. \ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \) \ atom z \ tm" +proof - + obtain z' and b where t:"\ = (\ z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \)" using infer_l_elims assms using infer_l.simps type_for_lit.simps + type_for_lit.cases by blast + obtain z::x where zf: "atom z \ tm" using obtain_fresh by metis + have "\ = \ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \" using type_e_eq ce.fresh v.fresh l.fresh + by (metis t type_l_eq) + thus ?thesis using zf by auto +qed + +lemma infer_l_form3: + fixes l::l + assumes "\ l \ \" + shows "\z. \ = (\ z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \)" + using infer_l_elims using assms using infer_l.simps type_for_lit.simps base_for_lit.simps by auto + +lemma infer_l_form4[simp]: + fixes \::\ + assumes "\ ; \ \\<^sub>w\<^sub>f \ " + shows "\z. \ l \ (\ z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \)" + using assms infer_l_form2 infer_l_form3 by metis + +lemma infer_v_unit_form: + fixes v::v + assumes "P ; \ ; \ \ v \ (\ z1 : B_unit | c1 \)" and "supp v = {}" + shows "v = V_lit L_unit" + using assms proof(nominal_induct \ v "\ z1 : B_unit | c1 \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ c x z) + then show ?case using supp_at_base by auto +next + case (infer_v_litI \ \ \ l) + from \\ l \ \ z1 : B_unit | c1 \\ show ?case by(nominal_induct "\ z1 : B_unit | c1 \" rule: infer_l.strong_induct,auto) +qed + +lemma base_for_lit_wf: + assumes "\\<^sub>w\<^sub>f \" + shows "\ ; \ \\<^sub>w\<^sub>f base_for_lit l" + using base_for_lit.simps using wfV_elims wf_intros assms l.exhaust by metis + +lemma infer_l_t_wf: + fixes \::\ + assumes "\ ; \ \\<^sub>w\<^sub>f \ \ atom z \ \" + shows "\; \; \ \\<^sub>w\<^sub>f \ z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \" +proof + show "atom z \ (\, \, \)" using wfG_fresh_x assms by auto + show "\ ; \ \\<^sub>w\<^sub>f base_for_lit l" using base_for_lit_wf assms wfX_wfY by metis + thus "\; \; (z, base_for_lit l, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var z) == CE_val (V_lit l)" using wfC_v_eq wfV_litI assms wfX_wfY by metis +qed + +lemma infer_l_wf: + fixes l::l and \::\ and \::\ and \::\ + assumes "\ l \ \" and "\ ; \ \\<^sub>w\<^sub>f \" + shows "\\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f \" +proof - + show *:"\ ; \ \\<^sub>w\<^sub>f \" using assms infer_l_elims by auto + thus "\\<^sub>w\<^sub>f \" using wfX_wfY by auto + show *:"\ ; \ ; \ \\<^sub>w\<^sub>f \" using infer_l_t_wf assms infer_l_form3 * + by (metis \\\<^sub>w\<^sub>f \\ fresh_GNil wfG_nilI wfT_weakening_nil) +qed + +lemma infer_l_uniqueness: + fixes l::l + assumes "\ l \ \" and "\ l \ \'" + shows "\ = \'" + using assms +proof - + obtain z and b where zt: "\ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \)" using infer_l_form assms by blast + obtain z' and b where z't: "\' = (\ z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \)" using infer_l_form assms by blast + thus ?thesis using type_l_eq zt z't assms infer_l.simps infer_l_elims l.distinct + by (metis infer_l_form3) +qed + +section \Values\ + +lemma type_v_eq: + assumes "\ z1 : b1 | c1 \ = \ z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \" and "atom z \ x" + shows "b = b1" and "c1 = C_eq (CE_val (V_var z1)) (CE_val (V_var x))" + using assms by (auto,metis Abs1_eq_iff \.eq_iff assms c.fresh ce.fresh type_e_eq v.fresh) + +lemma infer_var2 [elim]: + assumes "P; \ ; G \ V_var x \ \" + shows "\b c. Some (b,c) = lookup G x" + using assms infer_v_elims lookup_iff by (metis (no_types, lifting)) + +lemma infer_var3 [elim]: + assumes "\; \; \ \ V_var x \ \" + shows "\z b c. Some (b,c) = lookup \ x \ \ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \) \ atom z \ x \ atom z \ (\, \, \)" + using infer_v_elims(1)[OF assms(1)] by metis + +lemma infer_bool_options2: + fixes v::v + assumes "\; \; \ \ v \ \ z : b | c \" and "supp v = {} \ b = B_bool" + shows "v = V_lit L_true \ (v = (V_lit L_false))" + using assms +proof(nominal_induct "\ z : b | c \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ c x z) + then show ?case using v.supp supp_at_base by auto +next + case (infer_v_litI \ \ \ l) + from \ \ l \ \ z : b | c \\ show ?case proof(nominal_induct "\ z : b | c \" rule: infer_l.strong_induct) + case (infer_trueI z) + then show ?case by auto + next + case (infer_falseI z) + then show ?case by auto + next + case (infer_natI n z) + then show ?case using infer_v_litI by simp + next + case (infer_unitI z) + then show ?case using infer_v_litI by simp + next + case (infer_bitvecI bv z) + then show ?case using infer_v_litI by simp + qed +qed(auto+) + +lemma infer_bool_options: + fixes v::v + assumes "\; \; \ \ v \ \ z : B_bool | c \" and "supp v = {}" + shows "v = V_lit L_true \ (v = (V_lit L_false))" + using infer_bool_options2 assms by blast + +lemma infer_int2: + fixes v::v + assumes "\; \; \ \ v \ \ z : b | c \" + shows "supp v = {} \ b = B_int \ (\n. v = V_lit (L_num n))" + using assms +proof(nominal_induct "\ z : b | c \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ c x z) + then show ?case using v.supp supp_at_base by auto +next + case (infer_v_litI \ \ \ l) + from \ \ l \ \ z : b | c \\ show ?case proof(nominal_induct "\ z : b | c \" rule: infer_l.strong_induct) + case (infer_trueI z) + then show ?case by auto + next + case (infer_falseI z) + then show ?case by auto + next + case (infer_natI n z) + then show ?case using infer_v_litI by simp + next + case (infer_unitI z) + then show ?case using infer_v_litI by simp + next + case (infer_bitvecI bv z) + then show ?case using infer_v_litI by simp + qed +qed(auto+) + +lemma infer_bitvec: + fixes \::\ and v::v + assumes "\; \; \ \ v \ \ z' : B_bitvec | c' \" and "supp v = {}" + shows "\bv. v = V_lit (L_bitvec bv)" + using assms proof(nominal_induct v rule: v.strong_induct) + case (V_lit l) + then show ?case by(nominal_induct l rule: l.strong_induct,force+) +next + case (V_consp s dc b v) + then show ?case using infer_v_elims(7)[OF V_consp(2)] \.eq_iff by auto +next + case (V_var x) + then show ?case using supp_at_base by auto +qed(force+) + +lemma infer_int: + assumes "infer_v \ \ \ v (\ z : B_int | c \)" and "supp v= {}" + shows "\n. V_lit (L_num n) = v" + using assms infer_int2 by (metis (no_types, lifting)) + +lemma infer_lit: + assumes "infer_v \ \ \ v (\ z : b | c \)" and "supp v= {}" and "b \ { B_bool , B_int , B_unit }" + shows "\l. V_lit l = v" + using assms proof(nominal_induct v rule: v.strong_induct) + case (V_lit x) + then show ?case by (simp add: supp_at_base) +next + case (V_var x) + then show ?case + by (simp add: supp_at_base) +next + case (V_pair x1a x2a) + then show ?case using supp_at_base by auto +next + case (V_cons x1a x2a x3) + then show ?case using supp_at_base by auto +next + case (V_consp x1a x2a x3 x4) + then show ?case using supp_at_base by auto +qed + +lemma infer_v_form[simp]: + fixes v::v + assumes "\; \; \ \ v \ \" + shows "\z b. \ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val v)\) \ atom z \ v \ atom z \ (\, \, \)" + using assms +proof(nominal_induct rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ b c x z) + then show ?case by force +next + case (infer_v_litI \ \ \ l \) + then obtain z and b where "\ = \ z : b | CE_val (V_var z) == CE_val (V_lit l) \ \atom z \ (\, \, \) " + using infer_l_form by metis + moreover hence "atom z \ (V_lit l)" using supp_l_empty v.fresh(1) fresh_prod2 fresh_def by blast + ultimately show ?case by metis +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + then show ?case by force +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + moreover hence "atom z \ (V_cons s dc v)" using + Un_commute b.supp(3) fresh_def sup_bot.right_neutral supp_b_empty v.supp(4) pure_supp by metis + ultimately show ?case using fresh_prodN by metis +next + case (infer_v_conspI s bv dclist \ dc tc \ \ v tv b z) + moreover hence "atom z \ (V_consp s dc b v)" unfolding v.fresh using pure_fresh fresh_prodN * by metis + ultimately show ?case using fresh_prodN by metis +qed + +lemma infer_v_form2: + fixes v::v + assumes "\; \; \ \ v \ (\ z : b | c \)" and "atom z \ v" + shows "c = C_eq (CE_val (V_var z)) (CE_val v)" + using assms +proof - + obtain z' and b' where "(\ z : b | c \) = (\ z' : b' | CE_val (V_var z') == CE_val v \) \ atom z' \ v" + using infer_v_form assms by meson + thus ?thesis using Abs1_eq_iff(3) \.eq_iff type_e_eq + by (metis assms(2) ce.fresh(1)) +qed + +lemma infer_v_form3: + fixes v::v + assumes "\; \; \ \ v \ \" and "atom z \ (v,\)" + shows "\; \; \ \ v \ \ z : b_of \ | C_eq (CE_val (V_var z)) (CE_val v)\" +proof - + obtain z' and b' where "\ = \ z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\ \ atom z' \ v \ atom z' \ (\, \, \)" + using infer_v_form assms by metis + moreover hence "\ z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\ = \ z : b' | C_eq (CE_val (V_var z)) (CE_val v)\" + using assms type_e_eq fresh_Pair ce.fresh by auto + ultimately show ?thesis using b_of.simps assms by auto +qed + +lemma infer_v_form4: + fixes v::v + assumes "\; \; \ \ v \ \" and "atom z \ (v,\)" and "b = b_of \" + shows "\; \; \ \ v \ \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\" + using assms infer_v_form3 by simp + +lemma infer_v_v_wf: + fixes v::v + shows "\; \ ; G \ v \ \ \ \; \; G \\<^sub>w\<^sub>f v : (b_of \)" +proof(induct rule: infer_v.induct) + case (infer_v_varI \ \ \ b c x z) + then show ?case using wfC_elims wf_intros by auto +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + then show ?case using wfC_elims wf_intros by auto +next + case (infer_v_litI \ \ \ l \) + hence "b_of \ = base_for_lit l" using infer_l_form3 b_of.simps by metis + then show ?case using wfV_litI infer_l_wf infer_v_litI wfG_b_weakening + by (metis fempty_fsubsetI) +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + then show ?case using wfC_elims wf_intros + by (metis (no_types, lifting) b_of.simps has_fresh_z2 subtype_eq_base2) +next + case (infer_v_conspI s bv dclist \ dc tc \ \ v tv b z) + obtain z1 b1 c1 where t:"tc = \ z1 : b1 | c1 \" using obtain_fresh_z by metis + show ?case unfolding b_of.simps proof(rule wfV_conspI) + show \AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show \(dc, \ z1 : b1 | c1 \ ) \ set dclist\ using infer_v_conspI t by auto + show \ \ ; \ \\<^sub>w\<^sub>f b \ using infer_v_conspI by auto + show \atom bv \ (\, \, \, b, v)\ using infer_v_conspI by auto + have " b1[bv::=b]\<^sub>b\<^sub>b = b_of tv" using subtype_eq_base2[OF infer_v_conspI(5)] b_of.simps t subst_tb.simps by auto + thus \ \; \; \ \\<^sub>w\<^sub>f v : b1[bv::=b]\<^sub>b\<^sub>b \ using infer_v_conspI by auto + qed +qed + +lemma infer_v_t_form_wf: + 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_v_eq assms by auto + +lemma infer_v_t_wf: + fixes v::v + assumes "\; \; G \ v \ \" + shows "wfT \ \ G \ \ wfB \ \ (b_of \) " +proof - + obtain z and b where "\ = \ z : b | CE_val (V_var z) == CE_val v \ \ atom z \ v \ atom z \ (\, \, G)" using infer_v_form assms by metis + moreover have "wfB \ \ b" using infer_v_v_wf b_of.simps wfX_wfB(1) assms + using calculation by fastforce + ultimately show "wfT \ \ G \ \ wfB \ \ (b_of \)" using infer_v_v_wf infer_v_t_form_wf assms by fastforce +qed + +lemma infer_v_wf: + fixes v::v + assumes "\; \; G \ v \ \" + shows "\; \; G \\<^sub>w\<^sub>f v : (b_of \)" and "wfT \ \ G \" and "wfTh \" and "wfG \ \ G" +proof - + show "\; \; G \\<^sub>w\<^sub>f v : b_of \ " using infer_v_v_wf assms by auto + show "\; \; G \\<^sub>w\<^sub>f \" using infer_v_t_wf assms by auto + thus "\ ; \ \\<^sub>w\<^sub>f G" using wfX_wfY by auto + thus "\\<^sub>w\<^sub>f \" using wfX_wfY by auto +qed + +lemma check_bool_options: + assumes "\; \; \ \ v \ \ z : B_bool | TRUE \" and "supp v = {}" + shows "v = V_lit L_true \ v = V_lit L_false" +proof - + obtain t1 where " \; \; \ \ t1 \ \ z : B_bool | TRUE \ \ \; \; \ \ v \ t1" using check_v_elims + using assms by blast + thus ?thesis using infer_bool_options assms + by (metis \.exhaust b_of.simps subtype_eq_base2) +qed + +lemma check_v_wf: + fixes v::v and \::\ and \::\ + assumes "\; \; \ \ v \ \" + shows " \ ; \ \\<^sub>w\<^sub>f \" and "\; \;\ \\<^sub>w\<^sub>f v : b_of \" and "\; \;\ \\<^sub>w\<^sub>f \" +proof - + obtain \' where *: "\; \; \ \ \' \ \ \ \; \; \ \ v \ \'" using check_v_elims assms by auto + thus "\ ; \ \\<^sub>w\<^sub>f \ " and "\; \;\ \\<^sub>w\<^sub>f v : b_of \" and "\; \; \ \\<^sub>w\<^sub>f \" + using infer_v_wf infer_v_v_wf subtype_eq_base2 * subtype_wf by metis+ +qed + +lemma infer_v_form_fresh: + fixes v::v and t::"'a::fs" + assumes "\; \; \ \ v \ \" + shows "\z b. \ = \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\ \ atom z \ (t,v)" +proof - + obtain z' and b' where "\ = \ z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\" using infer_v_form assms by blast + moreover then obtain z and b and c where "\ = \ z : b | c \ \ atom z \ (t,v)" using obtain_fresh_z by metis + ultimately have "\ = \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\ \ atom z \ (t,v) " + using assms infer_v_form2 by auto + thus ?thesis by blast +qed + +text \ More generally, if support of a term is empty then any G will do \ + +lemma infer_v_form_consp: + assumes "\; \; \ \ V_consp s dc b v \ \" + shows "b_of \ = B_app s b" + using assms proof(nominal_induct "V_consp s dc b v" \ rule: infer_v.strong_induct) + case (infer_v_conspI bv dclist \ tc \ \ tv z) + then show ?case using b_of.simps by metis +qed + +lemma lookup_in_rig_b: + assumes "Some (b2, c2) = lookup (\[x\c']) x'" and + "Some (b1, c1) = lookup \ x'" + shows "b1 = b2" + using assms lookup_in_rig[OF assms(2)] + by (metis option.inject prod.inject) + +lemma infer_v_uniqueness_rig: + fixes x::x and c::c + assumes "infer_v P B G v \" and "infer_v P B (replace_in_g G x c') v \'" + shows "\ = \'" + using assms +proof(nominal_induct "v" arbitrary: \' \ rule: v.strong_induct) + case (V_lit l) + hence "infer_l l \" and "infer_l l \'" using assms(1) infer_v_elims(2) by auto + then show ?case using infer_l_uniqueness by presburger +next + case (V_var y) + + obtain b and c where bc: "Some (b,c) = lookup G y" + using assms(1) infer_v_elims(2) using V_var.prems(1) lookup_iff by force + then obtain c'' where bc':"Some (b,c'') = lookup (replace_in_g G x c') y" + using lookup_in_rig by blast + + obtain z where "\ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_var y)) \)" using infer_v_elims(1)[of P B G y \] V_var + bc option.inject prod.inject lookup_in_g by metis + moreover obtain z' where "\' = (\ z' : b | C_eq (CE_val (V_var z')) (CE_val (V_var y)) \)" using infer_v_elims(1)[of P B _ y \'] V_var + option.inject prod.inject lookup_in_rig by (metis bc') + ultimately show ?case using type_e_eq + by (metis V_var.prems(1) V_var.prems(2) \.eq_iff ce.fresh(1) finite.emptyI fresh_atom_at_base + fresh_finite_insert infer_v_elims(1) v.fresh(2)) +next + case (V_pair v1 v2) + obtain z and z1 and z2 and t1 and t2 and c1 and c2 where + t1: "\ = (\ z : [ b_of t1 , b_of t2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v1 v2) \) \ + atom z \ (v1, v2) \ P ; B ; G \ v1 \ t1 \ P ; B ; G \ v2 \ t2" + using infer_v_elims(3)[OF V_pair(3)] by metis + moreover obtain z' and z1' and z2' and t1' and t2' and c1' and c2' where + t2: "\' = (\ z' : [ b_of t1' , b_of t2' ]\<^sup>b | CE_val (V_var z') == CE_val (V_pair v1 v2) \) \ + atom z' \ (v1, v2) \ P ; B ; (replace_in_g G x c') \ v1 \ t1' \ + P ; B ; (replace_in_g G x c') \ v2 \ t2'" + using infer_v_elims(3)[OF V_pair(4)] by metis + ultimately have "t1 = t1' \ t2 = t2'" using V_pair.hyps(1) V_pair.hyps(2) \.eq_iff by blast + then show ?case using t1 t2 by simp +next + case (V_cons s dc v) + obtain x and z and tc and dclist where t1: "\ = (\ z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \) \ + AF_typedef s dclist \ set P \ + (dc, tc) \ set dclist \ atom z \ v" + using infer_v_elims(4)[OF V_cons(2)] by metis + moreover obtain x' and z' and tc' and dclist' where t2: "\' = (\ z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \) + \ AF_typedef s dclist' \ set P \ (dc, tc') \ set dclist' \ atom z' \ v" + using infer_v_elims(4)[OF V_cons(3)] by metis + moreover have a: "AF_typedef s dclist' \ set P" and b:"(dc,tc') \ set dclist'" and c:"AF_typedef s dclist \ set P" and + d:"(dc, tc) \ set dclist" using t1 t2 by auto + ultimately have "tc = tc'" using wfTh_dc_t_unique2 infer_v_wf(3)[OF V_cons(2)] by metis + + moreover have "atom z \ CE_val (V_cons s dc v) \ atom z' \ CE_val (V_cons s dc v)" + using e.fresh(1) v.fresh(4) t1 t2 pure_fresh by auto + ultimately have "(\ z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \) = (\ z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \)" + using type_e_eq by metis + thus ?case using t1 t2 by simp +next + case (V_consp s dc b v) + from V_consp(2) V_consp show ?case proof(nominal_induct "V_consp s dc b v" \ arbitrary: v rule:infer_v.strong_induct) + + case (infer_v_conspI bv dclist \ tc \ \ v tv z) + + obtain z3 and b3 where *:"\' = \ z3 : b3 | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \ \ atom z3 \ V_consp s dc b v" + using infer_v_form[OF \\; \; \[x\c'] \ V_consp s dc b v \ \'\ ] by metis + moreover then have "b3 = B_app s b" using infer_v_form_consp b_of.simps * infer_v_conspI by metis + + moreover have "\ z3 : B_app s b | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \ = \ z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \" + proof - + have "atom z3 \ [V_consp s dc b v]\<^sup>c\<^sup>e" using * ce.fresh by auto + moreover have "atom z \ [V_consp s dc b v]\<^sup>c\<^sup>e" using * infer_v_conspI ce.fresh v.fresh pure_fresh by metis + ultimately show ?thesis using type_e_eq infer_v_conspI v.fresh ce.fresh by metis + qed + ultimately show ?case using * by auto + qed +qed + +lemma infer_v_uniqueness: + assumes "infer_v P B G v \" and "infer_v P B G v \'" + shows "\ = \'" +proof - + obtain x::x where "atom x \ G" using obtain_fresh by metis + hence "G [ x \ C_true ] = G" using replace_in_g_forget assms infer_v_wf by fast + thus ?thesis using infer_v_uniqueness_rig assms by metis +qed + +lemma infer_v_tid_form: + fixes v::v + assumes "\ ; B ; \ \ v \ \ z : B_id tid | c \" and "AF_typedef tid dclist \ set \" and "supp v = {}" + shows "\dc v' t. v = V_cons tid dc v' \ (dc , t ) \ set dclist" + using assms proof(nominal_induct v "\ z : B_id tid | c \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ c x z) + then show ?case using v.supp supp_at_base by auto +next + case (infer_v_litI \ \ l) + then show ?case by auto +next + case (infer_v_consI dclist1 \ dc tc \ \ v tv z) + hence "supp v = {}" using v.supp by simp + then obtain dca and v' where *:"V_cons tid dc v = V_cons tid dca v'" using infer_v_consI by auto + hence "dca = dc" using v.eq_iff(4) by auto + hence "V_cons tid dc v = V_cons tid dca v' \ (dca, tc) \ set dclist1" using infer_v_consI * by auto + moreover have "dclist = dclist1" using wfTh_dclist_unique infer_v_consI wfX_wfY \dca=dc\ + proof - + show ?thesis + by (meson \AF_typedef tid dclist1 \ set \\ \\; \; \ \ v \ tv\ infer_v_consI.prems infer_v_wf(4) wfTh_dclist_unique wfX_wfY) + qed + ultimately show ?case by auto +qed + +lemma check_v_tid_form: + assumes "\ ; B ; \ \ v \ \ z : B_id tid | TRUE \" and "AF_typedef tid dclist \ set \" and "supp v = {}" + shows "\dc v' t. v = V_cons tid dc v' \ (dc , t ) \ set dclist" + using assms proof(nominal_induct v "\ z : B_id tid | TRUE \" rule: check_v.strong_induct) + case (check_v_subtypeI \ \ \ \1 v) + then obtain z and c where "\1 = \ z : B_id tid | c \" using subtype_eq_base2 b_of.simps + by (metis obtain_fresh_z2) + then show ?case using infer_v_tid_form check_v_subtypeI by simp +qed + +lemma check_v_num_leq: + fixes n::int and \::\ + assumes "0 \ n \ n \ int (length v)" and " \\<^sub>w\<^sub>f \ " and "\ ; {||} \\<^sub>w\<^sub>f \" + shows "\ ; {||} ; \ \ [ 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 ) \" +proof - + have "\ ; {||} ; \ \ [ L_num n ]\<^sup>v \ \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \" + using infer_v_litI infer_natI wfG_nilI assms by auto + thus ?thesis using subtype_range[OF assms(1) ] assms check_v_subtypeI by metis +qed + +lemma check_int: + assumes "check_v \ \ \ v (\ z : B_int | c \)" and "supp v = {}" + shows "\n. V_lit (L_num n) = v" + using assms infer_int check_v_elims by (metis b_of.simps infer_v_form subtype_eq_base2) + +definition sble :: "\ \ \ \ bool" where + "sble \ \ = (\i. i \ \ \ \ ; \ \ i)" + +lemma check_v_range: + assumes "\ ; B ; \ \ v2 \ \ 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 [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ " + (is "\ ; ?B ; \ \ v2 \ \ z : B_int | ?c1 \") + and "v1 = V_lit (L_bitvec bv) \ v2 = V_lit (L_num n) " and "atom z \ \" and "sble \ \" + shows "0 \ n \ n \ int (length bv)" +proof - + have "\ ; ?B ; \ \ \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ \ \ z : B_int | ?c1 \" + using check_v_elims assms + by (metis infer_l_uniqueness infer_natI infer_v_elims(2)) + moreover have "atom z \ \" using fresh_GNil assms by simp + ultimately have "\ ; ?B ; ((z, B_int, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ \) \ ?c1" + using subtype_valid_simple by auto + thus ?thesis using assms valid_range_length_inv check_v_wf wfX_wfY sble_def by metis +qed + +section \Expressions\ + +lemma infer_e_plus[elim]: + fixes v1::v and v2::v + assumes "\ ; \ ; \ ; \ ; \ \ AE_op Plus v1 v2 \ \" + shows "\z . (\ z : B_int | C_eq (CE_val (V_var z)) (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \ = \)" + using infer_e_elims assms by metis + +lemma infer_e_leq[elim]: + assumes "\ ; \ ; \ ; \ ; \ \ AE_op LEq v1 v2 \ \" + shows "\z . (\ z : B_bool | C_eq (CE_val (V_var z)) (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \ = \)" + using infer_e_elims assms by metis + +lemma infer_e_eq[elim]: + assumes "\ ; \ ; \ ; \ ; \ \ AE_op Eq v1 v2 \ \" + shows "\z . (\ z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \ = \)" + using infer_e_elims(25)[OF assms] by metis + +lemma infer_e_e_wf: + fixes e::e + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" + shows "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b_of \" + using assms proof(nominal_induct \ avoiding: \ rule: infer_e.strong_induct) + case (infer_e_valI \ \ \ \' \ v \) + then show ?case using infer_v_v_wf wf_intros by metis +next + case (infer_e_plusI \ \ \ \' \ v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_leqI \ \ \ \' v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_eqI \ \ \ \' v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_appI \ \ \ \ \ f x b c \' s' v \'') + have "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f AE_app f v : b_of \' " proof + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using infer_e_appI by auto + show \Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f\ using infer_e_appI by auto + show "\; \;\ \\<^sub>w\<^sub>f v : b" using infer_e_appI check_v_wf b_of.simps by metis + qed + moreover have "b_of \' = b_of (\'[x::=v]\<^sub>v)" using subst_tbase_eq subst_v_\_def by auto + ultimately show ?case using infer_e_appI subst_v_c_def subst_b_\_def by auto +next + case (infer_e_appPI \ \ \ \ \ b' f bv x b c \'' s' v \') + + have "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f AE_appP f b' v : (b_of \'')[bv::=b']\<^sub>b " proof + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appPI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using infer_e_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \'' s'))) = lookup_fun \ f\ using * infer_e_appPI by metis + show "\ ; \ \\<^sub>w\<^sub>f b'" using infer_e_appPI by auto + show "\; \;\ \\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b)" using infer_e_appPI check_v_wf b_of.simps subst_b_b_def by metis + have "atom bv \ (b_of \'')[bv::=b']\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def infer_e_appPI by metis + thus "atom bv \ (\, \, \, \, \, b', v, (b_of \'')[bv::=b']\<^sub>b)" using infer_e_appPI fresh_prodN subst_b_b_def by metis + qed + moreover have "b_of \' = (b_of \'')[bv::=b']\<^sub>b" + using \\''[bv::=b']\<^sub>b[x::=v]\<^sub>v = \'\ b_of_subst_bb_commute subst_tbase_eq subst_b_b_def subst_v_\_def subst_b_\_def by auto + ultimately show ?case using infer_e_appI by auto +next + case (infer_e_fstI \ \ \ \' \ v z' b1 b2 c z) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_sndI \ \ \ \' \ v z' b1 b2 c z) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_lenI \ \ \ \' \ v z' c z) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_mvarI \ \ \ \ u \) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_concatI \ \ \ \' \ v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + have " \ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec" + proof + show "\ \\<^sub>w\<^sub>f \" using infer_e_splitI by auto + show "\; \; \ \\<^sub>w\<^sub>f \" using infer_e_splitI by auto + show "\; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec" using infer_e_splitI b_of.simps infer_v_wf by metis + show "\; \; \ \\<^sub>w\<^sub>f v2 : B_int" using infer_e_splitI b_of.simps check_v_wf by metis + qed + then show ?case using b_of.simps by auto +qed + +lemma infer_e_t_wf: + fixes e::e and \::\ and \::\ and \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" + shows "\; \;\ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \" + using assms proof(induct rule: infer_e.induct) + case (infer_e_valI \ \ \ \' \ v \) + then show ?case using infer_v_t_wf by auto +next + case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence "\; \; \ \\<^sub>w\<^sub>f CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_int" using wfCE_plusI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_plusI by auto +next + case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence " \; \; \ \\<^sub>w\<^sub>f CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_leqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_leqI by auto +next + case (infer_e_eqI \ \ \ \ \ v1 z1 b c1 v2 z2 c2 z3) + hence " \; \; \ \\<^sub>w\<^sub>f CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_eqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_eqI by auto +next + case (infer_e_appI \ \ \ \ \ f x b c \ s' v \') + show ?case proof + show "\ \\<^sub>w\<^sub>f \" using infer_e_appI by auto + show "\; \; \ \\<^sub>w\<^sub>f \'" proof - + have *:" \; \; \ \\<^sub>w\<^sub>f v : b " using infer_e_appI check_v_wf(2) b_of.simps by metis + moreover have *:"\; \; (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f \" proof(rule wf_weakening1(4)) + show \ \; \; (x,b,c)#\<^sub>\GNil \\<^sub>w\<^sub>f \ \ using wfPhi_f_simple_wfT wfD_wf infer_e_appI wb_b_weakening by fastforce + have "\ ; \ ; \ \\<^sub>w\<^sub>f \ x : b | c \" using infer_e_appI check_v_wf(3) by auto + thus \ \ ; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \ \ using infer_e_appI + wfT_wfC[THEN wfG_consI[rotated 3] ] * wfT_wf_cons fresh_prodN by simp + show \toSet ((x,b,c)#\<^sub>\GNil) \ toSet ((x, b, c) #\<^sub>\ \)\ using toSet.simps by auto + qed + moreover have "((x, b, c) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v = \" using subst_gv.simps by auto + + ultimately show ?thesis using infer_e_appI wf_subst1(4)[OF *, of GNil x b c \ v] subst_v_\_def by auto + qed + qed +next + case (infer_e_appPI \ \ \ \ \ b' f bv x b c \' s' v \) + + have "\; \; ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v \\<^sub>w\<^sub>f (\'[bv::=b']\<^sub>b)[x::=v]\<^sub>\\<^sub>v" + proof(rule wf_subst(4)) + show \ \; \; (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \ \\<^sub>w\<^sub>f \'[bv::=b']\<^sub>b \ + proof(rule wf_weakening1(4)) + have \ \ ; {|bv|} ; (x,b,c)#\<^sub>\GNil \\<^sub>w\<^sub>f \' \ using wfPhi_f_poly_wfT infer_e_appI infer_e_appPI by simp + thus \ \; \; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\GNil \\<^sub>w\<^sub>f \'[bv::=b']\<^sub>b \ + using wfT_subst_wfT infer_e_appPI wb_b_weakening subst_b_\_def subst_v_\_def by presburger + have "\; \; \ \\<^sub>w\<^sub>f \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" + using infer_e_appPI check_v_wf(3) subst_b_b_def subst_b_c_def by metis + thus \ \ ; \ \\<^sub>w\<^sub>f (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \ \ + using infer_e_appPI wfT_wfC[THEN wfG_consI[rotated 3] ] * wfX_wfY wfT_wf_cons wb_b_weakening by metis + show \toSet ((x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\GNil) \ toSet ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \)\ using toSet.simps by auto + qed + show \(x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \ = GNil @ (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \\ using append_g.simps by auto + show \ \; \; \ \\<^sub>w\<^sub>f v :b[bv::=b']\<^sub>b\<^sub>b \ using infer_e_appPI check_v_wf(2) b_of.simps subst_b_b_def by metis + qed + moreover have "((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v = \" using subst_gv.simps by auto + ultimately show ?case using infer_e_appPI subst_v_\_def by simp +next + case (infer_e_fstI \ \ \ \ \ v z' b1 b2 c z) + hence "\; \; \ \\<^sub>w\<^sub>f CE_fst [v]\<^sup>c\<^sup>e: b1" using wfCE_fstI wfD_emptyI wfPhi_emptyI infer_v_v_wf + b_of.simps using wfCE_valI by fastforce + then show ?case using wfT_e_eq infer_e_fstI by auto +next + case (infer_e_sndI \ \ \ \ \ v z' b1 b2 c z) + hence "\; \; \ \\<^sub>w\<^sub>f CE_snd [v]\<^sup>c\<^sup>e: b2" using wfCE_sndI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_sndI by auto +next + case (infer_e_lenI \ \ \ \ \ v z' c z) + hence "\; \; \ \\<^sub>w\<^sub>f CE_len [v]\<^sup>c\<^sup>e: B_int" using wfCE_lenI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_lenI by auto +next + case (infer_e_mvarI \ \ \ \ u \) + then show ?case using wfD_wfT by blast +next + case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence "\; \; \ \\<^sub>w\<^sub>f CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e: B_bitvec" using wfCE_concatI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_concatI by auto +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + + hence wfg: " \ ; \ \\<^sub>w\<^sub>f (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \" + using infer_v_wf wfG_cons2I wfB_pairI wfB_bitvecI by simp + have wfz: "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e : [ B_bitvec , B_bitvec ]\<^sup>b " + apply(rule wfCE_valI , rule wfV_varI) + using wfg apply simp + using lookup.simps(2)[of z3 "[ B_bitvec , B_bitvec ]\<^sup>b" TRUE \ z3] by simp + have 1: "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v2 ]\<^sup>c\<^sup>e : B_int " + using check_v_wf[OF infer_e_splitI(4)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce + have 2: "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e : B_bitvec" + using infer_v_wf[OF infer_e_splitI(3)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce + + have "\; \; \ \\<^sub>w\<^sub>f \ z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \" + proof + show "atom z3 \ (\, \, \)" using infer_e_splitI wfTh_x_fresh wfX_wfY fresh_prod3 wfG_fresh_x by metis + show "\ ; \ \\<^sub>w\<^sub>f [ B_bitvec , B_bitvec ]\<^sup>b " using wfB_pairI wfB_bitvecI infer_e_splitI wfX_wfY by metis + show "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ + \ \\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e" + using wfg wfz 1 2 wf_intros by meson + qed + thus ?case using infer_e_splitI by auto +qed + +lemma infer_e_wf: + fixes e::e and \::\ and \::\ and \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" + shows "\; \; \ \\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f \" and "\ \\<^sub>w\<^sub>f \" and "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : (b_of \)" + using infer_e_t_wf infer_e_e_wf wfE_wf assms by metis+ + +lemma infer_e_fresh: + fixes x::x + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" and "atom x \ \" + shows "atom x \ (e,\)" +proof - + have "atom x \ e" using infer_e_e_wf[THEN wfE_x_fresh,OF assms(1)] assms(2) by auto + moreover have "atom x \ \" using assms infer_e_wf wfT_x_fresh by metis + ultimately show ?thesis using fresh_Pair by auto +qed + +inductive check_e :: "\ \ \ \ \ \ \ \ \ \ e \ \ \ bool" (" _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + check_e_subtypeI: "\ infer_e T P B G D e \' ; subtype T B G \' \ \ \ check_e T P B G D e \" +equivariance check_e +nominal_inductive check_e . + +inductive_cases check_e_elims[elim!]: + "check_e F D B G \ (AE_val v) \" + "check_e F D B G \ e \" + +lemma infer_e_fst_pair: + fixes v1::v + assumes "\ ; \ ; {||} ; GNil ; \ \ [#1[ v1 , v2 ]\<^sup>v]\<^sup>e \ \" + shows "\\'. \ ; \ ; {||} ; GNil ; \ \ [v1]\<^sup>e \ \' \ + \ ; {||} ; GNil \ \' \ \" +proof - + obtain z' and b1 and b2 and c and z where ** : "\ = (\ z : b1 | C_eq (CE_val (V_var z)) (CE_fst [(V_pair v1 v2)]\<^sup>c\<^sup>e) \) \ + wfD \ {||} GNil \ \ wfPhi \ \ \ + (\ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \) \ atom z \ V_pair v1 v2 " + using infer_e_elims assms by metis + hence *:" \ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \" by auto + + obtain t1a and t2a where + *: "\ ; {||} ; GNil \ v1 \ t1a \ \ ; {||} ; GNil \ v2 \ t2a \ B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)" + using infer_v_elims(5)[OF *] by metis + + hence suppv: "supp v1 = {} \ supp v2 = {} \ supp (V_pair v1 v2) = {}" using ** infer_v_v_wf wfV_supp atom_dom.simps toSet.simps supp_GNil + by (meson wfV_supp_nil) + + obtain z1 and b1' where "t1a = \ z1 : b1' | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v1 ]\<^sup>c\<^sup>e \ " + using infer_v_form[of \ "{||}" GNil v1 t1a] * by auto + moreover hence "b1' = b1" using * b_of.simps by simp + ultimately have "\ ; {||} ; GNil \ v1 \ \ z1 : b1 | CE_val (V_var z1) == CE_val v1 \" using * by auto + moreover have "\ ; {||} ; GNil \\<^sub>w\<^sub>f CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e : b1" using wfCE_fstI infer_v_wf(1) ** b_of.simps wfCE_valI by metis + moreover hence st: "\ ; {||} ; GNil \ \ z1 : b1 | CE_val (V_var z1) == CE_val v1 \ \ (\ z : b1 | CE_val (V_var z) == CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e \)" + using subtype_gnil_fst infer_v_v_wf by auto + moreover have "wfD \ {||} GNil \ \ wfPhi \ \" using ** by auto + ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis +qed + +lemma infer_e_snd_pair: + assumes "\ ; \ ; {||} ; GNil ; \ \ AE_snd (V_pair v1 v2) \ \" + shows "\\'. \ ; \ ; {||} ; GNil ; \ \ AE_val v2 \ \' \ \ ; {||} ; GNil \ \' \ \" +proof - + obtain z' and b1 and b2 and c and z where ** : "(\ = (\ z : b2 | C_eq (CE_val (V_var z)) (CE_snd [(V_pair v1 v2)]\<^sup>c\<^sup>e) \)) \ + (wfD \ {||} GNil \) \ (wfPhi \ \) \ + \ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \ \ atom z \ V_pair v1 v2 " + using infer_e_elims(9)[OF assms(1)] by metis + hence *:" \ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \" by auto + + obtain t1a and t2a where + *: "\ ; {||} ; GNil \ v1 \ t1a \ \ ; {||} ; GNil \ v2 \ t2a \ B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)" + using infer_v_elims(5)[OF *] by metis + + hence suppv: "supp v1 = {} \ supp v2 = {} \ supp (V_pair v1 v2) = {}" using infer_v_v_wf wfV.simps v.supp by (meson "**" wfV_supp_nil) + + obtain z2 and b2' where "t2a = \ z2 : b2' | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \ " + using infer_v_form[of \ "{||}" GNil v2 t2a] * by auto + moreover hence "b2' = b2" using * b_of.simps by simp + + ultimately have "\ ; {||} ; GNil \ v2 \ \ z2 : b2 | CE_val (V_var z2) == CE_val v2 \" using * by auto + moreover have "\ ; {||} ; GNil \\<^sub>w\<^sub>f CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e : b2" using wfCE_sndI infer_v_wf(1) ** b_of.simps wfCE_valI by metis + moreover hence st: "\ ; {||} ; GNil \ \ z2 : b2 | CE_val (V_var z2) == CE_val v2 \ \ (\ z : b2 | CE_val (V_var z) == CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e \)" + using subtype_gnil_snd infer_v_v_wf by auto + moreover have "wfD \ {||} GNil \ \ wfPhi \ \" using ** by metis + ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis +qed + +section \Statements\ + +lemma check_s_v_unit: + assumes "\; \; \ \ (\ z : B_unit | TRUE \) \ \" and "wfD \ \ \ \" and "wfPhi \ \" + shows "\ ; \ ; \ ; \ ; \ \ AS_val (V_lit L_unit ) \ \" +proof - + have "wfG \ \ \" using assms subtype_g_wf by meson + moreover hence "wfTh \" using wfG_wf by simp + moreover obtain z'::x where "atom z' \ \" using obtain_fresh by auto + ultimately have *:"\; \; \ \ V_lit L_unit \ \ z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \" + using infer_v_litI infer_unitI by simp + moreover have "wfT \ \ \ (\ z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \)" using infer_v_t_wf + by (meson calculation) + moreover then have "\; \; \ \ (\ z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \) \ \" using subtype_trans subtype_top assms + type_for_lit.simps(4) wfX_wfY by metis + ultimately show ?thesis using check_valI assms * by auto +qed + +lemma check_s_check_branch_s_wf: + fixes s::s and cs::branch_s and \::\ and \::\ and \::\ and \::\ and v::v and \::\ and css::branch_list + shows "\ ; \ ; B ; \ ; \ \ s \ \ \ \ ; B \\<^sub>w\<^sub>f \ \ wfTh \ \ wfD \ B \ \ \ wfT \ B \ \ \ wfPhi \ \ " and + "check_branch_s \ \ B \ \ tid cons const v cs \ \ \ ; B \\<^sub>w\<^sub>f \ \ wfTh \ \ wfD \ B \ \ \ wfT \ B \ \ \ wfPhi \ \ " + "check_branch_list \ \ B \ \ tid dclist v css \ \ \ ; B \\<^sub>w\<^sub>f \ \ wfTh \ \ wfD \ B \ \ \ wfT \ B \ \ \ wfPhi \ \ " +proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ B \ \ \ v \' \) + then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI + by (metis subtype_eq_base2) +next + case (check_letI x \ \ \ \ \ e \ z s b c) + then have *:"wfT \ \ ((x, b , c[z::=V_var x]\<^sub>v) #\<^sub>\ \) \" by force + moreover have "atom x \ \" using check_letI fresh_prodN by force + ultimately have "\; \;\ \\<^sub>w\<^sub>f \" using wfT_restrict2 by force + then show ?case using check_letI infer_e_wf wfS_letI wfX_wfY wfG_elims by metis +next + case (check_assertI x \ \ \ \ \ c \ s) + then have *:"wfT \ \ ((x, B_bool , c) #\<^sub>\ \) \" by force + moreover have "atom x \ \" using check_assertI fresh_prodN by force + ultimately have "\; \;\ \\<^sub>w\<^sub>f \" using wfT_restrict2 by force + then show ?case using check_assertI wfS_assertI wfX_wfY wfG_elims by metis +next + case (check_branch_s_branchI \ \ \ \ \ cons const x v \ s tid) + then show ?case using wfX_wfY by metis +next + case (check_branch_list_consI \ \ \ \ \ tid dclist' v cs \ css ) + then show ?case using wfX_wfY by metis +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist' v cs \ ) + then show ?case using wfX_wfY by metis +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + hence *:"wfT \ \ \ (\ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \)" (is "wfT \ \ \ ?tau") by auto + hence "wfT \ \ \ \" using wfT_wfT_if1 check_ifI fresh_prodN by metis + hence " \; \; \ \\<^sub>w\<^sub>f \ " using check_ifI b_of_c_of_eq fresh_prodN by auto + thus ?case using check_ifI by metis +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + then have "wfT \ \ ((x, b_of t, (c_of t x)) #\<^sub>\ G) \" by fastforce + moreover have "atom x \ \" using check_let2I by force + ultimately have "wfT \ \ G \" using wfT_restrict2 by metis + then show ?case using check_let2I by argo +next + case (check_varI u \ P G v \' \ s \) + then show ?case using wfG_elims wfD_elims + list.distinct list.inject by metis +next + case (check_assignI \ \ \ \ \ u \ v z \') + obtain z'::x where *:"atom z' \ \" using obtain_fresh by metis + moreover have "\ z : B_unit | TRUE \ = \ z' : B_unit | TRUE \" by auto + moreover hence "wfT \ \ \ \ z' : B_unit |TRUE \" using wfT_TRUE check_assignI check_v_wf * wfB_unitI wfG_wf by metis + ultimately show ?case using check_v.cases infer_v_wf subtype_wf check_assignI wfT_wf check_v_wf wfG_wf + by (meson subtype_wf) +next + case (check_whileI \ \ G P s1 z s2 \') + then show ?case using subtype_wf subtype_wf by auto +next + case (check_seqI \ G P s1 z s2 \) + then show ?case by fast +next + case (check_caseI \ \ \ \ \ dclist cs \ tid v z) + then show ?case by fast +qed + +lemma check_s_check_branch_s_wfS: + fixes s::s and cs::branch_s and \::\ and \::\ and \::\ and \::\ and v::v and \::\ and css::branch_list + shows "\ ; \ ; B ; \ ; \ \ s \ \ \ \ ; \ ; B ; \ ; \ \\<^sub>w\<^sub>f s : b_of \ " and + "check_branch_s \ \ B \ \ tid cons const v cs \ \ wfCS \ \ B \ \ tid cons const cs (b_of \)" + "check_branch_list \ \ B \ \ tid dclist v css \ \ wfCSS \ \ B \ \ tid dclist css (b_of \)" +proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ \ \ \ \ v \' \) + then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI + by (metis subtype_eq_base2) +next + case (check_letI x \ \ \ \ \ e \ z s b c) + show ?case proof + show \ \ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b \ using infer_e_wf check_letI b_of.simps by metis + show \ \ ; \ ; \ ; (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b_of \ \ + using check_letI b_of.simps wf_replace_true2(2)[OF check_letI(5)] append_g.simps by metis + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using infer_e_wf check_letI b_of.simps by metis + show \atom x \ (\, \, \, \, \, e, b_of \)\ + apply(simp add: fresh_prodN, intro conjI) + using check_letI(1) fresh_prod7 by simp+ + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + + show \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b_of \ \ using check_assertI by auto + next + show \ \; \; \ \\<^sub>w\<^sub>f c \ using check_assertI by auto + next + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using check_assertI by auto + next + show \atom x \ (\, \, \, \, \, c, b_of \, s)\ using check_assertI by auto + qed +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + show ?case proof + show \ \ ; \ ; \ ; (x, b_of const, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b_of \ \ + using wf_replace_true append_g.simps check_branch_s_branchI by metis + show \atom x \ (\, \, \, \, \, \, const)\ + using wf_replace_true append_g.simps check_branch_s_branchI fresh_prodN by metis + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wf_replace_true append_g.simps check_branch_s_branchI by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid cons const v cs \ dclist css) + then show ?case using wf_intros by metis +next + case (check_branch_list_finalI \ \ \ \ \ tid cons const v cs \) + then show ?case using wf_intros by metis +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case using wfS_ifI check_v_wf check_ifI b_of.simps by metis +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + show ?case proof + show \ \ ; \ ; \ ; G ; \ \\<^sub>w\<^sub>f s1 : b_of t \ using check_let2I b_of.simps by metis + show \ \; \; G \\<^sub>w\<^sub>f t \ using check_let2I check_s_check_branch_s_wf by metis + show \ \ ; \ ; \ ; (x, b_of t, TRUE) #\<^sub>\ G ; \ \\<^sub>w\<^sub>f s2 : b_of \ \ + using check_let2I(5) wf_replace_true2(2) append_g.simps check_let2I by metis + show \atom x \ (\, \, \, G, \, s1, b_of \, t)\ + apply(simp add: fresh_prodN, intro conjI) + using check_let2I(1) fresh_prod7 by simp+ + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using check_v_wf check_varI by metis + show \ \; \; \ \\<^sub>w\<^sub>f v : b_of \' \ using check_v_wf check_varI by metis + show \atom u \ (\, \, \, \, \, \', v, b_of \)\ using check_varI fresh_prodN u_fresh_b by metis + show \ \ ; \ ; \ ; \ ; (u, \') #\<^sub>\\ \\<^sub>w\<^sub>f s : b_of \ \ using check_varI by metis + qed +next + case (check_assignI \ \ \ \ \ u \ v z \') + then show ?case using wf_intros check_v_wf subtype_eq_base2 b_of.simps by metis +next + case (check_whileI \ \ \ \ \ s1 z s2 \') + thus ?case using wf_intros b_of.simps check_v_wf subtype_eq_base2 b_of.simps by metis +next + case (check_seqI \ \ \ \ \ s1 z s2 \) + thus ?case using wf_intros b_of.simps by metis +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + show ?case proof + show \ \; \; \ \\<^sub>w\<^sub>f v : B_id tid \ using check_caseI check_v_wf b_of.simps by metis + show \AF_typedef tid dclist \ set \\ using check_caseI by metis + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using check_caseI check_s_check_branch_s_wf by metis + show \ \ \\<^sub>w\<^sub>f \ \ using check_caseI check_s_check_branch_s_wf by metis + show \ \ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f cs : b_of \ \ using check_caseI by metis + qed +qed + +lemma check_s_wf: + fixes s::s + assumes "\ ; \ ; B ; \ ; \ \ s \ \" + shows "\ ; B \\<^sub>w\<^sub>f \ \ wfT \ B \ \ \ wfPhi \ \ \ wfTh \ \ wfD \ B \ \ \ wfS \ \ B \ \ s (b_of \)" + using check_s_check_branch_s_wf check_s_check_branch_s_wfS assms by meson + +lemma check_s_flip_u1: + fixes s::s and u::u and u'::u + assumes "\ ; \ ; \ ; \ ; \ \ s \ \" + shows "\ ; \ ; \ ; \ ; ( u \ u') \ \ \ (u \ u') \ s \ \" +proof - + have "(u \ u') \ \ ; (u \ u') \ \ ; (u \ u') \ \ ; (u \ u') \ \ ; (u \ u') \ \ \ (u \ u') \ s \ (u \ u') \ \" + using check_s.eqvt assms by blast + thus ?thesis using check_s_wf[OF assms] flip_u_eq phi_flip_eq by metis +qed + +lemma check_s_flip_u2: + fixes s::s and u::u and u'::u + assumes "\ ; \ ; B ; \ ; ( u \ u') \ \ \ (u \ u') \ s \ \" + shows "\ ; \ ; B ; \ ; \ \ s \ \" +proof - + have "\ ; \ ; B ; \ ; ( u \ u') \ ( u \ u') \ \ \ ( u \ u') \ (u \ u') \ s \ \" + using check_s_flip_u1 assms by blast + thus ?thesis using permute_flip_cancel by force +qed + +lemma check_s_flip_u: + fixes s::s and u::u and u'::u + shows "\ ; \ ; B ; \ ; ( u \ u') \ \ \ (u \ u') \ s \ \ = (\ ; \ ; B ; \ ; \ \ s \ \)" + using check_s_flip_u1 check_s_flip_u2 by metis + +lemma check_s_abs_u: + fixes s::s and s'::s and u::u and u'::u and \'::\ + assumes "[[atom u]]lst. s = [[atom u']]lst. s'" and "atom u \ \" and "atom u' \ \" + and "\ ; B ; \ \\<^sub>w\<^sub>f \'" + and "\ ; \ ; B ; \ ; ( u , \') #\<^sub>\\ \ s \ \" + shows "\ ; \ ; B ; \ ; ( u', \' ) #\<^sub>\\ \ s' \ \" +proof - + have "\ ; \ ; B ; \ ; ( u' \ u) \ (( u , \') #\<^sub>\\) \ ( u' \ u) \ s \ \" + using assms check_s_flip_u by metis + moreover have " ( u' \ u) \ (( u , \') #\<^sub>\ \) = ( u' , \') #\<^sub>\\" proof - + have " ( u' \ u) \ (( u , \') #\<^sub>\ \) = ((u' \ u) \ u, (u' \ u) \ \') #\<^sub>\ (u' \ u) \ \" + using DCons_eqvt Pair_eqvt by auto + also have "... = ( u' , (u' \ u) \ \') #\<^sub>\ \" + using assms flip_fresh_fresh by auto + also have "... = ( u' , \') #\<^sub>\ \" using + u_not_in_t fresh_def flip_fresh_fresh assms by metis + finally show ?thesis by auto + qed + moreover have "( u' \ u) \ s = s'" using assms Abs1_eq_iff(3)[of u' s' u s] by auto + ultimately show ?thesis by auto +qed + +section \Additional Elimination and Intros\ + +subsection \Values\ + +nominal_function b_for :: "opp \ b" where + "b_for Plus = B_int" +| "b_for LEq = B_bool" | "b_for Eq = B_bool" + apply(auto,simp add: eqvt_def b_for_graph_aux_def ) + by (meson opp.exhaust) +nominal_termination (eqvt) by lexicographic_order + + +lemma infer_v_pair2I: + fixes v\<^sub>1::v and v\<^sub>2::v + assumes "\; \; \ \ v\<^sub>1 \ \\<^sub>1" and "\; \; \ \ v\<^sub>2 \ \\<^sub>2" + shows "\\. \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ \ b_of \ = B_pair (b_of \\<^sub>1) (b_of \\<^sub>2)" +proof - + obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\\<^sub>1 = (\ z1 : b1 | c1 \) \ \\<^sub>2 = (\ z2 : b2 | c2 \)" + using \.exhaust by meson + obtain z::x where "atom z \ ( v\<^sub>1, v\<^sub>2,\, \,\)" using obtain_fresh + by blast + hence "atom z \ ( v\<^sub>1, v\<^sub>2) \ atom z \ (\, \,\)" using fresh_prodN by metis + hence " \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ z : [ b_of \\<^sub>1 , b_of \\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \" + using assms infer_v_pairI zbc by metis + moreover obtain \ where "\ = (\ z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \)" by blast + moreover hence "b_of \ = B_pair (b_of \\<^sub>1) (b_of \\<^sub>2)" using b_of.simps zbc by presburger + ultimately show ?thesis using b_of.simps by metis +qed + +lemma infer_v_pair2I_zbc: + fixes v\<^sub>1::v and v\<^sub>2::v + assumes "\; \; \ \ v\<^sub>1 \ \\<^sub>1" and "\; \; \ \ v\<^sub>2 \ \\<^sub>2" + shows "\z \. \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ \ \ = (\ z : B_pair (b_of \\<^sub>1) (b_of \\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \) \ atom z \ (v\<^sub>1,v\<^sub>2) \ atom z \ \" +proof - + obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\\<^sub>1 = (\ z1 : b1 | c1 \) \ \\<^sub>2 = (\ z2 : b2 | c2 \)" + using \.exhaust by meson + obtain z::x where * : "atom z \ ( v\<^sub>1, v\<^sub>2,\,\ , \ )" using obtain_fresh + by blast + hence vinf: " \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ z :[ b_of \\<^sub>1 , b_of \\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \" + using assms infer_v_pairI by simp + moreover obtain \ where "\ = (\ z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \)" by blast + moreover have "b_of \\<^sub>1 = b1 \ b_of \\<^sub>2 = b2" using zbc b_of.simps by auto + ultimately have "\; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ \ \ = (\ z : B_pair (b_of \\<^sub>1) (b_of \\<^sub>2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \)" by auto + thus ?thesis using * fresh_prod2 fresh_prod3 by metis +qed + +lemma infer_v_pair2E: + assumes "\; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \" + shows "\\\<^sub>1 \\<^sub>2 z . \; \; \ \ v\<^sub>1 \ \\<^sub>1 \ \; \; \ \ v\<^sub>2 \ \\<^sub>2 \ + \ = (\ z : B_pair (b_of \\<^sub>1) (b_of \\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \) \ atom z \ (v\<^sub>1, v\<^sub>2)" +proof - + obtain z and t1 and t2 where + "\ = (\ z : B_pair (b_of t1) (b_of t2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \) \ + atom z \ (v\<^sub>1, v\<^sub>2) \ \; \; \ \ v\<^sub>1 \ t1 \ \; \; \ \ v\<^sub>2 \ t2 " using infer_v_elims(3)[OF assms ] by metis + thus ?thesis using b_of.simps by metis +qed + +subsection \Expressions\ + +lemma infer_e_app2E: + fixes \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ AE_app f v \ \" + shows "\x b c s' \'. wfD \ \ \ \ \ Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f \ \ \\<^sub>w\<^sub>f \ \ + \; \; \ \ v \ \ x : b | c \ \ \ = \'[x::=v]\<^sub>\\<^sub>v \ atom x \ (\, \, \, \, \, v, \)" + using infer_e_elims(6)[OF assms] b_of.simps subst_defs by metis + +lemma infer_e_appP2E: + fixes \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ AE_appP f b v \ \" + shows "\bv x ba c s' \'. wfD \ \ \ \ \ Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \' s'))) = lookup_fun \ f \ \ \\<^sub>w\<^sub>f \ \ \ ; \ \\<^sub>w\<^sub>f b \ + (\; \; \ \ v \ \ x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \) \ (\ = \'[bv::=b]\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v) \ atom x \ \ \ atom bv \ v" +proof - + obtain bv x ba c s' \' where *:"wfD \ \ \ \ \ Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \' s'))) = lookup_fun \ f \ \ \\<^sub>w\<^sub>f \ \ \ ; \ \\<^sub>w\<^sub>f b \ + (\; \; \ \ v \ \ x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \) \ (\ = \'[bv::=b]\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v) \ atom x \ \ \ atom bv \ (\, \, \, \, \, b, v, \)" + using infer_e_elims(21)[OF assms] subst_defs by metis + moreover then have "atom bv \ v" using fresh_prodN by metis + ultimately show ?thesis by metis +qed + +section \Weakening\ + +text \ Lemmas showing that typing judgements hold when a context is extended \ + +lemma subtype_weakening: + fixes \'::\ + assumes "\; \; \ \ \1 \ \2" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\; \; \' \ \1 \ \2" + using assms proof(nominal_induct \2 avoiding: \' rule: subtype.strong_induct) + + case (subtype_baseI x \ \ \ z c z' c' b) + show ?case proof + show *:"\; \; \' \\<^sub>w\<^sub>f \ z : b | c \" using wfT_weakening subtype_baseI by metis + show "\; \; \' \\<^sub>w\<^sub>f \ z' : b | c' \" using wfT_weakening subtype_baseI by metis + show "atom x \ (\, \, \', z , c , z' , c' )" using subtype_baseI fresh_Pair by metis + have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \) \ toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \')" using subtype_baseI toSet.simps by blast + moreover have "\ ; \ \\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \'" using wfT_wf_cons3[OF *, of x] subtype_baseI fresh_Pair subst_defs by metis + ultimately show "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \' \ c'[z'::=V_var x]\<^sub>v" using valid_weakening subtype_baseI by metis + qed +qed + +method many_rules uses add = ( (rule+),((simp add: add)+)?) + +lemma infer_v_g_weakening: + fixes e::e and \'::\ and v::v + assumes "\; \ ; \ \ v \ \" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\; \ ; \' \ v \ \" + using assms proof(nominal_induct avoiding: \' rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ b c x' z) + show ?case proof + show \ \ ; \ \\<^sub>w\<^sub>f \' \ using infer_v_varI by auto + show \Some (b, c) = lookup \' x'\ using infer_v_varI lookup_weakening by metis + show \atom z \ x'\ using infer_v_varI by auto + show \atom z \ (\, \, \')\ using infer_v_varI by auto + qed +next + case (infer_v_litI \ \ \ l \) + then show ?case using infer_v.intros by simp +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + then show ?case using infer_v.intros by simp +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + show ?case proof + show \AF_typedef s dclist \ set \\ using infer_v_consI by auto + show \(dc, tc) \ set dclist\ using infer_v_consI by auto + show \ \; \; \' \ v \ tv\ using infer_v_consI by auto + show \\; \; \' \ tv \ tc\ using infer_v_consI subtype_weakening by auto + show \atom z \ v\ using infer_v_consI by auto + show \atom z \ (\, \, \')\ using infer_v_consI by auto + qed +next + case (infer_v_conspI s bv dclist \ dc tc \ \ v tv b z) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show \(dc, tc) \ set dclist\ using infer_v_conspI by auto + show \ \; \; \' \ v \ tv\ using infer_v_conspI by auto + show \\; \; \' \ tv \ tc[bv::=b]\<^sub>\\<^sub>b\ using infer_v_conspI subtype_weakening by auto + show \atom z \ (\, \, \', v, b)\ using infer_v_conspI by auto + show \atom bv \ (\, \, \', v, b)\ using infer_v_conspI by auto + show \ \ ; \ \\<^sub>w\<^sub>f b \ using infer_v_conspI by auto + qed +qed + +lemma check_v_g_weakening: + fixes e::e and \'::\ + assumes "\; \ ; \ \ v \ \" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\; \ ; \' \ v \ \" + using subtype_weakening infer_v_g_weakening check_v_elims check_v_subtypeI assms by metis + +lemma infer_e_g_weakening: + fixes e::e and \'::\ + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\ ; \ ; \ ; \'; \ \ e \ \" + using assms proof(nominal_induct \ avoiding: \' rule: infer_e.strong_induct) + case (infer_e_valI \ \ \ \' \ v \) + then show ?case using infer_v_g_weakening wf_weakening infer_e.intros by metis +next + case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + + obtain z'::x where z': "atom z' \ v1 \ atom z' \ v2 \ atom z' \ \'" using obtain_fresh fresh_prod3 by metis + moreover hence *:"\ z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = (\ z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \)" + using infer_e_plusI type_e_eq ce.fresh fresh_e_opp by auto + + have "\ ; \ ; \ ; \' ; \ \ AE_op Plus v1 v2 \ \ z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_plusI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_plusI by auto + show \ \ ; \ ; \' \ v1 \ \ z1 : B_int | c1 \\ using infer_v_g_weakening infer_e_plusI by auto + show \ \ ; \ ; \' \ v2 \ \ z2 : B_int | c2 \\ using infer_v_g_weakening infer_e_plusI by auto + show \atom z' \ AE_op Plus v1 v2\ using z' by auto + show \atom z' \ \'\ using z' by auto + qed + thus ?case using * by metis + +next + case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + obtain z'::x where z': "atom z' \ v1 \ atom z' \ v2 \ atom z' \ \'" using obtain_fresh fresh_prod3 by metis + + moreover hence *:"\ z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = (\ z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \)" + using infer_e_leqI type_e_eq ce.fresh fresh_e_opp by auto + + have "\ ; \ ; \ ; \' ; \ \ AE_op LEq v1 v2 \ \ z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_leqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_leqI by auto + show \ \ ; \ ; \' \ v1 \ \ z1 : B_int | c1 \\ using infer_v_g_weakening infer_e_leqI by auto + show \ \ ; \ ; \' \ v2 \ \ z2 : B_int | c2 \\ using infer_v_g_weakening infer_e_leqI by auto + show \atom z' \ AE_op LEq v1 v2\ using z' by auto + show \atom z' \ \'\ using z' by auto + qed + thus ?case using * by metis +next + case (infer_e_eqI \ \ \ \ \ v1 z1 bb c1 v2 z2 c2 z3) + obtain z'::x where z': "atom z' \ v1 \ atom z' \ v2 \ atom z' \ \'" using obtain_fresh fresh_prod3 by metis + + moreover hence *:"\ z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = (\ z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \)" + using infer_e_eqI type_e_eq ce.fresh fresh_e_opp by auto + + have "\ ; \ ; \ ; \' ; \ \ AE_op Eq v1 v2 \ \ z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_eqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_eqI by auto + show \ \ ; \ ; \' \ v1 \ \ z1 : bb | c1 \\ using infer_v_g_weakening infer_e_eqI by auto + show \ \ ; \ ; \' \ v2 \ \ z2 : bb | c2 \\ using infer_v_g_weakening infer_e_eqI by auto + show \atom z' \ AE_op Eq v1 v2\ using z' by auto + show \atom z' \ \'\ using z' by auto + show "bb \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed + thus ?case using * by metis +next + case (infer_e_appI \ \ \ \ \ f x b c \' s' v \) + show ?case proof + show "\; \; \' \\<^sub>w\<^sub>f \ " using wf_weakening infer_e_appI by auto + show "\ \\<^sub>w\<^sub>f \" using wf_weakening infer_e_appI by auto + show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f" using wf_weakening infer_e_appI by auto + show "\; \; \' \ v \ \ x : b | c \" using wf_weakening infer_e_appI check_v_g_weakening by auto + show "atom x \ (\, \, \, \', \, v, \)" using wf_weakening infer_e_appI by auto + show "\'[x::=v]\<^sub>v = \" using wf_weakening infer_e_appI by auto + qed + +next + case (infer_e_appPI \ \ \ \ \ b' f bv x b c \' s' v \) + + hence *:"\ ; \ ; \ ; \ ; \ \ AE_appP f b' v \ \" using Typing.infer_e_appPI by auto + + obtain x'::x where x':"atom x' \ (s', c, \', (\',v,\)) \ (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \' s'))) = (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \ x) \ c) ((x' \ x) \ \') ((x' \ x) \ s'))))" + using obtain_fresh_fun_def[of s' c \' "(\',v,\)" f x b] + by (metis fun_def.eq_iff fun_typ_q.eq_iff(2)) + + hence **: " \ x : b | c \ = \ x' : b | (x' \ x) \ c \" + using fresh_PairD(1) fresh_PairD(2) type_eq_flip by blast + have "atom x' \ \" using x' infer_e_appPI fresh_weakening fresh_Pair by metis + + show ?case proof(rule Typing.infer_e_appPI[where x=x' and bv=bv and b=b and c="(x' \ x) \ c" and \'="(x' \ x) \ \'"and s'="((x' \ x) \ s')"]) + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_appPI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_appPI by auto + show "\ ; \ \\<^sub>w\<^sub>f b'" using infer_e_appPI by auto + show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \ x) \ c) ((x' \ x) \ \') ((x' \ x) \ s')))) = lookup_fun \ f" using x' infer_e_appPI by argo + show "\; \; \' \ v \ \ x' : b[bv::=b']\<^sub>b | ((x' \ x) \ c)[bv::=b']\<^sub>b \" using ** + \.eq_iff check_v_g_weakening infer_e_appPI.hyps infer_e_appPI.prems(1) infer_e_appPI.prems subst_defs + subst_tb.simps by metis + show "atom x' \ \'" using x' fresh_prodN by metis + + have "atom x \ (v, \) \ atom x' \ (v, \)" using x' infer_e_fresh[OF *] e.fresh(2) fresh_Pair infer_e_appPI \atom x' \ \\ e.fresh by metis + moreover then have "((x' \ x) \ \')[bv::=b']\<^sub>\\<^sub>b = (x' \ x) \ (\'[bv::=b']\<^sub>\\<^sub>b)" using subst_b_x_flip + by (metis subst_b_\_def) + ultimately show "((x' \ x) \ \')[bv::=b']\<^sub>b[x'::=v]\<^sub>v = \" + using infer_e_appPI subst_tv_flip subst_defs by metis + + show "atom bv \ (\, \, \, \', \, b', v, \)" using infer_e_appPI fresh_prodN by metis + qed + +next + case (infer_e_fstI \ \ \ \ \ v z'' b1 b2 c z) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v \ atom z' \ c" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z : b1 | CE_val (V_var z) == CE_fst [v]\<^sup>c\<^sup>e \ = \ z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_fstI v.fresh e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_fst v \ \ z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_fstI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_fstI by auto + show "\ ; \ ; \' \ v \ \ z'' : B_pair b1 b2 | c \" using infer_v_g_weakening infer_e_fstI by metis + show "atom z' \ AE_fst v" using * ** e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using * ** by metis +next + case (infer_e_sndI \ \ \ \ \ v z'' b1 b2 c z) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v \ atom z' \ c" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z : b2 | CE_val (V_var z) == CE_snd [v]\<^sup>c\<^sup>e \ = \ z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_sndI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_snd v \ \ z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ;\' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_sndI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_sndI by auto + show "\ ; \ ; \' \ v \ \ z'' : B_pair b1 b2 | c \" using infer_v_g_weakening infer_e_sndI by metis + show "atom z' \ AE_snd v" using * e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using ** by metis +next + case (infer_e_lenI \ \ \ \ \ v z'' c z) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v \ atom z' \ c" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z : B_int | CE_val (V_var z) == CE_len [v]\<^sup>c\<^sup>e \ = \ z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_lenI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_len v \ \ z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \" proof + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_lenI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_lenI by auto + show "\ ; \ ; \' \ v \ \ z'' : B_bitvec | c \" using infer_v_g_weakening infer_e_lenI by metis + show "atom z' \ AE_len v" using * e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using * ** by metis +next + case (infer_e_mvarI \ \ \ \ u \) + then show ?case using wf_weakening infer_e.intros by metis +next + case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v1 \ atom z' \ v2" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = \ z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_concatI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_concat v1 v2 \ \ z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_concatI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_concatI by auto + show "\ ; \ ; \' \ v1 \ \ z1 : B_bitvec | c1 \" using infer_v_g_weakening infer_e_concatI by metis + show "\ ; \ ; \' \ v2 \ \ z2 : B_bitvec | c2 \" using infer_v_g_weakening infer_e_concatI by metis + show "atom z' \ AE_concat v1 v2" using * e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using * ** by metis +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + show ?case proof + show "\; \; \' \\<^sub>w\<^sub>f \" using infer_e_splitI wf_weakening by auto + show "\ \\<^sub>w\<^sub>f \ " using infer_e_splitI wf_weakening by auto + show "\; \; \' \ v1 \ \ z1 : B_bitvec | c1 \" using infer_v_g_weakening infer_e_splitI by metis + show "\; \; \' \ v2 \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e + AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \" + using check_v_g_weakening infer_e_splitI by metis + show "atom z1 \ AE_split v1 v2" using infer_e_splitI by auto + show "atom z1 \ \'" using infer_e_splitI by auto + show "atom z2 \ AE_split v1 v2" using infer_e_splitI by auto + show "atom z2 \ \'" using infer_e_splitI by auto + show "atom z3 \ AE_split v1 v2" using infer_e_splitI by auto + show "atom z3 \ \'" using infer_e_splitI by auto + qed +qed + +text \ Special cases proved explicitly, other cases at the end with method + \ + +lemma infer_e_d_weakening: + fixes e::e + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" and "setD \ \ setD \'" and "wfD \ \ \ \'" + shows "\; \ ; \ ; \ ; \' \ e \ \" + using assms by(nominal_induct \ avoiding: \' rule: infer_e.strong_induct,auto simp add:infer_e.intros) + +lemma wfG_x_fresh_in_v_simple: + fixes x::x and v :: v + assumes "\; \; \ \ v \ \" and "atom x \ \" + shows "atom x \ v" + using wfV_x_fresh infer_v_wf assms by metis + +lemma check_s_g_weakening: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \'::\ and \::\ and css::branch_list + shows "check_s \ \ \ \ \ s t \ toSet \ \ toSet \' \ \ ; \ \\<^sub>w\<^sub>f \' \ check_s \ \ \ \' \ s t" and + "check_branch_s \ \ \ \ \ tid cons const v cs t \ toSet \ \ toSet \' \ \ ; \ \\<^sub>w\<^sub>f \' \ check_branch_s \ \ \ \' \ tid cons const v cs t" and + "check_branch_list \ \ \ \ \ tid dclist v css t \ toSet \ \ toSet \' \ \ ; \ \\<^sub>w\<^sub>f \' \ check_branch_list \ \ \ \' \ tid dclist v css t" +proof(nominal_induct t and t and t avoiding: \' rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_valI \ \ \ \' \ v \' \) + then show ?case using Typing.check_valI infer_v_g_weakening wf_weakening subtype_weakening by metis +next + case (check_letI x \ \ \ \ \ e \ z s b c) + hence xf:"atom x \ \'" by metis + show ?case proof + show "atom x \ (\, \, \, \', \, e, \)" using check_letI using fresh_prod4 xf by metis + show "\ ; \ ; \ ; \' ; \ \ e \ \ z : b | c \" using infer_e_g_weakening check_letI by metis + show "atom z \ (x, \, \, \, \', \, e, \, s)" + by(unfold fresh_prodN,auto simp add: check_letI fresh_prodN) + have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \) \ toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \')" using check_letI toSet.simps + by (metis Un_commute Un_empty_right Un_insert_right insert_mono) + moreover hence "\ ; \ \\<^sub>w\<^sub>f ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \')" using check_letI wfG_cons_weakening check_s_wf by metis + ultimately show "\ ; \ ; \ ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \' ; \ \ s \ \" using check_letI by metis + qed +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + show ?case proof + show "atom x \ (\, \, \, \', \, t, s1, \)" using check_let2I using fresh_prod4 by auto + show "\ ; \ ; \ ; \' ; \ \ s1 \ t" using check_let2I by metis + have "toSet ((x, b_of t, c_of t x) #\<^sub>\ G) \ toSet ((x, b_of t, c_of t x ) #\<^sub>\ \')" using check_let2I by auto + moreover hence "\ ; \ \\<^sub>w\<^sub>f ((x, b_of t, c_of t x) #\<^sub>\ \')" using check_let2I wfG_cons_weakening check_s_wf by metis + ultimately show "\ ; \ ; \ ; (x, b_of t, c_of t x) #\<^sub>\ \' ; \ \ s2 \ \" using check_let2I by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist' v cs \ css dclist) + thus ?case using Typing.check_branch_list_consI by metis +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist' v cs \ dclist) + thus ?case using Typing.check_branch_list_finalI by metis +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + show ?case proof + show "\; \; \' \\<^sub>w\<^sub>f \ " using wf_weakening2(6) check_branch_s_branchI by metis + show "\\<^sub>w\<^sub>f \" using check_branch_s_branchI by auto + show "\; \; \' \\<^sub>w\<^sub>f \ " using check_branch_s_branchI wfT_weakening \wfG \ \ \'\ by presburger + + show "\ ; {||} ; GNil \\<^sub>w\<^sub>f const " using check_branch_s_branchI by auto + show "atom x \ (\, \, \, \', \, tid, cons, const, v, \)" using check_branch_s_branchI by auto + have "toSet ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) \ toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \')" + using check_branch_s_branchI by auto + moreover hence "\ ; \ \\<^sub>w\<^sub>f ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \')" + using check_branch_s_branchI wfG_cons_weakening check_s_wf by metis + ultimately show "\ ; \ ; \ ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \' ; \ \ s \ \" + using check_branch_s_branchI using fresh_dom_free by auto + qed +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case proof + show \atom z \ (\, \, \, \', \, v, s1, s2, \)\ using fresh_prodN check_ifI by auto + show \\; \; \' \ v \ \ z : B_bool | TRUE \\ using check_v_g_weakening check_ifI by auto + show \ \ ; \ ; \ ; \' ; \ \ s1 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_true) IMP c_of \ z \\ using check_ifI by auto + show \ \ ; \ ; \ ; \' ; \ \ s2 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \\ using check_ifI by auto + qed +next + case (check_whileI \ G P s1 z s2 \') + then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening + by (meson infer_v_g_weakening) +next + case (check_seqI \ G P s1 z s2 \) + then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening + by (meson infer_v_g_weakening) +next + case (check_varI u \ \ \ \ \ \' v \ s) + thus ?case using check_v_g_weakening check_s_check_branch_s_check_branch_list.intros by auto +next + case (check_assignI \ \ \ \ \ u \ v z \') + show ?case proof + show \\ \\<^sub>w\<^sub>f \ \ using check_assignI by auto + show \\; \; \' \\<^sub>w\<^sub>f \\ using check_assignI wf_weakening by auto + show \(u, \) \ setD \\ using check_assignI by auto + show \\; \; \' \ v \ \\ using check_assignI check_v_g_weakening by auto + show \\; \; \' \ \ z : B_unit | TRUE \ \ \'\ using subtype_weakening check_assignI by auto + qed +next + case (check_caseI \ \ \ dclist cs \ tid v z) + + then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening + by (meson infer_v_g_weakening) +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + show \atom x \ (\, \, \, \', \, c, \, s)\ using check_assertI by auto + + have " \ ; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \" using check_assertI check_s_wf by metis + hence *:" \ ; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \'" using wfG_cons_weakening check_assertI by metis + moreover have "toSet ((x, B_bool, c) #\<^sub>\ \) \ toSet ((x, B_bool, c) #\<^sub>\ \')" using check_assertI by auto + thus \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \' ; \ \ s \ \\ using check_assertI(11) [OF _ *] by auto + + show \\; \; \' \ c \ using check_assertI valid_weakening by metis + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using check_assertI wf_weakening by metis + qed +qed + +lemma wfG_xa_fresh_in_v: + fixes c::c and \::\ and G::\ and v::v and xa::x + assumes "\; \; \ \ v \ \" and "G=( \'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and "\ ; \ \\<^sub>w\<^sub>f G" + shows "atom xa \ v" +proof - + have "\; \; \ \\<^sub>w\<^sub>f v : b_of \" using infer_v_wf assms by metis + hence "supp v \ atom_dom \ \ supp \" using wfV_supp by simp + moreover have "atom xa \ atom_dom G" + using assms wfG_atoms_supp_eq[OF assms(4)] fresh_def by metis + ultimately show ?thesis using fresh_def + using assms infer_v_wf wfG_atoms_supp_eq + fresh_GCons fresh_append_g subsetCE + by (metis wfG_x_fresh_in_v_simple) +qed + +lemma fresh_z_subst_g: + fixes G::\ + assumes "atom z' \ (x,v)" and \atom z' \ G\ + shows "atom z' \ G[x::=v]\<^sub>\\<^sub>v" +proof - + have "atom z' \ v" using assms fresh_prod2 by auto + thus ?thesis using fresh_subst_gv assms by metis +qed + +lemma wfG_xa_fresh_in_subst_v: + fixes c::c and v::v and x::x and \::\ and G::\ and xa::x + assumes "\; \; \ \ v \ \" and "G=( \'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and "\ ; \ \\<^sub>w\<^sub>f G" + shows "atom xa \ (subst_gv G x v)" +proof - + have "atom xa \ v" using wfG_xa_fresh_in_v assms by metis + thus ?thesis using fresh_subst_gv assms by metis +qed + +subsection \Weakening Immutable Variable Context\ + +declare check_s_check_branch_s_check_branch_list.intros[simp] +declare check_s_check_branch_s_check_branch_list.intros[intro] + +lemma check_s_d_weakening: + fixes s::s and v::v and cs::branch_s and css::branch_list + shows " \ ; \ ; \ ; \ ; \ \ s \ \ \ setD \ \ setD \' \ wfD \ \ \ \' \ \ ; \ ; \ ; \ ; \' \ s \ \" and + "check_branch_s \ \ \ \ \ tid cons const v cs \ \ setD \ \ setD \' \ wfD \ \ \ \' \ check_branch_s \ \ \ \ \' tid cons const v cs \" and + "check_branch_list \ \ \ \ \ tid dclist v css \ \ setD \ \ setD \' \ wfD \ \ \ \' \ check_branch_list \ \ \ \ \' tid dclist v css \" +proof(nominal_induct \ and \ and \ avoiding: \' arbitrary: v rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_valI \ \ \ \ \ v \' \) + then show ?case using check_s_check_branch_s_check_branch_list.intros by blast +next + case (check_letI x \ \ \ \ \ e \ z s b c) + show ?case proof + show "atom x \ (\, \, \, \, \', e, \)" using check_letI by auto + show "atom z \ (x, \, \, \, \, \', e, \, s)" using check_letI by auto + show "\ ; \ ; \ ; \ ; \' \ e \ \ z : b | c \" using check_letI infer_e_d_weakening by auto + have "\ ; \ \\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \" using check_letI check_s_wf by metis + moreover have "\; \; \ \\<^sub>w\<^sub>f \'" using check_letI check_s_wf by metis + ultimately have "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) toSet.simps by fast + thus "\ ; \ ; \ ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ ; \' \ s \ \" using check_letI by simp + qed +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + moreover have "\ ;\ \\<^sub>w\<^sub>f (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \" + using check_s_wf[OF check_branch_s_branchI(16) ] by metis + moreover hence " \ ;\ ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" + using wf_weakening2(6) check_branch_s_branchI by fastforce + ultimately show ?case + using check_s_check_branch_s_check_branch_list.intros by simp +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case proof + show \atom z \ (\, \, \, \, \', v, s1, s2, \)\ using fresh_prodN check_ifI by auto + show \\; \; \ \ v \ \ z : B_bool | TRUE \\ using check_ifI by auto + show \ \ ; \ ; \ ; \ ; \' \ s1 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_true) IMP c_of \ z \\ using check_ifI by auto + show \ \ ; \ ; \ ; \ ; \' \ s2 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \\ using check_ifI by auto + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + show "atom x \ (\, \, \, \, \', c, \,s)" using fresh_prodN check_assertI by auto + show *:" \; \; \ \\<^sub>w\<^sub>f \' " using check_assertI by auto + hence "\; \; (x, B_bool, c) #\<^sub>\ \ \\<^sub>w\<^sub>f \' " using wf_weakening2(6)[OF *, of " (x, B_bool, c) #\<^sub>\ \"] check_assertI check_s_wf toSet.simps by auto + thus "\ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \' \ s \ \" + using check_assertI(11)[OF \setD \ \ setD \'\] by simp + (* wf_weakening2(6) check_s_wf check_assertI *) + show "\; \; \ \ c " using fresh_prodN check_assertI by auto + + qed +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + show ?case proof + show "atom x \ (\, \, \, G, \', t , s1, \)" using check_let2I by auto + + show "\ ; \ ; \ ; G ; \' \ s1 \ t" using check_let2I infer_e_d_weakening by auto + have "\; \; (x, b_of t, c_of t x ) #\<^sub>\ G \\<^sub>w\<^sub>f \'" using check_let2I wf_weakening2(6) check_s_wf by fastforce + thus "\ ; \ ; \ ; (x, b_of t, c_of t x) #\<^sub>\ G ; \' \ s2 \ \" using check_let2I by simp + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof + show "atom u \ (\, \, \, \, \', \', v, \)" using check_varI by auto + show "\; \; \ \ v \ \'" using check_varI by auto + have "setD ((u, \') #\<^sub>\\) \ setD ((u, \') #\<^sub>\\')" using setD.simps check_varI by auto + moreover have "u \ fst ` setD \'" using check_varI(1) setD.simps fresh_DCons by (simp add: fresh_d_not_in) + moreover hence "\; \; \ \\<^sub>w\<^sub>f (u, \') #\<^sub>\\' " using wfD_cons fresh_DCons setD.simps check_varI check_v_wf by metis + ultimately show "\ ; \ ; \ ; \ ; (u, \') #\<^sub>\\' \ s \ \" using check_varI by auto + qed +next + case (check_assignI \ \ \ \ \ u \ v z \') + moreover hence "(u, \) \ setD \'" by auto + ultimately show ?case using Typing.check_assignI by simp +next + case (check_whileI \ \ \ \ \ s1 z s2 \') + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_seqI \ \ \ \ \ s1 z s2 \) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson + +qed + +lemma valid_ce_eq: + fixes v::v and ce2::ce + assumes "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" and "wfV \ \ GNil v b" and "wfCE \ \ ((x, b, TRUE) #\<^sub>\ GNil) ce2 b'" and "wfCE \ \ GNil ce1 b'" + shows \\; \; (x, b, ([[x ]\<^sup>v]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\ GNil \ ce1 == ce2 \ + unfolding valid.simps proof + have wfg: "\ ; \ \\<^sub>w\<^sub>f (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil" + using wfG_cons1I wfG_nilI wfX_wfY assms wf_intros + by (meson fresh_GNil wfC_e_eq wfG_intros2) + + show wf: \\; \; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce1 == ce2 \ + apply(rule wfC_eqI[where b=b']) + using wfg toSet.simps assms wfCE_weakening apply simp + + using wfg assms wf_replace_inside1(8) assms + using wfC_trueI wf_trans(8) by auto + + show \\i. ((\ ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil \ i) \ (i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil) \ + (i \ (ce1 == ce2 )))\ proof(rule+,goal_cases) + fix i + assume as:"\ ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil \ i \ i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil" + have 1:"wfV \ \ ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil) v b" + using wf_weakening assms append_g.simps toSet.simps wf wfX_wfY + by (metis empty_subsetI) + hence "\s. i \ v \ ~ s" using eval_v_exist[OF _ 1] as by auto + then obtain s where iv:"i\ v \ ~ s" .. + + hence ix:"i x = Some s" proof - + have "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e" using is_satis_g.simps as by auto + hence "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ ~ True" using is_satis.simps by auto + hence "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s" using + iv eval_e_elims + by (metis eval_c_elims(7) eval_e_uniqueness eval_e_valI) + thus ?thesis using eval_v_elims(2) eval_e_elims(1) by metis + qed + + have 1:"wfCE \ \ ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil) ce1 b'" + using wfCE_weakening assms append_g.simps toSet.simps wf wfX_wfY + by (metis empty_subsetI) + hence "\s1. i \ ce1 \ ~ s1" using eval_e_exist assms as by auto + then obtain s1 where s1: "i\ce1\ ~ s1" .. + + moreover have "i\ ce2 \ ~ s1" proof - + have "i\ ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v \ ~ s1" using assms s1 by auto + moreover have "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" using subst_v_ce_def assms subst_v_simple_commute by auto + ultimately have "i(x \ s) \ ce2 \ ~ s1" + using ix subst_e_eval_v[of i ce1 s1 "ce2[z::=[ x ]\<^sup>v]\<^sub>v" x v s] iv s1 by auto + moreover have "i(x \ s) = i" using ix by auto + ultimately show ?thesis by auto + qed + ultimately show "i \ ce1 == ce2 \ ~ True " using eval_c_eqI by metis + qed +qed + + +lemma check_v_top: + fixes v::v + assumes "\; \; GNil \ v \ \" and "ce1 = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" and "\; \; GNil \\<^sub>w\<^sub>f \ z : b_of \ | ce1 == ce2 \" + and "supp ce1 \ supp \" + shows "\; \; GNil \ v \ \ z : b_of \ | ce1 == ce2 \" +proof - + obtain t where t: "\; \; GNil \ v \ t \ \; \; GNil \ t \ \" + using assms check_v_elims by metis + + then obtain z' and b' where *:"t = \ z' : b' | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ \ atom z' \ v \ atom z' \ (\, \,GNil)" + using assms infer_v_form by metis + have beq: "b_of t = b_of \" using subtype_eq_base2 b_of.simps t by auto + obtain x::x where xf: \atom x \ (\, \, GNil, z', [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e , z, ce1 == ce2 )\ + using obtain_fresh by metis + + have "\; \; (x, b_of \, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v )" + using wfT_wfC2[OF assms(3), of x] subst_cv.simps(6) subst_v_c_def subst_v_ce_def fresh_GNil by simp + + then obtain b2 where b2: "\; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \ + \; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2" using wfC_elims(3) + beq by metis + + from xf have "\; \; GNil \ \ z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ \ \ z : b_of t | ce1 == ce2 \" + proof + show \ \; \; GNil \\<^sub>w\<^sub>f \ z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ \ using b_of.simps assms infer_v_wf t * by auto + show \ \; \; GNil \\<^sub>w\<^sub>f \ z : b_of t | ce1 == ce2 \ \ using beq assms by auto + have \\; \; (x, b_of t, ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\ GNil \ (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v ) \ + proof(rule valid_ce_eq) + show \ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\ proof - + have "atom z \ ce1" using assms fresh_def x_not_in_b_set by fast + hence "ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce1" + using forget_subst_v by auto + also have "... = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" using assms by auto + also have "... = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v" proof - + have "atom x \ ce2" using xf fresh_prodN c.fresh by metis + thus ?thesis using subst_v_simple_commute subst_v_ce_def by simp + qed + finally show ?thesis by auto + qed + show \ \; \; GNil \\<^sub>w\<^sub>f v : b_of t \ using infer_v_wf t by simp + show \ \; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2 \ using b2 by auto + + have " \; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2" using b2 by auto + moreover have "atom x \ ce1[z::=[ x ]\<^sup>v]\<^sub>v" + using fresh_subst_v_if assms fresh_def + using \\; \; GNil \\<^sub>w\<^sub>f v : b_of t\ \ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\ + fresh_GNil subst_v_ce_def wfV_x_fresh by auto + ultimately show \ \; \; GNil \\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \ using + wf_restrict(8) by force + qed + moreover have v: "v[z'::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = v" + using forget_subst assms infer_v_wf wfV_supp x_not_in_b_set + by (simp add: "local.*") + ultimately show "\; \; (x, b_of t, ([ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ (ce1 == ce2 )[z::=[ x ]\<^sup>v]\<^sub>v" + unfolding subst_cv.simps subst_v_c_def subst_cev.simps subst_vv.simps + using subst_v_ce_def by simp + qed + thus ?thesis using b_of.simps assms * check_v_subtypeI t b_of.simps subtype_eq_base2 by metis +qed + +end + diff --git a/thys/MiniSail/Wellformed.thy b/thys/MiniSail/Wellformed.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Wellformed.thy @@ -0,0 +1,581 @@ +(*<*) +theory Wellformed + imports IVSubst BTVSubst +begin + (*>*) + +chapter \Wellformed Terms\ + +text \We require that expressions and values are well-formed. This includes a notion +of well-sortedness. We identify a sort with a basic type and +define the judgement as two clusters of mutually recursive +inductive predicates. Some of the proofs are across all of the predicates and +although they seemed at first to be daunting, they have all +worked out well.\ + +named_theorems ms_wb "Facts for helping with well-sortedness" + +section \Definitions\ + +inductive wfV :: "\ \ \ \ \ \ v \ b \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfC :: "\ \ \ \ \ \ c \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfG :: "\ \ \ \ \ \ bool" (" _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfT :: "\ \ \ \ \ \ \ \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfTs :: "\ \ \ \ \ \ (string*\) list \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfTh :: "\ \ bool" (" \\<^sub>w\<^sub>f _ " [50] 50) and + wfB :: "\ \ \ \ b \ bool" (" _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfCE :: "\ \ \ \ \ \ ce \ b \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfTD :: "\ \ type_def \ bool" (" _ \\<^sub>w\<^sub>f _ " [50,50] 50) + where + +wfB_intI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_int" +| wfB_boolI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_bool" +| wfB_unitI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_unit" +| wfB_bitvecI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_bitvec" +| wfB_pairI: "\ \; \ \\<^sub>w\<^sub>f b1 ; \; \ \\<^sub>w\<^sub>f b2 \ \ \; \ \\<^sub>w\<^sub>f B_pair b1 b2" + +| wfB_consI: "\ + \\<^sub>w\<^sub>f \; + (AF_typedef s dclist) \ set \ +\ \ + \; \ \\<^sub>w\<^sub>f B_id s" + +| wfB_appI: "\ + \\<^sub>w\<^sub>f \; + \; \ \\<^sub>w\<^sub>f b; + (AF_typedef_poly s bv dclist) \ set \ +\ \ + \; \ \\<^sub>w\<^sub>f B_app s b" + +| wfV_varI: "\ \; \ \\<^sub>w\<^sub>f \ ; Some (b,c) = lookup \ x \ \ \; \; \ \\<^sub>w\<^sub>f V_var x : b" +| wfV_litI: "\; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f V_lit l : base_for_lit l" + +| wfV_pairI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : b1 ; + \; \; \ \\<^sub>w\<^sub>f v2 : b2 +\ \ + \; \; \ \\<^sub>w\<^sub>f (V_pair v1 v2) : B_pair b1 b2" + +| wfV_consI: "\ + AF_typedef s dclist \ set \; + (dc, \ x : b' | c \) \ set dclist ; + \; \; \ \\<^sub>w\<^sub>f v : b' +\ \ + \; \; \ \\<^sub>w\<^sub>f V_cons s dc v : B_id s" + +| wfV_conspI: "\ + AF_typedef_poly s 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 +\ \ + \; \; \ \\<^sub>w\<^sub>f V_consp s dc b v : B_app s b" + +| wfCE_valI : "\ + \; \; \ \\<^sub>w\<^sub>f v : b +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_val v : b" + +| wfCE_plusI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_op Plus v1 v2 : B_int" + +| wfCE_leqI:"\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_op LEq v1 v2 : B_bool" + +| wfCE_eqI:"\ + \; \; \ \\<^sub>w\<^sub>f v1 : b; + \; \; \ \\<^sub>w\<^sub>f v2 : b +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_op Eq v1 v2 : B_bool" + +| wfCE_fstI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_fst v1 : b1" + +| wfCE_sndI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_snd v1 : b2" + +| wfCE_concatI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec ; + \; \; \ \\<^sub>w\<^sub>f v2 : B_bitvec +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_concat v1 v2 : B_bitvec" + +| wfCE_lenI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_len v1 : B_int" + +| wfTI : "\ + atom z \ (\, \, \) ; + \; \ \\<^sub>w\<^sub>f b; + \; \ ; (z,b,C_true) #\<^sub>\ \ \\<^sub>w\<^sub>f c +\ \ + \; \; \ \\<^sub>w\<^sub>f \ z : b | c \" + +| wfC_eqI: "\ + \; \; \ \\<^sub>w\<^sub>f e1 : b ; + \; \; \ \\<^sub>w\<^sub>f e2 : b \ \ + \; \; \ \\<^sub>w\<^sub>f C_eq e1 e2" +| wfC_trueI: " \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f C_true " +| wfC_falseI: " \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f C_false " + +| wfC_conjI: "\ \; \; \ \\<^sub>w\<^sub>f c1 ; \; \; \ \\<^sub>w\<^sub>f c2 \ \ \; \; \ \\<^sub>w\<^sub>f C_conj c1 c2" +| wfC_disjI: "\ \; \; \ \\<^sub>w\<^sub>f c1 ; \; \; \ \\<^sub>w\<^sub>f c2 \ \ \; \; \ \\<^sub>w\<^sub>f C_disj c1 c2" +| wfC_notI: "\ \; \; \ \\<^sub>w\<^sub>f c1 \ \ \; \; \ \\<^sub>w\<^sub>f C_not c1" +| wfC_impI: "\ \; \; \ \\<^sub>w\<^sub>f c1 ; + \; \; \ \\<^sub>w\<^sub>f c2 \ \ \; \; \ \\<^sub>w\<^sub>f C_imp c1 c2" + +| wfG_nilI: " \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f GNil" +| wfG_cons1I: "\ c \ { TRUE, FALSE } ; + \; \ \\<^sub>w\<^sub>f \ ; + atom x \ \ ; + \ ; \ ; (x,b,C_true)#\<^sub>\\ \\<^sub>w\<^sub>f c ; wfB \ \ b + \ \ \; \ \\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\\)" +| wfG_cons2I: "\ c \ { TRUE, FALSE } ; + \; \ \\<^sub>w\<^sub>f \ ; + atom x \ \ ; + wfB \ \ b + \ \ \; \ \\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\\)" + +| wfTh_emptyI: " \\<^sub>w\<^sub>f []" + +| wfTh_consI: "\ + (name_of_type tdef) \ name_of_type ` set \ ; + \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f tdef \ \ \\<^sub>w\<^sub>f tdef#\" + +| wfTD_simpleI: "\ + \ ; {||} ; GNil \\<^sub>w\<^sub>f lst + \ \ + \ \\<^sub>w\<^sub>f (AF_typedef s lst )" + +| wfTD_poly: "\ + \ ; {|bv|} ; GNil \\<^sub>w\<^sub>f lst + \ \ + \ \\<^sub>w\<^sub>f (AF_typedef_poly s bv lst)" + +| wfTs_nil: "\; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f []::(string*\) list" +| wfTs_cons: "\ \; \; \ \\<^sub>w\<^sub>f \ ; + dc \ fst ` set ts; + \; \; \ \\<^sub>w\<^sub>f ts::(string*\) list \ \ \; \; \ \\<^sub>w\<^sub>f ((dc,\)#ts)" + +inductive_cases wfC_elims: + "\; \; \ \\<^sub>w\<^sub>f C_true" + "\; \; \ \\<^sub>w\<^sub>f C_false" + "\; \; \ \\<^sub>w\<^sub>f C_eq e1 e2" + "\; \; \ \\<^sub>w\<^sub>f C_conj c1 c2" + "\; \; \ \\<^sub>w\<^sub>f C_disj c1 c2" + "\; \; \ \\<^sub>w\<^sub>f C_not c1" + "\; \; \ \\<^sub>w\<^sub>f C_imp c1 c2" + +inductive_cases wfV_elims: + "\; \; \ \\<^sub>w\<^sub>f V_var x : b" + "\; \; \ \\<^sub>w\<^sub>f V_lit l : b" + "\; \; \ \\<^sub>w\<^sub>f V_pair v1 v2 : b" + "\; \; \ \\<^sub>w\<^sub>f V_cons tyid dc v : b" + "\; \; \ \\<^sub>w\<^sub>f V_consp tyid dc b v : b'" + +inductive_cases wfCE_elims: + " \; \; \ \\<^sub>w\<^sub>f CE_val v : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op Plus v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op LEq v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_fst v1 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_snd v1 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_concat v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_len v1 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op opp v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op Eq v1 v2 : b" + +inductive_cases wfT_elims: + "\; \ ; \ \\<^sub>w\<^sub>f \::\" + "\; \ ; \ \\<^sub>w\<^sub>f \ z : b | c \" + +inductive_cases wfG_elims: + "\; \ \\<^sub>w\<^sub>f GNil" + "\; \ \\<^sub>w\<^sub>f (x,b,c)#\<^sub>\\" + "\; \ \\<^sub>w\<^sub>f (x,b,TRUE)#\<^sub>\\" + "\; \ \\<^sub>w\<^sub>f (x,b,FALSE)#\<^sub>\\" + +inductive_cases wfTh_elims: + " \\<^sub>w\<^sub>f []" + " \\<^sub>w\<^sub>f td#\" + +inductive_cases wfTD_elims: + "\ \\<^sub>w\<^sub>f (AF_typedef s lst )" + "\ \\<^sub>w\<^sub>f (AF_typedef_poly s bv lst )" + +inductive_cases wfTs_elims: + "\; \ ; GNil \\<^sub>w\<^sub>f ([]::((string*\) list))" + "\; \ ; GNil \\<^sub>w\<^sub>f ((t#ts)::((string*\) list))" + +inductive_cases wfB_elims: + " \; \ \\<^sub>w\<^sub>f B_pair b1 b2" + " \; \ \\<^sub>w\<^sub>f B_id s" + " \; \ \\<^sub>w\<^sub>f B_app s b" + +equivariance wfV + +text \This setup of 'avoiding' is not complete and for some of lemmas, such as weakening, +do it the hard way\ +nominal_inductive wfV + avoids wfV_conspI: bv | wfTI: z +proof(goal_cases) + case (1 s bv dclist \ dc x b' c \ b \ v) + + moreover hence "atom bv \ V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis + moreover have "atom bv \ B_app s b" using b.fresh fresh_prodN pure_fresh 1 by metis + ultimately show ?case using b.fresh v.fresh pure_fresh fresh_star_def fresh_prodN by fastforce +next + case (2 s bv dclist \ dc x b' c \ b \ v) + then show ?case by auto +next + case (3 z \ \ \ b c) + then show ?case using \.fresh fresh_star_def fresh_prodN by fastforce +next + case (4 z \ \ \ b c) + then show ?case by auto +qed + +inductive + wfE :: "\ \ \ \ \ \ \ \ \ \ e \ b \ bool" (" _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfS :: "\ \ \ \ \ \ \ \ \ \ s \ b \ bool" (" _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfCS :: "\ \ \ \ \ \ \ \ \ \ tyid \ string \ \ \ branch_s \ b \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and + wfCSS :: "\ \ \ \ \ \ \ \ \ \ tyid \ (string * \) list \ branch_list \ b \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and + wfPhi :: "\ \ \ \ bool" (" _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfD :: "\ \ \ \ \ \ \ \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfFTQ :: "\ \ \ \ fun_typ_q \ bool" (" _ ; _ \\<^sub>w\<^sub>f _ " [50] 50) and + wfFT :: "\ \ \ \ \ \ fun_typ \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50] 50) where + +wfE_valI : "\ + \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_val v : b" + +| wfE_plusI: "\ + \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_op Plus v1 v2 : B_int" + +| wfE_leqI:"\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_op LEq v1 v2 : B_bool" + +| wfE_eqI:"\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : b; + \; \; \ \\<^sub>w\<^sub>f v2 : b; + b \ { B_bool, B_int, B_unit } +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_op Eq v1 v2 : B_bool" + +| wfE_fstI: "\ + \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 + \ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_fst v1 : b1" + +| wfE_sndI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_snd v1 : b2" + +| wfE_concatI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec; + \; \; \ \\<^sub>w\<^sub>f v2 : B_bitvec +\ \ + \ ; \ ; \ ; \; \ \\<^sub>w\<^sub>f AE_concat v1 v2 : B_bitvec" + +| wfE_splitI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \ ; \ ; \ ; \; \ \\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec" + +| wfE_lenI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_len v1 : B_int" + +| wfE_appI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))) = lookup_fun \ f ; + \; \; \ \\<^sub>w\<^sub>f v : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_app f v : b_of \" + +| wfE_appPI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \ \\<^sub>w\<^sub>f b'; + atom bv \ (\, \, \, \, \, b', v, (b_of \)[bv::=b']\<^sub>b); + Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f; + \; \; \ \\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f (AE_appP f b' v) : ((b_of \)[bv::=b']\<^sub>b)" + +| wfE_mvarI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + (u,\) \ setD \ +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_mvar u : b_of \" + +| wfS_valI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f v : b ; + \; \; \ \\<^sub>w\<^sub>f \ +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f (AS_val v) : b" + +| wfS_letI: "\ + wfE \ \ \ \ \ e b' ; + \ ; \ ; \ ; (x,b',C_true) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b; + \; \; \ \\<^sub>w\<^sub>f \ ; + atom x \ (\, \, \, \, \, e, b) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f LET x = e IN s : b" + +| wfS_assertI: "\ + \ ; \ ; \ ; (x,B_bool,c) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b; + \; \; \ \\<^sub>w\<^sub>f c ; + \; \; \ \\<^sub>w\<^sub>f \ ; + atom x \ (\, \, \, \, \, c, b, s) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f ASSERT c IN s : b" + +| wfS_let2I: "\ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : b_of \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \ ; \ ; \ ; (x,b_of \,C_true) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : b ; + atom x \ (\, \, \, \, \, s1, b,\) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f LET x : \ = s1 IN s2 : b" + +| wfS_ifI: "\ + \; \; \ \\<^sub>w\<^sub>f v : B_bool; + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : b ; + \; \; \; \; \ \\<^sub>w\<^sub>f s2 : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f IF v THEN s1 ELSE s2 : b" + +| wfS_varI : "\ + wfT \ \ \ \ ; + \; \; \ \\<^sub>w\<^sub>f v : b_of \; + atom u \ (\, \, \, \, \, \, v, b); + \ ; \ ; \ ; \ ; (u,\)#\<^sub>\\ \\<^sub>w\<^sub>f s : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f VAR u : \ = v IN s : b " + +| wfS_assignI: "\ + (u,\) \ setD \ ; + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f v : b_of \ +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f u ::= v : B_unit" + +| wfS_whileI: "\ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : B_bool ; + \; \; \; \; \ \\<^sub>w\<^sub>f s2 : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f WHILE s1 DO { s2 } : b" + +| wfS_seqI: "\ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : B_unit ; + \; \; \; \; \ \\<^sub>w\<^sub>f s2 : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 ;; s2 : b" + +| wfS_matchI: "\ + wfV \ \ \ v (B_id tid) ; + (AF_typedef tid dclist ) \ set \; + wfD \ \ \ \ ; + \ \\<^sub>w\<^sub>f \ ; + \ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f cs : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AS_match v cs : b " + +| wfS_branchI: "\ + \ ; \ ; \ ; (x,b_of \,C_true) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b ; + atom x \ (\, \, \, \, \, \,\); + \; \; \ \\<^sub>w\<^sub>f \ +\ \ + \ ; \ ; \ ; \ ; \ ; tid ; dc ; \ \\<^sub>w\<^sub>f dc x \ s : b" + +| wfS_finalI: "\ + \; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b + \ \ + \ ; \ ; \ ; \ ; \ ; tid ; [(dc,t)] \\<^sub>w\<^sub>f AS_final cs : b " + +| wfS_cons: "\ + \; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b; + \; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b + \ \ + \ ; \ ; \ ; \ ; \ ; tid ; (dc,t)#dclist \\<^sub>w\<^sub>f AS_cons cs css : b " + +| wfD_emptyI: "\; \ \\<^sub>w\<^sub>f \ \ \ ; \ ; \ \\<^sub>w\<^sub>f []\<^sub>\" + +| wfD_cons: "\ + \ ; \ ; \ \\<^sub>w\<^sub>f \::\ ; + \ ; \ ; \ \\<^sub>w\<^sub>f \; + u \ fst ` setD \ +\ \ + \; \; \ \\<^sub>w\<^sub>f ((u,\) #\<^sub>\ \)" + +| wfPhi_emptyI: " \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f []" + +| wfPhi_consI: "\ + f \ name_of_fun ` set \; + \ ; \ \\<^sub>w\<^sub>f ft; + \ \\<^sub>w\<^sub>f \ +\ \ + \ \\<^sub>w\<^sub>f ((AF_fundef f ft)#\)" + +| wfFTNone: " \ ; \ ; {||} \\<^sub>w\<^sub>f ft \ \ ; \ \\<^sub>w\<^sub>f AF_fun_typ_none ft" +| wfFTSome: " \ ; \ ; {| bv |} \\<^sub>w\<^sub>f ft \ \ ; \ \\<^sub>w\<^sub>f AF_fun_typ_some bv ft" + +| wfFTI: "\ + \ ; B \\<^sub>w\<^sub>f b; + supp s \ {atom x} \ supp B ; + supp c \ { atom x } ; + \ ; B ; (x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \; + \ \\<^sub>w\<^sub>f \ +\ \ + \ ; \ ; B \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" + +inductive_cases wfE_elims: + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_val v : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op Plus v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op LEq v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_fst v1 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_snd v1 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_concat v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_len v1 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op opp v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_app f v: b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_appP f b' v: b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_mvar u : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op Eq v1 v2 : b" + +inductive_cases wfCS_elims: + "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f (cs::branch_s) : b" + "\; \; \; \; \ ; tid ; dc \\<^sub>w\<^sub>f (cs::branch_list) : b" + +inductive_cases wfPhi_elims: + " \ \\<^sub>w\<^sub>f []" + " \ \\<^sub>w\<^sub>f ((AF_fundef f ft)#\)" + " \ \\<^sub>w\<^sub>f (fd#\::\)" + +declare[[ simproc del: alpha_lst]] + +inductive_cases wfFTQ_elims: + "\ ; \ \\<^sub>w\<^sub>f AF_fun_typ_none ft" + "\ ; \ \\<^sub>w\<^sub>f AF_fun_typ_some bv ft" + "\ ; \ \\<^sub>w\<^sub>f AF_fun_typ_some bv (AF_fun_typ x b c \ s)" + +inductive_cases wfFT_elims: + "\ ; \ ; \ \\<^sub>w\<^sub>f AF_fun_typ x b c \ s" + +declare[[ simproc add: alpha_lst]] + +inductive_cases wfD_elims: + "\ ; \ ; (\::\) \\<^sub>w\<^sub>f []\<^sub>\" + "\ ; \ ; (\::\) \\<^sub>w\<^sub>f (u,\) #\<^sub>\ \::\" + +equivariance wfE + +nominal_inductive wfE + avoids wfE_appPI: bv | wfS_varI: u | wfS_letI: x | wfS_let2I: x | wfS_branchI: x | wfS_assertI: x + +proof(goal_cases) + case (1 \ \ \ \ \ b' bv v \ f x b c s) + moreover hence "atom bv \ AE_appP f b' v" using pure_fresh fresh_prodN e.fresh by auto + ultimately show ?case using fresh_star_def by fastforce +next + case (2 \ \ \ \ \ b' bv v \ f x b c s) + then show ?case by auto +next + case (3 \ \ \ \ \ e b' x s b) + moreover hence "atom x \ LET x = e IN s" using fresh_prodN by auto + ultimately show ?case using fresh_prodN fresh_star_def by fastforce +next + case (4 \ \ \ \ \ e b' x s b) + then show ?case by auto +next + case (5 \ \ \ x c \ \ s b) + hence "atom x \ ASSERT c IN s" using s_branch_s_branch_list.fresh by auto + then show ?case using fresh_prodN fresh_star_def 5 by fastforce +next + case (6 \ \ \ x c \ \ s b) + then show ?case by auto +next + case (7 \ \ \ \ \ s1 \ x s2 b) + hence "atom x \ \ \ atom x \ s1" using fresh_prodN by metis + moreover hence "atom x \ LET x : \ = s1 IN s2" + using s_branch_s_branch_list.fresh(3)[of "atom x" x "\" s1 s2 ] fresh_prodN by simp + ultimately show ?case using fresh_prodN fresh_star_def 7 by fastforce +next + case (8 \ \ \ \ \ s1 \ x s2 b) + then show ?case by auto +next + case (9 \ \ \ \ v u \ \ b s) + moreover hence " atom u \ AS_var u \ v s" using fresh_prodN s_branch_s_branch_list.fresh by simp + ultimately show ?case using fresh_star_def fresh_prodN s_branch_s_branch_list.fresh by fastforce +next + case (10 \ \ \ \ v u \ \ b s) + then show ?case by auto +next + case (11 \ \ \ x \ \ \ s b tid dc) + moreover have "atom x \ (dc x \ s)" using pure_fresh s_branch_s_branch_list.fresh by auto + ultimately show ?case using fresh_prodN fresh_star_def pure_fresh by fastforce +next + case (12 \ \ \ x \ \ \ s b tid dc) + then show ?case by auto +qed + +inductive wfVDs :: "var_def list \ bool" where + +wfVDs_nilI: "wfVDs []" + +| wfVDs_consI: "\ + atom u \ ts; + wfV ([]::\) {||} GNil v (b_of \); + wfT ([]::\) {||} GNil \; + wfVDs ts +\ \ wfVDs ((AV_def u \ v)#ts)" + +equivariance wfVDs +nominal_inductive wfVDs . + +end \ No newline at end of file diff --git a/thys/MiniSail/WellformedL.thy b/thys/MiniSail/WellformedL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/WellformedL.thy @@ -0,0 +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) + 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) + 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)) + 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 diff --git a/thys/MiniSail/document/root.bib b/thys/MiniSail/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/MiniSail/document/root.bib @@ -0,0 +1,29 @@ +@article{Armstrong2019, +abstract = {Architecture specifications notionally define the fundamental interface between hardware and software: the envelope of allowed behaviour for processor implementations, and the basic assumptions for software development and verification. But in practice, they are typically prose and pseudocode documents, not rigorous or executable artifacts, leaving software and verification on shaky ground. In this paper, we present rigorous semantic models for the sequential behaviour of large parts of the mainstream ARMv8-A, RISC-V, and MIPS architectures, and the research CHERI-MIPS architecture, that are complete enough to boot operating systems, variously Linux, FreeBSD, or seL4. Our ARMv8-A models are automatically translated from authoritative ARM-internal definitions, and (in one variant) tested against the ARM Architecture Validation Suite. We do this using a custom language for ISA semantics, Sail, with a lightweight dependent type system, that supports automatic generation of emulator code in C and OCaml, and automatic generation of proof-assistant definitions for Isabelle, HOL4, and (currently only for MIPS) Coq. We use the former for validation, and to assess specification coverage. To demonstrate the usability of the latter, we prove (in Isabelle) correctness of a purely functional characterisation of ARMv8-A address translation. We moreover integrate the RISC-V model into the RMEM tool for (user-mode) relaxed-memory concurrency exploration. We prove (on paper) the soundness of the core Sail type system. We thereby take a big step towards making the architectural abstraction actually well-defined, establishing foundations for verification and reasoning.}, +author = {Armstrong, Alasdair and Pulte, Christopher and Flur, Shaked and Stark, Ian and Krishnaswami, Neel and Sewell, Peter and Bauereiss, Thomas and Campbell, Brian and Reid, Alastair and Gray, Kathryn E. and Norton, Robert M. and Mundkur, Prashanth and Wassell, Mark and French, Jon}, +doi = {10.1145/3290384}, +file = {:E$\backslash$:/Mark/PhD/ReadingMaterial/SAIL/sail-popl2019.pdf:pdf}, +issn = {2475-1421}, +journal = {Proceedings of the ACM on Programming Languages}, +number = {POPL}, +pages = {1--31}, +title = {{ISA semantics for ARMv8-A, RISC-V, and CHERI-MIPS}}, +volume = {3}, +year = {2019} +} +@article{Vazou2014, +abstract = {SMT-based checking of refinement types for call-by-value languages is a well-studied subject. Unfortunately, the classical translation of refinement types to verification conditions is unsound under lazy evaluation. When checking an expression, such systems implicitly assume that all the free variables in the expression are bound to values. This property is trivially guaranteed by eager, but does not hold under lazy, evaluation. Thus, to be sound and precise, a refinement type system for Haskell and the corresponding verification conditions must take into account which subset of binders actually reduces to values. We present a stratified type system that labels binders as potentially diverging or not, and that (circularly) uses refinement types to verify the labeling. We have implemented our system in LIQUIDHASKELL and present an experimental evaluation of our approach on more than 10,000 lines of widely used Haskell libraries. We show that LIQUIDHASKELL is able to prove 96{\%} of all recursive functions terminating, while requiring a modest 1.7 lines of termination-annotations per 100 lines of code.}, +annote = {ICFP Video - https://www.youtube.com/watch?v=vqvNQixKr6w}, +archivePrefix = {arXiv}, +arxivId = {arXiv:1604.02480v1}, +author = {Vazou, Niki and Seidel, Eric L and Peyton-jones, Simon}, +doi = {10.1145/2628136.2628161}, +eprint = {arXiv:1604.02480v1}, +file = {:C$\backslash$:/Users/User/AppData/Local/Mendeley Ltd./Mendeley Desktop/Downloaded/Merz, Vanzetto - 2014 - Refinement types for Haskell.pdf:pdf}, +isbn = {9781450328739}, +issn = {15232867}, +journal = {ICFP '14 Proceedings of the 19th ACM SIGPLAN international conference on Functional programming}, +mendeley-groups = {Semantic Tools/RefinementTypes}, +title = {{Refinement Types For Haskell}}, +year = {2014} +} diff --git a/thys/MiniSail/document/root.tex b/thys/MiniSail/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/MiniSail/document/root.tex @@ -0,0 +1,71 @@ +\documentclass[11pt,a4paper]{report} +\usepackage{graphicx,isabelle,isabellesym} + +\usepackage[a4paper,includeheadfoot,margin=2.54cm]{geometry} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + +\begin{document} + +\title{MiniSail} +\author{Mark P. Wassell} +\maketitle + +\begin{abstract} +MiniSail is a kernel language for Sail~\cite{Armstrong2019}, +an instruction set architecture (ISA) specification language. +Sail is an imperative language with a light-weight dependent type system similar to refinement type systems such as~\cite{Vazou2014}. +From an ISA specification, the Sail compiler can generate theorem prover code and C (or OCaml) to give an executable +emulator for an architecture. +The idea behind MiniSail is to capture the key and novel features of Sail in terms of their syntax, +typing rules and operational semantics, +and to confirm that they work together by proving progress and preservation lemmas. +We use the Nominal2 library to handle binding. +\end{abstract} + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +\begin{center} + \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph} +\end{center} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,606 +1,607 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BNF_CC BNF_Operations BTree Banach_Steinhaus Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CSP_RefTK CYK CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties DFS_Framework DOM_Components DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpreter_Optimizations Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl IsaGeoCoq Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML +MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PCF PLM POPLmark-deBruijn PSemigroupsConvolution Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regression_Test_Selection Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL