src/LCF/ex/Ex3.thy
changeset 47025 b2b8ae61d6ad
parent 35762 af3ff2ba4c54
child 58889 5b7a9633cfa8
--- a/src/LCF/ex/Ex3.thy	Mon Mar 19 21:25:15 2012 +0100
+++ b/src/LCF/ex/Ex3.thy	Mon Mar 19 21:49:52 2012 +0100
@@ -4,20 +4,19 @@
 imports LCF
 begin
 
-consts
-  s     :: "'a => 'a"
+axiomatization
+  s     :: "'a => 'a" and
   p     :: "'a => 'a => 'a"
-
-axioms
-  p_strict:     "p(UU) = UU"
+where
+  p_strict:     "p(UU) = UU" and
   p_s:          "p(s(x),y) = s(p(x,y))"
 
 declare p_strict [simp] p_s [simp]
 
 lemma example: "p(FIX(s),y) = FIX(s)"
   apply (tactic {* induct_tac @{context} "s" 1 *})
-  apply (simp (no_asm))
-  apply (simp (no_asm))
+  apply simp
+  apply simp
   done
 
 end