src/FOL/ex/Locale_Test/Locale_Test.thy
author wenzelm
Sun, 12 Aug 2018 14:28:28 +0200
changeset 68743 91162dd89571
parent 61489 b8d375aee0df
child 69590 e65314985426
permissions -rw-r--r--
proper session dirs;
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_Test.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_Test
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     8
imports Locale_Test1 Locale_Test2 Locale_Test3
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     9
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    10
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 37134
diff changeset
    11
text \<open>Result of theory merge with distinct but identical interpretations\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    12
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    13
context mixin_thy_merge
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    14
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    15
lemmas less_mixin_thy_merge1 = le.less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    16
lemmas less_mixin_thy_merge2 = le'.less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    17
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    18
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
    19
lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y" (* mixin from first interpretation applied *)
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    20
  by (rule le1.less_mixin_thy_merge1)
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
    21
lemma "gless'(x, y) \<longleftrightarrow> gle'(x, y) \<and> x \<noteq> y" (* mixin from second interpretation applied *)
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    22
  by (rule le1.less_mixin_thy_merge2)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    23
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    24
end