Nominal/activities/tphols09/IDW/MW-Ex6.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Wed, 30 Mar 2016 17:27:34 +0100
changeset 415 f1be8028a4a9
permissions -rw-r--r--
updated

header {* Local theory data *}

theory Ex2
imports Main
begin

ML {*
  structure My_Rules =
    NamedThmsFun(val name = "my_rule" val description = "My collection of rules")
*}
setup My_Rules.setup

ML {*
  fun declare_my_rule th phi context =
    My_Rules.add_thm (Morphism.thm phi th) context
*}

locale foo =
  fixes x y z :: nat
  assumes 1: "x \<noteq> y" and 2: "y \<noteq> z" and 3: "x \<noteq> z"
begin

declaration {* declare_my_rule @{thm 1} *}
declaration {* declare_my_rule @{thm 2} *}
declaration {* declare_my_rule @{thm 3} *}

ML {* My_Rules.get @{context} *}

end

ML {* My_Rules.get @{context} *}

interpretation bar: foo 7 42 1 by unfold_locales simp_all

ML {* My_Rules.get @{context} *}

thm my_rule


lemma True
proof
  fix a b c d :: nat
  assume *: "distinct [a, b, c, d]"
  interpret foo a b c using * by unfold_locales simp_all
  interpret foo b c d using * by unfold_locales simp_all

  ML_prf {* My_Rules.get @{context} *}

qed  

end