src/HOL/Library/Numeral_Type.thy
author huffman
Thu, 19 Feb 2009 23:18:28 -0800
changeset 30001 dd27e16677b2
parent 29999 da85a244e328
child 30032 c7f0c1b8001b
permissions -rw-r--r--
cleaned up
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29629
5111ce425e7a tuned header
haftmann
parents: 29025
diff changeset
     1
(*  Title:      HOL/Library/Numeral_Type.thy
5111ce425e7a tuned header
haftmann
parents: 29025
diff changeset
     2
    Author:     Brian Huffman
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
     3
*)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
     4
29629
5111ce425e7a tuned header
haftmann
parents: 29025
diff changeset
     5
header {* Numeral Syntax for Types *}
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
     6
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
     7
theory Numeral_Type
27487
c8a6ce181805 absolute imports of HOL/*.thy theories
haftmann
parents: 27368
diff changeset
     8
imports Plain "~~/src/HOL/Presburger"
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
     9
begin
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    10
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    11
subsection {* Preliminary lemmas *}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    12
(* These should be moved elsewhere *)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    13
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    14
lemma (in type_definition) univ:
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    15
  "UNIV = Abs ` A"
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    16
proof
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    17
  show "Abs ` A \<subseteq> UNIV" by (rule subset_UNIV)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    18
  show "UNIV \<subseteq> Abs ` A"
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    19
  proof
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    20
    fix x :: 'b
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    21
    have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric])
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    22
    moreover have "Rep x \<in> A" by (rule Rep)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    23
    ultimately show "x \<in> Abs ` A" by (rule image_eqI)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    24
  qed
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    25
qed
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    26
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    27
lemma (in type_definition) card: "card (UNIV :: 'b set) = card A"
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    28
  by (simp add: univ card_image inj_on_def Abs_inject)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    29
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    30
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    31
subsection {* Cardinalities of types *}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    32
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    33
syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))")
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    34
28920
4ed4b8b1988d fix typed print translation for card UNIV
huffman
parents: 27487
diff changeset
    35
translations "CARD(t)" => "CONST card (CONST UNIV \<Colon> t set)"
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    36
24407
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    37
typed_print_translation {*
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    38
let
28920
4ed4b8b1988d fix typed print translation for card UNIV
huffman
parents: 27487
diff changeset
    39
  fun card_univ_tr' show_sorts _ [Const (@{const_name UNIV}, Type(_,[T,_]))] =
24407
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    40
    Syntax.const "_type_card" $ Syntax.term_of_typ show_sorts T;
28920
4ed4b8b1988d fix typed print translation for card UNIV
huffman
parents: 27487
diff changeset
    41
in [(@{const_syntax card}, card_univ_tr')]
24407
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    42
end
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    43
*}
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    44
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    45
lemma card_unit [simp]: "CARD(unit) = 1"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    46
  unfolding UNIV_unit by simp
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    47
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    48
lemma card_bool [simp]: "CARD(bool) = 2"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    49
  unfolding UNIV_bool by simp
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    50
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    51
lemma card_prod [simp]: "CARD('a \<times> 'b) = CARD('a::finite) * CARD('b::finite)"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    52
  unfolding UNIV_Times_UNIV [symmetric] by (simp only: card_cartesian_product)
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    53
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    54
lemma card_sum [simp]: "CARD('a + 'b) = CARD('a::finite) + CARD('b::finite)"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    55
  unfolding UNIV_Plus_UNIV [symmetric] by (simp only: finite card_Plus)
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    56
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    57
lemma card_option [simp]: "CARD('a option) = Suc CARD('a::finite)"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    58
  unfolding insert_None_conv_UNIV [symmetric]
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    59
  apply (subgoal_tac "(None::'a option) \<notin> range Some")
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
    60
  apply (simp add: card_image)
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    61
  apply fast
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    62
  done
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    63
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    64
lemma card_set [simp]: "CARD('a set) = 2 ^ CARD('a::finite)"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    65
  unfolding Pow_UNIV [symmetric]
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    66
  by (simp only: card_Pow finite numeral_2_eq_2)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    67
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    68
lemma card_nat [simp]: "CARD(nat) = 0"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    69
  by (simp add: infinite_UNIV_nat card_eq_0_iff)
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    70
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    71
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    72
subsection {* Classes with at least 1 and 2  *}
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    73
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    74
text {* Class finite already captures "at least 1" *}
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    75
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    76
lemma zero_less_card_finite [simp]: "0 < CARD('a::finite)"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
    77
  unfolding neq0_conv [symmetric] by simp
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
    78
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    79
lemma one_le_card_finite [simp]: "Suc 0 \<le> CARD('a::finite)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    80
  by (simp add: less_Suc_eq_le [symmetric])
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    81
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    82
text {* Class for cardinality "at least 2" *}
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    83
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    84
class card2 = finite + 
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    85
  assumes two_le_card: "2 \<le> CARD('a)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    86
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    87
lemma one_less_card: "Suc 0 < CARD('a::card2)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    88
  using two_le_card [where 'a='a] by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    89
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    90
lemma one_less_int_card: "1 < int CARD('a::card2)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    91
  using one_less_card [where 'a='a] by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    92
25378
dca691610489 tuned document;
wenzelm
parents: 24630
diff changeset
    93
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    94
subsection {* Numeral Types *}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    95
24406
d96eb21fc1bc rename type pls to num0
huffman
parents: 24332
diff changeset
    96
typedef (open) num0 = "UNIV :: nat set" ..
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    97
typedef (open) num1 = "UNIV :: unit set" ..
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
    98
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
    99
typedef (open) 'a bit0 = "{0 ..< 2 * int CARD('a::finite)}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   100
proof
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   101
  show "0 \<in> {0 ..< 2 * int CARD('a)}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   102
    by simp
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   103
qed
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   104
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   105
typedef (open) 'a bit1 = "{0 ..< 1 + 2 * int CARD('a::finite)}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   106
proof
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   107
  show "0 \<in> {0 ..< 1 + 2 * int CARD('a)}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   108
    by simp
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   109
qed
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   110
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   111
lemma card_num0 [simp]: "CARD (num0) = 0"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   112
  unfolding type_definition.card [OF type_definition_num0]
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   113
  by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   114
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   115
lemma card_num1 [simp]: "CARD(num1) = 1"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   116
  unfolding type_definition.card [OF type_definition_num1]
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   117
  by (simp only: card_unit)
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   118
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   119
lemma card_bit0 [simp]: "CARD('a bit0) = 2 * CARD('a::finite)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   120
  unfolding type_definition.card [OF type_definition_bit0]
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   121
  by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   122
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   123
lemma card_bit1 [simp]: "CARD('a bit1) = Suc (2 * CARD('a::finite))"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   124
  unfolding type_definition.card [OF type_definition_bit1]
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   125
  by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   126
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   127
instance num1 :: finite
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   128
proof
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   129
  show "finite (UNIV::num1 set)"
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   130
    unfolding type_definition.univ [OF type_definition_num1]
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   131
    using finite by (rule finite_imageI)
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   132
qed
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   133
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   134
instance bit0 :: (finite) card2
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   135
proof
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   136
  show "finite (UNIV::'a bit0 set)"
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   137
    unfolding type_definition.univ [OF type_definition_bit0]
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   138
    by simp
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   139
  show "2 \<le> CARD('a bit0)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   140
    by simp
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   141
qed
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   142
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   143
instance bit1 :: (finite) card2
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   144
proof
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   145
  show "finite (UNIV::'a bit1 set)"
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   146
    unfolding type_definition.univ [OF type_definition_bit1]
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   147
    by simp
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   148
  show "2 \<le> CARD('a bit1)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   149
    by simp
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   150
qed
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   151
25378
dca691610489 tuned document;
wenzelm
parents: 24630
diff changeset
   152
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   153
subsection {* Locale for modular arithmetic subtypes *}
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   154
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   155
locale mod_type =
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   156
  fixes n :: int
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   157
  and Rep :: "'a::{zero,one,plus,times,uminus,minus,power} \<Rightarrow> int"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   158
  and Abs :: "int \<Rightarrow> 'a::{zero,one,plus,times,uminus,minus,power}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   159
  assumes type: "type_definition Rep Abs {0..<n}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   160
  and size1: "1 < n"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   161
  and zero_def: "0 = Abs 0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   162
  and one_def:  "1 = Abs 1"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   163
  and add_def:  "x + y = Abs ((Rep x + Rep y) mod n)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   164
  and mult_def: "x * y = Abs ((Rep x * Rep y) mod n)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   165
  and diff_def: "x - y = Abs ((Rep x - Rep y) mod n)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   166
  and minus_def: "- x = Abs ((- Rep x) mod n)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   167
  and power_def: "x ^ k = Abs (Rep x ^ k mod n)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   168
begin
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   169
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   170
lemma size0: "0 < n"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   171
by (cut_tac size1, simp)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   172
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   173
lemmas definitions =
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   174
  zero_def one_def add_def mult_def minus_def diff_def power_def
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   175
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   176
lemma Rep_less_n: "Rep x < n"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   177
by (rule type_definition.Rep [OF type, simplified, THEN conjunct2])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   178
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   179
lemma Rep_le_n: "Rep x \<le> n"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   180
by (rule Rep_less_n [THEN order_less_imp_le])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   181
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   182
lemma Rep_inject_sym: "x = y \<longleftrightarrow> Rep x = Rep y"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   183
by (rule type_definition.Rep_inject [OF type, symmetric])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   184
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   185
lemma Rep_inverse: "Abs (Rep x) = x"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   186
by (rule type_definition.Rep_inverse [OF type])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   187
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   188
lemma Abs_inverse: "m \<in> {0..<n} \<Longrightarrow> Rep (Abs m) = m"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   189
by (rule type_definition.Abs_inverse [OF type])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   190
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   191
lemma Rep_Abs_mod: "Rep (Abs (m mod n)) = m mod n"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   192
by (simp add: Abs_inverse IntDiv.pos_mod_conj [OF size0])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   193
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   194
lemma Rep_Abs_0: "Rep (Abs 0) = 0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   195
by (simp add: Abs_inverse size0)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   196
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   197
lemma Rep_0: "Rep 0 = 0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   198
by (simp add: zero_def Rep_Abs_0)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   199
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   200
lemma Rep_Abs_1: "Rep (Abs 1) = 1"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   201
by (simp add: Abs_inverse size1)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   202
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   203
lemma Rep_1: "Rep 1 = 1"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   204
by (simp add: one_def Rep_Abs_1)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   205
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   206
lemma Rep_mod: "Rep x mod n = Rep x"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   207
apply (rule_tac x=x in type_definition.Abs_cases [OF type])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   208
apply (simp add: type_definition.Abs_inverse [OF type])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   209
apply (simp add: mod_pos_pos_trivial)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   210
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   211
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   212
lemmas Rep_simps =
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   213
  Rep_inject_sym Rep_inverse Rep_Abs_mod Rep_mod Rep_Abs_0 Rep_Abs_1
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   214
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   215
lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   216
apply (intro_classes, unfold definitions)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   217
apply (simp_all add: Rep_simps zmod_simps ring_simps)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   218
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   219
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   220
lemma recpower: "OFCLASS('a, recpower_class)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   221
apply (intro_classes, unfold definitions)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   222
apply (simp_all add: Rep_simps zmod_simps add_ac mult_assoc
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   223
                     mod_pos_pos_trivial size1)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   224
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   225
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   226
end
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   227
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   228
locale mod_ring = mod_type +
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   229
  constrains n :: int
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   230
  and Rep :: "'a::{number_ring,power} \<Rightarrow> int"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   231
  and Abs :: "int \<Rightarrow> 'a::{number_ring,power}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   232
begin
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   233
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   234
lemma of_nat_eq: "of_nat k = Abs (int k mod n)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   235
apply (induct k)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   236
apply (simp add: zero_def)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   237
apply (simp add: Rep_simps add_def one_def zmod_simps add_ac)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   238
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   239
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   240
lemma of_int_eq: "of_int z = Abs (z mod n)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   241
apply (cases z rule: int_diff_cases)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   242
apply (simp add: Rep_simps of_nat_eq diff_def zmod_simps)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   243
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   244
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   245
lemma Rep_number_of:
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   246
  "Rep (number_of w) = number_of w mod n"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   247
by (simp add: number_of_eq of_int_eq Rep_Abs_mod)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   248
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   249
lemma iszero_number_of:
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   250
  "iszero (number_of w::'a) \<longleftrightarrow> number_of w mod n = 0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   251
by (simp add: Rep_simps number_of_eq of_int_eq iszero_def zero_def)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   252
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   253
lemma cases:
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   254
  assumes 1: "\<And>z. \<lbrakk>(x::'a) = of_int z; 0 \<le> z; z < n\<rbrakk> \<Longrightarrow> P"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   255
  shows "P"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   256
apply (cases x rule: type_definition.Abs_cases [OF type])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   257
apply (rule_tac z="y" in 1)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   258
apply (simp_all add: of_int_eq mod_pos_pos_trivial)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   259
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   260
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   261
lemma induct:
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   262
  "(\<And>z. \<lbrakk>0 \<le> z; z < n\<rbrakk> \<Longrightarrow> P (of_int z)) \<Longrightarrow> P (x::'a)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   263
by (cases x rule: cases) simp
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   264
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   265
end
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   266
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   267
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   268
subsection {* Number ring instances *}
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   269
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   270
instantiation
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   271
  bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus,power}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   272
begin
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   273
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   274
definition Abs_bit0' :: "int \<Rightarrow> 'a bit0" where
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   275
  "Abs_bit0' x = Abs_bit0 (x mod int CARD('a bit0))"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   276
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   277
definition Abs_bit1' :: "int \<Rightarrow> 'a bit1" where
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   278
  "Abs_bit1' x = Abs_bit1 (x mod int CARD('a bit1))"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   279
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   280
definition "0 = Abs_bit0 0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   281
definition "1 = Abs_bit0 1"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   282
definition "x + y = Abs_bit0' (Rep_bit0 x + Rep_bit0 y)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   283
definition "x * y = Abs_bit0' (Rep_bit0 x * Rep_bit0 y)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   284
definition "x - y = Abs_bit0' (Rep_bit0 x - Rep_bit0 y)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   285
definition "- x = Abs_bit0' (- Rep_bit0 x)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   286
definition "x ^ k = Abs_bit0' (Rep_bit0 x ^ k)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   287
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   288
definition "0 = Abs_bit1 0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   289
definition "1 = Abs_bit1 1"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   290
definition "x + y = Abs_bit1' (Rep_bit1 x + Rep_bit1 y)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   291
definition "x * y = Abs_bit1' (Rep_bit1 x * Rep_bit1 y)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   292
definition "x - y = Abs_bit1' (Rep_bit1 x - Rep_bit1 y)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   293
definition "- x = Abs_bit1' (- Rep_bit1 x)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   294
definition "x ^ k = Abs_bit1' (Rep_bit1 x ^ k)"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   295
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   296
instance ..
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   297
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   298
end
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   299
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   300
interpretation bit0!:
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   301
  mod_type "int CARD('a::finite bit0)"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   302
           "Rep_bit0 :: 'a::finite bit0 \<Rightarrow> int"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   303
           "Abs_bit0 :: int \<Rightarrow> 'a::finite bit0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   304
apply (rule mod_type.intro)
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   305
apply (simp add: int_mult type_definition_bit0)
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   306
apply (rule one_less_int_card)
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   307
apply (rule zero_bit0_def)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   308
apply (rule one_bit0_def)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   309
apply (rule plus_bit0_def [unfolded Abs_bit0'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   310
apply (rule times_bit0_def [unfolded Abs_bit0'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   311
apply (rule minus_bit0_def [unfolded Abs_bit0'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   312
apply (rule uminus_bit0_def [unfolded Abs_bit0'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   313
apply (rule power_bit0_def [unfolded Abs_bit0'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   314
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   315
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   316
interpretation bit1!:
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   317
  mod_type "int CARD('a::finite bit1)"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   318
           "Rep_bit1 :: 'a::finite bit1 \<Rightarrow> int"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   319
           "Abs_bit1 :: int \<Rightarrow> 'a::finite bit1"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   320
apply (rule mod_type.intro)
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   321
apply (simp add: int_mult type_definition_bit1)
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   322
apply (rule one_less_int_card)
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   323
apply (rule zero_bit1_def)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   324
apply (rule one_bit1_def)
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   325
apply (rule plus_bit1_def [unfolded Abs_bit1'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   326
apply (rule times_bit1_def [unfolded Abs_bit1'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   327
apply (rule minus_bit1_def [unfolded Abs_bit1'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   328
apply (rule uminus_bit1_def [unfolded Abs_bit1'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   329
apply (rule power_bit1_def [unfolded Abs_bit1'_def])
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   330
done
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   331
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   332
instance bit0 :: (finite) "{comm_ring_1,recpower}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   333
  by (rule bit0.comm_ring_1 bit0.recpower)+
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   334
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   335
instance bit1 :: (finite) "{comm_ring_1,recpower}"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   336
  by (rule bit1.comm_ring_1 bit1.recpower)+
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   337
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   338
instantiation bit0 and bit1 :: (finite) number_ring
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   339
begin
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   340
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   341
definition "(number_of w :: _ bit0) = of_int w"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   342
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   343
definition "(number_of w :: _ bit1) = of_int w"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   344
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   345
instance proof
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   346
qed (rule number_of_bit0_def number_of_bit1_def)+
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   347
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   348
end
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   349
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   350
interpretation bit0!:
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   351
  mod_ring "int CARD('a::finite bit0)"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   352
           "Rep_bit0 :: 'a::finite bit0 \<Rightarrow> int"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   353
           "Abs_bit0 :: int \<Rightarrow> 'a::finite bit0"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   354
  ..
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   355
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   356
interpretation bit1!:
29998
19e1ef628b25 nicer induction/cases rules for numeral types
huffman
parents: 29997
diff changeset
   357
  mod_ring "int CARD('a::finite bit1)"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   358
           "Rep_bit1 :: 'a::finite bit1 \<Rightarrow> int"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   359
           "Abs_bit1 :: int \<Rightarrow> 'a::finite bit1"
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   360
  ..
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   361
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   362
text {* Set up cases, induction, and arithmetic *}
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   363
29999
da85a244e328 fix case_names
huffman
parents: 29998
diff changeset
   364
lemmas bit0_cases [case_names of_int, cases type: bit0] = bit0.cases
da85a244e328 fix case_names
huffman
parents: 29998
diff changeset
   365
lemmas bit1_cases [case_names of_int, cases type: bit1] = bit1.cases
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   366
29999
da85a244e328 fix case_names
huffman
parents: 29998
diff changeset
   367
lemmas bit0_induct [case_names of_int, induct type: bit0] = bit0.induct
da85a244e328 fix case_names
huffman
parents: 29998
diff changeset
   368
lemmas bit1_induct [case_names of_int, induct type: bit1] = bit1.induct
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   369
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   370
lemmas bit0_iszero_number_of [simp] = bit0.iszero_number_of
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   371
lemmas bit1_iszero_number_of [simp] = bit1.iszero_number_of
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   372
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   373
declare power_Suc [where ?'a="'a::finite bit0", standard, simp]
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   374
declare power_Suc [where ?'a="'a::finite bit1", standard, simp]
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   375
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   376
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   377
subsection {* Syntax *}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   378
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   379
syntax
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   380
  "_NumeralType" :: "num_const => type"  ("_")
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   381
  "_NumeralType0" :: type ("0")
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   382
  "_NumeralType1" :: type ("1")
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   383
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   384
translations
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   385
  "_NumeralType1" == (type) "num1"
24406
d96eb21fc1bc rename type pls to num0
huffman
parents: 24332
diff changeset
   386
  "_NumeralType0" == (type) "num0"
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   387
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   388
parse_translation {*
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   389
let
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   390
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   391
val num1_const = Syntax.const "Numeral_Type.num1";
24406
d96eb21fc1bc rename type pls to num0
huffman
parents: 24332
diff changeset
   392
val num0_const = Syntax.const "Numeral_Type.num0";
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   393
val B0_const = Syntax.const "Numeral_Type.bit0";
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   394
val B1_const = Syntax.const "Numeral_Type.bit1";
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   395
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   396
fun mk_bintype n =
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   397
  let
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   398
    fun mk_bit n = if n = 0 then B0_const else B1_const;
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   399
    fun bin_of n =
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   400
      if n = 1 then num1_const
24406
d96eb21fc1bc rename type pls to num0
huffman
parents: 24332
diff changeset
   401
      else if n = 0 then num0_const
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   402
      else if n = ~1 then raise TERM ("negative type numeral", [])
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   403
      else
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24407
diff changeset
   404
        let val (q, r) = Integer.div_mod n 2;
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   405
        in mk_bit r $ bin_of q end;
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   406
  in bin_of n end;
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   407
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   408
fun numeral_tr (*"_NumeralType"*) [Const (str, _)] =
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24407
diff changeset
   409
      mk_bintype (valOf (Int.fromString str))
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   410
  | numeral_tr (*"_NumeralType"*) ts = raise TERM ("numeral_tr", ts);
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   411
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   412
in [("_NumeralType", numeral_tr)] end;
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   413
*}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   414
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   415
print_translation {*
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   416
let
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   417
fun int_of [] = 0
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24407
diff changeset
   418
  | int_of (b :: bs) = b + 2 * int_of bs;
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   419
24406
d96eb21fc1bc rename type pls to num0
huffman
parents: 24332
diff changeset
   420
fun bin_of (Const ("num0", _)) = []
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   421
  | bin_of (Const ("num1", _)) = [1]
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   422
  | bin_of (Const ("bit0", _) $ bs) = 0 :: bin_of bs
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   423
  | bin_of (Const ("bit1", _) $ bs) = 1 :: bin_of bs
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   424
  | bin_of t = raise TERM("bin_of", [t]);
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   425
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   426
fun bit_tr' b [t] =
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   427
  let
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   428
    val rev_digs = b :: bin_of t handle TERM _ => raise Match
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   429
    val i = int_of rev_digs;
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24407
diff changeset
   430
    val num = string_of_int (abs i);
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   431
  in
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   432
    Syntax.const "_NumeralType" $ Syntax.free num
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   433
  end
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   434
  | bit_tr' b _ = raise Match;
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   435
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   436
in [("bit0", bit_tr' 0), ("bit1", bit_tr' 1)] end;
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   437
*}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   438
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   439
subsection {* Examples *}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   440
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   441
lemma "CARD(0) = 0" by simp
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   442
lemma "CARD(17) = 17" by simp
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   443
lemma "8 * 11 ^ 3 - 6 = (2::5)" by simp
28920
4ed4b8b1988d fix typed print translation for card UNIV
huffman
parents: 27487
diff changeset
   444
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
   445
end