201
|
1 |
theory LamEx
|
|
2 |
imports Nominal QuotMain
|
|
3 |
begin
|
|
4 |
|
|
5 |
atom_decl name
|
|
6 |
|
243
|
7 |
thm abs_fresh(1)
|
|
8 |
|
201
|
9 |
nominal_datatype rlam =
|
|
10 |
rVar "name"
|
|
11 |
| rApp "rlam" "rlam"
|
|
12 |
| rLam "name" "rlam"
|
|
13 |
|
259
|
14 |
print_theorems
|
201
|
15 |
|
243
|
16 |
function
|
|
17 |
rfv :: "rlam \<Rightarrow> name set"
|
|
18 |
where
|
|
19 |
rfv_var: "rfv (rVar a) = {a}"
|
|
20 |
| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
|
|
21 |
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
|
|
22 |
sorry
|
|
23 |
|
247
|
24 |
termination rfv sorry
|
243
|
25 |
|
271
|
26 |
inductive
|
246
|
27 |
alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
|
|
28 |
where
|
|
29 |
a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
|
|
30 |
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
|
|
31 |
| a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
|
|
32 |
|
259
|
33 |
print_theorems
|
|
34 |
|
271
|
35 |
lemma alpha_refl:
|
286
|
36 |
fixes t::"rlam"
|
|
37 |
shows "t \<approx> t"
|
|
38 |
apply(induct t rule: rlam.induct)
|
|
39 |
apply(simp add: a1)
|
|
40 |
apply(simp add: a2)
|
|
41 |
apply(rule a3)
|
|
42 |
apply(subst pt_swap_bij'')
|
|
43 |
apply(rule pt_name_inst)
|
|
44 |
apply(rule at_name_inst)
|
|
45 |
apply(simp)
|
|
46 |
apply(simp)
|
|
47 |
done
|
|
48 |
|
|
49 |
lemma alpha_EQUIV:
|
|
50 |
shows "EQUIV alpha"
|
271
|
51 |
sorry
|
|
52 |
|
201
|
53 |
quotient lam = rlam / alpha
|
286
|
54 |
apply(rule alpha_EQUIV)
|
|
55 |
done
|
201
|
56 |
|
203
|
57 |
print_quotients
|
|
58 |
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
59 |
quotient_def
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
60 |
Var :: "name \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
61 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
62 |
"Var \<equiv> rVar"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
63 |
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
64 |
quotient_def
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
65 |
App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
66 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
67 |
"App \<equiv> rApp"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
68 |
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
69 |
quotient_def
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
70 |
Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
71 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
72 |
"Lam \<equiv> rLam"
|
201
|
73 |
|
218
|
74 |
thm Var_def
|
|
75 |
thm App_def
|
|
76 |
thm Lam_def
|
|
77 |
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
78 |
quotient_def
|
243
|
79 |
fv :: "lam \<Rightarrow> name set"
|
|
80 |
where
|
|
81 |
"fv \<equiv> rfv"
|
|
82 |
|
|
83 |
thm fv_def
|
|
84 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
85 |
(* definition of overloaded permutation function *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
86 |
(* for the lifted type lam *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
87 |
overloading
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
88 |
perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam" (unchecked)
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
89 |
begin
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
90 |
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
91 |
quotient_def
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
92 |
perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
93 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
94 |
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
95 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
96 |
end
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
97 |
|
238
|
98 |
(*quotient_def (for lam)
|
|
99 |
abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
|
|
100 |
where
|
|
101 |
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*)
|
|
102 |
|
|
103 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
104 |
thm perm_lam_def
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
105 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
106 |
(* lemmas that need to lift *)
|
234
|
107 |
lemma pi_var_com:
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
108 |
fixes pi::"'x prm"
|
234
|
109 |
shows "(pi\<bullet>rVar a) \<approx> rVar (pi\<bullet>a)"
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
110 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
111 |
|
234
|
112 |
lemma pi_app_com:
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
113 |
fixes pi::"'x prm"
|
234
|
114 |
shows "(pi\<bullet>rApp t1 t2) \<approx> rApp (pi\<bullet>t1) (pi\<bullet>t2)"
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
115 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
116 |
|
234
|
117 |
lemma pi_lam_com:
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
118 |
fixes pi::"'x prm"
|
234
|
119 |
shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)"
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
120 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
121 |
|
243
|
122 |
lemma fv_var:
|
|
123 |
shows "fv (Var a) = {a}"
|
|
124 |
sorry
|
|
125 |
|
|
126 |
lemma fv_app:
|
|
127 |
shows "fv (App t1 t2) = (fv t1) \<union> (fv t2)"
|
|
128 |
sorry
|
|
129 |
|
|
130 |
lemma fv_lam:
|
|
131 |
shows "fv (Lam a t) = (fv t) - {a}"
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
132 |
sorry
|
243
|
133 |
|
201
|
134 |
lemma real_alpha:
|
242
|
135 |
assumes "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
|
201
|
136 |
shows "Lam a t = Lam b s"
|
217
|
137 |
sorry
|
|
138 |
|
286
|
139 |
lemma perm_rsp:
|
|
140 |
"(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
141 |
apply(auto)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
142 |
(* this is propably true if some type conditions are imposed ;o) *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
143 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
144 |
|
286
|
145 |
lemma fresh_rsp:
|
|
146 |
"(op = ===> alpha ===> op =) fresh fresh"
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
147 |
apply(auto)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
148 |
(* this is probably only true if some type conditions are imposed *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
149 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
150 |
|
234
|
151 |
lemma rVar_rsp: "(op = ===> alpha) rVar rVar"
|
|
152 |
apply(auto)
|
|
153 |
apply(rule a1)
|
|
154 |
apply(simp)
|
|
155 |
done
|
|
156 |
|
|
157 |
lemma rApp_rsp: "(alpha ===> alpha ===> alpha) rApp rApp"
|
|
158 |
apply(auto)
|
|
159 |
apply(rule a2)
|
|
160 |
apply (assumption)
|
|
161 |
apply (assumption)
|
|
162 |
done
|
|
163 |
|
|
164 |
lemma rLam_rsp: "(op = ===> alpha ===> alpha) rLam rLam"
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
165 |
apply(auto)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
166 |
apply(rule a3)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
167 |
apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
168 |
apply(rule sym)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
169 |
apply(rule trans)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
170 |
apply(rule pt_name3)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
171 |
apply(rule at_ds1[OF at_name_inst])
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
172 |
apply(simp add: pt_name1)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
173 |
apply(assumption)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
174 |
apply(simp add: abs_fresh)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
175 |
done
|
217
|
176 |
|
247
|
177 |
lemma rfv_rsp: "(alpha ===> op =) rfv rfv"
|
|
178 |
sorry
|
217
|
179 |
|
285
|
180 |
lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)"
|
|
181 |
apply (auto)
|
|
182 |
apply (erule alpha.cases)
|
|
183 |
apply (simp_all add: rlam.inject alpha_refl)
|
|
184 |
done
|
|
185 |
|
217
|
186 |
ML {* val qty = @{typ "lam"} *}
|
247
|
187 |
ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
|
271
|
188 |
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} @
|
|
189 |
@{thms ho_all_prs ho_ex_prs} *}
|
|
190 |
|
240
|
191 |
ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
|
237
|
192 |
|
249
|
193 |
ML {* val pi_var = lift_thm_lam @{context} @{thm pi_var_com} *}
|
|
194 |
ML {* val pi_app = lift_thm_lam @{context} @{thm pi_app_com} *}
|
|
195 |
ML {* val pi_lam = lift_thm_lam @{context} @{thm pi_lam_com} *}
|
|
196 |
|
|
197 |
ML {* val fv_var = lift_thm_lam @{context} @{thm rfv_var} *}
|
|
198 |
ML {* val fv_app = lift_thm_lam @{context} @{thm rfv_app} *}
|
|
199 |
ML {* val fv_lam = lift_thm_lam @{context} @{thm rfv_lam} *}
|
|
200 |
|
|
201 |
ML {* val a1 = lift_thm_lam @{context} @{thm a1} *}
|
259
|
202 |
ML {* val a2 = lift_thm_lam @{context} @{thm a2} *}
|
249
|
203 |
ML {* val a3 = lift_thm_lam @{context} @{thm a3} *}
|
|
204 |
|
286
|
205 |
thm a2
|
|
206 |
ML {* val t_a = atomize_thm @{thm a2} *}
|
|
207 |
ML {* val t_r = regularize t_a @{typ rlam} @{term alpha} @{thm alpha_EQUIV} @{thm alpha_refl} @{context} *}
|
|
208 |
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
209 |
ML {* val alpha_cases = lift_thm_lam @{context} @{thm alpha.cases} *}
|
271
|
210 |
ML {* val alpha_induct = lift_thm_lam @{context} @{thm alpha.induct} *}
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
211 |
|
285
|
212 |
ML {* val var_inject = lift_thm_lam @{context} @{thm rvar_inject} *}
|
|
213 |
|
249
|
214 |
local_setup {*
|
|
215 |
Quotient.note (@{binding "pi_var"}, pi_var) #> snd #>
|
|
216 |
Quotient.note (@{binding "pi_app"}, pi_app) #> snd #>
|
|
217 |
Quotient.note (@{binding "pi_lam"}, pi_lam) #> snd #>
|
|
218 |
Quotient.note (@{binding "a1"}, a1) #> snd #>
|
|
219 |
Quotient.note (@{binding "a2"}, a2) #> snd #>
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
220 |
Quotient.note (@{binding "a3"}, a3) #> snd #>
|
271
|
221 |
Quotient.note (@{binding "alpha_cases"}, alpha_cases) #> snd #>
|
285
|
222 |
Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd #>
|
|
223 |
Quotient.note (@{binding "var_inject"}, var_inject) #> snd
|
249
|
224 |
*}
|
237
|
225 |
|
249
|
226 |
thm alpha.cases
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
227 |
thm alpha_cases
|
271
|
228 |
thm alpha.induct
|
|
229 |
thm alpha_induct
|
249
|
230 |
|
|
231 |
lemma var_supp:
|
|
232 |
shows "supp (Var a) = ((supp a)::name set)"
|
|
233 |
apply(simp add: supp_def)
|
|
234 |
apply(simp add: pi_var)
|
|
235 |
apply(simp add: var_inject)
|
|
236 |
done
|
|
237 |
|
|
238 |
lemma var_fresh:
|
|
239 |
fixes a::"name"
|
|
240 |
shows "(a\<sharp>(Var b)) = (a\<sharp>b)"
|
|
241 |
apply(simp add: fresh_def)
|
|
242 |
apply(simp add: var_supp)
|
|
243 |
done
|
247
|
244 |
|
|
245 |
|
|
246 |
|
|
247 |
|
|
248 |
|
|
249 |
|
|
250 |
|
271
|
251 |
|
|
252 |
|
|
253 |
|
|
254 |
|
|
255 |
|
|
256 |
(* Construction Site code *)
|
|
257 |
|
|
258 |
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
|
|
259 |
ML {* val consts = lookup_quot_consts defs *}
|
|
260 |
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
|
|
261 |
|
292
|
262 |
ML {* val t_a = atomize_thm @{thm alpha.cases} *}
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
263 |
(* prove {* build_regularize_goal t_a rty rel @{context} *}
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
264 |
ML_prf {* fun tac ctxt =
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
265 |
(FIRST' [
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
266 |
rtac rel_refl,
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
267 |
atac,
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
268 |
rtac @{thm universal_twice},
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
269 |
(rtac @{thm impI} THEN' atac),
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
270 |
rtac @{thm implication_twice},
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
271 |
EqSubst.eqsubst_tac ctxt [0]
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
272 |
[(@{thm equiv_res_forall} OF [rel_eqv]),
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
273 |
(@{thm equiv_res_exists} OF [rel_eqv])],
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
274 |
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
275 |
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
276 |
]);
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
277 |
*}
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
278 |
apply (tactic {* tac @{context} 1 *}) *)
|
251
|
279 |
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
|
259
|
280 |
|
292
|
281 |
ML {*
|
257
|
282 |
val rt = build_repabs_term @{context} t_r consts rty qty
|
|
283 |
val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
|
|
284 |
*}
|
|
285 |
prove rg
|
|
286 |
apply(atomize(full))
|
292
|
287 |
(*ML_prf {*
|
257
|
288 |
fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms =
|
|
289 |
(FIRST' [
|
|
290 |
rtac trans_thm,
|
|
291 |
LAMBDA_RES_TAC ctxt,
|
|
292 |
res_forall_rsp_tac ctxt,
|
|
293 |
res_exists_rsp_tac ctxt,
|
|
294 |
(
|
|
295 |
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps rsp_thms))
|
|
296 |
THEN_ALL_NEW (fn _ => no_tac)
|
|
297 |
),
|
|
298 |
(instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
|
|
299 |
rtac refl,
|
|
300 |
(APPLY_RSP_TAC rty ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])),
|
|
301 |
Cong_Tac.cong_tac @{thm cong},
|
|
302 |
rtac @{thm ext},
|
|
303 |
rtac reflex_thm,
|
|
304 |
atac,
|
|
305 |
(
|
|
306 |
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
|
|
307 |
THEN_ALL_NEW (fn _ => no_tac)
|
|
308 |
),
|
|
309 |
WEAK_LAMBDA_RES_TAC ctxt
|
|
310 |
]);
|
292
|
311 |
*}*)
|
|
312 |
ML_prf {*
|
257
|
313 |
fun r_mk_comb_tac_lam lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms
|
|
314 |
*}
|
292
|
315 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
316 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
317 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
318 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
319 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
320 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
321 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
322 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
323 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
324 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
325 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
326 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
327 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
328 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
329 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
330 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
331 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
332 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
333 |
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
|
|
334 |
|
|
335 |
|
|
336 |
|
|
337 |
|
|
338 |
|
|
339 |
|
257
|
340 |
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_lam @{context}) 1 *})
|
292
|
341 |
|
257
|
342 |
|
251
|
343 |
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
344 |
ML {* val abs = findabs rty (prop_of (atomize_thm @{thm alpha.induct})) *}
|
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
345 |
ML {* val aps = findaps rty (prop_of (atomize_thm @{thm alpha.induct})) *}
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
346 |
ML {* val (alls, exs) = findallex rty qty (prop_of (atomize_thm @{thm alpha.induct})) *}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
347 |
ML {* val allthms = map (make_allex_prs_thm @{context} quot @{thm FORALL_PRS} ) alls *}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
348 |
ML {* val exthms = map (make_allex_prs_thm @{context} quot @{thm EXISTS_PRS} ) exs *}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
349 |
ML {* val t_a = MetaSimplifier.rewrite_rule allthms t_t *}
|
265
|
350 |
ML {* val simp_app_prs_thms = map (make_simp_prs_thm @{context} quot @{thm APP_PRS}) aps *}
|
|
351 |
ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
352 |
ML {* val t_l = repeat_eqsubst_thm @{context} (simp_lam_prs_thms) t_a *}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
353 |
ML {* val t_l1 = repeat_eqsubst_thm @{context} simp_app_prs_thms t_l *}
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
354 |
ML {* val defs_sym = add_lower_defs @{context} defs; *}
|
259
|
355 |
ML {* val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym *}
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
356 |
ML {* val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_l1 *}
|
259
|
357 |
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_d0 *}
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
358 |
ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
359 |
ML {* val t_r1 = repeat_eqsubst_thm @{context} @{thms fun_map.simps} t_r *}
|
259
|
360 |
ML {* val t_r2 = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_lam.thm10} t_r1 *}
|
|
361 |
ML {* val t_r3 = MetaSimplifier.rewrite_rule @{thms eq_reflection[OF id_apply]} t_r2 *}
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
362 |
ML {* val alpha_induct = ObjectLogic.rulify t_r3 *}
|
247
|
363 |
|
271
|
364 |
(*local_setup {*
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
365 |
Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd
|
271
|
366 |
*}*)
|
247
|
367 |
|
259
|
368 |
thm alpha_induct
|
|
369 |
thm alpha.induct
|
|
370 |
|
271
|
371 |
|
|
372 |
|
259
|
373 |
|
|
374 |
|
|
375 |
|
|
376 |
|
271
|
377 |
|
237
|
378 |
fun
|
|
379 |
option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
|
|
380 |
where
|
|
381 |
"option_map f (nSome x) = nSome (f x)"
|
|
382 |
| "option_map f nNone = nNone"
|
|
383 |
|
|
384 |
fun
|
|
385 |
option_rel
|
|
386 |
where
|
|
387 |
"option_rel r (nSome x) (nSome y) = r x y"
|
|
388 |
| "option_rel r _ _ = False"
|
|
389 |
|
|
390 |
declare [[map noption = (option_map, option_rel)]]
|
|
391 |
|
|
392 |
lemma OPT_QUOTIENT:
|
|
393 |
assumes q: "QUOTIENT R Abs Rep"
|
|
394 |
shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)"
|
|
395 |
apply (unfold QUOTIENT_def)
|
|
396 |
apply (auto)
|
|
397 |
using q
|
|
398 |
apply (unfold QUOTIENT_def)
|
|
399 |
apply (case_tac "a :: 'b noption")
|
|
400 |
apply (simp)
|
|
401 |
apply (simp)
|
|
402 |
apply (case_tac "a :: 'b noption")
|
|
403 |
apply (simp only: option_map.simps)
|
|
404 |
apply (subst option_rel.simps)
|
|
405 |
(* Simp starts hanging so don't know how to continue *)
|
|
406 |
sorry
|
|
407 |
|