src/HOL/Integ/bin_simprocs.ML
author paulson
Mon, 02 Aug 1999 11:24:30 +0200
changeset 7143 9c02848c5404
parent 7080 b551a5a8966c
child 7402 e53d5f0c7c94
permissions -rw-r--r--
the SVC link-up
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7080
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Integ/bin_simproc
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     2
    ID:         $Id$
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     3
    Author:    Lawrence C Paulson, Cambridge University Computer Laboratory
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     5
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     6
Installation of simprocs that work on numeric literals
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     7
*)
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     8
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
     9
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    10
(** Combining of literal coefficients in sums of products **)
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    11
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    12
Goal "(x < y) = (x-y < (#0::int))";
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    13
by (simp_tac (simpset() addsimps zcompare_rls) 1);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    14
qed "zless_iff_zdiff_zless_0";
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    15
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    16
Goal "(x = y) = (x-y = (#0::int))";
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    17
by (simp_tac (simpset() addsimps zcompare_rls) 1);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    18
qed "eq_iff_zdiff_eq_0";
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    19
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    20
Goal "(x <= y) = (x-y <= (#0::int))";
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    21
by (simp_tac (simpset() addsimps zcompare_rls) 1);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    22
qed "zle_iff_zdiff_zle_0";
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    23
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    24
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    25
structure Int_CC_Data : COMBINE_COEFF_DATA =
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    26
struct
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    27
  val ss		= HOL_ss
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    28
  val eq_reflection	= eq_reflection
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    29
  val thy	= Bin.thy
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    30
  val T		= Type ("IntDef.int", [])
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    31
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    32
  val trans		= trans
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    33
  val add_ac		= zadd_ac
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    34
  val diff_def		= zdiff_def
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    35
  val minus_add_distrib	= zminus_zadd_distrib
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    36
  val minus_minus	= zminus_zminus
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    37
  val mult_commute	= zmult_commute
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    38
  val mult_1_right	= zmult_1_right
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    39
  val add_mult_distrib = zadd_zmult_distrib2
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    40
  val diff_mult_distrib = zdiff_zmult_distrib2
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    41
  val mult_minus_right = zmult_zminus_right
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    42
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    43
  val rel_iff_rel_0_rls = [zless_iff_zdiff_zless_0, eq_iff_zdiff_eq_0, 
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    44
			   zle_iff_zdiff_zle_0]
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    45
  fun dest_eqI th = 
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    46
      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    47
	      (HOLogic.dest_Trueprop (concl_of th)))
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    48
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    49
end;
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    50
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    51
structure Int_CC = Combine_Coeff (Int_CC_Data);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    52
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    53
Addsimprocs [Int_CC.sum_conv, Int_CC.rel_conv];
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    54
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    55
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    56
(** Constant folding for integer plus and times **)
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    57
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    58
(*We do not need
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    59
    structure Int_Plus_Assoc = Assoc_Fold (Int_Plus_Assoc_Data);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    60
  because cancel_coeffs does the same thing*)
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    61
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    62
structure Int_Times_Assoc_Data : ASSOC_FOLD_DATA =
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    63
struct
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    64
  val ss		= HOL_ss
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    65
  val eq_reflection	= eq_reflection
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    66
  val thy    = Bin.thy
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    67
  val T	     = HOLogic.intT
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    68
  val plus   = Const ("op *", [HOLogic.intT,HOLogic.intT] ---> HOLogic.intT);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    69
  val add_ac = zmult_ac
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    70
end;
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    71
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    72
structure Int_Times_Assoc = Assoc_Fold (Int_Times_Assoc_Data);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    73
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    74
Addsimprocs [Int_Times_Assoc.conv];
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    75
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    76
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    77
(** The same for the naturals **)
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    78
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    79
structure Nat_Plus_Assoc_Data : ASSOC_FOLD_DATA =
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    80
struct
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    81
  val ss		= HOL_ss
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    82
  val eq_reflection	= eq_reflection
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    83
  val thy    = Bin.thy
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    84
  val T	     = HOLogic.natT
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    85
  val plus   = Const ("op +", [HOLogic.natT,HOLogic.natT] ---> HOLogic.natT);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    86
  val add_ac = add_ac
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    87
end;
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    88
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    89
structure Nat_Plus_Assoc = Assoc_Fold (Nat_Plus_Assoc_Data);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    90
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    91
structure Nat_Times_Assoc_Data : ASSOC_FOLD_DATA =
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    92
struct
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    93
  val ss		= HOL_ss
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    94
  val eq_reflection	= eq_reflection
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    95
  val thy    = Bin.thy
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    96
  val T	     = HOLogic.natT
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    97
  val plus   = Const ("op *", [HOLogic.natT,HOLogic.natT] ---> HOLogic.natT);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    98
  val add_ac = mult_ac
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
    99
end;
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
   100
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
   101
structure Nat_Times_Assoc = Assoc_Fold (Nat_Times_Assoc_Data);
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
   102
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
   103
Addsimprocs [Nat_Plus_Assoc.conv, Nat_Times_Assoc.conv];
b551a5a8966c new simprocs assoc_fold and combine_coeff
paulson
parents:
diff changeset
   104