# HG changeset patch # User nipkow # Date 1291659482 -3600 # Node ID 81d337539d57054b16b4ee72f2a767c427ed4cee # Parent 3efa0ec42ed4a98a0e83899e3b966fe7462ae667 moved coercion decl. for int diff -r 3efa0ec42ed4 -r 81d337539d57 src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Mon Dec 06 17:33:25 2010 +0100 +++ b/src/HOL/Decision_Procs/Approximation.thy Mon Dec 06 19:18:02 2010 +0100 @@ -10,7 +10,6 @@ declare [[coercion_map map]] declare [[coercion_map "% f g h . g o h o f"]] declare [[coercion_map "% f g (x,y) . (f x, g y)"]] -declare [[coercion int]] declare [[coercion "% x . Float x 0"]] declare [[coercion "real::float\real"]] diff -r 3efa0ec42ed4 -r 81d337539d57 src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Mon Dec 06 17:33:25 2010 +0100 +++ b/src/HOL/RealDef.thy Mon Dec 06 19:18:02 2010 +0100 @@ -1112,6 +1112,7 @@ declare [[coercion_enabled]] declare [[coercion "real::nat\real"]] declare [[coercion "real::int\real"]] +declare [[coercion "int"]] lemma real_eq_of_nat: "real = of_nat" unfolding real_of_nat_def ..