src/HOLCF/IOA/ABP/Correctness.ML
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 7958 f531589c9fc1
child 12218 6597093b77e7
permissions -rw-r--r--
final tuning;
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
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
    13
Delsimps [split_paired_All,Collect_empty_eq];
3072
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
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
(* auxiliary function *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    39
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
    40
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    41
(* lemmas about reduce *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    42
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
    43
Goal "(reduce(l)=[]) = (l=[])";  
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    44
by (induct_tac "l" 1);
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    45
by (auto_tac (claset(), simpset() addsplits [list.split]));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    46
val l_iff_red_nil = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    47
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
    48
Goal "s~=[] --> hd(s)=hd(reduce(s))";
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
    49
by (induct_tac "s" 1);
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    50
by (auto_tac (claset(), simpset() addsplits [list.split]));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    51
qed"hd_is_reduce_hd";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    52
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    53
(* to be used in the following Lemma *)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
    54
Goal "l~=[] --> reverse(reduce(l))~=[]";
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
    55
by (induct_tac "l" 1);
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    56
by (auto_tac (claset(), simpset() addsplits [list.split]));
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    57
qed_spec_mp "rev_red_not_nil";
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    58
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    59
(* shows applicability of the induction hypothesis of the following Lemma 1 *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    60
Goal "l~=[] ==> hd(reverse(reduce(a#l))) = hd(reverse(reduce(l)))";
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    61
 by (Simp_tac 1);
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    62
 by (asm_full_simp_tac (list_ss addsplits [list.split]
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    63
                                addsimps [reverse_Cons,hd_append,
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
    64
					  rev_red_not_nil])  1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    65
qed"last_ind_on_first";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    66
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
    67
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
    68
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    69
(* Main Lemma 1 for S_pkt in showing that reduce is refinement  *) 
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
    70
Goal 
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    71
   "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
    72
\      reduce(l@[x])=reduce(l) else                  \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    73
\      reduce(l@[x])=reduce(l)@[x]"; 
4833
2e53109d4bc8 Renamed expand_const -> split_const
nipkow
parents: 4681
diff changeset
    74
by (stac split_if 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    75
by (rtac conjI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    76
(* --> *)
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
    77
by (induct_tac "l" 1);
3072
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
by (case_tac "list=[]" 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
    80
 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
    81
 by (rtac impI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    82
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    83
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
    84
 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
    85
 by (REPEAT (etac exE 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    86
 by (hyp_subst_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    87
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
    88
(* <-- *)
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
    89
by (simp_tac (simpset() addsimps [and_de_morgan_and_absorbe,l_iff_red_nil]) 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
    90
by (induct_tac "l" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    91
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    92
by (case_tac "list=[]" 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    93
by (cut_inst_tac [("l","list")] cons_not_nil 2);
4833
2e53109d4bc8 Renamed expand_const -> split_const
nipkow
parents: 4681
diff changeset
    94
by (Asm_full_simp_tac 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
    95
by (auto_tac (claset(), 
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    96
	      impl_ss addsimps [last_ind_on_first,l_iff_red_nil] 
4833
2e53109d4bc8 Renamed expand_const -> split_const
nipkow
parents: 4681
diff changeset
    97
                      setloop split_tac [split_if]));
2e53109d4bc8 Renamed expand_const -> split_const
nipkow
parents: 4681
diff changeset
    98
by (Asm_full_simp_tac 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    99
qed"reduce_hd";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   100
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   101
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   102
(* Main Lemma 2 for R_pkt in showing that reduce is refinement *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   103
Goal "s~=[] ==>  \
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   104
\    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
   105
\      reduce(tl(s))=reduce(s) else      \
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   106
\      reduce(tl(s))=tl(reduce(s))"; 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   107
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
   108
by (Asm_full_simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   109
by (REPEAT (etac exE 1));
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   110
by (auto_tac (claset(), simpset() addsplits [list.split]));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   111
val reduce_tl =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   112
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
(*******************************************************************
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   115
          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
   116
 *******************************************************************)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   117
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   118
Delsplits [split_if]; 
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   119
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   120
Goal "is_weak_ref_map reduce ch_ioa ch_fin_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   121
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
   122
(* ---------------- main-part ------------------- *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   123
by (REPEAT (rtac allI 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   124
by (rtac imp_conj_lemma 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   125
by (induct_tac "a" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   126
(* ----------------- 2 cases ---------------------*)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   127
by (ALLGOALS (simp_tac (simpset() delcongs [if_weak_cong]
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   128
				  addsimps [externals_def])));
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   129
(* fst case --------------------------------------*)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   130
 by (rtac impI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   131
 by (rtac disjI2 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   132
by (rtac reduce_hd 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   133
(* snd case --------------------------------------*)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   134
 by (rtac impI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   135
 by (REPEAT (etac conjE 1));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   136
 by (etac disjE 1);
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   137
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
   138
by (etac (hd_is_reduce_hd RS mp) 1); 
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   139
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
   140
by (rtac conjI 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   141
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
   142
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
   143
by (etac reduce_tl 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   144
qed"channel_abstraction";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   145
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   146
Addsplits [split_if]; 
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   147
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   148
Goal "is_weak_ref_map reduce srch_ioa srch_fin_ioa";
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   149
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
   150
 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
   151
qed"sender_abstraction";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   152
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   153
Goal "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa";
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   154
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
   155
 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
   156
qed"receiver_abstraction";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   157
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   158
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   159
(* 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
   160
   the absence of internal actions. *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   161
Goal "is_weak_ref_map (%id. id) sender_ioa sender_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   162
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
   163
(* main-part *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   164
by (REPEAT (rtac allI 1));
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   165
by (induct_tac "a" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   166
(* 7 cases *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   167
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
   168
qed"sender_unchanged";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   169
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   170
(* 2 copies of before *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   171
Goal "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   172
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
   173
(* main-part *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   174
by (REPEAT (rtac allI 1));
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   175
by (induct_tac "a" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   176
(* 7 cases *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   177
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
   178
qed"receiver_unchanged";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   179
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   180
Goal "is_weak_ref_map (%id. id) env_ioa env_ioa";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   181
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
   182
(* main-part *)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   183
by (REPEAT (rtac allI 1));
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   184
by (induct_tac "a" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   185
(* 7 cases *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   186
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
   187
qed"env_unchanged";
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   188
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   189
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   190
Goal "compatible srch_ioa rsch_ioa"; 
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   191
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
   192
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   193
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   194
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   195
val compat_single_ch = result();
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
(* totally the same as before *)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   198
Goal "compatible srch_fin_ioa rsch_fin_ioa"; 
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   199
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
   200
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   201
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   202
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   203
val compat_single_fin_ch = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   204
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   205
val ss =
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   206
  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
   207
                      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
   208
                      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
   209
                      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
   210
                      Receiver.receiver_trans_def] @ set_lemmas);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   211
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   212
Goal "compatible receiver_ioa (srch_ioa || rsch_ioa)";
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   213
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   214
			   actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   215
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   216
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   217
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   218
by (ALLGOALS Simp_tac);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   219
val compat_rec = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   220
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   221
(* 5 proofs totally the same as before *)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   222
Goal "compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)";
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   223
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   224
			actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   225
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   226
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   227
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   228
by (ALLGOALS Simp_tac);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   229
val compat_rec_fin =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   230
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   231
Goal "compatible sender_ioa \
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   232
\      (receiver_ioa || srch_ioa || rsch_ioa)";
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   233
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   234
			actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   235
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   236
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   237
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   238
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   239
val compat_sen=result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   240
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   241
Goal "compatible sender_ioa\
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   242
\      (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   243
by (simp_tac (ss addsimps [empty_def,  compatible_def,asig_of_par,asig_comp_def,
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   244
				actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   245
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   246
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   247
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   248
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   249
val compat_sen_fin =result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   250
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   251
Goal "compatible env_ioa\
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   252
\      (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)";
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   253
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   254
				actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   255
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   256
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   257
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   258
by (ALLGOALS(Simp_tac));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   259
val compat_env=result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   260
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4833
diff changeset
   261
Goal "compatible env_ioa\
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   262
\      (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
7958
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   263
by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,
f531589c9fc1 added various little lemmas
oheimb
parents: 6916
diff changeset
   264
			actions_def,Int_def]) 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   265
by (Simp_tac 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   266
by (rtac set_ext 1);
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 5068
diff changeset
   267
by (induct_tac "x" 1);
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   268
by (ALLGOALS Simp_tac);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   269
val compat_env_fin=result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   270
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   271
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   272
(* lemmata about externals of channels *)
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   273
Goal "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) &  \
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   274
\ externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))";
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   275
   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
   276
val ext_single_ch = result();
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   277
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   278
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   279
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   280
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   281
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   282
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
   283
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
   284
         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
   285
val abstractions = [env_unchanged,sender_unchanged,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   286
          receiver_unchanged,sender_abstraction,receiver_abstraction];
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   287
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   288
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   289
(************************************************************************
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   290
            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
   291
*************************************************************************)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   292
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 4069
diff changeset
   293
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
   294
                             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
   295
                            impl_ioas @ env_ioas);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   296
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   297
(* 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
   298
        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
   299
6916
4957978b6f9e tidied proofs to cope with default if_weak_cong
paulson
parents: 5192
diff changeset
   300
Goal "is_weak_ref_map  abs  system_ioa  system_fin_ioa";
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   301
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   302
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
   303
                                 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
   304
                      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
   305
                                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
   306
                                sys_fin_IOA]) 1);
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   307
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   308
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
   309
                  simp_tac (ss addsimps abstractions) 1,
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   310
                  rtac conjI 1]));
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   311
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   312
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
   313
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   314
qed "system_refinement";
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   315
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
   316
3983
93ca73409df3 Deleted very odd Delsimps[Collect_False_empty] which made proofs fail!
nipkow
parents: 3842
diff changeset
   317
*)