| author | huffman | 
| Sat, 30 Sep 2006 17:10:55 +0200 | |
| changeset 20792 | add17d26151b | 
| parent 20713 | 823967ef47f1 | 
| child 21588 | cd0dc678a205 | 
| permissions | -rw-r--r-- | 
| 17516 | 1 | (* ID: $Id$ | 
| 2 | Author: Amine Chaieb | |
| 3 | ||
| 4 | Tactic for solving equalities over commutative rings. | |
| 5 | *) | |
| 6 | ||
| 7 | signature COMM_RING = | |
| 8 | sig | |
| 20623 | 9 | val comm_ring_tac : Proof.context -> int -> tactic | 
| 18708 | 10 | val setup : theory -> theory | 
| 17516 | 11 | end | 
| 12 | ||
| 13 | structure CommRing: COMM_RING = | |
| 14 | struct | |
| 15 | ||
| 16 | (* The Cring exception for erronous uses of cring_tac *) | |
| 17 | exception CRing of string; | |
| 18 | ||
| 19 | (* Zero and One of the commutative ring *) | |
| 20713 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20623diff
changeset | 20 | fun cring_zero T = Const("HOL.zero",T);
 | 
| 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20623diff
changeset | 21 | fun cring_one T = Const("HOL.one",T);
 | 
| 17516 | 22 | |
| 23 | (* reification functions *) | |
| 24 | (* add two polynom expressions *) | |
| 25 | fun polT t = Type ("Commutative_Ring.pol",[t]);
 | |
| 20623 | 26 | fun polexT t = Type("Commutative_Ring.polex",[t]);
 | 
| 17516 | 27 | val nT = HOLogic.natT; | 
| 28 | fun listT T = Type ("List.list",[T]);
 | |
| 29 | ||
| 30 | (* Reification of the constructors *) | |
| 31 | (* Nat*) | |
| 32 | val succ = Const("Suc",nT --> nT);
 | |
| 20713 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20623diff
changeset | 33 | val zero = Const("HOL.zero",nT);
 | 
| 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20623diff
changeset | 34 | val one = Const("HOL.one",nT);
 | 
| 17516 | 35 | |
| 36 | (* Lists *) | |
| 37 | fun reif_list T [] = Const("List.list.Nil",listT T)
 | |
| 38 |   | reif_list T (x::xs) = Const("List.list.Cons",[T,listT T] ---> listT T)
 | |
| 39 | $x$(reif_list T xs); | |
| 40 | ||
| 20623 | 41 | (* pol *) | 
| 17516 | 42 | fun pol_Pc t = Const("Commutative_Ring.pol.Pc",t --> polT t);
 | 
| 43 | fun pol_Pinj t = Const("Commutative_Ring.pol.Pinj",[nT,polT t] ---> polT t);
 | |
| 44 | fun pol_PX t = Const("Commutative_Ring.pol.PX",[polT t, nT, polT t] ---> polT t);
 | |
| 45 | ||
| 46 | (* polex *) | |
| 47 | fun polex_add t = Const("Commutative_Ring.polex.Add",[polexT t,polexT t] ---> polexT t);
 | |
| 48 | fun polex_sub t = Const("Commutative_Ring.polex.Sub",[polexT t,polexT t] ---> polexT t);
 | |
| 49 | fun polex_mul t = Const("Commutative_Ring.polex.Mul",[polexT t,polexT t] ---> polexT t);
 | |
| 50 | fun polex_neg t = Const("Commutative_Ring.polex.Neg",polexT t --> polexT t);
 | |
| 51 | fun polex_pol t = Const("Commutative_Ring.polex.Pol",polT t --> polexT t);
 | |
| 52 | fun polex_pow t = Const("Commutative_Ring.polex.Pow",[polexT t, nT] ---> polexT t);
 | |
| 20623 | 53 | |
| 17516 | 54 | (* reification of natural numbers *) | 
| 55 | fun reif_nat n = | |
| 56 | if n>0 then succ$(reif_nat (n-1)) | |
| 57 | else if n=0 then zero | |
| 58 | else raise CRing "ring_tac: reif_nat negative n"; | |
| 59 | ||
| 60 | (* reification of polynoms : primitive cring expressions *) | |
| 61 | fun reif_pol T vs t = | |
| 62 | case t of | |
| 63 | Free(_,_) => | |
| 64 | let val i = find_index_eq t vs | |
| 65 | in if i = 0 | |
| 66 | then (pol_PX T)$((pol_Pc T)$ (cring_one T)) | |
| 67 | $one$((pol_Pc T)$(cring_zero T)) | |
| 68 | else (pol_Pinj T)$(reif_nat i)$ | |
| 69 | ((pol_PX T)$((pol_Pc T)$ (cring_one T)) | |
| 70 | $one$ | |
| 71 | ((pol_Pc T)$(cring_zero T))) | |
| 72 | end | |
| 73 | | _ => (pol_Pc T)$ t; | |
| 74 | ||
| 75 | ||
| 76 | (* reification of polynom expressions *) | |
| 77 | fun reif_polex T vs t = | |
| 78 | case t of | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 79 |         Const("HOL.plus",_)$a$b => (polex_add T)
 | 
| 17516 | 80 | $ (reif_polex T vs a)$(reif_polex T vs b) | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 81 |       | Const("HOL.minus",_)$a$b => (polex_sub T)
 | 
| 17516 | 82 | $ (reif_polex T vs a)$(reif_polex T vs b) | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 83 |       | Const("HOL.times",_)$a$b =>  (polex_mul T)
 | 
| 17516 | 84 | $ (reif_polex T vs a)$ (reif_polex T vs b) | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 85 |       | Const("HOL.uminus",_)$a => (polex_neg T)
 | 
| 17516 | 86 | $ (reif_polex T vs a) | 
| 87 |       | (Const("Nat.power",_)$a$n) => (polex_pow T) $ (reif_polex T vs a) $ n
 | |
| 88 | ||
| 89 | | _ => (polex_pol T) $ (reif_pol T vs t); | |
| 90 | ||
| 91 | (* reification of the equation *) | |
| 92 | val cr_sort = Sign.read_sort (the_context ()) "{comm_ring,recpower}";
 | |
| 20623 | 93 | fun reif_eq thy (eq as Const("op =",Type("fun",a::_))$lhs$rhs) =
 | 
| 94 | if Sign.of_sort thy (a,cr_sort) | |
| 17516 | 95 | then | 
| 96 | let val fs = term_frees eq | |
| 20623 | 97 | val cvs = cterm_of thy (reif_list a fs) | 
| 98 | val clhs = cterm_of thy (reif_polex a fs lhs) | |
| 99 | val crhs = cterm_of thy (reif_polex a fs rhs) | |
| 100 | val ca = ctyp_of thy a | |
| 17516 | 101 | in (ca,cvs,clhs, crhs) | 
| 102 | end | |
| 103 | else raise CRing "reif_eq: not an equation over comm_ring + recpower" | |
| 20623 | 104 | | reif_eq _ _ = raise CRing "reif_eq: not an equation"; | 
| 17516 | 105 | |
| 106 | (*The cring tactic *) | |
| 107 | (* Attention: You have to make sure that no t^0 is in the goal!! *) | |
| 108 | (* Use simply rewriting t^0 = 1 *) | |
| 20623 | 109 | val cring_simps = | 
| 110 | map thm ["mkPX_def", "mkPinj_def","sub_def", "power_add","even_def","pow_if"] @ | |
| 111 | [sym OF [thm "power_add"]]; | |
| 17516 | 112 | |
| 113 | val norm_eq = thm "norm_eq" | |
| 114 | ||
| 20623 | 115 | fun comm_ring_tac ctxt = SUBGOAL (fn (g, i) => | 
| 116 | let | |
| 117 | val thy = ProofContext.theory_of ctxt | |
| 118 | val cring_ss = Simplifier.local_simpset_of ctxt (* FIXME really the full simpset!? *) | |
| 119 | addsimps cring_simps | |
| 120 | val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) | |
| 121 | val norm_eq_th = | |
| 122 | simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] norm_eq) | |
| 123 | in | |
| 124 | cut_rules_tac [norm_eq_th] i | |
| 125 | THEN (simp_tac cring_ss i) | |
| 126 | THEN (simp_tac cring_ss i) | |
| 127 | end); | |
| 128 | ||
| 129 | val comm_ring_meth = | |
| 130 | Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD' HEADGOAL (comm_ring_tac ctxt)); | |
| 17516 | 131 | |
| 132 | val setup = | |
| 20623 | 133 |   Method.add_method ("comm_ring", comm_ring_meth,
 | 
| 134 | "reflective decision procedure for equalities over commutative rings") #> | |
| 135 |   Method.add_method ("algebra", comm_ring_meth,
 | |
| 136 | "method for proving algebraic properties (same as comm_ring)"); | |
| 17516 | 137 | |
| 138 | end; |