13854
|
1 |
(*
|
|
2 |
Title: Normalisation method for locale cring
|
|
3 |
Id: $Id$
|
|
4 |
Author: Clemens Ballarin
|
|
5 |
Copyright: TU Muenchen
|
|
6 |
*)
|
|
7 |
|
|
8 |
local
|
|
9 |
|
|
10 |
(*** Lexicographic path order on terms.
|
|
11 |
|
|
12 |
See Baader & Nipkow, Term rewriting, CUP 1998.
|
|
13 |
Without variables. Const, Var, Bound, Free and Abs are treated all as
|
|
14 |
constants.
|
|
15 |
|
|
16 |
f_ord maps strings to integers and serves two purposes:
|
|
17 |
- Predicate on constant symbols. Those that are not recognised by f_ord
|
|
18 |
must be mapped to ~1.
|
|
19 |
- Order on the recognised symbols. These must be mapped to distinct
|
|
20 |
integers >= 0.
|
|
21 |
|
|
22 |
***)
|
|
23 |
|
|
24 |
fun dest_hd f_ord (Const (a, T)) =
|
|
25 |
let val ord = f_ord a in
|
|
26 |
if ord = ~1 then ((1, ((a, 0), T)), 0) else ((0, (("", ord), T)), 0)
|
|
27 |
end
|
|
28 |
| dest_hd _ (Free (a, T)) = ((1, ((a, 0), T)), 0)
|
|
29 |
| dest_hd _ (Var v) = ((1, v), 1)
|
|
30 |
| dest_hd _ (Bound i) = ((1, (("", i), Term.dummyT)), 2)
|
|
31 |
| dest_hd _ (Abs (_, T, _)) = ((1, (("", 0), T)), 3);
|
|
32 |
|
|
33 |
fun term_lpo f_ord (s, t) =
|
|
34 |
let val (f, ss) = strip_comb s and (g, ts) = strip_comb t in
|
|
35 |
if forall (fn si => term_lpo f_ord (si, t) = LESS) ss
|
|
36 |
then case hd_ord f_ord (f, g) of
|
|
37 |
GREATER =>
|
|
38 |
if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
|
|
39 |
then GREATER else LESS
|
|
40 |
| EQUAL =>
|
|
41 |
if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
|
|
42 |
then list_ord (term_lpo f_ord) (ss, ts)
|
|
43 |
else LESS
|
|
44 |
| LESS => LESS
|
|
45 |
else GREATER
|
|
46 |
end
|
|
47 |
and hd_ord f_ord (f, g) = case (f, g) of
|
|
48 |
(Abs (_, T, t), Abs (_, U, u)) =>
|
|
49 |
(case term_lpo f_ord (t, u) of EQUAL => Term.typ_ord (T, U) | ord => ord)
|
|
50 |
| (_, _) => prod_ord (prod_ord int_ord
|
|
51 |
(prod_ord Term.indexname_ord Term.typ_ord)) int_ord
|
|
52 |
(dest_hd f_ord f, dest_hd f_ord g)
|
|
53 |
|
|
54 |
in
|
|
55 |
|
|
56 |
(*** Term order for commutative rings ***)
|
|
57 |
|
|
58 |
fun ring_ord a =
|
|
59 |
find_index_eq a ["CRing.ring.zero", "CRing.ring.add", "Cring.ring.a_inv",
|
|
60 |
"CRing.ring.minus", "Group.monoid.one", "Group.semigroup.mult"];
|
|
61 |
|
|
62 |
fun termless_ring (a, b) = (term_lpo ring_ord (a, b) = LESS);
|
|
63 |
|
|
64 |
val cring_ss = HOL_ss settermless termless_ring;
|
|
65 |
|
|
66 |
fun cring_normalise ctxt = let
|
13864
|
67 |
fun filter t = (case HOLogic.dest_Trueprop t of
|
13854
|
68 |
Const ("CRing.cring_axioms", _) $ Free (s, _) => [s]
|
13864
|
69 |
| _ => [])
|
13854
|
70 |
handle TERM _ => [];
|
|
71 |
val insts = flat (map (filter o #prop o rep_thm)
|
|
72 |
(ProofContext.prems_of ctxt));
|
|
73 |
val _ = warning ("Rings in proof context: " ^ implode insts);
|
|
74 |
val simps =
|
|
75 |
flat (map (fn s => ProofContext.get_thms ctxt (s ^ ".cring_simprules"))
|
|
76 |
insts);
|
|
77 |
in Method.SIMPLE_METHOD' HEADGOAL (asm_full_simp_tac (cring_ss addsimps simps))
|
|
78 |
end;
|
|
79 |
|
|
80 |
(*
|
|
81 |
val ring_ss = HOL_basic_ss settermless termless_ring addsimps
|
|
82 |
[a_assoc, l_zero, l_neg, a_comm, m_assoc, l_one, l_distr, m_comm, minus_def,
|
|
83 |
r_zero, r_neg, r_neg2, r_neg1, minus_add, minus_minus, minus0,
|
|
84 |
a_lcomm, m_lcomm, (*r_one,*) r_distr, l_null, r_null, l_minus, r_minus];
|
|
85 |
*)
|
|
86 |
|
|
87 |
end; |