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