Nominal/activities/tphols09/IDW/MW-Ex6.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 header {* Local theory data *}
       
     2 
       
     3 theory Ex2
       
     4 imports Main
       
     5 begin
       
     6 
       
     7 ML {*
       
     8   structure My_Rules =
       
     9     NamedThmsFun(val name = "my_rule" val description = "My collection of rules")
       
    10 *}
       
    11 setup My_Rules.setup
       
    12 
       
    13 ML {*
       
    14   fun declare_my_rule th phi context =
       
    15     My_Rules.add_thm (Morphism.thm phi th) context
       
    16 *}
       
    17 
       
    18 locale foo =
       
    19   fixes x y z :: nat
       
    20   assumes 1: "x \<noteq> y" and 2: "y \<noteq> z" and 3: "x \<noteq> z"
       
    21 begin
       
    22 
       
    23 declaration {* declare_my_rule @{thm 1} *}
       
    24 declaration {* declare_my_rule @{thm 2} *}
       
    25 declaration {* declare_my_rule @{thm 3} *}
       
    26 
       
    27 ML {* My_Rules.get @{context} *}
       
    28 
       
    29 end
       
    30 
       
    31 ML {* My_Rules.get @{context} *}
       
    32 
       
    33 interpretation bar: foo 7 42 1 by unfold_locales simp_all
       
    34 
       
    35 ML {* My_Rules.get @{context} *}
       
    36 
       
    37 thm my_rule
       
    38 
       
    39 
       
    40 lemma True
       
    41 proof
       
    42   fix a b c d :: nat
       
    43   assume *: "distinct [a, b, c, d]"
       
    44   interpret foo a b c using * by unfold_locales simp_all
       
    45   interpret foo b c d using * by unfold_locales simp_all
       
    46 
       
    47   ML_prf {* My_Rules.get @{context} *}
       
    48 
       
    49 qed  
       
    50 
       
    51 end