# HG changeset patch # User haftmann # Date 1201598396 -3600 # Node ID b6a64fe386345522061e42294852db304705ea7e # Parent 24c82bef5696c542fb24b9d22f2f202d0d13e06c treating division by zero properly diff -r 24c82bef5696 -r b6a64fe38634 src/HOL/Library/Code_Index.thy --- a/src/HOL/Library/Code_Index.thy Mon Jan 28 22:27:29 2008 +0100 +++ b/src/HOL/Library/Code_Index.thy Tue Jan 29 10:19:56 2008 +0100 @@ -180,6 +180,36 @@ *} +subsection {* Specialized @{term "op - \ index \ index \ index"}, + @{term "op div \ index \ index \ index"} and @{term "op mod \ index \ index \ index"} + operations *} + +definition + minus_index_aux :: "index \ index \ index" +where + [code func del]: "minus_index_aux = op -" + +lemma [code func]: "op - = minus_index_aux" + using minus_index_aux_def .. + +definition + div_mod_index :: "index \ index \ index \ index" +where + [code func del]: "div_mod_index n m = (n div m, n mod m)" + +lemma [code func]: + "div_mod_index n m = (if m = 0 then (0, n) else (n div m, n mod m))" + unfolding div_mod_index_def by auto + +lemma [code func]: + "n div m = fst (div_mod_index n m)" + unfolding div_mod_index_def by simp + +lemma [code func]: + "n mod m = snd (div_mod_index n m)" + unfolding div_mod_index_def by simp + + subsection {* Code serialization *} text {* Implementation of indices by bounded integers *} @@ -205,7 +235,7 @@ (OCaml "Pervasives.( + )") (Haskell infixl 6 "+") -code_const "op - \ index \ index \ index" +code_const "minus_index_aux \ index \ index \ index" (SML "Int.max/ (_/ -/ _,/ 0 : int)") (OCaml "Pervasives.max/ (_/ -/ _)/ (0 : int) ") (Haskell "max/ (_/ -/ _)/ (0 :: Int)") @@ -215,15 +245,10 @@ (OCaml "Pervasives.( * )") (Haskell infixl 7 "*") -code_const "op div \ index \ index \ index" - (SML "Int.div/ ((_),/ (_))") - (OCaml "Pervasives.( / )") - (Haskell "div") - -code_const "op mod \ index \ index \ index" - (SML "Int.mod/ ((_),/ (_))") - (OCaml "Pervasives.( mod )") - (Haskell "mod") +code_const div_mod_index + (SML "(fn n => fn m =>/ (n div m, n mod m))") + (OCaml "(fun n -> fun m ->/ (n '/ m, n mod m))") + (Haskell "divMod") code_const "op = \ index \ index \ bool" (SML "!((_ : Int.int) = _)") diff -r 24c82bef5696 -r b6a64fe38634 src/HOL/Library/Code_Integer.thy --- a/src/HOL/Library/Code_Integer.thy Mon Jan 28 22:27:29 2008 +0100 +++ b/src/HOL/Library/Code_Integer.thy Tue Jan 29 10:19:56 2008 +0100 @@ -6,7 +6,7 @@ header {* Pretty integer literals for code generation *} theory Code_Integer -imports Int +imports Presburger begin text {* diff -r 24c82bef5696 -r b6a64fe38634 src/HOL/Library/Efficient_Nat.thy --- a/src/HOL/Library/Efficient_Nat.thy Mon Jan 28 22:27:29 2008 +0100 +++ b/src/HOL/Library/Efficient_Nat.thy Tue Jan 29 10:19:56 2008 +0100 @@ -53,13 +53,21 @@ "n * m = nat (of_nat n * of_nat m)" unfolding of_nat_mult [symmetric] by simp -lemma div_nat_code [code]: - "n div m = nat (of_nat n div of_nat m)" - unfolding zdiv_int [symmetric] by simp +text {* Specialized @{term "op div \ nat \ nat \ nat"} + and @{term "op mod \ nat \ nat \ nat"} operations. *} + +definition + div_mod_nat_aux :: "nat \ nat \ nat \ nat" +where + [code func del]: "div_mod_nat_aux = Divides.divmod" -lemma mod_nat_code [code]: - "n mod m = nat (of_nat n mod of_nat m)" - unfolding zmod_int [symmetric] by simp +lemma [code func]: + "Divides.divmod n m = (if m = 0 then (0, n) else div_mod_nat_aux n m)" + unfolding div_mod_nat_aux_def divmod_def by simp + +lemma div_mod_aux_code [code]: + "div_mod_nat_aux n m = (nat (of_nat n div of_nat m), nat (of_nat n mod of_nat m))" + unfolding div_mod_nat_aux_def divmod_def zdiv_int [symmetric] zmod_int [symmetric] by simp lemma eq_nat_code [code]: "n = m \ (of_nat n \ int) = of_nat m" @@ -380,15 +388,10 @@ (OCaml "Big'_int.mult'_big'_int") (Haskell infixl 7 "*") -code_const "op div \ nat \ nat \ nat" - (SML "IntInf.div/ ((_),/ (_))") - (OCaml "Big'_int.div'_big'_int") - (Haskell "div") - -code_const "op mod \ nat \ nat \ nat" - (SML "IntInf.mod/ ((_),/ (_))") - (OCaml "Big'_int.mod'_big'_int") - (Haskell "mod") +code_const div_mod_nat_aux + (SML "IntInf.divMod/ ((_),/ (_))") + (OCaml "Big'_int.quomod'_big'_int") + (Haskell "divMod") code_const "op = \ nat \ nat \ bool" (SML "!((_ : IntInf.int) = _)") @@ -410,8 +413,6 @@ Suc ("(_ +/ 1)") "op + \ nat \ nat \ nat" ("(_ +/ _)") "op * \ nat \ nat \ nat" ("(_ */ _)") - "op div \ nat \ nat \ nat" ("(_ div/ _)") - "op mod \ nat \ nat \ nat" ("(_ mod/ _)") "op \ \ nat \ nat \ bool" ("(_ <=/ _)") "op < \ nat \ nat \ bool" ("(_