| author | paulson | 
| Fri, 11 Aug 2000 13:27:17 +0200 | |
| changeset 9578 | ab26d6c8ebfe | 
| parent 9251 | bd57acd44fc1 | 
| permissions | -rw-r--r-- | 
| 1459 | 1 | (* Title: CTT/ex/equal | 
| 0 | 2 | ID: $Id$ | 
| 1459 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1991 University of Cambridge | 
| 5 | ||
| 6 | Equality reasoning by rewriting. | |
| 7 | *) | |
| 8 | ||
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 9 | Goal "p : Sum(A,B) ==> split(p,pair) = p : Sum(A,B)"; | 
| 0 | 10 | by (rtac EqE 1); | 
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
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: 
3837diff
changeset | 12 | by (rew_tac []); | 
| 1294 | 13 | qed "split_eq"; | 
| 0 | 14 | |
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 15 | Goal "[| A type; B type; p : A+B |] ==> when(p,inl,inr) = p : A + B"; | 
| 0 | 16 | by (rtac EqE 1); | 
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
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: 
3837diff
changeset | 18 | by (rew_tac []); | 
| 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 19 | by (ALLGOALS assume_tac); | 
| 1294 | 20 | qed "when_eq"; | 
| 0 | 21 | |
| 22 | ||
| 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: 
3837diff
changeset | 24 | Goal "p:N ==> rec(p,0, %y z. succ(y)) = p : N"; | 
| 0 | 25 | by (rtac EqE 1); | 
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
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: 
3837diff
changeset | 27 | by (rew_tac []); | 
| 0 | 28 | result(); | 
| 29 | ||
| 30 | ||
| 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: 
3837diff
changeset | 32 | Goal "p:N ==> rec(p,0, %y z. succ(z)) = p : N"; | 
| 0 | 33 | by (rtac EqE 1); | 
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
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: 
3837diff
changeset | 35 | by (hyp_rew_tac []); | 
| 0 | 36 | result(); | 
| 37 | ||
| 38 | ||
| 39 | (*Associativity of addition*) | |
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 40 | Goal "[| a:N; b:N; c:N |] \ | 
| 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
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: 
3837diff
changeset | 42 | \ rec(a, rec(b, c, %x y. succ(y)), %x y. succ(y)) : N"; | 
| 0 | 43 | by (NE_tac "a" 1); | 
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 44 | by (hyp_rew_tac []); | 
| 0 | 45 | result(); | 
| 46 | ||
| 47 | ||
| 48 | (*Martin-Lof (1984) page 62: pairing is surjective*) | |
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 49 | Goal "p : Sum(A,B) ==> <split(p,%x y. x), split(p,%x y. y)> = p : Sum(A,B)"; | 
| 0 | 50 | by (rtac EqE 1); | 
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
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: 
3837diff
changeset | 52 | by (DEPTH_SOLVE_1 (rew_tac [])); (*!!!!!!!*) | 
| 0 | 53 | result(); | 
| 54 | ||
| 55 | ||
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 56 | Goal "[| a : A; b : B |] ==> \ | 
| 3837 | 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: 
3837diff
changeset | 58 | by (rew_tac []); | 
| 0 | 59 | result(); | 
| 60 | ||
| 61 | ||
| 62 | (*a contrived, complicated simplication, requires sum-elimination also*) | |
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 63 | Goal "(lam f. lam x. f`(f`x)) ` (lam u. split(u, %v w.<w,v>)) = \ | 
| 3837 | 64 | \ lam x. x : PROD x:(SUM y:N. N). (SUM y:N. N)"; | 
| 0 | 65 | by (resolve_tac reduction_rls 1); | 
| 66 | by (resolve_tac intrL_rls 3); | |
| 67 | by (rtac EqE 4); | |
| 68 | by (rtac SumE 4 THEN assume_tac 4); | |
| 69 | (*order of unifiers is essential here*) | |
| 9251 
bd57acd44fc1
more tidying.  also generalized some tactics to prove "Type A" and
 paulson parents: 
3837diff
changeset | 70 | by (rew_tac []); | 
| 0 | 71 | result(); | 
| 72 | ||
| 73 | writeln"Reached end of file."; | |
| 74 | (*28 August 1988: loaded this file in 34 seconds*) | |
| 75 | (*2 September 1988: loaded this file in 48 seconds*) |