src/FOL/ex/Locale_Test/Locale_Test3.thy
author blanchet
Tue, 07 Dec 2010 11:56:01 +0100
changeset 41051 2ed1b971fc20
parent 37134 29bd6c2ffba8
child 61489 b8d375aee0df
permissions -rw-r--r--
give the inner timeout mechanism a chance, since it gives more precise information to the user

(*  Title:      FOL/ex/Locale_Test/Locale_Test3.thy
    Author:     Clemens Ballarin

Test environment for the locale implementation.
*)

theory Locale_Test3
imports Locale_Test1
begin

interpretation le2: mixin_thy_merge gle gle'
  where "reflexive.less(gle', x, y) <-> gless'(x, y)"
proof -
  show "mixin_thy_merge(gle, gle')" by unfold_locales
  note reflexive = this[unfolded mixin_thy_merge_def, THEN conjunct2]
  show "reflexive.less(gle', x, y) <-> gless'(x, y)"
    by (simp add: reflexive.less_def[OF reflexive] gless'_def)
qed

end