src/HOL/IOA/NTP/Impl.ML
changeset 3076 3e8d80cdd3e7
parent 3075 3c4fc62d494c
child 3077 750b766645b8
equal deleted inserted replaced
3075:3c4fc62d494c 3076:3e8d80cdd3e7
     1 (*  Title:      HOL/IOA/NTP/Impl.ML
       
     2     ID:         $Id$
       
     3     Author:     Tobias Nipkow & Konrad Slind
       
     4     Copyright   1994  TU Muenchen
       
     5 
       
     6 The implementation --- Invariants
       
     7 *)
       
     8 
       
     9 
       
    10 
       
    11 
       
    12 open Abschannel Impl;
       
    13 
       
    14 val impl_ioas =
       
    15   [Impl.impl_def,
       
    16    Sender.sender_ioa_def, 
       
    17    Receiver.receiver_ioa_def, 
       
    18    srch_ioa_thm RS eq_reflection,
       
    19    rsch_ioa_thm RS eq_reflection];
       
    20 
       
    21 val transitions = [Sender.sender_trans_def, Receiver.receiver_trans_def,
       
    22                    srch_trans_def, rsch_trans_def];
       
    23  
       
    24 
       
    25 Addsimps [ioa_triple_proj, starts_of_par, trans_of_par4,
       
    26           in_sender_asig, in_receiver_asig, in_srch_asig,
       
    27           in_rsch_asig];
       
    28 Addcongs [let_weak_cong];
       
    29 
       
    30 goal Impl.thy
       
    31   "fst(x) = sen(x)            & \
       
    32 \  fst(snd(x)) = rec(x)       & \
       
    33 \  fst(snd(snd(x))) = srch(x) & \
       
    34 \  snd(snd(snd(x))) = rsch(x)";
       
    35 by(simp_tac (!simpset addsimps
       
    36              [sen_def,rec_def,srch_def,rsch_def]) 1);
       
    37 Addsimps [result()];
       
    38 
       
    39 goal Impl.thy "a:actions(sender_asig)   \
       
    40 \            | a:actions(receiver_asig) \
       
    41 \            | a:actions(srch_asig)     \
       
    42 \            | a:actions(rsch_asig)";
       
    43   by(Action.action.induct_tac "a" 1);
       
    44   by(ALLGOALS (Simp_tac));
       
    45 Addsimps [result()];
       
    46 
       
    47 Delsimps [split_paired_All];
       
    48 
       
    49 
       
    50 (* Three Simp_sets in different sizes 
       
    51 ----------------------------------------------
       
    52 
       
    53 1) !simpset does not unfold the transition relations 
       
    54 2) ss unfolds transition relations 
       
    55 3) renname_ss unfolds transitions and the abstract channel *)
       
    56 
       
    57 
       
    58 val ss = (!simpset addcongs [if_weak_cong]
       
    59                    addsimps transitions);     
       
    60 val rename_ss = (ss addsimps unfold_renaming);
       
    61 
       
    62 
       
    63 
       
    64 val tac     = asm_simp_tac ((ss addcongs [conj_cong]) 
       
    65                             setloop (split_tac [expand_if]));
       
    66 val tac_ren = asm_simp_tac ((rename_ss addcongs [conj_cong]) 
       
    67                             setloop (split_tac [expand_if]));
       
    68 
       
    69 
       
    70 
       
    71 (* INVARIANT 1 *)
       
    72 
       
    73 goalw Impl.thy impl_ioas "invariant impl_ioa inv1";
       
    74 by (rtac invariantI 1);
       
    75 by(asm_full_simp_tac (!simpset
       
    76    addsimps [inv1_def, hdr_sum_def, Sender.srcvd_def,
       
    77              Sender.ssent_def, Receiver.rsent_def,Receiver.rrcvd_def]) 1);
       
    78 
       
    79 by(simp_tac (!simpset delsimps [trans_of_par4]
       
    80                 addsimps [fork_lemma,inv1_def]) 1);
       
    81 
       
    82 (* Split proof in two *)
       
    83 by (rtac conjI 1); 
       
    84 
       
    85 (* First half *)
       
    86 by(asm_full_simp_tac (!simpset addsimps [Impl.inv1_def]) 1);
       
    87 by (rtac Action.action.induct 1);
       
    88 
       
    89 by (EVERY1[tac, tac, tac, tac]);
       
    90 by (tac 1);
       
    91 by (tac_ren 1);
       
    92 
       
    93 (* 5 + 1 *)
       
    94 
       
    95 by (tac 1);
       
    96 by (tac_ren 1);
       
    97 
       
    98 (* 4 + 1 *)
       
    99 by(EVERY1[tac, tac, tac, tac]);
       
   100 
       
   101 
       
   102 (* Now the other half *)
       
   103 by(asm_full_simp_tac (!simpset addsimps [Impl.inv1_def]) 1);
       
   104 by (rtac Action.action.induct 1);
       
   105 by(EVERY1[tac, tac]);
       
   106 
       
   107 (* detour 1 *)
       
   108 by (tac 1);
       
   109 by (tac_ren 1);
       
   110 by (rtac impI 1);
       
   111 by (REPEAT (etac conjE 1));
       
   112 by (asm_simp_tac (!simpset addsimps [hdr_sum_def, Multiset.count_def,
       
   113                                Multiset.countm_nonempty_def]
       
   114                      setloop (split_tac [expand_if])) 1);
       
   115 (* detour 2 *)
       
   116 by (tac 1);
       
   117 by (tac_ren 1);
       
   118 by (rtac impI 1);
       
   119 by (REPEAT (etac conjE 1));
       
   120 by (asm_full_simp_tac (!simpset addsimps [Impl.hdr_sum_def, 
       
   121                                          Multiset.count_def,
       
   122                                          Multiset.countm_nonempty_def,
       
   123                                          Multiset.delm_nonempty_def,
       
   124                                          left_plus_cancel,
       
   125                                          left_plus_cancel_inside_succ,
       
   126                                          unzero_less]
       
   127                                setloop (split_tac [expand_if])) 1);
       
   128 by (rtac allI 1);
       
   129 by (rtac conjI 1);
       
   130 by (rtac impI 1);
       
   131 by (hyp_subst_tac 1);
       
   132 by (rtac (pred_suc RS mp RS sym RS iffD2) 1);
       
   133 by (dtac less_le_trans 1);
       
   134 by (cut_facts_tac [rewrite_rule[Packet.hdr_def]
       
   135                    eq_packet_imp_eq_hdr RS countm_props] 1);;
       
   136 by (assume_tac 1);
       
   137 by (assume_tac 1);
       
   138 
       
   139 by (rtac (countm_done_delm RS mp RS sym) 1);
       
   140 by (rtac refl 1);
       
   141 by (asm_simp_tac (!simpset addsimps [Multiset.count_def]) 1);
       
   142 
       
   143 by (rtac impI 1);
       
   144 by (asm_full_simp_tac (!simpset addsimps [neg_flip]) 1);
       
   145 by (hyp_subst_tac 1);
       
   146 by (rtac countm_spurious_delm 1);
       
   147 by (Simp_tac 1);
       
   148 
       
   149 by (EVERY1[tac, tac, tac, tac, tac, tac]);
       
   150 
       
   151 qed "inv1";
       
   152 
       
   153 
       
   154 
       
   155 (* INVARIANT 2 *)
       
   156 
       
   157   goal Impl.thy "invariant impl_ioa inv2";
       
   158 
       
   159   by (rtac invariantI1 1); 
       
   160   (* Base case *)
       
   161   by (asm_full_simp_tac (!simpset addsimps (inv2_def ::
       
   162                           (receiver_projections 
       
   163                            @ sender_projections @ impl_ioas)))
       
   164       1);
       
   165 
       
   166   by (asm_simp_tac (!simpset addsimps impl_ioas) 1);
       
   167   by (Action.action.induct_tac "a" 1);
       
   168 
       
   169   (* 10 cases. First 4 are simple, since state doesn't change *)
       
   170 
       
   171 val tac2 = asm_full_simp_tac (ss addsimps [inv2_def]);
       
   172 
       
   173   (* 10 - 7 *)
       
   174   by (EVERY1[tac2,tac2,tac2,tac2]);
       
   175   (* 6 *)
       
   176   by(forward_tac [rewrite_rule [Impl.inv1_def]
       
   177                                (inv1 RS invariantE) RS conjunct1] 1);
       
   178   (* 6 - 5 *)
       
   179   by (EVERY1[tac2,tac2]);
       
   180 
       
   181   (* 4 *)
       
   182   by (forward_tac [rewrite_rule [Impl.inv1_def]
       
   183                                 (inv1 RS invariantE) RS conjunct1] 1);
       
   184   by (tac2 1);
       
   185   by (fast_tac (!claset addDs [add_leD1,leD]) 1);
       
   186 
       
   187   (* 3 *)
       
   188   by (forward_tac [rewrite_rule [Impl.inv1_def] (inv1 RS invariantE)] 1);
       
   189 
       
   190   by (tac2 1);
       
   191   by (fold_tac [rewrite_rule [Packet.hdr_def]Impl.hdr_sum_def]);
       
   192   by (fast_tac (!claset addDs [add_leD1,leD]) 1);
       
   193 
       
   194   (* 2 *)
       
   195   by (tac2 1);
       
   196   by(forward_tac [rewrite_rule [Impl.inv1_def]
       
   197                                (inv1 RS invariantE) RS conjunct1] 1);
       
   198   by (rtac impI 1);
       
   199   by (rtac impI 1);
       
   200   by (REPEAT (etac conjE 1));
       
   201   by (dres_inst_tac [("k","count (rsch s) (~sbit(sen s))")] 
       
   202                      (standard(leq_add_leq RS mp)) 1);
       
   203   by (Asm_full_simp_tac 1);
       
   204 
       
   205   (* 1 *)
       
   206   by (tac2 1);
       
   207   by(forward_tac [rewrite_rule [Impl.inv1_def]
       
   208                                (inv1 RS invariantE) RS conjunct2] 1);
       
   209   by (rtac impI 1);
       
   210   by (rtac impI 1);
       
   211   by (REPEAT (etac conjE 1));
       
   212   by (fold_tac  [rewrite_rule[Packet.hdr_def]Impl.hdr_sum_def]);
       
   213   by (dres_inst_tac [("k","hdr_sum (srch s) (sbit(sen s))")] 
       
   214                      (standard(leq_add_leq RS mp)) 1);
       
   215   by (Asm_full_simp_tac 1);
       
   216 qed "inv2";
       
   217 
       
   218 
       
   219 (* INVARIANT 3 *)
       
   220 
       
   221 goal Impl.thy "invariant impl_ioa inv3";
       
   222 
       
   223   by (rtac invariantI 1); 
       
   224   (* Base case *)
       
   225   by (asm_full_simp_tac (!simpset addsimps 
       
   226                     (Impl.inv3_def :: (receiver_projections 
       
   227                                        @ sender_projections @ impl_ioas))) 1);
       
   228 
       
   229   by (asm_simp_tac (!simpset addsimps impl_ioas) 1);
       
   230   by (Action.action.induct_tac "a" 1);
       
   231 
       
   232 val tac3 = asm_full_simp_tac (ss addsimps [inv3_def] 
       
   233                               setloop (split_tac [expand_if]));
       
   234 
       
   235   (* 10 - 8 *)
       
   236 
       
   237   by (EVERY1[tac3,tac3,tac3]);
       
   238  
       
   239   by (tac_ren 1);
       
   240   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   241   by (hyp_subst_tac 1);
       
   242   by (etac exE 1);
       
   243   by (Asm_full_simp_tac 1);
       
   244 
       
   245   (* 7 *)
       
   246   by (tac3 1);
       
   247   by (tac_ren 1);
       
   248   by (Fast_tac 1);
       
   249 
       
   250   (* 6 - 3 *)
       
   251 
       
   252   by (EVERY1[tac3,tac3,tac3,tac3]);
       
   253 
       
   254   (* 2 *)
       
   255   by (asm_full_simp_tac ss 1);
       
   256   by (simp_tac (!simpset addsimps [inv3_def]) 1);
       
   257   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   258   by (rtac (imp_or_lem RS iffD2) 1);
       
   259   by (rtac impI 1);
       
   260   by (forward_tac [rewrite_rule [Impl.inv2_def] (inv2 RS invariantE)] 1);
       
   261   by (Asm_full_simp_tac 1);
       
   262   by (REPEAT (etac conjE 1));
       
   263   by (res_inst_tac [("j","count (ssent(sen s)) (~sbit(sen s))"),
       
   264                     ("k","count (rsent(rec s)) (sbit(sen s))")] le_trans 1);
       
   265   by (forward_tac [rewrite_rule [inv1_def]
       
   266                                 (inv1 RS invariantE) RS conjunct2] 1);
       
   267   by (asm_full_simp_tac (!simpset addsimps
       
   268                          [hdr_sum_def, Multiset.count_def]) 1);
       
   269   by (rtac (less_eq_add_cong RS mp RS mp) 1);
       
   270   by (rtac countm_props 1);
       
   271   by (Simp_tac 1);
       
   272   by (rtac countm_props 1);
       
   273   by (Simp_tac 1);
       
   274   by (assume_tac 1);
       
   275 
       
   276   (* 1 *)
       
   277   by (tac3 1);
       
   278   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   279   by (rtac (imp_or_lem RS iffD2) 1);
       
   280   by (rtac impI 1);
       
   281   by (forward_tac [rewrite_rule [Impl.inv2_def] (inv2 RS invariantE)] 1);
       
   282   by (Asm_full_simp_tac 1);
       
   283   by (REPEAT (etac conjE 1));
       
   284   by (dtac mp 1);
       
   285   by (assume_tac 1);
       
   286   by (etac allE 1);
       
   287   by (dtac (imp_or_lem RS iffD1) 1);
       
   288   by (dtac mp 1);
       
   289   by (assume_tac 1);
       
   290   by (assume_tac 1);
       
   291 qed "inv3";
       
   292 
       
   293 
       
   294 (* INVARIANT 4 *)
       
   295 
       
   296 goal Impl.thy "invariant impl_ioa inv4";
       
   297 
       
   298   by (rtac invariantI 1); 
       
   299   (* Base case *)
       
   300   by (asm_full_simp_tac (!simpset addsimps 
       
   301                     (Impl.inv4_def :: (receiver_projections 
       
   302                                        @ sender_projections @ impl_ioas))) 1);
       
   303 
       
   304   by (asm_simp_tac (!simpset addsimps impl_ioas) 1);
       
   305   by (Action.action.induct_tac "a" 1);
       
   306 
       
   307 val tac4 =  asm_full_simp_tac (ss addsimps [inv4_def]
       
   308                               setloop (split_tac [expand_if]));
       
   309 
       
   310   (* 10 - 2 *)
       
   311   
       
   312   by (EVERY1[tac4,tac4,tac4,tac4,tac4,tac4,tac4,tac4,tac4]);
       
   313  
       
   314  (* 2 b *)
       
   315  
       
   316   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   317   by(forward_tac [rewrite_rule [Impl.inv2_def]
       
   318                                (inv2 RS invariantE)] 1);
       
   319   by (tac4 1);
       
   320   by (Asm_full_simp_tac 1);
       
   321 
       
   322   (* 1 *)
       
   323   by (tac4 1);
       
   324   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   325   by (rtac ccontr 1);
       
   326   by(forward_tac [rewrite_rule [Impl.inv2_def]
       
   327                                (inv2 RS invariantE)] 1);
       
   328   by(forward_tac [rewrite_rule [Impl.inv3_def]
       
   329                                (inv3 RS invariantE)] 1);
       
   330   by (Asm_full_simp_tac 1);
       
   331   by (eres_inst_tac [("x","m")] allE 1);
       
   332   by (dtac less_le_trans 1);
       
   333   by (dtac (left_add_leq RS mp) 1);
       
   334   by (Asm_full_simp_tac 1);
       
   335   by (Asm_full_simp_tac 1);
       
   336 qed "inv4";
       
   337 
       
   338 
       
   339 (* rebind them *)
       
   340 
       
   341 val inv1 = rewrite_rule [Impl.inv1_def] (inv1 RS invariantE);
       
   342 val inv2 = rewrite_rule [Impl.inv2_def] (inv2 RS invariantE);
       
   343 val inv3 = rewrite_rule [Impl.inv3_def] (inv3 RS invariantE);
       
   344 val inv4 = rewrite_rule [Impl.inv4_def] (inv4 RS invariantE);