| author | huffman | 
| Tue, 24 Feb 2009 11:12:58 -0800 | |
| changeset 30082 | 43c5b7bfc791 | 
| parent 29269 | 5c25a2012975 | 
| child 30866 | dd5117e2d73e | 
| permissions | -rw-r--r-- | 
| 23252 | 1 | (* Title: HOL/Tools/Groebner_Basis/normalizer.ML | 
| 2 | Author: Amine Chaieb, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | signature NORMALIZER = | |
| 6 | sig | |
| 23485 | 7 | val semiring_normalize_conv : Proof.context -> conv | 
| 8 | val semiring_normalize_ord_conv : Proof.context -> (cterm -> cterm -> bool) -> conv | |
| 23252 | 9 | val semiring_normalize_tac : Proof.context -> int -> tactic | 
| 23485 | 10 | val semiring_normalize_wrapper : Proof.context -> NormalizerData.entry -> conv | 
| 27222 | 11 | val semiring_normalizers_ord_wrapper : | 
| 12 | Proof.context -> NormalizerData.entry -> (cterm -> cterm -> bool) -> | |
| 13 |       {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv}
 | |
| 23485 | 14 | val semiring_normalize_ord_wrapper : Proof.context -> NormalizerData.entry -> | 
| 15 | (cterm -> cterm -> bool) -> conv | |
| 23252 | 16 | val semiring_normalizers_conv : | 
| 17 | cterm list -> cterm list * thm list -> cterm list * thm list -> | |
| 23485 | 18 | (cterm -> bool) * conv * conv * conv -> (cterm -> cterm -> bool) -> | 
| 19 |        {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv}
 | |
| 23252 | 20 | end | 
| 21 | ||
| 22 | structure Normalizer: NORMALIZER = | |
| 23 | struct | |
| 23559 | 24 | |
| 25253 | 25 | open Conv; | 
| 23252 | 26 | |
| 27 | (* Very basic stuff for terms *) | |
| 25253 | 28 | |
| 29 | fun is_comb ct = | |
| 30 | (case Thm.term_of ct of | |
| 31 | _ $ _ => true | |
| 32 | | _ => false); | |
| 33 | ||
| 34 | val concl = Thm.cprop_of #> Thm.dest_arg; | |
| 35 | ||
| 36 | fun is_binop ct ct' = | |
| 37 | (case Thm.term_of ct' of | |
| 38 | c $ _ $ _ => term_of ct aconv c | |
| 39 | | _ => false); | |
| 40 | ||
| 41 | fun dest_binop ct ct' = | |
| 42 | if is_binop ct ct' then Thm.dest_binop ct' | |
| 43 |   else raise CTERM ("dest_binop: bad binop", [ct, ct'])
 | |
| 44 | ||
| 45 | fun inst_thm inst = Thm.instantiate ([], inst); | |
| 46 | ||
| 23252 | 47 | val dest_numeral = term_of #> HOLogic.dest_number #> snd; | 
| 48 | val is_numeral = can dest_numeral; | |
| 49 | ||
| 50 | val numeral01_conv = Simplifier.rewrite | |
| 25481 | 51 |                          (HOL_basic_ss addsimps [@{thm numeral_1_eq_1}, @{thm numeral_0_eq_0}]);
 | 
| 23252 | 52 | val zero1_numeral_conv = | 
| 25481 | 53 |  Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1} RS sym, @{thm numeral_0_eq_0} RS sym]);
 | 
| 23580 
998a6fda9bb6
moved mk_cnumeral/mk_cnumber to Tools/numeral.ML;
 wenzelm parents: 
23559diff
changeset | 54 | fun zerone_conv cv = zero1_numeral_conv then_conv cv then_conv numeral01_conv; | 
| 23252 | 55 | val natarith = [@{thm "add_nat_number_of"}, @{thm "diff_nat_number_of"},
 | 
| 56 |                 @{thm "mult_nat_number_of"}, @{thm "eq_nat_number_of"}, 
 | |
| 57 |                 @{thm "less_nat_number_of"}];
 | |
| 58 | val nat_add_conv = | |
| 59 | zerone_conv | |
| 60 | (Simplifier.rewrite | |
| 61 | (HOL_basic_ss | |
| 25481 | 62 |        addsimps @{thms arith_simps} @ natarith @ @{thms rel_simps}
 | 
| 63 |              @ [if_False, if_True, @{thm add_0}, @{thm add_Suc},
 | |
| 64 |                  @{thm add_number_of_left}, @{thm Suc_eq_add_numeral_1}]
 | |
| 65 |              @ map (fn th => th RS sym) @{thms numerals}));
 | |
| 23252 | 66 | |
| 67 | val nat_mul_conv = nat_add_conv; | |
| 68 | val zeron_tm = @{cterm "0::nat"};
 | |
| 69 | val onen_tm  = @{cterm "1::nat"};
 | |
| 70 | val true_tm = @{cterm "True"};
 | |
| 71 | ||
| 72 | ||
| 73 | (* The main function! *) | |
| 74 | fun semiring_normalizers_conv vars (sr_ops, sr_rules) (r_ops, r_rules) | |
| 75 | (is_semiring_constant, semiring_add_conv, semiring_mul_conv, semiring_pow_conv) = | |
| 76 | let | |
| 77 | ||
| 78 | val [pthm_02, pthm_03, pthm_04, pthm_05, pthm_07, pthm_08, | |
| 79 | pthm_09, pthm_10, pthm_11, pthm_12, pthm_13, pthm_14, pthm_15, pthm_16, | |
| 80 | pthm_17, pthm_18, pthm_19, pthm_21, pthm_22, pthm_23, pthm_24, | |
| 81 | pthm_25, pthm_26, pthm_27, pthm_28, pthm_29, pthm_30, pthm_31, pthm_32, | |
| 82 | pthm_33, pthm_34, pthm_35, pthm_36, pthm_37, pthm_38,pthm_39,pthm_40] = sr_rules; | |
| 83 | ||
| 84 | val [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry] = vars; | |
| 85 | val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; | |
| 86 | val [add_tm, mul_tm, pow_tm] = map (Thm.dest_fun o Thm.dest_fun) [add_pat, mul_pat, pow_pat]; | |
| 87 | ||
| 88 | val dest_add = dest_binop add_tm | |
| 89 | val dest_mul = dest_binop mul_tm | |
| 90 | fun dest_pow tm = | |
| 91 | let val (l,r) = dest_binop pow_tm tm | |
| 92 |  in if is_numeral r then (l,r) else raise CTERM ("dest_pow",[tm])
 | |
| 93 | end; | |
| 94 | val is_add = is_binop add_tm | |
| 95 | val is_mul = is_binop mul_tm | |
| 96 | fun is_pow tm = is_binop pow_tm tm andalso is_numeral(Thm.dest_arg tm); | |
| 97 | ||
| 98 | val (neg_mul,sub_add,sub_tm,neg_tm,dest_sub,is_sub,cx',cy') = | |
| 99 | (case (r_ops, r_rules) of | |
| 100 | ([], []) => (TrueI, TrueI, true_tm, true_tm, (fn t => (t,t)), K false, true_tm, true_tm) | |
| 101 | | ([sub_pat, neg_pat], [neg_mul, sub_add]) => | |
| 102 | let | |
| 103 | val sub_tm = Thm.dest_fun (Thm.dest_fun sub_pat) | |
| 104 | val neg_tm = Thm.dest_fun neg_pat | |
| 105 | val dest_sub = dest_binop sub_tm | |
| 106 | val is_sub = is_binop sub_tm | |
| 107 | in (neg_mul,sub_add,sub_tm,neg_tm,dest_sub,is_sub, neg_mul |> concl |> Thm.dest_arg, | |
| 108 | sub_add |> concl |> Thm.dest_arg |> Thm.dest_arg) | |
| 109 | end); | |
| 110 | in fn variable_order => | |
| 111 | let | |
| 112 | ||
| 113 | (* Conversion for "x^n * x^m", with either x^n = x and/or x^m = x possible. *) | |
| 114 | (* Also deals with "const * const", but both terms must involve powers of *) | |
| 115 | (* the same variable, or both be constants, or behaviour may be incorrect. *) | |
| 116 | ||
| 117 | fun powvar_mul_conv tm = | |
| 118 | let | |
| 119 | val (l,r) = dest_mul tm | |
| 120 | in if is_semiring_constant l andalso is_semiring_constant r | |
| 121 | then semiring_mul_conv tm | |
| 122 | else | |
| 123 | ((let | |
| 124 | val (lx,ln) = dest_pow l | |
| 125 | in | |
| 126 | ((let val (rx,rn) = dest_pow r | |
| 127 | val th1 = inst_thm [(cx,lx),(cp,ln),(cq,rn)] pthm_29 | |
| 128 | val (tm1,tm2) = Thm.dest_comb(concl th1) in | |
| 129 | transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end) | |
| 130 | handle CTERM _ => | |
| 131 | (let val th1 = inst_thm [(cx,lx),(cq,ln)] pthm_31 | |
| 132 | val (tm1,tm2) = Thm.dest_comb(concl th1) in | |
| 133 | transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)) end) | |
| 134 | handle CTERM _ => | |
| 135 | ((let val (rx,rn) = dest_pow r | |
| 136 | val th1 = inst_thm [(cx,rx),(cq,rn)] pthm_30 | |
| 137 | val (tm1,tm2) = Thm.dest_comb(concl th1) in | |
| 138 | transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end) | |
| 139 | handle CTERM _ => inst_thm [(cx,l)] pthm_32 | |
| 140 | ||
| 141 | )) | |
| 142 | end; | |
| 143 | ||
| 144 | (* Remove "1 * m" from a monomial, and just leave m. *) | |
| 145 | ||
| 146 | fun monomial_deone th = | |
| 147 | (let val (l,r) = dest_mul(concl th) in | |
| 148 | if l aconvc one_tm | |
| 149 | then transitive th (inst_thm [(ca,r)] pthm_13) else th end) | |
| 150 | handle CTERM _ => th; | |
| 151 | ||
| 152 | (* Conversion for "(monomial)^n", where n is a numeral. *) | |
| 153 | ||
| 154 | val monomial_pow_conv = | |
| 155 | let | |
| 156 | fun monomial_pow tm bod ntm = | |
| 157 | if not(is_comb bod) | |
| 158 | then reflexive tm | |
| 159 | else | |
| 160 | if is_semiring_constant bod | |
| 161 | then semiring_pow_conv tm | |
| 162 | else | |
| 163 | let | |
| 164 | val (lopr,r) = Thm.dest_comb bod | |
| 165 | in if not(is_comb lopr) | |
| 166 | then reflexive tm | |
| 167 | else | |
| 168 | let | |
| 169 | val (opr,l) = Thm.dest_comb lopr | |
| 170 | in | |
| 171 | if opr aconvc pow_tm andalso is_numeral r | |
| 172 | then | |
| 173 | let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34 | |
| 174 | val (l,r) = Thm.dest_comb(concl th1) | |
| 175 | in transitive th1 (Drule.arg_cong_rule l (nat_mul_conv r)) | |
| 176 | end | |
| 177 | else | |
| 178 | if opr aconvc mul_tm | |
| 179 | then | |
| 180 | let | |
| 181 | val th1 = inst_thm [(cx,l),(cy,r),(cq,ntm)] pthm_33 | |
| 182 | val (xy,z) = Thm.dest_comb(concl th1) | |
| 183 | val (x,y) = Thm.dest_comb xy | |
| 184 | val thl = monomial_pow y l ntm | |
| 185 | val thr = monomial_pow z r ntm | |
| 186 | in transitive th1 (combination (Drule.arg_cong_rule x thl) thr) | |
| 187 | end | |
| 188 | else reflexive tm | |
| 189 | end | |
| 190 | end | |
| 191 | in fn tm => | |
| 192 | let | |
| 193 | val (lopr,r) = Thm.dest_comb tm | |
| 194 | val (opr,l) = Thm.dest_comb lopr | |
| 195 | in if not (opr aconvc pow_tm) orelse not(is_numeral r) | |
| 196 |       then raise CTERM ("monomial_pow_conv", [tm])
 | |
| 197 | else if r aconvc zeron_tm | |
| 198 | then inst_thm [(cx,l)] pthm_35 | |
| 199 | else if r aconvc onen_tm | |
| 200 | then inst_thm [(cx,l)] pthm_36 | |
| 201 | else monomial_deone(monomial_pow tm l r) | |
| 202 | end | |
| 203 | end; | |
| 204 | ||
| 205 | (* Multiplication of canonical monomials. *) | |
| 206 | val monomial_mul_conv = | |
| 207 | let | |
| 208 | fun powvar tm = | |
| 209 | if is_semiring_constant tm then one_tm | |
| 210 | else | |
| 211 | ((let val (lopr,r) = Thm.dest_comb tm | |
| 212 | val (opr,l) = Thm.dest_comb lopr | |
| 213 | in if opr aconvc pow_tm andalso is_numeral r then l | |
| 214 |           else raise CTERM ("monomial_mul_conv",[tm]) end)
 | |
| 215 | handle CTERM _ => tm) (* FIXME !? *) | |
| 216 | fun vorder x y = | |
| 217 | if x aconvc y then 0 | |
| 218 | else | |
| 219 | if x aconvc one_tm then ~1 | |
| 220 | else if y aconvc one_tm then 1 | |
| 221 | else if variable_order x y then ~1 else 1 | |
| 222 | fun monomial_mul tm l r = | |
| 223 | ((let val (lx,ly) = dest_mul l val vl = powvar lx | |
| 224 | in | |
| 225 | ((let | |
| 226 | val (rx,ry) = dest_mul r | |
| 227 | val vr = powvar rx | |
| 228 | val ord = vorder vl vr | |
| 229 | in | |
| 230 | if ord = 0 | |
| 231 | then | |
| 232 | let | |
| 233 | val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] pthm_15 | |
| 234 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 235 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 236 | val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2 | |
| 237 | val th3 = transitive th1 th2 | |
| 238 | val (tm5,tm6) = Thm.dest_comb(concl th3) | |
| 239 | val (tm7,tm8) = Thm.dest_comb tm6 | |
| 240 | val th4 = monomial_mul tm6 (Thm.dest_arg tm7) tm8 | |
| 241 | in transitive th3 (Drule.arg_cong_rule tm5 th4) | |
| 242 | end | |
| 243 | else | |
| 244 | let val th0 = if ord < 0 then pthm_16 else pthm_17 | |
| 245 | val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] th0 | |
| 246 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 247 | val (tm3,tm4) = Thm.dest_comb tm2 | |
| 248 | in transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) | |
| 249 | end | |
| 250 | end) | |
| 251 | handle CTERM _ => | |
| 252 | (let val vr = powvar r val ord = vorder vl vr | |
| 253 | in | |
| 254 | if ord = 0 then | |
| 255 | let | |
| 256 | val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_18 | |
| 257 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 258 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 259 | val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2 | |
| 260 | in transitive th1 th2 | |
| 261 | end | |
| 262 | else | |
| 263 | if ord < 0 then | |
| 264 | let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_19 | |
| 265 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 266 | val (tm3,tm4) = Thm.dest_comb tm2 | |
| 267 | in transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) | |
| 268 | end | |
| 269 | else inst_thm [(ca,l),(cb,r)] pthm_09 | |
| 270 | end)) end) | |
| 271 | handle CTERM _ => | |
| 272 | (let val vl = powvar l in | |
| 273 | ((let | |
| 274 | val (rx,ry) = dest_mul r | |
| 275 | val vr = powvar rx | |
| 276 | val ord = vorder vl vr | |
| 277 | in if ord = 0 then | |
| 278 | let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_21 | |
| 279 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 280 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 281 | in transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2) | |
| 282 | end | |
| 283 | else if ord > 0 then | |
| 284 | let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_22 | |
| 285 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 286 | val (tm3,tm4) = Thm.dest_comb tm2 | |
| 287 | in transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) | |
| 288 | end | |
| 289 | else reflexive tm | |
| 290 | end) | |
| 291 | handle CTERM _ => | |
| 292 | (let val vr = powvar r | |
| 293 | val ord = vorder vl vr | |
| 294 | in if ord = 0 then powvar_mul_conv tm | |
| 295 | else if ord > 0 then inst_thm [(ca,l),(cb,r)] pthm_09 | |
| 296 | else reflexive tm | |
| 297 | end)) end)) | |
| 298 | in fn tm => let val (l,r) = dest_mul tm in monomial_deone(monomial_mul tm l r) | |
| 299 | end | |
| 300 | end; | |
| 301 | (* Multiplication by monomial of a polynomial. *) | |
| 302 | ||
| 303 | val polynomial_monomial_mul_conv = | |
| 304 | let | |
| 305 | fun pmm_conv tm = | |
| 306 | let val (l,r) = dest_mul tm | |
| 307 | in | |
| 308 | ((let val (y,z) = dest_add r | |
| 309 | val th1 = inst_thm [(cx,l),(cy,y),(cz,z)] pthm_37 | |
| 310 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 311 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 312 | val th2 = combination (Drule.arg_cong_rule tm3 (monomial_mul_conv tm4)) (pmm_conv tm2) | |
| 313 | in transitive th1 th2 | |
| 314 | end) | |
| 315 | handle CTERM _ => monomial_mul_conv tm) | |
| 316 | end | |
| 317 | in pmm_conv | |
| 318 | end; | |
| 319 | ||
| 320 | (* Addition of two monomials identical except for constant multiples. *) | |
| 321 | ||
| 322 | fun monomial_add_conv tm = | |
| 323 | let val (l,r) = dest_add tm | |
| 324 | in if is_semiring_constant l andalso is_semiring_constant r | |
| 325 | then semiring_add_conv tm | |
| 326 | else | |
| 327 | let val th1 = | |
| 328 | if is_mul l andalso is_semiring_constant(Thm.dest_arg1 l) | |
| 329 | then if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then | |
| 330 | inst_thm [(ca,Thm.dest_arg1 l),(cm,Thm.dest_arg r), (cb,Thm.dest_arg1 r)] pthm_02 | |
| 331 | else inst_thm [(ca,Thm.dest_arg1 l),(cm,r)] pthm_03 | |
| 332 | else if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) | |
| 333 | then inst_thm [(cm,l),(ca,Thm.dest_arg1 r)] pthm_04 | |
| 334 | else inst_thm [(cm,r)] pthm_05 | |
| 335 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 336 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 337 | val th2 = Drule.arg_cong_rule tm3 (semiring_add_conv tm4) | |
| 338 | val th3 = transitive th1 (Drule.fun_cong_rule th2 tm2) | |
| 339 | val tm5 = concl th3 | |
| 340 | in | |
| 341 | if (Thm.dest_arg1 tm5) aconvc zero_tm | |
| 342 | then transitive th3 (inst_thm [(ca,Thm.dest_arg tm5)] pthm_11) | |
| 343 | else monomial_deone th3 | |
| 344 | end | |
| 345 | end; | |
| 346 | ||
| 347 | (* Ordering on monomials. *) | |
| 348 | ||
| 349 | fun striplist dest = | |
| 350 | let fun strip x acc = | |
| 351 | ((let val (l,r) = dest x in | |
| 352 | strip l (strip r acc) end) | |
| 353 | handle CTERM _ => x::acc) (* FIXME !? *) | |
| 354 | in fn x => strip x [] | |
| 355 | end; | |
| 356 | ||
| 357 | ||
| 358 | fun powervars tm = | |
| 359 | let val ptms = striplist dest_mul tm | |
| 360 | in if is_semiring_constant (hd ptms) then tl ptms else ptms | |
| 361 | end; | |
| 362 | val num_0 = 0; | |
| 363 | val num_1 = 1; | |
| 364 | fun dest_varpow tm = | |
| 365 | ((let val (x,n) = dest_pow tm in (x,dest_numeral n) end) | |
| 366 | handle CTERM _ => | |
| 367 | (tm,(if is_semiring_constant tm then num_0 else num_1))); | |
| 368 | ||
| 369 | val morder = | |
| 370 | let fun lexorder l1 l2 = | |
| 371 | case (l1,l2) of | |
| 372 | ([],[]) => 0 | |
| 373 | | (vps,[]) => ~1 | |
| 374 | | ([],vps) => 1 | |
| 375 | | (((x1,n1)::vs1),((x2,n2)::vs2)) => | |
| 376 | if variable_order x1 x2 then 1 | |
| 377 | else if variable_order x2 x1 then ~1 | |
| 378 | else if n1 < n2 then ~1 | |
| 379 | else if n2 < n1 then 1 | |
| 380 | else lexorder vs1 vs2 | |
| 381 | in fn tm1 => fn tm2 => | |
| 382 | let val vdegs1 = map dest_varpow (powervars tm1) | |
| 383 | val vdegs2 = map dest_varpow (powervars tm2) | |
| 384 | val deg1 = fold_rev ((curry (op +)) o snd) vdegs1 num_0 | |
| 385 | val deg2 = fold_rev ((curry (op +)) o snd) vdegs2 num_0 | |
| 386 | in if deg1 < deg2 then ~1 else if deg1 > deg2 then 1 | |
| 387 | else lexorder vdegs1 vdegs2 | |
| 388 | end | |
| 389 | end; | |
| 390 | ||
| 391 | (* Addition of two polynomials. *) | |
| 392 | ||
| 393 | val polynomial_add_conv = | |
| 394 | let | |
| 395 | fun dezero_rule th = | |
| 396 | let | |
| 397 | val tm = concl th | |
| 398 | in | |
| 399 | if not(is_add tm) then th else | |
| 400 | let val (lopr,r) = Thm.dest_comb tm | |
| 401 | val l = Thm.dest_arg lopr | |
| 402 | in | |
| 403 | if l aconvc zero_tm | |
| 404 | then transitive th (inst_thm [(ca,r)] pthm_07) else | |
| 405 | if r aconvc zero_tm | |
| 406 | then transitive th (inst_thm [(ca,l)] pthm_08) else th | |
| 407 | end | |
| 408 | end | |
| 409 | fun padd tm = | |
| 410 | let | |
| 411 | val (l,r) = dest_add tm | |
| 412 | in | |
| 413 | if l aconvc zero_tm then inst_thm [(ca,r)] pthm_07 | |
| 414 | else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_08 | |
| 415 | else | |
| 416 | if is_add l | |
| 417 | then | |
| 418 | let val (a,b) = dest_add l | |
| 419 | in | |
| 420 | if is_add r then | |
| 421 | let val (c,d) = dest_add r | |
| 422 | val ord = morder a c | |
| 423 | in | |
| 424 | if ord = 0 then | |
| 425 | let val th1 = inst_thm [(ca,a),(cb,b),(cc,c),(cd,d)] pthm_23 | |
| 426 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 427 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 428 | val th2 = Drule.arg_cong_rule tm3 (monomial_add_conv tm4) | |
| 429 | in dezero_rule (transitive th1 (combination th2 (padd tm2))) | |
| 430 | end | |
| 431 | else (* ord <> 0*) | |
| 432 | let val th1 = | |
| 433 | if ord > 0 then inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 | |
| 434 | else inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 | |
| 435 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 436 | in dezero_rule (transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) | |
| 437 | end | |
| 438 | end | |
| 439 | else (* not (is_add r)*) | |
| 440 | let val ord = morder a r | |
| 441 | in | |
| 442 | if ord = 0 then | |
| 443 | let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_26 | |
| 444 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 445 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 446 | val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 | |
| 447 | in dezero_rule (transitive th1 th2) | |
| 448 | end | |
| 449 | else (* ord <> 0*) | |
| 450 | if ord > 0 then | |
| 451 | let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 | |
| 452 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 453 | in dezero_rule (transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) | |
| 454 | end | |
| 455 | else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) | |
| 456 | end | |
| 457 | end | |
| 458 | else (* not (is_add l)*) | |
| 459 | if is_add r then | |
| 460 | let val (c,d) = dest_add r | |
| 461 | val ord = morder l c | |
| 462 | in | |
| 463 | if ord = 0 then | |
| 464 | let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_28 | |
| 465 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 466 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 467 | val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 | |
| 468 | in dezero_rule (transitive th1 th2) | |
| 469 | end | |
| 470 | else | |
| 471 | if ord > 0 then reflexive tm | |
| 472 | else | |
| 473 | let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 | |
| 474 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 475 | in dezero_rule (transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) | |
| 476 | end | |
| 477 | end | |
| 478 | else | |
| 479 | let val ord = morder l r | |
| 480 | in | |
| 481 | if ord = 0 then monomial_add_conv tm | |
| 482 | else if ord > 0 then dezero_rule(reflexive tm) | |
| 483 | else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) | |
| 484 | end | |
| 485 | end | |
| 486 | in padd | |
| 487 | end; | |
| 488 | ||
| 489 | (* Multiplication of two polynomials. *) | |
| 490 | ||
| 491 | val polynomial_mul_conv = | |
| 492 | let | |
| 493 | fun pmul tm = | |
| 494 | let val (l,r) = dest_mul tm | |
| 495 | in | |
| 496 | if not(is_add l) then polynomial_monomial_mul_conv tm | |
| 497 | else | |
| 498 | if not(is_add r) then | |
| 499 | let val th1 = inst_thm [(ca,l),(cb,r)] pthm_09 | |
| 500 | in transitive th1 (polynomial_monomial_mul_conv(concl th1)) | |
| 501 | end | |
| 502 | else | |
| 503 | let val (a,b) = dest_add l | |
| 504 | val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_10 | |
| 505 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 506 | val (tm3,tm4) = Thm.dest_comb tm1 | |
| 507 | val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv tm4) | |
| 508 | val th3 = transitive th1 (combination th2 (pmul tm2)) | |
| 509 | in transitive th3 (polynomial_add_conv (concl th3)) | |
| 510 | end | |
| 511 | end | |
| 512 | in fn tm => | |
| 513 | let val (l,r) = dest_mul tm | |
| 514 | in | |
| 515 | if l aconvc zero_tm then inst_thm [(ca,r)] pthm_11 | |
| 516 | else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_12 | |
| 517 | else if l aconvc one_tm then inst_thm [(ca,r)] pthm_13 | |
| 518 | else if r aconvc one_tm then inst_thm [(ca,l)] pthm_14 | |
| 519 | else pmul tm | |
| 520 | end | |
| 521 | end; | |
| 522 | ||
| 523 | (* Power of polynomial (optimized for the monomial and trivial cases). *) | |
| 524 | ||
| 23580 
998a6fda9bb6
moved mk_cnumeral/mk_cnumber to Tools/numeral.ML;
 wenzelm parents: 
23559diff
changeset | 525 | fun num_conv n = | 
| 
998a6fda9bb6
moved mk_cnumeral/mk_cnumber to Tools/numeral.ML;
 wenzelm parents: 
23559diff
changeset | 526 |   nat_add_conv (Thm.capply @{cterm Suc} (Numeral.mk_cnumber @{ctyp nat} (dest_numeral n - 1)))
 | 
| 
998a6fda9bb6
moved mk_cnumeral/mk_cnumber to Tools/numeral.ML;
 wenzelm parents: 
23559diff
changeset | 527 | |> Thm.symmetric; | 
| 23252 | 528 | |
| 529 | ||
| 530 | val polynomial_pow_conv = | |
| 531 | let | |
| 532 | fun ppow tm = | |
| 533 | let val (l,n) = dest_pow tm | |
| 534 | in | |
| 535 | if n aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 | |
| 536 | else if n aconvc onen_tm then inst_thm [(cx,l)] pthm_36 | |
| 537 | else | |
| 538 | let val th1 = num_conv n | |
| 539 | val th2 = inst_thm [(cx,l),(cq,Thm.dest_arg (concl th1))] pthm_38 | |
| 540 | val (tm1,tm2) = Thm.dest_comb(concl th2) | |
| 541 | val th3 = transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2)) | |
| 542 | val th4 = transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3 | |
| 543 | in transitive th4 (polynomial_mul_conv (concl th4)) | |
| 544 | end | |
| 545 | end | |
| 546 | in fn tm => | |
| 547 | if is_add(Thm.dest_arg1 tm) then ppow tm else monomial_pow_conv tm | |
| 548 | end; | |
| 549 | ||
| 550 | (* Negation. *) | |
| 551 | ||
| 23580 
998a6fda9bb6
moved mk_cnumeral/mk_cnumber to Tools/numeral.ML;
 wenzelm parents: 
23559diff
changeset | 552 | fun polynomial_neg_conv tm = | 
| 23252 | 553 | let val (l,r) = Thm.dest_comb tm in | 
| 554 |         if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else
 | |
| 555 | let val th1 = inst_thm [(cx',r)] neg_mul | |
| 556 | val th2 = transitive th1 (arg1_conv semiring_mul_conv (concl th1)) | |
| 557 | in transitive th2 (polynomial_monomial_mul_conv (concl th2)) | |
| 558 | end | |
| 559 | end; | |
| 560 | ||
| 561 | ||
| 562 | (* Subtraction. *) | |
| 23580 
998a6fda9bb6
moved mk_cnumeral/mk_cnumber to Tools/numeral.ML;
 wenzelm parents: 
23559diff
changeset | 563 | fun polynomial_sub_conv tm = | 
| 23252 | 564 | let val (l,r) = dest_sub tm | 
| 565 | val th1 = inst_thm [(cx',l),(cy',r)] sub_add | |
| 566 | val (tm1,tm2) = Thm.dest_comb(concl th1) | |
| 567 | val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv tm2) | |
| 568 | in transitive th1 (transitive th2 (polynomial_add_conv (concl th2))) | |
| 569 | end; | |
| 570 | ||
| 571 | (* Conversion from HOL term. *) | |
| 572 | ||
| 573 | fun polynomial_conv tm = | |
| 23407 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 574 | if is_semiring_constant tm then semiring_add_conv tm | 
| 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 575 | else if not(is_comb tm) then reflexive tm | 
| 23252 | 576 | else | 
| 577 | let val (lopr,r) = Thm.dest_comb tm | |
| 578 | in if lopr aconvc neg_tm then | |
| 579 | let val th1 = Drule.arg_cong_rule lopr (polynomial_conv r) | |
| 580 | in transitive th1 (polynomial_neg_conv (concl th1)) | |
| 581 | end | |
| 582 | else | |
| 583 | if not(is_comb lopr) then reflexive tm | |
| 584 | else | |
| 585 | let val (opr,l) = Thm.dest_comb lopr | |
| 586 | in if opr aconvc pow_tm andalso is_numeral r | |
| 587 | then | |
| 588 | let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv l)) r | |
| 589 | in transitive th1 (polynomial_pow_conv (concl th1)) | |
| 590 | end | |
| 591 | else | |
| 592 | if opr aconvc add_tm orelse opr aconvc mul_tm orelse opr aconvc sub_tm | |
| 593 | then | |
| 594 | let val th1 = combination (Drule.arg_cong_rule opr (polynomial_conv l)) (polynomial_conv r) | |
| 595 | val f = if opr aconvc add_tm then polynomial_add_conv | |
| 596 | else if opr aconvc mul_tm then polynomial_mul_conv | |
| 597 | else polynomial_sub_conv | |
| 598 | in transitive th1 (f (concl th1)) | |
| 599 | end | |
| 600 | else reflexive tm | |
| 601 | end | |
| 602 | end; | |
| 603 | in | |
| 604 |    {main = polynomial_conv,
 | |
| 605 | add = polynomial_add_conv, | |
| 606 | mul = polynomial_mul_conv, | |
| 607 | pow = polynomial_pow_conv, | |
| 608 | neg = polynomial_neg_conv, | |
| 609 | sub = polynomial_sub_conv} | |
| 610 | end | |
| 611 | end; | |
| 612 | ||
| 613 | val nat_arith = @{thms "nat_arith"};
 | |
| 25481 | 614 | val nat_exp_ss = HOL_basic_ss addsimps (@{thms nat_number} @ nat_arith @ @{thms arith_simps} @ @{thms rel_simps})
 | 
| 23880 | 615 |                               addsimps [Let_def, if_False, if_True, @{thm add_0}, @{thm add_Suc}];
 | 
| 23252 | 616 | |
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
27222diff
changeset | 617 | fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS; | 
| 27222 | 618 | |
| 619 | fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, idom, ideal}, 
 | |
| 23407 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 620 |                                      {conv, dest_const, mk_const, is_const}) ord =
 | 
| 23252 | 621 | let | 
| 622 | val pow_conv = | |
| 623 | arg_conv (Simplifier.rewrite nat_exp_ss) | |
| 624 | then_conv Simplifier.rewrite | |
| 625 | (HOL_basic_ss addsimps [nth (snd semiring) 31, nth (snd semiring) 34]) | |
| 23330 
01c09922ce59
Conversion for computation on constants now depends on the context
 chaieb parents: 
23259diff
changeset | 626 | then_conv conv ctxt | 
| 
01c09922ce59
Conversion for computation on constants now depends on the context
 chaieb parents: 
23259diff
changeset | 627 | val dat = (is_const, conv ctxt, conv ctxt, pow_conv) | 
| 27222 | 628 | in semiring_normalizers_conv vars semiring ring dat ord end; | 
| 629 | ||
| 630 | fun semiring_normalize_ord_wrapper ctxt ({vars, semiring, ring, idom, ideal}, {conv, dest_const, mk_const, is_const}) ord =
 | |
| 631 |  #main (semiring_normalizers_ord_wrapper ctxt ({vars = vars, semiring = semiring, ring = ring, idom = idom, ideal = ideal},{conv = conv, dest_const = dest_const, mk_const = mk_const, is_const = is_const}) ord);
 | |
| 23252 | 632 | |
| 23407 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 633 | fun semiring_normalize_wrapper ctxt data = | 
| 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 634 | semiring_normalize_ord_wrapper ctxt data simple_cterm_ord; | 
| 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 635 | |
| 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 636 | fun semiring_normalize_ord_conv ctxt ord tm = | 
| 23252 | 637 | (case NormalizerData.match ctxt tm of | 
| 638 | NONE => reflexive tm | |
| 23407 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 639 | | SOME res => semiring_normalize_ord_wrapper ctxt res ord tm); | 
| 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 640 | |
| 23252 | 641 | |
| 23407 
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
 chaieb parents: 
23330diff
changeset | 642 | fun semiring_normalize_conv ctxt = semiring_normalize_ord_conv ctxt simple_cterm_ord; | 
| 23252 | 643 | |
| 644 | fun semiring_normalize_tac ctxt = SUBGOAL (fn (goal, i) => | |
| 645 | rtac (semiring_normalize_conv ctxt | |
| 646 | (cterm_of (ProofContext.theory_of ctxt) (fst (Logic.dest_equals goal)))) i); | |
| 647 | end; |