src/HOL/OrderedGroup.ML
author haftmann
Fri, 10 Mar 2006 15:33:48 +0100
changeset 19233 77ca20b0ed77
parent 18925 2e3d508ba8dc
child 19330 eaf569aa8fd4
permissions -rw-r--r--
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
     1
(*  Title:      HOL/OrderedGroup.ML
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
     2
    ID:         $Id$
18925
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
     3
    Author:     Steven Obua, Tobias Nipkow, Technische Universit� Mnchen
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
     4
*)
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
     5
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
     6
structure ab_group_add_cancel_data :> ABEL_CANCEL  = 
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
     7
struct
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
     8
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
     9
(*** Term order for abelian groups ***)
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    10
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 18925
diff changeset
    11
fun agrp_ord a = find_index (fn a' => a = a') ["0", "HOL.plus", "HOL.uminus", "HOL.minus"];
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    12
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    13
fun termless_agrp (a, b) = (Term.term_lpo agrp_ord (a, b) = LESS);
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    14
18925
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    15
local
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    16
  val ac1 = mk_meta_eq (thm "add_assoc");
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    17
  val ac2 = mk_meta_eq (thm "add_commute");
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    18
  val ac3 = mk_meta_eq (thm "add_left_commute");
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 18925
diff changeset
    19
  fun solve_add_ac thy _ (_ $ (Const ("HOL.plus",_) $ _ $ _) $ _) =
18925
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    20
        SOME ac1
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 18925
diff changeset
    21
    | solve_add_ac thy _ (_ $ x $ (Const ("HOL.plus",_) $ y $ z)) =
18925
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    22
        if termless_agrp (y, x) then SOME ac3 else NONE
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    23
    | solve_add_ac thy _ (_ $ x $ y) =
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    24
        if termless_agrp (y, x) then SOME ac2 else NONE
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    25
    | solve_add_ac thy _ _ = NONE
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    26
in
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    27
  val add_ac_proc = Simplifier.simproc (the_context ())
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    28
    "add_ac_proc" ["x + y::'a::ab_semigroup_add"] solve_add_ac;
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    29
end;
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    30
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    31
val cancel_ss = HOL_basic_ss settermless termless_agrp
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    32
  addsimprocs [add_ac_proc] addsimps
2e3d508ba8dc speedup: use simproc for AC rules
huffman
parents: 16568
diff changeset
    33
  [thm "add_0", thm "add_0_right",
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    34
   thm "diff_def", thm "minus_add_distrib",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    35
   thm "minus_minus", thm "minus_zero",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    36
   thm "right_minus", thm "left_minus",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    37
   thm "add_minus_cancel", thm "minus_add_cancel"];
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    38
  
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    39
  val eq_reflection = thm "eq_reflection"
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    40
  
16423
24abe4c0e4b4 renamed sg_ref to thy_ref;
wenzelm
parents: 14755
diff changeset
    41
  val thy_ref = Theory.self_ref (theory "OrderedGroup")
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    42
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    43
  val T = TFree("'a", ["OrderedGroup.ab_group_add"])
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    44
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    45
  val eqI_rules = [thm "less_eqI", thm "le_eqI", thm "eq_eqI"]
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    46
  fun dest_eqI th = 
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    47
      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    48
	      (HOLogic.dest_Trueprop (concl_of th)))  
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    49
end;
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    50
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    51
structure ab_group_add_cancel = Abel_Cancel (ab_group_add_cancel_data);
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    52
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    53
Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv];