Nominal/activities/tphols09/IDW/MW-Ex5.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 header {* Local theory specifications *}
       
     2 
       
     3 theory Ex1
       
     4 imports Main
       
     5 begin
       
     6 
       
     7 section {* Target context *}
       
     8 
       
     9 locale NAT =
       
    10   fixes zero :: 'n
       
    11     and succ :: "'n \<Rightarrow> 'n"
       
    12   assumes succ_inject: "(succ m = succ n) \<longleftrightarrow> (m = n)"
       
    13     and succ_neq_zero: "succ m \<noteq> zero"
       
    14     and induct "P zero \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (succ n)) \<Longrightarrow> P n"
       
    15 begin
       
    16 
       
    17 
       
    18 section {* Some definitions *}
       
    19 
       
    20 ML {* val lthy = @{context} *}
       
    21 
       
    22 ML {* val ((one, (_, one_def)), lthy1) = lthy |> LocalTheory.define ""
       
    23   ((@{binding one}, NoSyn), (Attrib.empty_binding, @{term "succ zero"})) *}
       
    24 
       
    25 ML {* val ((two, (_, two_def)), lthy2) = lthy1 |> LocalTheory.define ""
       
    26   ((@{binding two}, NoSyn), (Attrib.empty_binding, @{term "succ one"})) *}
       
    27 
       
    28 ML {* val export = singleton (ProofContext.export lthy2 (LocalTheory.target_of lthy2)) *}
       
    29 ML {* export one_def *}
       
    30 ML {* export two_def *}
       
    31 
       
    32 ML {* val global_export =
       
    33   singleton (ProofContext.export lthy2 (ProofContext.init (ProofContext.theory_of lthy2))) *}
       
    34 ML {* global_export one_def *}
       
    35 ML {* global_export two_def *}
       
    36 
       
    37 
       
    38 definition "one = succ zero"
       
    39 definition "two = succ one"
       
    40 
       
    41 
       
    42 section {* Proofs *}
       
    43 
       
    44 lemma "succ one = two" by (simp add: one_def two_def)
       
    45 
       
    46 ML {*
       
    47   val lthy = @{context};
       
    48 
       
    49   val th =
       
    50     Goal.prove lthy [] [] @{prop "succ one = two"}
       
    51       (fn _ => asm_full_simp_tac
       
    52         (local_simpset_of lthy addsimps [@{thm one_def}, @{thm two_def}]) 1);
       
    53 
       
    54   LocalTheory.note "" ((@{binding ex1}, []), [th]) lthy;
       
    55 *}
       
    56 
       
    57 local_setup {* fn lthy =>
       
    58   let
       
    59     val th =
       
    60       Goal.prove lthy [] [] @{prop "succ one = two"}
       
    61         (fn _ => asm_full_simp_tac
       
    62           (local_simpset_of lthy addsimps [@{thm one_def}, @{thm two_def}]) 1);
       
    63   in lthy |> LocalTheory.note "" ((@{binding ex1}, []), [th]) |> #2 end
       
    64 *}
       
    65 
       
    66 thm ex1
       
    67 
       
    68 
       
    69 section {* Derived definitional mechamisms *}
       
    70 
       
    71 inductive
       
    72   Rec :: "'a \<Rightarrow> ('n \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> 'a \<Rightarrow> bool"
       
    73   for e :: 'a and r :: "'n \<Rightarrow> 'a \<Rightarrow> 'a"
       
    74 where
       
    75     Rec_zero: "Rec e r zero e"
       
    76   | Rec_succ: "Rec e r m n \<Longrightarrow> Rec e r (succ m) (r m n)"
       
    77 
       
    78 thm Rec_def Rec.induct
       
    79 
       
    80 end
       
    81 
       
    82 thm NAT.Rec_def NAT.Rec.induct
       
    83 
       
    84 end