src/HOL/OrderedGroup.ML
author nipkow
Tue, 09 Aug 2005 11:44:38 +0200
changeset 17040 6682c93b7d9f
parent 16568 e02fe7ae212b
child 18925 2e3d508ba8dc
permissions -rw-r--r--
added finite(option) to Recdef.thy
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$
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
     3
    Author:     Steven Obua, Tobias Nipkow, Technische Universität München
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
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    11
fun agrp_ord a = find_index_eq a ["0", "op +", "uminus", "op -"];
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
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    15
val cancel_ss = HOL_basic_ss settermless termless_agrp addsimps
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    16
  [thm "add_assoc", thm "add_commute", thm "add_left_commute",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    17
   thm "add_0", thm "add_0_right",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    18
   thm "diff_def", thm "minus_add_distrib",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    19
   thm "minus_minus", thm "minus_zero",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    20
   thm "right_minus", thm "left_minus",
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    21
   thm "add_minus_cancel", thm "minus_add_cancel"];
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    22
  
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    23
  val eq_reflection = thm "eq_reflection"
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    24
  
16423
24abe4c0e4b4 renamed sg_ref to thy_ref;
wenzelm
parents: 14755
diff changeset
    25
  val thy_ref = Theory.self_ref (theory "OrderedGroup")
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    26
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    27
  val T = TFree("'a", ["OrderedGroup.ab_group_add"])
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 16423
diff changeset
    28
14755
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    29
  val eqI_rules = [thm "less_eqI", thm "le_eqI", thm "eq_eqI"]
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    30
  fun dest_eqI th = 
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    31
      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    32
	      (HOLogic.dest_Trueprop (concl_of th)))  
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    33
end;
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    34
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    35
structure ab_group_add_cancel = Abel_Cancel (ab_group_add_cancel_data);
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    36
5cc6e6b9e27a simplification for abelian groups
obua
parents:
diff changeset
    37
Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv];