diff -r 329111e1c4ba -r af951c8fb80a IntEx.thy --- a/IntEx.thy Wed Oct 28 15:25:36 2009 +0100 +++ b/IntEx.thy Wed Oct 28 15:48:38 2009 +0100 @@ -45,9 +45,10 @@ where "my_neg (x, y) = (y, x)" -local_setup {* - old_make_const_def @{binding NEG} @{term "my_neg"} NoSyn @{typ "nat \ nat"} @{typ "my_int"} #> snd -*} +quotient_def (for my_int) + NEG::"my_int \ my_int" +where + "NEG \ my_neg" term NEG thm NEG_def @@ -62,9 +63,10 @@ where "my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)" -local_setup {* - old_make_const_def @{binding MULT} @{term "my_mult"} NoSyn @{typ "nat \ nat"} @{typ "my_int"} #> snd -*} +quotient_def (for my_int) + MULT::"my_int \ my_int \ my_int" +where + "MULT \ my_mult" term MULT thm MULT_def @@ -75,9 +77,10 @@ where "my_le (x, y) (u, v) = (x+v \ u+y)" -local_setup {* - old_make_const_def @{binding LE} @{term "my_le"} NoSyn @{typ "nat \ nat"} @{typ "my_int"} #> snd -*} +quotient_def (for my_int) + LE :: "my_int \ my_int \ bool" +where + "LE \ my_le" term LE thm LE_def @@ -90,7 +93,6 @@ term LESS thm LESS_def - definition ABS :: "my_int \ my_int" where @@ -148,12 +150,6 @@ ML {* lift_thm_my_int @{context} @{thm plus_assoc_pre} *} - - - - - - text {* Below is the construction site code used if things do now work *} ML {* val t_a = atomize_thm @{thm plus_assoc_pre} *}