src/HOLCF/IOA/ABP/Correctness.ML
author nipkow
Fri, 06 Mar 1998 15:19:29 +0100
changeset 4681 a331c1f5a23e
parent 4477 b3e5857d8d99
child 4833 2e53109d4bc8
permissions -rw-r--r--
expand_if is now by default part of the simpset.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     1
 (*  Title:      HOLCF/IOA/ABP/Correctness.ML
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     2
    ID:         $Id$
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     3
    Author:     Olaf Mueller
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     4
    Copyright   1995  TU Muenchen
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     5
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     6
*)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     7
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     8
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3522
diff changeset
     9
goal Abschannel.thy "(? x. x=P & Q(x)) = Q(P)";
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    10
by (Fast_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    11
qed"exis_elim";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    12
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
Delsimps [split_paired_All];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    14
Addsimps 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    15
 ([srch_asig_def, rsch_asig_def, rsch_ioa_def, srch_ioa_def, ch_ioa_def, 
3522
a34c20f4bf44 changes neede for introducing fairness
mueller
parents: 3435
diff changeset
    16
   ch_asig_def, srch_actions_def, rsch_actions_def, rename_def, rename_set_def, asig_of_def, 
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    17
   actions_def, exis_elim, srch_trans_def, rsch_trans_def, ch_trans_def, 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    18
   trans_of_def] @ asig_projections @ set_lemmas);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    19
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    20
val abschannel_fin = [srch_fin_asig_def, rsch_fin_asig_def, 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    21
                      rsch_fin_ioa_def, srch_fin_ioa_def, 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    22
                      ch_fin_ioa_def,ch_fin_trans_def,ch_fin_asig_def];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    23
Addsimps abschannel_fin;
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    24
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    25
val impl_ioas =  [Sender.sender_ioa_def,Receiver.receiver_ioa_def];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    26
val impl_trans = [Sender.sender_trans_def,Receiver.receiver_trans_def];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    27
val impl_asigs = [Sender.sender_asig_def,Receiver.receiver_asig_def];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    28
Addcongs [let_weak_cong];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    29
Addsimps [Let_def, ioa_triple_proj, starts_of_par];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    30
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    31
val env_ioas = [Env.env_ioa_def,Env.env_asig_def,Env.env_trans_def];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    32
val hom_ioas = env_ioas @ impl_ioas @ impl_trans @ impl_asigs @ 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    33
               asig_projections @ set_lemmas;
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    34
Addsimps hom_ioas;
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    35
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    36
Addsimps [reduce_Nil,reduce_Cons];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    37
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    38
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    39
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    40
(* auxiliary function *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    41
fun rotate n i = EVERY(replicate n (etac revcut_rl i));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    42
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    43
(* lemmas about reduce *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    44
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    45
goal Correctness.thy "(reduce(l)=[]) = (l=[])";  
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    46
 by (rtac iffI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    47
 by (subgoal_tac "(l~=[]) --> (reduce(l)~=[])" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    48
 by (Fast_tac 1); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    49
 by (List.list.induct_tac "l" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    50
 by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    51
 by (Simp_tac 1);
4069
d6d06a03a2e9 expand_list_case -> split_list_case
nipkow
parents: 3983
diff changeset
    52
 by (rtac (split_list_case RS iffD2) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    53
 by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    54
 by (REPEAT (rtac allI 1)); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    55
 by (rtac impI 1); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    56
 by (hyp_subst_tac 1);
4423
a129b817b58a expandshort;
wenzelm
parents: 4098
diff changeset
    57
 by (stac expand_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    58
 by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    59
 by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    60
val l_iff_red_nil = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    61
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    62
goal Correctness.thy "s~=[] --> hd(s)=hd(reduce(s))";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    63
by (List.list.induct_tac "s" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    64
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    65
by (case_tac "list =[]" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    66
by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    67
(* main case *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    68
by (rotate 1 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    69
by (asm_full_simp_tac list_ss 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    70
by (Simp_tac 1);
4069
d6d06a03a2e9 expand_list_case -> split_list_case
nipkow
parents: 3983
diff changeset
    71
by (rtac (split_list_case RS iffD2) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    72
by (asm_full_simp_tac list_ss 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    73
by (REPEAT (rtac allI 1)); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    74
 by (rtac impI 1); 
4423
a129b817b58a expandshort;
wenzelm
parents: 4098
diff changeset
    75
by (stac expand_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    76
by (REPEAT(hyp_subst_tac 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    77
by (etac subst 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    78
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    79
qed"hd_is_reduce_hd";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    80
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    81
(* to be used in the following Lemma *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    82
goal Correctness.thy "l~=[] --> reverse(reduce(l))~=[]";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    83
by (List.list.induct_tac "l" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    84
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    85
by (case_tac "list =[]" 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
    86
by (asm_full_simp_tac (simpset() addsimps [reverse_Cons]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    87
(* main case *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    88
by (rotate 1 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    89
by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    90
by (cut_inst_tac [("l","list")] cons_not_nil 1); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    91
by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    92
by (REPEAT (etac exE 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    93
by (Asm_simp_tac 1);
4423
a129b817b58a expandshort;
wenzelm
parents: 4098
diff changeset
    94
by (stac expand_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    95
by (hyp_subst_tac 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
    96
by (asm_full_simp_tac (simpset() addsimps [reverse_Cons]) 1); 
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    97
qed"rev_red_not_nil";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    98
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    99
(* shows applicability of the induction hypothesis of the following Lemma 1 *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   100
goal Correctness.thy "!!l.[| l~=[] |] ==>   \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   101
\   hd(reverse(reduce(a#l))) = hd(reverse(reduce(l)))";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   102
 by (Simp_tac 1);
4069
d6d06a03a2e9 expand_list_case -> split_list_case
nipkow
parents: 3983
diff changeset
   103
 by (rtac (split_list_case RS iffD2) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   104
 by (asm_full_simp_tac list_ss 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   105
 by (REPEAT (rtac allI 1)); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   106
 by (rtac impI 1); 
4423
a129b817b58a expandshort;
wenzelm
parents: 4098
diff changeset
   107
 by (stac expand_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   108
 by (asm_full_simp_tac (list_ss addsimps [reverse_Cons,hd_append,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   109
                          (rev_red_not_nil RS mp)])  1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   110
qed"last_ind_on_first";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   111
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   112
val impl_ss = simpset() delsimps [reduce_Cons];
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   113
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   114
(* Main Lemma 1 for S_pkt in showing that reduce is refinement  *) 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   115
goal Correctness.thy 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   116
   "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then   \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   117
\      reduce(l@[x])=reduce(l) else                  \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   118
\      reduce(l@[x])=reduce(l)@[x]"; 
4423
a129b817b58a expandshort;
wenzelm
parents: 4098
diff changeset
   119
by (stac expand_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   120
by (rtac conjI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   121
(* --> *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   122
by (List.list.induct_tac "l" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   123
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   124
by (case_tac "list=[]" 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   125
 by (asm_full_simp_tac (simpset() addsimps [reverse_Nil,reverse_Cons]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   126
 by (rtac impI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   127
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   128
by (cut_inst_tac [("l","list")] cons_not_nil 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   129
 by (asm_full_simp_tac impl_ss 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   130
 by (REPEAT (etac exE 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   131
 by (hyp_subst_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   132
by (asm_full_simp_tac (impl_ss addsimps [last_ind_on_first,l_iff_red_nil]) 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   133
(* <-- *)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   134
by (simp_tac (simpset() addsimps [and_de_morgan_and_absorbe,l_iff_red_nil]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   135
by (List.list.induct_tac "l" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   136
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   137
by (case_tac "list=[]" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   138
by (cut_inst_tac [("l","list")] cons_not_nil 2);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   139
by (asm_full_simp_tac (simpset() setloop split_tac [expand_if]) 1);
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   140
by (auto_tac (claset(), 
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   141
	      impl_ss addsimps [last_ind_on_first,l_iff_red_nil] 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   142
                      setloop split_tac [expand_if]));
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   143
by (asm_full_simp_tac (simpset() setloop split_tac [expand_if]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   144
qed"reduce_hd";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   145
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   146
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   147
(* Main Lemma 2 for R_pkt in showing that reduce is refinement *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   148
goal Correctness.thy 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   149
  "!! s. [| s~=[] |] ==>  \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   150
\    if hd(s)=hd(tl(s)) & tl(s)~=[] then    \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   151
\      reduce(tl(s))=reduce(s) else      \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   152
\      reduce(tl(s))=tl(reduce(s))"; 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   153
by (cut_inst_tac [("l","s")] cons_not_nil 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   154
by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   155
by (REPEAT (etac exE 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   156
by (Asm_full_simp_tac 1);
4423
a129b817b58a expandshort;
wenzelm
parents: 4098
diff changeset
   157
by (stac expand_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   158
by (rtac conjI 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   159
by (simp_tac (simpset() addsimps [and_de_morgan_and_absorbe]) 2);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   160
by (Step_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   161
by (ALLGOALS (cut_inst_tac [("l","xs")] cons_not_nil));
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4423
diff changeset
   162
by (Auto_tac);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   163
val reduce_tl =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   164
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   165
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   166
(*******************************************************************
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   167
          C h a n n e l   A b s t r a c t i o n
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   168
 *******************************************************************)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   169
4681
a331c1f5a23e expand_if is now by default part of the simpset.
nipkow
parents: 4477
diff changeset
   170
Delsplits [expand_if];
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   171
goal Correctness.thy 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   172
      "is_weak_ref_map reduce ch_ioa ch_fin_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   173
by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   174
(* ---------------- main-part ------------------- *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   175
by (REPEAT (rtac allI 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   176
by (rtac imp_conj_lemma 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   177
by (abs_action.induct_tac "a" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   178
(* ----------------- 2 cases ---------------------*)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   179
by (ALLGOALS (simp_tac (simpset() addsimps [externals_def])));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   180
(* fst case --------------------------------------*)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   181
 by (rtac impI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   182
 by (rtac disjI2 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   183
by (rtac reduce_hd 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   184
(* snd case --------------------------------------*)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   185
 by (rtac impI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   186
 by (REPEAT (etac conjE 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   187
 by (etac disjE 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   188
by (asm_full_simp_tac (simpset() addsimps [l_iff_red_nil]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   189
by (etac (hd_is_reduce_hd RS mp) 1); 
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   190
by (asm_full_simp_tac (simpset() addsimps [l_iff_red_nil]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   191
by (rtac conjI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   192
by (etac (hd_is_reduce_hd RS mp) 1); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   193
by (rtac (bool_if_impl_or RS mp) 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   194
by (etac reduce_tl 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   195
qed"channel_abstraction";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   196
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   197
goal Correctness.thy 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   198
      "is_weak_ref_map reduce srch_ioa srch_fin_ioa";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   199
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   200
 srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   201
qed"sender_abstraction";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   202
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   203
goal Correctness.thy 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   204
      "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   205
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   206
 srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   207
qed"receiver_abstraction";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   208
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   209
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   210
(* 3 thms that do not hold generally! The lucky restriction here is 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   211
   the absence of internal actions. *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   212
goal Correctness.thy 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3522
diff changeset
   213
      "is_weak_ref_map (%id. id) sender_ioa sender_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   214
by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   215
by (TRY(
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   216
   (rtac conjI 1) THEN
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   217
(* start states *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   218
   (Fast_tac 1)));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   219
(* main-part *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   220
by (REPEAT (rtac allI 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   221
by (rtac imp_conj_lemma 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   222
by (Action.action.induct_tac "a" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   223
(* 7 cases *)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   224
by (ALLGOALS (simp_tac ((simpset() addsimps [externals_def]) setloop (split_tac [expand_if]))));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   225
qed"sender_unchanged";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   226
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   227
(* 2 copies of before *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   228
goal Correctness.thy 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3522
diff changeset
   229
      "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   230
by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   231
by (TRY(
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   232
   (rtac conjI 1) THEN
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   233
 (* start states *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   234
   (Fast_tac 1)));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   235
(* main-part *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   236
by (REPEAT (rtac allI 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   237
by (rtac imp_conj_lemma 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   238
by (Action.action.induct_tac "a" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   239
(* 7 cases *)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   240
by (ALLGOALS (simp_tac ((simpset() addsimps [externals_def]) setloop (split_tac [expand_if]))));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   241
qed"receiver_unchanged";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   242
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   243
goal Correctness.thy 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3522
diff changeset
   244
      "is_weak_ref_map (%id. id) env_ioa env_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   245
by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   246
by (TRY(
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   247
   (rtac conjI 1) THEN
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   248
(* start states *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   249
   (Fast_tac 1)));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   250
(* main-part *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   251
by (REPEAT (rtac allI 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   252
by (rtac imp_conj_lemma 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   253
by (Action.action.induct_tac "a" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   254
(* 7 cases *)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   255
by (ALLGOALS (simp_tac ((simpset() addsimps [externals_def]) setloop (split_tac [expand_if]))));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   256
qed"env_unchanged";
4681
a331c1f5a23e expand_if is now by default part of the simpset.
nipkow
parents: 4477
diff changeset
   257
Addsplits [expand_if];
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   258
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   259
goal Correctness.thy "compatible srch_ioa rsch_ioa"; 
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   260
by (simp_tac (simpset() addsimps [compatible_def,Int_def,empty_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   261
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   262
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   263
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   264
val compat_single_ch = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   265
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   266
(* totally the same as before *)
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   267
goal Correctness.thy "compatible srch_fin_ioa rsch_fin_ioa"; 
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   268
by (simp_tac (simpset() addsimps [compatible_def,Int_def,empty_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   269
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   270
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   271
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   272
val compat_single_fin_ch = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   273
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   274
val ss =
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   275
  simpset() delsimps ([trans_of_def, srch_asig_def,rsch_asig_def,
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   276
                      asig_of_def, actions_def, srch_trans_def, rsch_trans_def,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   277
                      srch_ioa_def, srch_fin_ioa_def, rsch_fin_ioa_def, 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   278
                      rsch_ioa_def, Sender.sender_trans_def,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   279
                      Receiver.receiver_trans_def] @ set_lemmas);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   280
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   281
goal Correctness.thy "compatible receiver_ioa (srch_ioa || rsch_ioa)";
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   282
by (simp_tac (ss addsimps [empty_def,compatible_def,
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   283
                           asig_of_par,asig_comp_def,actions_def,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   284
                           Int_def]) 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   285
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   286
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   287
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   288
by (ALLGOALS Simp_tac);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   289
val compat_rec = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   290
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   291
(* 5 proofs totally the same as before *)
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   292
goal Correctness.thy "compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)";
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   293
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   294
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   295
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   296
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   297
by (ALLGOALS Simp_tac);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   298
val compat_rec_fin =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   299
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   300
goal Correctness.thy "compatible sender_ioa \
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   301
\      (receiver_ioa || srch_ioa || rsch_ioa)";
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   302
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   303
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   304
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   305
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   306
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   307
val compat_sen=result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   308
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   309
goal Correctness.thy "compatible sender_ioa\
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   310
\      (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   311
by (simp_tac (ss addsimps [empty_def,  compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   312
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   313
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   314
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   315
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   316
val compat_sen_fin =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   317
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   318
goal Correctness.thy "compatible env_ioa\
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   319
\      (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)";
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   320
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   321
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   322
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   323
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   324
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   325
val compat_env=result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   326
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   327
goal Correctness.thy "compatible env_ioa\
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   328
\      (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
3435
ef4fe2a77e01 changed compatible definition;
mueller
parents: 3072
diff changeset
   329
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   330
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   331
by (rtac set_ext 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   332
by (Action.action.induct_tac "x" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   333
by (ALLGOALS Simp_tac);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   334
val compat_env_fin=result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   335
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   336
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   337
(* lemmata about externals of channels *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   338
goal Correctness.thy 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   339
 "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) &  \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   340
\ externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   341
   by (simp_tac (simpset() addsimps [externals_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   342
val ext_single_ch = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   343
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   344
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   345
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   346
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   347
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   348
val ext_ss = [externals_of_par,ext_single_ch];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   349
val compat_ss = [compat_single_ch,compat_single_fin_ch,compat_rec,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   350
         compat_rec_fin,compat_sen,compat_sen_fin,compat_env,compat_env_fin];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   351
val abstractions = [env_unchanged,sender_unchanged,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   352
          receiver_unchanged,sender_abstraction,receiver_abstraction];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   353
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   354
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   355
(************************************************************************
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   356
            S o u n d n e s s   o f   A b s t r a c t i o n        
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   357
*************************************************************************)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   358
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   359
val ss = simpset() delsimps ([asig_of_def, srch_ioa_def, rsch_ioa_def,
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   360
                             srch_fin_ioa_def, rsch_fin_ioa_def] @
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   361
                            impl_ioas @ env_ioas);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   362
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   363
(* FIX: this proof should be done with compositionality on trace level, not on 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   364
        weak_ref_map level, as done here with fxg_is_weak_ref_map_of_product_IOA 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   365
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   366
goal Correctness.thy 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   367
     "is_weak_ref_map  abs  system_ioa  system_fin_ioa";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   368
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   369
by (simp_tac (impl_ss delsimps ([srch_ioa_def, rsch_ioa_def, srch_fin_ioa_def,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   370
                                 rsch_fin_ioa_def] @ env_ioas @ impl_ioas)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   371
                      addsimps [system_def, system_fin_def, abs_def,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   372
                                impl_ioa_def, impl_fin_ioa_def, sys_IOA,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   373
                                sys_fin_IOA]) 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   374
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   375
by (REPEAT (EVERY[rtac fxg_is_weak_ref_map_of_product_IOA 1, 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   376
                  simp_tac (ss addsimps abstractions) 1,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   377
                  rtac conjI 1]));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   378
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   379
by (ALLGOALS (simp_tac (ss addsimps ext_ss @ compat_ss))); 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   380
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   381
qed "system_refinement";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   382
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   383
3983
93ca73409df3 Deleted very odd Delsimps[Collect_False_empty] which made proofs fail!
nipkow
parents: 3842
diff changeset
   384
*)