src/HOL/ex/Numeral.thy
author wenzelm
Wed, 23 Mar 2011 20:57:37 +0100
changeset 42078 d5bf0ce40bd7
parent 41489 8e2b8649507d
child 42245 29e3967550d5
permissions -rw-r--r--
isolate change of Proofterm.proofs in TPTP.thy from rest of session; qualified Proofterm.proofs to facilitate grepping;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/ex/Numeral.thy
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann
29947
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
     3
*)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
     4
29947
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
     5
header {* An experimental alternative numeral representation. *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
     6
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
     7
theory Numeral
33346
c5cd93763e71 recovered from 7a1f597f454e, simplified imports;
wenzelm
parents: 33297
diff changeset
     8
imports Main
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
     9
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    10
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    11
subsection {* The @{text num} type *}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    12
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    13
datatype num = One | Dig0 num | Dig1 num
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    14
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    15
text {* Increment function for type @{typ num} *}
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    16
31021
53642251a04f farewell to class recpower
haftmann
parents: 30792
diff changeset
    17
primrec inc :: "num \<Rightarrow> num" where
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    18
  "inc One = Dig0 One"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    19
| "inc (Dig0 x) = Dig1 x"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    20
| "inc (Dig1 x) = Dig0 (inc x)"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    21
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    22
text {* Converting between type @{typ num} and type @{typ nat} *}
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    23
31021
53642251a04f farewell to class recpower
haftmann
parents: 30792
diff changeset
    24
primrec nat_of_num :: "num \<Rightarrow> nat" where
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    25
  "nat_of_num One = Suc 0"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    26
| "nat_of_num (Dig0 x) = nat_of_num x + nat_of_num x"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    27
| "nat_of_num (Dig1 x) = Suc (nat_of_num x + nat_of_num x)"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    28
31021
53642251a04f farewell to class recpower
haftmann
parents: 30792
diff changeset
    29
primrec num_of_nat :: "nat \<Rightarrow> num" where
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    30
  "num_of_nat 0 = One"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    31
| "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    32
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
    33
lemma nat_of_num_pos: "0 < nat_of_num x"
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    34
  by (induct x) simp_all
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    35
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
    36
lemma nat_of_num_neq_0: "nat_of_num x \<noteq> 0"
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    37
  by (induct x) simp_all
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    38
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    39
lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    40
  by (induct x) simp_all
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    41
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    42
lemma num_of_nat_double:
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    43
  "0 < n \<Longrightarrow> num_of_nat (n + n) = Dig0 (num_of_nat n)"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    44
  by (induct n) simp_all
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    45
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    46
text {*
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    47
  Type @{typ num} is isomorphic to the strictly positive
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    48
  natural numbers.
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    49
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    50
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    51
lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x"
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
    52
  by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    53
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    54
lemma num_of_nat_inverse: "0 < n \<Longrightarrow> nat_of_num (num_of_nat n) = n"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    55
  by (induct n) (simp_all add: nat_of_num_inc)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    56
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
    57
lemma num_eq_iff: "x = y \<longleftrightarrow> nat_of_num x = nat_of_num y"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    58
proof
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    59
  assume "nat_of_num x = nat_of_num y"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    60
  then have "num_of_nat (nat_of_num x) = num_of_nat (nat_of_num y)" by simp
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    61
  then show "x = y" by (simp add: nat_of_num_inverse)
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    62
qed simp
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    63
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    64
lemma num_induct [case_names One inc]:
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    65
  fixes P :: "num \<Rightarrow> bool"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    66
  assumes One: "P One"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    67
    and inc: "\<And>x. P x \<Longrightarrow> P (inc x)"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    68
  shows "P x"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    69
proof -
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    70
  obtain n where n: "Suc n = nat_of_num x"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    71
    by (cases "nat_of_num x", simp_all add: nat_of_num_neq_0)
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    72
  have "P (num_of_nat (Suc n))"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    73
  proof (induct n)
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    74
    case 0 show ?case using One by simp
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    75
  next
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    76
    case (Suc n)
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    77
    then have "P (inc (num_of_nat (Suc n)))" by (rule inc)
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    78
    then show "P (num_of_nat (Suc (Suc n)))" by simp
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    79
  qed
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    80
  with n show "P x"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
    81
    by (simp add: nat_of_num_inverse)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    82
qed
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    83
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    84
text {*
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    85
  From now on, there are two possible models for @{typ num}: as
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    86
  positive naturals (rule @{text "num_induct"}) and as digit
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    87
  representation (rules @{text "num.induct"}, @{text "num.cases"}).
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    88
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    89
  It is not entirely clear in which context it is better to use the
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
    90
  one or the other, or whether the construction should be reversed.
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    91
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    92
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    93
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
    94
subsection {* Numeral operations *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    95
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
    96
ML {*
31902
862ae16a799d renamed NamedThmsFun to Named_Thms;
wenzelm
parents: 31029
diff changeset
    97
structure Dig_Simps = Named_Thms
862ae16a799d renamed NamedThmsFun to Named_Thms;
wenzelm
parents: 31029
diff changeset
    98
(
862ae16a799d renamed NamedThmsFun to Named_Thms;
wenzelm
parents: 31029
diff changeset
    99
  val name = "numeral"
38764
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38054
diff changeset
   100
  val description = "simplification rules for numerals"
31902
862ae16a799d renamed NamedThmsFun to Named_Thms;
wenzelm
parents: 31029
diff changeset
   101
)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   102
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   103
31902
862ae16a799d renamed NamedThmsFun to Named_Thms;
wenzelm
parents: 31029
diff changeset
   104
setup Dig_Simps.setup
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   105
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   106
instantiation num :: "{plus,times,ord}"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   107
begin
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   108
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   109
definition plus_num :: "num \<Rightarrow> num \<Rightarrow> num" where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 36176
diff changeset
   110
  "m + n = num_of_nat (nat_of_num m + nat_of_num n)"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   111
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   112
definition times_num :: "num \<Rightarrow> num \<Rightarrow> num" where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 36176
diff changeset
   113
  "m * n = num_of_nat (nat_of_num m * nat_of_num n)"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   114
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   115
definition less_eq_num :: "num \<Rightarrow> num \<Rightarrow> bool" where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 36176
diff changeset
   116
  "m \<le> n \<longleftrightarrow> nat_of_num m \<le> nat_of_num n"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   117
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   118
definition less_num :: "num \<Rightarrow> num \<Rightarrow> bool" where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 36176
diff changeset
   119
  "m < n \<longleftrightarrow> nat_of_num m < nat_of_num n"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   120
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   121
instance ..
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   122
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   123
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   124
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   125
lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   126
  unfolding plus_num_def
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   127
  by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos)
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   128
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   129
lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   130
  unfolding times_num_def
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   131
  by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   132
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   133
lemma Dig_plus [numeral, simp, code]:
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   134
  "One + One = Dig0 One"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   135
  "One + Dig0 m = Dig1 m"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   136
  "One + Dig1 m = Dig0 (m + One)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   137
  "Dig0 n + One = Dig1 n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   138
  "Dig0 n + Dig0 m = Dig0 (n + m)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   139
  "Dig0 n + Dig1 m = Dig1 (n + m)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   140
  "Dig1 n + One = Dig0 (n + One)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   141
  "Dig1 n + Dig0 m = Dig1 (n + m)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   142
  "Dig1 n + Dig1 m = Dig0 (n + m + One)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   143
  by (simp_all add: num_eq_iff nat_of_num_add)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   144
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   145
lemma Dig_times [numeral, simp, code]:
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   146
  "One * One = One"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   147
  "One * Dig0 n = Dig0 n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   148
  "One * Dig1 n = Dig1 n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   149
  "Dig0 n * One = Dig0 n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   150
  "Dig0 n * Dig0 m = Dig0 (n * Dig0 m)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   151
  "Dig0 n * Dig1 m = Dig0 (n * Dig1 m)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   152
  "Dig1 n * One = Dig1 n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   153
  "Dig1 n * Dig0 m = Dig0 (n * Dig0 m + m)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   154
  "Dig1 n * Dig1 m = Dig1 (n * Dig1 m + m)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   155
  by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   156
                    left_distrib right_distrib)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   157
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   158
lemma less_eq_num_code [numeral, simp, code]:
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   159
  "One \<le> n \<longleftrightarrow> True"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   160
  "Dig0 m \<le> One \<longleftrightarrow> False"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   161
  "Dig1 m \<le> One \<longleftrightarrow> False"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   162
  "Dig0 m \<le> Dig0 n \<longleftrightarrow> m \<le> n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   163
  "Dig0 m \<le> Dig1 n \<longleftrightarrow> m \<le> n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   164
  "Dig1 m \<le> Dig1 n \<longleftrightarrow> m \<le> n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   165
  "Dig1 m \<le> Dig0 n \<longleftrightarrow> m < n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   166
  using nat_of_num_pos [of n] nat_of_num_pos [of m]
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   167
  by (auto simp add: less_eq_num_def less_num_def)
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   168
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   169
lemma less_num_code [numeral, simp, code]:
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   170
  "m < One \<longleftrightarrow> False"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   171
  "One < One \<longleftrightarrow> False"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   172
  "One < Dig0 n \<longleftrightarrow> True"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   173
  "One < Dig1 n \<longleftrightarrow> True"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   174
  "Dig0 m < Dig0 n \<longleftrightarrow> m < n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   175
  "Dig0 m < Dig1 n \<longleftrightarrow> m \<le> n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   176
  "Dig1 m < Dig1 n \<longleftrightarrow> m < n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   177
  "Dig1 m < Dig0 n \<longleftrightarrow> m < n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   178
  using nat_of_num_pos [of n] nat_of_num_pos [of m]
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   179
  by (auto simp add: less_eq_num_def less_num_def)
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   180
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   181
text {* Rules using @{text One} and @{text inc} as constructors *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   182
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   183
lemma add_One: "x + One = inc x"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   184
  by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc)
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   185
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   186
lemma add_inc: "x + inc y = inc (x + y)"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   187
  by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc)
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   188
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   189
lemma mult_One: "x * One = x"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   190
  by (simp add: num_eq_iff nat_of_num_mult)
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   191
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   192
lemma mult_inc: "x * inc y = x * y + x"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   193
  by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc)
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   194
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   195
text {* A double-and-decrement function *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   196
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   197
primrec DigM :: "num \<Rightarrow> num" where
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   198
  "DigM One = One"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   199
| "DigM (Dig0 n) = Dig1 (DigM n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   200
| "DigM (Dig1 n) = Dig1 (Dig0 n)"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   201
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   202
lemma DigM_plus_one: "DigM n + One = Dig0 n"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   203
  by (induct n) simp_all
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   204
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   205
lemma add_One_commute: "One + n = n + One"
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   206
  by (induct n) simp_all
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   207
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   208
lemma one_plus_DigM: "One + DigM n = Dig0 n"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   209
  by (simp add: add_One_commute DigM_plus_one)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   210
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   211
text {* Squaring and exponentiation *}
29947
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
   212
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
   213
primrec square :: "num \<Rightarrow> num" where
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
   214
  "square One = One"
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
   215
| "square (Dig0 n) = Dig0 (Dig0 (square n))"
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
   216
| "square (Dig1 n) = Dig1 (Dig0 (square n + n))"
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
   217
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   218
primrec pow :: "num \<Rightarrow> num \<Rightarrow> num" where
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   219
  "pow x One = x"
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   220
| "pow x (Dig0 y) = square (pow x y)"
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   221
| "pow x (Dig1 y) = x * square (pow x y)"
29947
0a51765d2084 tune section headings; add square function
huffman
parents: 29946
diff changeset
   222
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   223
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   224
subsection {* Binary numerals *}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   225
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   226
text {*
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   227
  We embed binary representations into a generic algebraic
29934
5d81dd706206 tuned texts
haftmann
parents: 29667
diff changeset
   228
  structure using @{text of_num}.
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   229
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   230
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   231
class semiring_numeral = semiring + monoid_mult
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   232
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   233
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   234
primrec of_num :: "num \<Rightarrow> 'a" where
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   235
  of_num_One [numeral]: "of_num One = 1"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   236
| "of_num (Dig0 n) = of_num n + of_num n"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   237
| "of_num (Dig1 n) = of_num n + of_num n + 1"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   238
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   239
lemma of_num_inc: "of_num (inc n) = of_num n + 1"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   240
  by (induct n) (simp_all add: add_ac)
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
   241
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   242
lemma of_num_add: "of_num (m + n) = of_num m + of_num n"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   243
  by (induct n rule: num_induct) (simp_all add: add_One add_inc of_num_inc add_ac)
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   244
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   245
lemma of_num_mult: "of_num (m * n) = of_num m * of_num n"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   246
  by (induct n rule: num_induct) (simp_all add: mult_One mult_inc of_num_add of_num_inc right_distrib)
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   247
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   248
declare of_num.simps [simp del]
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   249
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   250
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   251
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   252
ML {*
31027
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   253
fun mk_num k =
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   254
  if k > 1 then
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   255
    let
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   256
      val (l, b) = Integer.div_mod k 2;
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   257
      val bit = (if b = 0 then @{term Dig0} else @{term Dig1});
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   258
    in bit $ (mk_num l) end
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   259
  else if k = 1 then @{term One}
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   260
  else error ("mk_num " ^ string_of_int k);
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   261
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
   262
fun dest_num @{term One} = 1
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   263
  | dest_num (@{term Dig0} $ n) = 2 * dest_num n
31027
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   264
  | dest_num (@{term Dig1} $ n) = 2 * dest_num n + 1
b5a35bfb3dab detect error cases in mk_num, dest_num
huffman
parents: 31026
diff changeset
   265
  | dest_num t = raise TERM ("dest_num", [t]);
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   266
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   267
fun mk_numeral phi T k = Morphism.term phi (Const (@{const_name of_num}, @{typ num} --> T))
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   268
  $ mk_num k
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   269
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   270
fun dest_numeral phi (u $ t) =
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   271
  if Term.aconv_untyped (u, Morphism.term phi (Const (@{const_name of_num}, dummyT)))
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   272
  then (range_type (fastype_of u), dest_num t)
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   273
  else raise TERM ("dest_numeral", [u, t]);
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   274
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   275
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   276
syntax
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   277
  "_Numerals" :: "xnum \<Rightarrow> 'a"    ("_")
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   278
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   279
parse_translation {*
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   280
let
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   281
  fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2)
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
   282
     of (0, 1) => Const (@{const_name One}, dummyT)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   283
      | (n, 0) => Const (@{const_name Dig0}, dummyT) $ num_of_int n
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   284
      | (n, 1) => Const (@{const_name Dig1}, dummyT) $ num_of_int n
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   285
    else raise Match;
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   286
  fun numeral_tr [Free (num, _)] =
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   287
        let
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   288
          val {leading_zeros, value, ...} = Syntax.read_xnum num;
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   289
          val _ = leading_zeros = 0 andalso value > 0
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   290
            orelse error ("Bad numeral: " ^ num);
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   291
        in Const (@{const_name of_num}, @{typ num} --> dummyT) $ num_of_int value end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   292
    | numeral_tr ts = raise TERM ("numeral_tr", ts);
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 35028
diff changeset
   293
in [(@{syntax_const "_Numerals"}, numeral_tr)] end
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   294
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   295
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38923
diff changeset
   296
typed_print_translation (advanced) {*
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   297
let
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   298
  fun dig b n = b + 2 * n; 
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   299
  fun int_of_num' (Const (@{const_syntax Dig0}, _) $ n) =
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   300
        dig 0 (int_of_num' n)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   301
    | int_of_num' (Const (@{const_syntax Dig1}, _) $ n) =
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   302
        dig 1 (int_of_num' n)
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
   303
    | int_of_num' (Const (@{const_syntax One}, _)) = 1;
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38923
diff changeset
   304
  fun num_tr' ctxt show_sorts T [n] =
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   305
    let
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   306
      val k = int_of_num' n;
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 35028
diff changeset
   307
      val t' = Syntax.const @{syntax_const "_Numerals"} $ Syntax.free ("#" ^ string_of_int k);
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38923
diff changeset
   308
    in
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38923
diff changeset
   309
      case T of
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38923
diff changeset
   310
        Type (@{type_name fun}, [_, T']) =>
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38923
diff changeset
   311
          if not (Config.get ctxt show_types) andalso can Term.dest_Type T' then t'
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38923
diff changeset
   312
          else Syntax.const Syntax.constrainC $ t' $ Syntax.term_of_typ show_sorts T'
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   313
      | T' => if T' = dummyT then t' else raise Match
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   314
    end;
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   315
in [(@{const_syntax of_num}, num_tr')] end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   316
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   317
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   318
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   319
subsection {* Class-specific numeral rules *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   320
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   321
subsubsection {* Class @{text semiring_numeral} *}
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   322
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   323
context semiring_numeral
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   324
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   325
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
   326
abbreviation "Num1 \<equiv> of_num One"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   327
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   328
text {*
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   329
  Alas, there is still the duplication of @{term 1}, although the
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   330
  duplicated @{term 0} has disappeared.  We could get rid of it by
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   331
  replacing the constructor @{term 1} in @{typ num} by two
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   332
  constructors @{text two} and @{text three}, resulting in a further
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   333
  blow-up.  But it could be worth the effort.
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   334
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   335
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   336
lemma of_num_plus_one [numeral]:
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
   337
  "of_num n + 1 = of_num (n + One)"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   338
  by (simp only: of_num_add of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   339
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   340
lemma of_num_one_plus [numeral]:
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   341
  "1 + of_num n = of_num (One + n)"
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   342
  by (simp only: of_num_add of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   343
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   344
lemma of_num_plus [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   345
  "of_num m + of_num n = of_num (m + n)"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   346
  by (simp only: of_num_add)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   347
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   348
lemma of_num_times_one [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   349
  "of_num n * 1 = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   350
  by simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   351
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   352
lemma of_num_one_times [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   353
  "1 * of_num n = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   354
  by simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   355
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   356
lemma of_num_times [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   357
  "of_num m * of_num n = of_num (m * n)"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   358
  unfolding of_num_mult ..
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   359
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   360
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   361
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   362
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   363
subsubsection {* Structures with a zero: class @{text semiring_1} *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   364
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   365
context semiring_1
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   366
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   367
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   368
subclass semiring_numeral ..
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   369
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   370
lemma of_nat_of_num [numeral]: "of_nat (of_num n) = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   371
  by (induct n)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   372
    (simp_all add: semiring_numeral_class.of_num.simps of_num.simps add_ac)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   373
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   374
declare of_nat_1 [numeral]
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   375
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   376
lemma Dig_plus_zero [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   377
  "0 + 1 = 1"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   378
  "0 + of_num n = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   379
  "1 + 0 = 1"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   380
  "of_num n + 0 = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   381
  by simp_all
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   382
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   383
lemma Dig_times_zero [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   384
  "0 * 1 = 0"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   385
  "0 * of_num n = 0"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   386
  "1 * 0 = 0"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   387
  "of_num n * 0 = 0"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   388
  by simp_all
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   389
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   390
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   391
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   392
lemma nat_of_num_of_num: "nat_of_num = of_num"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   393
proof
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   394
  fix n
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
   395
  have "of_num n = nat_of_num n"
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
   396
    by (induct n) (simp_all add: of_num.simps)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   397
  then show "nat_of_num n = of_num n" by simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   398
qed
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   399
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   400
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   401
subsubsection {* Equality: class @{text semiring_char_0} *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   402
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   403
context semiring_char_0
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   404
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   405
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   406
lemma of_num_eq_iff [numeral]: "of_num m = of_num n \<longleftrightarrow> m = n"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   407
  unfolding of_nat_of_num [symmetric] nat_of_num_of_num [symmetric]
29943
922b931fd2eb datatype num = One | Dig0 num | Dig1 num
huffman
parents: 29942
diff changeset
   408
    of_nat_eq_iff num_eq_iff ..
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   409
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   410
lemma of_num_eq_one_iff [numeral]: "of_num n = 1 \<longleftrightarrow> n = One"
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   411
  using of_num_eq_iff [of n One] by (simp add: of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   412
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   413
lemma one_eq_of_num_iff [numeral]: "1 = of_num n \<longleftrightarrow> One = n"
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   414
  using of_num_eq_iff [of One n] by (simp add: of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   415
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   416
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   417
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   418
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   419
subsubsection {* Comparisons: class @{text linordered_semidom} *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   420
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   421
text {*
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   422
  Perhaps the underlying structure could even 
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   423
  be more general than @{text linordered_semidom}.
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   424
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   425
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   426
context linordered_semidom
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   427
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   428
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   429
lemma of_num_pos [numeral]: "0 < of_num n"
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   430
  by (induct n) (simp_all add: of_num.simps add_pos_pos)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   431
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   432
lemma of_num_not_zero [numeral]: "of_num n \<noteq> 0"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   433
  using of_num_pos [of n] by simp
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   434
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   435
lemma of_num_less_eq_iff [numeral]: "of_num m \<le> of_num n \<longleftrightarrow> m \<le> n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   436
proof -
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   437
  have "of_nat (of_num m) \<le> of_nat (of_num n) \<longleftrightarrow> m \<le> n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   438
    unfolding less_eq_num_def nat_of_num_of_num of_nat_le_iff ..
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   439
  then show ?thesis by (simp add: of_nat_of_num)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   440
qed
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   441
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   442
lemma of_num_less_eq_one_iff [numeral]: "of_num n \<le> 1 \<longleftrightarrow> n \<le> One"
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   443
  using of_num_less_eq_iff [of n One] by (simp add: of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   444
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   445
lemma one_less_eq_of_num_iff [numeral]: "1 \<le> of_num n"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   446
  using of_num_less_eq_iff [of One n] by (simp add: of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   447
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   448
lemma of_num_less_iff [numeral]: "of_num m < of_num n \<longleftrightarrow> m < n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   449
proof -
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   450
  have "of_nat (of_num m) < of_nat (of_num n) \<longleftrightarrow> m < n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   451
    unfolding less_num_def nat_of_num_of_num of_nat_less_iff ..
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   452
  then show ?thesis by (simp add: of_nat_of_num)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   453
qed
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   454
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   455
lemma of_num_less_one_iff [numeral]: "\<not> of_num n < 1"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   456
  using of_num_less_iff [of n One] by (simp add: of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   457
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   458
lemma one_less_of_num_iff [numeral]: "1 < of_num n \<longleftrightarrow> One < n"
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   459
  using of_num_less_iff [of One n] by (simp add: of_num_One)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   460
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   461
lemma of_num_nonneg [numeral]: "0 \<le> of_num n"
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   462
  by (induct n) (simp_all add: of_num.simps add_nonneg_nonneg)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   463
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   464
lemma of_num_less_zero_iff [numeral]: "\<not> of_num n < 0"
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   465
  by (simp add: not_less of_num_nonneg)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   466
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   467
lemma of_num_le_zero_iff [numeral]: "\<not> of_num n \<le> 0"
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   468
  by (simp add: not_le of_num_pos)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   469
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   470
end
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   471
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   472
context linordered_idom
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   473
begin
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   474
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   475
lemma minus_of_num_less_of_num_iff: "- of_num m < of_num n"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   476
proof -
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   477
  have "- of_num m < 0" by (simp add: of_num_pos)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   478
  also have "0 < of_num n" by (simp add: of_num_pos)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   479
  finally show ?thesis .
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   480
qed
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   481
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   482
lemma minus_of_num_not_equal_of_num: "- of_num m \<noteq> of_num n"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   483
  using minus_of_num_less_of_num_iff [of m n] by simp
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   484
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   485
lemma minus_of_num_less_one_iff: "- of_num n < 1"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   486
  using minus_of_num_less_of_num_iff [of n One] by (simp add: of_num_One)
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   487
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   488
lemma minus_one_less_of_num_iff: "- 1 < of_num n"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   489
  using minus_of_num_less_of_num_iff [of One n] by (simp add: of_num_One)
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   490
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   491
lemma minus_one_less_one_iff: "- 1 < 1"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   492
  using minus_of_num_less_of_num_iff [of One One] by (simp add: of_num_One)
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   493
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   494
lemma minus_of_num_le_of_num_iff: "- of_num m \<le> of_num n"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   495
  by (simp add: less_imp_le minus_of_num_less_of_num_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   496
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   497
lemma minus_of_num_le_one_iff: "- of_num n \<le> 1"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   498
  by (simp add: less_imp_le minus_of_num_less_one_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   499
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   500
lemma minus_one_le_of_num_iff: "- 1 \<le> of_num n"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   501
  by (simp add: less_imp_le minus_one_less_of_num_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   502
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   503
lemma minus_one_le_one_iff: "- 1 \<le> 1"
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   504
  by (simp add: less_imp_le minus_one_less_one_iff)
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   505
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   506
lemma of_num_le_minus_of_num_iff: "\<not> of_num m \<le> - of_num n"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   507
  by (simp add: not_le minus_of_num_less_of_num_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   508
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   509
lemma one_le_minus_of_num_iff: "\<not> 1 \<le> - of_num n"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   510
  by (simp add: not_le minus_of_num_less_one_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   511
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   512
lemma of_num_le_minus_one_iff: "\<not> of_num n \<le> - 1"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   513
  by (simp add: not_le minus_one_less_of_num_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   514
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   515
lemma one_le_minus_one_iff: "\<not> 1 \<le> - 1"
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   516
  by (simp add: not_le minus_one_less_one_iff)
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   517
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   518
lemma of_num_less_minus_of_num_iff: "\<not> of_num m < - of_num n"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   519
  by (simp add: not_less minus_of_num_le_of_num_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   520
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   521
lemma one_less_minus_of_num_iff: "\<not> 1 < - of_num n"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   522
  by (simp add: not_less minus_of_num_le_one_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   523
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   524
lemma of_num_less_minus_one_iff: "\<not> of_num n < - 1"
29991
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   525
  by (simp add: not_less minus_one_le_of_num_iff)
c60ace776315 add more ordering lemmas
huffman
parents: 29954
diff changeset
   526
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   527
lemma one_less_minus_one_iff: "\<not> 1 < - 1"
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   528
  by (simp add: not_less minus_one_le_one_iff)
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   529
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   530
lemmas le_signed_numeral_special [numeral] =
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   531
  minus_of_num_le_of_num_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   532
  minus_of_num_le_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   533
  minus_one_le_of_num_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   534
  minus_one_le_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   535
  of_num_le_minus_of_num_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   536
  one_le_minus_of_num_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   537
  of_num_le_minus_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   538
  one_le_minus_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   539
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   540
lemmas less_signed_numeral_special [numeral] =
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   541
  minus_of_num_less_of_num_iff
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   542
  minus_of_num_not_equal_of_num
30792
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   543
  minus_of_num_less_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   544
  minus_one_less_of_num_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   545
  minus_one_less_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   546
  of_num_less_minus_of_num_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   547
  one_less_minus_of_num_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   548
  of_num_less_minus_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   549
  one_less_minus_one_iff
809c38c1a26c add more lemmas for signed comparisons
huffman
parents: 29991
diff changeset
   550
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   551
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   552
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   553
subsubsection {* Structures with subtraction: class @{text semiring_1_minus} *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   554
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   555
class semiring_minus = semiring + minus + zero +
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   556
  assumes minus_inverts_plus1: "a + b = c \<Longrightarrow> c - b = a"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   557
  assumes minus_minus_zero_inverts_plus1: "a + b = c \<Longrightarrow> b - c = 0 - a"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   558
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   559
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   560
lemma minus_inverts_plus2: "a + b = c \<Longrightarrow> c - a = b"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   561
  by (simp add: add_ac minus_inverts_plus1 [of b a])
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   562
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   563
lemma minus_minus_zero_inverts_plus2: "a + b = c \<Longrightarrow> a - c = 0 - b"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   564
  by (simp add: add_ac minus_minus_zero_inverts_plus1 [of b a])
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   565
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   566
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   567
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   568
class semiring_1_minus = semiring_1 + semiring_minus
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   569
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   570
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   571
lemma Dig_of_num_pos:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   572
  assumes "k + n = m"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   573
  shows "of_num m - of_num n = of_num k"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   574
  using assms by (simp add: of_num_plus minus_inverts_plus1)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   575
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   576
lemma Dig_of_num_zero:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   577
  shows "of_num n - of_num n = 0"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   578
  by (rule minus_inverts_plus1) simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   579
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   580
lemma Dig_of_num_neg:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   581
  assumes "k + m = n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   582
  shows "of_num m - of_num n = 0 - of_num k"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   583
  by (rule minus_minus_zero_inverts_plus1) (simp add: of_num_plus assms)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   584
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   585
lemmas Dig_plus_eval =
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
   586
  of_num_plus of_num_eq_iff Dig_plus refl [of One, THEN eqTrueI] num.inject
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   587
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   588
simproc_setup numeral_minus ("of_num m - of_num n") = {*
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   589
  let
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   590
    (*TODO proper implicit use of morphism via pattern antiquotations*)
41489
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41228
diff changeset
   591
    fun cdest_of_num ct = (List.last o snd o Drule.strip_comb) ct;
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   592
    fun cdest_minus ct = case (rev o snd o Drule.strip_comb) ct of [n, m] => (m, n);
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   593
    fun attach_num ct = (dest_num (Thm.term_of ct), ct);
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   594
    fun cdifference t = (pairself (attach_num o cdest_of_num) o cdest_minus) t;
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 39272
diff changeset
   595
    val simplify = Raw_Simplifier.rewrite false (map mk_meta_eq @{thms Dig_plus_eval});
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   596
    fun cert ck cl cj = @{thm eqTrueE} OF [@{thm meta_eq_to_obj_eq}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   597
      OF [simplify (Drule.list_comb (@{cterm "op = :: num \<Rightarrow> _"},
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   598
        [Drule.list_comb (@{cterm "op + :: num \<Rightarrow> _"}, [ck, cl]), cj]))]];
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   599
  in fn phi => fn _ => fn ct => case try cdifference ct
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   600
   of NONE => (NONE)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   601
    | SOME ((k, ck), (l, cl)) => SOME (let val j = k - l in if j = 0
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 39272
diff changeset
   602
        then Raw_Simplifier.rewrite false [mk_meta_eq (Morphism.thm phi @{thm Dig_of_num_zero})] ct
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   603
        else mk_meta_eq (let
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   604
          val cj = Thm.cterm_of (Thm.theory_of_cterm ct) (mk_num (abs j));
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   605
        in
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   606
          (if j > 0 then (Morphism.thm phi @{thm Dig_of_num_pos}) OF [cert cj cl ck]
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   607
          else (Morphism.thm phi @{thm Dig_of_num_neg}) OF [cert cj ck cl])
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   608
        end) end)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   609
  end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   610
*}
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   611
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   612
lemma Dig_of_num_minus_zero [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   613
  "of_num n - 0 = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   614
  by (simp add: minus_inverts_plus1)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   615
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   616
lemma Dig_one_minus_zero [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   617
  "1 - 0 = 1"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   618
  by (simp add: minus_inverts_plus1)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   619
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   620
lemma Dig_one_minus_one [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   621
  "1 - 1 = 0"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   622
  by (simp add: minus_inverts_plus1)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   623
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   624
lemma Dig_of_num_minus_one [numeral]:
29941
b951d80774d5 replace dec with double-and-decrement function
huffman
parents: 29667
diff changeset
   625
  "of_num (Dig0 n) - 1 = of_num (DigM n)"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   626
  "of_num (Dig1 n) - 1 = of_num (Dig0 n)"
29941
b951d80774d5 replace dec with double-and-decrement function
huffman
parents: 29667
diff changeset
   627
  by (auto intro: minus_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   628
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   629
lemma Dig_one_minus_of_num [numeral]:
29941
b951d80774d5 replace dec with double-and-decrement function
huffman
parents: 29667
diff changeset
   630
  "1 - of_num (Dig0 n) = 0 - of_num (DigM n)"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   631
  "1 - of_num (Dig1 n) = 0 - of_num (Dig0 n)"
29941
b951d80774d5 replace dec with double-and-decrement function
huffman
parents: 29667
diff changeset
   632
  by (auto intro: minus_minus_zero_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   633
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   634
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   635
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   636
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   637
subsubsection {* Structures with negation: class @{text ring_1} *}
29945
75df553549b1 rearrange subsections
huffman
parents: 29944
diff changeset
   638
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   639
context ring_1
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   640
begin
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   641
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   642
subclass semiring_1_minus proof
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   643
qed (simp_all add: algebra_simps)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   644
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   645
lemma Dig_zero_minus_of_num [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   646
  "0 - of_num n = - of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   647
  by simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   648
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   649
lemma Dig_zero_minus_one [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   650
  "0 - 1 = - 1"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   651
  by simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   652
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   653
lemma Dig_uminus_uminus [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   654
  "- (- of_num n) = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   655
  by simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   656
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   657
lemma Dig_plus_uminus [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   658
  "of_num m + - of_num n = of_num m - of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   659
  "- of_num m + of_num n = of_num n - of_num m"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   660
  "- of_num m + - of_num n = - (of_num m + of_num n)"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   661
  "of_num m - - of_num n = of_num m + of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   662
  "- of_num m - of_num n = - (of_num m + of_num n)"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   663
  "- of_num m - - of_num n = of_num n - of_num m"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   664
  by (simp_all add: diff_minus add_commute)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   665
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   666
lemma Dig_times_uminus [numeral]:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   667
  "- of_num n * of_num m = - (of_num n * of_num m)"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   668
  "of_num n * - of_num m = - (of_num n * of_num m)"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   669
  "- of_num n * - of_num m = of_num n * of_num m"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   670
  by simp_all
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   671
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   672
lemma of_int_of_num [numeral]: "of_int (of_num n) = of_num n"
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   673
by (induct n)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   674
  (simp_all only: of_num.simps semiring_numeral_class.of_num.simps of_int_add, simp_all)
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   675
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   676
declare of_int_1 [numeral]
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   677
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   678
end
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   679
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   680
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   681
subsubsection {* Structures with exponentiation *}
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   682
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   683
lemma of_num_square: "of_num (square x) = of_num x * of_num x"
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   684
by (induct x)
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   685
   (simp_all add: of_num.simps of_num_add algebra_simps)
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   686
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   687
lemma of_num_pow: "of_num (pow x y) = of_num x ^ of_num y"
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   688
by (induct y)
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   689
   (simp_all add: of_num.simps of_num_square of_num_mult power_add)
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   690
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   691
lemma power_of_num [numeral]: "of_num x ^ of_num y = of_num (pow x y)"
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   692
  unfolding of_num_pow ..
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   693
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   694
lemma power_zero_of_num [numeral]:
31029
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   695
  "0 ^ of_num n = (0::'a::semiring_1)"
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   696
  using of_num_pos [where n=n and ?'a=nat]
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   697
  by (simp add: power_0_left)
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   698
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   699
lemma power_minus_Dig0 [numeral]:
31029
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   700
  fixes x :: "'a::ring_1"
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   701
  shows "(- x) ^ of_num (Dig0 n) = x ^ of_num (Dig0 n)"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   702
  by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc)
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   703
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   704
lemma power_minus_Dig1 [numeral]:
31029
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   705
  fixes x :: "'a::ring_1"
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   706
  shows "(- x) ^ of_num (Dig1 n) = - (x ^ of_num (Dig1 n))"
31028
9c5b6a92da39 clean up unsigned numeral proofs
huffman
parents: 31027
diff changeset
   707
  by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc)
29954
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   708
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   709
declare power_one [numeral]
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   710
8a95050c7044 add lemmas for exponentiation
huffman
parents: 29947
diff changeset
   711
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   712
subsubsection {* Greetings to @{typ nat}. *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   713
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   714
instance nat :: semiring_1_minus proof
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   715
qed simp_all
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   716
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
   717
lemma Suc_of_num [numeral]: "Suc (of_num n) = of_num (n + One)"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   718
  unfolding of_num_plus_one [symmetric] by simp
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   719
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   720
lemma nat_number:
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   721
  "1 = Suc 0"
29942
31069b8d39df replace 1::num with One; remove monoid_mult instance
huffman
parents: 29941
diff changeset
   722
  "of_num One = Suc 0"
29941
b951d80774d5 replace dec with double-and-decrement function
huffman
parents: 29667
diff changeset
   723
  "of_num (Dig0 n) = Suc (of_num (DigM n))"
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   724
  "of_num (Dig1 n) = Suc (of_num (Dig0 n))"
29941
b951d80774d5 replace dec with double-and-decrement function
huffman
parents: 29667
diff changeset
   725
  by (simp_all add: of_num.simps DigM_plus_one Suc_of_num)
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   726
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   727
declare diff_0_eq_0 [numeral]
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   728
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   729
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   730
subsection {* Proof tools setup *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   731
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   732
subsubsection {* Numeral equations as default simplification rules *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   733
31029
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   734
declare (in semiring_numeral) of_num_One [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   735
declare (in semiring_numeral) of_num_plus_one [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   736
declare (in semiring_numeral) of_num_one_plus [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   737
declare (in semiring_numeral) of_num_plus [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   738
declare (in semiring_numeral) of_num_times [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   739
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   740
declare (in semiring_1) of_nat_of_num [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   741
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   742
declare (in semiring_char_0) of_num_eq_iff [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   743
declare (in semiring_char_0) of_num_eq_one_iff [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   744
declare (in semiring_char_0) one_eq_of_num_iff [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   745
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   746
declare (in linordered_semidom) of_num_pos [simp]
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   747
declare (in linordered_semidom) of_num_not_zero [simp]
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   748
declare (in linordered_semidom) of_num_less_eq_iff [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   749
declare (in linordered_semidom) of_num_less_eq_one_iff [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   750
declare (in linordered_semidom) one_less_eq_of_num_iff [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   751
declare (in linordered_semidom) of_num_less_iff [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   752
declare (in linordered_semidom) of_num_less_one_iff [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   753
declare (in linordered_semidom) one_less_of_num_iff [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   754
declare (in linordered_semidom) of_num_nonneg [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   755
declare (in linordered_semidom) of_num_less_zero_iff [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   756
declare (in linordered_semidom) of_num_le_zero_iff [simp]
31029
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   757
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   758
declare (in linordered_idom) le_signed_numeral_special [simp]
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33523
diff changeset
   759
declare (in linordered_idom) less_signed_numeral_special [simp]
31029
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   760
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   761
declare (in semiring_1_minus) Dig_of_num_minus_one [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   762
declare (in semiring_1_minus) Dig_one_minus_of_num [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   763
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   764
declare (in ring_1) Dig_plus_uminus [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   765
declare (in ring_1) of_int_of_num [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   766
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   767
declare power_of_num [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   768
declare power_zero_of_num [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   769
declare power_minus_Dig0 [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   770
declare power_minus_Dig1 [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   771
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   772
declare Suc_of_num [simp]
e564767f8f78 used named theorems for declaring numeral simps
huffman
parents: 31028
diff changeset
   773
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
   774
31026
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   775
subsubsection {* Reorientation of equalities *}
31025
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   776
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   777
setup {*
33523
96730ad673be modernized structure Reorient_Proc;
wenzelm
parents: 33346
diff changeset
   778
  Reorient_Proc.add
31025
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   779
    (fn Const(@{const_name of_num}, _) $ _ => true
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   780
      | Const(@{const_name uminus}, _) $
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   781
          (Const(@{const_name of_num}, _) $ _) => true
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   782
      | _ => false)
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   783
*}
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   784
33523
96730ad673be modernized structure Reorient_Proc;
wenzelm
parents: 33346
diff changeset
   785
simproc_setup reorient_num ("of_num n = x" | "- of_num m = y") = Reorient_Proc.proc
96730ad673be modernized structure Reorient_Proc;
wenzelm
parents: 33346
diff changeset
   786
31025
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   787
31026
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   788
subsubsection {* Constant folding for multiplication in semirings *}
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   789
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   790
context semiring_numeral
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   791
begin
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   792
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   793
lemma mult_of_num_commute: "x * of_num n = of_num n * x"
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   794
by (induct n)
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   795
  (simp_all only: of_num.simps left_distrib right_distrib mult_1_left mult_1_right)
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   796
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   797
definition
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   798
  "commutes_with a b \<longleftrightarrow> a * b = b * a"
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   799
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   800
lemma commutes_with_commute: "commutes_with a b \<Longrightarrow> a * b = b * a"
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   801
unfolding commutes_with_def .
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   802
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   803
lemma commutes_with_left_commute: "commutes_with a b \<Longrightarrow> a * (b * c) = b * (a * c)"
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   804
unfolding commutes_with_def by (simp only: mult_assoc [symmetric])
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   805
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   806
lemma commutes_with_numeral: "commutes_with x (of_num n)" "commutes_with (of_num n) x"
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   807
unfolding commutes_with_def by (simp_all add: mult_of_num_commute)
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   808
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   809
lemmas mult_ac_numeral =
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   810
  mult_assoc
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   811
  commutes_with_commute
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   812
  commutes_with_left_commute
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   813
  commutes_with_numeral
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   814
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   815
end
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   816
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   817
ML {*
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   818
structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA =
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   819
struct
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   820
  val assoc_ss = HOL_ss addsimps @{thms mult_ac_numeral}
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   821
  val eq_reflection = eq_reflection
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   822
  fun is_numeral (Const(@{const_name of_num}, _) $ _) = true
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   823
    | is_numeral _ = false;
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   824
end;
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   825
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   826
structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   827
*}
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   828
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   829
simproc_setup semiring_assoc_fold' ("(a::'a::semiring_numeral) * b") =
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   830
  {* fn phi => fn ss => fn ct =>
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   831
    Semiring_Times_Assoc.proc ss (Thm.term_of ct) *}
020efcbfe2d8 add semiring_assoc_fold simproc for unsigned numerals
huffman
parents: 31025
diff changeset
   832
31025
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
   833
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   834
subsection {* Code generator setup for @{typ int} *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   835
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   836
text {* Reversing standard setup *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   837
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   838
lemma [code_unfold del]: "(0::int) \<equiv> Numeral0" by simp
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   839
lemma [code_unfold del]: "(1::int) \<equiv> Numeral1" by simp
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   840
declare zero_is_num_zero [code_unfold del]
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   841
declare one_is_num_one [code_unfold del]
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   842
  
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   843
lemma [code, code del]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   844
  "(1 :: int) = 1"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   845
  "(op + :: int \<Rightarrow> int \<Rightarrow> int) = op +"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   846
  "(uminus :: int \<Rightarrow> int) = uminus"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   847
  "(op - :: int \<Rightarrow> int \<Rightarrow> int) = op -"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   848
  "(op * :: int \<Rightarrow> int \<Rightarrow> int) = op *"
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   849
  "(HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool) = HOL.equal"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   850
  "(op \<le> :: int \<Rightarrow> int \<Rightarrow> bool) = op \<le>"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   851
  "(op < :: int \<Rightarrow> int \<Rightarrow> bool) = op <"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   852
  by rule+
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   853
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   854
text {* Constructors *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   855
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   856
definition Pls :: "num \<Rightarrow> int" where
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   857
  [simp, code_post]: "Pls n = of_num n"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   858
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   859
definition Mns :: "num \<Rightarrow> int" where
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   860
  [simp, code_post]: "Mns n = - of_num n"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   861
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   862
code_datatype "0::int" Pls Mns
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   863
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   864
lemmas [code_unfold] = Pls_def [symmetric] Mns_def [symmetric]
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   865
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   866
text {* Auxiliary operations *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   867
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   868
definition dup :: "int \<Rightarrow> int" where
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   869
  [simp]: "dup k = k + k"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   870
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   871
lemma Dig_dup [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   872
  "dup 0 = 0"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   873
  "dup (Pls n) = Pls (Dig0 n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   874
  "dup (Mns n) = Mns (Dig0 n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   875
  by (simp_all add: of_num.simps)
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   876
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   877
definition sub :: "num \<Rightarrow> num \<Rightarrow> int" where
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   878
  [simp]: "sub m n = (of_num m - of_num n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   879
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   880
lemma Dig_sub [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   881
  "sub One One = 0"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   882
  "sub (Dig0 m) One = of_num (DigM m)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   883
  "sub (Dig1 m) One = of_num (Dig0 m)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   884
  "sub One (Dig0 n) = - of_num (DigM n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   885
  "sub One (Dig1 n) = - of_num (Dig0 n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   886
  "sub (Dig0 m) (Dig0 n) = dup (sub m n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   887
  "sub (Dig1 m) (Dig1 n) = dup (sub m n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   888
  "sub (Dig1 m) (Dig0 n) = dup (sub m n) + 1"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   889
  "sub (Dig0 m) (Dig1 n) = dup (sub m n) - 1"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   890
  by (simp_all add: algebra_simps num_eq_iff nat_of_num_add)
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   891
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   892
text {* Implementations *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   893
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   894
lemma one_int_code [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   895
  "1 = Pls One"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   896
  by (simp add: of_num_One)
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   897
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   898
lemma plus_int_code [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   899
  "k + 0 = (k::int)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   900
  "0 + l = (l::int)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   901
  "Pls m + Pls n = Pls (m + n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   902
  "Pls m + Mns n = sub m n"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   903
  "Mns m + Pls n = sub n m"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   904
  "Mns m + Mns n = Mns (m + n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   905
  by simp_all
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   906
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   907
lemma uminus_int_code [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   908
  "uminus 0 = (0::int)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   909
  "uminus (Pls m) = Mns m"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   910
  "uminus (Mns m) = Pls m"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   911
  by simp_all
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   912
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   913
lemma minus_int_code [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   914
  "k - 0 = (k::int)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   915
  "0 - l = uminus (l::int)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   916
  "Pls m - Pls n = sub m n"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   917
  "Pls m - Mns n = Pls (m + n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   918
  "Mns m - Pls n = Mns (m + n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   919
  "Mns m - Mns n = sub n m"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   920
  by simp_all
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   921
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   922
lemma times_int_code [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   923
  "k * 0 = (0::int)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   924
  "0 * l = (0::int)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   925
  "Pls m * Pls n = Pls (m * n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   926
  "Pls m * Mns n = Mns (m * n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   927
  "Mns m * Pls n = Mns (m * n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   928
  "Mns m * Mns n = Pls (m * n)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   929
  by simp_all
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   930
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   931
lemma eq_int_code [code]:
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   932
  "HOL.equal 0 (0::int) \<longleftrightarrow> True"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   933
  "HOL.equal 0 (Pls l) \<longleftrightarrow> False"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   934
  "HOL.equal 0 (Mns l) \<longleftrightarrow> False"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   935
  "HOL.equal (Pls k) 0 \<longleftrightarrow> False"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   936
  "HOL.equal (Pls k) (Pls l) \<longleftrightarrow> HOL.equal k l"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   937
  "HOL.equal (Pls k) (Mns l) \<longleftrightarrow> False"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   938
  "HOL.equal (Mns k) 0 \<longleftrightarrow> False"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   939
  "HOL.equal (Mns k) (Pls l) \<longleftrightarrow> False"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   940
  "HOL.equal (Mns k) (Mns l) \<longleftrightarrow> HOL.equal k l"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   941
  by (auto simp add: equal dest: sym)
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   942
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   943
lemma [code nbe]:
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   944
  "HOL.equal (k::int) k \<longleftrightarrow> True"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
   945
  by (fact equal_refl)
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   946
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   947
lemma less_eq_int_code [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   948
  "0 \<le> (0::int) \<longleftrightarrow> True"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   949
  "0 \<le> Pls l \<longleftrightarrow> True"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   950
  "0 \<le> Mns l \<longleftrightarrow> False"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   951
  "Pls k \<le> 0 \<longleftrightarrow> False"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   952
  "Pls k \<le> Pls l \<longleftrightarrow> k \<le> l"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   953
  "Pls k \<le> Mns l \<longleftrightarrow> False"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   954
  "Mns k \<le> 0 \<longleftrightarrow> True"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   955
  "Mns k \<le> Pls l \<longleftrightarrow> True"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   956
  "Mns k \<le> Mns l \<longleftrightarrow> l \<le> k"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   957
  by simp_all
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   958
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   959
lemma less_int_code [code]:
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   960
  "0 < (0::int) \<longleftrightarrow> False"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   961
  "0 < Pls l \<longleftrightarrow> True"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   962
  "0 < Mns l \<longleftrightarrow> False"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   963
  "Pls k < 0 \<longleftrightarrow> False"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   964
  "Pls k < Pls l \<longleftrightarrow> k < l"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   965
  "Pls k < Mns l \<longleftrightarrow> False"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   966
  "Mns k < 0 \<longleftrightarrow> True"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   967
  "Mns k < Pls l \<longleftrightarrow> True"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   968
  "Mns k < Mns l \<longleftrightarrow> l < k"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   969
  by simp_all
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   970
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   971
hide_const (open) sub dup
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   972
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   973
text {* Pretty literals *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   974
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   975
ML {*
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   976
local open Code_Thingol in
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   977
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   978
fun add_code print target =
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   979
  let
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   980
    fun dest_num one' dig0' dig1' thm =
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   981
      let
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   982
        fun dest_dig (IConst (c, _)) = if c = dig0' then 0
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   983
              else if c = dig1' then 1
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   984
              else Code_Printer.eqn_error thm "Illegal numeral expression: illegal dig"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   985
          | dest_dig _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal digit";
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   986
        fun dest_num (IConst (c, _)) = if c = one' then 1
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   987
              else Code_Printer.eqn_error thm "Illegal numeral expression: illegal leading digit"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   988
          | dest_num (t1 `$ t2) = 2 * dest_num t2 + dest_dig t1
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   989
          | dest_num _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal term";
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   990
      in dest_num end;
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   991
    fun pretty sgn literals [one', dig0', dig1'] _ thm _ _ [(t, _)] =
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   992
      (Code_Printer.str o print literals o sgn o dest_num one' dig0' dig1' thm) t
38923
79d7f2b4cf71 more coherent naming of syntax data structures
haftmann
parents: 38857
diff changeset
   993
    fun add_syntax (c, sgn) = Code_Target.add_const_syntax target c
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   994
      (SOME (Code_Printer.complex_const_syntax
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   995
        (1, ([@{const_name One}, @{const_name Dig0}, @{const_name Dig1}],
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   996
          pretty sgn))));
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   997
  in
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   998
    add_syntax (@{const_name Pls}, I)
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
   999
    #> add_syntax (@{const_name Mns}, (fn k => ~ k))
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1000
  end;
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1001
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1002
end
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1003
*}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1004
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1005
hide_const (open) One Dig0 Dig1
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1006
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1007
31025
6d2188134536 reorient simproc for unsigned numerals
huffman
parents: 31021
diff changeset
  1008
subsection {* Toy examples *}
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
  1009
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1010
definition "foo \<longleftrightarrow> #4 * #2 + #7 = (#8 :: nat)"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1011
definition "bar \<longleftrightarrow> #4 * #2 + #7 \<ge> (#8 :: int) - #3"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1012
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1013
code_thms foo bar
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1014
export_code foo bar checking SML OCaml? Haskell? Scala?
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1015
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1016
text {* This is an ad-hoc @{text Code_Integer} setup. *}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1017
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1018
setup {*
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1019
  fold (add_code Code_Printer.literal_numeral)
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1020
    [Code_ML.target_SML, Code_ML.target_OCaml, Code_Haskell.target, Code_Scala.target]
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1021
*}
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1022
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1023
code_type int
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1024
  (SML "IntInf.int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1025
  (OCaml "Big'_int.big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1026
  (Haskell "Integer")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1027
  (Scala "BigInt")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1028
  (Eval "int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1029
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1030
code_const "0::int"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1031
  (SML "0/ :/ IntInf.int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1032
  (OCaml "Big'_int.zero")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1033
  (Haskell "0")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1034
  (Scala "BigInt(0)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1035
  (Eval "0/ :/ int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1036
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1037
code_const Int.pred
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1038
  (SML "IntInf.- ((_), 1)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1039
  (OCaml "Big'_int.pred'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1040
  (Haskell "!(_/ -/ 1)")
38773
f9837065b5e8 prevent line breaks after Scala symbolic operators
haftmann
parents: 38054
diff changeset
  1041
  (Scala "!(_ -/ 1)")
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1042
  (Eval "!(_/ -/ 1)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1043
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1044
code_const Int.succ
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1045
  (SML "IntInf.+ ((_), 1)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1046
  (OCaml "Big'_int.succ'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1047
  (Haskell "!(_/ +/ 1)")
38773
f9837065b5e8 prevent line breaks after Scala symbolic operators
haftmann
parents: 38054
diff changeset
  1048
  (Scala "!(_ +/ 1)")
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1049
  (Eval "!(_/ +/ 1)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1050
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1051
code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1052
  (SML "IntInf.+ ((_), (_))")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1053
  (OCaml "Big'_int.add'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1054
  (Haskell infixl 6 "+")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1055
  (Scala infixl 7 "+")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1056
  (Eval infixl 8 "+")
37826
4c0a5e35931a avoid export_code ... file -
haftmann
parents: 37765
diff changeset
  1057
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1058
code_const "uminus \<Colon> int \<Rightarrow> int"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1059
  (SML "IntInf.~")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1060
  (OCaml "Big'_int.minus'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1061
  (Haskell "negate")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1062
  (Scala "!(- _)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1063
  (Eval "~/ _")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1064
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1065
code_const "op - \<Colon> int \<Rightarrow> int \<Rightarrow> int"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1066
  (SML "IntInf.- ((_), (_))")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1067
  (OCaml "Big'_int.sub'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1068
  (Haskell infixl 6 "-")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1069
  (Scala infixl 7 "-")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1070
  (Eval infixl 8 "-")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1071
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1072
code_const "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1073
  (SML "IntInf.* ((_), (_))")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1074
  (OCaml "Big'_int.mult'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1075
  (Haskell infixl 7 "*")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1076
  (Scala infixl 8 "*")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1077
  (Eval infixl 9 "*")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1078
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1079
code_const pdivmod
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1080
  (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1081
  (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1082
  (Haskell "divMod/ (abs _)/ (abs _)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1083
  (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1084
  (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
37826
4c0a5e35931a avoid export_code ... file -
haftmann
parents: 37765
diff changeset
  1085
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38797
diff changeset
  1086
code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1087
  (SML "!((_ : IntInf.int) = _)")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1088
  (OCaml "Big'_int.eq'_big'_int")
39272
0b61951d2682 Haskell == is infix, not infixl
haftmann
parents: 39134
diff changeset
  1089
  (Haskell infix 4 "==")
38054
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1090
  (Scala infixl 5 "==")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1091
  (Eval infixl 6 "=")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1092
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1093
code_const "op \<le> \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1094
  (SML "IntInf.<= ((_), (_))")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1095
  (OCaml "Big'_int.le'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1096
  (Haskell infix 4 "<=")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1097
  (Scala infixl 4 "<=")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1098
  (Eval infixl 6 "<=")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1099
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1100
code_const "op < \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1101
  (SML "IntInf.< ((_), (_))")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1102
  (OCaml "Big'_int.lt'_big'_int")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1103
  (Haskell infix 4 "<")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1104
  (Scala infixl 4 "<")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1105
  (Eval infixl 6 "<")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1106
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1107
code_const Code_Numeral.int_of
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1108
  (SML "IntInf.fromInt")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1109
  (OCaml "_")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1110
  (Haskell "toInteger")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1111
  (Scala "!_.as'_BigInt")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1112
  (Eval "_")
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1113
acd48ef85bfc tuned; added pretty numerals for code generation
haftmann
parents: 37826
diff changeset
  1114
export_code foo bar checking SML OCaml? Haskell? Scala?
28021
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
  1115
32acf3c6cd12 added HOL/ex/Numeral.thy
haftmann
parents:
diff changeset
  1116
end