Attic/Quot/Examples/IntEx2.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 27 Aug 2010 02:25:40 +0000
changeset 2441 fc3e8f79e698
parent 1974 13298f4b212b
permissions -rw-r--r--
Ball Bex can be lifted after unfolding.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
568
0384e039b7f2 added new example for Ints; regularise does not work in all instances
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory IntEx2
1974
13298f4b212b Cleaning of Int and FSet Examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1939
diff changeset
     2
imports "Quotient_Int"
568
0384e039b7f2 added new example for Ints; regularise does not work in all instances
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
0384e039b7f2 added new example for Ints; regularise does not work in all instances
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
570
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
     5
subsection {* Embedding of the Integers into any @{text ring_1}: @{text of_int}*}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
     6
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
     7
(*
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
     8
context ring_1
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
     9
begin
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    10
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    11
 
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    12
definition 
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    13
  of_int :: "int \<Rightarrow> 'a" 
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    14
where
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    15
  "of_int 
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    16
*)
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    17
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    18
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    19
subsection {* Binary representation *}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    20
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    21
text {*
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    22
  This formalization defines binary arithmetic in terms of the integers
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    23
  rather than using a datatype. This avoids multiple representations (leading
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    24
  zeroes, etc.)  See @{text "ZF/Tools/twos-compl.ML"}, function @{text
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    25
  int_of_binary}, for the numerical interpretation.
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    26
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    27
  The representation expects that @{text "(m mod 2)"} is 0 or 1,
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    28
  even if m is negative;
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    29
  For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    30
  @{text "-5 = (-3)*2 + 1"}.
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    31
  
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    32
  This two's complement binary representation derives from the paper 
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    33
  "An Efficient Representation of Arithmetic for Term Rewriting" by
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    34
  Dave Cohen and Phil Watson, Rewriting Techniques and Applications,
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    35
  Springer LNCS 488 (240-251), 1991.
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    36
*}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    37
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    38
subsubsection {* The constructors @{term Bit0}, @{term Bit1}, @{term Pls} and @{term Min} *}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    39
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    40
definition
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    41
  Pls :: int where
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    42
  [code del]: "Pls = 0"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    43
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    44
definition
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    45
  Min :: int where
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    46
  [code del]: "Min = - 1"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    47
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    48
definition
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    49
  Bit0 :: "int \<Rightarrow> int" where
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    50
  [code del]: "Bit0 k = k + k"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    51
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    52
definition
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    53
  Bit1 :: "int \<Rightarrow> int" where
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    54
  [code del]: "Bit1 k = 1 + k + k"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    55
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    56
class number = -- {* for numeric types: nat, int, real, \dots *}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    57
  fixes number_of :: "int \<Rightarrow> 'a"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    58
600
5d932e7a856c List moved after QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 597
diff changeset
    59
(*use "~~/src/HOL/Tools/numeral.ML"
570
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    60
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    61
syntax
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    62
  "_Numeral" :: "num_const \<Rightarrow> 'a"    ("_")
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    63
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    64
use "~~/src/HOL/Tools/numeral_syntax.ML"
600
5d932e7a856c List moved after QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 597
diff changeset
    65
570
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    66
setup NumeralSyntax.setup
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    67
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    68
abbreviation
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    69
  "Numeral0 \<equiv> number_of Pls"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    70
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    71
abbreviation
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    72
  "Numeral1 \<equiv> number_of (Bit1 Pls)"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    73
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    74
lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    75
  -- {* Unfold all @{text let}s involving constants *}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    76
  unfolding Let_def ..
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    77
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    78
definition
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    79
  succ :: "int \<Rightarrow> int" where
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    80
  [code del]: "succ k = k + 1"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    81
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    82
definition
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    83
  pred :: "int \<Rightarrow> int" where
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    84
  [code del]: "pred k = k - 1"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    85
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    86
lemmas
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    87
  max_number_of [simp] = max_def
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    88
    [of "number_of u" "number_of v", standard, simp]
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    89
and
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    90
  min_number_of [simp] = min_def 
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    91
    [of "number_of u" "number_of v", standard, simp]
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    92
  -- {* unfolding @{text minx} and @{text max} on numerals *}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    93
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    94
lemmas numeral_simps = 
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    95
  succ_def pred_def Pls_def Min_def Bit0_def Bit1_def
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    96
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    97
text {* Removal of leading zeroes *}
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    98
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
    99
lemma Bit0_Pls [simp, code_post]:
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   100
  "Bit0 Pls = Pls"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   101
  unfolding numeral_simps by simp
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   102
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   103
lemma Bit1_Min [simp, code_post]:
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   104
  "Bit1 Min = Min"
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   105
  unfolding numeral_simps by simp
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   106
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   107
lemmas normalize_bin_simps =
6a031829319a added more to IntEx2
Christian Urban <urbanc@in.tum.de>
parents: 568
diff changeset
   108
  Bit0_Pls Bit1_Min
600
5d932e7a856c List moved after QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 597
diff changeset
   109
*)
601
81f40b8bde7b added "end" to each example theory
Christian Urban <urbanc@in.tum.de>
parents: 600
diff changeset
   110
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 654
diff changeset
   111
end