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