author | haftmann |
Tue, 17 Aug 2010 14:19:11 +0200 | |
changeset 38458 | 2c46f628e6b7 |
parent 37744 | 3daaf23b9ab4 |
child 38864 | 4abe644fcea5 |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Decision_Procs/commutative_ring_tac.ML |
2 |
Author: Amine Chaieb |
|
17516 | 3 |
|
4 |
Tactic for solving equalities over commutative rings. |
|
5 |
*) |
|
6 |
||
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
7 |
signature COMMUTATIVE_RING_TAC = |
17516 | 8 |
sig |
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
9 |
val tac: Proof.context -> int -> tactic |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
10 |
val setup: theory -> theory |
17516 | 11 |
end |
12 |
||
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
13 |
structure Commutative_Ring_Tac: COMMUTATIVE_RING_TAC = |
17516 | 14 |
struct |
15 |
||
16 |
(* Zero and One of the commutative ring *) |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
17 |
fun cring_zero T = Const (@{const_name Groups.zero}, T); |
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
18 |
fun cring_one T = Const (@{const_name Groups.one}, T); |
17516 | 19 |
|
20 |
(* reification functions *) |
|
21 |
(* add two polynom expressions *) |
|
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
22 |
fun polT t = Type (@{type_name Commutative_Ring.pol}, [t]); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
23 |
fun polexT t = Type (@{type_name Commutative_Ring.polex}, [t]); |
17516 | 24 |
|
20623 | 25 |
(* pol *) |
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
26 |
fun pol_Pc t = Const (@{const_name Commutative_Ring.pol.Pc}, t --> polT t); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
27 |
fun pol_Pinj t = Const (@{const_name Commutative_Ring.pol.Pinj}, HOLogic.natT --> polT t --> polT t); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
28 |
fun pol_PX t = Const (@{const_name Commutative_Ring.pol.PX}, polT t --> HOLogic.natT --> polT t --> polT t); |
17516 | 29 |
|
30 |
(* polex *) |
|
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
31 |
fun polex_add t = Const (@{const_name Commutative_Ring.polex.Add}, polexT t --> polexT t --> polexT t); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
32 |
fun polex_sub t = Const (@{const_name Commutative_Ring.polex.Sub}, polexT t --> polexT t --> polexT t); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
33 |
fun polex_mul t = Const (@{const_name Commutative_Ring.polex.Mul}, polexT t --> polexT t --> polexT t); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
34 |
fun polex_neg t = Const (@{const_name Commutative_Ring.polex.Neg}, polexT t --> polexT t); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
35 |
fun polex_pol t = Const (@{const_name Commutative_Ring.polex.Pol}, polT t --> polexT t); |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
36 |
fun polex_pow t = Const (@{const_name Commutative_Ring.polex.Pow}, polexT t --> HOLogic.natT --> polexT t); |
17516 | 37 |
|
38 |
(* reification of polynoms : primitive cring expressions *) |
|
22950 | 39 |
fun reif_pol T vs (t as Free _) = |
40 |
let |
|
41 |
val one = @{term "1::nat"}; |
|
31986 | 42 |
val i = find_index (fn t' => t' = t) vs |
22950 | 43 |
in if i = 0 |
44 |
then pol_PX T $ (pol_Pc T $ cring_one T) |
|
45 |
$ one $ (pol_Pc T $ cring_zero T) |
|
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23261
diff
changeset
|
46 |
else pol_Pinj T $ HOLogic.mk_nat i |
22950 | 47 |
$ (pol_PX T $ (pol_Pc T $ cring_one T) |
48 |
$ one $ (pol_Pc T $ cring_zero T)) |
|
17516 | 49 |
end |
22950 | 50 |
| reif_pol T vs t = pol_Pc T $ t; |
17516 | 51 |
|
52 |
(* reification of polynom expressions *) |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
53 |
fun reif_polex T vs (Const (@{const_name Groups.plus}, _) $ a $ b) = |
22950 | 54 |
polex_add T $ reif_polex T vs a $ reif_polex T vs b |
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
55 |
| reif_polex T vs (Const (@{const_name Groups.minus}, _) $ a $ b) = |
22950 | 56 |
polex_sub T $ reif_polex T vs a $ reif_polex T vs b |
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
57 |
| reif_polex T vs (Const (@{const_name Groups.times}, _) $ a $ b) = |
22950 | 58 |
polex_mul T $ reif_polex T vs a $ reif_polex T vs b |
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
59 |
| reif_polex T vs (Const (@{const_name Groups.uminus}, _) $ a) = |
22950 | 60 |
polex_neg T $ reif_polex T vs a |
24996 | 61 |
| reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) = |
22950 | 62 |
polex_pow T $ reif_polex T vs a $ n |
63 |
| reif_polex T vs t = polex_pol T $ reif_pol T vs t; |
|
17516 | 64 |
|
65 |
(* reification of the equation *) |
|
31021 | 66 |
val cr_sort = @{sort "comm_ring_1"}; |
22950 | 67 |
|
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
68 |
fun reif_eq thy (eq as Const(@{const_name "op ="}, Type("fun", [T, _])) $ lhs $ rhs) = |
22950 | 69 |
if Sign.of_sort thy (T, cr_sort) then |
70 |
let |
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
26939
diff
changeset
|
71 |
val fs = OldTerm.term_frees eq; |
22950 | 72 |
val cvs = cterm_of thy (HOLogic.mk_list T fs); |
73 |
val clhs = cterm_of thy (reif_polex T fs lhs); |
|
74 |
val crhs = cterm_of thy (reif_polex T fs rhs); |
|
75 |
val ca = ctyp_of thy T; |
|
76 |
in (ca, cvs, clhs, crhs) end |
|
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
77 |
else error ("reif_eq: not an equation over " ^ Syntax.string_of_sort_global thy cr_sort) |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
78 |
| reif_eq _ _ = error "reif_eq: not an equation"; |
17516 | 79 |
|
22950 | 80 |
(* The cring tactic *) |
17516 | 81 |
(* Attention: You have to make sure that no t^0 is in the goal!! *) |
82 |
(* Use simply rewriting t^0 = 1 *) |
|
20623 | 83 |
val cring_simps = |
22950 | 84 |
[@{thm mkPX_def}, @{thm mkPinj_def}, @{thm sub_def}, @{thm power_add}, |
85 |
@{thm even_def}, @{thm pow_if}, sym OF [@{thm power_add}]]; |
|
17516 | 86 |
|
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
87 |
fun tac ctxt = SUBGOAL (fn (g, i) => |
20623 | 88 |
let |
22950 | 89 |
val thy = ProofContext.theory_of ctxt; |
32149
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents:
31986
diff
changeset
|
90 |
val cring_ss = Simplifier.simpset_of ctxt (*FIXME really the full simpset!?*) |
22950 | 91 |
addsimps cring_simps; |
20623 | 92 |
val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) |
93 |
val norm_eq_th = |
|
22950 | 94 |
simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq}) |
20623 | 95 |
in |
96 |
cut_rules_tac [norm_eq_th] i |
|
97 |
THEN (simp_tac cring_ss i) |
|
98 |
THEN (simp_tac cring_ss i) |
|
99 |
end); |
|
100 |
||
17516 | 101 |
val setup = |
33356
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
102 |
Method.setup @{binding comm_ring} (Scan.succeed (SIMPLE_METHOD' o tac)) |
9157d0f9f00e
moved Commutative_Ring into session Decision_Procs
haftmann
parents:
32149
diff
changeset
|
103 |
"reflective decision procedure for equalities over commutative rings"; |
17516 | 104 |
|
105 |
end; |