src/HOL/Code_Numeral.thy
author haftmann
Mon, 23 Jul 2012 09:28:03 +0200
changeset 49446 6efff142bb54
parent 48116 30a1692557b0
child 50849 b27bbb021df1
permissions -rw-r--r--
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
     1 (* Author: Florian Haftmann, TU Muenchen *)
     2 
     3 header {* Type of target language numerals *}
     4 
     5 theory Code_Numeral
     6 imports Nat_Transfer Divides
     7 begin
     8 
     9 text {*
    10   Code numerals are isomorphic to HOL @{typ nat} but
    11   mapped to target-language builtin numerals.
    12 *}
    13 
    14 subsection {* Datatype of target language numerals *}
    15 
    16 typedef (open) code_numeral = "UNIV \<Colon> nat set"
    17   morphisms nat_of of_nat ..
    18 
    19 lemma of_nat_nat_of [simp]:
    20   "of_nat (nat_of k) = k"
    21   by (rule nat_of_inverse)
    22 
    23 lemma nat_of_of_nat [simp]:
    24   "nat_of (of_nat n) = n"
    25   by (rule of_nat_inverse) (rule UNIV_I)
    26 
    27 lemma [measure_function]:
    28   "is_measure nat_of" by (rule is_measure_trivial)
    29 
    30 lemma code_numeral:
    31   "(\<And>n\<Colon>code_numeral. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (of_nat n))"
    32 proof
    33   fix n :: nat
    34   assume "\<And>n\<Colon>code_numeral. PROP P n"
    35   then show "PROP P (of_nat n)" .
    36 next
    37   fix n :: code_numeral
    38   assume "\<And>n\<Colon>nat. PROP P (of_nat n)"
    39   then have "PROP P (of_nat (nat_of n))" .
    40   then show "PROP P n" by simp
    41 qed
    42 
    43 lemma code_numeral_case:
    44   assumes "\<And>n. k = of_nat n \<Longrightarrow> P"
    45   shows P
    46   by (rule assms [of "nat_of k"]) simp
    47 
    48 lemma code_numeral_induct_raw:
    49   assumes "\<And>n. P (of_nat n)"
    50   shows "P k"
    51 proof -
    52   from assms have "P (of_nat (nat_of k))" .
    53   then show ?thesis by simp
    54 qed
    55 
    56 lemma nat_of_inject [simp]:
    57   "nat_of k = nat_of l \<longleftrightarrow> k = l"
    58   by (rule nat_of_inject)
    59 
    60 lemma of_nat_inject [simp]:
    61   "of_nat n = of_nat m \<longleftrightarrow> n = m"
    62   by (rule of_nat_inject) (rule UNIV_I)+
    63 
    64 instantiation code_numeral :: zero
    65 begin
    66 
    67 definition [simp, code del]:
    68   "0 = of_nat 0"
    69 
    70 instance ..
    71 
    72 end
    73 
    74 definition Suc where [simp]:
    75   "Suc k = of_nat (Nat.Suc (nat_of k))"
    76 
    77 rep_datatype "0 \<Colon> code_numeral" Suc
    78 proof -
    79   fix P :: "code_numeral \<Rightarrow> bool"
    80   fix k :: code_numeral
    81   assume "P 0" then have init: "P (of_nat 0)" by simp
    82   assume "\<And>k. P k \<Longrightarrow> P (Suc k)"
    83     then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc (of_nat n))" .
    84     then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Nat.Suc n))" by simp
    85   from init step have "P (of_nat (nat_of k))"
    86     by (induct ("nat_of k")) simp_all
    87   then show "P k" by simp
    88 qed simp_all
    89 
    90 declare code_numeral_case [case_names nat, cases type: code_numeral]
    91 declare code_numeral.induct [case_names nat, induct type: code_numeral]
    92 
    93 lemma code_numeral_decr [termination_simp]:
    94   "k \<noteq> of_nat 0 \<Longrightarrow> nat_of k - Nat.Suc 0 < nat_of k"
    95   by (cases k) simp
    96 
    97 lemma [simp, code]:
    98   "code_numeral_size = nat_of"
    99 proof (rule ext)
   100   fix k
   101   have "code_numeral_size k = nat_size (nat_of k)"
   102     by (induct k rule: code_numeral.induct) (simp_all del: zero_code_numeral_def Suc_def, simp_all)
   103   also have "nat_size (nat_of k) = nat_of k" by (induct ("nat_of k")) simp_all
   104   finally show "code_numeral_size k = nat_of k" .
   105 qed
   106 
   107 lemma [simp, code]:
   108   "size = nat_of"
   109 proof (rule ext)
   110   fix k
   111   show "size k = nat_of k"
   112   by (induct k) (simp_all del: zero_code_numeral_def Suc_def, simp_all)
   113 qed
   114 
   115 lemmas [code del] = code_numeral.recs code_numeral.cases
   116 
   117 lemma [code]:
   118   "HOL.equal k l \<longleftrightarrow> HOL.equal (nat_of k) (nat_of l)"
   119   by (cases k, cases l) (simp add: equal)
   120 
   121 lemma [code nbe]:
   122   "HOL.equal (k::code_numeral) k \<longleftrightarrow> True"
   123   by (rule equal_refl)
   124 
   125 
   126 subsection {* Basic arithmetic *}
   127 
   128 instantiation code_numeral :: "{minus, linordered_semidom, semiring_div, linorder}"
   129 begin
   130 
   131 definition [simp, code del]:
   132   "(1\<Colon>code_numeral) = of_nat 1"
   133 
   134 definition [simp, code del]:
   135   "n + m = of_nat (nat_of n + nat_of m)"
   136 
   137 definition [simp, code del]:
   138   "n - m = of_nat (nat_of n - nat_of m)"
   139 
   140 definition [simp, code del]:
   141   "n * m = of_nat (nat_of n * nat_of m)"
   142 
   143 definition [simp, code del]:
   144   "n div m = of_nat (nat_of n div nat_of m)"
   145 
   146 definition [simp, code del]:
   147   "n mod m = of_nat (nat_of n mod nat_of m)"
   148 
   149 definition [simp, code del]:
   150   "n \<le> m \<longleftrightarrow> nat_of n \<le> nat_of m"
   151 
   152 definition [simp, code del]:
   153   "n < m \<longleftrightarrow> nat_of n < nat_of m"
   154 
   155 instance proof
   156 qed (auto simp add: code_numeral left_distrib intro: mult_commute)
   157 
   158 end
   159 
   160 lemma nat_of_numeral [simp]: "nat_of (numeral k) = numeral k"
   161   by (induct k rule: num_induct) (simp_all add: numeral_inc)
   162 
   163 definition Num :: "num \<Rightarrow> code_numeral"
   164   where [simp, code_abbrev]: "Num = numeral"
   165 
   166 code_datatype "0::code_numeral" Num
   167 
   168 lemma one_code_numeral_code [code]:
   169   "(1\<Colon>code_numeral) = Numeral1"
   170   by simp
   171 
   172 lemma [code_abbrev]: "Numeral1 = (1\<Colon>code_numeral)"
   173   using one_code_numeral_code ..
   174 
   175 lemma plus_code_numeral_code [code nbe]:
   176   "of_nat n + of_nat m = of_nat (n + m)"
   177   by simp
   178 
   179 lemma minus_code_numeral_code [code nbe]:
   180   "of_nat n - of_nat m = of_nat (n - m)"
   181   by simp
   182 
   183 lemma times_code_numeral_code [code nbe]:
   184   "of_nat n * of_nat m = of_nat (n * m)"
   185   by simp
   186 
   187 lemma less_eq_code_numeral_code [code nbe]:
   188   "of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m"
   189   by simp
   190 
   191 lemma less_code_numeral_code [code nbe]:
   192   "of_nat n < of_nat m \<longleftrightarrow> n < m"
   193   by simp
   194 
   195 lemma code_numeral_zero_minus_one:
   196   "(0::code_numeral) - 1 = 0"
   197   by simp
   198 
   199 lemma Suc_code_numeral_minus_one:
   200   "Suc n - 1 = n"
   201   by simp
   202 
   203 lemma of_nat_code [code]:
   204   "of_nat = Nat.of_nat"
   205 proof
   206   fix n :: nat
   207   have "Nat.of_nat n = of_nat n"
   208     by (induct n) simp_all
   209   then show "of_nat n = Nat.of_nat n"
   210     by (rule sym)
   211 qed
   212 
   213 lemma code_numeral_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1"
   214   by (cases i) auto
   215 
   216 definition nat_of_aux :: "code_numeral \<Rightarrow> nat \<Rightarrow> nat" where
   217   "nat_of_aux i n = nat_of i + n"
   218 
   219 lemma nat_of_aux_code [code]:
   220   "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Nat.Suc n))"
   221   by (auto simp add: nat_of_aux_def code_numeral_not_eq_zero)
   222 
   223 lemma nat_of_code [code]:
   224   "nat_of i = nat_of_aux i 0"
   225   by (simp add: nat_of_aux_def)
   226 
   227 definition div_mod :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where
   228   [code del]: "div_mod n m = (n div m, n mod m)"
   229 
   230 lemma [code]:
   231   "div_mod n m = (if m = 0 then (0, n) else (n div m, n mod m))"
   232   unfolding div_mod_def by auto
   233 
   234 lemma [code]:
   235   "n div m = fst (div_mod n m)"
   236   unfolding div_mod_def by simp
   237 
   238 lemma [code]:
   239   "n mod m = snd (div_mod n m)"
   240   unfolding div_mod_def by simp
   241 
   242 definition int_of :: "code_numeral \<Rightarrow> int" where
   243   "int_of = Nat.of_nat o nat_of"
   244 
   245 lemma int_of_code [code]:
   246   "int_of k = (if k = 0 then 0
   247     else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"
   248 proof -
   249   have "(nat_of k div 2) * 2 + nat_of k mod 2 = nat_of k" 
   250     by (rule mod_div_equality)
   251   then have "int ((nat_of k div 2) * 2 + nat_of k mod 2) = int (nat_of k)" 
   252     by simp
   253   then have "int (nat_of k) = int (nat_of k div 2) * 2 + int (nat_of k mod 2)" 
   254     unfolding of_nat_mult of_nat_add by simp
   255   then show ?thesis by (auto simp add: int_of_def mult_ac)
   256 qed
   257 
   258 
   259 hide_const (open) of_nat nat_of Suc int_of
   260 
   261 
   262 subsection {* Code generator setup *}
   263 
   264 text {* Implementation of code numerals by bounded integers *}
   265 
   266 code_type code_numeral
   267   (SML "int")
   268   (OCaml "Big'_int.big'_int")
   269   (Haskell "Integer")
   270   (Scala "BigInt")
   271 
   272 code_instance code_numeral :: equal
   273   (Haskell -)
   274 
   275 setup {*
   276   Numeral.add_code @{const_name Num}
   277     false Code_Printer.literal_naive_numeral "SML"
   278   #> fold (Numeral.add_code @{const_name Num}
   279     false Code_Printer.literal_numeral) ["OCaml", "Haskell", "Scala"]
   280 *}
   281 
   282 code_reserved SML Int int
   283 code_reserved Eval Integer
   284 
   285 code_const "0::code_numeral"
   286   (SML "0")
   287   (OCaml "Big'_int.zero'_big'_int")
   288   (Haskell "0")
   289   (Scala "BigInt(0)")
   290 
   291 code_const "plus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
   292   (SML "Int.+/ ((_),/ (_))")
   293   (OCaml "Big'_int.add'_big'_int")
   294   (Haskell infixl 6 "+")
   295   (Scala infixl 7 "+")
   296   (Eval infixl 8 "+")
   297 
   298 code_const "minus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
   299   (SML "Int.max/ (0 : int,/ Int.-/ ((_),/ (_)))")
   300   (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)")
   301   (Haskell "Prelude.max/ (0 :: Integer)/ (_/ -/ _)")
   302   (Scala "!(_/ -/ _).max(0)")
   303   (Eval "Integer.max/ 0/ (_/ -/ _)")
   304 
   305 code_const "times \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
   306   (SML "Int.*/ ((_),/ (_))")
   307   (OCaml "Big'_int.mult'_big'_int")
   308   (Haskell infixl 7 "*")
   309   (Scala infixl 8 "*")
   310   (Eval infixl 8 "*")
   311 
   312 code_const Code_Numeral.div_mod
   313   (SML "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Int.div (n, m), Int.mod (n, m)))")
   314   (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
   315   (Haskell "divMod")
   316   (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
   317   (Eval "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Integer.div'_mod n m))")
   318 
   319 code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   320   (SML "!((_ : Int.int) = _)")
   321   (OCaml "Big'_int.eq'_big'_int")
   322   (Haskell infix 4 "==")
   323   (Scala infixl 5 "==")
   324   (Eval "!((_ : int) = _)")
   325 
   326 code_const "less_eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   327   (SML "Int.<=/ ((_),/ (_))")
   328   (OCaml "Big'_int.le'_big'_int")
   329   (Haskell infix 4 "<=")
   330   (Scala infixl 4 "<=")
   331   (Eval infixl 6 "<=")
   332 
   333 code_const "less \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   334   (SML "Int.</ ((_),/ (_))")
   335   (OCaml "Big'_int.lt'_big'_int")
   336   (Haskell infix 4 "<")
   337   (Scala infixl 4 "<")
   338   (Eval infixl 6 "<")
   339 
   340 code_modulename SML
   341   Code_Numeral Arith
   342 
   343 code_modulename OCaml
   344   Code_Numeral Arith
   345 
   346 code_modulename Haskell
   347   Code_Numeral Arith
   348 
   349 end
   350