src/FOL/ex/Locale_Test/Locale_Test3.thy
author wenzelm
Mon, 15 Feb 2016 14:55:44 +0100
changeset 62337 d3996d5873dd
parent 61566 c3d6e570ccef
child 69590 e65314985426
permissions -rw-r--r--
proper syntax;

(*  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'
  rewrites "reflexive.less(gle', x, y) \<longleftrightarrow> 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) \<longleftrightarrow> gless'(x, y)"
    by (simp add: reflexive.less_def[OF reflexive] gless'_def)
qed

end