haftmann@29752: (* Author: Florian Haftmann, TU Muenchen *) haftmann@24999: haftmann@31205: header {* Type of target language numerals *} haftmann@24999: haftmann@31205: theory Code_Numeral huffman@48116: imports Nat_Transfer Divides haftmann@24999: begin haftmann@24999: haftmann@24999: text {* haftmann@31205: Code numerals are isomorphic to HOL @{typ nat} but haftmann@31205: mapped to target-language builtin numerals. haftmann@24999: *} haftmann@24999: haftmann@31205: subsection {* Datatype of target language numerals *} haftmann@24999: haftmann@31205: typedef (open) code_numeral = "UNIV \ nat set" wenzelm@46567: morphisms nat_of of_nat .. haftmann@24999: haftmann@29752: lemma of_nat_nat_of [simp]: haftmann@29752: "of_nat (nat_of k) = k" haftmann@29752: by (rule nat_of_inverse) haftmann@25967: haftmann@29752: lemma nat_of_of_nat [simp]: haftmann@29752: "nat_of (of_nat n) = n" haftmann@29752: by (rule of_nat_inverse) (rule UNIV_I) haftmann@24999: haftmann@28708: lemma [measure_function]: haftmann@29752: "is_measure nat_of" by (rule is_measure_trivial) haftmann@28708: haftmann@31205: lemma code_numeral: haftmann@31205: "(\n\code_numeral. PROP P n) \ (\n\nat. PROP P (of_nat n))" haftmann@24999: proof haftmann@25767: fix n :: nat haftmann@31205: assume "\n\code_numeral. PROP P n" haftmann@29752: then show "PROP P (of_nat n)" . haftmann@24999: next haftmann@31205: fix n :: code_numeral haftmann@29752: assume "\n\nat. PROP P (of_nat n)" haftmann@29752: then have "PROP P (of_nat (nat_of n))" . haftmann@25767: then show "PROP P n" by simp haftmann@24999: qed haftmann@24999: haftmann@31205: lemma code_numeral_case: haftmann@29752: assumes "\n. k = of_nat n \ P" haftmann@26140: shows P haftmann@29752: by (rule assms [of "nat_of k"]) simp haftmann@26140: haftmann@31205: lemma code_numeral_induct_raw: haftmann@29752: assumes "\n. P (of_nat n)" haftmann@26140: shows "P k" haftmann@26140: proof - haftmann@29752: from assms have "P (of_nat (nat_of k))" . haftmann@26140: then show ?thesis by simp haftmann@26140: qed haftmann@26140: haftmann@29752: lemma nat_of_inject [simp]: haftmann@29752: "nat_of k = nat_of l \ k = l" haftmann@29752: by (rule nat_of_inject) haftmann@26140: haftmann@29752: lemma of_nat_inject [simp]: haftmann@29752: "of_nat n = of_nat m \ n = m" haftmann@29752: by (rule of_nat_inject) (rule UNIV_I)+ haftmann@26140: haftmann@31205: instantiation code_numeral :: zero haftmann@26140: begin haftmann@26140: haftmann@28562: definition [simp, code del]: haftmann@29752: "0 = of_nat 0" haftmann@26140: haftmann@26140: instance .. haftmann@26140: haftmann@26140: end haftmann@26140: huffman@47418: definition Suc where [simp]: huffman@47418: "Suc k = of_nat (Nat.Suc (nat_of k))" haftmann@26140: huffman@47418: rep_datatype "0 \ code_numeral" Suc haftmann@26140: proof - haftmann@31205: fix P :: "code_numeral \ bool" haftmann@31205: fix k :: code_numeral haftmann@29752: assume "P 0" then have init: "P (of_nat 0)" by simp huffman@47418: assume "\k. P k \ P (Suc k)" huffman@47418: then have "\n. P (of_nat n) \ P (Suc (of_nat n))" . huffman@47418: then have step: "\n. P (of_nat n) \ P (of_nat (Nat.Suc n))" by simp haftmann@29752: from init step have "P (of_nat (nat_of k))" berghofe@34915: by (induct ("nat_of k")) simp_all haftmann@26140: then show "P k" by simp haftmann@27104: qed simp_all haftmann@26140: haftmann@31205: declare code_numeral_case [case_names nat, cases type: code_numeral] haftmann@31205: declare code_numeral.induct [case_names nat, induct type: code_numeral] haftmann@26140: haftmann@31205: lemma code_numeral_decr [termination_simp]: huffman@47418: "k \ of_nat 0 \ nat_of k - Nat.Suc 0 < nat_of k" haftmann@30245: by (cases k) simp haftmann@30245: haftmann@30245: lemma [simp, code]: haftmann@31205: "code_numeral_size = nat_of" haftmann@26140: proof (rule ext) haftmann@26140: fix k haftmann@31205: have "code_numeral_size k = nat_size (nat_of k)" huffman@47418: by (induct k rule: code_numeral.induct) (simp_all del: zero_code_numeral_def Suc_def, simp_all) berghofe@34915: also have "nat_size (nat_of k) = nat_of k" by (induct ("nat_of k")) simp_all haftmann@31205: finally show "code_numeral_size k = nat_of k" . haftmann@26140: qed haftmann@26140: haftmann@30245: lemma [simp, code]: haftmann@29752: "size = nat_of" haftmann@26140: proof (rule ext) haftmann@26140: fix k haftmann@29752: show "size k = nat_of k" huffman@47418: by (induct k) (simp_all del: zero_code_numeral_def Suc_def, simp_all) haftmann@26140: qed haftmann@26140: haftmann@31205: lemmas [code del] = code_numeral.recs code_numeral.cases haftmann@30245: haftmann@28562: lemma [code]: haftmann@39086: "HOL.equal k l \ HOL.equal (nat_of k) (nat_of l)" haftmann@39086: by (cases k, cases l) (simp add: equal) haftmann@24999: haftmann@28351: lemma [code nbe]: haftmann@39086: "HOL.equal (k::code_numeral) k \ True" haftmann@39086: by (rule equal_refl) haftmann@28351: haftmann@24999: haftmann@24999: subsection {* Basic arithmetic *} haftmann@24999: haftmann@35028: instantiation code_numeral :: "{minus, linordered_semidom, semiring_div, linorder}" haftmann@25767: begin haftmann@24999: haftmann@28708: definition [simp, code del]: haftmann@31205: "(1\code_numeral) = of_nat 1" haftmann@28708: haftmann@28708: definition [simp, code del]: haftmann@29752: "n + m = of_nat (nat_of n + nat_of m)" haftmann@28708: haftmann@28708: definition [simp, code del]: haftmann@29752: "n - m = of_nat (nat_of n - nat_of m)" haftmann@28708: haftmann@28708: definition [simp, code del]: haftmann@29752: "n * m = of_nat (nat_of n * nat_of m)" haftmann@28708: haftmann@28708: definition [simp, code del]: haftmann@29752: "n div m = of_nat (nat_of n div nat_of m)" haftmann@28708: haftmann@28708: definition [simp, code del]: haftmann@29752: "n mod m = of_nat (nat_of n mod nat_of m)" haftmann@28708: haftmann@28708: definition [simp, code del]: haftmann@29752: "n \ m \ nat_of n \ nat_of m" haftmann@28708: haftmann@28708: definition [simp, code del]: haftmann@29752: "n < m \ nat_of n < nat_of m" haftmann@28708: haftmann@29752: instance proof haftmann@33335: qed (auto simp add: code_numeral left_distrib intro: mult_commute) haftmann@28708: haftmann@28708: end haftmann@28708: huffman@47978: lemma nat_of_numeral [simp]: "nat_of (numeral k) = numeral k" huffman@47978: by (induct k rule: num_induct) (simp_all add: numeral_inc) haftmann@46899: huffman@47978: definition Num :: "num \ code_numeral" huffman@47978: where [simp, code_abbrev]: "Num = numeral" huffman@47978: huffman@47978: code_datatype "0::code_numeral" Num haftmann@25767: haftmann@46899: lemma one_code_numeral_code [code]: haftmann@31205: "(1\code_numeral) = Numeral1" huffman@47978: by simp haftmann@46899: haftmann@46899: lemma [code_abbrev]: "Numeral1 = (1\code_numeral)" haftmann@31205: using one_code_numeral_code .. haftmann@25767: haftmann@31205: lemma plus_code_numeral_code [code nbe]: haftmann@29752: "of_nat n + of_nat m = of_nat (n + m)" haftmann@24999: by simp haftmann@24999: huffman@47978: lemma minus_code_numeral_code [code nbe]: huffman@47978: "of_nat n - of_nat m = of_nat (n - m)" haftmann@28708: by simp haftmann@28708: haftmann@31205: lemma times_code_numeral_code [code nbe]: haftmann@29752: "of_nat n * of_nat m = of_nat (n * m)" haftmann@25767: by simp haftmann@25335: haftmann@31205: lemma less_eq_code_numeral_code [code nbe]: haftmann@29752: "of_nat n \ of_nat m \ n \ m" haftmann@25767: by simp haftmann@24999: haftmann@31205: lemma less_code_numeral_code [code nbe]: haftmann@29752: "of_nat n < of_nat m \ n < m" haftmann@25767: by simp haftmann@24999: haftmann@31259: lemma code_numeral_zero_minus_one: haftmann@31259: "(0::code_numeral) - 1 = 0" haftmann@31259: by simp haftmann@31259: haftmann@31259: lemma Suc_code_numeral_minus_one: huffman@47418: "Suc n - 1 = n" haftmann@31259: by simp haftmann@26140: haftmann@29752: lemma of_nat_code [code]: haftmann@29752: "of_nat = Nat.of_nat" haftmann@25918: proof haftmann@25918: fix n :: nat haftmann@29752: have "Nat.of_nat n = of_nat n" haftmann@25918: by (induct n) simp_all haftmann@29752: then show "of_nat n = Nat.of_nat n" haftmann@25918: by (rule sym) haftmann@25918: qed haftmann@25918: haftmann@31205: lemma code_numeral_not_eq_zero: "i \ of_nat 0 \ i \ 1" haftmann@25928: by (cases i) auto haftmann@25928: haftmann@31205: definition nat_of_aux :: "code_numeral \ nat \ nat" where haftmann@29752: "nat_of_aux i n = nat_of i + n" haftmann@25928: haftmann@29752: lemma nat_of_aux_code [code]: huffman@47418: "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Nat.Suc n))" haftmann@31205: by (auto simp add: nat_of_aux_def code_numeral_not_eq_zero) haftmann@25928: haftmann@29752: lemma nat_of_code [code]: haftmann@29752: "nat_of i = nat_of_aux i 0" haftmann@29752: by (simp add: nat_of_aux_def) haftmann@25918: huffman@47418: definition div_mod :: "code_numeral \ code_numeral \ code_numeral \ code_numeral" where huffman@47418: [code del]: "div_mod n m = (n div m, n mod m)" haftmann@26009: haftmann@28562: lemma [code]: huffman@47418: "div_mod n m = (if m = 0 then (0, n) else (n div m, n mod m))" huffman@47418: unfolding div_mod_def by auto haftmann@26009: haftmann@28562: lemma [code]: huffman@47418: "n div m = fst (div_mod n m)" huffman@47418: unfolding div_mod_def by simp haftmann@26009: haftmann@28562: lemma [code]: huffman@47418: "n mod m = snd (div_mod n m)" huffman@47418: unfolding div_mod_def by simp haftmann@26009: haftmann@31205: definition int_of :: "code_numeral \ int" where haftmann@31192: "int_of = Nat.of_nat o nat_of" haftmann@26009: haftmann@31192: lemma int_of_code [code]: haftmann@31192: "int_of k = (if k = 0 then 0 haftmann@31192: else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))" haftmann@33335: proof - haftmann@33335: have "(nat_of k div 2) * 2 + nat_of k mod 2 = nat_of k" haftmann@33335: by (rule mod_div_equality) haftmann@33335: then have "int ((nat_of k div 2) * 2 + nat_of k mod 2) = int (nat_of k)" haftmann@33335: by simp haftmann@33335: then have "int (nat_of k) = int (nat_of k div 2) * 2 + int (nat_of k mod 2)" huffman@45692: unfolding of_nat_mult of_nat_add by simp haftmann@33335: then show ?thesis by (auto simp add: int_of_def mult_ac) haftmann@33335: qed haftmann@28708: haftmann@28708: huffman@47978: hide_const (open) of_nat nat_of Suc int_of huffman@47418: haftmann@28708: haftmann@28228: subsection {* Code generator setup *} haftmann@24999: haftmann@38195: text {* Implementation of code numerals by bounded integers *} haftmann@25767: haftmann@31205: code_type code_numeral haftmann@24999: (SML "int") haftmann@31377: (OCaml "Big'_int.big'_int") haftmann@38185: (Haskell "Integer") haftmann@38195: (Scala "BigInt") haftmann@24999: haftmann@39086: code_instance code_numeral :: equal haftmann@24999: (Haskell -) haftmann@24999: haftmann@24999: setup {* huffman@47978: Numeral.add_code @{const_name Num} haftmann@38195: false Code_Printer.literal_naive_numeral "SML" huffman@47978: #> fold (Numeral.add_code @{const_name Num} haftmann@38195: false Code_Printer.literal_numeral) ["OCaml", "Haskell", "Scala"] haftmann@24999: *} haftmann@24999: haftmann@25918: code_reserved SML Int int haftmann@38195: code_reserved Eval Integer haftmann@24999: huffman@47978: code_const "0::code_numeral" huffman@47978: (SML "0") huffman@47978: (OCaml "Big'_int.zero'_big'_int") huffman@47978: (Haskell "0") huffman@47978: (Scala "BigInt(0)") huffman@47978: huffman@47418: code_const "plus \ code_numeral \ code_numeral \ code_numeral" haftmann@25928: (SML "Int.+/ ((_),/ (_))") haftmann@31377: (OCaml "Big'_int.add'_big'_int") haftmann@24999: (Haskell infixl 6 "+") haftmann@34886: (Scala infixl 7 "+") haftmann@38195: (Eval infixl 8 "+") haftmann@24999: huffman@47978: code_const "minus \ code_numeral \ code_numeral \ code_numeral" huffman@47978: (SML "Int.max/ (0 : int,/ Int.-/ ((_),/ (_)))") huffman@47978: (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)") haftmann@49446: (Haskell "Prelude.max/ (0 :: Integer)/ (_/ -/ _)") haftmann@34886: (Scala "!(_/ -/ _).max(0)") huffman@47978: (Eval "Integer.max/ 0/ (_/ -/ _)") haftmann@24999: huffman@47418: code_const "times \ code_numeral \ code_numeral \ code_numeral" haftmann@25928: (SML "Int.*/ ((_),/ (_))") haftmann@31377: (OCaml "Big'_int.mult'_big'_int") haftmann@24999: (Haskell infixl 7 "*") haftmann@34886: (Scala infixl 8 "*") haftmann@38195: (Eval infixl 8 "*") haftmann@24999: huffman@47418: code_const Code_Numeral.div_mod haftmann@38195: (SML "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Int.div (n, m), Int.mod (n, m)))") haftmann@34898: (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") haftmann@26009: (Haskell "divMod") haftmann@38195: (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") haftmann@40060: (Eval "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Integer.div'_mod n m))") haftmann@25928: haftmann@39086: code_const "HOL.equal \ code_numeral \ code_numeral \ bool" haftmann@24999: (SML "!((_ : Int.int) = _)") haftmann@31377: (OCaml "Big'_int.eq'_big'_int") haftmann@39499: (Haskell infix 4 "==") haftmann@34886: (Scala infixl 5 "==") haftmann@38195: (Eval "!((_ : int) = _)") haftmann@24999: huffman@47418: code_const "less_eq \ code_numeral \ code_numeral \ bool" haftmann@25928: (SML "Int.<=/ ((_),/ (_))") haftmann@31377: (OCaml "Big'_int.le'_big'_int") haftmann@24999: (Haskell infix 4 "<=") haftmann@34898: (Scala infixl 4 "<=") haftmann@38195: (Eval infixl 6 "<=") haftmann@24999: huffman@47418: code_const "less \ code_numeral \ code_numeral \ bool" haftmann@25928: (SML "Int.