Nominal/activities/tphols09/IDW/MW-Ex2.thy
changeset 415 f1be8028a4a9
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/activities/tphols09/IDW/MW-Ex2.thy	Wed Mar 30 17:27:34 2016 +0100
@@ -0,0 +1,75 @@
+theory Ex2
+imports Main
+begin
+
+section {* Simple methods *}
+
+ML Method.setup
+ML SIMPLE_METHOD'
+ML SIMPLE_METHOD
+
+method_setup my_method1 = {*
+  Scan.succeed (K (SIMPLE_METHOD' (fn i: int => no_tac)))
+*}  ""
+
+method_setup my_method2 = {*
+  Scan.succeed (fn ctxt: Proof.context =>
+    SIMPLE_METHOD' (fn i: int => no_tac))
+*}  ""
+
+method_setup my_method3 = {*
+  Attrib.thms >> (fn thms: thm list => fn ctxt: Proof.context =>
+    SIMPLE_METHOD' (fn i: int => no_tac))
+*}  ""
+
+
+section {* My simplification method *}
+
+ML {*
+structure My_Simps = NamedThmsFun
+(
+  val name = "my_simp";
+  val description = "My simp declaration";
+)
+*}
+setup My_Simps.setup
+
+method_setup my_simp = {*
+  Scan.succeed (fn ctxt =>
+    SIMPLE_METHOD' (fn i =>
+      CHANGED (asm_full_simp_tac
+        (HOL_basic_ss addsimps (My_Simps.get ctxt)) i)))
+*} ""
+
+lemma True
+proof
+
+  fix a b c
+  assume [my_simp]: "a = b"
+  assume [my_simp]: "b = c"
+  have "a = c" by my_simp
+
+qed
+
+
+method_setup my_simp_all = {*
+  Scan.succeed (fn ctxt =>
+    SIMPLE_METHOD
+      (CHANGED
+        (ALLGOALS (asm_full_simp_tac
+          (HOL_basic_ss addsimps (My_Simps.get ctxt))))))
+*} ""
+
+
+lemma True
+proof
+
+  fix a b c
+  assume [my_simp]: "a = b"
+  assume [my_simp]: "b = c"
+  have "a = c" and "c = b" by my_simp_all
+
+qed
+
+
+end