src/HOL/ex/Numeral.thy
author haftmann
Thu, 28 Aug 2008 22:08:11 +0200
changeset 28053 a2106c0d8c45
parent 28021 32acf3c6cd12
child 28367 10ea34297962
permissions -rw-r--r--
no parameter prefix for class interpretation
haftmann@28021
     1
(*  Title:      HOL/ex/Numeral.thy
haftmann@28021
     2
    ID:         $Id$
haftmann@28021
     3
    Author:     Florian Haftmann
haftmann@28021
     4
haftmann@28021
     5
An experimental alternative numeral representation.
haftmann@28021
     6
*)
haftmann@28021
     7
haftmann@28021
     8
theory Numeral
haftmann@28021
     9
imports Int Inductive
haftmann@28021
    10
begin
haftmann@28021
    11
haftmann@28021
    12
subsection {* The @{text num} type *}
haftmann@28021
    13
haftmann@28021
    14
text {*
haftmann@28021
    15
  We construct @{text num} as a copy of strictly positive
haftmann@28021
    16
  natural numbers.
haftmann@28021
    17
*}
haftmann@28021
    18
haftmann@28021
    19
typedef (open) num = "\<lambda>n\<Colon>nat. n > 0"
haftmann@28021
    20
  morphisms nat_of_num num_of_nat_abs
haftmann@28021
    21
  by (auto simp add: mem_def)
haftmann@28021
    22
haftmann@28021
    23
text {*
haftmann@28021
    24
  A totalized abstraction function.  It is not entirely clear
haftmann@28021
    25
  whether this is really useful.
haftmann@28021
    26
*}
haftmann@28021
    27
haftmann@28021
    28
definition num_of_nat :: "nat \<Rightarrow> num" where
haftmann@28021
    29
  "num_of_nat n = (if n = 0 then num_of_nat_abs 1 else num_of_nat_abs n)"
haftmann@28021
    30
haftmann@28021
    31
lemma num_cases [case_names nat, cases type: num]:
haftmann@28021
    32
  assumes "(\<And>n\<Colon>nat. m = num_of_nat n \<Longrightarrow> 0 < n \<Longrightarrow> P)"
haftmann@28021
    33
  shows P
haftmann@28021
    34
apply (rule num_of_nat_abs_cases)
haftmann@28021
    35
apply (unfold mem_def)
haftmann@28021
    36
using assms unfolding num_of_nat_def
haftmann@28021
    37
apply auto
haftmann@28021
    38
done
haftmann@28021
    39
haftmann@28021
    40
lemma num_of_nat_zero: "num_of_nat 0 = num_of_nat 1"
haftmann@28021
    41
  by (simp add: num_of_nat_def)
haftmann@28021
    42
haftmann@28021
    43
lemma num_of_nat_inverse: "nat_of_num (num_of_nat n) = (if n = 0 then 1 else n)"
haftmann@28021
    44
  apply (simp add: num_of_nat_def)
haftmann@28021
    45
  apply (subst num_of_nat_abs_inverse)
haftmann@28021
    46
  apply (auto simp add: mem_def num_of_nat_abs_inverse)
haftmann@28021
    47
  done
haftmann@28021
    48
haftmann@28021
    49
lemma num_of_nat_inject:
haftmann@28021
    50
  "num_of_nat m = num_of_nat n \<longleftrightarrow> m = n \<or> (m = 0 \<or> m = 1) \<and> (n = 0 \<or> n = 1)"
haftmann@28021
    51
by (auto simp add: num_of_nat_def num_of_nat_abs_inject [unfolded mem_def])
haftmann@28021
    52
haftmann@28021
    53
lemma split_num_all:
haftmann@28021
    54
  "(\<And>m. PROP P m) \<equiv> (\<And>n. PROP P (num_of_nat n))"
haftmann@28021
    55
proof
haftmann@28021
    56
  fix n
haftmann@28021
    57
  assume "\<And>m\<Colon>num. PROP P m"
haftmann@28021
    58
  then show "PROP P (num_of_nat n)" .
haftmann@28021
    59
next
haftmann@28021
    60
  fix m
haftmann@28021
    61
  have nat_of_num: "\<And>m. nat_of_num m \<noteq> 0"
haftmann@28021
    62
    using nat_of_num by (auto simp add: mem_def)
haftmann@28021
    63
  have nat_of_num_inverse: "\<And>m. num_of_nat (nat_of_num m) = m"
haftmann@28021
    64
    by (auto simp add: num_of_nat_def nat_of_num_inverse nat_of_num)
haftmann@28021
    65
  assume "\<And>n. PROP P (num_of_nat n)"
haftmann@28021
    66
  then have "PROP P (num_of_nat (nat_of_num m))" .
haftmann@28021
    67
  then show "PROP P m" unfolding nat_of_num_inverse .
haftmann@28021
    68
qed
haftmann@28021
    69
haftmann@28021
    70
haftmann@28021
    71
subsection {* Digit representation for @{typ num} *}
haftmann@28021
    72
haftmann@28021
    73
instantiation num :: "{semiring, monoid_mult}"
haftmann@28021
    74
begin
haftmann@28021
    75
haftmann@28021
    76
definition one_num :: num where
haftmann@28021
    77
  [code func del]: "1 = num_of_nat 1"
haftmann@28021
    78
haftmann@28021
    79
definition plus_num :: "num \<Rightarrow> num \<Rightarrow> num" where
haftmann@28021
    80
  [code func del]: "m + n = num_of_nat (nat_of_num m + nat_of_num n)"
haftmann@28021
    81
haftmann@28021
    82
definition times_num :: "num \<Rightarrow> num \<Rightarrow> num" where
haftmann@28021
    83
  [code func del]: "m * n = num_of_nat (nat_of_num m * nat_of_num n)"
haftmann@28021
    84
haftmann@28021
    85
definition Dig0 :: "num \<Rightarrow> num" where
haftmann@28021
    86
  [code func del]: "Dig0 n = n + n"
haftmann@28021
    87
haftmann@28021
    88
definition Dig1 :: "num \<Rightarrow> num" where
haftmann@28021
    89
  [code func del]: "Dig1 n = n + n + 1"
haftmann@28021
    90
haftmann@28021
    91
instance proof
haftmann@28021
    92
qed (simp_all add: one_num_def plus_num_def times_num_def
haftmann@28021
    93
  split_num_all num_of_nat_inverse num_of_nat_zero add_ac mult_ac nat_distrib)
haftmann@28021
    94
haftmann@28021
    95
end
haftmann@28021
    96
haftmann@28021
    97
text {*
haftmann@28021
    98
  The following proofs seem horribly complicated.
haftmann@28021
    99
  Any room for simplification!?
haftmann@28021
   100
*}
haftmann@28021
   101
haftmann@28021
   102
lemma nat_dig_cases [case_names 0 1 dig0 dig1]:
haftmann@28021
   103
  fixes n :: nat
haftmann@28021
   104
  assumes "n = 0 \<Longrightarrow> P"
haftmann@28021
   105
  and "n = 1 \<Longrightarrow> P"
haftmann@28021
   106
  and "\<And>m. m > 0 \<Longrightarrow> n = m + m \<Longrightarrow> P"
haftmann@28021
   107
  and "\<And>m. m > 0 \<Longrightarrow> n = Suc (m + m) \<Longrightarrow> P"
haftmann@28021
   108
  shows P
haftmann@28021
   109
using assms proof (induct n)
haftmann@28021
   110
  case 0 then show ?case by simp
haftmann@28021
   111
next
haftmann@28021
   112
  case (Suc n)
haftmann@28021
   113
  show P proof (rule Suc.hyps)
haftmann@28021
   114
    assume "n = 0"
haftmann@28021
   115
    then have "Suc n = 1" by simp
haftmann@28021
   116
    then show P by (rule Suc.prems(2))
haftmann@28021
   117
  next
haftmann@28021
   118
    assume "n = 1"
haftmann@28021
   119
    have "1 > (0\<Colon>nat)" by simp
haftmann@28021
   120
    moreover from `n = 1` have "Suc n = 1 + 1" by simp
haftmann@28021
   121
    ultimately show P by (rule Suc.prems(3))
haftmann@28021
   122
  next
haftmann@28021
   123
    fix m
haftmann@28021
   124
    assume "0 < m" and "n = m + m"
haftmann@28021
   125
    note `0 < m`
haftmann@28021
   126
    moreover from `n = m + m` have "Suc n = Suc (m + m)" by simp
haftmann@28021
   127
    ultimately show P by (rule Suc.prems(4))
haftmann@28021
   128
  next
haftmann@28021
   129
    fix m
haftmann@28021
   130
    assume "0 < m" and "n = Suc (m + m)"
haftmann@28021
   131
    have "0 < Suc m" by simp
haftmann@28021
   132
    moreover from `n = Suc (m + m)` have "Suc n = Suc m + Suc m" by simp
haftmann@28021
   133
    ultimately show P by (rule Suc.prems(3))
haftmann@28021
   134
  qed
haftmann@28021
   135
qed
haftmann@28021
   136
haftmann@28021
   137
lemma num_induct_raw:
haftmann@28021
   138
  fixes n :: nat
haftmann@28021
   139
  assumes not0: "n > 0"
haftmann@28021
   140
  assumes "P 1"
haftmann@28021
   141
  and "\<And>n. n > 0 \<Longrightarrow> P n \<Longrightarrow> P (n + n)"
haftmann@28021
   142
  and "\<And>n. n > 0 \<Longrightarrow> P n \<Longrightarrow> P (Suc (n + n))"
haftmann@28021
   143
  shows "P n"
haftmann@28021
   144
using not0 proof (induct n rule: less_induct)
haftmann@28021
   145
  case (less n)
haftmann@28021
   146
  show "P n" proof (cases n rule: nat_dig_cases)
haftmann@28021
   147
    case 0 then show ?thesis using less by simp
haftmann@28021
   148
  next
haftmann@28021
   149
    case 1 then show ?thesis using assms by simp
haftmann@28021
   150
  next
haftmann@28021
   151
    case (dig0 m)
haftmann@28021
   152
    then show ?thesis apply simp
haftmann@28021
   153
      apply (rule assms(3)) apply assumption
haftmann@28021
   154
      apply (rule less)
haftmann@28021
   155
      apply simp_all
haftmann@28021
   156
    done
haftmann@28021
   157
  next
haftmann@28021
   158
    case (dig1 m)
haftmann@28021
   159
    then show ?thesis apply simp
haftmann@28021
   160
      apply (rule assms(4)) apply assumption
haftmann@28021
   161
      apply (rule less)
haftmann@28021
   162
      apply simp_all
haftmann@28021
   163
    done
haftmann@28021
   164
  qed
haftmann@28021
   165
qed
haftmann@28021
   166
haftmann@28021
   167
lemma num_of_nat_Suc: "num_of_nat (Suc n) = (if n = 0 then 1 else num_of_nat n + 1)"
haftmann@28021
   168
  by (cases n) (auto simp add: one_num_def plus_num_def num_of_nat_inverse)
haftmann@28021
   169
haftmann@28021
   170
lemma num_induct [case_names 1 Suc, induct type: num]:
haftmann@28021
   171
  fixes P :: "num \<Rightarrow> bool"
haftmann@28021
   172
  assumes 1: "P 1"
haftmann@28021
   173
    and Suc: "\<And>n. P n \<Longrightarrow> P (n + 1)"
haftmann@28021
   174
  shows "P n"
haftmann@28021
   175
proof (cases n)
haftmann@28021
   176
  case (nat m) then show ?thesis by (induct m arbitrary: n)
haftmann@28021
   177
    (auto simp: num_of_nat_Suc intro: 1 Suc split: split_if_asm)
haftmann@28021
   178
qed
haftmann@28021
   179
haftmann@28021
   180
rep_datatype "1::num" Dig0 Dig1 proof -
haftmann@28021
   181
  fix P m
haftmann@28021
   182
  assume 1: "P 1"
haftmann@28021
   183
    and Dig0: "\<And>m. P m \<Longrightarrow> P (Dig0 m)"
haftmann@28021
   184
    and Dig1: "\<And>m. P m \<Longrightarrow> P (Dig1 m)"
haftmann@28021
   185
  obtain n where "0 < n" and m: "m = num_of_nat n"
haftmann@28021
   186
    by (cases m) auto
haftmann@28021
   187
  from `0 < n` have "P (num_of_nat n)" proof (induct n rule: num_induct_raw)
haftmann@28021
   188
    case 1 from `0 < n` show ?case .
haftmann@28021
   189
  next
haftmann@28021
   190
    case 2 with 1 show ?case by (simp add: one_num_def)
haftmann@28021
   191
  next
haftmann@28021
   192
    case (3 n) then have "P (num_of_nat n)" by auto
haftmann@28021
   193
    then have "P (Dig0 (num_of_nat n))" by (rule Dig0)
haftmann@28021
   194
    with 3 show ?case by (simp add: Dig0_def plus_num_def num_of_nat_inverse)
haftmann@28021
   195
  next
haftmann@28021
   196
    case (4 n) then have "P (num_of_nat n)" by auto
haftmann@28021
   197
    then have "P (Dig1 (num_of_nat n))" by (rule Dig1)
haftmann@28021
   198
    with 4 show ?case by (simp add: Dig1_def one_num_def plus_num_def num_of_nat_inverse)
haftmann@28021
   199
  qed
haftmann@28021
   200
  with m show "P m" by simp
haftmann@28021
   201
next
haftmann@28021
   202
  fix m n
haftmann@28021
   203
  show "Dig0 m = Dig0 n \<longleftrightarrow> m = n"
haftmann@28021
   204
    apply (cases m) apply (cases n)
haftmann@28021
   205
    by (auto simp add: Dig0_def plus_num_def num_of_nat_inverse num_of_nat_inject)
haftmann@28021
   206
next
haftmann@28021
   207
  fix m n
haftmann@28021
   208
  show "Dig1 m = Dig1 n \<longleftrightarrow> m = n"
haftmann@28021
   209
    apply (cases m) apply (cases n)
haftmann@28021
   210
    by (auto simp add: Dig1_def plus_num_def num_of_nat_inverse num_of_nat_inject)
haftmann@28021
   211
next
haftmann@28021
   212
  fix n
haftmann@28021
   213
  show "1 \<noteq> Dig0 n"
haftmann@28021
   214
    apply (cases n)
haftmann@28021
   215
    by (auto simp add: Dig0_def one_num_def plus_num_def num_of_nat_inverse num_of_nat_inject)
haftmann@28021
   216
next
haftmann@28021
   217
  fix n
haftmann@28021
   218
  show "1 \<noteq> Dig1 n"
haftmann@28021
   219
    apply (cases n)
haftmann@28021
   220
    by (auto simp add: Dig1_def one_num_def plus_num_def num_of_nat_inverse num_of_nat_inject)
haftmann@28021
   221
next
haftmann@28021
   222
  fix m n
haftmann@28021
   223
  have "\<And>n m. n + n \<noteq> Suc (m + m)"
haftmann@28021
   224
  proof -
haftmann@28021
   225
    fix n m
haftmann@28021
   226
    show "n + n \<noteq> Suc (m + m)"
haftmann@28021
   227
    proof (induct m arbitrary: n)
haftmann@28021
   228
      case 0 then show ?case by (cases n) simp_all
haftmann@28021
   229
    next
haftmann@28021
   230
      case (Suc m) then show ?case by (cases n) simp_all
haftmann@28021
   231
    qed
haftmann@28021
   232
  qed
haftmann@28021
   233
  then show "Dig0 n \<noteq> Dig1 m"
haftmann@28021
   234
    apply (cases n) apply (cases m)
haftmann@28021
   235
    by (auto simp add: Dig0_def Dig1_def one_num_def plus_num_def num_of_nat_inverse num_of_nat_inject)
haftmann@28021
   236
qed
haftmann@28021
   237
haftmann@28021
   238
text {*
haftmann@28021
   239
  From now on, there are two possible models for @{typ num}:
haftmann@28021
   240
  as positive naturals (rules @{text "num_induct"}, @{text "num_cases"})
haftmann@28021
   241
  and as digit representation (rules @{text "num.induct"}, @{text "num.cases"}).
haftmann@28021
   242
haftmann@28021
   243
  It is not entirely clear in which context it is better to use
haftmann@28021
   244
  the one or the other, or whether the construction should be reversed.
haftmann@28021
   245
*}
haftmann@28021
   246
haftmann@28021
   247
haftmann@28021
   248
subsection {* Binary numerals *}
haftmann@28021
   249
haftmann@28021
   250
text {*
haftmann@28021
   251
  We embed binary representations into a generic algebraic
haftmann@28021
   252
  structure using @{text of_num}
haftmann@28021
   253
*}
haftmann@28021
   254
haftmann@28021
   255
ML {*
haftmann@28021
   256
structure DigSimps =
haftmann@28021
   257
  NamedThmsFun(val name = "numeral"; val description = "Simplification rules for numerals")
haftmann@28021
   258
*}
haftmann@28021
   259
haftmann@28021
   260
setup DigSimps.setup
haftmann@28021
   261
haftmann@28021
   262
class semiring_numeral = semiring + monoid_mult
haftmann@28021
   263
begin
haftmann@28021
   264
haftmann@28021
   265
primrec of_num :: "num \<Rightarrow> 'a" where
haftmann@28021
   266
  of_num_one [numeral]: "of_num 1 = 1"
haftmann@28021
   267
  | "of_num (Dig0 n) = of_num n + of_num n"
haftmann@28021
   268
  | "of_num (Dig1 n) = of_num n + of_num n + 1"
haftmann@28021
   269
haftmann@28021
   270
declare of_num.simps [simp del]
haftmann@28021
   271
haftmann@28021
   272
end
haftmann@28021
   273
haftmann@28021
   274
text {*
haftmann@28021
   275
  ML stuff and syntax.
haftmann@28021
   276
*}
haftmann@28021
   277
haftmann@28021
   278
ML {*
haftmann@28021
   279
fun mk_num 1 = @{term "1::num"}
haftmann@28021
   280
  | mk_num k =
haftmann@28021
   281
      let
haftmann@28021
   282
        val (l, b) = Integer.div_mod k 2;
haftmann@28021
   283
        val bit = (if b = 0 then @{term Dig0} else @{term Dig1});
haftmann@28021
   284
      in bit $ (mk_num l) end;
haftmann@28021
   285
haftmann@28021
   286
fun dest_num @{term "1::num"} = 1
haftmann@28021
   287
  | dest_num (@{term Dig0} $ n) = 2 * dest_num n
haftmann@28021
   288
  | dest_num (@{term Dig1} $ n) = 2 * dest_num n + 1;
haftmann@28021
   289
haftmann@28021
   290
(*FIXME these have to gain proper context via morphisms phi*)
haftmann@28021
   291
haftmann@28021
   292
fun mk_numeral T k = Const (@{const_name of_num}, @{typ num} --> T)
haftmann@28021
   293
  $ mk_num k
haftmann@28021
   294
haftmann@28021
   295
fun dest_numeral (Const (@{const_name of_num}, Type ("fun", [@{typ num}, T])) $ t) =
haftmann@28021
   296
  (T, dest_num t)
haftmann@28021
   297
*}
haftmann@28021
   298
haftmann@28021
   299
syntax
haftmann@28021
   300
  "_Numerals" :: "xnum \<Rightarrow> 'a"    ("_")
haftmann@28021
   301
haftmann@28021
   302
parse_translation {*
haftmann@28021
   303
let
haftmann@28021
   304
  fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2)
haftmann@28021
   305
     of (0, 1) => Const (@{const_name HOL.one}, dummyT)
haftmann@28021
   306
      | (n, 0) => Const (@{const_name Dig0}, dummyT) $ num_of_int n
haftmann@28021
   307
      | (n, 1) => Const (@{const_name Dig1}, dummyT) $ num_of_int n
haftmann@28021
   308
    else raise Match;
haftmann@28021
   309
  fun numeral_tr [Free (num, _)] =
haftmann@28021
   310
        let
haftmann@28021
   311
          val {leading_zeros, value, ...} = Syntax.read_xnum num;
haftmann@28021
   312
          val _ = leading_zeros = 0 andalso value > 0
haftmann@28021
   313
            orelse error ("Bad numeral: " ^ num);
haftmann@28021
   314
        in Const (@{const_name of_num}, @{typ num} --> dummyT) $ num_of_int value end
haftmann@28021
   315
    | numeral_tr ts = raise TERM ("numeral_tr", ts);
haftmann@28021
   316
in [("_Numerals", numeral_tr)] end
haftmann@28021
   317
*}
haftmann@28021
   318
haftmann@28021
   319
typed_print_translation {*
haftmann@28021
   320
let
haftmann@28021
   321
  fun dig b n = b + 2 * n; 
haftmann@28021
   322
  fun int_of_num' (Const (@{const_syntax Dig0}, _) $ n) =
haftmann@28021
   323
        dig 0 (int_of_num' n)
haftmann@28021
   324
    | int_of_num' (Const (@{const_syntax Dig1}, _) $ n) =
haftmann@28021
   325
        dig 1 (int_of_num' n)
haftmann@28021
   326
    | int_of_num' (Const (@{const_syntax HOL.one}, _)) = 1;
haftmann@28021
   327
  fun num_tr' show_sorts T [n] =
haftmann@28021
   328
    let
haftmann@28021
   329
      val k = int_of_num' n;
haftmann@28021
   330
      val t' = Syntax.const "_Numerals" $ Syntax.free ("#" ^ string_of_int k);
haftmann@28021
   331
    in case T
haftmann@28021
   332
     of Type ("fun", [_, T']) =>
haftmann@28021
   333
         if not (! show_types) andalso can Term.dest_Type T' then t'
haftmann@28021
   334
         else Syntax.const Syntax.constrainC $ t' $ Syntax.term_of_typ show_sorts T'
haftmann@28021
   335
      | T' => if T' = dummyT then t' else raise Match
haftmann@28021
   336
    end;
haftmann@28021
   337
in [(@{const_syntax of_num}, num_tr')] end
haftmann@28021
   338
*}
haftmann@28021
   339
haftmann@28021
   340
haftmann@28021
   341
subsection {* Numeral operations *}
haftmann@28021
   342
haftmann@28021
   343
text {*
haftmann@28021
   344
  First, addition and multiplication on digits.
haftmann@28021
   345
*}
haftmann@28021
   346
haftmann@28021
   347
lemma Dig_plus [numeral, simp, code]:
haftmann@28021
   348
  "1 + 1 = Dig0 1"
haftmann@28021
   349
  "1 + Dig0 m = Dig1 m"
haftmann@28021
   350
  "1 + Dig1 m = Dig0 (m + 1)"
haftmann@28021
   351
  "Dig0 n + 1 = Dig1 n"
haftmann@28021
   352
  "Dig0 n + Dig0 m = Dig0 (n + m)"
haftmann@28021
   353
  "Dig0 n + Dig1 m = Dig1 (n + m)"
haftmann@28021
   354
  "Dig1 n + 1 = Dig0 (n + 1)"
haftmann@28021
   355
  "Dig1 n + Dig0 m = Dig1 (n + m)"
haftmann@28021
   356
  "Dig1 n + Dig1 m = Dig0 (n + m + 1)"
haftmann@28021
   357
  by (simp_all add: add_ac Dig0_def Dig1_def)
haftmann@28021
   358
haftmann@28021
   359
lemma Dig_times [numeral, simp, code]:
haftmann@28021
   360
  "1 * 1 = (1::num)"
haftmann@28021
   361
  "1 * Dig0 n = Dig0 n"
haftmann@28021
   362
  "1 * Dig1 n = Dig1 n"
haftmann@28021
   363
  "Dig0 n * 1 = Dig0 n"
haftmann@28021
   364
  "Dig0 n * Dig0 m = Dig0 (n * Dig0 m)"
haftmann@28021
   365
  "Dig0 n * Dig1 m = Dig0 (n * Dig1 m)"
haftmann@28021
   366
  "Dig1 n * 1 = Dig1 n"
haftmann@28021
   367
  "Dig1 n * Dig0 m = Dig0 (n * Dig0 m + m)"
haftmann@28021
   368
  "Dig1 n * Dig1 m = Dig1 (n * Dig1 m + m)"
haftmann@28021
   369
  by (simp_all add: left_distrib right_distrib add_ac Dig0_def Dig1_def)
haftmann@28021
   370
haftmann@28021
   371
text {*
haftmann@28021
   372
  @{const of_num} is a morphism.
haftmann@28021
   373
*}
haftmann@28021
   374
haftmann@28021
   375
context semiring_numeral
haftmann@28021
   376
begin
haftmann@28021
   377
haftmann@28021
   378
abbreviation "Num1 \<equiv> of_num 1"
haftmann@28021
   379
haftmann@28021
   380
text {*
haftmann@28021
   381
  Alas, there is still the duplication of @{term 1},
haftmann@28021
   382
  thought the duplicated @{term 0} has disappeared.
haftmann@28021
   383
  We could get rid of it by replacing the constructor
haftmann@28021
   384
  @{term 1} in @{typ num} by two constructors
haftmann@28021
   385
  @{text two} and @{text three}, resulting in a further
haftmann@28021
   386
  blow-up.  But it could be worth the effort.
haftmann@28021
   387
*}
haftmann@28021
   388
haftmann@28021
   389
lemma of_num_plus_one [numeral]:
haftmann@28021
   390
  "of_num n + 1 = of_num (n + 1)"
haftmann@28021
   391
  by (rule sym, induct n) (simp_all add: Dig_plus of_num.simps add_ac)
haftmann@28021
   392
haftmann@28021
   393
lemma of_num_one_plus [numeral]:
haftmann@28021
   394
  "1 + of_num n = of_num (n + 1)"
haftmann@28021
   395
  unfolding of_num_plus_one [symmetric] add_commute ..
haftmann@28021
   396
haftmann@28021
   397
lemma of_num_plus [numeral]:
haftmann@28021
   398
  "of_num m + of_num n = of_num (m + n)"
haftmann@28021
   399
  by (induct n rule: num_induct)
haftmann@28053
   400
    (simp_all add: Dig_plus of_num_one semigroup_add_class.add_assoc [symmetric, of m]
haftmann@28021
   401
    add_ac of_num_plus_one [symmetric])
haftmann@28021
   402
haftmann@28021
   403
lemma of_num_times_one [numeral]:
haftmann@28021
   404
  "of_num n * 1 = of_num n"
haftmann@28021
   405
  by simp
haftmann@28021
   406
haftmann@28021
   407
lemma of_num_one_times [numeral]:
haftmann@28021
   408
  "1 * of_num n = of_num n"
haftmann@28021
   409
  by simp
haftmann@28021
   410
haftmann@28021
   411
lemma of_num_times [numeral]:
haftmann@28021
   412
  "of_num m * of_num n = of_num (m * n)"
haftmann@28021
   413
  by (induct n rule: num_induct)
haftmann@28021
   414
    (simp_all add: of_num_plus [symmetric]
haftmann@28053
   415
    semiring_class.right_distrib right_distrib of_num_one)
haftmann@28021
   416
haftmann@28021
   417
end
haftmann@28021
   418
haftmann@28021
   419
text {*
haftmann@28021
   420
  Structures with a @{term 0}.
haftmann@28021
   421
*}
haftmann@28021
   422
haftmann@28021
   423
context semiring_1
haftmann@28021
   424
begin
haftmann@28021
   425
haftmann@28021
   426
subclass semiring_numeral ..
haftmann@28021
   427
haftmann@28021
   428
lemma of_nat_of_num [numeral]: "of_nat (of_num n) = of_num n"
haftmann@28021
   429
  by (induct n)
haftmann@28021
   430
    (simp_all add: semiring_numeral_class.of_num.simps of_num.simps add_ac)
haftmann@28021
   431
haftmann@28021
   432
declare of_nat_1 [numeral]
haftmann@28021
   433
haftmann@28021
   434
lemma Dig_plus_zero [numeral]:
haftmann@28021
   435
  "0 + 1 = 1"
haftmann@28021
   436
  "0 + of_num n = of_num n"
haftmann@28021
   437
  "1 + 0 = 1"
haftmann@28021
   438
  "of_num n + 0 = of_num n"
haftmann@28021
   439
  by simp_all
haftmann@28021
   440
haftmann@28021
   441
lemma Dig_times_zero [numeral]:
haftmann@28021
   442
  "0 * 1 = 0"
haftmann@28021
   443
  "0 * of_num n = 0"
haftmann@28021
   444
  "1 * 0 = 0"
haftmann@28021
   445
  "of_num n * 0 = 0"
haftmann@28021
   446
  by simp_all
haftmann@28021
   447
haftmann@28021
   448
end
haftmann@28021
   449
haftmann@28021
   450
lemma nat_of_num_of_num: "nat_of_num = of_num"
haftmann@28021
   451
proof
haftmann@28021
   452
  fix n
haftmann@28021
   453
  have "of_num n = nat_of_num n" apply (induct n)
haftmann@28021
   454
    apply (simp_all add: of_num.simps)
haftmann@28021
   455
    using nat_of_num
haftmann@28021
   456
    apply (simp_all add: one_num_def plus_num_def Dig0_def Dig1_def num_of_nat_inverse mem_def)
haftmann@28021
   457
    done
haftmann@28021
   458
  then show "nat_of_num n = of_num n" by simp
haftmann@28021
   459
qed
haftmann@28021
   460
haftmann@28021
   461
text {*
haftmann@28021
   462
  Equality.
haftmann@28021
   463
*}
haftmann@28021
   464
haftmann@28021
   465
context semiring_char_0
haftmann@28021
   466
begin
haftmann@28021
   467
haftmann@28021
   468
lemma of_num_eq_iff [numeral]:
haftmann@28021
   469
  "of_num m = of_num n \<longleftrightarrow> m = n"
haftmann@28021
   470
  unfolding of_nat_of_num [symmetric] nat_of_num_of_num [symmetric]
haftmann@28021
   471
    of_nat_eq_iff nat_of_num_inject ..
haftmann@28021
   472
haftmann@28021
   473
lemma of_num_eq_one_iff [numeral]:
haftmann@28021
   474
  "of_num n = 1 \<longleftrightarrow> n = 1"
haftmann@28021
   475
proof -
haftmann@28021
   476
  have "of_num n = of_num 1 \<longleftrightarrow> n = 1" unfolding of_num_eq_iff ..
haftmann@28021
   477
  then show ?thesis by (simp add: of_num_one)
haftmann@28021
   478
qed
haftmann@28021
   479
haftmann@28021
   480
lemma one_eq_of_num_iff [numeral]:
haftmann@28021
   481
  "1 = of_num n \<longleftrightarrow> n = 1"
haftmann@28021
   482
  unfolding of_num_eq_one_iff [symmetric] by auto
haftmann@28021
   483
haftmann@28021
   484
end
haftmann@28021
   485
haftmann@28021
   486
text {*
haftmann@28021
   487
  Comparisons.  Could be perhaps more general than here.
haftmann@28021
   488
*}
haftmann@28021
   489
haftmann@28021
   490
lemma (in ordered_semidom) of_num_pos: "0 < of_num n"
haftmann@28021
   491
proof -
haftmann@28021
   492
  have "(0::nat) < of_num n"
haftmann@28021
   493
    by (induct n) (simp_all add: semiring_numeral_class.of_num.simps)
haftmann@28021
   494
  then have "of_nat 0 \<noteq> of_nat (of_num n)" 
haftmann@28021
   495
    by (cases n) (simp_all only: semiring_numeral_class.of_num.simps of_nat_eq_iff)
haftmann@28021
   496
  then have "0 \<noteq> of_num n"
haftmann@28021
   497
    by (simp add: of_nat_of_num)
haftmann@28021
   498
  moreover have "0 \<le> of_nat (of_num n)" by simp
haftmann@28021
   499
  ultimately show ?thesis by (simp add: of_nat_of_num)
haftmann@28021
   500
qed
haftmann@28021
   501
haftmann@28021
   502
instantiation num :: linorder
haftmann@28021
   503
begin
haftmann@28021
   504
haftmann@28021
   505
definition less_eq_num :: "num \<Rightarrow> num \<Rightarrow> bool" where
haftmann@28021
   506
  [code func del]: "m \<le> n \<longleftrightarrow> nat_of_num m \<le> nat_of_num n"
haftmann@28021
   507
haftmann@28021
   508
definition less_num :: "num \<Rightarrow> num \<Rightarrow> bool" where
haftmann@28021
   509
  [code func del]: "m < n \<longleftrightarrow> nat_of_num m < nat_of_num n"
haftmann@28021
   510
haftmann@28021
   511
instance proof
haftmann@28021
   512
qed (auto simp add: less_eq_num_def less_num_def
haftmann@28021
   513
  split_num_all num_of_nat_inverse num_of_nat_inject split: split_if_asm)
haftmann@28021
   514
haftmann@28021
   515
end
haftmann@28021
   516
haftmann@28021
   517
lemma less_eq_num_code [numeral, simp, code]:
haftmann@28021
   518
  "(1::num) \<le> n \<longleftrightarrow> True"
haftmann@28021
   519
  "Dig0 m \<le> 1 \<longleftrightarrow> False"
haftmann@28021
   520
  "Dig1 m \<le> 1 \<longleftrightarrow> False"
haftmann@28021
   521
  "Dig0 m \<le> Dig0 n \<longleftrightarrow> m \<le> n"
haftmann@28021
   522
  "Dig0 m \<le> Dig1 n \<longleftrightarrow> m \<le> n"
haftmann@28021
   523
  "Dig1 m \<le> Dig1 n \<longleftrightarrow> m \<le> n"
haftmann@28021
   524
  "Dig1 m \<le> Dig0 n \<longleftrightarrow> m < n"
haftmann@28021
   525
  using of_num_pos [of n, where ?'a = nat] of_num_pos [of m, where ?'a = nat]
haftmann@28021
   526
  by (auto simp add: less_eq_num_def less_num_def nat_of_num_of_num of_num.simps)
haftmann@28021
   527
haftmann@28021
   528
lemma less_num_code [numeral, simp, code]:
haftmann@28021
   529
  "m < (1::num) \<longleftrightarrow> False"
haftmann@28021
   530
  "(1::num) < 1 \<longleftrightarrow> False"
haftmann@28021
   531
  "1 < Dig0 n \<longleftrightarrow> True"
haftmann@28021
   532
  "1 < Dig1 n \<longleftrightarrow> True"
haftmann@28021
   533
  "Dig0 m < Dig0 n \<longleftrightarrow> m < n"
haftmann@28021
   534
  "Dig0 m < Dig1 n \<longleftrightarrow> m \<le> n"
haftmann@28021
   535
  "Dig1 m < Dig1 n \<longleftrightarrow> m < n"
haftmann@28021
   536
  "Dig1 m < Dig0 n \<longleftrightarrow> m < n"
haftmann@28021
   537
  using of_num_pos [of n, where ?'a = nat] of_num_pos [of m, where ?'a = nat]
haftmann@28021
   538
  by (auto simp add: less_eq_num_def less_num_def nat_of_num_of_num of_num.simps)
haftmann@28021
   539
  
haftmann@28021
   540
context ordered_semidom
haftmann@28021
   541
begin
haftmann@28021
   542
haftmann@28021
   543
lemma of_num_less_eq_iff [numeral]: "of_num m \<le> of_num n \<longleftrightarrow> m \<le> n"
haftmann@28021
   544
proof -
haftmann@28021
   545
  have "of_nat (of_num m) \<le> of_nat (of_num n) \<longleftrightarrow> m \<le> n"
haftmann@28021
   546
    unfolding less_eq_num_def nat_of_num_of_num of_nat_le_iff ..
haftmann@28021
   547
  then show ?thesis by (simp add: of_nat_of_num)
haftmann@28021
   548
qed
haftmann@28021
   549
haftmann@28021
   550
lemma of_num_less_eq_one_iff [numeral]: "of_num n \<le> 1 \<longleftrightarrow> n = 1"
haftmann@28021
   551
proof -
haftmann@28021
   552
  have "of_num n \<le> of_num 1 \<longleftrightarrow> n = 1"
haftmann@28021
   553
    by (cases n) (simp_all add: of_num_less_eq_iff)
haftmann@28021
   554
  then show ?thesis by (simp add: of_num_one)
haftmann@28021
   555
qed
haftmann@28021
   556
haftmann@28021
   557
lemma one_less_eq_of_num_iff [numeral]: "1 \<le> of_num n"
haftmann@28021
   558
proof -
haftmann@28021
   559
  have "of_num 1 \<le> of_num n"
haftmann@28021
   560
    by (cases n) (simp_all add: of_num_less_eq_iff)
haftmann@28021
   561
  then show ?thesis by (simp add: of_num_one)
haftmann@28021
   562
qed
haftmann@28021
   563
haftmann@28021
   564
lemma of_num_less_iff [numeral]: "of_num m < of_num n \<longleftrightarrow> m < n"
haftmann@28021
   565
proof -
haftmann@28021
   566
  have "of_nat (of_num m) < of_nat (of_num n) \<longleftrightarrow> m < n"
haftmann@28021
   567
    unfolding less_num_def nat_of_num_of_num of_nat_less_iff ..
haftmann@28021
   568
  then show ?thesis by (simp add: of_nat_of_num)
haftmann@28021
   569
qed
haftmann@28021
   570
haftmann@28021
   571
lemma of_num_less_one_iff [numeral]: "\<not> of_num n < 1"
haftmann@28021
   572
proof -
haftmann@28021
   573
  have "\<not> of_num n < of_num 1"
haftmann@28021
   574
    by (cases n) (simp_all add: of_num_less_iff)
haftmann@28021
   575
  then show ?thesis by (simp add: of_num_one)
haftmann@28021
   576
qed
haftmann@28021
   577
haftmann@28021
   578
lemma one_less_of_num_iff [numeral]: "1 < of_num n \<longleftrightarrow> n \<noteq> 1"
haftmann@28021
   579
proof -
haftmann@28021
   580
  have "of_num 1 < of_num n \<longleftrightarrow> n \<noteq> 1"
haftmann@28021
   581
    by (cases n) (simp_all add: of_num_less_iff)
haftmann@28021
   582
  then show ?thesis by (simp add: of_num_one)
haftmann@28021
   583
qed
haftmann@28021
   584
haftmann@28021
   585
end
haftmann@28021
   586
haftmann@28021
   587
text {*
haftmann@28021
   588
  Structures with subtraction @{term "op -"}.
haftmann@28021
   589
*}
haftmann@28021
   590
haftmann@28021
   591
text {* A decrement function *}
haftmann@28021
   592
haftmann@28021
   593
primrec dec :: "num \<Rightarrow> num" where
haftmann@28021
   594
  "dec 1 = 1"
haftmann@28021
   595
  | "dec (Dig0 n) = (case n of 1 \<Rightarrow> 1 | _ \<Rightarrow> Dig1 (dec n))"
haftmann@28021
   596
  | "dec (Dig1 n) = Dig0 n"
haftmann@28021
   597
haftmann@28021
   598
declare dec.simps [simp del, code del]
haftmann@28021
   599
haftmann@28021
   600
lemma Dig_dec [numeral, simp, code]:
haftmann@28021
   601
  "dec 1 = 1"
haftmann@28021
   602
  "dec (Dig0 1) = 1"
haftmann@28021
   603
  "dec (Dig0 (Dig0 n)) = Dig1 (dec (Dig0 n))"
haftmann@28021
   604
  "dec (Dig0 (Dig1 n)) = Dig1 (Dig0 n)"
haftmann@28021
   605
  "dec (Dig1 n) = Dig0 n"
haftmann@28021
   606
  by (simp_all add: dec.simps)
haftmann@28021
   607
haftmann@28021
   608
lemma Dig_dec_plus_one:
haftmann@28021
   609
  "dec n + 1 = (if n = 1 then Dig0 1 else n)"
haftmann@28021
   610
  by (induct n)
haftmann@28021
   611
    (auto simp add: Dig_plus dec.simps,
haftmann@28021
   612
     auto simp add: Dig_plus split: num.splits)
haftmann@28021
   613
haftmann@28021
   614
lemma Dig_one_plus_dec:
haftmann@28021
   615
  "1 + dec n = (if n = 1 then Dig0 1 else n)"
haftmann@28021
   616
  unfolding add_commute [of 1] Dig_dec_plus_one ..
haftmann@28021
   617
haftmann@28021
   618
class semiring_minus = semiring + minus + zero +
haftmann@28021
   619
  assumes minus_inverts_plus1: "a + b = c \<Longrightarrow> c - b = a"
haftmann@28021
   620
  assumes minus_minus_zero_inverts_plus1: "a + b = c \<Longrightarrow> b - c = 0 - a"
haftmann@28021
   621
begin
haftmann@28021
   622
haftmann@28021
   623
lemma minus_inverts_plus2: "a + b = c \<Longrightarrow> c - a = b"
haftmann@28021
   624
  by (simp add: add_ac minus_inverts_plus1 [of b a])
haftmann@28021
   625
haftmann@28021
   626
lemma minus_minus_zero_inverts_plus2: "a + b = c \<Longrightarrow> a - c = 0 - b"
haftmann@28021
   627
  by (simp add: add_ac minus_minus_zero_inverts_plus1 [of b a])
haftmann@28021
   628
haftmann@28021
   629
end
haftmann@28021
   630
haftmann@28021
   631
class semiring_1_minus = semiring_1 + semiring_minus
haftmann@28021
   632
begin
haftmann@28021
   633
haftmann@28021
   634
lemma Dig_of_num_pos:
haftmann@28021
   635
  assumes "k + n = m"
haftmann@28021
   636
  shows "of_num m - of_num n = of_num k"
haftmann@28021
   637
  using assms by (simp add: of_num_plus minus_inverts_plus1)
haftmann@28021
   638
haftmann@28021
   639
lemma Dig_of_num_zero:
haftmann@28021
   640
  shows "of_num n - of_num n = 0"
haftmann@28021
   641
  by (rule minus_inverts_plus1) simp
haftmann@28021
   642
haftmann@28021
   643
lemma Dig_of_num_neg:
haftmann@28021
   644
  assumes "k + m = n"
haftmann@28021
   645
  shows "of_num m - of_num n = 0 - of_num k"
haftmann@28021
   646
  by (rule minus_minus_zero_inverts_plus1) (simp add: of_num_plus assms)
haftmann@28021
   647
haftmann@28021
   648
lemmas Dig_plus_eval =
haftmann@28021
   649
  of_num_plus of_num_eq_iff Dig_plus refl [of "1::num", THEN eqTrueI] num.inject
haftmann@28021
   650
haftmann@28021
   651
simproc_setup numeral_minus ("of_num m - of_num n") = {*
haftmann@28021
   652
  let
haftmann@28021
   653
    (*TODO proper implicit use of morphism via pattern antiquotations*)
haftmann@28021
   654
    fun cdest_of_num ct = (snd o split_last o snd o Drule.strip_comb) ct;
haftmann@28021
   655
    fun cdest_minus ct = case (rev o snd o Drule.strip_comb) ct of [n, m] => (m, n);
haftmann@28021
   656
    fun attach_num ct = (dest_num (Thm.term_of ct), ct);
haftmann@28021
   657
    fun cdifference t = (pairself (attach_num o cdest_of_num) o cdest_minus) t;
haftmann@28021
   658
    val simplify = MetaSimplifier.rewrite false (map mk_meta_eq @{thms Dig_plus_eval});
haftmann@28021
   659
    fun cert ck cl cj = @{thm eqTrueE} OF [@{thm meta_eq_to_obj_eq} OF [simplify (Drule.list_comb (@{cterm "op = :: num \<Rightarrow> _"},
haftmann@28021
   660
      [Drule.list_comb (@{cterm "op + :: num \<Rightarrow> _"}, [ck, cl]), cj]))]];
haftmann@28021
   661
  in fn phi => fn _ => fn ct => case try cdifference ct
haftmann@28021
   662
   of NONE => (NONE)
haftmann@28021
   663
    | SOME ((k, ck), (l, cl)) => SOME (let val j = k - l in if j = 0
haftmann@28021
   664
        then MetaSimplifier.rewrite false [mk_meta_eq (Morphism.thm phi @{thm Dig_of_num_zero})] ct
haftmann@28021
   665
        else mk_meta_eq (let
haftmann@28021
   666
          val cj = Thm.cterm_of (Thm.theory_of_cterm ct) (mk_num (abs j));
haftmann@28021
   667
        in
haftmann@28021
   668
          (if j > 0 then (Morphism.thm phi @{thm Dig_of_num_pos}) OF [cert cj cl ck]
haftmann@28021
   669
          else (Morphism.thm phi @{thm Dig_of_num_neg}) OF [cert cj ck cl])
haftmann@28021
   670
        end) end)
haftmann@28021
   671
  end
haftmann@28021
   672
*}
haftmann@28021
   673
haftmann@28021
   674
lemma Dig_of_num_minus_zero [numeral]:
haftmann@28021
   675
  "of_num n - 0 = of_num n"
haftmann@28021
   676
  by (simp add: minus_inverts_plus1)
haftmann@28021
   677
haftmann@28021
   678
lemma Dig_one_minus_zero [numeral]:
haftmann@28021
   679
  "1 - 0 = 1"
haftmann@28021
   680
  by (simp add: minus_inverts_plus1)
haftmann@28021
   681
haftmann@28021
   682
lemma Dig_one_minus_one [numeral]:
haftmann@28021
   683
  "1 - 1 = 0"
haftmann@28021
   684
  by (simp add: minus_inverts_plus1)
haftmann@28021
   685
haftmann@28021
   686
lemma Dig_of_num_minus_one [numeral]:
haftmann@28021
   687
  "of_num (Dig0 n) - 1 = of_num (dec (Dig0 n))"
haftmann@28021
   688
  "of_num (Dig1 n) - 1 = of_num (Dig0 n)"
haftmann@28021
   689
  by (auto intro: minus_inverts_plus1 simp add: Dig_dec_plus_one of_num.simps of_num_plus_one)
haftmann@28021
   690
haftmann@28021
   691
lemma Dig_one_minus_of_num [numeral]:
haftmann@28021
   692
  "1 - of_num (Dig0 n) = 0 - of_num (dec (Dig0 n))"
haftmann@28021
   693
  "1 - of_num (Dig1 n) = 0 - of_num (Dig0 n)"
haftmann@28021
   694
  by (auto intro: minus_minus_zero_inverts_plus1 simp add: Dig_dec_plus_one of_num.simps of_num_plus_one)
haftmann@28021
   695
haftmann@28021
   696
end
haftmann@28021
   697
haftmann@28021
   698
context ring_1
haftmann@28021
   699
begin
haftmann@28021
   700
haftmann@28021
   701
subclass semiring_1_minus
haftmann@28021
   702
  by unfold_locales (simp_all add: ring_simps)
haftmann@28021
   703
haftmann@28021
   704
lemma Dig_zero_minus_of_num [numeral]:
haftmann@28021
   705
  "0 - of_num n = - of_num n"
haftmann@28021
   706
  by simp
haftmann@28021
   707
haftmann@28021
   708
lemma Dig_zero_minus_one [numeral]:
haftmann@28021
   709
  "0 - 1 = - 1"
haftmann@28021
   710
  by simp
haftmann@28021
   711
haftmann@28021
   712
lemma Dig_uminus_uminus [numeral]:
haftmann@28021
   713
  "- (- of_num n) = of_num n"
haftmann@28021
   714
  by simp
haftmann@28021
   715
haftmann@28021
   716
lemma Dig_plus_uminus [numeral]:
haftmann@28021
   717
  "of_num m + - of_num n = of_num m - of_num n"
haftmann@28021
   718
  "- of_num m + of_num n = of_num n - of_num m"
haftmann@28021
   719
  "- of_num m + - of_num n = - (of_num m + of_num n)"
haftmann@28021
   720
  "of_num m - - of_num n = of_num m + of_num n"
haftmann@28021
   721
  "- of_num m - of_num n = - (of_num m + of_num n)"
haftmann@28021
   722
  "- of_num m - - of_num n = of_num n - of_num m"
haftmann@28021
   723
  by (simp_all add: diff_minus add_commute)
haftmann@28021
   724
haftmann@28021
   725
lemma Dig_times_uminus [numeral]:
haftmann@28021
   726
  "- of_num n * of_num m = - (of_num n * of_num m)"
haftmann@28021
   727
  "of_num n * - of_num m = - (of_num n * of_num m)"
haftmann@28021
   728
  "- of_num n * - of_num m = of_num n * of_num m"
haftmann@28021
   729
  by (simp_all add: minus_mult_left [symmetric] minus_mult_right [symmetric])
haftmann@28021
   730
haftmann@28021
   731
lemma of_int_of_num [numeral]: "of_int (of_num n) = of_num n"
haftmann@28021
   732
by (induct n)
haftmann@28021
   733
  (simp_all only: of_num.simps semiring_numeral_class.of_num.simps of_int_add, simp_all)
haftmann@28021
   734
haftmann@28021
   735
declare of_int_1 [numeral]
haftmann@28021
   736
haftmann@28021
   737
end
haftmann@28021
   738
haftmann@28021
   739
text {*
haftmann@28021
   740
  Greetings to @{typ nat}.
haftmann@28021
   741
*}
haftmann@28021
   742
haftmann@28021
   743
instance nat :: semiring_1_minus proof qed simp_all
haftmann@28021
   744
haftmann@28021
   745
lemma Suc_of_num [numeral]: "Suc (of_num n) = of_num (n + 1)"
haftmann@28021
   746
  unfolding of_num_plus_one [symmetric] by simp
haftmann@28021
   747
haftmann@28021
   748
lemma nat_number:
haftmann@28021
   749
  "1 = Suc 0"
haftmann@28021
   750
  "of_num 1 = Suc 0"
haftmann@28021
   751
  "of_num (Dig0 n) = Suc (of_num (dec (Dig0 n)))"
haftmann@28021
   752
  "of_num (Dig1 n) = Suc (of_num (Dig0 n))"
haftmann@28021
   753
  by (simp_all add: of_num.simps Dig_dec_plus_one Suc_of_num)
haftmann@28021
   754
haftmann@28021
   755
declare diff_0_eq_0 [numeral]
haftmann@28021
   756
haftmann@28021
   757
haftmann@28021
   758
subsection {* Code generator setup for @{typ int} *}
haftmann@28021
   759
haftmann@28021
   760
definition Pls :: "num \<Rightarrow> int" where
haftmann@28021
   761
  [simp, code post]: "Pls n = of_num n"
haftmann@28021
   762
haftmann@28021
   763
definition Mns :: "num \<Rightarrow> int" where
haftmann@28021
   764
  [simp, code post]: "Mns n = - of_num n"
haftmann@28021
   765
haftmann@28021
   766
code_datatype "0::int" Pls Mns
haftmann@28021
   767
haftmann@28021
   768
lemmas [code inline] = Pls_def [symmetric] Mns_def [symmetric]
haftmann@28021
   769
haftmann@28021
   770
definition sub :: "num \<Rightarrow> num \<Rightarrow> int" where
haftmann@28021
   771
  [simp, code func del]: "sub m n = (of_num m - of_num n)"
haftmann@28021
   772
haftmann@28021
   773
definition dup :: "int \<Rightarrow> int" where
haftmann@28021
   774
  [code func del]: "dup k = 2 * k"
haftmann@28021
   775
haftmann@28021
   776
lemma Dig_sub [code]:
haftmann@28021
   777
  "sub 1 1 = 0"
haftmann@28021
   778
  "sub (Dig0 m) 1 = of_num (dec (Dig0 m))"
haftmann@28021
   779
  "sub (Dig1 m) 1 = of_num (Dig0 m)"
haftmann@28021
   780
  "sub 1 (Dig0 n) = - of_num (dec (Dig0 n))"
haftmann@28021
   781
  "sub 1 (Dig1 n) = - of_num (Dig0 n)"
haftmann@28021
   782
  "sub (Dig0 m) (Dig0 n) = dup (sub m n)"
haftmann@28021
   783
  "sub (Dig1 m) (Dig1 n) = dup (sub m n)"
haftmann@28021
   784
  "sub (Dig1 m) (Dig0 n) = dup (sub m n) + 1"
haftmann@28021
   785
  "sub (Dig0 m) (Dig1 n) = dup (sub m n) - 1"
haftmann@28021
   786
  apply (simp_all add: dup_def ring_simps)
haftmann@28021
   787
  apply (simp_all add: of_num_plus Dig_one_plus_dec)[4]
haftmann@28021
   788
  apply (simp_all add: of_num.simps)
haftmann@28021
   789
  done
haftmann@28021
   790
haftmann@28021
   791
lemma dup_code [code]:
haftmann@28021
   792
  "dup 0 = 0"
haftmann@28021
   793
  "dup (Pls n) = Pls (Dig0 n)"
haftmann@28021
   794
  "dup (Mns n) = Mns (Dig0 n)"
haftmann@28021
   795
  by (simp_all add: dup_def of_num.simps)
haftmann@28021
   796
  
haftmann@28021
   797
lemma [code func, code func del]:
haftmann@28021
   798
  "(1 :: int) = 1"
haftmann@28021
   799
  "(op + :: int \<Rightarrow> int \<Rightarrow> int) = op +"
haftmann@28021
   800
  "(uminus :: int \<Rightarrow> int) = uminus"
haftmann@28021
   801
  "(op - :: int \<Rightarrow> int \<Rightarrow> int) = op -"
haftmann@28021
   802
  "(op * :: int \<Rightarrow> int \<Rightarrow> int) = op *"
haftmann@28021
   803
  "(op = :: int \<Rightarrow> int \<Rightarrow> bool) = op ="
haftmann@28021
   804
  "(op \<le> :: int \<Rightarrow> int \<Rightarrow> bool) = op \<le>"
haftmann@28021
   805
  "(op < :: int \<Rightarrow> int \<Rightarrow> bool) = op <"
haftmann@28021
   806
  by rule+
haftmann@28021
   807
haftmann@28021
   808
lemma one_int_code [code]:
haftmann@28021
   809
  "1 = Pls 1"
haftmann@28021
   810
  by (simp add: of_num_one)
haftmann@28021
   811
haftmann@28021
   812
lemma plus_int_code [code]:
haftmann@28021
   813
  "k + 0 = (k::int)"
haftmann@28021
   814
  "0 + l = (l::int)"
haftmann@28021
   815
  "Pls m + Pls n = Pls (m + n)"
haftmann@28021
   816
  "Pls m - Pls n = sub m n"
haftmann@28021
   817
  "Mns m + Mns n = Mns (m + n)"
haftmann@28021
   818
  "Mns m - Mns n = sub n m"
haftmann@28021
   819
  by (simp_all add: of_num_plus [symmetric])
haftmann@28021
   820
haftmann@28021
   821
lemma uminus_int_code [code]:
haftmann@28021
   822
  "uminus 0 = (0::int)"
haftmann@28021
   823
  "uminus (Pls m) = Mns m"
haftmann@28021
   824
  "uminus (Mns m) = Pls m"
haftmann@28021
   825
  by simp_all
haftmann@28021
   826
haftmann@28021
   827
lemma minus_int_code [code]:
haftmann@28021
   828
  "k - 0 = (k::int)"
haftmann@28021
   829
  "0 - l = uminus (l::int)"
haftmann@28021
   830
  "Pls m - Pls n = sub m n"
haftmann@28021
   831
  "Pls m - Mns n = Pls (m + n)"
haftmann@28021
   832
  "Mns m - Pls n = Mns (m + n)"
haftmann@28021
   833
  "Mns m - Mns n = sub n m"
haftmann@28021
   834
  by (simp_all add: of_num_plus [symmetric])
haftmann@28021
   835
haftmann@28021
   836
lemma times_int_code [code]:
haftmann@28021
   837
  "k * 0 = (0::int)"
haftmann@28021
   838
  "0 * l = (0::int)"
haftmann@28021
   839
  "Pls m * Pls n = Pls (m * n)"
haftmann@28021
   840
  "Pls m * Mns n = Mns (m * n)"
haftmann@28021
   841
  "Mns m * Pls n = Mns (m * n)"
haftmann@28021
   842
  "Mns m * Mns n = Pls (m * n)"
haftmann@28021
   843
  by (simp_all add: of_num_times [symmetric])
haftmann@28021
   844
haftmann@28021
   845
lemma eq_int_code [code]:
haftmann@28021
   846
  "0 = (0::int) \<longleftrightarrow> True"
haftmann@28021
   847
  "0 = Pls l \<longleftrightarrow> False"
haftmann@28021
   848
  "0 = Mns l \<longleftrightarrow> False"
haftmann@28021
   849
  "Pls k = 0 \<longleftrightarrow> False"
haftmann@28021
   850
  "Pls k = Pls l \<longleftrightarrow> k = l"
haftmann@28021
   851
  "Pls k = Mns l \<longleftrightarrow> False"
haftmann@28021
   852
  "Mns k = 0 \<longleftrightarrow> False"
haftmann@28021
   853
  "Mns k = Pls l \<longleftrightarrow> False"
haftmann@28021
   854
  "Mns k = Mns l \<longleftrightarrow> k = l"
haftmann@28021
   855
  using of_num_pos [of l, where ?'a = int] of_num_pos [of k, where ?'a = int]
haftmann@28021
   856
  by (simp_all add: of_num_eq_iff)
haftmann@28021
   857
haftmann@28021
   858
lemma less_eq_int_code [code]:
haftmann@28021
   859
  "0 \<le> (0::int) \<longleftrightarrow> True"
haftmann@28021
   860
  "0 \<le> Pls l \<longleftrightarrow> True"
haftmann@28021
   861
  "0 \<le> Mns l \<longleftrightarrow> False"
haftmann@28021
   862
  "Pls k \<le> 0 \<longleftrightarrow> False"
haftmann@28021
   863
  "Pls k \<le> Pls l \<longleftrightarrow> k \<le> l"
haftmann@28021
   864
  "Pls k \<le> Mns l \<longleftrightarrow> False"
haftmann@28021
   865
  "Mns k \<le> 0 \<longleftrightarrow> True"
haftmann@28021
   866
  "Mns k \<le> Pls l \<longleftrightarrow> True"
haftmann@28021
   867
  "Mns k \<le> Mns l \<longleftrightarrow> l \<le> k"
haftmann@28021
   868
  using of_num_pos [of l, where ?'a = int] of_num_pos [of k, where ?'a = int]
haftmann@28021
   869
  by (simp_all add: of_num_less_eq_iff)
haftmann@28021
   870
haftmann@28021
   871
lemma less_int_code [code]:
haftmann@28021
   872
  "0 < (0::int) \<longleftrightarrow> False"
haftmann@28021
   873
  "0 < Pls l \<longleftrightarrow> True"
haftmann@28021
   874
  "0 < Mns l \<longleftrightarrow> False"
haftmann@28021
   875
  "Pls k < 0 \<longleftrightarrow> False"
haftmann@28021
   876
  "Pls k < Pls l \<longleftrightarrow> k < l"
haftmann@28021
   877
  "Pls k < Mns l \<longleftrightarrow> False"
haftmann@28021
   878
  "Mns k < 0 \<longleftrightarrow> True"
haftmann@28021
   879
  "Mns k < Pls l \<longleftrightarrow> True"
haftmann@28021
   880
  "Mns k < Mns l \<longleftrightarrow> l < k"
haftmann@28021
   881
  using of_num_pos [of l, where ?'a = int] of_num_pos [of k, where ?'a = int]
haftmann@28021
   882
  by (simp_all add: of_num_less_iff)
haftmann@28021
   883
haftmann@28021
   884
lemma [code inline del]: "(0::int) \<equiv> Numeral0" by simp
haftmann@28021
   885
lemma [code inline del]: "(1::int) \<equiv> Numeral1" by simp
haftmann@28021
   886
declare zero_is_num_zero [code inline del]
haftmann@28021
   887
declare one_is_num_one [code inline del]
haftmann@28021
   888
haftmann@28021
   889
hide (open) const sub dup
haftmann@28021
   890
haftmann@28021
   891
haftmann@28021
   892
subsection {* Numeral equations as default simplification rules *}
haftmann@28021
   893
haftmann@28021
   894
text {* TODO.  Be more precise here with respect to subsumed facts. *}
haftmann@28021
   895
declare (in semiring_numeral) numeral [simp]
haftmann@28021
   896
declare (in semiring_1) numeral [simp]
haftmann@28021
   897
declare (in semiring_char_0) numeral [simp]
haftmann@28021
   898
declare (in ring_1) numeral [simp]
haftmann@28021
   899
thm numeral
haftmann@28021
   900
haftmann@28021
   901
haftmann@28021
   902
text {* Toy examples *}
haftmann@28021
   903
haftmann@28021
   904
definition "bar \<longleftrightarrow> #4 * #2 + #7 = (#8 :: nat) \<and> #4 * #2 + #7 \<ge> (#8 :: int) - #3"
haftmann@28021
   905
code_thms bar
haftmann@28021
   906
export_code bar in Haskell file -
haftmann@28021
   907
export_code bar in OCaml module_name Foo file -
haftmann@28021
   908
ML {* @{code bar} *}
haftmann@28021
   909
haftmann@28021
   910
end