# HG changeset patch # User haftmann # Date 1352365358 -3600 # Node ID 28f3263d4d1b0907109b7c805c771263915e68c2 # Parent 286dfcab9833a06fc8c571425f50572af8223273 refined stack of library theories implementing int and/or nat by target language numerals diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Code_Binary_Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Binary_Nat.thy Thu Nov 08 10:02:38 2012 +0100 @@ -0,0 +1,259 @@ +(* Title: HOL/Library/Code_Binary_Nat.thy + Author: Stefan Berghofer, Florian Haftmann, TU Muenchen +*) + +header {* Implementation of natural numbers as binary numerals *} + +theory Code_Binary_Nat +imports Main +begin + +text {* + When generating code for functions on natural numbers, the + canonical representation using @{term "0::nat"} and + @{term Suc} is unsuitable for computations involving large + numbers. This theory refines the representation of + natural numbers for code generation to use binary + numerals, which do not grow linear in size but logarithmic. +*} + +subsection {* Representation *} + +code_datatype "0::nat" nat_of_num + +lemma [code_abbrev]: + "nat_of_num = numeral" + by (fact nat_of_num_numeral) + +lemma [code]: + "num_of_nat 0 = Num.One" + "num_of_nat (nat_of_num k) = k" + by (simp_all add: nat_of_num_inverse) + +lemma [code]: + "(1\nat) = Numeral1" + by simp + +lemma [code_abbrev]: "Numeral1 = (1\nat)" + by simp + +lemma [code]: + "Suc n = n + 1" + by simp + + +subsection {* Basic arithmetic *} + +lemma [code, code del]: + "(plus :: nat \ _) = plus" .. + +lemma plus_nat_code [code]: + "nat_of_num k + nat_of_num l = nat_of_num (k + l)" + "m + 0 = (m::nat)" + "0 + n = (n::nat)" + by (simp_all add: nat_of_num_numeral) + +text {* Bounded subtraction needs some auxiliary *} + +definition dup :: "nat \ nat" where + "dup n = n + n" + +lemma dup_code [code]: + "dup 0 = 0" + "dup (nat_of_num k) = nat_of_num (Num.Bit0 k)" + by (simp_all add: dup_def numeral_Bit0) + +definition sub :: "num \ num \ nat option" where + "sub k l = (if k \ l then Some (numeral k - numeral l) else None)" + +lemma sub_code [code]: + "sub Num.One Num.One = Some 0" + "sub (Num.Bit0 m) Num.One = Some (nat_of_num (Num.BitM m))" + "sub (Num.Bit1 m) Num.One = Some (nat_of_num (Num.Bit0 m))" + "sub Num.One (Num.Bit0 n) = None" + "sub Num.One (Num.Bit1 n) = None" + "sub (Num.Bit0 m) (Num.Bit0 n) = Option.map dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit1 n) = Option.map dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit0 n) = Option.map (\q. dup q + 1) (sub m n)" + "sub (Num.Bit0 m) (Num.Bit1 n) = (case sub m n of None \ None + | Some q \ if q = 0 then None else Some (dup q - 1))" + apply (auto simp add: nat_of_num_numeral + Num.dbl_def Num.dbl_inc_def Num.dbl_dec_def + Let_def le_imp_diff_is_add BitM_plus_one sub_def dup_def) + apply (simp_all add: sub_non_positive) + apply (simp_all add: sub_non_negative [symmetric, where ?'a = int]) + done + +lemma [code, code del]: + "(minus :: nat \ _) = minus" .. + +lemma minus_nat_code [code]: + "nat_of_num k - nat_of_num l = (case sub k l of None \ 0 | Some j \ j)" + "m - 0 = (m::nat)" + "0 - n = (0::nat)" + by (simp_all add: nat_of_num_numeral sub_non_positive sub_def) + +lemma [code, code del]: + "(times :: nat \ _) = times" .. + +lemma times_nat_code [code]: + "nat_of_num k * nat_of_num l = nat_of_num (k * l)" + "m * 0 = (0::nat)" + "0 * n = (0::nat)" + by (simp_all add: nat_of_num_numeral) + +lemma [code, code del]: + "(HOL.equal :: nat \ _) = HOL.equal" .. + +lemma equal_nat_code [code]: + "HOL.equal 0 (0::nat) \ True" + "HOL.equal 0 (nat_of_num l) \ False" + "HOL.equal (nat_of_num k) 0 \ False" + "HOL.equal (nat_of_num k) (nat_of_num l) \ HOL.equal k l" + by (simp_all add: nat_of_num_numeral equal) + +lemma equal_nat_refl [code nbe]: + "HOL.equal (n::nat) n \ True" + by (rule equal_refl) + +lemma [code, code del]: + "(less_eq :: nat \ _) = less_eq" .. + +lemma less_eq_nat_code [code]: + "0 \ (n::nat) \ True" + "nat_of_num k \ 0 \ False" + "nat_of_num k \ nat_of_num l \ k \ l" + by (simp_all add: nat_of_num_numeral) + +lemma [code, code del]: + "(less :: nat \ _) = less" .. + +lemma less_nat_code [code]: + "(m::nat) < 0 \ False" + "0 < nat_of_num l \ True" + "nat_of_num k < nat_of_num l \ k < l" + by (simp_all add: nat_of_num_numeral) + + +subsection {* Conversions *} + +lemma [code, code del]: + "of_nat = of_nat" .. + +lemma of_nat_code [code]: + "of_nat 0 = 0" + "of_nat (nat_of_num k) = numeral k" + by (simp_all add: nat_of_num_numeral) + + +subsection {* Case analysis *} + +text {* + Case analysis on natural numbers is rephrased using a conditional + expression: +*} + +lemma [code, code_unfold]: + "nat_case = (\f g n. if n = 0 then f else g (n - 1))" + by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc) + + +subsection {* Preprocessors *} + +text {* + The term @{term "Suc n"} is no longer a valid pattern. + Therefore, all occurrences of this term in a position + where a pattern is expected (i.e.~on the left-hand side of a recursion + equation) must be eliminated. + This can be accomplished by applying the following transformation rules: +*} + +lemma Suc_if_eq: "(\n. f (Suc n) \ h n) \ f 0 \ g \ + f n \ if n = 0 then g else h (n - 1)" + by (rule eq_reflection) (cases n, simp_all) + +text {* + The rules above are built into a preprocessor that is plugged into + the code generator. Since the preprocessor for introduction rules + does not know anything about modes, some of the modes that worked + for the canonical representation of natural numbers may no longer work. +*} + +(*<*) +setup {* +let + +fun remove_suc thy thms = + let + val vname = singleton (Name.variant_list (map fst + (fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n"; + val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT)); + fun lhs_of th = snd (Thm.dest_comb + (fst (Thm.dest_comb (cprop_of th)))); + fun rhs_of th = snd (Thm.dest_comb (cprop_of th)); + fun find_vars ct = (case term_of ct of + (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))] + | _ $ _ => + let val (ct1, ct2) = Thm.dest_comb ct + in + map (apfst (fn ct => Thm.apply ct ct2)) (find_vars ct1) @ + map (apfst (Thm.apply ct1)) (find_vars ct2) + end + | _ => []); + val eqs = maps + (fn th => map (pair th) (find_vars (lhs_of th))) thms; + fun mk_thms (th, (ct, cv')) = + let + val th' = + Thm.implies_elim + (Conv.fconv_rule (Thm.beta_conversion true) + (Drule.instantiate' + [SOME (ctyp_of_term ct)] [SOME (Thm.lambda cv ct), + SOME (Thm.lambda cv' (rhs_of th)), NONE, SOME cv'] + @{thm Suc_if_eq})) (Thm.forall_intr cv' th) + in + case map_filter (fn th'' => + SOME (th'', singleton + (Variable.trade (K (fn [th'''] => [th''' RS th'])) + (Variable.global_thm_context th'')) th'') + handle THM _ => NONE) thms of + [] => NONE + | thps => + let val (ths1, ths2) = split_list thps + in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end + end + in get_first mk_thms eqs end; + +fun eqn_suc_base_preproc thy thms = + let + val dest = fst o Logic.dest_equals o prop_of; + val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc}); + in + if forall (can dest) thms andalso exists (contains_suc o dest) thms + then thms |> perhaps_loop (remove_suc thy) |> (Option.map o map) Drule.zero_var_indexes + else NONE + end; + +val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc; + +in + + Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc) + +end; +*} +(*>*) + +code_modulename SML + Code_Binary_Nat Arith + +code_modulename OCaml + Code_Binary_Nat Arith + +code_modulename Haskell + Code_Binary_Nat Arith + +hide_const (open) dup sub + +end + diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Code_Nat.thy --- a/src/HOL/Library/Code_Nat.thy Wed Nov 07 20:48:04 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,258 +0,0 @@ -(* Title: HOL/Library/Code_Nat.thy - Author: Stefan Berghofer, Florian Haftmann, TU Muenchen -*) - -header {* Implementation of natural numbers as binary numerals *} - -theory Code_Nat -imports Main -begin - -text {* - When generating code for functions on natural numbers, the - canonical representation using @{term "0::nat"} and - @{term Suc} is unsuitable for computations involving large - numbers. This theory refines the representation of - natural numbers for code generation to use binary - numerals, which do not grow linear in size but logarithmic. -*} - -subsection {* Representation *} - -lemma [code_abbrev]: - "nat_of_num = numeral" - by (fact nat_of_num_numeral) - -code_datatype "0::nat" nat_of_num - -lemma [code]: - "num_of_nat 0 = Num.One" - "num_of_nat (nat_of_num k) = k" - by (simp_all add: nat_of_num_inverse) - -lemma [code]: - "(1\nat) = Numeral1" - by simp - -lemma [code_abbrev]: "Numeral1 = (1\nat)" - by simp - -lemma [code]: - "Suc n = n + 1" - by simp - - -subsection {* Basic arithmetic *} - -lemma [code, code del]: - "(plus :: nat \ _) = plus" .. - -lemma plus_nat_code [code]: - "nat_of_num k + nat_of_num l = nat_of_num (k + l)" - "m + 0 = (m::nat)" - "0 + n = (n::nat)" - by (simp_all add: nat_of_num_numeral) - -text {* Bounded subtraction needs some auxiliary *} - -definition dup :: "nat \ nat" where - "dup n = n + n" - -lemma dup_code [code]: - "dup 0 = 0" - "dup (nat_of_num k) = nat_of_num (Num.Bit0 k)" - unfolding Num_def by (simp_all add: dup_def numeral_Bit0) - -definition sub :: "num \ num \ nat option" where - "sub k l = (if k \ l then Some (numeral k - numeral l) else None)" - -lemma sub_code [code]: - "sub Num.One Num.One = Some 0" - "sub (Num.Bit0 m) Num.One = Some (nat_of_num (Num.BitM m))" - "sub (Num.Bit1 m) Num.One = Some (nat_of_num (Num.Bit0 m))" - "sub Num.One (Num.Bit0 n) = None" - "sub Num.One (Num.Bit1 n) = None" - "sub (Num.Bit0 m) (Num.Bit0 n) = Option.map dup (sub m n)" - "sub (Num.Bit1 m) (Num.Bit1 n) = Option.map dup (sub m n)" - "sub (Num.Bit1 m) (Num.Bit0 n) = Option.map (\q. dup q + 1) (sub m n)" - "sub (Num.Bit0 m) (Num.Bit1 n) = (case sub m n of None \ None - | Some q \ if q = 0 then None else Some (dup q - 1))" - apply (auto simp add: nat_of_num_numeral - Num.dbl_def Num.dbl_inc_def Num.dbl_dec_def - Let_def le_imp_diff_is_add BitM_plus_one sub_def dup_def) - apply (simp_all add: sub_non_positive) - apply (simp_all add: sub_non_negative [symmetric, where ?'a = int]) - done - -lemma [code, code del]: - "(minus :: nat \ _) = minus" .. - -lemma minus_nat_code [code]: - "nat_of_num k - nat_of_num l = (case sub k l of None \ 0 | Some j \ j)" - "m - 0 = (m::nat)" - "0 - n = (0::nat)" - by (simp_all add: nat_of_num_numeral sub_non_positive sub_def) - -lemma [code, code del]: - "(times :: nat \ _) = times" .. - -lemma times_nat_code [code]: - "nat_of_num k * nat_of_num l = nat_of_num (k * l)" - "m * 0 = (0::nat)" - "0 * n = (0::nat)" - by (simp_all add: nat_of_num_numeral) - -lemma [code, code del]: - "(HOL.equal :: nat \ _) = HOL.equal" .. - -lemma equal_nat_code [code]: - "HOL.equal 0 (0::nat) \ True" - "HOL.equal 0 (nat_of_num l) \ False" - "HOL.equal (nat_of_num k) 0 \ False" - "HOL.equal (nat_of_num k) (nat_of_num l) \ HOL.equal k l" - by (simp_all add: nat_of_num_numeral equal) - -lemma equal_nat_refl [code nbe]: - "HOL.equal (n::nat) n \ True" - by (rule equal_refl) - -lemma [code, code del]: - "(less_eq :: nat \ _) = less_eq" .. - -lemma less_eq_nat_code [code]: - "0 \ (n::nat) \ True" - "nat_of_num k \ 0 \ False" - "nat_of_num k \ nat_of_num l \ k \ l" - by (simp_all add: nat_of_num_numeral) - -lemma [code, code del]: - "(less :: nat \ _) = less" .. - -lemma less_nat_code [code]: - "(m::nat) < 0 \ False" - "0 < nat_of_num l \ True" - "nat_of_num k < nat_of_num l \ k < l" - by (simp_all add: nat_of_num_numeral) - - -subsection {* Conversions *} - -lemma [code, code del]: - "of_nat = of_nat" .. - -lemma of_nat_code [code]: - "of_nat 0 = 0" - "of_nat (nat_of_num k) = numeral k" - by (simp_all add: nat_of_num_numeral) - - -subsection {* Case analysis *} - -text {* - Case analysis on natural numbers is rephrased using a conditional - expression: -*} - -lemma [code, code_unfold]: - "nat_case = (\f g n. if n = 0 then f else g (n - 1))" - by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc) - - -subsection {* Preprocessors *} - -text {* - The term @{term "Suc n"} is no longer a valid pattern. - Therefore, all occurrences of this term in a position - where a pattern is expected (i.e.~on the left-hand side of a recursion - equation) must be eliminated. - This can be accomplished by applying the following transformation rules: -*} - -lemma Suc_if_eq: "(\n. f (Suc n) \ h n) \ f 0 \ g \ - f n \ if n = 0 then g else h (n - 1)" - by (rule eq_reflection) (cases n, simp_all) - -text {* - The rules above are built into a preprocessor that is plugged into - the code generator. Since the preprocessor for introduction rules - does not know anything about modes, some of the modes that worked - for the canonical representation of natural numbers may no longer work. -*} - -(*<*) -setup {* -let - -fun remove_suc thy thms = - let - val vname = singleton (Name.variant_list (map fst - (fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n"; - val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT)); - fun lhs_of th = snd (Thm.dest_comb - (fst (Thm.dest_comb (cprop_of th)))); - fun rhs_of th = snd (Thm.dest_comb (cprop_of th)); - fun find_vars ct = (case term_of ct of - (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))] - | _ $ _ => - let val (ct1, ct2) = Thm.dest_comb ct - in - map (apfst (fn ct => Thm.apply ct ct2)) (find_vars ct1) @ - map (apfst (Thm.apply ct1)) (find_vars ct2) - end - | _ => []); - val eqs = maps - (fn th => map (pair th) (find_vars (lhs_of th))) thms; - fun mk_thms (th, (ct, cv')) = - let - val th' = - Thm.implies_elim - (Conv.fconv_rule (Thm.beta_conversion true) - (Drule.instantiate' - [SOME (ctyp_of_term ct)] [SOME (Thm.lambda cv ct), - SOME (Thm.lambda cv' (rhs_of th)), NONE, SOME cv'] - @{thm Suc_if_eq})) (Thm.forall_intr cv' th) - in - case map_filter (fn th'' => - SOME (th'', singleton - (Variable.trade (K (fn [th'''] => [th''' RS th'])) - (Variable.global_thm_context th'')) th'') - handle THM _ => NONE) thms of - [] => NONE - | thps => - let val (ths1, ths2) = split_list thps - in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end - end - in get_first mk_thms eqs end; - -fun eqn_suc_base_preproc thy thms = - let - val dest = fst o Logic.dest_equals o prop_of; - val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc}); - in - if forall (can dest) thms andalso exists (contains_suc o dest) thms - then thms |> perhaps_loop (remove_suc thy) |> (Option.map o map) Drule.zero_var_indexes - else NONE - end; - -val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc; - -in - - Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc) - -end; -*} -(*>*) - -code_modulename SML - Code_Nat Arith - -code_modulename OCaml - Code_Nat Arith - -code_modulename Haskell - Code_Nat Arith - -hide_const (open) dup sub - -end diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Code_Numeral_Types.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Numeral_Types.thy Thu Nov 08 10:02:38 2012 +0100 @@ -0,0 +1,870 @@ +(* Title: HOL/Library/Code_Numeral_Types.thy + Author: Florian Haftmann, TU Muenchen +*) + +header {* Numeric types for code generation onto target language numerals only *} + +theory Code_Numeral_Types +imports Main Nat_Transfer Divides Lifting +begin + +subsection {* Type of target language integers *} + +typedef integer = "UNIV \ int set" + morphisms int_of_integer integer_of_int .. + +lemma integer_eq_iff: + "k = l \ int_of_integer k = int_of_integer l" + using int_of_integer_inject [of k l] .. + +lemma integer_eqI: + "int_of_integer k = int_of_integer l \ k = l" + using integer_eq_iff [of k l] by simp + +lemma int_of_integer_integer_of_int [simp]: + "int_of_integer (integer_of_int k) = k" + using integer_of_int_inverse [of k] by simp + +lemma integer_of_int_int_of_integer [simp]: + "integer_of_int (int_of_integer k) = k" + using int_of_integer_inverse [of k] by simp + +instantiation integer :: ring_1 +begin + +definition + "0 = integer_of_int 0" + +lemma int_of_integer_zero [simp]: + "int_of_integer 0 = 0" + by (simp add: zero_integer_def) + +definition + "1 = integer_of_int 1" + +lemma int_of_integer_one [simp]: + "int_of_integer 1 = 1" + by (simp add: one_integer_def) + +definition + "k + l = integer_of_int (int_of_integer k + int_of_integer l)" + +lemma int_of_integer_plus [simp]: + "int_of_integer (k + l) = int_of_integer k + int_of_integer l" + by (simp add: plus_integer_def) + +definition + "- k = integer_of_int (- int_of_integer k)" + +lemma int_of_integer_uminus [simp]: + "int_of_integer (- k) = - int_of_integer k" + by (simp add: uminus_integer_def) + +definition + "k - l = integer_of_int (int_of_integer k - int_of_integer l)" + +lemma int_of_integer_minus [simp]: + "int_of_integer (k - l) = int_of_integer k - int_of_integer l" + by (simp add: minus_integer_def) + +definition + "k * l = integer_of_int (int_of_integer k * int_of_integer l)" + +lemma int_of_integer_times [simp]: + "int_of_integer (k * l) = int_of_integer k * int_of_integer l" + by (simp add: times_integer_def) + +instance proof +qed (auto simp add: integer_eq_iff algebra_simps) + +end + +lemma int_of_integer_of_nat [simp]: + "int_of_integer (of_nat n) = of_nat n" + by (induct n) simp_all + +definition nat_of_integer :: "integer \ nat" +where + "nat_of_integer k = Int.nat (int_of_integer k)" + +lemma nat_of_integer_of_nat [simp]: + "nat_of_integer (of_nat n) = n" + by (simp add: nat_of_integer_def) + +lemma int_of_integer_of_int [simp]: + "int_of_integer (of_int k) = k" + by (induct k) (simp_all, simp only: neg_numeral_def numeral_One int_of_integer_uminus int_of_integer_one) + +lemma integer_integer_of_int_eq_of_integer_integer_of_int [simp, code_abbrev]: + "integer_of_int = of_int" + by rule (simp add: integer_eq_iff) + +lemma of_int_integer_of [simp]: + "of_int (int_of_integer k) = (k :: integer)" + by (simp add: integer_eq_iff) + +lemma int_of_integer_numeral [simp]: + "int_of_integer (numeral k) = numeral k" + using int_of_integer_of_int [of "numeral k"] by simp + +lemma int_of_integer_neg_numeral [simp]: + "int_of_integer (neg_numeral k) = neg_numeral k" + by (simp only: neg_numeral_def int_of_integer_uminus) simp + +lemma int_of_integer_sub [simp]: + "int_of_integer (Num.sub k l) = Num.sub k l" + by (simp only: Num.sub_def int_of_integer_minus int_of_integer_numeral) + +instantiation integer :: "{ring_div, equal, linordered_idom}" +begin + +definition + "k div l = of_int (int_of_integer k div int_of_integer l)" + +lemma int_of_integer_div [simp]: + "int_of_integer (k div l) = int_of_integer k div int_of_integer l" + by (simp add: div_integer_def) + +definition + "k mod l = of_int (int_of_integer k mod int_of_integer l)" + +lemma int_of_integer_mod [simp]: + "int_of_integer (k mod l) = int_of_integer k mod int_of_integer l" + by (simp add: mod_integer_def) + +definition + "\k\ = of_int \int_of_integer k\" + +lemma int_of_integer_abs [simp]: + "int_of_integer \k\ = \int_of_integer k\" + by (simp add: abs_integer_def) + +definition + "sgn k = of_int (sgn (int_of_integer k))" + +lemma int_of_integer_sgn [simp]: + "int_of_integer (sgn k) = sgn (int_of_integer k)" + by (simp add: sgn_integer_def) + +definition + "k \ l \ int_of_integer k \ int_of_integer l" + +definition + "k < l \ int_of_integer k < int_of_integer l" + +definition + "HOL.equal k l \ HOL.equal (int_of_integer k) (int_of_integer l)" + +instance proof +qed (auto simp add: integer_eq_iff algebra_simps + less_eq_integer_def less_integer_def equal_integer_def equal + intro: mult_strict_right_mono) + +end + +lemma int_of_integer_min [simp]: + "int_of_integer (min k l) = min (int_of_integer k) (int_of_integer l)" + by (simp add: min_def less_eq_integer_def) + +lemma int_of_integer_max [simp]: + "int_of_integer (max k l) = max (int_of_integer k) (int_of_integer l)" + by (simp add: max_def less_eq_integer_def) + +lemma nat_of_integer_non_positive [simp]: + "k \ 0 \ nat_of_integer k = 0" + by (simp add: nat_of_integer_def less_eq_integer_def) + +lemma of_nat_of_integer [simp]: + "of_nat (nat_of_integer k) = max 0 k" + by (simp add: nat_of_integer_def integer_eq_iff less_eq_integer_def max_def) + + +subsection {* Code theorems for target language integers *} + +text {* Constructors *} + +definition Pos :: "num \ integer" +where + [simp, code_abbrev]: "Pos = numeral" + +definition Neg :: "num \ integer" +where + [simp, code_abbrev]: "Neg = neg_numeral" + +code_datatype "0::integer" Pos Neg + + +text {* Auxiliary operations *} + +definition dup :: "integer \ integer" +where + [simp]: "dup k = k + k" + +lemma dup_code [code]: + "dup 0 = 0" + "dup (Pos n) = Pos (Num.Bit0 n)" + "dup (Neg n) = Neg (Num.Bit0 n)" + unfolding Pos_def Neg_def neg_numeral_def + by (simp_all add: numeral_Bit0) + +definition sub :: "num \ num \ integer" +where + [simp]: "sub m n = numeral m - numeral n" + +lemma sub_code [code]: + "sub Num.One Num.One = 0" + "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" + "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" + "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" + "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" + "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" + "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" + unfolding sub_def dup_def numeral.simps Pos_def Neg_def + neg_numeral_def numeral_BitM + by (simp_all only: algebra_simps add.comm_neutral) + + +text {* Implementations *} + +lemma one_integer_code [code, code_unfold]: + "1 = Pos Num.One" + by simp + +lemma plus_integer_code [code]: + "k + 0 = (k::integer)" + "0 + l = (l::integer)" + "Pos m + Pos n = Pos (m + n)" + "Pos m + Neg n = sub m n" + "Neg m + Pos n = sub n m" + "Neg m + Neg n = Neg (m + n)" + by simp_all + +lemma uminus_integer_code [code]: + "uminus 0 = (0::integer)" + "uminus (Pos m) = Neg m" + "uminus (Neg m) = Pos m" + by simp_all + +lemma minus_integer_code [code]: + "k - 0 = (k::integer)" + "0 - l = uminus (l::integer)" + "Pos m - Pos n = sub m n" + "Pos m - Neg n = Pos (m + n)" + "Neg m - Pos n = Neg (m + n)" + "Neg m - Neg n = sub n m" + by simp_all + +lemma abs_integer_code [code]: + "\k\ = (if (k::integer) < 0 then - k else k)" + by simp + +lemma sgn_integer_code [code]: + "sgn k = (if k = 0 then 0 else if (k::integer) < 0 then - 1 else 1)" + by simp + +lemma times_integer_code [code]: + "k * 0 = (0::integer)" + "0 * l = (0::integer)" + "Pos m * Pos n = Pos (m * n)" + "Pos m * Neg n = Neg (m * n)" + "Neg m * Pos n = Neg (m * n)" + "Neg m * Neg n = Pos (m * n)" + by simp_all + +definition divmod_integer :: "integer \ integer \ integer \ integer" +where + "divmod_integer k l = (k div l, k mod l)" + +lemma fst_divmod [simp]: + "fst (divmod_integer k l) = k div l" + by (simp add: divmod_integer_def) + +lemma snd_divmod [simp]: + "snd (divmod_integer k l) = k mod l" + by (simp add: divmod_integer_def) + +definition divmod_abs :: "integer \ integer \ integer \ integer" +where + "divmod_abs k l = (\k\ div \l\, \k\ mod \l\)" + +lemma fst_divmod_abs [simp]: + "fst (divmod_abs k l) = \k\ div \l\" + by (simp add: divmod_abs_def) + +lemma snd_divmod_abs [simp]: + "snd (divmod_abs k l) = \k\ mod \l\" + by (simp add: divmod_abs_def) + +lemma divmod_abs_terminate_code [code]: + "divmod_abs (Neg k) (Neg l) = divmod_abs (Pos k) (Pos l)" + "divmod_abs (Neg k) (Pos l) = divmod_abs (Pos k) (Pos l)" + "divmod_abs (Pos k) (Neg l) = divmod_abs (Pos k) (Pos l)" + "divmod_abs j 0 = (0, \j\)" + "divmod_abs 0 j = (0, 0)" + by (simp_all add: prod_eq_iff) + +lemma divmod_abs_rec_code [code]: + "divmod_abs (Pos k) (Pos l) = + (let j = sub k l in + if j < 0 then (0, Pos k) + else let (q, r) = divmod_abs j (Pos l) in (q + 1, r))" + by (auto simp add: prod_eq_iff integer_eq_iff Let_def prod_case_beta + sub_non_negative sub_negative div_pos_pos_trivial mod_pos_pos_trivial div_pos_geq mod_pos_geq) + +lemma divmod_integer_code [code]: "divmod_integer k l = + (if k = 0 then (0, 0) else if l = 0 then (0, k) else + (apsnd \ times \ sgn) l (if sgn k = sgn l + then divmod_abs k l + else (let (r, s) = divmod_abs k l in + if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" +proof - + have aux1: "\k l::int. sgn k = sgn l \ k = 0 \ l = 0 \ 0 < l \ 0 < k \ l < 0 \ k < 0" + by (auto simp add: sgn_if) + have aux2: "\q::int. - int_of_integer k = int_of_integer l * q \ int_of_integer k = int_of_integer l * - q" by auto + show ?thesis + by (simp add: prod_eq_iff integer_eq_iff prod_case_beta aux1) + (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right aux2) +qed + +lemma div_integer_code [code]: + "k div l = fst (divmod_integer k l)" + by simp + +lemma mod_integer_code [code]: + "k mod l = snd (divmod_integer k l)" + by simp + +lemma equal_integer_code [code]: + "HOL.equal 0 (0::integer) \ True" + "HOL.equal 0 (Pos l) \ False" + "HOL.equal 0 (Neg l) \ False" + "HOL.equal (Pos k) 0 \ False" + "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" + "HOL.equal (Pos k) (Neg l) \ False" + "HOL.equal (Neg k) 0 \ False" + "HOL.equal (Neg k) (Pos l) \ False" + "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" + by (simp_all add: equal integer_eq_iff) + +lemma equal_integer_refl [code nbe]: + "HOL.equal (k::integer) k \ True" + by (fact equal_refl) + +lemma less_eq_integer_code [code]: + "0 \ (0::integer) \ True" + "0 \ Pos l \ True" + "0 \ Neg l \ False" + "Pos k \ 0 \ False" + "Pos k \ Pos l \ k \ l" + "Pos k \ Neg l \ False" + "Neg k \ 0 \ True" + "Neg k \ Pos l \ True" + "Neg k \ Neg l \ l \ k" + by (simp_all add: less_eq_integer_def) + +lemma less_integer_code [code]: + "0 < (0::integer) \ False" + "0 < Pos l \ True" + "0 < Neg l \ False" + "Pos k < 0 \ False" + "Pos k < Pos l \ k < l" + "Pos k < Neg l \ False" + "Neg k < 0 \ True" + "Neg k < Pos l \ True" + "Neg k < Neg l \ l < k" + by (simp_all add: less_integer_def) + +definition integer_of_num :: "num \ integer" +where + "integer_of_num = numeral" + +lemma integer_of_num [code]: + "integer_of_num num.One = 1" + "integer_of_num (num.Bit0 n) = (let k = integer_of_num n in k + k)" + "integer_of_num (num.Bit1 n) = (let k = integer_of_num n in k + k + 1)" + by (simp_all only: Let_def) (simp_all only: integer_of_num_def numeral.simps) + +definition num_of_integer :: "integer \ num" +where + "num_of_integer = num_of_nat \ nat_of_integer" + +lemma num_of_integer_code [code]: + "num_of_integer k = (if k \ 1 then Num.One + else let + (l, j) = divmod_integer k 2; + l' = num_of_integer l; + l'' = l' + l' + in if j = 0 then l'' else l'' + Num.One)" +proof - + { + assume "int_of_integer k mod 2 = 1" + then have "nat (int_of_integer k mod 2) = nat 1" by simp + moreover assume *: "1 < int_of_integer k" + ultimately have **: "nat (int_of_integer k) mod 2 = 1" by (simp add: nat_mod_distrib) + have "num_of_nat (nat (int_of_integer k)) = + num_of_nat (2 * (nat (int_of_integer k) div 2) + nat (int_of_integer k) mod 2)" + by simp + then have "num_of_nat (nat (int_of_integer k)) = + num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + nat (int_of_integer k) mod 2)" + by (simp add: mult_2) + with ** have "num_of_nat (nat (int_of_integer k)) = + num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + 1)" + by simp + } + note aux = this + show ?thesis + by (auto simp add: num_of_integer_def nat_of_integer_def Let_def prod_case_beta + not_le integer_eq_iff less_eq_integer_def + nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib + mult_2 [where 'a=nat] aux add_One) +qed + +lemma nat_of_integer_code [code]: + "nat_of_integer k = (if k \ 0 then 0 + else let + (l, j) = divmod_integer k 2; + l' = nat_of_integer l; + l'' = l' + l' + in if j = 0 then l'' else l'' + 1)" +proof - + obtain j where "k = integer_of_int j" + proof + show "k = integer_of_int (int_of_integer k)" by simp + qed + moreover have "2 * (j div 2) = j - j mod 2" + by (simp add: zmult_div_cancel mult_commute) + ultimately show ?thesis + by (auto simp add: split_def Let_def mod_integer_def nat_of_integer_def not_le + nat_add_distrib [symmetric] Suc_nat_eq_nat_zadd1) +qed + +lemma int_of_integer_code [code]: + "int_of_integer k = (if k < 0 then - (int_of_integer (- k)) + else if k = 0 then 0 + else let + (l, j) = divmod_integer k 2; + l' = 2 * int_of_integer l + in if j = 0 then l' else l' + 1)" + by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel) + +lemma integer_of_int_code [code]: + "integer_of_int k = (if k < 0 then - (integer_of_int (- k)) + else if k = 0 then 0 + else let + (l, j) = divmod_int k 2; + l' = 2 * integer_of_int l + in if j = 0 then l' else l' + 1)" + by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel) + +hide_const (open) Pos Neg sub dup divmod_abs + + +subsection {* Serializer setup for target language integers *} + +code_reserved Eval abs + +code_type integer + (SML "IntInf.int") + (OCaml "Big'_int.big'_int") + (Haskell "Integer") + (Scala "BigInt") + (Eval "int") + +code_instance integer :: equal + (Haskell -) + +code_const "0::integer" + (SML "0") + (OCaml "Big'_int.zero'_big'_int") + (Haskell "0") + (Scala "BigInt(0)") + +setup {* + fold (Numeral.add_code @{const_name Code_Numeral_Types.Pos} + false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] +*} + +setup {* + fold (Numeral.add_code @{const_name Code_Numeral_Types.Neg} + true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] +*} + +code_const "plus :: integer \ _ \ _" + (SML "IntInf.+ ((_), (_))") + (OCaml "Big'_int.add'_big'_int") + (Haskell infixl 6 "+") + (Scala infixl 7 "+") + (Eval infixl 8 "+") + +code_const "uminus :: integer \ _" + (SML "IntInf.~") + (OCaml "Big'_int.minus'_big'_int") + (Haskell "negate") + (Scala "!(- _)") + (Eval "~/ _") + +code_const "minus :: integer \ _" + (SML "IntInf.- ((_), (_))") + (OCaml "Big'_int.sub'_big'_int") + (Haskell infixl 6 "-") + (Scala infixl 7 "-") + (Eval infixl 8 "-") + +code_const Code_Numeral_Types.dup + (SML "IntInf.*/ (2,/ (_))") + (OCaml "Big'_int.mult'_big'_int/ (Big'_int.big'_int'_of'_int/ 2)") + (Haskell "!(2 * _)") + (Scala "!(2 * _)") + (Eval "!(2 * _)") + +code_const Code_Numeral_Types.sub + (SML "!(raise/ Fail/ \"sub\")") + (OCaml "failwith/ \"sub\"") + (Haskell "error/ \"sub\"") + (Scala "!sys.error(\"sub\")") + +code_const "times :: integer \ _ \ _" + (SML "IntInf.* ((_), (_))") + (OCaml "Big'_int.mult'_big'_int") + (Haskell infixl 7 "*") + (Scala infixl 8 "*") + (Eval infixl 9 "*") + +code_const Code_Numeral_Types.divmod_abs + (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") + (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") + (Haskell "divMod/ (abs _)/ (abs _)") + (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") + (Eval "Integer.div'_mod/ (abs _)/ (abs _)") + +code_const "HOL.equal :: integer \ _ \ bool" + (SML "!((_ : IntInf.int) = _)") + (OCaml "Big'_int.eq'_big'_int") + (Haskell infix 4 "==") + (Scala infixl 5 "==") + (Eval infixl 6 "=") + +code_const "less_eq :: integer \ _ \ bool" + (SML "IntInf.<= ((_), (_))") + (OCaml "Big'_int.le'_big'_int") + (Haskell infix 4 "<=") + (Scala infixl 4 "<=") + (Eval infixl 6 "<=") + +code_const "less :: integer \ _ \ bool" + (SML "IntInf.< ((_), (_))") + (OCaml "Big'_int.lt'_big'_int") + (Haskell infix 4 "<") + (Scala infixl 4 "<") + (Eval infixl 6 "<") + +code_modulename SML + Code_Numeral_Types Arith + +code_modulename OCaml + Code_Numeral_Types Arith + +code_modulename Haskell + Code_Numeral_Types Arith + + +subsection {* Type of target language naturals *} + +typedef natural = "UNIV \ nat set" + morphisms nat_of_natural natural_of_nat .. + +lemma natural_eq_iff [termination_simp]: + "m = n \ nat_of_natural m = nat_of_natural n" + using nat_of_natural_inject [of m n] .. + +lemma natural_eqI: + "nat_of_natural m = nat_of_natural n \ m = n" + using natural_eq_iff [of m n] by simp + +lemma nat_of_natural_of_nat_inverse [simp]: + "nat_of_natural (natural_of_nat n) = n" + using natural_of_nat_inverse [of n] by simp + +lemma natural_of_nat_of_natural_inverse [simp]: + "natural_of_nat (nat_of_natural n) = n" + using nat_of_natural_inverse [of n] by simp + +instantiation natural :: "{comm_monoid_diff, semiring_1}" +begin + +definition + "0 = natural_of_nat 0" + +lemma nat_of_natural_zero [simp]: + "nat_of_natural 0 = 0" + by (simp add: zero_natural_def) + +definition + "1 = natural_of_nat 1" + +lemma nat_of_natural_one [simp]: + "nat_of_natural 1 = 1" + by (simp add: one_natural_def) + +definition + "m + n = natural_of_nat (nat_of_natural m + nat_of_natural n)" + +lemma nat_of_natural_plus [simp]: + "nat_of_natural (m + n) = nat_of_natural m + nat_of_natural n" + by (simp add: plus_natural_def) + +definition + "m - n = natural_of_nat (nat_of_natural m - nat_of_natural n)" + +lemma nat_of_natural_minus [simp]: + "nat_of_natural (m - n) = nat_of_natural m - nat_of_natural n" + by (simp add: minus_natural_def) + +definition + "m * n = natural_of_nat (nat_of_natural m * nat_of_natural n)" + +lemma nat_of_natural_times [simp]: + "nat_of_natural (m * n) = nat_of_natural m * nat_of_natural n" + by (simp add: times_natural_def) + +instance proof +qed (auto simp add: natural_eq_iff algebra_simps) + +end + +lemma nat_of_natural_of_nat [simp]: + "nat_of_natural (of_nat n) = n" + by (induct n) simp_all + +lemma natural_of_nat_of_nat [simp, code_abbrev]: + "natural_of_nat = of_nat" + by rule (simp add: natural_eq_iff) + +lemma of_nat_of_natural [simp]: + "of_nat (nat_of_natural n) = n" + using natural_of_nat_of_natural_inverse [of n] by simp + +lemma nat_of_natural_numeral [simp]: + "nat_of_natural (numeral k) = numeral k" + using nat_of_natural_of_nat [of "numeral k"] by simp + +instantiation natural :: "{semiring_div, equal, linordered_semiring}" +begin + +definition + "m div n = natural_of_nat (nat_of_natural m div nat_of_natural n)" + +lemma nat_of_natural_div [simp]: + "nat_of_natural (m div n) = nat_of_natural m div nat_of_natural n" + by (simp add: div_natural_def) + +definition + "m mod n = natural_of_nat (nat_of_natural m mod nat_of_natural n)" + +lemma nat_of_natural_mod [simp]: + "nat_of_natural (m mod n) = nat_of_natural m mod nat_of_natural n" + by (simp add: mod_natural_def) + +definition + [termination_simp]: "m \ n \ nat_of_natural m \ nat_of_natural n" + +definition + [termination_simp]: "m < n \ nat_of_natural m < nat_of_natural n" + +definition + "HOL.equal m n \ HOL.equal (nat_of_natural m) (nat_of_natural n)" + +instance proof +qed (auto simp add: natural_eq_iff algebra_simps + less_eq_natural_def less_natural_def equal_natural_def equal) + +end + +lemma nat_of_natural_min [simp]: + "nat_of_natural (min k l) = min (nat_of_natural k) (nat_of_natural l)" + by (simp add: min_def less_eq_natural_def) + +lemma nat_of_natural_max [simp]: + "nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)" + by (simp add: max_def less_eq_natural_def) + +definition natural_of_integer :: "integer \ natural" +where + "natural_of_integer = of_nat \ nat_of_integer" + +definition integer_of_natural :: "natural \ integer" +where + "integer_of_natural = of_nat \ nat_of_natural" + +lemma natural_of_integer_of_natural [simp]: + "natural_of_integer (integer_of_natural n) = n" + by (simp add: natural_of_integer_def integer_of_natural_def natural_eq_iff) + +lemma integer_of_natural_of_integer [simp]: + "integer_of_natural (natural_of_integer k) = max 0 k" + by (simp add: natural_of_integer_def integer_of_natural_def integer_eq_iff) + +lemma int_of_integer_of_natural [simp]: + "int_of_integer (integer_of_natural n) = of_nat (nat_of_natural n)" + by (simp add: integer_of_natural_def) + +lemma integer_of_natural_of_nat [simp]: + "integer_of_natural (of_nat n) = of_nat n" + by (simp add: integer_eq_iff) + +lemma [measure_function]: + "is_measure nat_of_natural" by (rule is_measure_trivial) + + +subsection {* Inductive represenation of target language naturals *} + +definition Suc :: "natural \ natural" +where + "Suc = natural_of_nat \ Nat.Suc \ nat_of_natural" + +lemma nat_of_natural_Suc [simp]: + "nat_of_natural (Suc n) = Nat.Suc (nat_of_natural n)" + by (simp add: Suc_def) + +rep_datatype "0::natural" Suc +proof - + fix P :: "natural \ bool" + fix n :: natural + assume "P 0" then have init: "P (natural_of_nat 0)" by simp + assume "\n. P n \ P (Suc n)" + then have "\n. P (natural_of_nat n) \ P (Suc (natural_of_nat n))" . + then have step: "\n. P (natural_of_nat n) \ P (natural_of_nat (Nat.Suc n))" + by (simp add: Suc_def) + from init step have "P (natural_of_nat (nat_of_natural n))" + by (rule nat.induct) + with natural_of_nat_of_natural_inverse show "P n" by simp +qed (simp_all add: natural_eq_iff) + +lemma natural_case [case_names nat, cases type: natural]: + fixes m :: natural + assumes "\n. m = of_nat n \ P" + shows P + by (rule assms [of "nat_of_natural m"]) simp + +lemma [simp, code]: + "natural_size = nat_of_natural" +proof (rule ext) + fix n + show "natural_size n = nat_of_natural n" + by (induct n) simp_all +qed + +lemma [simp, code]: + "size = nat_of_natural" +proof (rule ext) + fix n + show "size n = nat_of_natural n" + by (induct n) simp_all +qed + +lemma natural_decr [termination_simp]: + "n \ 0 \ nat_of_natural n - Nat.Suc 0 < nat_of_natural n" + by (simp add: natural_eq_iff) + +lemma natural_zero_minus_one: + "(0::natural) - 1 = 0" + by simp + +lemma Suc_natural_minus_one: + "Suc n - 1 = n" + by (simp add: natural_eq_iff) + +hide_const (open) Suc + + +subsection {* Code refinement for target language naturals *} + +definition Nat :: "integer \ natural" +where + "Nat = natural_of_integer" + +lemma [code abstype]: + "Nat (integer_of_natural n) = n" + by (unfold Nat_def) (fact natural_of_integer_of_natural) + +lemma [code abstract]: + "integer_of_natural (natural_of_nat n) = of_nat n" + by simp + +lemma [code abstract]: + "integer_of_natural (natural_of_integer k) = max 0 k" + by simp + +lemma [code_abbrev]: + "natural_of_integer (Code_Numeral_Types.Pos k) = numeral k" + by (simp add: nat_of_integer_def natural_of_integer_def) + +lemma [code abstract]: + "integer_of_natural 0 = 0" + by (simp add: integer_eq_iff) + +lemma [code abstract]: + "integer_of_natural 1 = 1" + by (simp add: integer_eq_iff) + +lemma [code abstract]: + "integer_of_natural (Code_Numeral_Types.Suc n) = integer_of_natural n + 1" + by (simp add: integer_eq_iff) + +lemma [code]: + "nat_of_natural = nat_of_integer \ integer_of_natural" + by (simp add: integer_of_natural_def fun_eq_iff) + +lemma [code, code_unfold]: + "natural_case f g n = (if n = 0 then f else g (n - 1))" + by (cases n rule: natural.exhaust) (simp_all, simp add: Suc_def) + +declare natural.recs [code del] + +lemma [code abstract]: + "integer_of_natural (m + n) = integer_of_natural m + integer_of_natural n" + by (simp add: integer_eq_iff) + +lemma [code abstract]: + "integer_of_natural (m - n) = max 0 (integer_of_natural m - integer_of_natural n)" + by (simp add: integer_eq_iff) + +lemma [code abstract]: + "integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n" + by (simp add: integer_eq_iff of_nat_mult) + +lemma [code abstract]: + "integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n" + by (simp add: integer_eq_iff zdiv_int) + +lemma [code abstract]: + "integer_of_natural (m mod n) = integer_of_natural m mod integer_of_natural n" + by (simp add: integer_eq_iff zmod_int) + +lemma [code]: + "HOL.equal m n \ HOL.equal (integer_of_natural m) (integer_of_natural n)" + by (simp add: equal natural_eq_iff integer_eq_iff) + +lemma [code nbe]: + "HOL.equal n (n::natural) \ True" + by (simp add: equal) + +lemma [code]: + "m \ n \ (integer_of_natural m) \ integer_of_natural n" + by (simp add: less_eq_natural_def less_eq_integer_def) + +lemma [code]: + "m < n \ (integer_of_natural m) < integer_of_natural n" + by (simp add: less_natural_def less_integer_def) + +hide_const (open) Nat + + +code_reflect Code_Numeral_Types + datatypes natural = _ + functions integer_of_natural natural_of_integer + +end + diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Code_Target_Int.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Target_Int.thy Thu Nov 08 10:02:38 2012 +0100 @@ -0,0 +1,122 @@ +(* Title: HOL/Library/Code_Target_Int.thy + Author: Florian Haftmann, TU Muenchen +*) + +header {* Implementation of integer numbers by target-language integers *} + +theory Code_Target_Int +imports Main "~~/src/HOL/Library/Code_Numeral_Types" +begin + +code_datatype int_of_integer + +lemma [code, code del]: + "integer_of_int = integer_of_int" .. + +lemma [code]: + "integer_of_int (int_of_integer k) = k" + by (simp add: integer_eq_iff) + +lemma [code]: + "Int.Pos = int_of_integer \ integer_of_num" + by (simp add: integer_of_num_def fun_eq_iff) + +lemma [code]: + "Int.Neg = int_of_integer \ uminus \ integer_of_num" + by (simp add: integer_of_num_def fun_eq_iff) + +lemma [code_abbrev]: + "int_of_integer (Code_Numeral_Types.Pos k) = Int.Pos k" + by simp + +lemma [code_abbrev]: + "int_of_integer (Code_Numeral_Types.Neg k) = Int.Neg k" + by simp + +lemma [code]: + "0 = int_of_integer 0" + by simp + +lemma [code]: + "1 = int_of_integer 1" + by simp + +lemma [code]: + "k + l = int_of_integer (of_int k + of_int l)" + by simp + +lemma [code]: + "- k = int_of_integer (- of_int k)" + by simp + +lemma [code]: + "k - l = int_of_integer (of_int k - of_int l)" + by simp + +lemma [code]: + "Int.dup k = int_of_integer (Code_Numeral_Types.dup (of_int k))" + by simp + +lemma [code, code del]: + "Int.sub = Int.sub" .. + +lemma [code]: + "k * l = int_of_integer (of_int k * of_int l)" + by simp + +lemma [code]: + "pdivmod k l = map_pair int_of_integer int_of_integer + (Code_Numeral_Types.divmod_abs (of_int k) (of_int l))" + by (simp add: prod_eq_iff pdivmod_def) + +lemma [code]: + "k div l = int_of_integer (of_int k div of_int l)" + by simp + +lemma [code]: + "k mod l = int_of_integer (of_int k mod of_int l)" + by simp + +lemma [code]: + "HOL.equal k l = HOL.equal (of_int k :: integer) (of_int l)" + by (simp add: equal integer_eq_iff) + +lemma [code]: + "k \ l \ (of_int k :: integer) \ of_int l" + by (simp add: less_eq_int_def) + +lemma [code]: + "k < l \ (of_int k :: integer) < of_int l" + by (simp add: less_int_def) + +lemma (in ring_1) of_int_code: + "of_int k = (if k = 0 then 0 + else if k < 0 then - of_int (- k) + else let + (l, j) = divmod_int k 2; + l' = 2 * of_int l + in if j = 0 then l' else l' + 1)" +proof - + from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp + show ?thesis + by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int + of_int_add [symmetric]) (simp add: * mult_commute) +qed + +declare of_int_code [code] + +lemma [code]: + "nat = nat_of_integer \ of_int" + by (simp add: fun_eq_iff nat_of_integer_def) + +code_modulename SML + Code_Target_Int Arith + +code_modulename OCaml + Code_Target_Int Arith + +code_modulename Haskell + Code_Target_Int Arith + +end + diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Code_Target_Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Target_Nat.thy Thu Nov 08 10:02:38 2012 +0100 @@ -0,0 +1,137 @@ +(* Title: HOL/Library/Code_Target_Nat.thy + Author: Stefan Berghofer, Florian Haftmann, TU Muenchen +*) + +header {* Implementation of natural numbers by target-language integers *} + +theory Code_Target_Nat +imports Main Code_Numeral_Types Code_Binary_Nat +begin + +subsection {* Implementation for @{typ nat} *} + +definition Nat :: "integer \ nat" +where + "Nat = nat_of_integer" + +definition integer_of_nat :: "nat \ integer" +where + [code_abbrev]: "integer_of_nat = of_nat" + +lemma int_of_integer_integer_of_nat [simp]: + "int_of_integer (integer_of_nat n) = of_nat n" + by (simp add: integer_of_nat_def) + +lemma [code_unfold]: + "Int.nat (int_of_integer k) = nat_of_integer k" + by (simp add: nat_of_integer_def) + +lemma [code abstype]: + "Code_Target_Nat.Nat (integer_of_nat n) = n" + by (simp add: Nat_def integer_of_nat_def) + +lemma [code abstract]: + "integer_of_nat (nat_of_integer k) = max 0 k" + by (simp add: integer_of_nat_def) + +lemma [code_abbrev]: + "nat_of_integer (Code_Numeral_Types.Pos k) = nat_of_num k" + by (simp add: nat_of_integer_def nat_of_num_numeral) + +lemma [code abstract]: + "integer_of_nat (nat_of_num n) = integer_of_num n" + by (simp add: integer_eq_iff integer_of_num_def nat_of_num_numeral) + +lemma [code abstract]: + "integer_of_nat 0 = 0" + by (simp add: integer_eq_iff integer_of_nat_def) + +lemma [code abstract]: + "integer_of_nat 1 = 1" + by (simp add: integer_eq_iff integer_of_nat_def) + +lemma [code abstract]: + "integer_of_nat (m + n) = of_nat m + of_nat n" + by (simp add: integer_eq_iff integer_of_nat_def) + +lemma [code abstract]: + "integer_of_nat (Code_Binary_Nat.dup n) = Code_Numeral_Types.dup (of_nat n)" + by (simp add: integer_eq_iff Code_Binary_Nat.dup_def integer_of_nat_def) + +lemma [code, code del]: + "Code_Binary_Nat.sub = Code_Binary_Nat.sub" .. + +lemma [code abstract]: + "integer_of_nat (m - n) = max 0 (of_nat m - of_nat n)" + by (simp add: integer_eq_iff integer_of_nat_def) + +lemma [code abstract]: + "integer_of_nat (m * n) = of_nat m * of_nat n" + by (simp add: integer_eq_iff of_nat_mult integer_of_nat_def) + +lemma [code abstract]: + "integer_of_nat (m div n) = of_nat m div of_nat n" + by (simp add: integer_eq_iff zdiv_int integer_of_nat_def) + +lemma [code abstract]: + "integer_of_nat (m mod n) = of_nat m mod of_nat n" + by (simp add: integer_eq_iff zmod_int integer_of_nat_def) + +lemma [code]: + "Divides.divmod_nat m n = (m div n, m mod n)" + by (simp add: prod_eq_iff) + +lemma [code]: + "HOL.equal m n = HOL.equal (of_nat m :: integer) (of_nat n)" + by (simp add: equal integer_eq_iff) + +lemma [code]: + "m \ n \ (of_nat m :: integer) \ of_nat n" + by simp + +lemma [code]: + "m < n \ (of_nat m :: integer) < of_nat n" + by simp + +lemma num_of_nat_code [code]: + "num_of_nat = num_of_integer \ of_nat" + by (simp add: fun_eq_iff num_of_integer_def integer_of_nat_def) + +lemma (in semiring_1) of_nat_code: + "of_nat n = (if n = 0 then 0 + else let + (m, q) = divmod_nat n 2; + m' = 2 * of_nat m + in if q = 0 then m' else m' + 1)" +proof - + from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp + show ?thesis + by (simp add: Let_def divmod_nat_div_mod mod_2_not_eq_zero_eq_one_nat + of_nat_add [symmetric]) + (simp add: * mult_commute of_nat_mult add_commute) +qed + +declare of_nat_code [code] + +definition int_of_nat :: "nat \ int" where + [code_abbrev]: "int_of_nat = of_nat" + +lemma [code]: + "int_of_nat n = int_of_integer (of_nat n)" + by (simp add: int_of_nat_def) + +lemma [code abstract]: + "integer_of_nat (nat k) = max 0 (integer_of_int k)" + by (simp add: integer_of_nat_def of_int_of_nat max_def) + +code_modulename SML + Code_Target_Nat Arith + +code_modulename OCaml + Code_Target_Nat Arith + +code_modulename Haskell + Code_Target_Nat Arith + +end + diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Code_Target_Numeral.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Target_Numeral.thy Thu Nov 08 10:02:38 2012 +0100 @@ -0,0 +1,12 @@ +(* Title: HOL/Library/Code_Target_Numeral.thy + Author: Florian Haftmann, TU Muenchen +*) + +header {* Implementation of natural and integer numbers by target-language integers *} + +theory Code_Target_Numeral +imports Code_Target_Int Code_Target_Nat +begin + +end + diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Efficient_Nat.thy --- a/src/HOL/Library/Efficient_Nat.thy Wed Nov 07 20:48:04 2012 +0100 +++ b/src/HOL/Library/Efficient_Nat.thy Thu Nov 08 10:02:38 2012 +0100 @@ -5,7 +5,7 @@ header {* Implementation of natural numbers by target-language integers *} theory Efficient_Nat -imports Code_Nat Code_Integer Main +imports Code_Binary_Nat Code_Integer Main begin text {* @@ -217,14 +217,14 @@ (Scala infixl 7 "-") (Eval "Integer.max/ 0/ (_ -/ _)") -code_const Code_Nat.dup +code_const Code_Binary_Nat.dup (SML "IntInf.*/ (2,/ (_))") (OCaml "Big'_int.mult'_big'_int/ 2") (Haskell "!(2 * _)") (Scala "!(2 * _)") (Eval "!(2 * _)") -code_const Code_Nat.sub +code_const Code_Binary_Nat.sub (SML "!(raise/ Fail/ \"sub\")") (OCaml "failwith/ \"sub\"") (Haskell "error/ \"sub\"") @@ -302,3 +302,4 @@ hide_const (open) int end + diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/Library/Target_Numeral.thy --- a/src/HOL/Library/Target_Numeral.thy Wed Nov 07 20:48:04 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,742 +0,0 @@ -theory Target_Numeral -imports Main Code_Nat -begin - -subsection {* Type of target language numerals *} - -typedef int = "UNIV \ int set" - morphisms int_of of_int .. - -hide_type (open) int -hide_const (open) of_int - -lemma int_eq_iff: - "k = l \ int_of k = int_of l" - using int_of_inject [of k l] .. - -lemma int_eqI: - "int_of k = int_of l \ k = l" - using int_eq_iff [of k l] by simp - -lemma int_of_int [simp]: - "int_of (Target_Numeral.of_int k) = k" - using of_int_inverse [of k] by simp - -lemma of_int_of [simp]: - "Target_Numeral.of_int (int_of k) = k" - using int_of_inverse [of k] by simp - -hide_fact (open) int_eq_iff int_eqI - -instantiation Target_Numeral.int :: ring_1 -begin - -definition - "0 = Target_Numeral.of_int 0" - -lemma int_of_zero [simp]: - "int_of 0 = 0" - by (simp add: zero_int_def) - -definition - "1 = Target_Numeral.of_int 1" - -lemma int_of_one [simp]: - "int_of 1 = 1" - by (simp add: one_int_def) - -definition - "k + l = Target_Numeral.of_int (int_of k + int_of l)" - -lemma int_of_plus [simp]: - "int_of (k + l) = int_of k + int_of l" - by (simp add: plus_int_def) - -definition - "- k = Target_Numeral.of_int (- int_of k)" - -lemma int_of_uminus [simp]: - "int_of (- k) = - int_of k" - by (simp add: uminus_int_def) - -definition - "k - l = Target_Numeral.of_int (int_of k - int_of l)" - -lemma int_of_minus [simp]: - "int_of (k - l) = int_of k - int_of l" - by (simp add: minus_int_def) - -definition - "k * l = Target_Numeral.of_int (int_of k * int_of l)" - -lemma int_of_times [simp]: - "int_of (k * l) = int_of k * int_of l" - by (simp add: times_int_def) - -instance proof -qed (auto simp add: Target_Numeral.int_eq_iff algebra_simps) - -end - -lemma int_of_of_nat [simp]: - "int_of (of_nat n) = of_nat n" - by (induct n) simp_all - -definition nat_of :: "Target_Numeral.int \ nat" where - "nat_of k = Int.nat (int_of k)" - -lemma nat_of_of_nat [simp]: - "nat_of (of_nat n) = n" - by (simp add: nat_of_def) - -lemma int_of_of_int [simp]: - "int_of (of_int k) = k" - by (induct k) (simp_all, simp only: neg_numeral_def numeral_One int_of_uminus int_of_one) - -lemma of_int_of_int [simp, code_abbrev]: - "Target_Numeral.of_int = of_int" - by rule (simp add: Target_Numeral.int_eq_iff) - -lemma int_of_numeral [simp]: - "int_of (numeral k) = numeral k" - using int_of_of_int [of "numeral k"] by simp - -lemma int_of_neg_numeral [simp]: - "int_of (neg_numeral k) = neg_numeral k" - by (simp only: neg_numeral_def int_of_uminus) simp - -lemma int_of_sub [simp]: - "int_of (Num.sub k l) = Num.sub k l" - by (simp only: Num.sub_def int_of_minus int_of_numeral) - -instantiation Target_Numeral.int :: "{ring_div, equal, linordered_idom}" -begin - -definition - "k div l = of_int (int_of k div int_of l)" - -lemma int_of_div [simp]: - "int_of (k div l) = int_of k div int_of l" - by (simp add: div_int_def) - -definition - "k mod l = of_int (int_of k mod int_of l)" - -lemma int_of_mod [simp]: - "int_of (k mod l) = int_of k mod int_of l" - by (simp add: mod_int_def) - -definition - "\k\ = of_int \int_of k\" - -lemma int_of_abs [simp]: - "int_of \k\ = \int_of k\" - by (simp add: abs_int_def) - -definition - "sgn k = of_int (sgn (int_of k))" - -lemma int_of_sgn [simp]: - "int_of (sgn k) = sgn (int_of k)" - by (simp add: sgn_int_def) - -definition - "k \ l \ int_of k \ int_of l" - -definition - "k < l \ int_of k < int_of l" - -definition - "HOL.equal k l \ HOL.equal (int_of k) (int_of l)" - -instance proof -qed (auto simp add: Target_Numeral.int_eq_iff algebra_simps - less_eq_int_def less_int_def equal_int_def equal) - -end - -lemma int_of_min [simp]: - "int_of (min k l) = min (int_of k) (int_of l)" - by (simp add: min_def less_eq_int_def) - -lemma int_of_max [simp]: - "int_of (max k l) = max (int_of k) (int_of l)" - by (simp add: max_def less_eq_int_def) - -lemma of_nat_nat_of [simp]: - "of_nat (nat_of k) = max 0 k" - by (simp add: nat_of_def Target_Numeral.int_eq_iff less_eq_int_def max_def) - - -subsection {* Code theorems for target language numerals *} - -text {* Constructors *} - -definition Pos :: "num \ Target_Numeral.int" where - [simp, code_abbrev]: "Pos = numeral" - -definition Neg :: "num \ Target_Numeral.int" where - [simp, code_abbrev]: "Neg = neg_numeral" - -code_datatype "0::Target_Numeral.int" Pos Neg - - -text {* Auxiliary operations *} - -definition dup :: "Target_Numeral.int \ Target_Numeral.int" where - [simp]: "dup k = k + k" - -lemma dup_code [code]: - "dup 0 = 0" - "dup (Pos n) = Pos (Num.Bit0 n)" - "dup (Neg n) = Neg (Num.Bit0 n)" - unfolding Pos_def Neg_def neg_numeral_def - by (simp_all add: numeral_Bit0) - -definition sub :: "num \ num \ Target_Numeral.int" where - [simp]: "sub m n = numeral m - numeral n" - -lemma sub_code [code]: - "sub Num.One Num.One = 0" - "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" - "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" - "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" - "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" - "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" - "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" - "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" - "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" - unfolding sub_def dup_def numeral.simps Pos_def Neg_def - neg_numeral_def numeral_BitM - by (simp_all only: algebra_simps add.comm_neutral) - - -text {* Implementations *} - -lemma one_int_code [code, code_unfold]: - "1 = Pos Num.One" - by simp - -lemma plus_int_code [code]: - "k + 0 = (k::Target_Numeral.int)" - "0 + l = (l::Target_Numeral.int)" - "Pos m + Pos n = Pos (m + n)" - "Pos m + Neg n = sub m n" - "Neg m + Pos n = sub n m" - "Neg m + Neg n = Neg (m + n)" - by simp_all - -lemma uminus_int_code [code]: - "uminus 0 = (0::Target_Numeral.int)" - "uminus (Pos m) = Neg m" - "uminus (Neg m) = Pos m" - by simp_all - -lemma minus_int_code [code]: - "k - 0 = (k::Target_Numeral.int)" - "0 - l = uminus (l::Target_Numeral.int)" - "Pos m - Pos n = sub m n" - "Pos m - Neg n = Pos (m + n)" - "Neg m - Pos n = Neg (m + n)" - "Neg m - Neg n = sub n m" - by simp_all - -lemma times_int_code [code]: - "k * 0 = (0::Target_Numeral.int)" - "0 * l = (0::Target_Numeral.int)" - "Pos m * Pos n = Pos (m * n)" - "Pos m * Neg n = Neg (m * n)" - "Neg m * Pos n = Neg (m * n)" - "Neg m * Neg n = Pos (m * n)" - by simp_all - -definition divmod :: "Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int" where - "divmod k l = (k div l, k mod l)" - -lemma fst_divmod [simp]: - "fst (divmod k l) = k div l" - by (simp add: divmod_def) - -lemma snd_divmod [simp]: - "snd (divmod k l) = k mod l" - by (simp add: divmod_def) - -definition divmod_abs :: "Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int" where - "divmod_abs k l = (\k\ div \l\, \k\ mod \l\)" - -lemma fst_divmod_abs [simp]: - "fst (divmod_abs k l) = \k\ div \l\" - by (simp add: divmod_abs_def) - -lemma snd_divmod_abs [simp]: - "snd (divmod_abs k l) = \k\ mod \l\" - by (simp add: divmod_abs_def) - -lemma divmod_abs_terminate_code [code]: - "divmod_abs (Neg k) (Neg l) = divmod_abs (Pos k) (Pos l)" - "divmod_abs (Neg k) (Pos l) = divmod_abs (Pos k) (Pos l)" - "divmod_abs (Pos k) (Neg l) = divmod_abs (Pos k) (Pos l)" - "divmod_abs j 0 = (0, \j\)" - "divmod_abs 0 j = (0, 0)" - by (simp_all add: prod_eq_iff) - -lemma divmod_abs_rec_code [code]: - "divmod_abs (Pos k) (Pos l) = - (let j = sub k l in - if j < 0 then (0, Pos k) - else let (q, r) = divmod_abs j (Pos l) in (q + 1, r))" - by (auto simp add: prod_eq_iff Target_Numeral.int_eq_iff Let_def prod_case_beta - sub_non_negative sub_negative div_pos_pos_trivial mod_pos_pos_trivial div_pos_geq mod_pos_geq) - -lemma divmod_code [code]: "divmod k l = - (if k = 0 then (0, 0) else if l = 0 then (0, k) else - (apsnd \ times \ sgn) l (if sgn k = sgn l - then divmod_abs k l - else (let (r, s) = divmod_abs k l in - if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" -proof - - have aux1: "\k l::int. sgn k = sgn l \ k = 0 \ l = 0 \ 0 < l \ 0 < k \ l < 0 \ k < 0" - by (auto simp add: sgn_if) - have aux2: "\q::int. - int_of k = int_of l * q \ int_of k = int_of l * - q" by auto - show ?thesis - by (simp add: prod_eq_iff Target_Numeral.int_eq_iff prod_case_beta aux1) - (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right aux2) -qed - -lemma div_int_code [code]: - "k div l = fst (divmod k l)" - by simp - -lemma div_mod_code [code]: - "k mod l = snd (divmod k l)" - by simp - -lemma equal_int_code [code]: - "HOL.equal 0 (0::Target_Numeral.int) \ True" - "HOL.equal 0 (Pos l) \ False" - "HOL.equal 0 (Neg l) \ False" - "HOL.equal (Pos k) 0 \ False" - "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" - "HOL.equal (Pos k) (Neg l) \ False" - "HOL.equal (Neg k) 0 \ False" - "HOL.equal (Neg k) (Pos l) \ False" - "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" - by (simp_all add: equal Target_Numeral.int_eq_iff) - -lemma equal_int_refl [code nbe]: - "HOL.equal (k::Target_Numeral.int) k \ True" - by (fact equal_refl) - -lemma less_eq_int_code [code]: - "0 \ (0::Target_Numeral.int) \ True" - "0 \ Pos l \ True" - "0 \ Neg l \ False" - "Pos k \ 0 \ False" - "Pos k \ Pos l \ k \ l" - "Pos k \ Neg l \ False" - "Neg k \ 0 \ True" - "Neg k \ Pos l \ True" - "Neg k \ Neg l \ l \ k" - by (simp_all add: less_eq_int_def) - -lemma less_int_code [code]: - "0 < (0::Target_Numeral.int) \ False" - "0 < Pos l \ True" - "0 < Neg l \ False" - "Pos k < 0 \ False" - "Pos k < Pos l \ k < l" - "Pos k < Neg l \ False" - "Neg k < 0 \ True" - "Neg k < Pos l \ True" - "Neg k < Neg l \ l < k" - by (simp_all add: less_int_def) - -lemma nat_of_code [code]: - "nat_of (Neg k) = 0" - "nat_of 0 = 0" - "nat_of (Pos k) = nat_of_num k" - by (simp_all add: nat_of_def nat_of_num_numeral) - -lemma int_of_code [code]: - "int_of (Neg k) = neg_numeral k" - "int_of 0 = 0" - "int_of (Pos k) = numeral k" - by simp_all - -lemma of_int_code [code]: - "Target_Numeral.of_int (Int.Neg k) = neg_numeral k" - "Target_Numeral.of_int 0 = 0" - "Target_Numeral.of_int (Int.Pos k) = numeral k" - by simp_all - -definition num_of_int :: "Target_Numeral.int \ num" where - "num_of_int = num_of_nat \ nat_of" - -lemma num_of_int_code [code]: - "num_of_int k = (if k \ 1 then Num.One - else let - (l, j) = divmod k 2; - l' = num_of_int l + num_of_int l - in if j = 0 then l' else l' + Num.One)" -proof - - { - assume "int_of k mod 2 = 1" - then have "nat (int_of k mod 2) = nat 1" by simp - moreover assume *: "1 < int_of k" - ultimately have **: "nat (int_of k) mod 2 = 1" by (simp add: nat_mod_distrib) - have "num_of_nat (nat (int_of k)) = - num_of_nat (2 * (nat (int_of k) div 2) + nat (int_of k) mod 2)" - by simp - then have "num_of_nat (nat (int_of k)) = - num_of_nat (nat (int_of k) div 2 + nat (int_of k) div 2 + nat (int_of k) mod 2)" - by (simp add: mult_2) - with ** have "num_of_nat (nat (int_of k)) = - num_of_nat (nat (int_of k) div 2 + nat (int_of k) div 2 + 1)" - by simp - } - note aux = this - show ?thesis - by (auto simp add: num_of_int_def nat_of_def Let_def prod_case_beta - not_le Target_Numeral.int_eq_iff less_eq_int_def - nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib - mult_2 [where 'a=nat] aux add_One) -qed - -hide_const (open) int_of nat_of Pos Neg sub dup divmod_abs num_of_int - - -subsection {* Serializer setup for target language numerals *} - -code_type Target_Numeral.int - (SML "IntInf.int") - (OCaml "Big'_int.big'_int") - (Haskell "Integer") - (Scala "BigInt") - (Eval "int") - -code_instance Target_Numeral.int :: equal - (Haskell -) - -code_const "0::Target_Numeral.int" - (SML "0") - (OCaml "Big'_int.zero'_big'_int") - (Haskell "0") - (Scala "BigInt(0)") - -setup {* - fold (Numeral.add_code @{const_name Target_Numeral.Pos} - false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] -*} - -setup {* - fold (Numeral.add_code @{const_name Target_Numeral.Neg} - true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] -*} - -code_const "plus :: Target_Numeral.int \ _ \ _" - (SML "IntInf.+ ((_), (_))") - (OCaml "Big'_int.add'_big'_int") - (Haskell infixl 6 "+") - (Scala infixl 7 "+") - (Eval infixl 8 "+") - -code_const "uminus :: Target_Numeral.int \ _" - (SML "IntInf.~") - (OCaml "Big'_int.minus'_big'_int") - (Haskell "negate") - (Scala "!(- _)") - (Eval "~/ _") - -code_const "minus :: Target_Numeral.int \ _" - (SML "IntInf.- ((_), (_))") - (OCaml "Big'_int.sub'_big'_int") - (Haskell infixl 6 "-") - (Scala infixl 7 "-") - (Eval infixl 8 "-") - -code_const Target_Numeral.dup - (SML "IntInf.*/ (2,/ (_))") - (OCaml "Big'_int.mult'_big'_int/ 2") - (Haskell "!(2 * _)") - (Scala "!(2 * _)") - (Eval "!(2 * _)") - -code_const Target_Numeral.sub - (SML "!(raise/ Fail/ \"sub\")") - (OCaml "failwith/ \"sub\"") - (Haskell "error/ \"sub\"") - (Scala "!sys.error(\"sub\")") - -code_const "times :: Target_Numeral.int \ _ \ _" - (SML "IntInf.* ((_), (_))") - (OCaml "Big'_int.mult'_big'_int") - (Haskell infixl 7 "*") - (Scala infixl 8 "*") - (Eval infixl 9 "*") - -code_const Target_Numeral.divmod_abs - (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") - (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") - (Haskell "divMod/ (abs _)/ (abs _)") - (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") - (Eval "Integer.div'_mod/ (abs _)/ (abs _)") - -code_const "HOL.equal :: Target_Numeral.int \ _ \ bool" - (SML "!((_ : IntInf.int) = _)") - (OCaml "Big'_int.eq'_big'_int") - (Haskell infix 4 "==") - (Scala infixl 5 "==") - (Eval infixl 6 "=") - -code_const "less_eq :: Target_Numeral.int \ _ \ bool" - (SML "IntInf.<= ((_), (_))") - (OCaml "Big'_int.le'_big'_int") - (Haskell infix 4 "<=") - (Scala infixl 4 "<=") - (Eval infixl 6 "<=") - -code_const "less :: Target_Numeral.int \ _ \ bool" - (SML "IntInf.< ((_), (_))") - (OCaml "Big'_int.lt'_big'_int") - (Haskell infix 4 "<") - (Scala infixl 4 "<") - (Eval infixl 6 "<") - -ML {* -structure Target_Numeral = -struct - -val T = @{typ "Target_Numeral.int"}; - -end; -*} - -code_reserved Eval Target_Numeral - -code_const "Code_Evaluation.term_of \ Target_Numeral.int \ term" - (Eval "HOLogic.mk'_number/ Target'_Numeral.T") - -code_modulename SML - Target_Numeral Arith - -code_modulename OCaml - Target_Numeral Arith - -code_modulename Haskell - Target_Numeral Arith - - -subsection {* Implementation for @{typ int} *} - -code_datatype Target_Numeral.int_of - -lemma [code, code del]: - "Target_Numeral.of_int = Target_Numeral.of_int" .. - -lemma [code]: - "Target_Numeral.of_int (Target_Numeral.int_of k) = k" - by (simp add: Target_Numeral.int_eq_iff) - -declare Int.Pos_def [code] - -lemma [code_abbrev]: - "Target_Numeral.int_of (Target_Numeral.Pos k) = Int.Pos k" - by simp - -declare Int.Neg_def [code] - -lemma [code_abbrev]: - "Target_Numeral.int_of (Target_Numeral.Neg k) = Int.Neg k" - by simp - -lemma [code]: - "0 = Target_Numeral.int_of 0" - by simp - -lemma [code]: - "1 = Target_Numeral.int_of 1" - by simp - -lemma [code]: - "k + l = Target_Numeral.int_of (of_int k + of_int l)" - by simp - -lemma [code]: - "- k = Target_Numeral.int_of (- of_int k)" - by simp - -lemma [code]: - "k - l = Target_Numeral.int_of (of_int k - of_int l)" - by simp - -lemma [code]: - "Int.dup k = Target_Numeral.int_of (Target_Numeral.dup (of_int k))" - by simp - -lemma [code, code del]: - "Int.sub = Int.sub" .. - -lemma [code]: - "k * l = Target_Numeral.int_of (of_int k * of_int l)" - by simp - -lemma [code]: - "pdivmod k l = map_pair Target_Numeral.int_of Target_Numeral.int_of - (Target_Numeral.divmod_abs (of_int k) (of_int l))" - by (simp add: prod_eq_iff pdivmod_def) - -lemma [code]: - "k div l = Target_Numeral.int_of (of_int k div of_int l)" - by simp - -lemma [code]: - "k mod l = Target_Numeral.int_of (of_int k mod of_int l)" - by simp - -lemma [code]: - "HOL.equal k l = HOL.equal (of_int k :: Target_Numeral.int) (of_int l)" - by (simp add: equal Target_Numeral.int_eq_iff) - -lemma [code]: - "k \ l \ (of_int k :: Target_Numeral.int) \ of_int l" - by (simp add: less_eq_int_def) - -lemma [code]: - "k < l \ (of_int k :: Target_Numeral.int) < of_int l" - by (simp add: less_int_def) - -lemma (in ring_1) of_int_code: - "of_int k = (if k = 0 then 0 - else if k < 0 then - of_int (- k) - else let - (l, j) = divmod_int k 2; - l' = 2 * of_int l - in if j = 0 then l' else l' + 1)" -proof - - from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp - show ?thesis - by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int - of_int_add [symmetric]) (simp add: * mult_commute) -qed - -declare of_int_code [code] - - -subsection {* Implementation for @{typ nat} *} - -definition Nat :: "Target_Numeral.int \ nat" where - "Nat = Target_Numeral.nat_of" - -definition of_nat :: "nat \ Target_Numeral.int" where - [code_abbrev]: "of_nat = Nat.of_nat" - -hide_const (open) of_nat Nat - -lemma [code_unfold]: - "Int.nat (Target_Numeral.int_of k) = Target_Numeral.nat_of k" - by (simp add: nat_of_def) - -lemma int_of_nat [simp]: - "Target_Numeral.int_of (Target_Numeral.of_nat n) = of_nat n" - by (simp add: of_nat_def) - -lemma [code abstype]: - "Target_Numeral.Nat (Target_Numeral.of_nat n) = n" - by (simp add: Nat_def nat_of_def) - -lemma [code abstract]: - "Target_Numeral.of_nat (Target_Numeral.nat_of k) = max 0 k" - by (simp add: of_nat_def) - -lemma [code_abbrev]: - "nat (Int.Pos k) = nat_of_num k" - by (simp add: nat_of_num_numeral) - -lemma [code abstract]: - "Target_Numeral.of_nat 0 = 0" - by (simp add: Target_Numeral.int_eq_iff) - -lemma [code abstract]: - "Target_Numeral.of_nat 1 = 1" - by (simp add: Target_Numeral.int_eq_iff) - -lemma [code abstract]: - "Target_Numeral.of_nat (m + n) = of_nat m + of_nat n" - by (simp add: Target_Numeral.int_eq_iff) - -lemma [code abstract]: - "Target_Numeral.of_nat (Code_Nat.dup n) = Target_Numeral.dup (of_nat n)" - by (simp add: Target_Numeral.int_eq_iff Code_Nat.dup_def) - -lemma [code, code del]: - "Code_Nat.sub = Code_Nat.sub" .. - -lemma [code abstract]: - "Target_Numeral.of_nat (m - n) = max 0 (of_nat m - of_nat n)" - by (simp add: Target_Numeral.int_eq_iff) - -lemma [code abstract]: - "Target_Numeral.of_nat (m * n) = of_nat m * of_nat n" - by (simp add: Target_Numeral.int_eq_iff of_nat_mult) - -lemma [code abstract]: - "Target_Numeral.of_nat (m div n) = of_nat m div of_nat n" - by (simp add: Target_Numeral.int_eq_iff zdiv_int) - -lemma [code abstract]: - "Target_Numeral.of_nat (m mod n) = of_nat m mod of_nat n" - by (simp add: Target_Numeral.int_eq_iff zmod_int) - -lemma [code]: - "Divides.divmod_nat m n = (m div n, m mod n)" - by (simp add: prod_eq_iff) - -lemma [code]: - "HOL.equal m n = HOL.equal (of_nat m :: Target_Numeral.int) (of_nat n)" - by (simp add: equal Target_Numeral.int_eq_iff) - -lemma [code]: - "m \ n \ (of_nat m :: Target_Numeral.int) \ of_nat n" - by (simp add: less_eq_int_def) - -lemma [code]: - "m < n \ (of_nat m :: Target_Numeral.int) < of_nat n" - by (simp add: less_int_def) - -lemma num_of_nat_code [code]: - "num_of_nat = Target_Numeral.num_of_int \ of_nat" - by (simp add: fun_eq_iff num_of_int_def of_nat_def) - -lemma (in semiring_1) of_nat_code: - "of_nat n = (if n = 0 then 0 - else let - (m, q) = divmod_nat n 2; - m' = 2 * of_nat m - in if q = 0 then m' else m' + 1)" -proof - - from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp - show ?thesis - by (simp add: Let_def divmod_nat_div_mod mod_2_not_eq_zero_eq_one_nat - of_nat_add [symmetric]) - (simp add: * mult_commute of_nat_mult add_commute) -qed - -declare of_nat_code [code] - -text {* Conversions between @{typ nat} and @{typ int} *} - -definition int :: "nat \ int" where - [code_abbrev]: "int = of_nat" - -hide_const (open) int - -lemma [code]: - "Target_Numeral.int n = Target_Numeral.int_of (of_nat n)" - by (simp add: int_def) - -lemma [code abstract]: - "Target_Numeral.of_nat (nat k) = max 0 (Target_Numeral.of_int k)" - by (simp add: of_nat_def of_int_of_nat max_def) - -end - diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/ROOT --- a/src/HOL/ROOT Wed Nov 07 20:48:04 2012 +0100 +++ b/src/HOL/ROOT Thu Nov 08 10:02:38 2012 +0100 @@ -48,7 +48,7 @@ Efficient_Nat (* Code_Prolog FIXME cf. 76965c356d2a *) Code_Real_Approx_By_Float - Target_Numeral + Code_Target_Numeral Refute theories [condition = ISABELLE_FULL_TEST] Sum_of_Squares_Remote @@ -415,7 +415,7 @@ options [timeout = 3600, condition = ISABELLE_POLYML] theories [document = false] "~~/src/HOL/Library/State_Monad" - Code_Nat_examples + Code_Binary_Nat_examples "~~/src/HOL/Library/FuncSet" Eval_Examples Normalization_by_Evaluation diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/ex/Code_Binary_Nat_examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Code_Binary_Nat_examples.thy Thu Nov 08 10:02:38 2012 +0100 @@ -0,0 +1,57 @@ +(* Title: HOL/ex/Code_Binary_Nat_examples.thy + Author: Florian Haftmann, TU Muenchen +*) + +header {* Simple examples for natural numbers implemented in binary representation theory. *} + +theory Code_Binary_Nat_examples +imports Complex_Main "~~/src/HOL/Library/Efficient_Nat" +begin + +fun to_n :: "nat \ nat list" +where + "to_n 0 = []" +| "to_n (Suc 0) = []" +| "to_n (Suc (Suc 0)) = []" +| "to_n (Suc n) = n # to_n n" + +definition naive_prime :: "nat \ bool" +where + "naive_prime n \ n \ 2 \ filter (\m. n mod m = 0) (to_n n) = []" + +primrec fac :: "nat \ nat" +where + "fac 0 = 1" +| "fac (Suc n) = Suc n * fac n" + +primrec harmonic :: "nat \ rat" +where + "harmonic 0 = 0" +| "harmonic (Suc n) = 1 / of_nat (Suc n) + harmonic n" + +lemma "harmonic 200 \ 5" + by eval + +lemma "(let (q, r) = quotient_of (harmonic 8) in q div r) \ 2" + by normalization + +lemma "naive_prime 89" + by eval + +lemma "naive_prime 89" + by normalization + +lemma "\ naive_prime 87" + by eval + +lemma "\ naive_prime 87" + by normalization + +lemma "fac 10 > 3000000" + by eval + +lemma "fac 10 > 3000000" + by normalization + +end + diff -r 286dfcab9833 -r 28f3263d4d1b src/HOL/ex/Code_Nat_examples.thy --- a/src/HOL/ex/Code_Nat_examples.thy Wed Nov 07 20:48:04 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,57 +0,0 @@ -(* Title: HOL/ex/Code_Nat_examples.thy - Author: Florian Haftmann, TU Muenchen -*) - -header {* Simple examples for Code\_Numeral\_Nat theory. *} - -theory Code_Nat_examples -imports Complex_Main "~~/src/HOL/Library/Efficient_Nat" -begin - -fun to_n :: "nat \ nat list" -where - "to_n 0 = []" -| "to_n (Suc 0) = []" -| "to_n (Suc (Suc 0)) = []" -| "to_n (Suc n) = n # to_n n" - -definition naive_prime :: "nat \ bool" -where - "naive_prime n \ n \ 2 \ filter (\m. n mod m = 0) (to_n n) = []" - -primrec fac :: "nat \ nat" -where - "fac 0 = 1" -| "fac (Suc n) = Suc n * fac n" - -primrec harmonic :: "nat \ rat" -where - "harmonic 0 = 0" -| "harmonic (Suc n) = 1 / of_nat (Suc n) + harmonic n" - -lemma "harmonic 200 \ 5" - by eval - -lemma "(let (q, r) = quotient_of (harmonic 8) in q div r) \ 2" - by normalization - -lemma "naive_prime 89" - by eval - -lemma "naive_prime 89" - by normalization - -lemma "\ naive_prime 87" - by eval - -lemma "\ naive_prime 87" - by normalization - -lemma "fac 10 > 3000000" - by eval - -lemma "fac 10 > 3000000" - by normalization - -end -