Nominal/Ex/TypeSchemes.thy
changeset 2301 8732ff59068b
parent 2181 b997c22805ae
child 2308 387fcbd33820
--- a/Nominal/Ex/TypeSchemes.thy	Wed May 26 15:34:54 2010 +0200
+++ b/Nominal/Ex/TypeSchemes.thy	Wed May 26 15:37:56 2010 +0200
@@ -152,6 +152,121 @@
   apply auto
   done
 
+fun
+  lookup :: "(name \<times> ty) list \<Rightarrow> name \<Rightarrow> ty"
+where
+  "lookup [] n = Var n"
+| "lookup ((p, s) # t) n = (if p = n then s else lookup t n)"
+
+locale subst_loc =
+fixes
+    subst  :: "(name \<times> ty) list \<Rightarrow> ty \<Rightarrow> ty"
+and substs :: "(name \<times> ty) list \<Rightarrow> tys \<Rightarrow> tys"
+assumes
+    s1: "subst \<theta> (Var n) = lookup \<theta> n"
+and s2: "subst \<theta> (Fun l r) = Fun (subst \<theta> l) (subst \<theta> r)"
+and s3: "fset_to_set (fmap atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All xs t) = All xs (subst \<theta> t)"
+begin
+
+lemma subst_ty:
+  assumes x: "atom x \<sharp> t"
+  shows "subst [(x, S)] t = t"
+  using x
+  apply (induct t rule: ty_tys.induct[of _ "\<lambda>t. True" _ , simplified])
+  by (simp_all add: s1 s2 fresh_def ty_tys.fv[simplified ty_tys.supp] supp_at_base)
+
+lemma subst_tyS:
+  shows "atom x \<sharp> T \<longrightarrow> substs [(x, S)] T = T"
+  apply (rule strong_induct[of
+    "\<lambda>a t. True" "\<lambda>(x, S) T. (atom x \<sharp> T \<longrightarrow> substs [(x, S)] T = T)" _ "t" "(x, S)", simplified])
+  apply clarify
+  apply (subst s3)
+  apply (simp add: fresh_star_def fresh_Cons fresh_Nil)
+  apply (subst subst_ty)
+  apply (simp_all add: fresh_star_prod_elim)
+  apply (drule fresh_star_atom)
+  apply (simp add: fresh_def ty_tys.fv[simplified ty_tys.supp])
+  apply (subgoal_tac "atom a \<notin> fset_to_set (fmap atom fset)")
+  apply blast
+  apply (metis supp_finite_atom_set finite_fset)
+  done
+
+lemma subst_lemma_pre:
+  "z \<sharp> (N,L) \<longrightarrow> z \<sharp> (subst [(y, L)] N)"
+  apply (induct N rule: ty_tys.induct[of _ "\<lambda>t. True" _ , simplified])
+  apply (simp add: s1)
+  apply (auto simp add: fresh_Pair)
+  apply (auto simp add: fresh_def ty_tys.fv[simplified ty_tys.supp])[3]
+  apply (simp add: s2)
+  apply (auto simp add: fresh_def ty_tys.fv[simplified ty_tys.supp])
+  done
+
+lemma substs_lemma_pre:
+  "atom z \<sharp> (N,L) \<longrightarrow> atom z \<sharp> (substs [(y, L)] N)"
+  apply (rule strong_induct[of
+    "\<lambda>a t. True" "\<lambda>(z, y, L) N. (atom z \<sharp> (N, L) \<longrightarrow> atom z \<sharp> (substs [(y, L)] N))" _ _ "(z, y, L)", simplified])
+  apply clarify
+  apply (subst s3)
+  apply (simp add: fresh_star_def fresh_Cons fresh_Nil fresh_Pair)
+  apply (simp_all add: fresh_star_prod_elim fresh_Pair)
+  apply clarify
+  apply (drule fresh_star_atom)
+  apply (drule fresh_star_atom)
+  apply (simp add: fresh_def)
+  apply (simp only: ty_tys.fv[simplified ty_tys.supp])
+  apply (subgoal_tac "atom a \<notin> supp (subst [(aa, b)] t)")
+  apply blast
+  apply (subgoal_tac "atom a \<notin> supp t")
+  apply (fold fresh_def)[1]
+  apply (rule mp[OF subst_lemma_pre])
+  apply (simp add: fresh_Pair)
+  apply (subgoal_tac "atom a \<notin> (fset_to_set (fmap atom fset))")
+  apply blast
+  apply (metis supp_finite_atom_set finite_fset)
+  done
+
+lemma subst_lemma:
+  shows "x \<noteq> y \<and> atom x \<sharp> L \<longrightarrow>
+    subst [(y, L)] (subst [(x, N)] M) =
+    subst [(x, (subst [(y, L)] N))] (subst [(y, L)] M)"
+  apply (induct M rule: ty_tys.induct[of _ "\<lambda>t. True" _ , simplified])
+  apply (simp_all add: s1 s2)
+  apply clarify
+  apply (subst (2) subst_ty)
+  apply simp_all
+  done
+
+lemma substs_lemma:
+  shows "x \<noteq> y \<and> atom x \<sharp> L \<longrightarrow>
+    substs [(y, L)] (substs [(x, N)] M) =
+    substs [(x, (subst [(y, L)] N))] (substs [(y, L)] M)"
+  apply (rule strong_induct[of
+    "\<lambda>a t. True" "\<lambda>(x, y, N, L) M. x \<noteq> y \<and> atom x \<sharp> L \<longrightarrow>
+    substs [(y, L)] (substs [(x, N)] M) =
+    substs [(x, (subst [(y, L)] N))] (substs [(y, L)] M)" _ _ "(x, y, N, L)", simplified])
+  apply clarify
+  apply (simp_all add: fresh_star_prod_elim fresh_Pair)
+  apply (subst s3)
+  apply (unfold fresh_star_def)[1]
+  apply (simp add: fresh_Cons fresh_Nil fresh_Pair)
+  apply (subst s3)
+  apply (unfold fresh_star_def)[1]
+  apply (simp add: fresh_Cons fresh_Nil fresh_Pair)
+  apply (subst s3)
+  apply (unfold fresh_star_def)[1]
+  apply (simp add: fresh_Cons fresh_Nil fresh_Pair)
+  apply (subst s3)
+  apply (unfold fresh_star_def)[1]
+  apply (simp add: fresh_Cons fresh_Nil fresh_Pair)
+  apply (rule ballI)
+  apply (rule mp[OF subst_lemma_pre])
+  apply (simp add: fresh_Pair)
+  apply (subst subst_lemma)
+  apply simp_all
+  done
+
+end
+
 (* PROBLEM:
 Type schemes with separate datatypes