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