author | ballarin |
Wed, 08 Jun 2005 16:11:09 +0200 | |
changeset 16325 | a6431098a929 |
parent 16168 | adb83939177f |
child 16620 | 2a7f46324218 |
permissions | -rw-r--r-- |
15596 | 1 |
(* Title: FOL/ex/LocaleTest.thy |
2 |
ID: $Id$ |
|
3 |
Author: Clemens Ballarin |
|
4 |
Copyright (c) 2005 by Clemens Ballarin |
|
5 |
||
6 |
Collection of regression tests for locales. |
|
7 |
*) |
|
8 |
||
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
9 |
header {* Test of Locale Interpretation *} |
15596 | 10 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
11 |
theory LocaleTest |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
12 |
imports FOL |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
13 |
begin |
15596 | 14 |
|
15696 | 15 |
ML {* set quick_and_dirty *} (* allow for thm command in batch mode *) |
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
16 |
ML {* set Toplevel.debug *} |
15696 | 17 |
ML {* set show_hyps *} |
18 |
ML {* set show_sorts *} |
|
19 |
||
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
20 |
section {* Renaming with Syntax *} |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
21 |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
22 |
|
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
23 |
locale (open) S = var mult + |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
24 |
assumes "mult(x, y) = mult(y, x)" |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
25 |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
26 |
print_locale S |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
27 |
|
16325
a6431098a929
Fixed "axiom" generation for mixed locales with and without predicates.
ballarin
parents:
16168
diff
changeset
|
28 |
locale S' = S mult (infixl "**" 60) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
29 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
30 |
print_locale S' |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
31 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
32 |
locale T = var mult (infixl "**" 60) + |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
33 |
assumes "x ** y = y ** x" |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
34 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
35 |
locale U = T mult (infixl "**" 60) + T add (infixl "++" 55) + var h + |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
36 |
assumes hom: "h(x ** y) = h(x) ++ h(y)" |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
37 |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
38 |
locale V = U _ add |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
39 |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
40 |
|
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
41 |
section {* Constrains *} |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
42 |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
43 |
locale Z = fixes a (structure) |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
44 |
locale Z' = Z + |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
45 |
constrains a :: "'a => 'b" |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
46 |
assumes "a (x :: 'a) = a (y)" |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16102
diff
changeset
|
47 |
print_locale Z' |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
48 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
49 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
50 |
section {* Interpretation *} |
15696 | 51 |
|
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
52 |
(* interpretation input syntax *) |
15596 | 53 |
|
54 |
locale L |
|
55 |
locale M = fixes a and b and c |
|
56 |
||
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
57 |
interpretation test [simp]: L + M a b c [x y z] . |
15596 | 58 |
|
15837 | 59 |
print_interps L (* output: test *) |
60 |
print_interps M (* output: test *) |
|
15596 | 61 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
62 |
interpretation test [simp]: L print_interps M . |
15596 | 63 |
|
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
64 |
interpretation L . |
15596 | 65 |
|
66 |
(* processing of locale expression *) |
|
67 |
||
68 |
locale A = fixes a assumes asm_A: "a = a" |
|
69 |
||
70 |
locale (open) B = fixes b assumes asm_B [simp]: "b = b" |
|
71 |
||
72 |
locale C = A + B + assumes asm_C: "c = c" |
|
73 |
(* TODO: independent type var in c, prohibit locale declaration *) |
|
74 |
||
75 |
locale D = A + B + fixes d defines def_D: "d == (a = b)" |
|
76 |
||
15696 | 77 |
theorem (in A) |
78 |
includes D |
|
79 |
shows True .. |
|
80 |
||
81 |
theorem (in D) True .. |
|
15596 | 82 |
|
83 |
typedecl i |
|
84 |
arities i :: "term" |
|
85 |
||
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
86 |
|
15837 | 87 |
interpretation p1: C ["X::i" "Y::i"] by (auto intro: A.intro C_axioms.intro) |
15596 | 88 |
|
15837 | 89 |
print_interps A (* output: p1 *) |
15596 | 90 |
|
15696 | 91 |
(* possible accesses *) |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
92 |
thm p1.a.asm_A thm LocaleTest.p1.a.asm_A |
15763
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
ballarin
parents:
15696
diff
changeset
|
93 |
thm p1.asm_A thm LocaleTest.p1.asm_A |
15696 | 94 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
95 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
96 |
(* without prefix *) |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
97 |
|
15837 | 98 |
interpretation C ["W::i" "Z::i"] . (* subsumed by p1: C *) |
99 |
interpretation C ["W::'a" "Z::i"] by (auto intro: A.intro C_axioms.intro) |
|
100 |
(* subsumes p1: A and p1: C *) |
|
15696 | 101 |
|
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
102 |
|
15837 | 103 |
print_interps A (* output: <no prefix>, p1 *) |
15596 | 104 |
|
15837 | 105 |
(* possible accesses *) |
106 |
thm asm_C thm a_b.asm_C thm LocaleTest.a_b.asm_C thm LocaleTest.a_b.asm_C |
|
15696 | 107 |
|
15596 | 108 |
|
15837 | 109 |
interpretation p2: D [X "Y::i" "Y = X"] by (simp add: eq_commute) |
110 |
||
111 |
print_interps A (* output: <no prefix>, p1 *) |
|
112 |
print_interps D (* output: p2 *) |
|
113 |
||
114 |
||
115 |
interpretation p3: D [X "Y::i"] . |
|
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
116 |
|
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
117 |
(* duplicate: not registered *) |
15696 | 118 |
|
119 |
(* thm p3.a.asm_A *) |
|
120 |
||
15596 | 121 |
|
15837 | 122 |
print_interps A (* output: <no prefix>, p1 *) |
123 |
print_interps B (* output: p1 *) |
|
124 |
print_interps C (* output: <no name>, p1 *) |
|
125 |
print_interps D (* output: p2, p3 *) |
|
15596 | 126 |
|
15837 | 127 |
(* schematic vars in instantiation not permitted *) |
128 |
(* |
|
15696 | 129 |
interpretation p4: A ["?x::?'a1"] apply (rule A.intro) apply rule done |
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
130 |
print_interps A |
15696 | 131 |
*) |
15596 | 132 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
133 |
interpretation p10: D + D a' b' d' [X "Y::i" _ u "v::i" _] . |
15696 | 134 |
|
135 |
corollary (in D) th_x: True .. |
|
136 |
||
137 |
(* possible accesses: for each registration *) |
|
138 |
||
15837 | 139 |
thm p2.th_x thm p3.th_x |
15696 | 140 |
|
141 |
lemma (in D) th_y: "d == (a = b)" . |
|
142 |
||
15837 | 143 |
thm p2.th_y thm p3.th_y |
15696 | 144 |
|
145 |
lemmas (in D) th_z = th_y |
|
146 |
||
147 |
thm p2.th_z |
|
148 |
||
149 |
section {* Interpretation in proof contexts *} |
|
15596 | 150 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
151 |
locale F = fixes f assumes asm_F: "f & f --> f" |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
152 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
153 |
theorem True |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
154 |
proof - |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
155 |
fix alpha::i and beta::'a and gamma::o |
15837 | 156 |
(* FIXME: omitting type of beta leads to error later at interpret p6 *) |
15696 | 157 |
have alpha_A: "A(alpha)" by (auto intro: A.intro) |
15837 | 158 |
interpret p5: A [alpha] . (* subsumed *) |
159 |
print_interps A (* output: <no prefix>, p1 *) |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
160 |
interpret p6: C [alpha beta] by (auto intro: C_axioms.intro) |
15837 | 161 |
print_interps A (* output: <no prefix>, p1 *) |
162 |
print_interps C (* output: <no prefix>, p1, p6 *) |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
163 |
interpret p11: F [gamma] by (fast intro: F.intro) |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
164 |
thm p11.asm_F (* gamma is a Free *) |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
165 |
qed rule |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
166 |
|
15696 | 167 |
theorem (in A) True |
168 |
proof - |
|
169 |
print_interps A |
|
170 |
fix beta and gamma |
|
171 |
interpret p9: D [a beta _] |
|
172 |
(* no proof obligation for A !!! *) |
|
173 |
apply - apply (rule refl) apply assumption done |
|
174 |
qed rule |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
175 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15598
diff
changeset
|
176 |
|
15696 | 177 |
(* Definition involving free variable *) |
178 |
||
179 |
ML {* reset show_sorts *} |
|
180 |
||
181 |
locale E = fixes e defines e_def: "e(x) == x & x" |
|
182 |
notes e_def2 = e_def |
|
183 |
||
184 |
lemma (in E) True thm e_def by fast |
|
185 |
||
186 |
interpretation p7: E ["(%x. x)"] by simp |
|
187 |
||
188 |
(* TODO: goal mustn't be beta-reduced here, is doesn't match meta-hyp *) |
|
189 |
||
190 |
thm p7.e_def2 |
|
191 |
||
192 |
locale E' = fixes e defines e_def: "e == (%x. x & x)" |
|
193 |
notes e_def2 = e_def |
|
194 |
||
195 |
interpretation p7': E' ["(%x. x)"] by simp |
|
196 |
||
197 |
thm p7'.e_def2 |
|
198 |
||
199 |
(* Definition involving free variable in assm *) |
|
200 |
||
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
201 |
locale (open) G = fixes g assumes asm_G: "g --> x" |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
202 |
notes asm_G2 = asm_G |
15696 | 203 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
204 |
interpretation p8: G ["False"] by fast |
15696 | 205 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
15837
diff
changeset
|
206 |
thm p8.asm_G2 |
15696 | 207 |
|
208 |
subsection {* Locale without assumptions *} |
|
209 |
||
210 |
locale L1 = notes rev_conjI [intro] = conjI [THEN iffD1 [OF conj_commute]] |
|
211 |
||
212 |
lemma "[| P; Q |] ==> P & Q" |
|
213 |
proof - |
|
214 |
interpret my: L1 . txt {* No chained fact required. *} |
|
215 |
assume Q and P txt {* order reversed *} |
|
216 |
then show "P & Q" .. txt {* Applies @{thm my.rev_conjI}. *} |
|
217 |
qed |
|
218 |
||
219 |
locale L11 = notes rev_conjI = conjI [THEN iffD1 [OF conj_commute]] |
|
220 |
||
221 |
lemma "[| P; Q |] ==> P & Q" |
|
222 |
proof - |
|
223 |
interpret [intro]: L11 . txt {* Attribute supplied at instantiation. *} |
|
224 |
assume Q and P |
|
225 |
then show "P & Q" .. |
|
226 |
qed |
|
227 |
||
228 |
subsection {* Simple locale with assumptions *} |
|
229 |
||
230 |
consts bin :: "[i, i] => i" (infixl "#" 60) |
|
231 |
||
232 |
axioms i_assoc: "(x # y) # z = x # (y # z)" |
|
233 |
i_comm: "x # y = y # x" |
|
234 |
||
235 |
locale L2 = |
|
236 |
fixes OP (infixl "+" 60) |
|
237 |
assumes assoc: "(x + y) + z = x + (y + z)" |
|
238 |
and comm: "x + y = y + x" |
|
239 |
||
240 |
lemma (in L2) lcomm: "x + (y + z) = y + (x + z)" |
|
241 |
proof - |
|
242 |
have "x + (y + z) = (x + y) + z" by (simp add: assoc) |
|
243 |
also have "... = (y + x) + z" by (simp add: comm) |
|
244 |
also have "... = y + (x + z)" by (simp add: assoc) |
|
245 |
finally show ?thesis . |
|
246 |
qed |
|
247 |
||
248 |
lemmas (in L2) AC = comm assoc lcomm |
|
249 |
||
250 |
lemma "(x::i) # y # z # w = y # x # w # z" |
|
251 |
proof - |
|
252 |
interpret my: L2 ["op #"] by (rule L2.intro [of "op #", OF i_assoc i_comm]) |
|
253 |
txt {* Chained fact required to discharge assumptions of @{text L2} |
|
254 |
and instantiate parameters. *} |
|
255 |
show ?thesis by (simp only: my.OP.AC) (* or simply AC *) |
|
256 |
qed |
|
257 |
||
258 |
subsection {* Nested locale with assumptions *} |
|
259 |
||
260 |
locale L3 = |
|
261 |
fixes OP (infixl "+" 60) |
|
262 |
assumes assoc: "(x + y) + z = x + (y + z)" |
|
263 |
||
264 |
locale L4 = L3 + |
|
265 |
assumes comm: "x + y = y + x" |
|
266 |
||
267 |
lemma (in L4) lcomm: "x + (y + z) = y + (x + z)" |
|
268 |
proof - |
|
269 |
have "x + (y + z) = (x + y) + z" by (simp add: assoc) |
|
270 |
also have "... = (y + x) + z" by (simp add: comm) |
|
271 |
also have "... = y + (x + z)" by (simp add: assoc) |
|
272 |
finally show ?thesis . |
|
273 |
qed |
|
274 |
||
275 |
lemmas (in L4) AC = comm assoc lcomm |
|
276 |
||
277 |
lemma "(x::i) # y # z # w = y # x # w # z" |
|
278 |
proof - |
|
279 |
interpret my: L4 ["op #"] |
|
280 |
by (auto intro: L3.intro L4_axioms.intro i_assoc i_comm) |
|
281 |
show ?thesis by (simp only: my.OP.AC) (* or simply AC *) |
|
282 |
qed |
|
283 |
||
284 |
subsection {* Locale with definition *} |
|
285 |
||
286 |
text {* This example is admittedly not very creative :-) *} |
|
287 |
||
288 |
locale L5 = L4 + var A + |
|
289 |
defines A_def: "A == True" |
|
290 |
||
291 |
lemma (in L5) lem: A |
|
292 |
by (unfold A_def) rule |
|
293 |
||
294 |
lemma "L5(op #) ==> True" |
|
295 |
proof - |
|
296 |
assume "L5(op #)" |
|
297 |
then interpret L5 ["op #"] by (auto intro: L5.axioms) |
|
298 |
show ?thesis by (rule lem) (* lem instantiated to True *) |
|
299 |
qed |
|
300 |
||
301 |
subsection {* Instantiation in a context with target *} |
|
302 |
||
303 |
lemma (in L4) |
|
304 |
fixes A (infixl "$" 60) |
|
305 |
assumes A: "L4(A)" |
|
306 |
shows "(x::i) $ y $ z $ w = y $ x $ w $ z" |
|
307 |
proof - |
|
308 |
from A interpret A: L4 ["A"] by (auto intro: L4.axioms) |
|
309 |
show ?thesis by (simp only: A.OP.AC) |
|
310 |
qed |
|
311 |
||
15596 | 312 |
end |