src/FOL/ex/Locale_Test/Locale_Test3.thy
author wenzelm
Fri, 06 Mar 2015 16:09:51 +0100
changeset 59622 deae170e24a6
parent 37134 29bd6c2ffba8
child 61489 b8d375aee0df
permissions -rw-r--r--
updated to scala-2.11.6;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     1
(*  Title:      FOL/ex/Locale_Test/Locale_Test3.thy
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     2
    Author:     Clemens Ballarin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     3
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     4
Test environment for the locale implementation.
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     5
*)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     6
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     7
theory Locale_Test3
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     8
imports Locale_Test1
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     9
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    10
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    11
interpretation le2: mixin_thy_merge gle gle'
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    12
  where "reflexive.less(gle', x, y) <-> gless'(x, y)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    13
proof -
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    14
  show "mixin_thy_merge(gle, gle')" by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    15
  note reflexive = this[unfolded mixin_thy_merge_def, THEN conjunct2]
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    16
  show "reflexive.less(gle', x, y) <-> gless'(x, y)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    17
    by (simp add: reflexive.less_def[OF reflexive] gless'_def)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    18
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    19
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    20
end