author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17274 | 746bb4c56800 |
child 19233 | 77ca20b0ed77 |
permissions | -rw-r--r-- |
13735 | 1 |
(* |
2 |
Title: Term order, needed for normal forms in rings |
|
3 |
Id: $Id$ |
|
13783 | 4 |
Author: Clemens Ballarin |
13735 | 5 |
Copyright: TU Muenchen |
6 |
*) |
|
7 |
||
13783 | 8 |
(*** Term order for commutative rings ***) |
9 |
||
10 |
fun ring_ord a = |
|
11 |
find_index_eq a ["0", "op +", "uminus", "op -", "1", "op *"]; |
|
12 |
||
16706 | 13 |
fun termless_ring (a, b) = (Term.term_lpo ring_ord (a, b) = LESS); |
13783 | 14 |
|
15661
9ef583b08647
reverted renaming of Some/None in comments and strings;
wenzelm
parents:
15531
diff
changeset
|
15 |
(* Some code useful for debugging |
13735 | 16 |
|
13783 | 17 |
val intT = HOLogic.intT; |
18 |
val a = Free ("a", intT); |
|
19 |
val b = Free ("b", intT); |
|
20 |
val c = Free ("c", intT); |
|
21 |
val plus = Const ("op +", [intT, intT]--->intT); |
|
22 |
val mult = Const ("op *", [intT, intT]--->intT); |
|
23 |
val uminus = Const ("uminus", intT-->intT); |
|
24 |
val one = Const ("1", intT); |
|
25 |
val f = Const("f", intT-->intT); |
|
26 |
||
27 |
*) |
|
13735 | 28 |
|
29 |
(*** Rewrite rules ***) |
|
30 |
||
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
31 |
val a_assoc = thm "ring_class.a_assoc"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
32 |
val l_zero = thm "ring_class.l_zero"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
33 |
val l_neg = thm "ring_class.l_neg"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
34 |
val a_comm = thm "ring_class.a_comm"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
35 |
val m_assoc = thm "ring_class.m_assoc"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
36 |
val l_one = thm "ring_class.l_one"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
37 |
val l_distr = thm "ring_class.l_distr"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
38 |
val m_comm = thm "ring_class.m_comm"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
39 |
val minus_def = thm "ring_class.minus_def"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
40 |
val inverse_def = thm "ring_class.inverse_def"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
41 |
val divide_def = thm "ring_class.divide_def"; |
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16706
diff
changeset
|
42 |
val power_def = thm "ring_class.power_def"; |
13735 | 43 |
|
44 |
(* These are the following axioms: |
|
45 |
||
46 |
a_assoc: "(a + b) + c = a + (b + c)" |
|
47 |
l_zero: "0 + a = a" |
|
48 |
l_neg: "(-a) + a = 0" |
|
49 |
a_comm: "a + b = b + a" |
|
50 |
||
51 |
m_assoc: "(a * b) * c = a * (b * c)" |
|
52 |
l_one: "1 * a = a" |
|
53 |
||
54 |
l_distr: "(a + b) * c = a * c + b * c" |
|
55 |
||
56 |
m_comm: "a * b = b * a" |
|
57 |
||
58 |
-- {* Definition of derived operations *} |
|
59 |
||
60 |
minus_def: "a - b = a + (-b)" |
|
61 |
inverse_def: "inverse a = (if a dvd 1 then THE x. a*x = 1 else 0)" |
|
62 |
divide_def: "a / b = a * inverse b" |
|
63 |
power_def: "a ^ n = nat_rec 1 (%u b. b * a) n" |
|
64 |
*) |
|
65 |
||
66 |
(* These lemmas are needed in the proofs *) |
|
67 |
val trans = thm "trans"; |
|
68 |
val sym = thm "sym"; |
|
69 |
val subst = thm "subst"; |
|
70 |
val box_equals = thm "box_equals"; |
|
71 |
val arg_cong = thm "arg_cong"; |
|
72 |
||
73 |
(* derived rewrite rules *) |
|
16706 | 74 |
val a_lcomm = prove_goal (the_context ()) "(a::'a::ring)+(b+c) = b+(a+c)" |
13735 | 75 |
(fn _ => [rtac (a_comm RS trans) 1, rtac (a_assoc RS trans) 1, |
76 |
rtac (a_comm RS arg_cong) 1]); |
|
16706 | 77 |
val r_zero = prove_goal (the_context ()) "(a::'a::ring) + 0 = a" |
13735 | 78 |
(fn _ => [rtac (a_comm RS trans) 1, rtac l_zero 1]); |
16706 | 79 |
val r_neg = prove_goal (the_context ()) "(a::'a::ring) + (-a) = 0" |
13735 | 80 |
(fn _ => [rtac (a_comm RS trans) 1, rtac l_neg 1]); |
16706 | 81 |
val r_neg2 = prove_goal (the_context ()) "(a::'a::ring) + (-a + b) = b" |
13735 | 82 |
(fn _ => [rtac (a_assoc RS sym RS trans) 1, |
83 |
simp_tac (simpset() addsimps [r_neg, l_zero]) 1]); |
|
16706 | 84 |
val r_neg1 = prove_goal (the_context ()) "-(a::'a::ring) + (a + b) = b" |
13735 | 85 |
(fn _ => [rtac (a_assoc RS sym RS trans) 1, |
86 |
simp_tac (simpset() addsimps [l_neg, l_zero]) 1]); |
|
87 |
(* auxiliary *) |
|
16706 | 88 |
val a_lcancel = prove_goal (the_context ()) "!! a::'a::ring. a + b = a + c ==> b = c" |
13735 | 89 |
(fn _ => [rtac box_equals 1, rtac l_zero 2, rtac l_zero 2, |
90 |
res_inst_tac [("a1", "a")] (l_neg RS subst) 1, |
|
91 |
asm_simp_tac (simpset() addsimps [a_assoc]) 1]); |
|
16706 | 92 |
val minus_add = prove_goal (the_context ()) "-((a::'a::ring) + b) = (-a) + (-b)" |
13735 | 93 |
(fn _ => [res_inst_tac [("a", "a+b")] a_lcancel 1, |
94 |
simp_tac (simpset() addsimps [r_neg, l_neg, l_zero, |
|
95 |
a_assoc, a_comm, a_lcomm]) 1]); |
|
16706 | 96 |
val minus_minus = prove_goal (the_context ()) "-(-(a::'a::ring)) = a" |
13735 | 97 |
(fn _ => [rtac a_lcancel 1, rtac (r_neg RS trans) 1, rtac (l_neg RS sym) 1]); |
16706 | 98 |
val minus0 = prove_goal (the_context ()) "- 0 = (0::'a::ring)" |
13735 | 99 |
(fn _ => [rtac a_lcancel 1, rtac (r_neg RS trans) 1, |
100 |
rtac (l_zero RS sym) 1]); |
|
101 |
||
102 |
(* derived rules for multiplication *) |
|
16706 | 103 |
val m_lcomm = prove_goal (the_context ()) "(a::'a::ring)*(b*c) = b*(a*c)" |
13735 | 104 |
(fn _ => [rtac (m_comm RS trans) 1, rtac (m_assoc RS trans) 1, |
105 |
rtac (m_comm RS arg_cong) 1]); |
|
16706 | 106 |
val r_one = prove_goal (the_context ()) "(a::'a::ring) * 1 = a" |
13735 | 107 |
(fn _ => [rtac (m_comm RS trans) 1, rtac l_one 1]); |
16706 | 108 |
val r_distr = prove_goal (the_context ()) "(a::'a::ring) * (b + c) = a * b + a * c" |
13735 | 109 |
(fn _ => [rtac (m_comm RS trans) 1, rtac (l_distr RS trans) 1, |
110 |
simp_tac (simpset() addsimps [m_comm]) 1]); |
|
111 |
(* the following proof is from Jacobson, Basic Algebra I, pp. 88-89 *) |
|
16706 | 112 |
val l_null = prove_goal (the_context ()) "0 * (a::'a::ring) = 0" |
13735 | 113 |
(fn _ => [rtac a_lcancel 1, rtac (l_distr RS sym RS trans) 1, |
114 |
simp_tac (simpset() addsimps [r_zero]) 1]); |
|
16706 | 115 |
val r_null = prove_goal (the_context ()) "(a::'a::ring) * 0 = 0" |
13735 | 116 |
(fn _ => [rtac (m_comm RS trans) 1, rtac l_null 1]); |
16706 | 117 |
val l_minus = prove_goal (the_context ()) "(-(a::'a::ring)) * b = - (a * b)" |
13735 | 118 |
(fn _ => [rtac a_lcancel 1, rtac (r_neg RS sym RSN (2, trans)) 1, |
119 |
rtac (l_distr RS sym RS trans) 1, |
|
120 |
simp_tac (simpset() addsimps [l_null, r_neg]) 1]); |
|
16706 | 121 |
val r_minus = prove_goal (the_context ()) "(a::'a::ring) * (-b) = - (a * b)" |
13735 | 122 |
(fn _ => [rtac a_lcancel 1, rtac (r_neg RS sym RSN (2, trans)) 1, |
123 |
rtac (r_distr RS sym RS trans) 1, |
|
124 |
simp_tac (simpset() addsimps [r_null, r_neg]) 1]); |
|
125 |
||
126 |
val ring_ss = HOL_basic_ss settermless termless_ring addsimps |
|
127 |
[a_assoc, l_zero, l_neg, a_comm, m_assoc, l_one, l_distr, m_comm, minus_def, |
|
128 |
r_zero, r_neg, r_neg2, r_neg1, minus_add, minus_minus, minus0, |
|
13783 | 129 |
a_lcomm, m_lcomm, (*r_one,*) r_distr, l_null, r_null, l_minus, r_minus]; |
130 |
||
131 |
(* Note: r_one is not necessary in ring_ss *) |
|
13735 | 132 |
|
13783 | 133 |
val x = bind_thms ("ring_simps", |
134 |
[l_zero, r_zero, l_neg, r_neg, minus_minus, minus0, |
|
13735 | 135 |
l_one, r_one, l_null, r_null, l_minus, r_minus]); |
136 |
||
137 |
(* note: not added (and not proved): |
|
138 |
a_lcancel_eq, a_rcancel_eq, power_one, power_Suc, power_zero, power_one, |
|
139 |
m_lcancel_eq, m_rcancel_eq, |
|
140 |
thms involving dvd, integral domains, fields |
|
141 |
*) |