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