src/HOLCF/IOA/NTP/Correctness.ML
author wenzelm
Sun, 25 Oct 1998 12:33:27 +0100
changeset 5769 6a422b22ba02
parent 5192 704dd3a6d47d
child 5857 701498a38a76
permissions -rw-r--r--
tuned checklist;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     1
(*  Title:      HOL/IOA/NTP/Correctness.ML
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     2
    ID:         $Id$
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Konrad Slind
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     4
    Copyright   1994  TU Muenchen
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     5
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     6
The main correctness proof: Impl implements Spec
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     7
*)
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     8
4815
b8a32ef742d9 removed split_all_tac from claset() globally within IOA
oheimb
parents: 4681
diff changeset
     9
(* repeated from Traces.ML *)
b8a32ef742d9 removed split_all_tac from claset() globally within IOA
oheimb
parents: 4681
diff changeset
    10
claset_ref() := claset() delSWrapper "split_all_tac";
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    11
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    12
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
val hom_ioas = [Spec.ioa_def, Spec.trans_def,
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    14
                Sender.sender_trans_def,Receiver.receiver_trans_def]
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    15
               @ impl_ioas;
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    16
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    17
val impl_asigs = [Sender.sender_asig_def,Receiver.receiver_asig_def,
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    18
                  Abschannel.srch_asig_def,Abschannel.rsch_asig_def];
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    19
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    20
(* Two simpsets: - simpset() is basic, ss' unfolds hom_ioas *)
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    21
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    22
Delsimps [split_paired_All];
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    23
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    24
val ss' = (simpset() addsimps hom_ioas);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    25
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    26
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    27
(* A lemma about restricting the action signature of the implementation
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    28
 * to that of the specification.
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    29
 ****************************)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
    30
Goal
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    31
 "a:externals(asig_of(restrict impl_ioa (externals spec_sig))) = \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    32
\ (case a of                  \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    33
\     S_msg(m) => True        \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    34
\   | R_msg(m) => True        \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    35
\   | S_pkt(pkt) => False  \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    36
\   | R_pkt(pkt) => False  \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    37
\   | S_ack(b) => False    \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    38
\   | R_ack(b) => False    \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    39
\   | C_m_s => False          \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    40
\   | C_m_r => False          \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    41
\   | C_r_s => False          \
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    42
\   | C_r_r(m) => False)";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    43
 by (simp_tac (simpset() addsimps ([externals_def, restrict_def,
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    44
                            restrict_asig_def, Spec.sig_def]
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    45
                            @asig_projections)) 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    46
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
    47
  by (induct_tac "a" 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    48
  by (ALLGOALS(simp_tac (simpset() addsimps [actions_def]@asig_projections)));
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    49
 (* 2 *)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    50
  by (simp_tac (simpset() addsimps impl_ioas) 1);
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    51
  by (simp_tac (simpset() addsimps impl_asigs) 1);
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    52
  by (simp_tac (simpset() addsimps 
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    53
             [asig_of_par, asig_comp_def]@asig_projections) 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    54
  by (simp_tac rename_ss 1); 
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    55
 (* 1 *)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    56
  by (simp_tac (simpset() addsimps impl_ioas) 1);
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    57
  by (simp_tac (simpset() addsimps impl_asigs) 1);
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    58
  by (simp_tac (simpset() addsimps 
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    59
             [asig_of_par, asig_comp_def]@asig_projections) 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    60
qed "externals_lemma"; 
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    61
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    62
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    63
val sels = [Sender.sbit_def, Sender.sq_def, Sender.ssending_def,
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    64
            Receiver.rbit_def, Receiver.rq_def, Receiver.rsending_def];
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    65
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    66
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    67
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    68
(* Proof of correctness *)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
    69
Goalw [Spec.ioa_def, is_weak_ref_map_def]
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    70
  "is_weak_ref_map hom (restrict impl_ioa (externals spec_sig)) spec_ioa";
4833
2e53109d4bc8 Renamed expand_const -> split_const
nipkow
parents: 4815
diff changeset
    71
by (simp_tac (simpset() delsplits [split_if] addsimps 
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    72
    [Correctness.hom_def, cancel_restrict, externals_lemma]) 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    73
by (rtac conjI 1);
4681
a331c1f5a23e expand_if is now by default part of the simpset.
nipkow
parents: 4098
diff changeset
    74
 by (simp_tac ss' 1);
a331c1f5a23e expand_if is now by default part of the simpset.
nipkow
parents: 4098
diff changeset
    75
 by (asm_simp_tac (simpset() addsimps sels) 1);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    76
by (REPEAT(rtac allI 1));
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    77
by (rtac imp_conj_lemma 1);   (* from lemmas.ML *)
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    78
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
    79
by (induct_tac "a"  1);
4833
2e53109d4bc8 Renamed expand_const -> split_const
nipkow
parents: 4815
diff changeset
    80
by (asm_simp_tac (ss' addsplits [split_if]) 1);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    81
by (forward_tac [inv4] 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    82
by (fast_tac (claset() addss (simpset() addsimps [neq_Nil_conv])) 1);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    83
by (simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    84
by (simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    85
by (simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    86
by (simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    87
by (simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    88
by (simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    89
by (simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    90
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    91
by (asm_simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    92
by (forward_tac [inv4] 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    93
by (forward_tac [inv2] 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    94
by (etac disjE 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    95
by (Asm_simp_tac 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
    96
by (fast_tac (claset() addss (simpset() addsimps [neq_Nil_conv])) 1);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    97
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    98
by (asm_simp_tac ss' 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    99
by (forward_tac [inv2] 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   100
by (etac disjE 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   101
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   102
by (forward_tac [inv3] 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   103
by (case_tac "sq(sen(s))=[]" 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   104
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   105
by (asm_full_simp_tac ss' 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
   106
by (fast_tac (claset() addSDs [add_leD1 RS leD]) 1);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   107
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   108
by (case_tac "m = hd(sq(sen(s)))" 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   109
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
   110
by (fast_tac (claset() addss (simpset() addsimps [neq_Nil_conv])) 1);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   111
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   112
by (Asm_full_simp_tac 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 3221
diff changeset
   113
by (fast_tac (claset() addSDs [add_leD1 RS leD]) 1);
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   114
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   115
by (Asm_full_simp_tac 1);
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   116
qed"ntp_correct";