src/ZF/ex/Integ.ML
author paulson
Thu, 26 Sep 1996 15:49:54 +0200
changeset 2034 5079fdf938dd
parent 1677 99044cda4ef3
child 2469 b50b8c0eec01
permissions -rw-r--r--
Ran expandshort; used stac instead of ssubst
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
     1
(*  Title:      ZF/ex/integ.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
1614
c9f0fc335b12 Rewriting changes due to new arith_ss
paulson
parents: 1461
diff changeset
     6
The integers as equivalence classes over nat*nat.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
Could also prove...
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
"znegative(z) ==> $# zmagnitude(z) = $~ z"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
"~ znegative(z) ==> $# zmagnitude(z) = z"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
$< is a linear ordering
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
$+ and $* are monotonic wrt $<
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    15
open Integ;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
(*** Proving that intrel is an equivalence relation ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
(*By luck, requires no typing premises for y1, y2,y3*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
val eqa::eqb::prems = goal Arith.thy 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
    "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2;  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
\       x1: nat; x2: nat; x3: nat |]    ==>    x1 #+ y3 = x3 #+ y1";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
by (res_inst_tac [("k","x2")] add_left_cancel 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    24
by (resolve_tac prems 2);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    25
by (rtac (add_left_commute RS trans) 1 THEN typechk_tac prems);
2034
5079fdf938dd Ran expandshort; used stac instead of ssubst
paulson
parents: 1677
diff changeset
    26
by (stac eqb 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    27
by (rtac (add_left_commute RS trans) 1 THEN typechk_tac prems);
2034
5079fdf938dd Ran expandshort; used stac instead of ssubst
paulson
parents: 1677
diff changeset
    28
by (stac eqa 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    29
by (rtac (add_left_commute) 1 THEN typechk_tac prems);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    30
qed "integ_trans_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
(** Natural deduction for intrel **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    34
goalw Integ.thy [intrel_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    35
    "<<x1,y1>,<x2,y2>>: intrel <-> \
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    36
\    x1: nat & y1: nat & x2: nat & y2: nat & x1#+y2 = x2#+y1";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    37
by (fast_tac ZF_cs 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
    38
qed "intrel_iff";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    39
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    40
goalw Integ.thy [intrel_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    41
    "!!x1 x2. [| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    42
\             <<x1,y1>,<x2,y2>>: intrel";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
by (fast_tac (ZF_cs addIs prems) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
    44
qed "intrelI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
goalw Integ.thy [intrel_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
  "p: intrel --> (EX x1 y1 x2 y2. \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
\                  p = <<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1 & \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
\                  x1: nat & y1: nat & x2: nat & y2: nat)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
by (fast_tac ZF_cs 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    52
qed "intrelE_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
val [major,minor] = goal Integ.thy
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
  "[| p: intrel;  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
\     !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>;  x1#+y2 = x2#+y1; \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
\                       x1: nat; y1: nat; x2: nat; y2: nat |] ==> Q |] \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
\  ==> Q";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
    61
qed "intrelE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
val intrel_cs = ZF_cs addSIs [intrelI] addSEs [intrelE];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    65
goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    66
    "equiv(nat*nat, intrel)";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    67
by (fast_tac (intrel_cs addSEs [sym, integ_trans_lemma]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
    68
qed "equiv_intrel";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    71
val intrel_ss = 
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    72
    arith_ss addsimps [equiv_intrel RS eq_equiv_class_iff, intrel_iff,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
    73
                       add_0_right, add_succ_right]
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    74
             addcongs [conj_cong];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    76
val eq_intrelD = equiv_intrel RSN (2,eq_equiv_class);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
(** znat: the injection from nat to integ **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    80
goalw Integ.thy [integ_def,quotient_def,znat_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    81
    "!!m. m : nat ==> $#m : integ";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    82
by (fast_tac (ZF_cs addSIs [nat_0I]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
    83
qed "znat_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    85
goalw Integ.thy [znat_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    86
    "!!m n. [| $#m = $#n;  n: nat |] ==> m=n";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    87
by (dtac eq_intrelD 1);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    88
by (typechk_tac [nat_0I, SigmaI]);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    89
by (asm_full_simp_tac intrel_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    90
qed "znat_inject";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
(**** zminus: unary negation on integ ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
goalw Integ.thy [congruent_def]
1614
c9f0fc335b12 Rewriting changes due to new arith_ss
paulson
parents: 1461
diff changeset
    96
    "congruent(intrel, %<x,y>. intrel``{<y,x>})";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
by (safe_tac intrel_cs);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
    98
by (asm_full_simp_tac (intrel_ss addsimps add_ac) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    99
qed "zminus_congruent";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
(*Resolve th against the corresponding facts for zminus*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
val zminus_ize = RSLIST [equiv_intrel, zminus_congruent];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   104
goalw Integ.thy [integ_def,zminus_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   105
    "!!z. z : integ ==> $~z : integ";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   106
by (typechk_tac [split_type, SigmaI, zminus_ize UN_equiv_class_type,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   107
                 quotientI]);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   108
qed "zminus_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   110
goalw Integ.thy [integ_def,zminus_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   111
    "!!z w. [| $~z = $~w;  z: integ;  w: integ |] ==> z=w";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   112
by (etac (zminus_ize UN_equiv_class_inject) 1);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   113
by (safe_tac intrel_cs);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   114
(*The setloop is only needed because assumptions are in the wrong order!*)
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   115
by (asm_full_simp_tac (intrel_ss addsimps add_ac
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   116
                       setloop dtac eq_intrelD) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   117
qed "zminus_inject";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   119
goalw Integ.thy [zminus_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   120
    "!!x y.[| x: nat;  y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   121
by (asm_simp_tac (ZF_ss addsimps [zminus_ize UN_equiv_class, SigmaI]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   122
qed "zminus";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
goalw Integ.thy [integ_def] "!!z. z : integ ==> $~ ($~ z) = z";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   125
by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1));
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   126
by (asm_simp_tac (ZF_ss addsimps [zminus]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   127
qed "zminus_zminus";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
goalw Integ.thy [integ_def, znat_def] "$~ ($#0) = $#0";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   130
by (simp_tac (arith_ss addsimps [zminus]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   131
qed "zminus_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
(**** znegative: the test for negative integers ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
1021
9aa52d5d614f Modified proofs for new hyp_subst_tac.
lcp
parents: 904
diff changeset
   136
(*No natural number is negative!*)
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   137
goalw Integ.thy [znegative_def, znat_def]  "~ znegative($# n)";
1021
9aa52d5d614f Modified proofs for new hyp_subst_tac.
lcp
parents: 904
diff changeset
   138
by (safe_tac intrel_cs);
9aa52d5d614f Modified proofs for new hyp_subst_tac.
lcp
parents: 904
diff changeset
   139
by (dres_inst_tac [("psi", "?lhs=?rhs")] asm_rl 1);
9aa52d5d614f Modified proofs for new hyp_subst_tac.
lcp
parents: 904
diff changeset
   140
by (dres_inst_tac [("psi", "?lhs<?rhs")] asm_rl 1);
9aa52d5d614f Modified proofs for new hyp_subst_tac.
lcp
parents: 904
diff changeset
   141
by (fast_tac (FOL_cs addss
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   142
              (intrel_ss addsimps [add_le_self2 RS le_imp_not_lt])) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   143
qed "not_znegative_znat";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
29
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   145
goalw Integ.thy [znegative_def, znat_def]
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   146
    "!!n. n: nat ==> znegative($~ $# succ(n))";
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   147
by (asm_simp_tac (intrel_ss addsimps [zminus]) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
by (REPEAT 
29
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   149
    (ares_tac [refl, exI, conjI, nat_0_le,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   150
               refl RS intrelI RS imageI, consI1, nat_0I, nat_succI] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   151
qed "znegative_zminus_znat";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
(**** zmagnitude: magnitide of an integer, as a natural number ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
goalw Integ.thy [congruent_def]
1614
c9f0fc335b12 Rewriting changes due to new arith_ss
paulson
parents: 1461
diff changeset
   157
    "congruent(intrel, %<x,y>. (y#-x) #+ (x#-y))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
by (safe_tac intrel_cs);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   159
by (ALLGOALS (asm_simp_tac intrel_ss));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
by (etac rev_mp 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   161
by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1 THEN 
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   162
    REPEAT (assume_tac 1));
1614
c9f0fc335b12 Rewriting changes due to new arith_ss
paulson
parents: 1461
diff changeset
   163
by (asm_simp_tac intrel_ss 3);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   164
by (asm_simp_tac  (*this one's very sensitive to order of rewrites*)
1614
c9f0fc335b12 Rewriting changes due to new arith_ss
paulson
parents: 1461
diff changeset
   165
    (arith_ss0 addsimps [diff_add_inverse,diff_add_0,add_0_right]) 2);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   166
by (asm_simp_tac intrel_ss 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
by (rtac impI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   168
by (etac subst 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   169
by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1 THEN
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   170
    REPEAT (assume_tac 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   171
by (asm_simp_tac (arith_ss addsimps [diff_add_inverse, diff_add_0]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   172
qed "zmagnitude_congruent";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   174
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
(*Resolve th against the corresponding facts for zmagnitude*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   177
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   178
goalw Integ.thy [integ_def,zmagnitude_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   179
    "!!z. z : integ ==> zmagnitude(z) : nat";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   180
by (typechk_tac [split_type, zmagnitude_ize UN_equiv_class_type,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   181
                 add_type, diff_type]);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   182
qed "zmagnitude_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   183
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   184
goalw Integ.thy [zmagnitude_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   185
    "!!x y. [| x: nat;  y: nat |] ==> \
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   186
\           zmagnitude (intrel``{<x,y>}) = (y #- x) #+ (x #- y)";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   187
by (asm_simp_tac (ZF_ss addsimps [zmagnitude_ize UN_equiv_class, SigmaI]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   188
qed "zmagnitude";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
29
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   190
goalw Integ.thy [znat_def]
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   191
    "!!n. n: nat ==> zmagnitude($# n) = n";
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   192
by (asm_simp_tac (intrel_ss addsimps [zmagnitude]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   193
qed "zmagnitude_znat";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
29
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   195
goalw Integ.thy [znat_def]
4ec9b266ccd1 Modification of examples for the new operators, < and le.
lcp
parents: 16
diff changeset
   196
    "!!n. n: nat ==> zmagnitude($~ $# n) = n";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   197
by (asm_simp_tac (intrel_ss addsimps [zmagnitude, zminus]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   198
qed "zmagnitude_zminus_znat";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   200
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
(**** zadd: addition on integ ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   202
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
(** Congruence property for addition **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   204
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
goalw Integ.thy [congruent2_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   206
    "congruent2(intrel, %z1 z2.                      \
1110
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 1021
diff changeset
   207
\         let <x1,y1>=z1; <x2,y2>=z2                 \
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 1021
diff changeset
   208
\                           in intrel``{<x1#+x2, y1#+y2>})";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   209
(*Proof via congruent2_commuteI seems longer*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
by (safe_tac intrel_cs);
1110
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 1021
diff changeset
   211
by (asm_simp_tac (intrel_ss addsimps [add_assoc, Let_def]) 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   212
(*The rest should be trivial, but rearranging terms is hard;
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   213
  add_ac does not help rewriting with the assumptions.*)
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   214
by (res_inst_tac [("m1","x1a")] (add_left_commute RS ssubst) 1);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   215
by (res_inst_tac [("m1","x2a")] (add_left_commute RS ssubst) 3);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
by (typechk_tac [add_type]);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   217
by (asm_simp_tac (arith_ss addsimps [add_assoc RS sym]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   218
qed "zadd_congruent2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   219
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
(*Resolve th against the corresponding facts for zadd*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   221
val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   222
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   223
goalw Integ.thy [integ_def,zadd_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   224
    "!!z w. [| z: integ;  w: integ |] ==> z $+ w : integ";
1110
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 1021
diff changeset
   225
by (rtac (zadd_ize UN_equiv_class_type2) 1);
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 1021
diff changeset
   226
by (simp_tac (ZF_ss addsimps [Let_def]) 3);
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 1021
diff changeset
   227
by (REPEAT (ares_tac [split_type, add_type, quotientI, SigmaI] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   228
qed "zadd_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   229
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   230
goalw Integ.thy [zadd_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   231
  "!!x1 y1. [| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>       \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   232
\           (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =        \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   233
\           intrel `` {<x1#+x2, y1#+y2>}";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   234
by (asm_simp_tac (ZF_ss addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1);
1110
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 1021
diff changeset
   235
by (simp_tac (ZF_ss addsimps [Let_def]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   236
qed "zadd";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   238
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   239
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   240
by (asm_simp_tac (arith_ss addsimps [zadd]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   241
qed "zadd_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   242
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   243
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   244
    "!!z w. [| z: integ;  w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   245
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   246
by (asm_simp_tac (arith_ss addsimps [zminus,zadd]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   247
qed "zminus_zadd_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   249
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
    "!!z w. [| z: integ;  w: integ |] ==> z $+ w = w $+ z";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   251
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   252
by (asm_simp_tac (intrel_ss addsimps (add_ac @ [zadd])) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   253
qed "zadd_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   254
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   255
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   256
    "!!z1 z2 z3. [| z1: integ;  z2: integ;  z3: integ |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   257
\                (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   258
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
(*rewriting is much faster without intrel_iff, etc.*)
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   260
by (asm_simp_tac (arith_ss addsimps [zadd, add_assoc]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   261
qed "zadd_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   262
904
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   263
(*For AC rewriting*)
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   264
qed_goal "zadd_left_commute" Integ.thy
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   265
    "!!z1 z2 z3. [| z1:integ;  z2:integ;  z3: integ |] ==> \
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   266
\                z1$+(z2$+z3) = z2$+(z1$+z3)"
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   267
 (fn _ => [asm_simp_tac (ZF_ss addsimps [zadd_assoc RS sym, zadd_commute]) 1]);
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   268
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   269
(*Integer addition is an AC operator*)
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   270
val zadd_ac = [zadd_assoc, zadd_commute, zadd_left_commute];
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   271
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   272
goalw Integ.thy [znat_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   273
    "!!m n. [| m: nat;  n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   274
by (asm_simp_tac (arith_ss addsimps [zadd]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   275
qed "znat_add";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   277
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   278
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   279
by (asm_simp_tac (intrel_ss addsimps [zminus, zadd, add_commute]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   280
qed "zadd_zminus_inverse";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   281
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   282
goal Integ.thy "!!z. z : integ ==> ($~ z) $+ z = $#0";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   283
by (asm_simp_tac
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   284
    (ZF_ss addsimps [zadd_commute, zminus_type, zadd_zminus_inverse]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   285
qed "zadd_zminus_inverse2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   287
goal Integ.thy "!!z. z:integ ==> z $+ $#0 = z";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
by (rtac (zadd_commute RS trans) 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   289
by (REPEAT (ares_tac [znat_type, nat_0I, zadd_0] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   290
qed "zadd_0_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   291
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
(*Need properties of $- ???  Or use $- just as an abbreviation?
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   294
     [| m: nat;  n: nat;  m>=n |] ==> $# (m #- n) = ($#m) $- ($#n)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   295
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   296
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   297
(**** zmult: multiplication on integ ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   298
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   299
(** Congruence property for multiplication **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   300
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   301
goal Integ.thy 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   302
    "congruent2(intrel, %p1 p2.                 \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   303
\               split(%x1 y1. split(%x2 y2.     \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
\                   intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   305
by (rtac (equiv_intrel RS congruent2_commuteI) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   306
by (safe_tac intrel_cs);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   307
by (ALLGOALS (asm_simp_tac intrel_ss));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   308
(*Proof that zmult is congruent in one argument*)
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   309
by (asm_simp_tac 
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   310
    (arith_ss addsimps (add_ac @ [add_mult_distrib_left RS sym])) 2);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   311
by (asm_simp_tac
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   312
    (arith_ss addsimps ([add_assoc RS sym, add_mult_distrib_left RS sym])) 2);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
(*Proof that zmult is commutative on representatives*)
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   314
by (asm_simp_tac (arith_ss addsimps (mult_ac@add_ac)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   315
qed "zmult_congruent2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   316
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   317
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
(*Resolve th against the corresponding facts for zmult*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   319
val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   320
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   321
goalw Integ.thy [integ_def,zmult_def]
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   322
    "!!z w. [| z: integ;  w: integ |] ==> z $* w : integ";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   323
by (REPEAT (ares_tac [zmult_ize UN_equiv_class_type2,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   324
                      split_type, add_type, mult_type, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   325
                      quotientI, SigmaI] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   326
qed "zmult_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   327
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   328
goalw Integ.thy [zmult_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   329
     "!!x1 x2. [| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>    \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   330
\              (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =     \
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   331
\              intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}";
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   332
by (asm_simp_tac (ZF_ss addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   333
qed "zmult";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   334
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   335
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   336
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   337
by (asm_simp_tac (arith_ss addsimps [zmult]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   338
qed "zmult_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   339
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
   340
goalw Integ.thy [integ_def,znat_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   341
    "!!z. z : integ ==> $#1 $* z = z";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   342
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   343
by (asm_simp_tac (arith_ss addsimps [zmult, add_0_right]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   344
qed "zmult_1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   345
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   346
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   347
    "!!z w. [| z: integ;  w: integ |] ==> ($~ z) $* w = $~ (z $* w)";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   348
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   349
by (asm_simp_tac (intrel_ss addsimps ([zminus, zmult] @ add_ac)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   350
qed "zmult_zminus";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   351
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   352
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   353
    "!!z w. [| z: integ;  w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   354
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   355
by (asm_simp_tac (intrel_ss addsimps ([zminus, zmult] @ add_ac)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   356
qed "zmult_zminus_zminus";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   357
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   358
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   359
    "!!z w. [| z: integ;  w: integ |] ==> z $* w = w $* z";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   360
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   361
by (asm_simp_tac (intrel_ss addsimps ([zmult] @ add_ac @ mult_ac)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   362
qed "zmult_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   363
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   364
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   365
    "!!z1 z2 z3. [| z1: integ;  z2: integ;  z3: integ |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   366
\                (z1 $* z2) $* z3 = z1 $* (z2 $* z3)";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   367
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   368
by (asm_simp_tac 
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   369
    (intrel_ss addsimps ([zmult, add_mult_distrib_left, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   370
                          add_mult_distrib] @ add_ac @ mult_ac)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   371
qed "zmult_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   372
904
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   373
(*For AC rewriting*)
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   374
qed_goal "zmult_left_commute" Integ.thy
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   375
    "!!z1 z2 z3. [| z1:integ;  z2:integ;  z3: integ |] ==> \
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   376
\                z1$*(z2$*z3) = z2$*(z1$*z3)"
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   377
 (fn _ => [asm_simp_tac (ZF_ss addsimps [zmult_assoc RS sym, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   378
                                         zmult_commute]) 1]);
904
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   379
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   380
(*Integer multiplication is an AC operator*)
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   381
val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute];
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   382
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   383
goalw Integ.thy [integ_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   384
    "!!z1 z2 z3. [| z1: integ;  z2: integ;  w: integ |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   385
\                (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 29
diff changeset
   386
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
1677
99044cda4ef3 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1614
diff changeset
   387
by (asm_simp_tac (intrel_ss addsimps [zadd, zmult, add_mult_distrib]) 1);
99044cda4ef3 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1614
diff changeset
   388
by (asm_simp_tac (intrel_ss addsimps (add_ac @ mult_ac)) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 438
diff changeset
   389
qed "zadd_zmult_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   390
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   391
val integ_typechecks =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   392
    [znat_type, zminus_type, zmagnitude_type, zadd_type, zmult_type];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   393
904
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   394
val zadd_simps = 
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   395
    [zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2];
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   396
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   397
val zminus_simps = [zminus_zminus, zminus_0];
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   398
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   399
val zmult_simps = [zmult_0, zmult_1, zmult_zminus];
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   400
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   401
val integ_ss =
904
5240dcd12adb Proved zadd_left_commute and zmult_left_commute to define
lcp
parents: 804
diff changeset
   402
    arith_ss addsimps (zadd_simps @ zminus_simps @ zmult_simps @ 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   403
                       [zmagnitude_znat, zmagnitude_zminus_znat] @ 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   404
                       integ_typechecks);