src/HOL/UNITY/Simple/Common.thy
changeset 44871 fbfdc5ac86be
parent 37936 1e4c5015a72e
child 54863 82acc20ded73
--- a/src/HOL/UNITY/Simple/Common.thy	Sat Sep 10 21:47:55 2011 +0200
+++ b/src/HOL/UNITY/Simple/Common.thy	Sat Sep 10 22:11:55 2011 +0200
@@ -17,11 +17,11 @@
   ftime :: "nat=>nat"
   gtime :: "nat=>nat"
 
-axioms
-  fmono: "m \<le> n ==> ftime m \<le> ftime n"
-  gmono: "m \<le> n ==> gtime m \<le> gtime n"
+axiomatization where
+  fmono: "m \<le> n ==> ftime m \<le> ftime n" and
+  gmono: "m \<le> n ==> gtime m \<le> gtime n" and
 
-  fasc:  "m \<le> ftime n"
+  fasc:  "m \<le> ftime n" and
   gasc:  "m \<le> gtime n"
 
 definition common :: "nat set" where