author | haftmann |
Mon, 01 Jun 2015 18:59:21 +0200 | |
changeset 60352 | d46de31a50c4 |
parent 59582 | 0fbed69ff081 |
child 60642 | 48dd1cefb4ae |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Tools/semiring_normalizer.ML |
23252 | 2 |
Author: Amine Chaieb, TU Muenchen |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
3 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
4 |
Normalization of expressions in semirings. |
23252 | 5 |
*) |
6 |
||
36753
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents:
36751
diff
changeset
|
7 |
signature SEMIRING_NORMALIZER = |
23252 | 8 |
sig |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
9 |
type entry |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
10 |
val match: Proof.context -> cterm -> entry option |
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
11 |
val the_semiring: Proof.context -> thm -> cterm list * thm list |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
12 |
val the_ring: Proof.context -> thm -> cterm list * thm list |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
13 |
val the_field: Proof.context -> thm -> cterm list * thm list |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
14 |
val the_idom: Proof.context -> thm -> thm list |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
15 |
val the_ideal: Proof.context -> thm -> thm list |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
16 |
val declare: thm -> {semiring: cterm list * thm list, ring: cterm list * thm list, |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
17 |
field: cterm list * thm list, idom: thm list, ideal: thm list} -> declaration |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
18 |
|
36711 | 19 |
val semiring_normalize_conv: Proof.context -> conv |
20 |
val semiring_normalize_ord_conv: Proof.context -> (cterm -> cterm -> bool) -> conv |
|
21 |
val semiring_normalize_wrapper: Proof.context -> entry -> conv |
|
22 |
val semiring_normalize_ord_wrapper: Proof.context -> entry |
|
23 |
-> (cterm -> cterm -> bool) -> conv |
|
24 |
val semiring_normalizers_conv: cterm list -> cterm list * thm list |
|
25 |
-> cterm list * thm list -> cterm list * thm list -> |
|
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
26 |
(cterm -> bool) * conv * conv * conv -> (cterm -> cterm -> bool) -> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
27 |
{add: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
28 |
mul: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
29 |
neg: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
30 |
main: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
31 |
pow: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
32 |
sub: Proof.context -> conv} |
36711 | 33 |
val semiring_normalizers_ord_wrapper: Proof.context -> entry -> |
34 |
(cterm -> cterm -> bool) -> |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
35 |
{add: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
36 |
mul: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
37 |
neg: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
38 |
main: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
39 |
pow: Proof.context -> conv, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
40 |
sub: Proof.context -> conv} |
23252 | 41 |
end |
42 |
||
36753
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents:
36751
diff
changeset
|
43 |
structure Semiring_Normalizer: SEMIRING_NORMALIZER = |
23252 | 44 |
struct |
23559 | 45 |
|
36708 | 46 |
(** data **) |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
47 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
48 |
type entry = |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
49 |
{vars: cterm list, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
50 |
semiring: cterm list * thm list, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
51 |
ring: cterm list * thm list, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
52 |
field: cterm list * thm list, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
53 |
idom: thm list, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
54 |
ideal: thm list} * |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
55 |
{is_const: cterm -> bool, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
56 |
dest_const: cterm -> Rat.rat, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
57 |
mk_const: ctyp -> Rat.rat -> cterm, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
58 |
conv: Proof.context -> cterm -> thm}; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
59 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
60 |
structure Data = Generic_Data |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
61 |
( |
36705 | 62 |
type T = (thm * entry) list; |
63 |
val empty = []; |
|
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
64 |
val extend = I; |
36771 | 65 |
fun merge data = AList.merge Thm.eq_thm (K true) data; |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
66 |
); |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
67 |
|
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
68 |
fun the_rules ctxt = fst o the o AList.lookup Thm.eq_thm (Data.get (Context.Proof ctxt)) |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
69 |
|
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
70 |
val the_semiring = #semiring oo the_rules |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
71 |
val the_ring = #ring oo the_rules |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
72 |
val the_field = #field oo the_rules |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
73 |
val the_idom = #idom oo the_rules |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
74 |
val the_ideal = #ideal oo the_rules |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
75 |
|
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
76 |
fun match ctxt tm = |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
77 |
let |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
78 |
fun match_inst |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
79 |
({vars, semiring = (sr_ops, sr_rules), |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
80 |
ring = (r_ops, r_rules), field = (f_ops, f_rules), idom, ideal}, |
59321 | 81 |
fns) pat = |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
82 |
let |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
83 |
fun h instT = |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
84 |
let |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
85 |
val substT = Thm.instantiate (instT, []); |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
86 |
val substT_cterm = Drule.cterm_rule substT; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
87 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
88 |
val vars' = map substT_cterm vars; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
89 |
val semiring' = (map substT_cterm sr_ops, map substT sr_rules); |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
90 |
val ring' = (map substT_cterm r_ops, map substT r_rules); |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
91 |
val field' = (map substT_cterm f_ops, map substT f_rules); |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
92 |
val idom' = map substT idom; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
93 |
val ideal' = map substT ideal; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
94 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
95 |
val result = ({vars = vars', semiring = semiring', |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
96 |
ring = ring', field = field', idom = idom', ideal = ideal'}, fns); |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
97 |
in SOME result end |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
98 |
in (case try Thm.match (pat, tm) of |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
99 |
NONE => NONE |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
100 |
| SOME (instT, _) => h instT) |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
101 |
end; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
102 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
103 |
fun match_struct (_, |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
104 |
entry as ({semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), ...}, _): entry) = |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
105 |
get_first (match_inst entry) (sr_ops @ r_ops @ f_ops); |
59549
6e685f9c9aa5
more direct expression of syntactic function records
haftmann
parents:
59548
diff
changeset
|
106 |
in get_first match_struct (Data.get (Context.Proof ctxt)) end; |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
107 |
|
59539 | 108 |
|
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
109 |
(* extra-logical functions *) |
59539 | 110 |
|
111 |
val semiring_norm_ss = |
|
112 |
simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms semiring_norm}); |
|
113 |
||
59549
6e685f9c9aa5
more direct expression of syntactic function records
haftmann
parents:
59548
diff
changeset
|
114 |
val semiring_funs = |
59548 | 115 |
{is_const = can HOLogic.dest_number o Thm.term_of, |
116 |
dest_const = (fn ct => |
|
59539 | 117 |
Rat.rat_of_int (snd |
118 |
(HOLogic.dest_number (Thm.term_of ct) |
|
59540 | 119 |
handle TERM _ => error "ring_dest_const"))), |
59548 | 120 |
mk_const = (fn cT => fn x => Numeral.mk_cnumber cT |
59540 | 121 |
(case Rat.quotient_of_rat x of (i, 1) => i | _ => error "int_of_rat: bad int")), |
59548 | 122 |
conv = (fn ctxt => |
59539 | 123 |
Simplifier.rewrite (put_simpset semiring_norm_ss ctxt) |
59540 | 124 |
then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms numeral_1_eq_1}))}; |
59539 | 125 |
|
59549
6e685f9c9aa5
more direct expression of syntactic function records
haftmann
parents:
59548
diff
changeset
|
126 |
val field_funs = |
59539 | 127 |
let |
128 |
fun numeral_is_const ct = |
|
59582 | 129 |
case Thm.term_of ct of |
60352
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
haftmann
parents:
59582
diff
changeset
|
130 |
Const (@{const_name Rings.divide},_) $ a $ b => |
59539 | 131 |
can HOLogic.dest_number a andalso can HOLogic.dest_number b |
132 |
| Const (@{const_name Fields.inverse},_)$t => can HOLogic.dest_number t |
|
133 |
| t => can HOLogic.dest_number t |
|
59582 | 134 |
fun dest_const ct = ((case Thm.term_of ct of |
60352
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
haftmann
parents:
59582
diff
changeset
|
135 |
Const (@{const_name Rings.divide},_) $ a $ b=> |
59539 | 136 |
Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) |
137 |
| Const (@{const_name Fields.inverse},_)$t => |
|
138 |
Rat.inv (Rat.rat_of_int (snd (HOLogic.dest_number t))) |
|
139 |
| t => Rat.rat_of_int (snd (HOLogic.dest_number t))) |
|
140 |
handle TERM _ => error "ring_dest_const") |
|
59540 | 141 |
fun mk_const cT x = |
59539 | 142 |
let val (a, b) = Rat.quotient_of_rat x |
143 |
in if b = 1 then Numeral.mk_cnumber cT a |
|
144 |
else Thm.apply |
|
145 |
(Thm.apply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"}) |
|
146 |
(Numeral.mk_cnumber cT a)) |
|
147 |
(Numeral.mk_cnumber cT b) |
|
148 |
end |
|
59549
6e685f9c9aa5
more direct expression of syntactic function records
haftmann
parents:
59548
diff
changeset
|
149 |
in |
59548 | 150 |
{is_const = numeral_is_const, |
151 |
dest_const = dest_const, |
|
152 |
mk_const = mk_const, |
|
153 |
conv = Numeral_Simprocs.field_comp_conv} |
|
59539 | 154 |
end; |
155 |
||
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
156 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
157 |
(* logical content *) |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
158 |
|
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
159 |
val semiringN = "semiring"; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
160 |
val ringN = "ring"; |
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
161 |
val fieldN = "field"; |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
162 |
val idomN = "idom"; |
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
163 |
|
59562 | 164 |
fun declare raw_key |
165 |
{semiring = raw_semiring, ring = raw_ring, field = raw_field, idom = raw_idom, ideal = raw_ideal} |
|
166 |
phi context = |
|
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
167 |
let |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
168 |
val ctxt = Context.proof_of context; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
169 |
val key = Morphism.thm phi raw_key; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
170 |
fun morphism_ops_rules (ops, rules) = (map (Morphism.cterm phi) ops, Morphism.fact phi rules); |
59562 | 171 |
val (sr_ops, sr_rules) = morphism_ops_rules raw_semiring; |
172 |
val (r_ops, r_rules) = morphism_ops_rules raw_ring; |
|
173 |
val (f_ops, f_rules) = morphism_ops_rules raw_field; |
|
174 |
val idom = Morphism.fact phi raw_idom; |
|
175 |
val ideal = Morphism.fact phi raw_ideal; |
|
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
176 |
|
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
177 |
fun check kind name xs n = |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
178 |
null xs orelse length xs = n orelse |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
179 |
error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name); |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
180 |
val check_ops = check "operations"; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
181 |
val check_rules = check "rules"; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
182 |
val _ = |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
183 |
check_ops semiringN sr_ops 5 andalso |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
184 |
check_rules semiringN sr_rules 36 andalso |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
185 |
check_ops ringN r_ops 2 andalso |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
186 |
check_rules ringN r_rules 2 andalso |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
187 |
check_ops fieldN f_ops 2 andalso |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
188 |
check_rules fieldN f_rules 2 andalso |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
189 |
check_rules idomN idom 2; |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
190 |
|
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
191 |
val mk_meta = Local_Defs.meta_rewrite_rule ctxt; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
192 |
val sr_rules' = map mk_meta sr_rules; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
193 |
val r_rules' = map mk_meta r_rules; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
194 |
val f_rules' = map mk_meta f_rules; |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
195 |
|
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
196 |
fun rule i = nth sr_rules' (i - 1); |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
197 |
|
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
198 |
val (cx, cy) = Thm.dest_binop (hd sr_ops); |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
199 |
val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
200 |
val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
201 |
val ((clx, crx), (cly, cry)) = |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
202 |
rule 13 |> Thm.rhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
203 |
val ((ca, cb), (cc, cd)) = |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
204 |
rule 20 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
205 |
val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
206 |
val (cp, cq) = rule 26 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_arg; |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
207 |
|
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
208 |
val vars = [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry]; |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
209 |
|
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
210 |
val semiring = (sr_ops, sr_rules'); |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
211 |
val ring = (r_ops, r_rules'); |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
212 |
val field = (f_ops, f_rules'); |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
213 |
val ideal' = map (Thm.symmetric o mk_meta) ideal |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
214 |
|
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
215 |
in |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
216 |
context |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
217 |
|> Data.map (AList.update Thm.eq_thm (key, |
59539 | 218 |
({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal'}, |
59553
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
219 |
(if null f_ops then semiring_funs else field_funs)))) |
e87974cd9b86
explicit declaration allows cumulative declaration
haftmann
parents:
59551
diff
changeset
|
220 |
end |
36720 | 221 |
|
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
222 |
|
36710 | 223 |
(** auxiliary **) |
25253 | 224 |
|
225 |
fun is_comb ct = |
|
226 |
(case Thm.term_of ct of |
|
227 |
_ $ _ => true |
|
228 |
| _ => false); |
|
229 |
||
230 |
val concl = Thm.cprop_of #> Thm.dest_arg; |
|
231 |
||
232 |
fun is_binop ct ct' = |
|
233 |
(case Thm.term_of ct' of |
|
59582 | 234 |
c $ _ $ _ => Thm.term_of ct aconv c |
25253 | 235 |
| _ => false); |
236 |
||
237 |
fun dest_binop ct ct' = |
|
238 |
if is_binop ct ct' then Thm.dest_binop ct' |
|
239 |
else raise CTERM ("dest_binop: bad binop", [ct, ct']) |
|
240 |
||
241 |
fun inst_thm inst = Thm.instantiate ([], inst); |
|
242 |
||
59582 | 243 |
val dest_number = Thm.term_of #> HOLogic.dest_number #> snd; |
59538 | 244 |
val is_number = can dest_number; |
23252 | 245 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
246 |
fun numeral01_conv ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
247 |
Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_1_eq_1}]); |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
248 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
249 |
fun zero1_numeral_conv ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
250 |
Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_1_eq_1} RS sym]); |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
251 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
252 |
fun zerone_conv ctxt cv = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
253 |
zero1_numeral_conv ctxt then_conv cv then_conv numeral01_conv ctxt; |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
254 |
|
59547 | 255 |
val nat_add_ss = simpset_of |
256 |
(put_simpset HOL_basic_ss @{context} |
|
257 |
addsimps @{thms arith_simps} @ @{thms diff_nat_numeral} @ @{thms rel_simps} |
|
258 |
@ @{thms if_False if_True Nat.add_0 add_Suc add_numeral_left Suc_eq_plus1} |
|
259 |
@ map (fn th => th RS sym) @{thms numerals}); |
|
53078
cc06f17d8057
more static simpsets, which also avoids spurious warnings due to duplicate rules provided here;
wenzelm
parents:
51717
diff
changeset
|
260 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
261 |
fun nat_add_conv ctxt = |
53078
cc06f17d8057
more static simpsets, which also avoids spurious warnings due to duplicate rules provided here;
wenzelm
parents:
51717
diff
changeset
|
262 |
zerone_conv ctxt (Simplifier.rewrite (put_simpset nat_add_ss ctxt)); |
23252 | 263 |
|
264 |
val zeron_tm = @{cterm "0::nat"}; |
|
265 |
val onen_tm = @{cterm "1::nat"}; |
|
266 |
val true_tm = @{cterm "True"}; |
|
267 |
||
268 |
||
36710 | 269 |
(** normalizing conversions **) |
270 |
||
271 |
(* core conversion *) |
|
272 |
||
30866 | 273 |
fun semiring_normalizers_conv vars (sr_ops, sr_rules) (r_ops, r_rules) (f_ops, f_rules) |
23252 | 274 |
(is_semiring_constant, semiring_add_conv, semiring_mul_conv, semiring_pow_conv) = |
275 |
let |
|
276 |
||
277 |
val [pthm_02, pthm_03, pthm_04, pthm_05, pthm_07, pthm_08, |
|
278 |
pthm_09, pthm_10, pthm_11, pthm_12, pthm_13, pthm_14, pthm_15, pthm_16, |
|
279 |
pthm_17, pthm_18, pthm_19, pthm_21, pthm_22, pthm_23, pthm_24, |
|
280 |
pthm_25, pthm_26, pthm_27, pthm_28, pthm_29, pthm_30, pthm_31, pthm_32, |
|
59550 | 281 |
pthm_33, pthm_34, pthm_35, pthm_36, pthm_37, pthm_38, _] = sr_rules; |
23252 | 282 |
|
283 |
val [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry] = vars; |
|
284 |
val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; |
|
285 |
val [add_tm, mul_tm, pow_tm] = map (Thm.dest_fun o Thm.dest_fun) [add_pat, mul_pat, pow_pat]; |
|
286 |
||
287 |
val dest_add = dest_binop add_tm |
|
288 |
val dest_mul = dest_binop mul_tm |
|
289 |
fun dest_pow tm = |
|
290 |
let val (l,r) = dest_binop pow_tm tm |
|
59538 | 291 |
in if is_number r then (l,r) else raise CTERM ("dest_pow",[tm]) |
23252 | 292 |
end; |
293 |
val is_add = is_binop add_tm |
|
294 |
val is_mul = is_binop mul_tm |
|
295 |
||
59321 | 296 |
val (neg_mul, sub_add, sub_tm, neg_tm, dest_sub, cx', cy') = |
23252 | 297 |
(case (r_ops, r_rules) of |
30866 | 298 |
([sub_pat, neg_pat], [neg_mul, sub_add]) => |
23252 | 299 |
let |
300 |
val sub_tm = Thm.dest_fun (Thm.dest_fun sub_pat) |
|
301 |
val neg_tm = Thm.dest_fun neg_pat |
|
302 |
val dest_sub = dest_binop sub_tm |
|
59321 | 303 |
in (neg_mul,sub_add,sub_tm,neg_tm,dest_sub, neg_mul |> concl |> Thm.dest_arg, |
23252 | 304 |
sub_add |> concl |> Thm.dest_arg |> Thm.dest_arg) |
30866 | 305 |
end |
59321 | 306 |
| _ => (TrueI, TrueI, true_tm, true_tm, (fn t => (t,t)), true_tm, true_tm)); |
30866 | 307 |
|
59321 | 308 |
val (divide_inverse, divide_tm, inverse_tm) = |
30866 | 309 |
(case (f_ops, f_rules) of |
59321 | 310 |
([divide_pat, inverse_pat], [div_inv, _]) => |
30866 | 311 |
let val div_tm = funpow 2 Thm.dest_fun divide_pat |
312 |
val inv_tm = Thm.dest_fun inverse_pat |
|
59321 | 313 |
in (div_inv, div_tm, inv_tm) |
30866 | 314 |
end |
59321 | 315 |
| _ => (TrueI, true_tm, true_tm)); |
30866 | 316 |
|
23252 | 317 |
in fn variable_order => |
318 |
let |
|
319 |
||
320 |
(* Conversion for "x^n * x^m", with either x^n = x and/or x^m = x possible. *) |
|
321 |
(* Also deals with "const * const", but both terms must involve powers of *) |
|
322 |
(* the same variable, or both be constants, or behaviour may be incorrect. *) |
|
323 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
324 |
fun powvar_mul_conv ctxt tm = |
23252 | 325 |
let |
326 |
val (l,r) = dest_mul tm |
|
327 |
in if is_semiring_constant l andalso is_semiring_constant r |
|
328 |
then semiring_mul_conv tm |
|
329 |
else |
|
330 |
((let |
|
331 |
val (lx,ln) = dest_pow l |
|
332 |
in |
|
59321 | 333 |
((let val (_, rn) = dest_pow r |
23252 | 334 |
val th1 = inst_thm [(cx,lx),(cp,ln),(cq,rn)] pthm_29 |
335 |
val (tm1,tm2) = Thm.dest_comb(concl th1) in |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
336 |
Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) |
23252 | 337 |
handle CTERM _ => |
338 |
(let val th1 = inst_thm [(cx,lx),(cq,ln)] pthm_31 |
|
339 |
val (tm1,tm2) = Thm.dest_comb(concl th1) in |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
340 |
Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end)) end) |
23252 | 341 |
handle CTERM _ => |
342 |
((let val (rx,rn) = dest_pow r |
|
343 |
val th1 = inst_thm [(cx,rx),(cq,rn)] pthm_30 |
|
344 |
val (tm1,tm2) = Thm.dest_comb(concl th1) in |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
345 |
Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) |
23252 | 346 |
handle CTERM _ => inst_thm [(cx,l)] pthm_32 |
347 |
||
348 |
)) |
|
349 |
end; |
|
350 |
||
351 |
(* Remove "1 * m" from a monomial, and just leave m. *) |
|
352 |
||
353 |
fun monomial_deone th = |
|
354 |
(let val (l,r) = dest_mul(concl th) in |
|
355 |
if l aconvc one_tm |
|
36945 | 356 |
then Thm.transitive th (inst_thm [(ca,r)] pthm_13) else th end) |
23252 | 357 |
handle CTERM _ => th; |
358 |
||
359 |
(* Conversion for "(monomial)^n", where n is a numeral. *) |
|
360 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
361 |
fun monomial_pow_conv ctxt = |
23252 | 362 |
let |
363 |
fun monomial_pow tm bod ntm = |
|
364 |
if not(is_comb bod) |
|
36945 | 365 |
then Thm.reflexive tm |
23252 | 366 |
else |
367 |
if is_semiring_constant bod |
|
368 |
then semiring_pow_conv tm |
|
369 |
else |
|
370 |
let |
|
371 |
val (lopr,r) = Thm.dest_comb bod |
|
372 |
in if not(is_comb lopr) |
|
36945 | 373 |
then Thm.reflexive tm |
23252 | 374 |
else |
375 |
let |
|
376 |
val (opr,l) = Thm.dest_comb lopr |
|
377 |
in |
|
59538 | 378 |
if opr aconvc pow_tm andalso is_number r |
23252 | 379 |
then |
380 |
let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34 |
|
381 |
val (l,r) = Thm.dest_comb(concl th1) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
382 |
in Thm.transitive th1 (Drule.arg_cong_rule l (nat_add_conv ctxt r)) |
23252 | 383 |
end |
384 |
else |
|
385 |
if opr aconvc mul_tm |
|
386 |
then |
|
387 |
let |
|
388 |
val th1 = inst_thm [(cx,l),(cy,r),(cq,ntm)] pthm_33 |
|
389 |
val (xy,z) = Thm.dest_comb(concl th1) |
|
390 |
val (x,y) = Thm.dest_comb xy |
|
391 |
val thl = monomial_pow y l ntm |
|
392 |
val thr = monomial_pow z r ntm |
|
36945 | 393 |
in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule x thl) thr) |
23252 | 394 |
end |
36945 | 395 |
else Thm.reflexive tm |
23252 | 396 |
end |
397 |
end |
|
398 |
in fn tm => |
|
399 |
let |
|
400 |
val (lopr,r) = Thm.dest_comb tm |
|
401 |
val (opr,l) = Thm.dest_comb lopr |
|
59538 | 402 |
in if not (opr aconvc pow_tm) orelse not(is_number r) |
23252 | 403 |
then raise CTERM ("monomial_pow_conv", [tm]) |
404 |
else if r aconvc zeron_tm |
|
405 |
then inst_thm [(cx,l)] pthm_35 |
|
406 |
else if r aconvc onen_tm |
|
407 |
then inst_thm [(cx,l)] pthm_36 |
|
408 |
else monomial_deone(monomial_pow tm l r) |
|
409 |
end |
|
410 |
end; |
|
411 |
||
412 |
(* Multiplication of canonical monomials. *) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
413 |
fun monomial_mul_conv ctxt = |
23252 | 414 |
let |
415 |
fun powvar tm = |
|
416 |
if is_semiring_constant tm then one_tm |
|
417 |
else |
|
418 |
((let val (lopr,r) = Thm.dest_comb tm |
|
419 |
val (opr,l) = Thm.dest_comb lopr |
|
59538 | 420 |
in if opr aconvc pow_tm andalso is_number r then l |
23252 | 421 |
else raise CTERM ("monomial_mul_conv",[tm]) end) |
422 |
handle CTERM _ => tm) (* FIXME !? *) |
|
423 |
fun vorder x y = |
|
424 |
if x aconvc y then 0 |
|
425 |
else |
|
426 |
if x aconvc one_tm then ~1 |
|
427 |
else if y aconvc one_tm then 1 |
|
428 |
else if variable_order x y then ~1 else 1 |
|
429 |
fun monomial_mul tm l r = |
|
430 |
((let val (lx,ly) = dest_mul l val vl = powvar lx |
|
431 |
in |
|
432 |
((let |
|
433 |
val (rx,ry) = dest_mul r |
|
434 |
val vr = powvar rx |
|
435 |
val ord = vorder vl vr |
|
436 |
in |
|
437 |
if ord = 0 |
|
438 |
then |
|
439 |
let |
|
440 |
val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] pthm_15 |
|
441 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
442 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
443 |
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 |
36945 | 444 |
val th3 = Thm.transitive th1 th2 |
23252 | 445 |
val (tm5,tm6) = Thm.dest_comb(concl th3) |
446 |
val (tm7,tm8) = Thm.dest_comb tm6 |
|
447 |
val th4 = monomial_mul tm6 (Thm.dest_arg tm7) tm8 |
|
36945 | 448 |
in Thm.transitive th3 (Drule.arg_cong_rule tm5 th4) |
23252 | 449 |
end |
450 |
else |
|
451 |
let val th0 = if ord < 0 then pthm_16 else pthm_17 |
|
452 |
val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] th0 |
|
453 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
454 |
val (tm3,tm4) = Thm.dest_comb tm2 |
|
36945 | 455 |
in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) |
23252 | 456 |
end |
457 |
end) |
|
458 |
handle CTERM _ => |
|
459 |
(let val vr = powvar r val ord = vorder vl vr |
|
460 |
in |
|
461 |
if ord = 0 then |
|
462 |
let |
|
463 |
val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_18 |
|
464 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
465 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
466 |
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 |
36945 | 467 |
in Thm.transitive th1 th2 |
23252 | 468 |
end |
469 |
else |
|
470 |
if ord < 0 then |
|
471 |
let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_19 |
|
472 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
473 |
val (tm3,tm4) = Thm.dest_comb tm2 |
|
36945 | 474 |
in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) |
23252 | 475 |
end |
476 |
else inst_thm [(ca,l),(cb,r)] pthm_09 |
|
477 |
end)) end) |
|
478 |
handle CTERM _ => |
|
479 |
(let val vl = powvar l in |
|
480 |
((let |
|
481 |
val (rx,ry) = dest_mul r |
|
482 |
val vr = powvar rx |
|
483 |
val ord = vorder vl vr |
|
484 |
in if ord = 0 then |
|
485 |
let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_21 |
|
486 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
487 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
488 |
in Thm.transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2) |
23252 | 489 |
end |
490 |
else if ord > 0 then |
|
491 |
let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_22 |
|
492 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
493 |
val (tm3,tm4) = Thm.dest_comb tm2 |
|
36945 | 494 |
in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) |
23252 | 495 |
end |
36945 | 496 |
else Thm.reflexive tm |
23252 | 497 |
end) |
498 |
handle CTERM _ => |
|
499 |
(let val vr = powvar r |
|
500 |
val ord = vorder vl vr |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
501 |
in if ord = 0 then powvar_mul_conv ctxt tm |
23252 | 502 |
else if ord > 0 then inst_thm [(ca,l),(cb,r)] pthm_09 |
36945 | 503 |
else Thm.reflexive tm |
23252 | 504 |
end)) end)) |
505 |
in fn tm => let val (l,r) = dest_mul tm in monomial_deone(monomial_mul tm l r) |
|
506 |
end |
|
507 |
end; |
|
508 |
(* Multiplication by monomial of a polynomial. *) |
|
509 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
510 |
fun polynomial_monomial_mul_conv ctxt = |
23252 | 511 |
let |
512 |
fun pmm_conv tm = |
|
513 |
let val (l,r) = dest_mul tm |
|
514 |
in |
|
515 |
((let val (y,z) = dest_add r |
|
516 |
val th1 = inst_thm [(cx,l),(cy,y),(cz,z)] pthm_37 |
|
517 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
518 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
519 |
val th2 = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
520 |
Thm.combination (Drule.arg_cong_rule tm3 (monomial_mul_conv ctxt tm4)) (pmm_conv tm2) |
36945 | 521 |
in Thm.transitive th1 th2 |
23252 | 522 |
end) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
523 |
handle CTERM _ => monomial_mul_conv ctxt tm) |
23252 | 524 |
end |
525 |
in pmm_conv |
|
526 |
end; |
|
527 |
||
528 |
(* Addition of two monomials identical except for constant multiples. *) |
|
529 |
||
530 |
fun monomial_add_conv tm = |
|
531 |
let val (l,r) = dest_add tm |
|
532 |
in if is_semiring_constant l andalso is_semiring_constant r |
|
533 |
then semiring_add_conv tm |
|
534 |
else |
|
535 |
let val th1 = |
|
536 |
if is_mul l andalso is_semiring_constant(Thm.dest_arg1 l) |
|
537 |
then if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then |
|
538 |
inst_thm [(ca,Thm.dest_arg1 l),(cm,Thm.dest_arg r), (cb,Thm.dest_arg1 r)] pthm_02 |
|
539 |
else inst_thm [(ca,Thm.dest_arg1 l),(cm,r)] pthm_03 |
|
540 |
else if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) |
|
541 |
then inst_thm [(cm,l),(ca,Thm.dest_arg1 r)] pthm_04 |
|
542 |
else inst_thm [(cm,r)] pthm_05 |
|
543 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
544 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
545 |
val th2 = Drule.arg_cong_rule tm3 (semiring_add_conv tm4) |
|
36945 | 546 |
val th3 = Thm.transitive th1 (Drule.fun_cong_rule th2 tm2) |
23252 | 547 |
val tm5 = concl th3 |
548 |
in |
|
549 |
if (Thm.dest_arg1 tm5) aconvc zero_tm |
|
36945 | 550 |
then Thm.transitive th3 (inst_thm [(ca,Thm.dest_arg tm5)] pthm_11) |
23252 | 551 |
else monomial_deone th3 |
552 |
end |
|
553 |
end; |
|
554 |
||
555 |
(* Ordering on monomials. *) |
|
556 |
||
557 |
fun striplist dest = |
|
558 |
let fun strip x acc = |
|
559 |
((let val (l,r) = dest x in |
|
560 |
strip l (strip r acc) end) |
|
561 |
handle CTERM _ => x::acc) (* FIXME !? *) |
|
562 |
in fn x => strip x [] |
|
563 |
end; |
|
564 |
||
565 |
||
566 |
fun powervars tm = |
|
567 |
let val ptms = striplist dest_mul tm |
|
568 |
in if is_semiring_constant (hd ptms) then tl ptms else ptms |
|
569 |
end; |
|
570 |
val num_0 = 0; |
|
571 |
val num_1 = 1; |
|
572 |
fun dest_varpow tm = |
|
59538 | 573 |
((let val (x,n) = dest_pow tm in (x,dest_number n) end) |
23252 | 574 |
handle CTERM _ => |
575 |
(tm,(if is_semiring_constant tm then num_0 else num_1))); |
|
576 |
||
577 |
val morder = |
|
578 |
let fun lexorder l1 l2 = |
|
579 |
case (l1,l2) of |
|
580 |
([],[]) => 0 |
|
59321 | 581 |
| (_ ,[]) => ~1 |
582 |
| ([], _) => 1 |
|
23252 | 583 |
| (((x1,n1)::vs1),((x2,n2)::vs2)) => |
584 |
if variable_order x1 x2 then 1 |
|
585 |
else if variable_order x2 x1 then ~1 |
|
586 |
else if n1 < n2 then ~1 |
|
587 |
else if n2 < n1 then 1 |
|
588 |
else lexorder vs1 vs2 |
|
589 |
in fn tm1 => fn tm2 => |
|
590 |
let val vdegs1 = map dest_varpow (powervars tm1) |
|
591 |
val vdegs2 = map dest_varpow (powervars tm2) |
|
33002 | 592 |
val deg1 = fold (Integer.add o snd) vdegs1 num_0 |
593 |
val deg2 = fold (Integer.add o snd) vdegs2 num_0 |
|
23252 | 594 |
in if deg1 < deg2 then ~1 else if deg1 > deg2 then 1 |
595 |
else lexorder vdegs1 vdegs2 |
|
596 |
end |
|
597 |
end; |
|
598 |
||
599 |
(* Addition of two polynomials. *) |
|
600 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
601 |
fun polynomial_add_conv ctxt = |
23252 | 602 |
let |
603 |
fun dezero_rule th = |
|
604 |
let |
|
605 |
val tm = concl th |
|
606 |
in |
|
607 |
if not(is_add tm) then th else |
|
608 |
let val (lopr,r) = Thm.dest_comb tm |
|
609 |
val l = Thm.dest_arg lopr |
|
610 |
in |
|
611 |
if l aconvc zero_tm |
|
36945 | 612 |
then Thm.transitive th (inst_thm [(ca,r)] pthm_07) else |
23252 | 613 |
if r aconvc zero_tm |
36945 | 614 |
then Thm.transitive th (inst_thm [(ca,l)] pthm_08) else th |
23252 | 615 |
end |
616 |
end |
|
617 |
fun padd tm = |
|
618 |
let |
|
619 |
val (l,r) = dest_add tm |
|
620 |
in |
|
621 |
if l aconvc zero_tm then inst_thm [(ca,r)] pthm_07 |
|
622 |
else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_08 |
|
623 |
else |
|
624 |
if is_add l |
|
625 |
then |
|
626 |
let val (a,b) = dest_add l |
|
627 |
in |
|
628 |
if is_add r then |
|
629 |
let val (c,d) = dest_add r |
|
630 |
val ord = morder a c |
|
631 |
in |
|
632 |
if ord = 0 then |
|
633 |
let val th1 = inst_thm [(ca,a),(cb,b),(cc,c),(cd,d)] pthm_23 |
|
634 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
635 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
636 |
val th2 = Drule.arg_cong_rule tm3 (monomial_add_conv tm4) |
|
36945 | 637 |
in dezero_rule (Thm.transitive th1 (Thm.combination th2 (padd tm2))) |
23252 | 638 |
end |
639 |
else (* ord <> 0*) |
|
640 |
let val th1 = |
|
641 |
if ord > 0 then inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 |
|
642 |
else inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 |
|
643 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
36945 | 644 |
in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) |
23252 | 645 |
end |
646 |
end |
|
647 |
else (* not (is_add r)*) |
|
648 |
let val ord = morder a r |
|
649 |
in |
|
650 |
if ord = 0 then |
|
651 |
let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_26 |
|
652 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
653 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
654 |
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 |
|
36945 | 655 |
in dezero_rule (Thm.transitive th1 th2) |
23252 | 656 |
end |
657 |
else (* ord <> 0*) |
|
658 |
if ord > 0 then |
|
659 |
let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 |
|
660 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
36945 | 661 |
in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) |
23252 | 662 |
end |
663 |
else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) |
|
664 |
end |
|
665 |
end |
|
666 |
else (* not (is_add l)*) |
|
667 |
if is_add r then |
|
668 |
let val (c,d) = dest_add r |
|
669 |
val ord = morder l c |
|
670 |
in |
|
671 |
if ord = 0 then |
|
672 |
let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_28 |
|
673 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
674 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
675 |
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 |
|
36945 | 676 |
in dezero_rule (Thm.transitive th1 th2) |
23252 | 677 |
end |
678 |
else |
|
36945 | 679 |
if ord > 0 then Thm.reflexive tm |
23252 | 680 |
else |
681 |
let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 |
|
682 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
36945 | 683 |
in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) |
23252 | 684 |
end |
685 |
end |
|
686 |
else |
|
687 |
let val ord = morder l r |
|
688 |
in |
|
689 |
if ord = 0 then monomial_add_conv tm |
|
36945 | 690 |
else if ord > 0 then dezero_rule(Thm.reflexive tm) |
23252 | 691 |
else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) |
692 |
end |
|
693 |
end |
|
694 |
in padd |
|
695 |
end; |
|
696 |
||
697 |
(* Multiplication of two polynomials. *) |
|
698 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
699 |
fun polynomial_mul_conv ctxt = |
23252 | 700 |
let |
701 |
fun pmul tm = |
|
702 |
let val (l,r) = dest_mul tm |
|
703 |
in |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
704 |
if not(is_add l) then polynomial_monomial_mul_conv ctxt tm |
23252 | 705 |
else |
706 |
if not(is_add r) then |
|
707 |
let val th1 = inst_thm [(ca,l),(cb,r)] pthm_09 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
708 |
in Thm.transitive th1 (polynomial_monomial_mul_conv ctxt (concl th1)) |
23252 | 709 |
end |
710 |
else |
|
711 |
let val (a,b) = dest_add l |
|
712 |
val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_10 |
|
713 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
714 |
val (tm3,tm4) = Thm.dest_comb tm1 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
715 |
val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv ctxt tm4) |
36945 | 716 |
val th3 = Thm.transitive th1 (Thm.combination th2 (pmul tm2)) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
717 |
in Thm.transitive th3 (polynomial_add_conv ctxt (concl th3)) |
23252 | 718 |
end |
719 |
end |
|
720 |
in fn tm => |
|
721 |
let val (l,r) = dest_mul tm |
|
722 |
in |
|
723 |
if l aconvc zero_tm then inst_thm [(ca,r)] pthm_11 |
|
724 |
else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_12 |
|
725 |
else if l aconvc one_tm then inst_thm [(ca,r)] pthm_13 |
|
726 |
else if r aconvc one_tm then inst_thm [(ca,l)] pthm_14 |
|
727 |
else pmul tm |
|
728 |
end |
|
729 |
end; |
|
730 |
||
731 |
(* Power of polynomial (optimized for the monomial and trivial cases). *) |
|
732 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
733 |
fun num_conv ctxt n = |
59538 | 734 |
nat_add_conv ctxt (Thm.apply @{cterm Suc} (Numeral.mk_cnumber @{ctyp nat} (dest_number n - 1))) |
23580
998a6fda9bb6
moved mk_cnumeral/mk_cnumber to Tools/numeral.ML;
wenzelm
parents:
23559
diff
changeset
|
735 |
|> Thm.symmetric; |
23252 | 736 |
|
737 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
738 |
fun polynomial_pow_conv ctxt = |
23252 | 739 |
let |
740 |
fun ppow tm = |
|
741 |
let val (l,n) = dest_pow tm |
|
742 |
in |
|
743 |
if n aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 |
|
744 |
else if n aconvc onen_tm then inst_thm [(cx,l)] pthm_36 |
|
745 |
else |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
746 |
let val th1 = num_conv ctxt n |
23252 | 747 |
val th2 = inst_thm [(cx,l),(cq,Thm.dest_arg (concl th1))] pthm_38 |
748 |
val (tm1,tm2) = Thm.dest_comb(concl th2) |
|
36945 | 749 |
val th3 = Thm.transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2)) |
750 |
val th4 = Thm.transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
751 |
in Thm.transitive th4 (polynomial_mul_conv ctxt (concl th4)) |
23252 | 752 |
end |
753 |
end |
|
754 |
in fn tm => |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
755 |
if is_add(Thm.dest_arg1 tm) then ppow tm else monomial_pow_conv ctxt tm |
23252 | 756 |
end; |
757 |
||
758 |
(* Negation. *) |
|
759 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
760 |
fun polynomial_neg_conv ctxt tm = |
23252 | 761 |
let val (l,r) = Thm.dest_comb tm in |
762 |
if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else |
|
763 |
let val th1 = inst_thm [(cx',r)] neg_mul |
|
36945 | 764 |
val th2 = Thm.transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1)) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
765 |
in Thm.transitive th2 (polynomial_monomial_mul_conv ctxt (concl th2)) |
23252 | 766 |
end |
767 |
end; |
|
768 |
||
769 |
||
770 |
(* Subtraction. *) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
771 |
fun polynomial_sub_conv ctxt tm = |
23252 | 772 |
let val (l,r) = dest_sub tm |
773 |
val th1 = inst_thm [(cx',l),(cy',r)] sub_add |
|
774 |
val (tm1,tm2) = Thm.dest_comb(concl th1) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
775 |
val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv ctxt tm2) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
776 |
in Thm.transitive th1 (Thm.transitive th2 (polynomial_add_conv ctxt (concl th2))) |
23252 | 777 |
end; |
778 |
||
779 |
(* Conversion from HOL term. *) |
|
780 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
781 |
fun polynomial_conv ctxt 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:
23330
diff
changeset
|
782 |
if is_semiring_constant tm then semiring_add_conv tm |
36945 | 783 |
else if not(is_comb tm) then Thm.reflexive tm |
23252 | 784 |
else |
785 |
let val (lopr,r) = Thm.dest_comb tm |
|
786 |
in if lopr aconvc neg_tm then |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
787 |
let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
788 |
in Thm.transitive th1 (polynomial_neg_conv ctxt (concl th1)) |
23252 | 789 |
end |
30866 | 790 |
else if lopr aconvc inverse_tm then |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
791 |
let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) |
36945 | 792 |
in Thm.transitive th1 (semiring_mul_conv (concl th1)) |
30866 | 793 |
end |
23252 | 794 |
else |
36945 | 795 |
if not(is_comb lopr) then Thm.reflexive tm |
23252 | 796 |
else |
797 |
let val (opr,l) = Thm.dest_comb lopr |
|
59538 | 798 |
in if opr aconvc pow_tm andalso is_number r |
23252 | 799 |
then |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
800 |
let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) r |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
801 |
in Thm.transitive th1 (polynomial_pow_conv ctxt (concl th1)) |
23252 | 802 |
end |
30866 | 803 |
else if opr aconvc divide_tm |
804 |
then |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
805 |
let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
806 |
(polynomial_conv ctxt r) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
807 |
val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv ctxt) |
30866 | 808 |
(Thm.rhs_of th1) |
36945 | 809 |
in Thm.transitive th1 th2 |
30866 | 810 |
end |
23252 | 811 |
else |
812 |
if opr aconvc add_tm orelse opr aconvc mul_tm orelse opr aconvc sub_tm |
|
813 |
then |
|
36945 | 814 |
let val th1 = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
815 |
Thm.combination |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
816 |
(Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
817 |
val f = if opr aconvc add_tm then polynomial_add_conv ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
818 |
else if opr aconvc mul_tm then polynomial_mul_conv ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
819 |
else polynomial_sub_conv ctxt |
36945 | 820 |
in Thm.transitive th1 (f (concl th1)) |
23252 | 821 |
end |
36945 | 822 |
else Thm.reflexive tm |
23252 | 823 |
end |
824 |
end; |
|
825 |
in |
|
826 |
{main = polynomial_conv, |
|
827 |
add = polynomial_add_conv, |
|
828 |
mul = polynomial_mul_conv, |
|
829 |
pow = polynomial_pow_conv, |
|
830 |
neg = polynomial_neg_conv, |
|
831 |
sub = polynomial_sub_conv} |
|
832 |
end |
|
833 |
end; |
|
834 |
||
35410 | 835 |
val nat_exp_ss = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
836 |
simpset_of |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
837 |
(put_simpset HOL_basic_ss @{context} |
54249 | 838 |
addsimps (@{thms eval_nat_numeral} @ @{thms diff_nat_numeral} @ @{thms arith_simps} @ @{thms rel_simps}) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
839 |
addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]); |
23252 | 840 |
|
59582 | 841 |
fun simple_cterm_ord t u = Term_Ord.term_ord (Thm.term_of t, Thm.term_of u) = LESS; |
27222 | 842 |
|
36710 | 843 |
|
844 |
(* various normalizing conversions *) |
|
845 |
||
30866 | 846 |
fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, field, 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:
23330
diff
changeset
|
847 |
{conv, dest_const, mk_const, is_const}) ord = |
23252 | 848 |
let |
849 |
val pow_conv = |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
850 |
Conv.arg_conv (Simplifier.rewrite (put_simpset nat_exp_ss ctxt)) |
23252 | 851 |
then_conv Simplifier.rewrite |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
852 |
(put_simpset HOL_basic_ss ctxt addsimps [nth (snd semiring) 31, nth (snd semiring) 34]) |
23330
01c09922ce59
Conversion for computation on constants now depends on the context
chaieb
parents:
23259
diff
changeset
|
853 |
then_conv conv ctxt |
01c09922ce59
Conversion for computation on constants now depends on the context
chaieb
parents:
23259
diff
changeset
|
854 |
val dat = (is_const, conv ctxt, conv ctxt, pow_conv) |
30866 | 855 |
in semiring_normalizers_conv vars semiring ring field dat ord end; |
27222 | 856 |
|
30866 | 857 |
fun semiring_normalize_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) ord = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
858 |
#main (semiring_normalizers_ord_wrapper ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
859 |
({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal}, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47108
diff
changeset
|
860 |
{conv = conv, dest_const = dest_const, mk_const = mk_const, is_const = is_const}) ord) ctxt; |
23252 | 861 |
|
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:
23330
diff
changeset
|
862 |
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:
23330
diff
changeset
|
863 |
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:
23330
diff
changeset
|
864 |
|
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
chaieb
parents:
23330
diff
changeset
|
865 |
fun semiring_normalize_ord_conv ctxt ord tm = |
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
35410
diff
changeset
|
866 |
(case match ctxt tm of |
36945 | 867 |
NONE => Thm.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:
23330
diff
changeset
|
868 |
| 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:
23330
diff
changeset
|
869 |
|
0e4452fcbeb8
normalizer conversions depend on the proof context; added functions for conversion and wrapper that sill depend on the variable ordering (_ord)
chaieb
parents:
23330
diff
changeset
|
870 |
fun semiring_normalize_conv ctxt = semiring_normalize_ord_conv ctxt simple_cterm_ord; |
23252 | 871 |
|
872 |
end; |