201
|
1 |
theory LamEx
|
|
2 |
imports Nominal QuotMain
|
|
3 |
begin
|
|
4 |
|
|
5 |
atom_decl name
|
|
6 |
|
|
7 |
nominal_datatype rlam =
|
|
8 |
rVar "name"
|
|
9 |
| rApp "rlam" "rlam"
|
|
10 |
| rLam "name" "rlam"
|
|
11 |
|
|
12 |
inductive
|
|
13 |
alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
|
|
14 |
where
|
|
15 |
a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
|
|
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>
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
|
18 |
|
|
19 |
quotient lam = rlam / alpha
|
203
|
20 |
apply -
|
201
|
21 |
sorry
|
|
22 |
|
203
|
23 |
print_quotients
|
|
24 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
25 |
quotient_def (for lam)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
26 |
Var :: "name \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
27 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
28 |
"Var \<equiv> rVar"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
29 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
30 |
quotient_def (for lam)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
31 |
App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
32 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
33 |
"App \<equiv> rApp"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
34 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
35 |
quotient_def (for lam)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
36 |
Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
37 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
38 |
"Lam \<equiv> rLam"
|
201
|
39 |
|
218
|
40 |
thm Var_def
|
|
41 |
thm App_def
|
|
42 |
thm Lam_def
|
|
43 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
44 |
(* definition of overloaded permutation function *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
45 |
(* for the lifted type lam *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
46 |
overloading
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
47 |
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>
diff
changeset
|
48 |
begin
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
49 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
50 |
quotient_def (for lam)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
51 |
perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
52 |
where
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
53 |
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
54 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
55 |
end
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
56 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
57 |
thm perm_lam_def
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
58 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
59 |
(* lemmas that need to lift *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
60 |
lemma
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
61 |
fixes pi::"'x prm"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
62 |
shows "(pi\<bullet>Var a) = Var (pi\<bullet>a)"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
63 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
64 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
65 |
lemma
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
66 |
fixes pi::"'x prm"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
67 |
shows "(pi\<bullet>App t1 t2) = App (pi\<bullet>t1) (pi\<bullet>t2)"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
68 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
69 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
70 |
lemma
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
71 |
fixes pi::"'x prm"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
72 |
shows "(pi\<bullet>Lam a t) = Lam (pi\<bullet>a) (pi\<bullet>t)"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
73 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
74 |
|
201
|
75 |
lemma real_alpha:
|
|
76 |
assumes "t = ([(a,b)]\<bullet>s)" "a\<sharp>s"
|
|
77 |
shows "Lam a t = Lam b s"
|
217
|
78 |
sorry
|
|
79 |
|
|
80 |
|
|
81 |
|
|
82 |
|
|
83 |
|
|
84 |
(* Construction Site code *)
|
|
85 |
|
229
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
86 |
lemma perm_rsp: "op = ===> alpha ===> alpha op \<bullet> op \<bullet>"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
87 |
apply(auto)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
88 |
(* this is propably true if some type conditions are imposed ;o) *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
89 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
90 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
91 |
lemma fresh_rsp: "op = ===> (alpha ===> op =) fresh fresh"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
92 |
apply(auto)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
93 |
(* this is probably only true if some type conditions are imposed *)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
94 |
sorry
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
95 |
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
96 |
lemma rLam_rsp: "op = ===> (alpha ===> alpha) rLam rLam"
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
97 |
apply(auto)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
98 |
apply(rule a3)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
99 |
apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
100 |
apply(rule sym)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
101 |
apply(rule trans)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
102 |
apply(rule pt_name3)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
103 |
apply(rule at_ds1[OF at_name_inst])
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
104 |
apply(simp add: pt_name1)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
105 |
apply(assumption)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
106 |
apply(simp add: abs_fresh)
|
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
107 |
done
|
217
|
108 |
|
|
109 |
ML {* val defs = @{thms Var_def App_def Lam_def} *}
|
|
110 |
ML {* val consts = [@{const_name "rVar"}, @{const_name "rApp"}, @{const_name "rLam"}]; *}
|
|
111 |
|
|
112 |
ML {* val rty = @{typ "rlam"} *}
|
|
113 |
ML {* val qty = @{typ "lam"} *}
|
|
114 |
ML {* val rel = @{term "alpha"} *}
|
225
|
115 |
ML {* val rel_eqv = (#equiv_thm o hd) (quotdata_lookup @{context}) *}
|
217
|
116 |
ML {* val rel_refl = @{thm EQUIV_REFL} OF [rel_eqv] *}
|
|
117 |
ML {* val quot = @{thm QUOTIENT_lam} *}
|
|
118 |
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rLam_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
|
|
119 |
ML {* val trans2 = @{thm QUOT_TYPE_I_lam.R_trans2} *}
|
|
120 |
ML {* val reps_same = @{thm QUOT_TYPE_I_lam.REPS_same} *}
|
|
121 |
|
|
122 |
thm a3
|
|
123 |
ML {* val t_a = atomize_thm @{thm a3} *}
|
|
124 |
ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}
|
|
125 |
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
|
|
126 |
ML {* val abs = findabs rty (prop_of t_a) *}
|
|
127 |
ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *}
|
|
128 |
ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
|
|
129 |
ML {* val t_c = simp_allex_prs @{context} quot t_l *}
|
|
130 |
ML {* val t_defs_sym = add_lower_defs @{context} defs *}
|
|
131 |
ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_c *}
|
|
132 |
ML {* val t_b = MetaSimplifier.rewrite_rule [reps_same] t_d *}
|
|
133 |
ML {* ObjectLogic.rulify t_b *}
|
|
134 |
|
221
|
135 |
thm fresh_def
|
|
136 |
thm supp_def
|
|
137 |
|
217
|
138 |
local_setup {*
|
225
|
139 |
old_make_const_def @{binding lperm} @{term "perm :: ('a \<times> 'a) list \<Rightarrow> rlam \<Rightarrow> rlam"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd
|
217
|
140 |
*}
|
|
141 |
|
221
|
142 |
ML {* val consts = @{const_name perm} :: consts *}
|
|
143 |
ML {* val defs = @{thms lperm_def} @ defs *}
|
|
144 |
ML {* val t_u = MetaSimplifier.rewrite_rule @{thms fresh_def supp_def} @{thm a3} *}
|
225
|
145 |
ML {* val t_a = atomize_thm @{thm a3} *}
|
221
|
146 |
ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}
|
|
147 |
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
|
217
|
148 |
ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
|
|
149 |
ML {* val t_c = simp_allex_prs @{context} quot t_l *}
|
|
150 |
ML {* val t_defs_sym = add_lower_defs @{context} defs *}
|
|
151 |
ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_c *}
|
|
152 |
ML {* val t_b = MetaSimplifier.rewrite_rule [reps_same] t_d *}
|
|
153 |
ML {* val rr = (add_lower_defs @{context} @{thms lperm_def}) *}
|
225
|
154 |
ML {* val rrr = @{thm eq_reflection} OF [hd (rev rr)] *}
|
|
155 |
lemma prod_fun_id: "prod_fun id id = id"
|
|
156 |
apply (simp add: prod_fun_def)
|
|
157 |
done
|
|
158 |
lemma map_id: "map id x = x"
|
|
159 |
apply (induct x)
|
|
160 |
apply (simp_all add: map.simps)
|
|
161 |
done
|
217
|
162 |
|
225
|
163 |
ML {* val rrrr = repeat_eqsubst_thm @{context} @{thms prod_fun_id map_id} rrr *}
|
|
164 |
ML {* val t_b' = eqsubst_thm @{context} [rrrr] t_b *}
|
|
165 |
ML {* ObjectLogic.rulify t_b' *}
|
217
|
166 |
|
|
167 |
|
|
168 |
|
221
|
169 |
local_setup {*
|
|
170 |
make_const_def @{binding lfresh} @{term "fresh :: 'a \<Rightarrow> rlam \<Rightarrow> bool"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd #>
|
|
171 |
*}
|
|
172 |
@{const_name fresh} ::
|
|
173 |
lfresh_def
|
217
|
174 |
ML {*
|
|
175 |
fun lift_thm_lam lthy t =
|
|
176 |
lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same defs t
|
|
177 |
*}
|
|
178 |
|
|
179 |
ML {* Toplevel.program (fn () => lift_thm_lam @{context} @{thm a3}) *}
|
|
180 |
|