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
451
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 122
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 123
201
+ − 124
lemma real_alpha:
451
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 125
assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
201
+ − 126
shows "Lam a t = Lam b s"
451
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 127
using a
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 128
unfolding fresh_def supp_def
217
+ − 129
sorry
+ − 130
508
+ − 131
lemma perm_rsp[quotient_rsp]:
286
+ − 132
"(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
+ − 133
apply(auto)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 134
(* 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
+ − 135
sorry
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 136
451
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 137
lemma fresh_rsp:
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 138
"(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
+ − 139
apply(auto)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 140
(* 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
+ − 141
sorry
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 142
508
+ − 143
lemma rVar_rsp[quotient_rsp]:
451
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 144
"(op = ===> alpha) rVar rVar"
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 145
by (auto intro:a1)
234
+ − 146
508
+ − 147
lemma rApp_rsp[quotient_rsp]: "(alpha ===> alpha ===> alpha) rApp rApp"
451
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 148
by (auto intro:a2)
234
+ − 149
508
+ − 150
lemma rLam_rsp[quotient_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
+ − 151
apply(auto)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 152
apply(rule a3)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 153
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
+ − 154
apply(rule sym)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 155
apply(rule trans)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
apply(rule pt_name3)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 157
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
+ − 158
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
+ − 159
apply(assumption)
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 160
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
+ − 161
done
217
+ − 162
508
+ − 163
lemma rfv_rsp[quotient_rsp]: "(alpha ===> op =) rfv rfv"
247
+ − 164
sorry
217
+ − 165
285
+ − 166
lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)"
+ − 167
apply (auto)
+ − 168
apply (erule alpha.cases)
+ − 169
apply (simp_all add: rlam.inject alpha_refl)
+ − 170
done
+ − 171
217
+ − 172
ML {* val qty = @{typ "lam"} *}
458
+ − 173
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *}
271
+ − 174
370
+ − 175
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
+ − 176
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
508
+ − 177
ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] *}
237
+ − 178
376
+ − 179
lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
370
+ − 180
apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *})
+ − 181
done
+ − 182
376
+ − 183
lemma pi_app: "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
370
+ − 184
apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *})
+ − 185
done
376
+ − 186
+ − 187
lemma pi_lam: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
370
+ − 188
apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *})
+ − 189
done
249
+ − 190
376
+ − 191
lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
370
+ − 192
apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *})
+ − 193
done
376
+ − 194
+ − 195
lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa"
+ − 196
apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *})
+ − 197
done
249
+ − 198
376
+ − 199
lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
+ − 200
apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *})
+ − 201
done
+ − 202
+ − 203
lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
370
+ − 204
apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *})
+ − 205
done
376
+ − 206
+ − 207
lemma a2: "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
370
+ − 208
apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *})
+ − 209
done
376
+ − 210
+ − 211
lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa"
+ − 212
apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *})
+ − 213
done
286
+ − 214
378
+ − 215
lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
+ − 216
\<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
+ − 217
\<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
370
+ − 218
\<Longrightarrow> P"
378
+ − 219
apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *})
+ − 220
done
+ − 221
+ − 222
lemma alpha_induct: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b);
370
+ − 223
\<And>(x\<Colon>lam) (xa\<Colon>lam) (xb\<Colon>lam) xc\<Colon>lam. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
+ − 224
\<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam.
+ − 225
\<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
+ − 226
\<Longrightarrow> qxb qx qxa"
378
+ − 227
apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *})
+ − 228
done
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 229
376
+ − 230
lemma var_inject: "(Var a = Var b) = (a = b)"
370
+ − 231
apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *})
+ − 232
done
285
+ − 233
451
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 234
lemma lam_induct:" \<lbrakk>\<And>name. P (Var name); \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 235
\<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> \<Longrightarrow> P lam"
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 236
apply (tactic {* lift_tac_lam @{context} @{thm rlam.induct} 1 *})
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 237
done
586e3dc4afdb
Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 238
249
+ − 239
lemma var_supp:
+ − 240
shows "supp (Var a) = ((supp a)::name set)"
+ − 241
apply(simp add: supp_def)
+ − 242
apply(simp add: pi_var)
+ − 243
apply(simp add: var_inject)
+ − 244
done
+ − 245
+ − 246
lemma var_fresh:
+ − 247
fixes a::"name"
+ − 248
shows "(a\<sharp>(Var b)) = (a\<sharp>b)"
+ − 249
apply(simp add: fresh_def)
+ − 250
apply(simp add: var_supp)
+ − 251
done
247
+ − 252
+ − 253
+ − 254
+ − 255
+ − 256
+ − 257
+ − 258
271
+ − 259
+ − 260
+ − 261
+ − 262
+ − 263
+ − 264
(* Construction Site code *)
+ − 265
+ − 266
237
+ − 267
fun
+ − 268
option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
+ − 269
where
+ − 270
"option_map f (nSome x) = nSome (f x)"
+ − 271
| "option_map f nNone = nNone"
+ − 272
+ − 273
fun
+ − 274
option_rel
+ − 275
where
+ − 276
"option_rel r (nSome x) (nSome y) = r x y"
+ − 277
| "option_rel r _ _ = False"
+ − 278
+ − 279
declare [[map noption = (option_map, option_rel)]]
+ − 280
379
+ − 281
lemma "option_map id = id"
+ − 282
sorry
+ − 283
237
+ − 284
lemma OPT_QUOTIENT:
+ − 285
assumes q: "QUOTIENT R Abs Rep"
+ − 286
shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)"
+ − 287
apply (unfold QUOTIENT_def)
+ − 288
apply (auto)
+ − 289
using q
+ − 290
apply (unfold QUOTIENT_def)
+ − 291
apply (case_tac "a :: 'b noption")
+ − 292
apply (simp)
+ − 293
apply (simp)
+ − 294
apply (case_tac "a :: 'b noption")
+ − 295
apply (simp only: option_map.simps)
+ − 296
apply (subst option_rel.simps)
+ − 297
(* Simp starts hanging so don't know how to continue *)
+ − 298
sorry
+ − 299
487
+ − 300
lemma real_alpha_pre:
+ − 301
assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
+ − 302
shows "rLam a t \<approx> rLam b s"
+ − 303
sorry
+ − 304
+ − 305
lemma perm_prs : "(id ---> option_map ABS_lam) ([b].REP_lam s) = [b].s"
+ − 306
sorry
+ − 307
+ − 308
lemma perm_prs2: "(id ---> option_map ABS_lam) ((id ---> option_map REP_lam) ([b].(s::lam))) = [b].s"
+ − 309
sorry
+ − 310
+ − 311
lemma real_alpha:
+ − 312
"\<lbrakk>t = [(a, b)] \<bullet> s; a \<sharp> [b].s\<rbrakk> \<Longrightarrow> Lam a t = Lam b s"
+ − 313
apply (tactic {* procedure_tac @{context} @{thm real_alpha_pre} 1 *})
+ − 314
prefer 2
+ − 315
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 316
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 317
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 318
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 319
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 320
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 321
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 322
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 323
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 324
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 325
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 326
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 327
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 328
apply (tactic {* all_inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 329
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 330
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 331
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 332
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 333
apply (tactic {* inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 334
apply (simp only: perm_prs)
+ − 335
prefer 2
+ − 336
apply (tactic {* all_inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
+ − 337
prefer 3
500
+ − 338
apply (tactic {* clean_tac @{context} [quot] 1 *})
487
+ − 339
+ − 340
thm all_prs
+ − 341
thm REP_ABS_RSP
+ − 342
thm ball_reg_eqv_range
+ − 343
+ − 344
+ − 345
thm perm_lam_def
+ − 346
apply (simp only: perm_prs)
+ − 347
+ − 348
apply (tactic {* regularize_tac @{context} [quot] 1 *})
+ − 349
+ − 350
done
+ − 351