src/ZF/Integ/Int.ML
author oheimb
Mon, 26 Jun 2000 16:18:51 +0200 (2000-06-26)
changeset 9147 9a64807da023
parent 8201 a81d18b0a9b1
child 9333 5cacc383157a
permissions -rw-r--r--
corrected specifications and simplified proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     1
(*  Title:      ZF/Integ/Int.ML
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     2
    ID:         $Id$
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     5
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     6
The integers as equivalence classes over nat*nat.
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     7
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     8
Could also prove...
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     9
"znegative(z) ==> $# zmagnitude(z) = $~ z"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    10
"~ znegative(z) ==> $# zmagnitude(z) = z"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    11
$< is a linear ordering
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    12
$+ and $* are monotonic wrt $<
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    13
*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    14
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    15
AddSEs [quotientE];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    16
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    17
(*** Proving that intrel is an equivalence relation ***)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    18
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    19
(*By luck, requires no typing premises for y1, y2,y3*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    20
val eqa::eqb::prems = goal Arith.thy 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    21
    "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2;  \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    22
\       x1: nat; x2: nat; x3: nat |]    ==>    x1 #+ y3 = x3 #+ y1";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    23
by (cut_facts_tac prems 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    24
by (res_inst_tac [("k","x2")] add_left_cancel 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    25
by (rtac (add_left_commute RS trans) 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    26
by Auto_tac;
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    27
by (stac eqb 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    28
by (rtac (add_left_commute RS trans) 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    29
by (stac eqa 3);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    30
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_left_commute])));
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    31
qed "int_trans_lemma";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    32
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    33
(** Natural deduction for intrel **)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    34
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    35
Goalw [intrel_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    36
    "<<x1,y1>,<x2,y2>>: intrel <-> \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    37
\    x1: nat & y1: nat & x2: nat & y2: nat & x1#+y2 = x2#+y1";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    38
by (Fast_tac 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    39
qed "intrel_iff";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    40
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    41
Goalw [intrel_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    42
    "[| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    43
\             <<x1,y1>,<x2,y2>>: intrel";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    44
by (fast_tac (claset() addIs prems) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    45
qed "intrelI";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    46
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    47
(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    48
Goalw [intrel_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    49
  "p: intrel --> (EX x1 y1 x2 y2. \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    50
\                  p = <<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1 & \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    51
\                  x1: nat & y1: nat & x2: nat & y2: nat)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    52
by (Fast_tac 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    53
qed "intrelE_lemma";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    54
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    55
val [major,minor] = goal thy
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    56
  "[| p: intrel;  \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    57
\     !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>;  x1#+y2 = x2#+y1; \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    58
\                       x1: nat; y1: nat; x2: nat; y2: nat |] ==> Q |] \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    59
\  ==> Q";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    60
by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    61
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    62
qed "intrelE";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    63
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    64
AddSIs [intrelI];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    65
AddSEs [intrelE];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    66
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    67
Goalw [equiv_def, refl_def, sym_def, trans_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    68
    "equiv(nat*nat, intrel)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    69
by (fast_tac (claset() addSEs [sym, int_trans_lemma]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    70
qed "equiv_intrel";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    71
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    72
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    73
Addsimps [equiv_intrel RS eq_equiv_class_iff, intrel_iff,
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    74
	  add_0_right, add_succ_right];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    75
Addcongs [conj_cong];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    76
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    77
val eq_intrelD = equiv_intrel RSN (2,eq_equiv_class);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    78
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    79
(** int_of: the injection from nat to int **)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    80
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    81
Goalw [int_def,quotient_def,int_of_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    82
    "m : nat ==> $#m : int";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    83
by Auto_tac;
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    84
qed "int_of_type";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    85
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    86
Addsimps [int_of_type];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    87
AddTCs   [int_of_type];
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    88
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    89
Goalw [int_of_def] "[| $#m = $#n;  m: nat |] ==> m=n";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    90
by (dtac (sym RS eq_intrelD) 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
    91
by Auto_tac;
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    92
qed "int_of_inject";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    93
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    94
AddSDs [int_of_inject];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    95
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    96
Goal "m: nat ==> ($# m = $# n) <-> (m = n)"; 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    97
by (Blast_tac 1); 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    98
qed "int_of_eq"; 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    99
Addsimps [int_of_eq]; 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   100
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   101
(**** zminus: unary negation on int ****)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   102
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   103
Goalw [congruent_def] "congruent(intrel, %<x,y>. intrel``{<y,x>})";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   104
by Safe_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   105
by (asm_full_simp_tac (simpset() addsimps add_ac) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   106
qed "zminus_congruent";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   107
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   108
(*Resolve th against the corresponding facts for zminus*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   109
val zminus_ize = RSLIST [equiv_intrel, zminus_congruent];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   110
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   111
Goalw [int_def,zminus_def] "z : int ==> $~z : int";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   112
by (typecheck_tac (tcset() addTCs [zminus_ize UN_equiv_class_type]));
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   113
qed "zminus_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   114
AddTCs [zminus_type];
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   115
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   116
Goalw [int_def,zminus_def] "[| $~z = $~w;  z: int;  w: int |] ==> z=w";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   117
by (etac (zminus_ize UN_equiv_class_inject) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   118
by Safe_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   119
(*The setloop is only needed because assumptions are in the wrong order!*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   120
by (asm_full_simp_tac (simpset() addsimps add_ac
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   121
                       setloop dtac eq_intrelD) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   122
qed "zminus_inject";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   123
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   124
Goalw [zminus_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   125
    "[| x: nat;  y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   126
by (asm_simp_tac (simpset() addsimps [zminus_ize UN_equiv_class, SigmaI]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   127
qed "zminus";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   128
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   129
Goalw [int_def] "z : int ==> $~ ($~ z) = z";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   130
by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   131
by (asm_simp_tac (simpset() addsimps [zminus]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   132
qed "zminus_zminus";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   133
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   134
Goalw [int_def, int_of_def] "$~ ($#0) = $#0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   135
by (simp_tac (simpset() addsimps [zminus]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   136
qed "zminus_0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   137
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   138
Addsimps [zminus_zminus, zminus_0];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   139
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   140
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   141
(**** znegative: the test for negative integers ****)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   142
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   143
(*No natural number is negative!*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   144
Goalw [znegative_def, int_of_def]  "~ znegative($# n)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   145
by Safe_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   146
by (dres_inst_tac [("psi", "?lhs=?rhs")] asm_rl 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   147
by (dres_inst_tac [("psi", "?lhs<?rhs")] asm_rl 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   148
by (force_tac (claset(),
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   149
	       simpset() addsimps [add_le_self2 RS le_imp_not_lt]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   150
qed "not_znegative_int_of";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   151
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   152
Addsimps [not_znegative_int_of];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   153
AddSEs   [not_znegative_int_of RS notE];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   154
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   155
Goalw [znegative_def, int_of_def] "n: nat ==> znegative($~ $# succ(n))";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   156
by (asm_simp_tac (simpset() addsimps [zminus]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   157
by (blast_tac (claset() addIs [nat_0_le]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   158
qed "znegative_zminus_int_of";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   159
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   160
Addsimps [znegative_zminus_int_of];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   161
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   162
Goalw [znegative_def, int_of_def] "[| n: nat; ~ znegative($~ $# n) |] ==> n=0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   163
by (asm_full_simp_tac (simpset() addsimps [zminus, image_singleton_iff]) 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   164
by (etac natE 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   165
by (dres_inst_tac [("x","0")] spec 2);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   166
by Auto_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   167
qed "not_znegative_imp_zero";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   168
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   169
(**** zmagnitude: magnitide of an integer, as a natural number ****)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   170
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   171
Goalw [zmagnitude_def] "n: nat ==> zmagnitude($# n) = n";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   172
by Auto_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   173
qed "zmagnitude_int_of";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   174
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   175
Goalw [zmagnitude_def] "n: nat ==> zmagnitude($~ $# n) = n";
5758
27a2b36efd95 corrected auto_tac (applications of unsafe wrappers)
oheimb
parents: 5561
diff changeset
   176
by (force_tac(claset() addDs [not_znegative_imp_zero], simpset())1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   177
qed "zmagnitude_zminus_int_of";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   178
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   179
Addsimps [zmagnitude_int_of, zmagnitude_zminus_int_of];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   180
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   181
Goalw [zmagnitude_def] "zmagnitude(z) : nat";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   182
by (rtac theI2 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   183
by Auto_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   184
qed "zmagnitude_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   185
AddTCs [zmagnitude_type];
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   186
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   187
Goalw [int_def, znegative_def, int_of_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   188
     "[| z: int; ~ znegative(z) |] ==> EX n:nat. z = $# n"; 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   189
by (auto_tac(claset() , simpset() addsimps [image_singleton_iff]));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   190
by (rename_tac "i j" 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   191
by (dres_inst_tac [("x", "i")] spec 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   192
by (dres_inst_tac [("x", "j")] spec 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   193
by (rtac bexI 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   194
by (rtac (add_diff_inverse2 RS sym) 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   195
by Auto_tac;
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 6153
diff changeset
   196
by (asm_full_simp_tac (simpset() addsimps [not_lt_iff_le]) 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   197
qed "not_zneg_int_of";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   198
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   199
Goal "[| z: int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z"; 
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   200
by (dtac not_zneg_int_of 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   201
by Auto_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   202
qed "not_zneg_mag"; 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   203
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   204
Addsimps [not_zneg_mag];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   205
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   206
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   207
Goalw [int_def, znegative_def, int_of_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   208
     "[| z: int; znegative(z) |] ==> EX n:nat. z = $~ ($# succ(n))"; 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   209
by (auto_tac(claset() addSDs [less_imp_Suc_add], 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   210
	     simpset() addsimps [zminus, image_singleton_iff]));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   211
by (rename_tac "m n j k" 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   212
by (subgoal_tac "j #+ succ(m #+ k) = j #+ n" 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   213
by (rotate_tac ~2 2);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   214
by (asm_full_simp_tac (simpset() addsimps add_ac) 2);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   215
by (blast_tac (claset() addSDs [add_left_cancel]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   216
qed "zneg_int_of";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   217
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   218
Goal "[| z: int; znegative(z) |] ==> $# (zmagnitude(z)) = $~ z"; 
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   219
by (dtac zneg_int_of 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   220
by Auto_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   221
qed "zneg_mag"; 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   222
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   223
Addsimps [zneg_mag];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   224
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   225
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   226
(**** zadd: addition on int ****)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   227
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   228
(** Congruence property for addition **)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   229
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   230
Goalw [congruent2_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   231
    "congruent2(intrel, %z1 z2.                      \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   232
\         let <x1,y1>=z1; <x2,y2>=z2                 \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   233
\                           in intrel``{<x1#+x2, y1#+y2>})";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   234
(*Proof via congruent2_commuteI seems longer*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   235
by Safe_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   236
by (asm_simp_tac (simpset() addsimps [add_assoc, Let_def]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   237
(*The rest should be trivial, but rearranging terms is hard;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   238
  add_ac does not help rewriting with the assumptions.*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   239
by (res_inst_tac [("m1","x1a")] (add_left_commute RS ssubst) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   240
by (res_inst_tac [("m1","x2a")] (add_left_commute RS ssubst) 3);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   241
by Auto_tac;
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   242
by (asm_simp_tac (simpset() addsimps [add_assoc RS sym]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   243
qed "zadd_congruent2";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   244
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   245
(*Resolve th against the corresponding facts for zadd*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   246
val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   247
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   248
Goalw [int_def,zadd_def] "[| z: int;  w: int |] ==> z $+ w : int";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   249
by (rtac (zadd_ize UN_equiv_class_type2) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   250
by (simp_tac (simpset() addsimps [Let_def]) 3);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   251
by (REPEAT (ares_tac [split_type, add_type, quotientI, SigmaI] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   252
qed "zadd_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   253
AddTCs [zadd_type];
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   254
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   255
Goalw [zadd_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   256
  "[| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>       \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   257
\           (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =        \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   258
\           intrel `` {<x1#+x2, y1#+y2>}";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   259
by (asm_simp_tac (simpset() addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   260
by (simp_tac (simpset() addsimps [Let_def]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   261
qed "zadd";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   262
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   263
Goalw [int_def,int_of_def] "z : int ==> $#0 $+ z = z";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   264
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   265
by (asm_simp_tac (simpset() addsimps [zadd]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   266
qed "zadd_0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   267
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   268
Goalw [int_def] "[| z: int;  w: int |] ==> $~ (z $+ w) = $~ z $+ $~ w";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   269
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   270
by (asm_simp_tac (simpset() addsimps [zminus,zadd]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   271
qed "zminus_zadd_distrib";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   272
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   273
Goalw [int_def] "[| z: int;  w: int |] ==> z $+ w = w $+ z";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   274
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   275
by (asm_simp_tac (simpset() addsimps add_ac @ [zadd]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   276
qed "zadd_commute";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   277
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   278
Goalw [int_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   279
    "[| z1: int;  z2: int;  z3: int |]   \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   280
\    ==> (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   281
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   282
(*rewriting is much faster without intrel_iff, etc.*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   283
by (asm_simp_tac (simpset() addsimps [zadd, add_assoc]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   284
qed "zadd_assoc";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   285
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   286
(*For AC rewriting*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   287
Goal "[| z1:int;  z2:int;  z3: int |] ==> z1$+(z2$+z3) = z2$+(z1$+z3)";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   288
by (asm_simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   289
by (asm_simp_tac (simpset() addsimps [zadd_commute]) 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   290
qed "zadd_left_commute";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   291
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   292
(*Integer addition is an AC operator*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   293
val zadd_ac = [zadd_assoc, zadd_commute, zadd_left_commute];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   294
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   295
Goalw [int_of_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   296
    "[| m: nat;  n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   297
by (asm_simp_tac (simpset() addsimps [zadd]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   298
qed "int_of_add";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   299
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   300
Goalw [int_def,int_of_def] "z : int ==> z $+ ($~ z) = $#0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   301
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   302
by (asm_simp_tac (simpset() addsimps [zminus, zadd, add_commute]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   303
qed "zadd_zminus_inverse";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   304
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   305
Goal "z : int ==> ($~ z) $+ z = $#0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   306
by (asm_simp_tac
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   307
    (simpset() addsimps [zadd_commute, zminus_type, zadd_zminus_inverse]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   308
qed "zadd_zminus_inverse2";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   309
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   310
Goal "z:int ==> z $+ $#0 = z";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   311
by (rtac (zadd_commute RS trans) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   312
by (REPEAT (ares_tac [int_of_type, nat_0I, zadd_0] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   313
qed "zadd_0_right";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   314
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   315
Addsimps [zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   316
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   317
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   318
(*Need properties of $- ???  Or use $- just as an abbreviation?
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   319
     [| m: nat;  n: nat;  m>=n |] ==> $# (m #- n) = ($#m) $- ($#n)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   320
*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   321
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   322
(**** zmult: multiplication on int ****)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   323
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   324
(** Congruence property for multiplication **)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   325
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   326
Goal "congruent2(intrel, %p1 p2.                 \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   327
\               split(%x1 y1. split(%x2 y2.     \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   328
\                   intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   329
by (rtac (equiv_intrel RS congruent2_commuteI) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   330
by Safe_tac;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   331
by (ALLGOALS Asm_simp_tac);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   332
(*Proof that zmult is congruent in one argument*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   333
by (asm_simp_tac 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   334
    (simpset() addsimps add_ac @ [add_mult_distrib_left RS sym]) 2);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   335
by (asm_simp_tac
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   336
    (simpset() addsimps [add_assoc RS sym, add_mult_distrib_left RS sym]) 2);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   337
(*Proof that zmult is commutative on representatives*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   338
by (asm_simp_tac (simpset() addsimps mult_ac@add_ac) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   339
qed "zmult_congruent2";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   340
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   341
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   342
(*Resolve th against the corresponding facts for zmult*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   343
val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   344
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   345
Goalw [int_def,zmult_def] "[| z: int;  w: int |] ==> z $* w : int";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   346
by (REPEAT (ares_tac [zmult_ize UN_equiv_class_type2,
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   347
                      split_type, add_type, mult_type, 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   348
                      quotientI, SigmaI] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   349
qed "zmult_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   350
AddTCs [zmult_type];
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   351
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   352
Goalw [zmult_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   353
     "[| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>    \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   354
\              (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =     \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   355
\              intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   356
by (asm_simp_tac (simpset() addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   357
qed "zmult";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   358
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   359
Goalw [int_def,int_of_def] "z : int ==> $#0 $* z = $#0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   360
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   361
by (asm_simp_tac (simpset() addsimps [zmult]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   362
qed "zmult_0";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   363
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   364
Goalw [int_def,int_of_def] "z : int ==> $#1 $* z = z";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   365
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   366
by (asm_simp_tac (simpset() addsimps [zmult, add_0_right]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   367
qed "zmult_1";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   368
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   369
Goalw [int_def] "[| z: int;  w: int |] ==> ($~ z) $* w = $~ (z $* w)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   370
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   371
by (asm_simp_tac (simpset() addsimps [zminus, zmult] @ add_ac) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   372
qed "zmult_zminus";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   373
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   374
Addsimps [zmult_0, zmult_1, zmult_zminus];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   375
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   376
Goalw [int_def] "[| z: int;  w: int |] ==> ($~ z) $* ($~ w) = (z $* w)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   377
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   378
by (asm_simp_tac (simpset() addsimps [zminus, zmult] @ add_ac) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   379
qed "zmult_zminus_zminus";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   380
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   381
Goalw [int_def] "[| z: int;  w: int |] ==> z $* w = w $* z";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   382
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   383
by (asm_simp_tac (simpset() addsimps [zmult] @ add_ac @ mult_ac) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   384
qed "zmult_commute";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   385
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   386
Goalw [int_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   387
    "[| z1: int;  z2: int;  z3: int |]     \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   388
\    ==> (z1 $* z2) $* z3 = z1 $* (z2 $* z3)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   389
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   390
by (asm_simp_tac 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   391
    (simpset() addsimps [zmult, add_mult_distrib_left, 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   392
                         add_mult_distrib] @ add_ac @ mult_ac) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   393
qed "zmult_assoc";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   394
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   395
(*For AC rewriting*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   396
Goal "[| z1:int;  z2:int;  z3: int |] ==> z1$*(z2$*z3) = z2$*(z1$*z3)";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   397
by (asm_simp_tac (simpset() addsimps [zmult_assoc RS sym]) 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 5758
diff changeset
   398
by (asm_simp_tac (simpset() addsimps [zmult_commute]) 1);
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   399
qed "zmult_left_commute";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   400
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   401
(*Integer multiplication is an AC operator*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   402
val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   403
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   404
Goalw [int_def]
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   405
    "[| z1: int;  z2: int;  w: int |] ==> \
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   406
\                (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   407
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   408
by (asm_simp_tac (simpset() addsimps [zadd, zmult, add_mult_distrib]) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   409
by (asm_simp_tac (simpset() addsimps add_ac @ mult_ac) 1);
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   410
qed "zadd_zmult_distrib";
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   411
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   412
val int_typechecks =
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   413
    [int_of_type, zminus_type, zmagnitude_type, zadd_type, zmult_type];
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   414
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   415
Addsimps int_typechecks;
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   416
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   417
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   418