diff -r 2406350c358f -r 8b5a1ad60487 Nominal/ExTySch.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/ExTySch.thy Tue Mar 23 08:51:43 2010 +0100 @@ -0,0 +1,166 @@ +theory ExTySch +imports "Parser" +begin + +(* Type Schemes *) +atom_decl name + +nominal_datatype t = + Var "name" +| Fun "t" "t" +and tyS = + All xs::"name fset" ty::"t" bind xs in ty + +lemma size_eqvt_raw: + "size (pi \ t :: t_raw) = size t" + "size (pi \ ts :: tyS_raw) = size ts" + apply (induct rule: t_raw_tyS_raw.inducts) + apply simp_all + done + +instantiation t and tyS :: size begin + +quotient_definition + "size_t :: t \ nat" +is + "size :: t_raw \ nat" + +quotient_definition + "size_tyS :: tyS \ nat" +is + "size :: tyS_raw \ nat" + +lemma size_rsp: + "alpha_t_raw x y \ size x = size y" + "alpha_tyS_raw a b \ size a = size b" + apply (induct rule: alpha_t_raw_alpha_tyS_raw.inducts) + apply (simp_all only: t_raw_tyS_raw.size) + apply (simp_all only: alpha_gen) + apply clarify + apply (simp_all only: size_eqvt_raw) + done + +lemma [quot_respect]: + "(alpha_t_raw ===> op =) size size" + "(alpha_tyS_raw ===> op =) size size" + by (simp_all add: size_rsp) + +lemma [quot_preserve]: + "(rep_t ---> id) size = size" + "(rep_tyS ---> id) size = size" + by (simp_all add: size_t_def size_tyS_def) + +instance + by default + +end + +thm t_raw_tyS_raw.size(4)[quot_lifted] +thm t_raw_tyS_raw.size(5)[quot_lifted] +thm t_raw_tyS_raw.size(6)[quot_lifted] + + +thm t_tyS.fv +thm t_tyS.eq_iff +thm t_tyS.bn +thm t_tyS.perm +thm t_tyS.inducts +thm t_tyS.distinct +ML {* Sign.of_sort @{theory} (@{typ t}, @{sort fs}) *} + +lemmas t_tyS_supp = t_tyS.fv[simplified t_tyS.supp] + +lemma induct: + assumes a1: "\name b. P b (Var name)" + and a2: "\t1 t2 b. \\c. P c t1; \c. P c t2\ \ P b (Fun t1 t2)" + and a3: "\fset t b. \\c. P c t; fset_to_set (fmap atom fset) \* b\ \ P' b (All fset t)" + shows "P (a :: 'a :: pt) t \ P' (d :: 'b :: {fs}) ts " +proof - + have " (\p a. P a (p \ t)) \ (\p d. P' d (p \ ts))" + apply (rule t_tyS.induct) + apply (simp add: a1) + apply (simp) + apply (rule allI)+ + apply (rule a2) + apply simp + apply simp + apply (rule allI) + apply (rule allI) + apply(subgoal_tac "\pa. ((pa \ (fset_to_set (fmap atom (p \ fset)))) \* d \ supp (p \ TySch.All fset t) \* pa)") + apply clarify + apply(rule_tac t="p \ TySch.All fset t" and + s="pa \ (p \ TySch.All fset t)" in subst) + apply (rule supp_perm_eq) + apply assumption + apply (simp only: t_tyS.perm) + apply (rule a3) + apply(erule_tac x="(pa + p)" in allE) + apply simp + apply (simp add: eqvts eqvts_raw) + apply (rule at_set_avoiding2) + apply (simp add: fin_fset_to_set) + apply (simp add: finite_supp) + apply (simp add: eqvts finite_supp) + apply (subst atom_eqvt_raw[symmetric]) + apply (subst fmap_eqvt[symmetric]) + apply (subst fset_to_set_eqvt[symmetric]) + apply (simp only: fresh_star_permute_iff) + apply (simp add: fresh_star_def) + apply clarify + apply (simp add: fresh_def) + apply (simp add: t_tyS_supp) + done + then have "P a (0 \ t) \ P' d (0 \ ts)" by blast + then show ?thesis by simp +qed + +lemma + shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|b, a|} (Fun (Var a) (Var b))" + apply(simp add: t_tyS.eq_iff) + apply(rule_tac x="0::perm" in exI) + apply(simp add: alpha_gen) + apply(auto) + apply(simp add: fresh_star_def fresh_zero_perm) + done + +lemma + shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var b) (Var a))" + apply(simp add: t_tyS.eq_iff) + apply(rule_tac x="(atom a \ atom b)" in exI) + apply(simp add: alpha_gen fresh_star_def eqvts) + apply auto + done + +lemma + shows "All {|a, b, c|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var a) (Var b))" + apply(simp add: t_tyS.eq_iff) + apply(rule_tac x="0::perm" in exI) + apply(simp add: alpha_gen fresh_star_def eqvts t_tyS.eq_iff) +oops + +lemma + assumes a: "a \ b" + shows "\(All {|a, b|} (Fun (Var a) (Var b)) = All {|c|} (Fun (Var c) (Var c)))" + using a + apply(simp add: t_tyS.eq_iff) + apply(clarify) + apply(simp add: alpha_gen fresh_star_def eqvts t_tyS.eq_iff) + apply auto + done + +(* PROBLEM: +Type schemes with separate datatypes + +nominal_datatype T = + TVar "name" +| TFun "T" "T" +nominal_datatype TyS = + TAll xs::"name list" ty::"T" bind xs in ty + +*** exception Datatype raised +*** (line 218 of "/usr/local/src/Isabelle_16-Mar-2010/src/HOL/Tools/Datatype/datatype_aux.ML") +*** At command "nominal_datatype". +*) + + +end