author | bauerg |
Thu, 28 Apr 2005 17:08:08 +0200 | |
changeset 15871 | e524119dbf19 |
parent 15570 | 8d8c70b41bab |
child 16486 | 1a12cdb6ee6b |
permissions | -rw-r--r-- |
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 = |
|
13936 | 59 |
find_index_eq a ["CRing.ring.zero", "CRing.ring.add", "CRing.a_inv", |
14963 | 60 |
"CRing.minus", "Group.monoid.one", "Group.monoid.mult"]; |
13854 | 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 |
|
14399
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
67 |
fun ring_filter t = (case HOLogic.dest_Trueprop t of |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
68 |
Const ("CRing.ring_axioms", _) $ Free (s, _) => [s] |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
69 |
| _ => []) |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
70 |
handle TERM _ => []; |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
71 |
fun comm_filter t = (case HOLogic.dest_Trueprop t of |
14963 | 72 |
Const ("Group.comm_monoid_axioms", _) $ Free (s, _) => [s] |
13864 | 73 |
| _ => []) |
13854 | 74 |
handle TERM _ => []; |
14399
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
75 |
|
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
76 |
val prems = ProofContext.prems_of ctxt; |
15570 | 77 |
val rings = List.concat (map (ring_filter o #prop o rep_thm) prems); |
78 |
val comms = List.concat (map (comm_filter o #prop o rep_thm) prems); |
|
14399
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
79 |
val non_comm_rings = rings \\ comms; |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
80 |
val comm_rings = rings inter_string comms; |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
81 |
val _ = tracing |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
82 |
("Non-commutative rings in proof context: " ^ commas non_comm_rings ^ |
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
83 |
"\nCommutative rings in proof context: " ^ commas comm_rings); |
13854 | 84 |
val simps = |
15570 | 85 |
List.concat (map (fn s => ProofContext.get_thms ctxt (s ^ ".ring_simprules", NONE)) |
14399
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
86 |
non_comm_rings) @ |
15570 | 87 |
List.concat (map (fn s => ProofContext.get_thms ctxt (s ^ ".cring_simprules", NONE)) |
14399
dc677b35e54f
New lemmas about inversion of restricted functions.
ballarin
parents:
13936
diff
changeset
|
88 |
comm_rings); |
13854 | 89 |
in Method.SIMPLE_METHOD' HEADGOAL (asm_full_simp_tac (cring_ss addsimps simps)) |
90 |
end; |
|
91 |
||
92 |
(* |
|
93 |
val ring_ss = HOL_basic_ss settermless termless_ring addsimps |
|
94 |
[a_assoc, l_zero, l_neg, a_comm, m_assoc, l_one, l_distr, m_comm, minus_def, |
|
95 |
r_zero, r_neg, r_neg2, r_neg1, minus_add, minus_minus, minus0, |
|
96 |
a_lcomm, m_lcomm, (*r_one,*) r_distr, l_null, r_null, l_minus, r_minus]; |
|
97 |
*) |
|
98 |
||
99 |
end; |