diff -r 9fb315392493 -r 8732ff59068b Nominal/Ex/TypeSchemes.thy --- 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 \ ty) list \ name \ 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 \ ty) list \ ty \ ty" +and substs :: "(name \ ty) list \ tys \ tys" +assumes + s1: "subst \ (Var n) = lookup \ n" +and s2: "subst \ (Fun l r) = Fun (subst \ l) (subst \ r)" +and s3: "fset_to_set (fmap atom xs) \* \ \ substs \ (All xs t) = All xs (subst \ t)" +begin + +lemma subst_ty: + assumes x: "atom x \ t" + shows "subst [(x, S)] t = t" + using x + apply (induct t rule: ty_tys.induct[of _ "\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 \ T \ substs [(x, S)] T = T" + apply (rule strong_induct[of + "\a t. True" "\(x, S) T. (atom x \ T \ 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 \ fset_to_set (fmap atom fset)") + apply blast + apply (metis supp_finite_atom_set finite_fset) + done + +lemma subst_lemma_pre: + "z \ (N,L) \ z \ (subst [(y, L)] N)" + apply (induct N rule: ty_tys.induct[of _ "\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 \ (N,L) \ atom z \ (substs [(y, L)] N)" + apply (rule strong_induct[of + "\a t. True" "\(z, y, L) N. (atom z \ (N, L) \ atom z \ (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 \ supp (subst [(aa, b)] t)") + apply blast + apply (subgoal_tac "atom a \ supp t") + apply (fold fresh_def)[1] + apply (rule mp[OF subst_lemma_pre]) + apply (simp add: fresh_Pair) + apply (subgoal_tac "atom a \ (fset_to_set (fmap atom fset))") + apply blast + apply (metis supp_finite_atom_set finite_fset) + done + +lemma subst_lemma: + shows "x \ y \ atom x \ L \ + 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 _ "\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 \ y \ atom x \ L \ + substs [(y, L)] (substs [(x, N)] M) = + substs [(x, (subst [(y, L)] N))] (substs [(y, L)] M)" + apply (rule strong_induct[of + "\a t. True" "\(x, y, N, L) M. x \ y \ atom x \ L \ + 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