src/HOL/Integ/simproc.ML
author paulson
Fri, 23 Jul 1999 17:28:18 +0200
changeset 7076 a30e024791c6
parent 6917 eba301caceea
child 7585 dca904d4ce4c
permissions -rw-r--r--
because intT is now defined in HOLogic
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5501
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Integ/simproc
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     2
    ID:         $Id$
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     5
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
     6
Apply Abel_Cancel to the integers
5501
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     7
*)
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     8
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
     9
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    10
(*** Two lemmas needed for the simprocs ***)
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    11
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    12
(*Deletion of other terms in the formula, seeking the -x at the front of z*)
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    13
val zadd_cancel_21 = prove_goal IntDef.thy
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    14
    "((x::int) + (y + z) = y + u) = ((x + z) = u)"
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    15
  (fn _ => [stac zadd_left_commute 1,
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    16
	    rtac zadd_left_cancel 1]);
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    17
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    18
(*A further rule to deal with the case that
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    19
  everything gets cancelled on the right.*)
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    20
val zadd_cancel_end = prove_goal IntDef.thy
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    21
    "((x::int) + (y + z) = y) = (x = -z)"
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    22
  (fn _ => [stac zadd_left_commute 1,
6917
eba301caceea Introduction of integer division algorithm
paulson
parents: 5610
diff changeset
    23
	    res_inst_tac [("t", "y")] (zadd_int0_right RS subst) 1,
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    24
	    stac zadd_left_cancel 1,
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    25
	    simp_tac (simpset() addsimps [eq_zdiff_eq RS sym]) 1]);
5501
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
    26
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
    27
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    28
structure Int_Cancel_Data =
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    29
struct
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    30
  val ss		= HOL_ss
5610
377acd99d74c simpler interface for Abel_Cancel
paulson
parents: 5595
diff changeset
    31
  val eq_reflection	= eq_reflection
5501
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
    32
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    33
  val thy	= IntDef.thy
7076
a30e024791c6 because intT is now defined in HOLogic
paulson
parents: 6917
diff changeset
    34
  val T		= HOLogic.intT
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    35
  val zero	= Const ("IntDef.int", HOLogic.natT --> T) $ HOLogic.zero
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    36
  val add_cancel_21	= zadd_cancel_21
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    37
  val add_cancel_end	= zadd_cancel_end
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    38
  val add_left_cancel	= zadd_left_cancel
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    39
  val add_assoc		= zadd_assoc
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    40
  val add_commute	= zadd_commute
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    41
  val add_left_commute	= zadd_left_commute
6917
eba301caceea Introduction of integer division algorithm
paulson
parents: 5610
diff changeset
    42
  val add_0		= zadd_int0
eba301caceea Introduction of integer division algorithm
paulson
parents: 5610
diff changeset
    43
  val add_0_right	= zadd_int0_right
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    44
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    45
  val eq_diff_eq	= eq_zdiff_eq
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    46
  val eqI_rules		= [zless_eqI, zeq_eqI, zle_eqI]
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    47
  fun dest_eqI th = 
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    48
      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    49
	      (HOLogic.dest_Trueprop (concl_of th)))
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    50
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    51
  val diff_def		= zdiff_def
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    52
  val minus_add_distrib	= zminus_zadd_distrib
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    53
  val minus_minus	= zminus_zminus
6917
eba301caceea Introduction of integer division algorithm
paulson
parents: 5610
diff changeset
    54
  val minus_0		= zminus_int0
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    55
  val add_inverses	= [zadd_zminus_inverse, zadd_zminus_inverse2];
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    56
  val cancel_simps	= [zadd_zminus_cancel, zminus_zadd_cancel]
5501
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
    57
end;
a63e0c326e6c new files in Integ
paulson
parents:
diff changeset
    58
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    59
structure Int_Cancel = Abel_Cancel (Int_Cancel_Data);
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    60
5595
d283d6120548 now invokes functor
paulson
parents: 5582
diff changeset
    61
Addsimprocs [Int_Cancel.sum_conv, Int_Cancel.rel_conv];
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    62
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5553
diff changeset
    63