src/HOL/IOA/ABP/Correctness.ML
author nipkow
Mon, 04 Mar 1996 14:38:30 +0100
changeset 1532 769a4517ad7b
parent 1465 5d7a7e439cec
child 1570 fd1b9c721ac7
permissions -rw-r--r--
Proof modification.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1138
mueller
parents: 1050
diff changeset
     1
 (*  Title:      HOL/IOA/example/Correctness.ML
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1138
mueller
parents: 1050
diff changeset
     3
    Author:     Tobias Nipkow & Konrad Slind
mueller
parents: 1050
diff changeset
     4
    Copyright   1994  TU Muenchen
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     5
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     6
*)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     7
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     8
open Correctness; open Abschannel; open Abschannel_finite;
1138
mueller
parents: 1050
diff changeset
     9
open Impl; open Impl_finite;
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    10
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    11
goal Abschannel.thy "(? x.x=P & Q(x)) = Q(P)";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    12
by (fast_tac HOL_cs 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    13
qed"exis_elim";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    14
1328
9a449a91425d *** empty log message ***
mueller
parents: 1266
diff changeset
    15
Delsimps [split_paired_All];
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    16
Addsimps 
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    17
 ([srch_asig_def, rsch_asig_def, rsch_ioa_def, srch_ioa_def, ch_ioa_def, 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    18
   ch_asig_def, srch_actions_def, rsch_actions_def, rename_def, asig_of_def, 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    19
   actions_def, exis_elim, srch_trans_def, rsch_trans_def, ch_trans_def, 
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    20
   trans_of_def] @ asig_projections @ set_lemmas);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    21
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    22
val abschannel_fin = [srch_fin_asig_def, rsch_fin_asig_def, 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    23
                      rsch_fin_ioa_def, srch_fin_ioa_def, 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    24
                      ch_fin_ioa_def,ch_fin_trans_def,ch_fin_asig_def];
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    25
Addsimps abschannel_fin;
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    26
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    27
val impl_ioas =  [Sender.sender_ioa_def,Receiver.receiver_ioa_def];
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    28
val impl_trans = [Sender.sender_trans_def,Receiver.receiver_trans_def];
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    29
val impl_asigs = [Sender.sender_asig_def,Receiver.receiver_asig_def];
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    30
Addcongs [let_weak_cong];
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    31
Addsimps [Let_def, ioa_triple_proj, starts_of_par];
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    32
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    33
val env_ioas = [Env.env_ioa_def,Env.env_asig_def,Env.env_trans_def];
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    34
val hom_ioas = env_ioas @ impl_ioas @ impl_trans @ impl_asigs @ 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    35
               asig_projections @ set_lemmas;
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    36
Addsimps hom_ioas;
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    37
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    38
Addsimps [reduce_Nil,reduce_Cons];
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    39
1138
mueller
parents: 1050
diff changeset
    40
mueller
parents: 1050
diff changeset
    41
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    42
(* auxiliary function *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    43
fun rotate n i = EVERY(replicate n (etac revcut_rl i));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    44
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    45
(* lemmas about reduce *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    46
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    47
goal Correctness.thy "(reduce(l)=[]) = (l=[])";  
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    48
 by (rtac iffI 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    49
 by (subgoal_tac "(l~=[]) --> (reduce(l)~=[])" 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    50
 by (fast_tac HOL_cs 1); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    51
 by (List.list.induct_tac "l" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    52
 by (Simp_tac 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    53
 by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    54
 by (rtac (expand_list_case RS iffD2) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    55
 by (Asm_full_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    56
 by (REPEAT (rtac allI 1)); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    57
 by (rtac impI 1); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    58
 by (hyp_subst_tac 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    59
 by (rtac (expand_if RS ssubst) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    60
 by (Asm_full_simp_tac 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    61
 by (Asm_full_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    62
val l_iff_red_nil = result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    63
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    64
goal Correctness.thy "s~=[] --> hd(s)=hd(reduce(s))";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    65
by (List.list.induct_tac "s" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    66
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    67
by (case_tac "list =[]" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    68
by (Asm_full_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    69
(* main case *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    70
by (rotate 1 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    71
by (asm_full_simp_tac list_ss 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    72
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    73
by (rtac (expand_list_case RS iffD2) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    74
by (asm_full_simp_tac list_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    75
by (REPEAT (rtac allI 1)); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    76
 by (rtac impI 1); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    77
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    78
by (REPEAT(hyp_subst_tac 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    79
by (etac subst 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    80
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    81
qed"hd_is_reduce_hd";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    82
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    83
(* to be used in the following Lemma *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    84
goal Correctness.thy "l~=[] --> reverse(reduce(l))~=[]";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    85
by (List.list.induct_tac "l" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    86
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    87
by (case_tac "list =[]" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    88
by (asm_full_simp_tac (!simpset addsimps [reverse_Cons]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    89
(* main case *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    90
by (rotate 1 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    91
by (Asm_full_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    92
by (cut_inst_tac [("l","list")] cons_not_nil 1); 
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    93
by (Asm_full_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    94
by (REPEAT (etac exE 1));
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    95
by (Asm_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    96
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    97
by (hyp_subst_tac 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
    98
by (asm_full_simp_tac (!simpset addsimps [reverse_Cons]) 1); 
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    99
qed"rev_red_not_nil";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   100
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   101
(* shows applicability of the induction hypothesis of the following Lemma 1 *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   102
goal Correctness.thy "!!l.[| l~=[] |] ==>   \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   103
\   hd(reverse(reduce(a#l))) = hd(reverse(reduce(l)))";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   104
 by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   105
 by (rtac (expand_list_case RS iffD2) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   106
 by (asm_full_simp_tac list_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   107
 by (REPEAT (rtac allI 1)); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   108
 by (rtac impI 1); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   109
 by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   110
 by (asm_full_simp_tac (list_ss addsimps [reverse_Cons,hd_append,
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   111
                          (rev_red_not_nil RS mp)])  1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   112
qed"last_ind_on_first";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   113
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   114
val impl_ss = !simpset delsimps [reduce_Cons];
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   115
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   116
(* Main Lemma 1 for S_pkt in showing that reduce is refinement  *) 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   117
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   118
   "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then   \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   119
\      reduce(l@[x])=reduce(l) else                  \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   120
\      reduce(l@[x])=reduce(l)@[x]"; 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   121
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   122
by (rtac conjI 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   123
(* --> *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   124
by (List.list.induct_tac "l" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   125
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   126
by (case_tac "list=[]" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   127
 by (asm_full_simp_tac (!simpset addsimps [reverse_Nil,reverse_Cons]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   128
 by (rtac impI 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   129
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   130
by (cut_inst_tac [("l","list")] cons_not_nil 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   131
 by (asm_full_simp_tac impl_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   132
 by (REPEAT (etac exE 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   133
 by (hyp_subst_tac 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   134
by (asm_full_simp_tac (impl_ss addsimps [last_ind_on_first,l_iff_red_nil]) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   135
(* <-- *)
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   136
by (simp_tac (!simpset addsimps [and_de_morgan_and_absorbe,l_iff_red_nil]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   137
by (List.list.induct_tac "l" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   138
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   139
by (case_tac "list=[]" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   140
 by (asm_full_simp_tac (!simpset addsimps [reverse_Nil,reverse_Cons]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   141
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   142
 by (fast_tac HOL_cs 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   143
 by (rtac impI 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   144
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   145
by (cut_inst_tac [("l","list")] cons_not_nil 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   146
 by (asm_full_simp_tac impl_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   147
 by (REPEAT (etac exE 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   148
 by (hyp_subst_tac 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   149
by (asm_full_simp_tac (impl_ss addsimps [last_ind_on_first,l_iff_red_nil]) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   150
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   151
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   152
by (asm_full_simp_tac impl_ss 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   153
qed"reduce_hd";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   154
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   155
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   156
(* Main Lemma 2 for R_pkt in showing that reduce is refinement *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   157
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   158
  "!! s. [| s~=[] |] ==>  \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   159
\    if hd(s)=hd(tl(s)) & tl(s)~=[] then    \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   160
\      reduce(tl(s))=reduce(s) else      \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   161
\      reduce(tl(s))=tl(reduce(s))"; 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   162
by (cut_inst_tac [("l","s")] cons_not_nil 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   163
by (Asm_full_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   164
by (REPEAT (etac exE 1));
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   165
by (Asm_full_simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   166
by (rtac (expand_if RS ssubst) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   167
by (rtac conjI 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   168
by (simp_tac (!simpset addsimps [and_de_morgan_and_absorbe]) 2);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   169
by (ALLGOALS (EVERY'[rtac impI,etac conjE,cut_inst_tac [("l","xs")] cons_not_nil,Asm_full_simp_tac]));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   170
by (REPEAT (etac exE 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   171
by (REPEAT (etac exE 2));
1138
mueller
parents: 1050
diff changeset
   172
by (REPEAT(hyp_subst_tac 2));
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   173
by (ALLGOALS (Asm_full_simp_tac));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   174
val reduce_tl =result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   175
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   176
1138
mueller
parents: 1050
diff changeset
   177
(*******************************************************************
mueller
parents: 1050
diff changeset
   178
          C h a n n e l   A b s t r a c t i o n
mueller
parents: 1050
diff changeset
   179
 *******************************************************************)
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   180
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   181
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   182
      "is_weak_pmap reduce ch_ioa ch_fin_ioa";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   183
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   184
by (rtac conjI 1);
1138
mueller
parents: 1050
diff changeset
   185
(* --------------  start states ----------------- *)
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   186
by (Simp_tac 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1328
diff changeset
   187
by (rtac ballI 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   188
by (Asm_full_simp_tac 1);
1138
mueller
parents: 1050
diff changeset
   189
(* ---------------- main-part ------------------- *)
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   190
by (REPEAT (rtac allI 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   191
by (rtac imp_conj_lemma 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   192
by (act.induct_tac "a" 1);
1138
mueller
parents: 1050
diff changeset
   193
(* ----------------- 2 cases ---------------------*)
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   194
by (ALLGOALS (simp_tac (!simpset addsimps [externals_def])));
1138
mueller
parents: 1050
diff changeset
   195
(* fst case --------------------------------------*)
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   196
 by (rtac impI 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   197
 by (rtac disjI2 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   198
by (rtac reduce_hd 1);
1138
mueller
parents: 1050
diff changeset
   199
(* snd case --------------------------------------*)
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   200
 by (rtac impI 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   201
 by (REPEAT (etac conjE 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   202
 by (etac disjE 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   203
by (asm_full_simp_tac (!simpset addsimps [l_iff_red_nil]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   204
by (etac (hd_is_reduce_hd RS mp) 1); 
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   205
by (asm_full_simp_tac (!simpset addsimps [l_iff_red_nil]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   206
by (rtac conjI 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   207
by (etac (hd_is_reduce_hd RS mp) 1); 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   208
by (rtac (bool_if_impl_or RS mp) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   209
by (etac reduce_tl 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   210
qed"channel_abstraction";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   211
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   212
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   213
      "is_weak_pmap reduce srch_ioa srch_fin_ioa";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   214
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   215
 srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   216
qed"sender_abstraction";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   217
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   218
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   219
      "is_weak_pmap reduce rsch_ioa rsch_fin_ioa";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   220
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   221
 srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   222
qed"receiver_abstraction";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   223
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   224
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   225
(* 3 thms that do not hold generally! The lucky restriction here is 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   226
   the absence of internal actions. *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   227
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   228
      "is_weak_pmap (%id.id) sender_ioa sender_ioa";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   229
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   230
by (rtac conjI 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   231
(* start states *)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1328
diff changeset
   232
by (rtac ballI 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   233
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   234
(* main-part *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   235
by (REPEAT (rtac allI 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   236
by (rtac imp_conj_lemma 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   237
by (Action.action.induct_tac "a" 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   238
(* 7 cases *)
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   239
by (ALLGOALS (simp_tac ((!simpset addsimps [externals_def]) setloop (split_tac [expand_if]))));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   240
qed"sender_unchanged";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   241
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   242
(* 2 copies of before *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   243
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   244
      "is_weak_pmap (%id.id) receiver_ioa receiver_ioa";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   245
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   246
by (rtac conjI 1);
1138
mueller
parents: 1050
diff changeset
   247
 (* start states *)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1328
diff changeset
   248
by (rtac ballI 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   249
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   250
(* main-part *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   251
by (REPEAT (rtac allI 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   252
by (rtac imp_conj_lemma 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   253
by (Action.action.induct_tac "a" 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   254
(* 7 cases *)
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   255
by (ALLGOALS (simp_tac ((!simpset addsimps [externals_def]) setloop (split_tac [expand_if]))));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   256
qed"receiver_unchanged";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   257
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   258
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   259
      "is_weak_pmap (%id.id) env_ioa env_ioa";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   260
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   261
by (rtac conjI 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   262
(* start states *)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1328
diff changeset
   263
by (rtac ballI 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   264
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   265
(* main-part *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   266
by (REPEAT (rtac allI 1));
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   267
by (rtac imp_conj_lemma 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   268
by (Action.action.induct_tac "a" 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   269
(* 7 cases *)
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   270
by (ALLGOALS (simp_tac ((!simpset addsimps [externals_def]) setloop (split_tac [expand_if]))));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   271
qed"env_unchanged";
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   272
1532
769a4517ad7b Proof modification.
nipkow
parents: 1465
diff changeset
   273
Delsimps [Collect_False_empty];
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   274
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   275
goal Correctness.thy "compat_ioas srch_ioa rsch_ioa"; 
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   276
by (simp_tac (!simpset addsimps [compat_ioas_def,compat_asigs_def,Int_def,empty_def]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   277
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   278
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   279
by (ALLGOALS(Simp_tac));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   280
val compat_single_ch = result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   281
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   282
(* totally the same as before *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   283
goal Correctness.thy "compat_ioas srch_fin_ioa rsch_fin_ioa"; 
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   284
by (simp_tac (!simpset addsimps [compat_ioas_def,compat_asigs_def,Int_def,empty_def]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   285
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   286
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   287
by (ALLGOALS(Simp_tac));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   288
val compat_single_fin_ch = result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   289
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   290
val ss =
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   291
  !simpset delsimps ([trans_of_def, srch_asig_def,rsch_asig_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   292
                      asig_of_def, actions_def, srch_trans_def, rsch_trans_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   293
                      srch_ioa_def, srch_fin_ioa_def, rsch_fin_ioa_def, 
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   294
                      rsch_ioa_def, Sender.sender_trans_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   295
                      Receiver.receiver_trans_def] @ set_lemmas);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   296
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   297
goal Correctness.thy "compat_ioas receiver_ioa (srch_ioa || rsch_ioa)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   298
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   299
                           asig_of_par,asig_comp_def,actions_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   300
                           Int_def]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   301
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   302
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   303
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   304
by (ALLGOALS Simp_tac);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   305
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); 
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   306
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set])));
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   307
val compat_rec = result();
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   308
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   309
(* 5 proofs totally the same as before *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   310
goal Correctness.thy "compat_ioas receiver_ioa (srch_fin_ioa || rsch_fin_ioa)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   311
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   312
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   313
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   314
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   315
by (ALLGOALS Simp_tac);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   316
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); 
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   317
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set])));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   318
val compat_rec_fin =result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   319
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   320
goal Correctness.thy "compat_ioas sender_ioa \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   321
\      (receiver_ioa || srch_ioa || rsch_ioa)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   322
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   323
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   324
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   325
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   326
by (ALLGOALS(Simp_tac));
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   327
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); 
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   328
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set])));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   329
val compat_sen=result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   330
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   331
goal Correctness.thy "compat_ioas sender_ioa\
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   332
\      (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   333
by (simp_tac (ss addsimps [empty_def,  compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   334
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   335
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   336
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   337
by (ALLGOALS(Simp_tac));
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   338
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); 
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   339
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set])));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   340
val compat_sen_fin =result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   341
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   342
goal Correctness.thy "compat_ioas env_ioa\
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   343
\      (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   344
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   345
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   346
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   347
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   348
by (ALLGOALS(Simp_tac));
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   349
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def])));
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   350
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set])));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   351
val compat_env=result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   352
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   353
goal Correctness.thy "compat_ioas env_ioa\
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   354
\      (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   355
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   356
by (Simp_tac 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   357
by (rtac set_ext 1);
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   358
by (Action.action.induct_tac "x" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   359
by (ALLGOALS Simp_tac);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   360
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); 
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   361
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set])));
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   362
val compat_env_fin=result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   363
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   364
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   365
(* lemmata about externals of channels *)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   366
goal Correctness.thy 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   367
 "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) &  \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   368
\ externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   369
by (simp_tac (!simpset addsimps [externals_def]) 1);
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   370
val ext_single_ch = result();
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   371
1138
mueller
parents: 1050
diff changeset
   372
mueller
parents: 1050
diff changeset
   373
mueller
parents: 1050
diff changeset
   374
mueller
parents: 1050
diff changeset
   375
mueller
parents: 1050
diff changeset
   376
val ext_ss = [externals_of_par,ext_single_ch];
mueller
parents: 1050
diff changeset
   377
val compat_ss = [compat_single_ch,compat_single_fin_ch,compat_rec,
mueller
parents: 1050
diff changeset
   378
         compat_rec_fin,compat_sen,compat_sen_fin,compat_env,compat_env_fin];
mueller
parents: 1050
diff changeset
   379
val abstractions = [env_unchanged,sender_unchanged,
mueller
parents: 1050
diff changeset
   380
          receiver_unchanged,sender_abstraction,receiver_abstraction];
mueller
parents: 1050
diff changeset
   381
mueller
parents: 1050
diff changeset
   382
mueller
parents: 1050
diff changeset
   383
(************************************************************************
mueller
parents: 1050
diff changeset
   384
            S o u n d n e s s   o f   A b s t r a c t i o n        
mueller
parents: 1050
diff changeset
   385
*************************************************************************)
mueller
parents: 1050
diff changeset
   386
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   387
val ss = !simpset delsimps ([asig_of_def, srch_ioa_def, rsch_ioa_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   388
                             srch_fin_ioa_def, rsch_fin_ioa_def] @
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   389
                            impl_ioas @ env_ioas);
1138
mueller
parents: 1050
diff changeset
   390
mueller
parents: 1050
diff changeset
   391
goal Correctness.thy 
mueller
parents: 1050
diff changeset
   392
     "is_weak_pmap  abs  system_ioa  system_fin_ioa";
mueller
parents: 1050
diff changeset
   393
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   394
by (simp_tac (impl_ss delsimps ([srch_ioa_def, rsch_ioa_def, srch_fin_ioa_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   395
                                 rsch_fin_ioa_def] @ env_ioas @ impl_ioas)
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   396
                      addsimps [system_def, system_fin_def, abs_def, impl_def,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   397
                                impl_fin_def, sys_IOA, sys_fin_IOA]) 1);
1138
mueller
parents: 1050
diff changeset
   398
mueller
parents: 1050
diff changeset
   399
by (REPEAT (EVERY[rtac fxg_is_weak_pmap_of_product_IOA 1, 
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   400
                  simp_tac (ss addsimps abstractions) 1,
1138
mueller
parents: 1050
diff changeset
   401
                  rtac conjI 1]));
mueller
parents: 1050
diff changeset
   402
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1252
diff changeset
   403
by (ALLGOALS (simp_tac (ss addsimps ext_ss @ compat_ss))); 
1138
mueller
parents: 1050
diff changeset
   404
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
   405
qed "system_refinement";