src/HOL/Library/ExecutableRat.thy
author haftmann
Wed, 27 Dec 2006 19:10:00 +0100
changeset 21911 e29bcab0c81c
parent 21545 54cc492d80a9
child 22067 39d5d42116c4
permissions -rw-r--r--
added OCaml code generation (without dictionaries)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/ExecutableRat.thy
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     2
    ID:         $Id$
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     4
*)
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     5
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     6
header {* Executable implementation of rational numbers in HOL *}
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     7
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     8
theory ExecutableRat
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
     9
imports "~~/src/HOL/Real/Rational" "~~/src/HOL/NumberTheory/IntPrimes"
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    10
begin
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    11
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    12
text {*
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    13
  Actually nothing is proved about the implementation.
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    14
*}
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    15
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    16
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    17
section {* HOL definitions *}
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    18
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    19
datatype erat = Rat bool int int
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    20
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19137
diff changeset
    21
instance erat :: zero ..
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19137
diff changeset
    22
instance erat :: one ..
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19137
diff changeset
    23
instance erat :: plus ..
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19137
diff changeset
    24
instance erat :: minus ..
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19137
diff changeset
    25
instance erat :: times ..
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19137
diff changeset
    26
instance erat :: inverse ..
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19137
diff changeset
    27
instance erat :: ord ..
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    28
19609
a677ac8c9b10 fixed codegen bug, cleanup
haftmann
parents: 19233
diff changeset
    29
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    30
  norm :: "erat \<Rightarrow> erat" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    31
  "norm r = (case r of (Rat a p q) \<Rightarrow>
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    32
     if p = 0 then Rat True 0 1
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    33
     else
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    34
       let
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    35
         absp = abs p
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    36
       in let
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    37
         m = zgcd (absp, q)
19609
a677ac8c9b10 fixed codegen bug, cleanup
haftmann
parents: 19233
diff changeset
    38
       in Rat (a = (0 <= p)) (absp div m) (q div m))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    39
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    40
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    41
  common :: "(int * int) * (int * int) \<Rightarrow> (int * int) * int" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    42
  "common r = (case r of ((p1, q1), (p2, q2)) \<Rightarrow>
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    43
       let q' = q1 * q2 div int (gcd (nat q1, nat q2))
19609
a677ac8c9b10 fixed codegen bug, cleanup
haftmann
parents: 19233
diff changeset
    44
       in ((p1 * (q' div q1), p2 * (q' div q2)), q'))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    45
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    46
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    47
  of_quotient :: "int \<Rightarrow> int \<Rightarrow> erat" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    48
  "of_quotient a b = norm (Rat True a b)"
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    49
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    50
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    51
  of_rat :: "rat \<Rightarrow> erat" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    52
  "of_rat r = split of_quotient (SOME s. s : Rep_Rat r)"
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    53
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    54
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    55
  to_rat :: "erat \<Rightarrow> rat" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    56
  "to_rat r = (case r of (Rat a p q) \<Rightarrow>
19609
a677ac8c9b10 fixed codegen bug, cleanup
haftmann
parents: 19233
diff changeset
    57
       if a then Fract p q else Fract (uminus p) q)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    58
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    59
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21191
diff changeset
    60
  eq_erat :: "erat \<Rightarrow> erat \<Rightarrow> bool" where
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20453
diff changeset
    61
  "eq_erat r s = (norm r = norm s)"
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    62
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    63
axiomatization
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    64
  div_zero :: erat
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    65
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    66
defs (overloaded)
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    67
  zero_rat_def: "0 == Rat True 0 1"
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    68
  one_rat_def: "1 == Rat True 1 1"
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    69
  add_rat_def: "r + s == case r of Rat a1 p1 q1 \<Rightarrow> case s of Rat a2 p2 q2 \<Rightarrow>
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    70
        let
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    71
          ((r1, r2), den) = common ((p1, q1), (p2, q2))
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    72
        in let
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    73
          num = (if a1 then r1 else -r1) + (if a2 then r2 else -r2)
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    74
        in norm (Rat True num den)"
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    75
  uminus_rat_def: "- r == case r of Rat a p q \<Rightarrow>
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    76
        if p = 0 then Rat a p q
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    77
        else Rat (\<not> a) p q"
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    78
  times_rat_def: "r * s == case r of Rat a1 p1 q1 \<Rightarrow> case s of Rat a2 p2 q2 \<Rightarrow>
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    79
        norm (Rat (a1 = a2) (p1 * p2) (q1 * q2))"
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    80
  inverse_rat_def: "inverse r == case r of Rat a p q \<Rightarrow>
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    81
        if p = 0 then div_zero
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    82
        else Rat a q p"
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    83
  le_rat_def: "r <= s == case r of Rat a1 p1 q1 \<Rightarrow> case s of Rat a2 p2 q2 \<Rightarrow>
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    84
        (\<not> a1 \<and> a2) \<or>
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    85
        (\<not> (a1 \<and> \<not> a2) \<and>
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    86
          (let
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    87
            ((r1, r2), dummy) = common ((p1, q1), (p2, q2))
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    88
          in if a1 then r1 <= r2 else r2 <= r1))"
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
    89
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
    90
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    91
section {* code lemmas *}
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    92
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
    93
lemma number_of_rat [code unfold]:
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
    94
  "(number_of k \<Colon> rat) \<equiv> Fract k 1"
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    95
  unfolding Fract_of_int_eq rat_number_of_def by simp
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    96
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
    97
declare divide_inverse [where ?'a = rat, code func]
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
    98
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
    99
instance rat :: eq ..
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   100
instance erat :: eq ..
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
   101
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
   102
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
   103
section {* code names *}
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
   104
21191
c00161fbf990 code generator module naming improved
haftmann
parents: 21125
diff changeset
   105
code_modulename SML
c00161fbf990 code generator module naming improved
haftmann
parents: 21125
diff changeset
   106
  ExecutableRat Rational
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
   107
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21545
diff changeset
   108
code_modulename OCaml
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21545
diff changeset
   109
  ExecutableRat Rational
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21545
diff changeset
   110
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
   111
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   112
section {* rat as abstype *}
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   113
21191
c00161fbf990 code generator module naming improved
haftmann
parents: 21125
diff changeset
   114
lemma [code func]: -- {* prevents eq constraint *}
c00161fbf990 code generator module naming improved
haftmann
parents: 21125
diff changeset
   115
  shows "All = All"
c00161fbf990 code generator module naming improved
haftmann
parents: 21125
diff changeset
   116
    and "contents = contents" by rule+
c00161fbf990 code generator module naming improved
haftmann
parents: 21125
diff changeset
   117
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   118
code_abstype rat erat where
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   119
  Fract \<equiv> of_quotient
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   120
  "0 \<Colon> rat" \<equiv> "0 \<Colon> erat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   121
  "1 \<Colon> rat" \<equiv> "1 \<Colon> erat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   122
  "op + \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat" \<equiv> "op + \<Colon> erat \<Rightarrow> erat \<Rightarrow> erat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   123
  "uminus \<Colon> rat \<Rightarrow> rat" \<equiv> "uminus \<Colon> erat \<Rightarrow> erat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   124
  "op * \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat" \<equiv> "op * \<Colon> erat \<Rightarrow> erat \<Rightarrow> erat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   125
  "inverse \<Colon> rat \<Rightarrow> rat" \<equiv> "inverse \<Colon> erat \<Rightarrow> erat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   126
  "op \<le> \<Colon> rat \<Rightarrow> rat \<Rightarrow> bool" \<equiv>  "op \<le> \<Colon> erat \<Rightarrow> erat \<Rightarrow> bool"
21454
a1937c51ed88 dropped eq const
haftmann
parents: 21404
diff changeset
   127
  "op = \<Colon> rat \<Rightarrow> rat \<Rightarrow> bool" \<equiv> eq_erat
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   128
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   129
code_const div_zero
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21545
diff changeset
   130
  (SML "raise/ Fail/ \"Division by zero\"")
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21545
diff changeset
   131
  (OCaml "failwith \"Division by zero\"")
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   132
  (Haskell "error/ \"Division by zero\"")
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   133
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   134
code_gen
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   135
  Fract
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   136
  "0 \<Colon> rat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   137
  "1 \<Colon> rat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   138
  "op + \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   139
  "uminus \<Colon> rat \<Rightarrow> rat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   140
  "op * \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   141
  "inverse \<Colon> rat \<Rightarrow> rat"
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   142
  "op \<le> \<Colon> rat \<Rightarrow> rat \<Rightarrow> bool"
21454
a1937c51ed88 dropped eq const
haftmann
parents: 21404
diff changeset
   143
  "op = \<Colon> rat \<Rightarrow> rat \<Rightarrow> bool"
21545
54cc492d80a9 adjusted syntax for internal code generation
haftmann
parents: 21454
diff changeset
   144
  (SML #)
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   145
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   146
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   147
section {* type serializations *}
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   148
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   149
types_code
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   150
  rat ("{*erat*}")
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   151
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   152
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   153
section {* const serializations *}
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   154
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   155
consts_code
20701
17a625996bb0 better handling for div by zero
haftmann
parents: 20597
diff changeset
   156
  div_zero ("raise/ (Fail/ \"non-defined rational number\")")
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   157
  Fract ("{*of_quotient*}")
20713
823967ef47f1 renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents: 20701
diff changeset
   158
  HOL.zero :: rat ("{*0::erat*}")
823967ef47f1 renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents: 20701
diff changeset
   159
  HOL.one :: rat ("{*1::erat*}")
19791
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   160
  HOL.plus :: "rat \<Rightarrow> rat \<Rightarrow> rat" ("{*op + :: erat \<Rightarrow> erat \<Rightarrow> erat*}")
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   161
  uminus :: "rat \<Rightarrow> rat" ("{*uminus :: erat \<Rightarrow> erat*}")
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   162
  HOL.times :: "rat \<Rightarrow> rat \<Rightarrow> rat" ("{*op * :: erat \<Rightarrow> erat \<Rightarrow> erat*}")
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   163
  inverse :: "rat \<Rightarrow> rat" ("{*inverse :: erat \<Rightarrow> erat*}")
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   164
  divide :: "rat \<Rightarrow> rat \<Rightarrow> rat" ("{*op * :: erat \<Rightarrow> erat \<Rightarrow> erat*}/ _/ ({*inverse :: erat \<Rightarrow> erat*}/ _)")
ab326de16ad5 refined code generation
haftmann
parents: 19609
diff changeset
   165
  Orderings.less_eq :: "rat \<Rightarrow> rat \<Rightarrow> bool" ("{*op <= :: erat \<Rightarrow> erat \<Rightarrow> bool*}")
20934
2b872c161747 added code_abstype
haftmann
parents: 20713
diff changeset
   166
  "op =" :: "rat \<Rightarrow> rat \<Rightarrow> bool" ("{*eq_erat*}")
19889
2202a5648897 slight adaptions for code generator
haftmann
parents: 19791
diff changeset
   167
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents:
diff changeset
   168
end