src/CTT/ex/equal.ML
author nipkow
Tue, 17 Oct 2000 08:00:46 +0200
changeset 10228 e653cb933293
parent 9251 bd57acd44fc1
permissions -rw-r--r--
added intermediate value thms
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1459
d12da312eff4 expanded tabs
clasohm
parents: 1294
diff changeset
     1
(*  Title:      CTT/ex/equal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1459
d12da312eff4 expanded tabs
clasohm
parents: 1294
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Equality reasoning by rewriting.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
     9
Goal "p : Sum(A,B) ==> split(p,pair) = p : Sum(A,B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
by (rtac EqE 1);
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    11
by (resolve_tac elim_rls 1  THEN  assume_tac 1);
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    12
by (rew_tac []);
1294
1358dc040edb added calls of init_html and make_chart;
clasohm
parents: 0
diff changeset
    13
qed "split_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    15
Goal "[| A type;  B type;  p : A+B |] ==> when(p,inl,inr) = p : A + B";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
by (rtac EqE 1);
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    17
by (resolve_tac elim_rls 1  THEN  assume_tac 1);
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    18
by (rew_tac []);
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    19
by (ALLGOALS assume_tac);
1294
1358dc040edb added calls of init_html and make_chart;
clasohm
parents: 0
diff changeset
    20
qed "when_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
(*in the "rec" formulation of addition, 0+n=n *)
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    24
Goal "p:N ==> rec(p,0, %y z. succ(y)) = p : N";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
by (rtac EqE 1);
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    26
by (resolve_tac elim_rls 1  THEN  assume_tac 1);
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    27
by (rew_tac []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
(*the harder version, n+0=n: recursive, uses induction hypothesis*)
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    32
Goal "p:N ==> rec(p,0, %y z. succ(z)) = p : N";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
by (rtac EqE 1);
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    34
by (resolve_tac elim_rls 1  THEN  assume_tac 1);
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    35
by (hyp_rew_tac []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
(*Associativity of addition*)
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    40
Goal "[| a:N;  b:N;  c:N |] \
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    41
\     ==> rec(rec(a, b, %x y. succ(y)), c, %x y. succ(y)) = \
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    42
\         rec(a, rec(b, c, %x y. succ(y)), %x y. succ(y)) : N";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
by (NE_tac "a" 1);
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    44
by (hyp_rew_tac []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
(*Martin-Lof (1984) page 62: pairing is surjective*)
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    49
Goal "p : Sum(A,B) ==> <split(p,%x y. x), split(p,%x y. y)> = p : Sum(A,B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
by (rtac EqE 1);
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    51
by (resolve_tac elim_rls 1  THEN  assume_tac 1);
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    52
by (DEPTH_SOLVE_1 (rew_tac []));   (*!!!!!!!*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    56
Goal "[| a : A;  b : B |] ==> \
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
    57
\    (lam u. split(u, %v w.<w,v>)) ` <a,b> = <b,a> : SUM x:B. A";
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    58
by (rew_tac []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
(*a contrived, complicated simplication, requires sum-elimination also*)
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    63
Goal "(lam f. lam x. f`(f`x)) ` (lam u. split(u, %v w.<w,v>)) =  \
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
    64
\     lam x. x  :  PROD x:(SUM y:N. N). (SUM y:N. N)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
by (resolve_tac reduction_rls 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
by (resolve_tac intrL_rls 3);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
by (rtac EqE 4);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
by (rtac SumE 4  THEN  assume_tac 4);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
(*order of unifiers is essential here*)
9251
bd57acd44fc1 more tidying. also generalized some tactics to prove "Type A" and
paulson
parents: 3837
diff changeset
    70
by (rew_tac []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
writeln"Reached end of file.";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
(*28 August 1988: loaded this file in 34 seconds*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
(*2 September 1988: loaded this file in 48 seconds*)