author | Christian Urban <urbanc@in.tum.de> |
Thu, 13 May 2010 15:12:34 +0100 | |
changeset 2124 | a17b25cb94a6 |
parent 2120 | 2786ff1df475 |
child 2179 | 7687f97eca53 |
permissions | -rw-r--r-- |
1795 | 1 |
theory TypeSchemes |
2040
94e24da9ae75
Move TypeSchemes to NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
2 |
imports "../NewParser" |
1795 | 3 |
begin |
4 |
||
5 |
section {*** Type Schemes ***} |
|
6 |
||
7 |
atom_decl name |
|
8 |
||
9 |
nominal_datatype ty = |
|
10 |
Var "name" |
|
11 |
| Fun "ty" "ty" |
|
12 |
and tys = |
|
2040
94e24da9ae75
Move TypeSchemes to NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
13 |
All xs::"name fset" ty::"ty" bind_res xs in ty |
1795 | 14 |
|
15 |
lemmas ty_tys_supp = ty_tys.fv[simplified ty_tys.supp] |
|
16 |
||
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2040
diff
changeset
|
17 |
|
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2040
diff
changeset
|
18 |
|
1795 | 19 |
(* below we define manually the function for size *) |
20 |
||
21 |
lemma size_eqvt_raw: |
|
22 |
"size (pi \<bullet> t :: ty_raw) = size t" |
|
23 |
"size (pi \<bullet> ts :: tys_raw) = size ts" |
|
24 |
apply (induct rule: ty_raw_tys_raw.inducts) |
|
25 |
apply simp_all |
|
26 |
done |
|
27 |
||
28 |
instantiation ty and tys :: size |
|
29 |
begin |
|
30 |
||
31 |
quotient_definition |
|
32 |
"size_ty :: ty \<Rightarrow> nat" |
|
33 |
is |
|
34 |
"size :: ty_raw \<Rightarrow> nat" |
|
35 |
||
36 |
quotient_definition |
|
37 |
"size_tys :: tys \<Rightarrow> nat" |
|
38 |
is |
|
39 |
"size :: tys_raw \<Rightarrow> nat" |
|
40 |
||
41 |
lemma size_rsp: |
|
42 |
"alpha_ty_raw x y \<Longrightarrow> size x = size y" |
|
43 |
"alpha_tys_raw a b \<Longrightarrow> size a = size b" |
|
44 |
apply (induct rule: alpha_ty_raw_alpha_tys_raw.inducts) |
|
45 |
apply (simp_all only: ty_raw_tys_raw.size) |
|
46 |
apply (simp_all only: alphas) |
|
47 |
apply clarify |
|
48 |
apply (simp_all only: size_eqvt_raw) |
|
49 |
done |
|
50 |
||
51 |
lemma [quot_respect]: |
|
52 |
"(alpha_ty_raw ===> op =) size size" |
|
53 |
"(alpha_tys_raw ===> op =) size size" |
|
54 |
by (simp_all add: size_rsp) |
|
55 |
||
56 |
lemma [quot_preserve]: |
|
57 |
"(rep_ty ---> id) size = size" |
|
58 |
"(rep_tys ---> id) size = size" |
|
59 |
by (simp_all add: size_ty_def size_tys_def) |
|
60 |
||
61 |
instance |
|
62 |
by default |
|
63 |
||
64 |
end |
|
65 |
||
66 |
thm ty_raw_tys_raw.size(4)[quot_lifted] |
|
67 |
thm ty_raw_tys_raw.size(5)[quot_lifted] |
|
68 |
thm ty_raw_tys_raw.size(6)[quot_lifted] |
|
69 |
||
70 |
||
71 |
thm ty_tys.fv |
|
72 |
thm ty_tys.eq_iff |
|
73 |
thm ty_tys.bn |
|
74 |
thm ty_tys.perm |
|
75 |
thm ty_tys.inducts |
|
76 |
thm ty_tys.distinct |
|
77 |
||
78 |
ML {* Sign.of_sort @{theory} (@{typ ty}, @{sort fs}) *} |
|
79 |
||
80 |
lemma strong_induct: |
|
81 |
assumes a1: "\<And>name b. P b (Var name)" |
|
82 |
and a2: "\<And>t1 t2 b. \<lbrakk>\<And>c. P c t1; \<And>c. P c t2\<rbrakk> \<Longrightarrow> P b (Fun t1 t2)" |
|
83 |
and a3: "\<And>fset t b. \<lbrakk>\<And>c. P c t; fset_to_set (fmap atom fset) \<sharp>* b\<rbrakk> \<Longrightarrow> P' b (All fset t)" |
|
84 |
shows "P (a :: 'a :: pt) t \<and> P' (d :: 'b :: {fs}) ts " |
|
85 |
proof - |
|
86 |
have " (\<forall>p a. P a (p \<bullet> t)) \<and> (\<forall>p d. P' d (p \<bullet> ts))" |
|
87 |
apply (rule ty_tys.induct) |
|
88 |
apply (simp add: a1) |
|
89 |
apply (simp) |
|
90 |
apply (rule allI)+ |
|
91 |
apply (rule a2) |
|
92 |
apply simp |
|
93 |
apply simp |
|
94 |
apply (rule allI) |
|
95 |
apply (rule allI) |
|
96 |
apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (fset_to_set (fmap atom (p \<bullet> fset)))) \<sharp>* d \<and> supp (p \<bullet> All fset ty) \<sharp>* pa)") |
|
97 |
apply clarify |
|
98 |
apply(rule_tac t="p \<bullet> All fset ty" and |
|
99 |
s="pa \<bullet> (p \<bullet> All fset ty)" in subst) |
|
100 |
apply (rule supp_perm_eq) |
|
101 |
apply assumption |
|
102 |
apply (simp only: ty_tys.perm) |
|
103 |
apply (rule a3) |
|
104 |
apply(erule_tac x="(pa + p)" in allE) |
|
105 |
apply simp |
|
106 |
apply (simp add: eqvts eqvts_raw) |
|
107 |
apply (rule at_set_avoiding2) |
|
108 |
apply (simp add: fin_fset_to_set) |
|
109 |
apply (simp add: finite_supp) |
|
110 |
apply (simp add: eqvts finite_supp) |
|
1933
9eab1dfc14d2
moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents:
1795
diff
changeset
|
111 |
apply (rule_tac p=" -p" in permute_boolE) |
9eab1dfc14d2
moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents:
1795
diff
changeset
|
112 |
apply(simp add: eqvts) |
9eab1dfc14d2
moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents:
1795
diff
changeset
|
113 |
apply(simp add: permute_fun_def atom_eqvt) |
1795 | 114 |
apply (simp add: fresh_star_def) |
115 |
apply clarify |
|
116 |
apply (simp add: fresh_def) |
|
117 |
apply (simp add: ty_tys_supp) |
|
118 |
done |
|
119 |
then have "P a (0 \<bullet> t) \<and> P' d (0 \<bullet> ts)" by blast |
|
120 |
then show ?thesis by simp |
|
121 |
qed |
|
122 |
||
123 |
lemma |
|
124 |
shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|b, a|} (Fun (Var a) (Var b))" |
|
125 |
apply(simp add: ty_tys.eq_iff) |
|
126 |
apply(rule_tac x="0::perm" in exI) |
|
127 |
apply(simp add: alphas) |
|
2040
94e24da9ae75
Move TypeSchemes to NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
128 |
apply(simp add: fresh_star_def fresh_zero_perm supp_at_base) |
1795 | 129 |
done |
130 |
||
131 |
lemma |
|
132 |
shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var b) (Var a))" |
|
133 |
apply(simp add: ty_tys.eq_iff) |
|
134 |
apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI) |
|
2040
94e24da9ae75
Move TypeSchemes to NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
135 |
apply(simp add: alphas fresh_star_def eqvts supp_at_base) |
1795 | 136 |
done |
137 |
||
138 |
lemma |
|
139 |
shows "All {|a, b, c|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var a) (Var b))" |
|
140 |
apply(simp add: ty_tys.eq_iff) |
|
141 |
apply(rule_tac x="0::perm" in exI) |
|
2040
94e24da9ae75
Move TypeSchemes to NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
142 |
apply(simp add: alphas fresh_star_def eqvts ty_tys.eq_iff supp_at_base) |
1795 | 143 |
done |
144 |
||
145 |
lemma |
|
146 |
assumes a: "a \<noteq> b" |
|
147 |
shows "\<not>(All {|a, b|} (Fun (Var a) (Var b)) = All {|c|} (Fun (Var c) (Var c)))" |
|
148 |
using a |
|
149 |
apply(simp add: ty_tys.eq_iff) |
|
150 |
apply(clarify) |
|
2040
94e24da9ae75
Move TypeSchemes to NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
151 |
apply(simp add: alphas fresh_star_def eqvts ty_tys.eq_iff supp_at_base) |
1795 | 152 |
apply auto |
153 |
done |
|
154 |
||
155 |
(* PROBLEM: |
|
156 |
Type schemes with separate datatypes |
|
157 |
||
158 |
nominal_datatype T = |
|
159 |
TVar "name" |
|
160 |
| TFun "T" "T" |
|
161 |
nominal_datatype TyS = |
|
162 |
TAll xs::"name list" ty::"T" bind xs in ty |
|
163 |
||
164 |
*** exception Datatype raised |
|
165 |
*** (line 218 of "/usr/local/src/Isabelle_16-Mar-2010/src/HOL/Tools/Datatype/datatype_aux.ML") |
|
166 |
*** At command "nominal_datatype". |
|
167 |
*) |
|
168 |
||
169 |
||
170 |
end |