author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Fri, 04 Dec 2009 09:33:32 +0100 | |
changeset 515 | b00a9b58264d |
parent 514 | 6b3be083229c |
child 534 | 051bd9e90e92 |
permissions | -rw-r--r-- |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory LamEx |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
imports Nominal QuotMain |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
atom_decl name |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
|
243 | 7 |
thm abs_fresh(1) |
8 |
||
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
nominal_datatype rlam = |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
rVar "name" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
| rApp "rlam" "rlam" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
| rLam "name" "rlam" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
|
259 | 14 |
print_theorems |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
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
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
26 |
inductive |
246
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
27 |
alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100) |
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
28 |
where |
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
29 |
a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)" |
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
30 |
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2" |
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
31 |
| a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s" |
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
32 |
|
259 | 33 |
print_theorems |
34 |
||
271
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
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
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
51 |
sorry |
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
52 |
|
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
quotient lam = rlam / alpha |
286 | 54 |
apply(rule alpha_EQUIV) |
55 |
done |
|
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
|
203
7384115df9fd
added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents:
201
diff
changeset
|
57 |
print_quotients |
7384115df9fd
added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents:
201
diff
changeset
|
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>
parents:
267
diff
changeset
|
59 |
quotient_def |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
61 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
62 |
"Var \<equiv> rVar" |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
267
diff
changeset
|
64 |
quotient_def |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
66 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
67 |
"App \<equiv> rApp" |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
267
diff
changeset
|
69 |
quotient_def |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
71 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
72 |
"Lam \<equiv> rLam" |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
|
218
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
74 |
thm Var_def |
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
75 |
thm App_def |
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
76 |
thm Lam_def |
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
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>
parents:
267
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>
parents:
225
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>
parents:
225
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>
parents:
225
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>
parents:
267
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>
parents:
225
diff
changeset
|
89 |
begin |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
267
diff
changeset
|
91 |
quotient_def |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
93 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
95 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
96 |
end |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
97 |
|
238
e9cc3a3aa5d1
Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
237
diff
changeset
|
98 |
(*quotient_def (for lam) |
e9cc3a3aa5d1
Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
237
diff
changeset
|
99 |
abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam" |
e9cc3a3aa5d1
Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
237
diff
changeset
|
100 |
where |
e9cc3a3aa5d1
Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
237
diff
changeset
|
101 |
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*) |
e9cc3a3aa5d1
Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
237
diff
changeset
|
102 |
|
e9cc3a3aa5d1
Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
237
diff
changeset
|
103 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
105 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
106 |
(* lemmas that need to lift *) |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
107 |
lemma pi_var_com: |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
108 |
fixes pi::"'x prm" |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
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>
parents:
225
diff
changeset
|
110 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
111 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
112 |
lemma pi_app_com: |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
113 |
fixes pi::"'x prm" |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
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>
parents:
225
diff
changeset
|
115 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
116 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
117 |
lemma pi_lam_com: |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
118 |
fixes pi::"'x prm" |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
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>
parents:
225
diff
changeset
|
120 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
445
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>
parents:
445
diff
changeset
|
123 |
|
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
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>
parents:
445
diff
changeset
|
125 |
assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s" |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
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>
parents:
445
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>
parents:
445
diff
changeset
|
128 |
unfolding fresh_def supp_def |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
129 |
sorry |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
130 |
|
508
fac6069d8e80
Made your version work again with LIST_REL_EQ.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
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>
parents:
225
diff
changeset
|
133 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
135 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
445
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>
parents:
445
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>
parents:
225
diff
changeset
|
139 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
141 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
142 |
|
508
fac6069d8e80
Made your version work again with LIST_REL_EQ.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
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>
parents:
445
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>
parents:
445
diff
changeset
|
145 |
by (auto intro:a1) |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
146 |
|
508
fac6069d8e80
Made your version work again with LIST_REL_EQ.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
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>
parents:
445
diff
changeset
|
148 |
by (auto intro:a2) |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
149 |
|
508
fac6069d8e80
Made your version work again with LIST_REL_EQ.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
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>
parents:
225
diff
changeset
|
151 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
152 |
apply(rule a3) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
154 |
apply(rule sym) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
155 |
apply(rule trans) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
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>
parents:
225
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>
parents:
225
diff
changeset
|
159 |
apply(assumption) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
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>
parents:
225
diff
changeset
|
161 |
done |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
162 |
|
508
fac6069d8e80
Made your version work again with LIST_REL_EQ.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
163 |
lemma rfv_rsp[quotient_rsp]: "(alpha ===> op =) rfv rfv" |
247 | 164 |
sorry |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
165 |
|
285
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
166 |
lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)" |
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
167 |
apply (auto) |
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
168 |
apply (erule alpha.cases) |
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
169 |
apply (simp_all add: rlam.inject alpha_refl) |
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
170 |
done |
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
171 |
|
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
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
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
174 |
|
370
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
175 |
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *} |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
176 |
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *} |
508
fac6069d8e80
Made your version work again with LIST_REL_EQ.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
177 |
ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] *} |
237 | 178 |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
179 |
lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)" |
370
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
180 |
apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *}) |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
181 |
done |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
182 |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
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
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
184 |
apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *}) |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
185 |
done |
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
186 |
|
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
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
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
188 |
apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *}) |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
189 |
done |
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
190 |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
191 |
lemma fv_var: "fv (Var (a\<Colon>name)) = {a}" |
370
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
192 |
apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *}) |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
193 |
done |
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
194 |
|
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
195 |
lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa" |
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
196 |
apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *}) |
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
197 |
done |
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
198 |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
199 |
lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}" |
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
200 |
apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *}) |
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
201 |
done |
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
202 |
|
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
203 |
lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b" |
370
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
204 |
apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *}) |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
205 |
done |
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
206 |
|
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
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
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
208 |
apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *}) |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
209 |
done |
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
210 |
|
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
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" |
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
212 |
apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *}) |
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
213 |
done |
286 | 214 |
|
378
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
215 |
lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P; |
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
216 |
\<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P; |
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
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
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
218 |
\<Longrightarrow> P" |
378
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
219 |
apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *}) |
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
220 |
done |
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
221 |
|
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
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
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
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); |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
224 |
\<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam. |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
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> |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
226 |
\<Longrightarrow> qxb qx qxa" |
378
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
227 |
apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *}) |
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
228 |
done |
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
251
diff
changeset
|
229 |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
230 |
lemma var_inject: "(Var a = Var b) = (a = b)" |
370
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
231 |
apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *}) |
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
232 |
done |
285
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
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>
parents:
445
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>
parents:
445
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>
parents:
445
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>
parents:
445
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>
parents:
445
diff
changeset
|
238 |
|
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
239 |
lemma var_supp: |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
240 |
shows "supp (Var a) = ((supp a)::name set)" |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
241 |
apply(simp add: supp_def) |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
242 |
apply(simp add: pi_var) |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
243 |
apply(simp add: var_inject) |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
244 |
done |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
245 |
|
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
246 |
lemma var_fresh: |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
247 |
fixes a::"name" |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
248 |
shows "(a\<sharp>(Var b)) = (a\<sharp>b)" |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
249 |
apply(simp add: fresh_def) |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
250 |
apply(simp add: var_supp) |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
251 |
done |
247 | 252 |
|
253 |
||
254 |
||
255 |
||
256 |
||
257 |
||
258 |
||
271
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
259 |
|
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
260 |
|
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
261 |
|
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
262 |
|
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
263 |
|
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
264 |
(* Construction Site code *) |
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
265 |
|
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
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
57bde65f6eb2
Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
378
diff
changeset
|
281 |
lemma "option_map id = id" |
57bde65f6eb2
Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
378
diff
changeset
|
282 |
sorry |
57bde65f6eb2
Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
378
diff
changeset
|
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
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
300 |
lemma real_alpha_pre: |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
301 |
assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s" |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
302 |
shows "rLam a t \<approx> rLam b s" |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
303 |
sorry |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
304 |
|
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
305 |
lemma perm_prs : "(id ---> option_map ABS_lam) ([b].REP_lam s) = [b].s" |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
306 |
sorry |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
307 |
|
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
308 |
lemma perm_prs2: "(id ---> option_map ABS_lam) ((id ---> option_map REP_lam) ([b].(s::lam))) = [b].s" |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
309 |
sorry |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
310 |
|
513
eed5d55ea9a6
Testing the new tactic everywhere
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
311 |
lemma real_alpha_lift: |
487
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
312 |
"\<lbrakk>t = [(a, b)] \<bullet> s; a \<sharp> [b].s\<rbrakk> \<Longrightarrow> Lam a t = Lam b s" |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
313 |
apply (tactic {* procedure_tac @{context} @{thm real_alpha_pre} 1 *}) |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
314 |
prefer 2 |
510
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
315 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
316 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
317 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
318 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
319 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
320 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
321 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
322 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
323 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
324 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
325 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
326 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
327 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
328 |
apply (tactic {* all_inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
329 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
330 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
331 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
332 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
333 |
apply (tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
487
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
334 |
apply (simp only: perm_prs) |
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
335 |
prefer 2 |
510
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
336 |
apply (tactic {* all_inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) |
487
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
337 |
prefer 3 |
510
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
338 |
apply (tactic {* clean_tac @{context} 1 *}) |
487
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
339 |
apply (simp only: perm_prs) |
510
8dbc521ee97f
Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
508
diff
changeset
|
340 |
(*apply (tactic {* regularize_tac @{context} 1 *})*) |
513
eed5d55ea9a6
Testing the new tactic everywhere
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
501
diff
changeset
|
341 |
sorry |
487
f5db9ede89b0
Experiments with OPTION_map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
458
diff
changeset
|
342 |