src/ZF/IMP/Equiv.ML
author paulson
Wed, 05 Nov 1997 13:14:15 +0100
changeset 4152 451104c223e2
parent 4091 771b1f6422a8
child 4298 b69eedd3aa6c
permissions -rw-r--r--
Ran expandshort, especially to introduce Safe_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
     1
(*  Title:      ZF/IMP/Equiv.ML
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
     3
    Author:     Heiko Loetzbeyer & Robert Sandner, TUM
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
     4
    Copyright   1994 TUM
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
     5
*)
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
     6
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
     7
val prems = goal Equiv.thy "[| a: aexp; sigma: loc -> nat |] ==> \
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
     8
\ <a,sigma> -a-> n <-> A(a,sigma) = n";
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
     9
by (res_inst_tac [("x","n")] spec 1);                       (* quantify n *)
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    10
by (res_inst_tac [("x","a")] aexp.induct 1);                (* struct. ind. *)
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    11
by (resolve_tac prems 1);                                   (* type prem. *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    12
by (rewrite_goals_tac A_rewrite_rules);                     (* rewr. Den.   *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    13
by (TRYALL (fast_tac (claset() addSIs (evala.intrs@prems)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    14
                              addSEs aexp_elim_cases)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 672
diff changeset
    15
qed "aexp_iff";
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    16
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    17
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    18
val aexp1 = prove_goal Equiv.thy
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    19
        "[| <a,sigma> -a-> n; a: aexp; sigma: loc -> nat |] \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    20
        \ ==> A(a,sigma) = n"       (* destruction rule *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    21
     (fn prems => [(fast_tac (claset() addSIs ((aexp_iff RS iffD1)::prems)) 1)]);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    22
val aexp2 = aexp_iff RS iffD2;
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    23
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    24
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    25
val bexp_elim_cases = 
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    26
   [
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    27
    evalb.mk_cases bexp.con_defs "<true,sigma> -b-> x",
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    28
    evalb.mk_cases bexp.con_defs "<false,sigma> -b-> x",
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    29
    evalb.mk_cases bexp.con_defs "<ROp(f,a0,a1),sigma> -b-> x",
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    30
    evalb.mk_cases bexp.con_defs "<noti(b),sigma> -b-> x",
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    31
    evalb.mk_cases bexp.con_defs "<b0 andi b1,sigma> -b-> x",
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    32
    evalb.mk_cases bexp.con_defs "<b0 ori b1,sigma> -b-> x"
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    33
   ];
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    34
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    35
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    36
val prems = goal Equiv.thy "[| b: bexp; sigma: loc -> nat |] ==> \
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    37
\ <b,sigma> -b-> w <-> B(b,sigma) = w";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    38
by (res_inst_tac [("x","w")] spec 1);                   (* quantify w *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    39
by (res_inst_tac [("x","b")] bexp.induct 1);            (* struct. ind. *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    40
by (resolve_tac prems 1);                               (* type prem. *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    41
by (rewrite_goals_tac B_rewrite_rules);                 (* rewr. Den.   *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    42
by (TRYALL (fast_tac (claset() addSIs (evalb.intrs@prems@[aexp2])
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    43
                            addSDs [aexp1] addSEs bexp_elim_cases)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 672
diff changeset
    44
qed "bexp_iff";
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    45
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    46
val bexp1 = prove_goal Equiv.thy
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    47
        "[| <b,sigma> -b-> w; b: bexp; sigma: loc -> nat |]\
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    48
        \ ==> B(b,sigma) = w"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    49
     (fn prems => [(fast_tac (claset() addSIs ((bexp_iff RS iffD1)::prems)) 1)]);
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    50
val bexp2 = bexp_iff RS iffD2;
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    51
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    52
goal Equiv.thy "!!c. <c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    53
500
0842a38074e7 some small simplifications
nipkow
parents: 482
diff changeset
    54
(* start with rule induction *)
1742
328fb06a1648 Updated for new form of induction rules
paulson
parents: 1461
diff changeset
    55
by (etac evalc.induct 1);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    56
500
0842a38074e7 some small simplifications
nipkow
parents: 482
diff changeset
    57
by (rewrite_tac (Gamma_def::C_rewrite_rules));
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    58
(* skip *)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    59
by (Fast_tac 1);
500
0842a38074e7 some small simplifications
nipkow
parents: 482
diff changeset
    60
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    61
(* assign *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    62
by (asm_full_simp_tac (simpset() addsimps [aexp1,assign_type] @ op_type_intrs) 1);
672
1922f98b8f7e com1,2: added simplifier calls to remove use of ssubst in fast_tac
lcp
parents: 518
diff changeset
    63
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    64
(* comp *)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    65
by (Fast_tac 1);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    66
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    67
(* if *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    68
by (asm_simp_tac (simpset() addsimps [bexp1]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    69
by (asm_simp_tac (simpset() addsimps [bexp1]) 1);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    70
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    71
(* while *)
672
1922f98b8f7e com1,2: added simplifier calls to remove use of ssubst in fast_tac
lcp
parents: 518
diff changeset
    72
by (etac (rewrite_rule [Gamma_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    73
          (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    74
by (asm_simp_tac (simpset() addsimps [bexp1]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    75
by (fast_tac (claset() addSIs [bexp1,idI]@evalb_type_intrs) 1);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    76
672
1922f98b8f7e com1,2: added simplifier calls to remove use of ssubst in fast_tac
lcp
parents: 518
diff changeset
    77
by (etac (rewrite_rule [Gamma_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 808
diff changeset
    78
          (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    79
by (asm_simp_tac (simpset() addsimps [bexp1]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    80
by (fast_tac (claset() addSIs [bexp1,compI]@evalb_type_intrs) 1);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    81
500
0842a38074e7 some small simplifications
nipkow
parents: 482
diff changeset
    82
val com1 = result();
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    83
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    84
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    85
AddSIs [aexp2,bexp2,B_type,A_type];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    86
AddIs evalc.intrs;
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    87
AddEs [C_type,C_type_fst];
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    88
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
    89
val [prem] = goal Equiv.thy
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
    90
    "c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";
808
c51c1f59e59e ran expandshort script
lcp
parents: 760
diff changeset
    91
by (rtac (prem RS com.induct) 1);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    92
by (rewrite_tac C_rewrite_rules);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    93
by Safe_tac;
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    94
by (ALLGOALS Asm_full_simp_tac);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    95
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    96
(* skip *)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
    97
by (Fast_tac 1);
500
0842a38074e7 some small simplifications
nipkow
parents: 482
diff changeset
    98
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
    99
(* assign *)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
   100
by (Fast_tac 1);
500
0842a38074e7 some small simplifications
nipkow
parents: 482
diff changeset
   101
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
   102
(* comp *)
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
   103
by (REPEAT (EVERY [(dtac bspec 1),(atac 1)]));
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
   104
by (Asm_full_simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
   105
by (Fast_tac 1);
500
0842a38074e7 some small simplifications
nipkow
parents: 482
diff changeset
   106
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
   107
(* while *)
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
   108
by (EVERY1 [forward_tac [Gamma_bnd_mono], etac induct, atac]);
808
c51c1f59e59e ran expandshort script
lcp
parents: 760
diff changeset
   109
by (rewtac Gamma_def);  
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   110
by Safe_tac;
518
4530c45370b4 Proof beautification
nipkow
parents: 511
diff changeset
   111
by (EVERY1 [dtac bspec, atac]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
   112
by (ALLGOALS Asm_full_simp_tac);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
   113
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
   114
(* while, if *)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
   115
by (ALLGOALS Fast_tac);
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
   116
val com2 = result();
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
   117
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
   118
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
   119
(**** Proof of Equivalence ****)
482
3a4e092ba69c Initial revision
nipkow
parents:
diff changeset
   120
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
   121
goal Equiv.thy
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
   122
    "ALL c:com. C(c) = {io:(loc->nat)*(loc->nat). <c,fst(io)> -c-> snd(io)}";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   123
by (fast_tac (claset() addIs [C_subset RS subsetD]
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
   124
	              addEs [com2 RS bspec]
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1742
diff changeset
   125
		      addDs [com1]
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   126
		      addss (simpset())) 1);
511
b2be4790da7a re-organized using new theory sections
lcp
parents: 510
diff changeset
   127
val com_equivalence = result();