--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/activities/tphols09/IDW/MW-Ex5.thy Wed Mar 30 17:27:34 2016 +0100
@@ -0,0 +1,84 @@
+header {* Local theory specifications *}
+
+theory Ex1
+imports Main
+begin
+
+section {* Target context *}
+
+locale NAT =
+ fixes zero :: 'n
+ and succ :: "'n \<Rightarrow> 'n"
+ assumes succ_inject: "(succ m = succ n) \<longleftrightarrow> (m = n)"
+ and succ_neq_zero: "succ m \<noteq> zero"
+ and induct "P zero \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (succ n)) \<Longrightarrow> P n"
+begin
+
+
+section {* Some definitions *}
+
+ML {* val lthy = @{context} *}
+
+ML {* val ((one, (_, one_def)), lthy1) = lthy |> LocalTheory.define ""
+ ((@{binding one}, NoSyn), (Attrib.empty_binding, @{term "succ zero"})) *}
+
+ML {* val ((two, (_, two_def)), lthy2) = lthy1 |> LocalTheory.define ""
+ ((@{binding two}, NoSyn), (Attrib.empty_binding, @{term "succ one"})) *}
+
+ML {* val export = singleton (ProofContext.export lthy2 (LocalTheory.target_of lthy2)) *}
+ML {* export one_def *}
+ML {* export two_def *}
+
+ML {* val global_export =
+ singleton (ProofContext.export lthy2 (ProofContext.init (ProofContext.theory_of lthy2))) *}
+ML {* global_export one_def *}
+ML {* global_export two_def *}
+
+
+definition "one = succ zero"
+definition "two = succ one"
+
+
+section {* Proofs *}
+
+lemma "succ one = two" by (simp add: one_def two_def)
+
+ML {*
+ val lthy = @{context};
+
+ val th =
+ Goal.prove lthy [] [] @{prop "succ one = two"}
+ (fn _ => asm_full_simp_tac
+ (local_simpset_of lthy addsimps [@{thm one_def}, @{thm two_def}]) 1);
+
+ LocalTheory.note "" ((@{binding ex1}, []), [th]) lthy;
+*}
+
+local_setup {* fn lthy =>
+ let
+ val th =
+ Goal.prove lthy [] [] @{prop "succ one = two"}
+ (fn _ => asm_full_simp_tac
+ (local_simpset_of lthy addsimps [@{thm one_def}, @{thm two_def}]) 1);
+ in lthy |> LocalTheory.note "" ((@{binding ex1}, []), [th]) |> #2 end
+*}
+
+thm ex1
+
+
+section {* Derived definitional mechamisms *}
+
+inductive
+ Rec :: "'a \<Rightarrow> ('n \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> 'a \<Rightarrow> bool"
+ for e :: 'a and r :: "'n \<Rightarrow> 'a \<Rightarrow> 'a"
+where
+ Rec_zero: "Rec e r zero e"
+ | Rec_succ: "Rec e r m n \<Longrightarrow> Rec e r (succ m) (r m n)"
+
+thm Rec_def Rec.induct
+
+end
+
+thm NAT.Rec_def NAT.Rec.induct
+
+end