src/HOL/Code_Numeral.thy
author haftmann
Wed, 27 May 2009 22:11:05 +0200
changeset 31259 55e70b6d812e
parent 31205 98370b26c2ce
child 31377 a48f9ef9de15
permissions -rw-r--r--
added lemma about 0 - 1
     1 (* Author: Florian Haftmann, TU Muenchen *)
     2 
     3 header {* Type of target language numerals *}
     4 
     5 theory Code_Numeral
     6 imports Nat_Numeral
     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 by rule
    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 [simp]:
    75   "Suc_code_numeral k = of_nat (Suc (nat_of k))"
    76 
    77 rep_datatype "0 \<Colon> code_numeral" Suc_code_numeral
    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_code_numeral k)"
    83     then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc_code_numeral (of_nat n))" .
    84     then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_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 - 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_code_numeral_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_code_numeral_def, simp_all)
   113 qed
   114 
   115 lemmas [code del] = code_numeral.recs code_numeral.cases
   116 
   117 lemma [code]:
   118   "eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of k) (nat_of l)"
   119   by (cases k, cases l) (simp add: eq)
   120 
   121 lemma [code nbe]:
   122   "eq_class.eq (k::code_numeral) k \<longleftrightarrow> True"
   123   by (rule HOL.eq_refl)
   124 
   125 
   126 subsection {* Indices as datatype of ints *}
   127 
   128 instantiation code_numeral :: number
   129 begin
   130 
   131 definition
   132   "number_of = of_nat o nat"
   133 
   134 instance ..
   135 
   136 end
   137 
   138 lemma nat_of_number [simp]:
   139   "nat_of (number_of k) = number_of k"
   140   by (simp add: number_of_code_numeral_def nat_number_of_def number_of_is_id)
   141 
   142 code_datatype "number_of \<Colon> int \<Rightarrow> code_numeral"
   143 
   144 
   145 subsection {* Basic arithmetic *}
   146 
   147 instantiation code_numeral :: "{minus, ordered_semidom, semiring_div, linorder}"
   148 begin
   149 
   150 definition [simp, code del]:
   151   "(1\<Colon>code_numeral) = of_nat 1"
   152 
   153 definition [simp, code del]:
   154   "n + m = of_nat (nat_of n + nat_of m)"
   155 
   156 definition [simp, code del]:
   157   "n - m = of_nat (nat_of n - nat_of m)"
   158 
   159 definition [simp, code del]:
   160   "n * m = of_nat (nat_of n * nat_of m)"
   161 
   162 definition [simp, code del]:
   163   "n div m = of_nat (nat_of n div nat_of m)"
   164 
   165 definition [simp, code del]:
   166   "n mod m = of_nat (nat_of n mod nat_of m)"
   167 
   168 definition [simp, code del]:
   169   "n \<le> m \<longleftrightarrow> nat_of n \<le> nat_of m"
   170 
   171 definition [simp, code del]:
   172   "n < m \<longleftrightarrow> nat_of n < nat_of m"
   173 
   174 instance proof
   175 qed (auto simp add: code_numeral left_distrib div_mult_self1)
   176 
   177 end
   178 
   179 lemma zero_code_numeral_code [code inline, code]:
   180   "(0\<Colon>code_numeral) = Numeral0"
   181   by (simp add: number_of_code_numeral_def Pls_def)
   182 lemma [code post]: "Numeral0 = (0\<Colon>code_numeral)"
   183   using zero_code_numeral_code ..
   184 
   185 lemma one_code_numeral_code [code inline, code]:
   186   "(1\<Colon>code_numeral) = Numeral1"
   187   by (simp add: number_of_code_numeral_def Pls_def Bit1_def)
   188 lemma [code post]: "Numeral1 = (1\<Colon>code_numeral)"
   189   using one_code_numeral_code ..
   190 
   191 lemma plus_code_numeral_code [code nbe]:
   192   "of_nat n + of_nat m = of_nat (n + m)"
   193   by simp
   194 
   195 definition subtract_code_numeral :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
   196   [simp, code del]: "subtract_code_numeral = op -"
   197 
   198 lemma subtract_code_numeral_code [code nbe]:
   199   "subtract_code_numeral (of_nat n) (of_nat m) = of_nat (n - m)"
   200   by simp
   201 
   202 lemma minus_code_numeral_code [code]:
   203   "n - m = subtract_code_numeral n m"
   204   by simp
   205 
   206 lemma times_code_numeral_code [code nbe]:
   207   "of_nat n * of_nat m = of_nat (n * m)"
   208   by simp
   209 
   210 lemma less_eq_code_numeral_code [code nbe]:
   211   "of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m"
   212   by simp
   213 
   214 lemma less_code_numeral_code [code nbe]:
   215   "of_nat n < of_nat m \<longleftrightarrow> n < m"
   216   by simp
   217 
   218 lemma code_numeral_zero_minus_one:
   219   "(0::code_numeral) - 1 = 0"
   220   by simp
   221 
   222 lemma Suc_code_numeral_minus_one:
   223   "Suc_code_numeral n - 1 = n"
   224   by simp
   225 
   226 lemma of_nat_code [code]:
   227   "of_nat = Nat.of_nat"
   228 proof
   229   fix n :: nat
   230   have "Nat.of_nat n = of_nat n"
   231     by (induct n) simp_all
   232   then show "of_nat n = Nat.of_nat n"
   233     by (rule sym)
   234 qed
   235 
   236 lemma code_numeral_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1"
   237   by (cases i) auto
   238 
   239 definition nat_of_aux :: "code_numeral \<Rightarrow> nat \<Rightarrow> nat" where
   240   "nat_of_aux i n = nat_of i + n"
   241 
   242 lemma nat_of_aux_code [code]:
   243   "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Suc n))"
   244   by (auto simp add: nat_of_aux_def code_numeral_not_eq_zero)
   245 
   246 lemma nat_of_code [code]:
   247   "nat_of i = nat_of_aux i 0"
   248   by (simp add: nat_of_aux_def)
   249 
   250 definition div_mod_code_numeral ::  "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where
   251   [code del]: "div_mod_code_numeral n m = (n div m, n mod m)"
   252 
   253 lemma [code]:
   254   "div_mod_code_numeral n m = (if m = 0 then (0, n) else (n div m, n mod m))"
   255   unfolding div_mod_code_numeral_def by auto
   256 
   257 lemma [code]:
   258   "n div m = fst (div_mod_code_numeral n m)"
   259   unfolding div_mod_code_numeral_def by simp
   260 
   261 lemma [code]:
   262   "n mod m = snd (div_mod_code_numeral n m)"
   263   unfolding div_mod_code_numeral_def by simp
   264 
   265 definition int_of :: "code_numeral \<Rightarrow> int" where
   266   "int_of = Nat.of_nat o nat_of"
   267 
   268 lemma int_of_code [code]:
   269   "int_of k = (if k = 0 then 0
   270     else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"
   271   by (auto simp add: int_of_def mod_div_equality')
   272 
   273 hide (open) const of_nat nat_of int_of
   274 
   275 
   276 subsection {* Code generator setup *}
   277 
   278 text {* Implementation of indices by bounded integers *}
   279 
   280 code_type code_numeral
   281   (SML "int")
   282   (OCaml "int")
   283   (Haskell "Int")
   284 
   285 code_instance code_numeral :: eq
   286   (Haskell -)
   287 
   288 setup {*
   289   fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
   290     false false) ["SML", "OCaml", "Haskell"]
   291 *}
   292 
   293 code_reserved SML Int int
   294 code_reserved OCaml Pervasives int
   295 
   296 code_const "op + \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
   297   (SML "Int.+/ ((_),/ (_))")
   298   (OCaml "Pervasives.( + )")
   299   (Haskell infixl 6 "+")
   300 
   301 code_const "subtract_code_numeral \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
   302   (SML "Int.max/ (_/ -/ _,/ 0 : int)")
   303   (OCaml "Pervasives.max/ (_/ -/ _)/ (0 : int) ")
   304   (Haskell "max/ (_/ -/ _)/ (0 :: Int)")
   305 
   306 code_const "op * \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
   307   (SML "Int.*/ ((_),/ (_))")
   308   (OCaml "Pervasives.( * )")
   309   (Haskell infixl 7 "*")
   310 
   311 code_const div_mod_code_numeral
   312   (SML "(fn n => fn m =>/ if m = 0/ then (0, n) else/ (n div m, n mod m))")
   313   (OCaml "(fun n -> fun m ->/ if m = 0/ then (0, n) else/ (n '/ m, n mod m))")
   314   (Haskell "divMod")
   315 
   316 code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   317   (SML "!((_ : Int.int) = _)")
   318   (OCaml "!((_ : int) = _)")
   319   (Haskell infixl 4 "==")
   320 
   321 code_const "op \<le> \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   322   (SML "Int.<=/ ((_),/ (_))")
   323   (OCaml "!((_ : int) <= _)")
   324   (Haskell infix 4 "<=")
   325 
   326 code_const "op < \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   327   (SML "Int.</ ((_),/ (_))")
   328   (OCaml "!((_ : int) < _)")
   329   (Haskell infix 4 "<")
   330 
   331 end