author | wenzelm |
Wed, 04 Oct 2017 12:00:53 +0200 | |
changeset 66787 | 64b47495676d |
parent 61566 | c3d6e570ccef |
child 69590 | e65314985426 |
permissions | -rw-r--r-- |
37134 | 1 |
(* Title: FOL/ex/Locale_Test/Locale_Test3.thy |
2 |
Author: Clemens Ballarin |
|
3 |
||
4 |
Test environment for the locale implementation. |
|
5 |
*) |
|
6 |
||
7 |
theory Locale_Test3 |
|
8 |
imports Locale_Test1 |
|
9 |
begin |
|
10 |
||
11 |
interpretation le2: mixin_thy_merge gle gle' |
|
61566
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents:
61489
diff
changeset
|
12 |
rewrites "reflexive.less(gle', x, y) \<longleftrightarrow> gless'(x, y)" |
37134 | 13 |
proof - |
14 |
show "mixin_thy_merge(gle, gle')" by unfold_locales |
|
15 |
note reflexive = this[unfolded mixin_thy_merge_def, THEN conjunct2] |
|
61489 | 16 |
show "reflexive.less(gle', x, y) \<longleftrightarrow> gless'(x, y)" |
37134 | 17 |
by (simp add: reflexive.less_def[OF reflexive] gless'_def) |
18 |
qed |
|
19 |
||
20 |
end |