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