author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Thu, 14 Jan 2010 17:53:23 +0100 | |
changeset 879 | f2a1ebba9bdc |
parent 878 | c3662f845129 |
child 880 | cd3f1409780a |
permissions | -rw-r--r-- |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory LamEx |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
705
diff
changeset
|
2 |
imports Nominal "../QuotMain" "../QuotList" |
201
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 |
|
804 | 7 |
datatype rlam = |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
rVar "name" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
| rApp "rlam" "rlam" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
| rLam "name" "rlam" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
|
804 | 12 |
fun |
243 | 13 |
rfv :: "rlam \<Rightarrow> name set" |
14 |
where |
|
15 |
rfv_var: "rfv (rVar a) = {a}" |
|
16 |
| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)" |
|
17 |
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}" |
|
804 | 18 |
|
19 |
overloading |
|
876 | 20 |
perm_rlam \<equiv> "perm :: 'x prm \<Rightarrow> rlam \<Rightarrow> rlam" (unchecked) |
804 | 21 |
begin |
243 | 22 |
|
804 | 23 |
fun |
24 |
perm_rlam |
|
25 |
where |
|
26 |
"perm_rlam pi (rVar a) = rVar (pi \<bullet> a)" |
|
27 |
| "perm_rlam pi (rApp t1 t2) = rApp (perm_rlam pi t1) (perm_rlam pi t2)" |
|
28 |
| "perm_rlam pi (rLam a t) = rLam (pi \<bullet> a) (perm_rlam pi t)" |
|
29 |
||
30 |
end |
|
243 | 31 |
|
271
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
32 |
inductive |
246
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
33 |
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
|
34 |
where |
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
35 |
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
|
36 |
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2" |
804 | 37 |
| a3: "\<lbrakk>t \<approx> ([(a,b)] \<bullet> s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s" |
246
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents:
245
diff
changeset
|
38 |
|
804 | 39 |
lemma helper: |
40 |
fixes t::"rlam" |
|
41 |
and a::"name" |
|
42 |
shows "[(a, a)] \<bullet> t = t" |
|
43 |
by (induct t) |
|
44 |
(auto simp add: calc_atm) |
|
259 | 45 |
|
271
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
46 |
lemma alpha_refl: |
286 | 47 |
fixes t::"rlam" |
48 |
shows "t \<approx> t" |
|
49 |
apply(induct t rule: rlam.induct) |
|
50 |
apply(simp add: a1) |
|
51 |
apply(simp add: a2) |
|
52 |
apply(rule a3) |
|
804 | 53 |
apply(simp add: helper) |
286 | 54 |
apply(simp) |
55 |
done |
|
56 |
||
534 | 57 |
lemma alpha_equivp: |
58 |
shows "equivp alpha" |
|
271
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
59 |
sorry |
1b57f99737fe
Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
270
diff
changeset
|
60 |
|
766
df053507edba
renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
61 |
quotient_type lam = rlam / alpha |
804 | 62 |
by (rule alpha_equivp) |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
|
203
7384115df9fd
added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents:
201
diff
changeset
|
64 |
|
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
766
diff
changeset
|
65 |
quotient_definition |
876 | 66 |
"Var :: name \<Rightarrow> lam" |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
663
diff
changeset
|
67 |
as |
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
636
diff
changeset
|
68 |
"rVar" |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
69 |
|
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
766
diff
changeset
|
70 |
quotient_definition |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
663
diff
changeset
|
71 |
"App :: lam \<Rightarrow> lam \<Rightarrow> lam" |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
663
diff
changeset
|
72 |
as |
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
636
diff
changeset
|
73 |
"rApp" |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
74 |
|
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
766
diff
changeset
|
75 |
quotient_definition |
876 | 76 |
"Lam :: name \<Rightarrow> lam \<Rightarrow> lam" |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
663
diff
changeset
|
77 |
as |
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
636
diff
changeset
|
78 |
"rLam" |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
79 |
|
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
766
diff
changeset
|
80 |
quotient_definition |
876 | 81 |
"fv :: lam \<Rightarrow> name set" |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
663
diff
changeset
|
82 |
as |
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
636
diff
changeset
|
83 |
"rfv" |
243 | 84 |
|
804 | 85 |
thm Var_def |
86 |
thm App_def |
|
87 |
thm Lam_def |
|
243 | 88 |
thm fv_def |
89 |
||
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
90 |
(* 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
|
91 |
(* 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
|
92 |
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
|
93 |
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
|
94 |
begin |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
95 |
|
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
766
diff
changeset
|
96 |
quotient_definition |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
663
diff
changeset
|
97 |
"perm_lam :: 'x prm \<Rightarrow> lam \<Rightarrow> lam" |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
663
diff
changeset
|
98 |
as |
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
636
diff
changeset
|
99 |
"perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam" |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
100 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
101 |
end |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
102 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
103 |
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
|
104 |
|
804 | 105 |
(* lemmas that need to be lifted *) |
106 |
lemma pi_var_eqvt1: |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
107 |
fixes pi::"'x prm" |
804 | 108 |
shows "(pi \<bullet> rVar a) \<approx> rVar (pi \<bullet> a)" |
109 |
by (simp add: alpha_refl) |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
110 |
|
804 | 111 |
lemma pi_var_eqvt2: |
112 |
fixes pi::"'x prm" |
|
113 |
shows "(pi \<bullet> rVar a) = rVar (pi \<bullet> a)" |
|
114 |
by (simp) |
|
115 |
||
116 |
lemma pi_app_eqvt1: |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
117 |
fixes pi::"'x prm" |
804 | 118 |
shows "(pi \<bullet> rApp t1 t2) \<approx> rApp (pi \<bullet> t1) (pi \<bullet> t2)" |
119 |
by (simp add: alpha_refl) |
|
120 |
||
121 |
lemma pi_app_eqvt2: |
|
122 |
fixes pi::"'x prm" |
|
123 |
shows "(pi \<bullet> rApp t1 t2) = rApp (pi \<bullet> t1) (pi \<bullet> t2)" |
|
124 |
by (simp) |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
125 |
|
804 | 126 |
lemma pi_lam_eqvt1: |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
127 |
fixes pi::"'x prm" |
804 | 128 |
shows "(pi \<bullet> rLam a t) \<approx> rLam (pi \<bullet> a) (pi \<bullet> t)" |
129 |
by (simp add: alpha_refl) |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
130 |
|
804 | 131 |
lemma pi_lam_eqvt2: |
132 |
fixes pi::"'x prm" |
|
133 |
shows "(pi \<bullet> rLam a t) = rLam (pi \<bullet> a) (pi \<bullet> t)" |
|
134 |
by (simp add: 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
|
135 |
|
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
|
136 |
|
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
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
|
138 |
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
|
139 |
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
|
140 |
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
|
141 |
unfolding fresh_def supp_def |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
142 |
sorry |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
143 |
|
636
520a4084d064
changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents:
610
diff
changeset
|
144 |
lemma perm_rsp[quot_respect]: |
286 | 145 |
"(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
|
146 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
147 |
(* 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
|
148 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
149 |
|
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
|
150 |
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
|
151 |
"(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
|
152 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
153 |
(* 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
|
154 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
155 |
|
636
520a4084d064
changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents:
610
diff
changeset
|
156 |
lemma rVar_rsp[quot_respect]: |
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
|
157 |
"(op = ===> alpha) rVar rVar" |
876 | 158 |
by (auto intro: a1) |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
159 |
|
636
520a4084d064
changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents:
610
diff
changeset
|
160 |
lemma rApp_rsp[quot_respect]: "(alpha ===> alpha ===> alpha) rApp rApp" |
876 | 161 |
by (auto intro: a2) |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
162 |
|
636
520a4084d064
changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents:
610
diff
changeset
|
163 |
lemma rLam_rsp[quot_respect]: "(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
|
164 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
165 |
apply(rule a3) |
804 | 166 |
apply(simp add: helper) |
167 |
apply(simp) |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
168 |
done |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
169 |
|
804 | 170 |
lemma rfv_rsp[quot_respect]: |
171 |
"(alpha ===> op =) rfv rfv" |
|
247 | 172 |
sorry |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
173 |
|
285
8ebdef196fd5
Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
273
diff
changeset
|
174 |
lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)" |
876 | 175 |
apply (auto) |
176 |
apply (erule alpha.cases) |
|
177 |
apply (simp_all add: rlam.inject alpha_refl) |
|
178 |
done |
|
237 | 179 |
|
804 | 180 |
lemma pi_var1: |
181 |
fixes pi::"'x prm" |
|
182 |
shows "pi \<bullet> Var a = Var (pi \<bullet> a)" |
|
183 |
by (lifting pi_var_eqvt1) |
|
370
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
184 |
|
804 | 185 |
lemma pi_var2: |
186 |
fixes pi::"'x prm" |
|
187 |
shows "pi \<bullet> Var a = Var (pi \<bullet> a)" |
|
188 |
by (lifting pi_var_eqvt2) |
|
189 |
||
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
190 |
|
804 | 191 |
lemma pi_app: |
192 |
fixes pi::"'x prm" |
|
193 |
shows "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)" |
|
194 |
by (lifting pi_app_eqvt2) |
|
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
195 |
|
804 | 196 |
lemma pi_lam: |
197 |
fixes pi::"'x prm" |
|
198 |
shows "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)" |
|
199 |
by (lifting pi_lam_eqvt2) |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
200 |
|
804 | 201 |
lemma fv_var: |
202 |
shows "fv (Var a) = {a}" |
|
203 |
by (lifting rfv_var) |
|
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
204 |
|
804 | 205 |
lemma fv_app: |
206 |
shows "fv (App t1 t2) = fv t1 \<union> fv t2" |
|
207 |
by (lifting rfv_app) |
|
208 |
||
209 |
lemma fv_lam: |
|
210 |
shows "fv (Lam a t) = fv t - {a}" |
|
211 |
by (lifting rfv_lam) |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
212 |
|
804 | 213 |
lemma a1: |
214 |
"a = b \<Longrightarrow> Var a = Var b" |
|
215 |
by (lifting a1) |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
216 |
|
804 | 217 |
lemma a2: |
218 |
"\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc" |
|
219 |
by (lifting a2) |
|
376
e99c0334d8bf
lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
370
diff
changeset
|
220 |
|
804 | 221 |
lemma a3: |
222 |
"\<lbrakk>x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa" |
|
223 |
by (lifting a3) |
|
286 | 224 |
|
804 | 225 |
lemma alpha_cases: |
226 |
"\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P; |
|
227 |
\<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P; |
|
228 |
\<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
|
229 |
\<Longrightarrow> P" |
804 | 230 |
by (lifting alpha.cases) |
378
86fba2c4eeef
All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
376
diff
changeset
|
231 |
|
804 | 232 |
lemma alpha_induct: |
233 |
"\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b); |
|
234 |
\<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc); |
|
235 |
\<And>x a b xa. |
|
370
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
292
diff
changeset
|
236 |
\<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
|
237 |
\<Longrightarrow> qxb qx qxa" |
804 | 238 |
by (lifting alpha.induct) |
239 |
||
240 |
lemma var_inject: |
|
241 |
"(Var a = Var b) = (a = b)" |
|
242 |
by (lifting rvar_inject) |
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
251
diff
changeset
|
243 |
|
804 | 244 |
lemma lam_induct: |
877
09a64cb04851
exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
876
diff
changeset
|
245 |
"\<lbrakk>\<And>name. P (Var name); |
804 | 246 |
\<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2); |
878
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
247 |
\<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
248 |
\<Longrightarrow> P lam" |
804 | 249 |
by (lifting rlam.induct) |
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
|
250 |
|
878
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
251 |
instance lam::pt_name |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
252 |
apply(default) |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
253 |
sorry |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
254 |
|
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
255 |
instance lam::fs_name |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
256 |
apply(default) |
877
09a64cb04851
exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
876
diff
changeset
|
257 |
sorry |
09a64cb04851
exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
876
diff
changeset
|
258 |
|
09a64cb04851
exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
876
diff
changeset
|
259 |
lemma lam_induct_strong: |
878
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
260 |
fixes a::"'a::fs_name" |
879 | 261 |
assumes a1: "\<And>name b. P b (Var name)" |
262 |
and a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)" |
|
263 |
and a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; name \<sharp> lam\<rbrakk> \<Longrightarrow> P b (Lam name lam)" |
|
264 |
shows "P a lam" |
|
878
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
265 |
proof - |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
266 |
have "\<And>(pi::name prm). P a (pi \<bullet> lam)" |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
267 |
proof (induct lam rule: lam_induct) |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
268 |
case (1 name pi) |
879 | 269 |
show "P a (pi \<bullet> Var name)" |
270 |
apply (simp only: pi_var1) |
|
271 |
apply (rule a1) |
|
272 |
done |
|
878
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
273 |
next |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
274 |
case (2 lam1 lam2 pi) |
879 | 275 |
have b1: "P a (pi \<bullet> lam1)" by fact |
276 |
have b2: "P a (pi \<bullet> lam2)" by fact |
|
277 |
show "P a (pi \<bullet> App lam1 lam2)" |
|
278 |
apply (simp only: pi_app) |
|
279 |
apply (rule a2) |
|
280 |
using b1 b2 |
|
281 |
sorry |
|
878
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
282 |
next |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
283 |
case (3 name lam pi) |
879 | 284 |
have b: "P a (pi \<bullet> lam)" by fact |
285 |
show "P a (pi \<bullet> Lam name lam)" |
|
286 |
apply (simp add: pi_lam) |
|
287 |
using b |
|
288 |
sorry |
|
878
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
289 |
qed |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
290 |
then have "P a (([]::name prm) \<bullet> lam)" by blast |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
291 |
then show "P a lam" by simp |
c3662f845129
setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents:
877
diff
changeset
|
292 |
qed |
877
09a64cb04851
exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
876
diff
changeset
|
293 |
|
09a64cb04851
exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
876
diff
changeset
|
294 |
|
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
295 |
lemma var_supp: |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
296 |
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
|
297 |
apply(simp add: supp_def) |
804 | 298 |
apply(simp add: pi_var2) |
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
299 |
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
|
300 |
done |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
301 |
|
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
302 |
lemma var_fresh: |
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
303 |
fixes a::"name" |
804 | 304 |
shows "(a \<sharp> (Var b)) = (a \<sharp> b)" |
249
7dec34d12328
added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents:
247
diff
changeset
|
305 |
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
|
306 |
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
|
307 |
done |
247 | 308 |
|
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
636
diff
changeset
|
309 |
end |