author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Thu, 29 Oct 2009 13:30:11 +0100 | |
changeset 237 | 80f1df49b940 |
parent 234 | 811f17de4031 |
child 238 | e9cc3a3aa5d1 |
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 |
|
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
nominal_datatype rlam = |
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 |
|
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
inductive |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100) |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
where |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2" |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
17 |
| a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a\<sharp>[b].s\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s" |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
|
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
quotient lam = rlam / alpha |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
sorry |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
|
203
7384115df9fd
added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents:
201
diff
changeset
|
22 |
print_quotients |
7384115df9fd
added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents:
201
diff
changeset
|
23 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
24 |
quotient_def (for lam) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
25 |
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
|
26 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
27 |
"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
|
28 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
29 |
quotient_def (for lam) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
30 |
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
|
31 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
32 |
"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
|
33 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
34 |
quotient_def (for lam) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
35 |
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
|
36 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
37 |
"Lam \<equiv> rLam" |
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
|
218
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
39 |
thm Var_def |
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
40 |
thm App_def |
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
41 |
thm Lam_def |
df05cd030d2f
added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents:
203
diff
changeset
|
42 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
43 |
(* 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
|
44 |
(* 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
|
45 |
overloading |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
46 |
perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam" (unchecked) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
47 |
begin |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
48 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
49 |
quotient_def (for lam) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
50 |
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
|
51 |
where |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
52 |
"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
|
53 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
54 |
end |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
55 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
56 |
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
|
57 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
58 |
(* lemmas that need to lift *) |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
59 |
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
|
60 |
fixes pi::"'x prm" |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
61 |
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
|
62 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
63 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
64 |
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
|
65 |
fixes pi::"'x prm" |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
66 |
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
|
67 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
68 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
69 |
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
|
70 |
fixes pi::"'x prm" |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
71 |
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
|
72 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
73 |
|
201
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
lemma real_alpha: |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
assumes "t = ([(a,b)]\<bullet>s)" "a\<sharp>s" |
1ac36993cc71
added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
shows "Lam a t = Lam b s" |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
77 |
sorry |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
78 |
|
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
79 |
|
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
80 |
|
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
81 |
|
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
82 |
|
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
83 |
(* Construction Site code *) |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
84 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
85 |
lemma perm_rsp: "(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
|
86 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
87 |
(* 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
|
88 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
89 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
90 |
lemma fresh_rsp: "(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
|
91 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
92 |
(* 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
|
93 |
sorry |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
94 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
95 |
lemma rVar_rsp: "(op = ===> alpha) rVar rVar" |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
96 |
apply(auto) |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
97 |
apply(rule a1) |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
98 |
apply(simp) |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
99 |
done |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
100 |
|
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
101 |
lemma rApp_rsp: "(alpha ===> alpha ===> alpha) rApp rApp" |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
102 |
apply(auto) |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
103 |
apply(rule a2) |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
104 |
apply (assumption) |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
105 |
apply (assumption) |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
106 |
done |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
107 |
|
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
108 |
lemma rLam_rsp: "(op = ===> alpha ===> alpha) rLam rLam" |
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
109 |
apply(auto) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
110 |
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
|
111 |
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
|
112 |
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
|
113 |
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
|
114 |
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
|
115 |
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
|
116 |
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
|
117 |
apply(assumption) |
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents:
225
diff
changeset
|
118 |
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
|
119 |
done |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
120 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
121 |
ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def} *} |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
122 |
ML {* val consts = [@{const_name "rVar"}, @{const_name "rApp"}, @{const_name "rLam"}, @{const_name "perm"}]; *} |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
123 |
|
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
124 |
ML {* val rty = @{typ "rlam"} *} |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
125 |
ML {* val qty = @{typ "lam"} *} |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
126 |
ML {* val rel = @{term "alpha"} *} |
225 | 127 |
ML {* val rel_eqv = (#equiv_thm o hd) (quotdata_lookup @{context}) *} |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
128 |
ML {* val rel_refl = @{thm EQUIV_REFL} OF [rel_eqv] *} |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
129 |
ML {* val quot = @{thm QUOTIENT_lam} *} |
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
130 |
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp} @ @{thms ho_all_prs ho_ex_prs} *} |
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
131 |
ML {* val trans2 = @{thm QUOT_TYPE_I_lam.R_trans2} *} |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
132 |
ML {* val reps_same = @{thm QUOT_TYPE_I_lam.REPS_same} *} |
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
133 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
134 |
ML {* add_lower_defs @{context} @{thms perm_lam_def} *} |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
135 |
ML {* val rr = @{thm eq_reflection} OF [hd (rev (add_lower_defs @{context} @{thms perm_lam_def}))] *} |
237 | 136 |
|
234
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
137 |
ML {* val rrr = repeat_eqsubst_thm @{context} @{thms prod_fun_id map_id} rr *} |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
138 |
ML {* |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
139 |
fun lift_thm_lam lthy t = |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
140 |
lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same defs t |
811f17de4031
Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
233
diff
changeset
|
141 |
*} |
237 | 142 |
|
143 |
ML {* lift_thm_lam @{context} @{thm pi_var_com} *} |
|
144 |
ML {* lift_thm_lam @{context} @{thm pi_app_com} *} |
|
145 |
ML {* lift_thm_lam @{context} @{thm pi_lam_com} *} |
|
146 |
||
147 |
fun |
|
148 |
option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)" |
|
149 |
where |
|
150 |
"option_map f (nSome x) = nSome (f x)" |
|
151 |
| "option_map f nNone = nNone" |
|
152 |
||
153 |
fun |
|
154 |
option_rel |
|
155 |
where |
|
156 |
"option_rel r (nSome x) (nSome y) = r x y" |
|
157 |
| "option_rel r _ _ = False" |
|
158 |
||
159 |
declare [[map noption = (option_map, option_rel)]] |
|
160 |
||
161 |
lemma OPT_QUOTIENT: |
|
162 |
assumes q: "QUOTIENT R Abs Rep" |
|
163 |
shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)" |
|
164 |
apply (unfold QUOTIENT_def) |
|
165 |
apply (auto) |
|
166 |
using q |
|
167 |
apply (unfold QUOTIENT_def) |
|
168 |
apply (case_tac "a :: 'b noption") |
|
169 |
apply (simp) |
|
170 |
apply (simp) |
|
171 |
apply (case_tac "a :: 'b noption") |
|
172 |
apply (simp only: option_map.simps) |
|
173 |
apply (subst option_rel.simps) |
|
174 |
(* Simp starts hanging so don't know how to continue *) |
|
175 |
sorry |
|
176 |
||
177 |
(* Christian: Does it make sense? *) |
|
178 |
lemma abs_fun_rsp: "(op = ===> alpha ===> op = ===> op =) abs_fun abs_fun" |
|
179 |
sorry |
|
180 |
||
181 |
(* Should not be needed *) |
|
182 |
lemma eq_rsp2: "((op = ===> op =) ===> (op = ===> op =) ===> op =) op = op =" |
|
183 |
apply auto |
|
184 |
apply (rule ext) |
|
185 |
apply auto |
|
186 |
apply (rule ext) |
|
187 |
apply auto |
|
188 |
done |
|
217
9cc87d02190a
First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
203
diff
changeset
|
189 |
|
237 | 190 |
(* Should not be needed *) |
191 |
lemma perm_rsp_eq: "(op = ===> (op = ===> op =) ===> op = ===> op =) op \<bullet> op \<bullet>" |
|
192 |
apply auto |
|
193 |
thm arg_cong2 |
|
194 |
apply (rule_tac f="perm x" in arg_cong2) |
|
195 |
apply (auto) |
|
196 |
apply (rule ext) |
|
197 |
apply (auto) |
|
198 |
done |
|
199 |
||
200 |
(* Should not be needed *) |
|
201 |
lemma fresh_rsp_eq: "(op = ===> (op = ===> op =) ===> op =) fresh fresh" |
|
202 |
apply (simp add: FUN_REL.simps) |
|
203 |
apply (metis ext) |
|
204 |
done |
|
205 |
||
206 |
(* It is just a test, it doesn't seem true... *) |
|
207 |
lemma quotient_cheat: "QUOTIENT op = (option_map ABS_lam) (option_map REP_lam)" |
|
208 |
sorry |
|
209 |
||
210 |
||
211 |
ML {* val rsp_thms = @{thms abs_fun_rsp OPT_QUOTIENT eq_rsp2 quotient_cheat perm_rsp_eq fresh_rsp_eq} @ rsp_thms *} |
|
212 |
ML {* |
|
213 |
fun lift_thm_lam lthy t = |
|
214 |
lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same defs t |
|
215 |
*} |
|
216 |
||
217 |
thm a3 |
|
218 |
ML {* Toplevel.program (fn () => lift_thm_lam @{context} @{thm a3}) *} |
|
219 |
ML {* val t_u = MetaSimplifier.rewrite_rule @{thms fresh_def supp_def} @{thm a3} *} |
|
220 |
ML {* Toplevel.program (fn () => lift_thm_lam @{context} t_u) *} |
|
221 |
||
222 |
ML t_u |
|
223 |
ML {* val t_a = atomize_thm t_u *} |
|
224 |
ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *} |
|
225 |
ML {* val t = (fst o Thm.dest_comb o snd o Thm.dest_comb) (cprop_of t_u) *} |
|
226 |
ML {* val t = (snd o Thm.dest_comb o snd o Thm.dest_comb) t *} |
|
227 |
ML {* val t = (snd o Thm.dest_comb o snd o Thm.dest_comb) t *} |
|
228 |
ML {* val t = (snd o (Thm.dest_abs NONE) o snd o Thm.dest_comb) t *} |
|
229 |
ML {* val t = (snd o Thm.dest_comb o snd o Thm.dest_comb) t *} |
|
230 |
ML {* val t = (snd o (Thm.dest_abs NONE) o snd o Thm.dest_comb) t *} |
|
231 |
ML {* val t = (fst o Thm.dest_comb o snd o Thm.dest_comb) t *} |
|
232 |
ML {* term_of t *} |
|
233 |
term "[b].(s::rlam)" |
|
234 |
thm abs_fun_def |
|
235 |
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *} |
|
236 |