| author | wenzelm | 
| Tue, 14 Apr 2009 15:41:40 +0200 | |
| changeset 30890 | 0214d179c2be | 
| parent 30510 | 4120fc59dd85 | 
| child 31021 | 53642251a04f | 
| permissions | -rw-r--r-- | 
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
26939diff
changeset | 1 | (* Author: Amine Chaieb | 
| 17516 | 2 | |
| 3 | Tactic for solving equalities over commutative rings. | |
| 4 | *) | |
| 5 | ||
| 6 | signature COMM_RING = | |
| 7 | sig | |
| 20623 | 8 | val comm_ring_tac : Proof.context -> int -> tactic | 
| 18708 | 9 | val setup : theory -> theory | 
| 17516 | 10 | end | 
| 11 | ||
| 12 | structure CommRing: COMM_RING = | |
| 13 | struct | |
| 14 | ||
| 15 | (* The Cring exception for erronous uses of cring_tac *) | |
| 16 | exception CRing of string; | |
| 17 | ||
| 18 | (* Zero and One of the commutative ring *) | |
| 22997 | 19 | fun cring_zero T = Const (@{const_name HOL.zero}, T);
 | 
| 20 | fun cring_one T = Const (@{const_name HOL.one}, T);
 | |
| 17516 | 21 | |
| 22 | (* reification functions *) | |
| 23 | (* add two polynom expressions *) | |
| 22950 | 24 | fun polT t = Type ("Commutative_Ring.pol", [t]);
 | 
| 25 | fun polexT t = Type ("Commutative_Ring.polex", [t]);
 | |
| 17516 | 26 | |
| 20623 | 27 | (* pol *) | 
| 22950 | 28 | fun pol_Pc t = Const ("Commutative_Ring.pol.Pc", t --> polT t);
 | 
| 29 | fun pol_Pinj t = Const ("Commutative_Ring.pol.Pinj", HOLogic.natT --> polT t --> polT t);
 | |
| 30 | fun pol_PX t = Const ("Commutative_Ring.pol.PX", polT t --> HOLogic.natT --> polT t --> polT t);
 | |
| 17516 | 31 | |
| 32 | (* polex *) | |
| 22950 | 33 | fun polex_add t = Const ("Commutative_Ring.polex.Add", polexT t --> polexT t --> polexT t);
 | 
| 34 | fun polex_sub t = Const ("Commutative_Ring.polex.Sub", polexT t --> polexT t --> polexT t);
 | |
| 35 | fun polex_mul t = Const ("Commutative_Ring.polex.Mul", polexT t --> polexT t --> polexT t);
 | |
| 36 | fun polex_neg t = Const ("Commutative_Ring.polex.Neg", polexT t --> polexT t);
 | |
| 37 | fun polex_pol t = Const ("Commutative_Ring.polex.Pol", polT t --> polexT t);
 | |
| 38 | fun polex_pow t = Const ("Commutative_Ring.polex.Pow", polexT t --> HOLogic.natT --> polexT t);
 | |
| 17516 | 39 | |
| 40 | (* reification of polynoms : primitive cring expressions *) | |
| 22950 | 41 | fun reif_pol T vs (t as Free _) = | 
| 42 | let | |
| 43 |         val one = @{term "1::nat"};
 | |
| 44 | val i = find_index_eq t vs | |
| 45 | in if i = 0 | |
| 46 | then pol_PX T $ (pol_Pc T $ cring_one T) | |
| 47 | $ one $ (pol_Pc T $ cring_zero T) | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
23261diff
changeset | 48 | else pol_Pinj T $ HOLogic.mk_nat i | 
| 22950 | 49 | $ (pol_PX T $ (pol_Pc T $ cring_one T) | 
| 50 | $ one $ (pol_Pc T $ cring_zero T)) | |
| 17516 | 51 | end | 
| 22950 | 52 | | reif_pol T vs t = pol_Pc T $ t; | 
| 17516 | 53 | |
| 54 | (* reification of polynom expressions *) | |
| 22997 | 55 | fun reif_polex T vs (Const (@{const_name HOL.plus}, _) $ a $ b) =
 | 
| 22950 | 56 | polex_add T $ reif_polex T vs a $ reif_polex T vs b | 
| 22997 | 57 |   | reif_polex T vs (Const (@{const_name HOL.minus}, _) $ a $ b) =
 | 
| 22950 | 58 | polex_sub T $ reif_polex T vs a $ reif_polex T vs b | 
| 22997 | 59 |   | reif_polex T vs (Const (@{const_name HOL.times}, _) $ a $ b) =
 | 
| 22950 | 60 | polex_mul T $ reif_polex T vs a $ reif_polex T vs b | 
| 22997 | 61 |   | reif_polex T vs (Const (@{const_name HOL.uminus}, _) $ a) =
 | 
| 22950 | 62 | polex_neg T $ reif_polex T vs a | 
| 24996 | 63 |   | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) =
 | 
| 22950 | 64 | polex_pow T $ reif_polex T vs a $ n | 
| 65 | | reif_polex T vs t = polex_pol T $ reif_pol T vs t; | |
| 17516 | 66 | |
| 67 | (* reification of the equation *) | |
| 22950 | 68 | val TFree (_, cr_sort) = @{typ "'a :: {comm_ring, recpower}"};
 | 
| 69 | ||
| 70 | fun reif_eq thy (eq as Const("op =", Type("fun", [T, _])) $ lhs $ rhs) =
 | |
| 71 | if Sign.of_sort thy (T, cr_sort) then | |
| 72 | let | |
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
26939diff
changeset | 73 | val fs = OldTerm.term_frees eq; | 
| 22950 | 74 | val cvs = cterm_of thy (HOLogic.mk_list T fs); | 
| 75 | val clhs = cterm_of thy (reif_polex T fs lhs); | |
| 76 | val crhs = cterm_of thy (reif_polex T fs rhs); | |
| 77 | val ca = ctyp_of thy T; | |
| 78 | in (ca, cvs, clhs, crhs) end | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
24996diff
changeset | 79 |       else raise CRing ("reif_eq: not an equation over " ^ Syntax.string_of_sort_global thy cr_sort)
 | 
| 20623 | 80 | | reif_eq _ _ = raise CRing "reif_eq: not an equation"; | 
| 17516 | 81 | |
| 22950 | 82 | (* The cring tactic *) | 
| 17516 | 83 | (* Attention: You have to make sure that no t^0 is in the goal!! *) | 
| 84 | (* Use simply rewriting t^0 = 1 *) | |
| 20623 | 85 | val cring_simps = | 
| 22950 | 86 |   [@{thm mkPX_def}, @{thm mkPinj_def}, @{thm sub_def}, @{thm power_add},
 | 
| 87 |     @{thm even_def}, @{thm pow_if}, sym OF [@{thm power_add}]];
 | |
| 17516 | 88 | |
| 20623 | 89 | fun comm_ring_tac ctxt = SUBGOAL (fn (g, i) => | 
| 90 | let | |
| 22950 | 91 | val thy = ProofContext.theory_of ctxt; | 
| 92 | val cring_ss = Simplifier.local_simpset_of ctxt (*FIXME really the full simpset!?*) | |
| 93 | addsimps cring_simps; | |
| 20623 | 94 | val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) | 
| 95 | val norm_eq_th = | |
| 22950 | 96 |       simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq})
 | 
| 20623 | 97 | in | 
| 98 | cut_rules_tac [norm_eq_th] i | |
| 99 | THEN (simp_tac cring_ss i) | |
| 100 | THEN (simp_tac cring_ss i) | |
| 101 | end); | |
| 102 | ||
| 103 | val comm_ring_meth = | |
| 30510 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 wenzelm parents: 
29265diff
changeset | 104 | Method.ctxt_args (SIMPLE_METHOD' o comm_ring_tac); | 
| 17516 | 105 | |
| 106 | val setup = | |
| 20623 | 107 |   Method.add_method ("comm_ring", comm_ring_meth,
 | 
| 108 | "reflective decision procedure for equalities over commutative rings") #> | |
| 109 |   Method.add_method ("algebra", comm_ring_meth,
 | |
| 110 | "method for proving algebraic properties (same as comm_ring)"); | |
| 17516 | 111 | |
| 112 | end; |