src/HOL/Library/Commutative_Ring.thy
author wenzelm
Thu, 19 Jan 2006 21:22:08 +0100
changeset 18708 4b3dadb4fe33
parent 17516 45164074dad4
child 19736 d8d0f8f51d69
permissions -rw-r--r--
setup: theory -> theory;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     1
(*  ID:         $Id$
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     2
    Author:     Bernhard Haeupler
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     3
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     4
Proving equalities in commutative rings done "right" in Isabelle/HOL.
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     5
*)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     6
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     7
header {* Proving equalities in commutative rings *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     8
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     9
theory Commutative_Ring
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    10
imports Main
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    11
uses ("comm_ring.ML")
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    12
begin
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    13
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    14
text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    15
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    16
datatype 'a pol =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    17
    Pc 'a
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    18
  | Pinj nat "'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    19
  | PX "'a pol" nat "'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    20
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    21
datatype 'a polex =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    22
  Pol "'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    23
  | Add "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    24
  | Sub "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    25
  | Mul "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    26
  | Pow "'a polex" nat
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    27
  | Neg "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    28
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    29
text {* Interpretation functions for the shadow syntax. *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    30
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    31
consts
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    32
  Ipol :: "'a::{comm_ring,recpower} list \<Rightarrow> 'a pol \<Rightarrow> 'a"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    33
  Ipolex :: "'a::{comm_ring,recpower} list \<Rightarrow> 'a polex \<Rightarrow> 'a"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    34
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    35
primrec
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    36
  "Ipol l (Pc c) = c"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    37
  "Ipol l (Pinj i P) = Ipol (drop i l) P"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    38
  "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    39
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    40
primrec
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    41
  "Ipolex l (Pol P) = Ipol l P"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    42
  "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    43
  "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    44
  "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    45
  "Ipolex l (Pow p n) = Ipolex l p ^ n"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    46
  "Ipolex l (Neg P) = - Ipolex l P"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    47
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    48
text {* Create polynomial normalized polynomials given normalized inputs. *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    49
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    50
constdefs
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    51
  mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    52
  "mkPinj x P \<equiv> (case P of
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    53
    Pc c \<Rightarrow> Pc c |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    54
    Pinj y P \<Rightarrow> Pinj (x + y) P |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    55
    PX p1 y p2 \<Rightarrow> Pinj x P)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    56
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    57
constdefs
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    58
  mkPX :: "'a::{comm_ring,recpower} pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    59
  "mkPX P i Q == (case P of
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    60
    Pc c \<Rightarrow> (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    61
    Pinj j R \<Rightarrow> PX P i Q |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    62
    PX P2 i2 Q2 \<Rightarrow> (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    63
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    64
text {* Defining the basic ring operations on normalized polynomials *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    65
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    66
consts
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    67
  add :: "'a::{comm_ring,recpower} pol \<times> 'a pol \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    68
  mul :: "'a::{comm_ring,recpower} pol \<times> 'a pol \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    69
  neg :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    70
  sqr :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    71
  pow :: "'a::{comm_ring,recpower} pol \<times> nat \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    72
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    73
text {* Addition *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    74
recdef add "measure (\<lambda>(x, y). size x + size y)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    75
  "add (Pc a, Pc b) = Pc (a + b)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    76
  "add (Pc c, Pinj i P) = Pinj i (add (P, Pc c))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    77
  "add (Pinj i P, Pc c) = Pinj i (add (P, Pc c))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    78
  "add (Pc c, PX P i Q) = PX P i (add (Q, Pc c))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    79
  "add (PX P i Q, Pc c) = PX P i (add (Q, Pc c))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    80
  "add (Pinj x P, Pinj y Q) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    81
  (if x=y then mkPinj x (add (P, Q))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    82
   else (if x>y then mkPinj y (add (Pinj (x-y) P, Q))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    83
         else mkPinj x (add (Pinj (y-x) Q, P)) ))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    84
  "add (Pinj x P, PX Q y R) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    85
  (if x=0 then add(P, PX Q y R)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    86
   else (if x=1 then PX Q y (add (R, P))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    87
         else PX Q y (add (R, Pinj (x - 1) P))))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    88
  "add (PX P x R, Pinj y Q) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    89
  (if y=0 then add(PX P x R, Q)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    90
   else (if y=1 then PX P x (add (R, Q))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    91
         else PX P x (add (R, Pinj (y - 1) Q))))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    92
  "add (PX P1 x P2, PX Q1 y Q2) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    93
  (if x=y then mkPX (add (P1, Q1)) x (add (P2, Q2))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    94
  else (if x>y then mkPX (add (PX P1 (x-y) (Pc 0), Q1)) y (add (P2,Q2))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    95
        else mkPX (add (PX Q1 (y-x) (Pc 0), P1)) x (add (P2,Q2)) ))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    96
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    97
text {* Multiplication *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    98
recdef mul "measure (\<lambda>(x, y). size x + size y)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    99
  "mul (Pc a, Pc b) = Pc (a*b)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   100
  "mul (Pc c, Pinj i P) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   101
  "mul (Pinj i P, Pc c) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   102
  "mul (Pc c, PX P i Q) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   103
  (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   104
  "mul (PX P i Q, Pc c) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   105
  (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   106
  "mul (Pinj x P, Pinj y Q) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   107
  (if x=y then mkPinj x (mul (P, Q))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   108
   else (if x>y then mkPinj y (mul (Pinj (x-y) P, Q))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   109
         else mkPinj x (mul (Pinj (y-x) Q, P)) ))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   110
  "mul (Pinj x P, PX Q y R) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   111
  (if x=0 then mul(P, PX Q y R)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   112
   else (if x=1 then mkPX (mul (Pinj x P, Q)) y (mul (R, P))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   113
         else mkPX (mul (Pinj x P, Q)) y (mul (R, Pinj (x - 1) P))))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   114
  "mul (PX P x R, Pinj y Q) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   115
  (if y=0 then mul(PX P x R, Q)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   116
   else (if y=1 then mkPX (mul (Pinj y Q, P)) x (mul (R, Q))
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   117
         else mkPX (mul (Pinj y Q, P)) x (mul (R, Pinj (y - 1) Q))))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   118
  "mul (PX P1 x P2, PX Q1 y Q2) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   119
  add (mkPX (mul (P1, Q1)) (x+y) (mul (P2, Q2)),
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   120
  add (mkPX (mul (P1, mkPinj 1 Q2)) x (Pc 0), mkPX (mul (Q1, mkPinj 1 P2)) y (Pc 0)) )"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   121
(hints simp add: mkPinj_def split: pol.split)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   122
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   123
text {* Negation*}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   124
primrec
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   125
  "neg (Pc c) = Pc (-c)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   126
  "neg (Pinj i P) = Pinj i (neg P)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   127
  "neg (PX P x Q) = PX (neg P) x (neg Q)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   128
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   129
text {* Substraction *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   130
constdefs
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   131
  sub :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   132
  "sub p q \<equiv> add (p, neg q)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   133
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   134
text {* Square for Fast Exponentation *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   135
primrec
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   136
  "sqr (Pc c) = Pc (c * c)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   137
  "sqr (Pinj i P) = mkPinj i (sqr P)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   138
  "sqr (PX A x B) = add (mkPX (sqr A) (x + x) (sqr B),
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   139
    mkPX (mul (mul (Pc (1 + 1), A), mkPinj 1 B)) x (Pc 0))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   140
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   141
text {* Fast Exponentation *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   142
lemma pow_wf:"odd n \<Longrightarrow> (n::nat) div 2 < n" by (cases n) auto
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   143
recdef pow "measure (\<lambda>(x, y). y)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   144
  "pow (p, 0) = Pc 1"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   145
  "pow (p, n) = (if even n then (pow (sqr p, n div 2)) else mul (p, pow (sqr p, n div 2)))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   146
(hints simp add: pow_wf)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   147
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   148
lemma pow_if:
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   149
  "pow (p,n) =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   150
   (if n = 0 then Pc 1 else if even n then pow (sqr p, n div 2)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   151
    else mul (p, pow (sqr p, n div 2)))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   152
  by (cases n) simp_all
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   153
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   154
(*
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   155
lemma number_of_nat_B0: "(number_of (w BIT bit.B0) ::nat) = 2* (number_of w)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   156
by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   157
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   158
lemma number_of_nat_even: "even (number_of (w BIT bit.B0)::nat)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   159
by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   160
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   161
lemma pow_even : "pow (p, number_of(w BIT bit.B0)) = pow (sqr p, number_of w)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   162
  ( is "pow(?p,?n) = pow (_,?n2)")
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   163
proof-
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   164
  have "even ?n" by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   165
  hence "pow (p, ?n) = pow (sqr p, ?n div 2)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   166
    apply simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   167
    apply (cases "IntDef.neg (number_of w)")
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   168
    apply simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   169
    done
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   170
*)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   171
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   172
text {* Normalization of polynomial expressions *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   173
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   174
consts norm :: "'a::{comm_ring,recpower} polex \<Rightarrow> 'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   175
primrec
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   176
  "norm (Pol P) = P"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   177
  "norm (Add P Q) = add (norm P, norm Q)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   178
  "norm (Sub p q) = sub (norm p) (norm q)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   179
  "norm (Mul P Q) = mul (norm P, norm Q)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   180
  "norm (Pow p n) = pow (norm p, n)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   181
  "norm (Neg P) = neg (norm P)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   182
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   183
text {* mkPinj preserve semantics *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   184
lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   185
  by (induct B) (auto simp add: mkPinj_def ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   186
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   187
text {* mkPX preserves semantics *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   188
lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   189
  by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   190
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   191
text {* Correctness theorems for the implemented operations *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   192
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   193
text {* Negation *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   194
lemma neg_ci: "\<And>l. Ipol l (neg P) = -(Ipol l P)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   195
  by (induct P) auto
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   196
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   197
text {* Addition *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   198
lemma add_ci: "\<And>l. Ipol l (add (P, Q)) = Ipol l P + Ipol l Q"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   199
proof (induct P Q rule: add.induct)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   200
  case (6 x P y Q)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   201
  show ?case
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   202
  proof (rule linorder_cases)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   203
    assume "x < y"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   204
    with 6 show ?case by (simp add: mkPinj_ci ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   205
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   206
    assume "x = y"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   207
    with 6 show ?case by (simp add: mkPinj_ci)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   208
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   209
    assume "x > y"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   210
    with 6 show ?case by (simp add: mkPinj_ci ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   211
  qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   212
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   213
  case (7 x P Q y R)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   214
  have "x = 0 \<or> x = 1 \<or> x > 1" by arith
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   215
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   216
  { assume "x = 0" with 7 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   217
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   218
  { assume "x = 1" with 7 have ?case by (simp add: ring_eq_simps) }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   219
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   220
  { assume "x > 1" from 7 have ?case by (cases x) simp_all }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   221
  ultimately show ?case by blast
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   222
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   223
  case (8 P x R y Q)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   224
  have "y = 0 \<or> y = 1 \<or> y > 1" by arith
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   225
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   226
  { assume "y = 0" with 8 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   227
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   228
  { assume "y = 1" with 8 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   229
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   230
  { assume "y > 1" with 8 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   231
  ultimately show ?case by blast
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   232
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   233
  case (9 P1 x P2 Q1 y Q2)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   234
  show ?case
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   235
  proof (rule linorder_cases)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   236
    assume a: "x < y" hence "EX d. d + x = y" by arith
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   237
    with 9 a show ?case by (auto simp add: mkPX_ci power_add ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   238
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   239
    assume a: "y < x" hence "EX d. d + y = x" by arith
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   240
    with 9 a show ?case by (auto simp add: power_add mkPX_ci ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   241
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   242
    assume "x = y"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   243
    with 9 show ?case by (simp add: mkPX_ci ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   244
  qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   245
qed (auto simp add: ring_eq_simps)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   246
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   247
text {* Multiplication *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   248
lemma mul_ci: "\<And>l. Ipol l (mul (P, Q)) = Ipol l P * Ipol l Q"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   249
  by (induct P Q rule: mul.induct)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   250
    (simp_all add: mkPX_ci mkPinj_ci ring_eq_simps add_ci power_add)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   251
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   252
text {* Substraction *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   253
lemma sub_ci: "Ipol l (sub p q) = Ipol l p - Ipol l q"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   254
  by (simp add: add_ci neg_ci sub_def)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   255
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   256
text {* Square *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   257
lemma sqr_ci:"\<And>ls. Ipol ls (sqr p) = Ipol ls p * Ipol ls p"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   258
  by (induct p) (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci ring_eq_simps power_add)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   259
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   260
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   261
text {* Power *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   262
lemma even_pow:"even n \<Longrightarrow> pow (p, n) = pow (sqr p, n div 2)" by (induct n) simp_all
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   263
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   264
lemma pow_ci: "\<And>p. Ipol ls (pow (p, n)) = (Ipol ls p) ^ n"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   265
proof (induct n rule: nat_less_induct)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   266
  case (1 k)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   267
  have two:"2 = Suc (Suc 0)" by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   268
  show ?case
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   269
  proof (cases k)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   270
    case (Suc l)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   271
    show ?thesis
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   272
    proof cases
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   273
      assume EL: "even l"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   274
      have "Suc l div 2 = l div 2"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   275
        by (simp add: nat_number even_nat_plus_one_div_two [OF EL])
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   276
      moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   277
      from Suc have "l < k" by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   278
      with 1 have "\<forall>p. Ipol ls (pow (p, l)) = Ipol ls p ^ l" by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   279
      moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   280
      note Suc EL even_nat_plus_one_div_two [OF EL]
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   281
      ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   282
    next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   283
      assume OL: "odd l"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   284
      with prems have "\<lbrakk>\<forall>m<Suc l. \<forall>p. Ipol ls (pow (p, m)) = Ipol ls p ^ m; k = Suc l; odd l\<rbrakk> \<Longrightarrow> \<forall>p. Ipol ls (sqr p) ^ (Suc l div 2) = Ipol ls p ^ Suc l"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   285
      proof(cases l)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   286
        case (Suc w)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   287
        from prems have EW: "even w" by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   288
        from two have two_times:"(2 * (w div 2))= w"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   289
          by (simp only: even_nat_div_two_times_two[OF EW])
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   290
        have A: "\<And>p. (Ipol ls p * Ipol ls p) = (Ipol ls p) ^ (Suc (Suc 0))"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   291
          by (simp add: power_Suc)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   292
        from A two [symmetric] have "ALL p.(Ipol ls p * Ipol ls p) = (Ipol ls p) ^ 2"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   293
          by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   294
        with prems show ?thesis
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   295
          by (auto simp add: power_mult[symmetric, of _ 2 _] two_times mul_ci sqr_ci)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   296
      qed simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   297
      with prems show ?thesis by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   298
    qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   299
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   300
    case 0
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   301
    then show ?thesis by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   302
  qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   303
qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   304
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   305
text {* Normalization preserves semantics  *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   306
lemma norm_ci:"Ipolex l Pe = Ipol l (norm Pe)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   307
  by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   308
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   309
text {* Reflection lemma: Key to the (incomplete) decision procedure *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   310
lemma norm_eq:
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   311
  assumes eq: "norm P1  = norm P2"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   312
  shows "Ipolex l P1 = Ipolex l P2"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   313
proof -
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   314
  from eq have "Ipol l (norm P1) = Ipol l (norm P2)" by simp
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   315
  thus ?thesis by (simp only: norm_ci)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   316
qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   317
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   318
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   319
text {* Code generation *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   320
(*
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   321
Does not work, since no generic ring operations implementation is there
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   322
generate_code ("ring.ML") test = "norm"*)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   323
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   324
use "comm_ring.ML"
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 17516
diff changeset
   325
setup CommRing.setup
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   326
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   327
end