IOA/example/Impl.ML
changeset 156 fd1be45b64bf
child 168 44ff2275d44f
equal deleted inserted replaced
155:722bf1319be5 156:fd1be45b64bf
       
     1 val impl_ioas =
       
     2   [Impl.impl_def,
       
     3    Sender.sender_ioa_def, 
       
     4    Receiver.receiver_ioa_def,
       
     5    Channels.srch_ioa_def,
       
     6    Channels.rsch_ioa_def];
       
     7 
       
     8 val transitions = [Sender.sender_trans_def, Receiver.receiver_trans_def,
       
     9                    Channels.srch_trans_def, Channels.rsch_trans_def];
       
    10 
       
    11 
       
    12 val impl_ss = merge_ss(action_ss,list_ss) 
       
    13               addcongs [let_weak_cong] 
       
    14               addsimps [Let_def, ioa_triple_proj, starts_of_par, trans_of_par4,
       
    15                        in_sender_asig, in_receiver_asig, in_srch_asig,
       
    16                        in_rsch_asig, count_addm_simp, count_delm_simp,
       
    17                        Multiset.countm_empty_def, Multiset.delm_empty_def,
       
    18                        (* Multiset.count_def, *) count_empty,
       
    19                        Packet.hdr_def, Packet.msg_def];
       
    20 
       
    21 goal Impl.thy
       
    22   "fst(x) = sen(x)            & \
       
    23 \  fst(snd(x)) = rec(x)       & \
       
    24 \  fst(snd(snd(x))) = srch(x) & \
       
    25 \  snd(snd(snd(x))) = rsch(x)";
       
    26 by(simp_tac (HOL_ss addsimps
       
    27         [Impl.sen_def,Impl.rec_def,Impl.srch_def,Impl.rsch_def]) 1);
       
    28 val impl_ss = impl_ss addsimps [result()];
       
    29 
       
    30 goal Impl.thy "a:actions(sender_asig)   \
       
    31 \            | a:actions(receiver_asig) \
       
    32 \            | a:actions(srch_asig)     \
       
    33 \            | a:actions(rsch_asig)";
       
    34   by(Action.action.induct_tac "a" 1);
       
    35   by(ALLGOALS(simp_tac impl_ss));
       
    36 val impl_ss = impl_ss addsimps [result()];
       
    37 
       
    38 
       
    39 (* Instantiation of a tautology? *)
       
    40 goal Packet.thy "!x. x = packet --> hdr(x) = hdr(packet)";
       
    41  by (simp_tac (HOL_ss addsimps [Packet.hdr_def]) 1);
       
    42 val eq_packet_imp_eq_hdr = result();
       
    43 
       
    44 
       
    45 (* INVARIANT 1 *)
       
    46 val ss = impl_ss addcongs [if_weak_cong] addsimps transitions;
       
    47 
       
    48 goalw Impl.thy impl_ioas "invariant(impl_ioa,inv1)";
       
    49 br invariantI 1;
       
    50 by(asm_full_simp_tac (impl_ss addsimps 
       
    51        [Impl.inv1_def, Impl.hdr_sum_def,
       
    52         Sender.srcvd_def, Sender.ssent_def,
       
    53         Receiver.rsent_def,Receiver.rrcvd_def]) 1);
       
    54 
       
    55 by(simp_tac (HOL_ss addsimps [fork_lemma,Impl.inv1_def]) 1);
       
    56 
       
    57 (* Split proof in two *)
       
    58 by (rtac conjI 1);
       
    59 
       
    60 (* First half *)
       
    61 by(asm_full_simp_tac (impl_ss addsimps [Impl.inv1_def]) 1);
       
    62 br Action.action.induct 1;
       
    63 
       
    64 val tac = asm_simp_tac (ss addcongs [conj_cong] 
       
    65                            addsimps [Suc_pred_lemma]
       
    66                            setloop (split_tac [expand_if]));
       
    67 
       
    68 by(EVERY1[tac, tac, tac, tac, tac, tac, tac, tac, tac, tac]);
       
    69 
       
    70 (* Now the other half *)
       
    71 by(asm_full_simp_tac (impl_ss addsimps [Impl.inv1_def]) 1);
       
    72 br Action.action.induct 1;
       
    73 by(EVERY1[tac, tac]);
       
    74 
       
    75 (* detour 1 *)
       
    76 by (tac 1);
       
    77 by (rtac impI 1);
       
    78 by (REPEAT (etac conjE 1));
       
    79 by (asm_simp_tac (impl_ss addsimps [Impl.hdr_sum_def, Multiset.count_def,
       
    80                                     Multiset.countm_nonempty_def]
       
    81                           setloop (split_tac [expand_if])) 1);
       
    82 (* detour 2 *)
       
    83 by (tac 1);
       
    84 by (rtac impI 1);
       
    85 by (REPEAT (etac conjE 1));
       
    86 by (asm_full_simp_tac (impl_ss addsimps 
       
    87                        [Impl.hdr_sum_def, 
       
    88                         Multiset.count_def,
       
    89                         Multiset.countm_nonempty_def,
       
    90                         Multiset.delm_nonempty_def,
       
    91                         left_plus_cancel,left_plus_cancel_inside_succ,
       
    92                         unzero_less]
       
    93                  setloop (split_tac [expand_if])) 1);
       
    94 by (rtac allI 1);
       
    95 by (rtac conjI 1);
       
    96 by (rtac impI 1);
       
    97 by (hyp_subst_tac 1);
       
    98 
       
    99 by (rtac (pred_suc RS mp RS sym RS iffD2) 1);
       
   100 by (dmatch_tac [less_leq_less RS mp] 1);
       
   101 by (cut_facts_tac [rewrite_rule[Packet.hdr_def]
       
   102                     eq_packet_imp_eq_hdr RS countm_props] 1);;
       
   103 by (dtac mp 1);
       
   104 by (assume_tac 1);
       
   105 by (assume_tac 1);
       
   106 
       
   107 by (rtac (countm_done_delm RS mp RS sym) 1);
       
   108 by (rtac refl 1);
       
   109 by (asm_simp_tac (HOL_ss addsimps [Multiset.count_def]) 1);
       
   110 
       
   111 by (rtac impI 1);
       
   112 by (asm_full_simp_tac (HOL_ss addsimps [neg_flip]) 1);
       
   113 by (hyp_subst_tac 1);
       
   114 by (rtac countm_spurious_delm 1);
       
   115 by (simp_tac HOL_ss 1);
       
   116 
       
   117 by (EVERY1[tac, tac, tac, tac, tac, tac]);
       
   118 
       
   119 val inv1 = result();
       
   120 
       
   121 
       
   122 
       
   123 (* INVARIANT 2 *)
       
   124 
       
   125   goal Impl.thy "invariant(impl_ioa, inv2)";
       
   126 
       
   127   by (rtac invariantI1 1); 
       
   128   (* Base case *)
       
   129   by (asm_full_simp_tac (impl_ss addsimps 
       
   130                     (Impl.inv2_def :: (receiver_projections 
       
   131                                        @ sender_projections @ impl_ioas))) 1);
       
   132 
       
   133   by (asm_simp_tac (impl_ss addsimps impl_ioas) 1);
       
   134   by (Action.action.induct_tac "a" 1);
       
   135 
       
   136   (* 10 cases. First 4 are simple, since state doesn't change wrt. invariant *)
       
   137   (* 10 *)
       
   138   by (asm_simp_tac (impl_ss addsimps (Impl.inv2_def::transitions)) 1);
       
   139   (* 9 *)
       
   140   by (asm_simp_tac (impl_ss addsimps (Impl.inv2_def::transitions)) 1);
       
   141   (* 8 *)
       
   142   by (asm_simp_tac (impl_ss addsimps (Impl.inv2_def::transitions)) 2);
       
   143   (* 7 *)
       
   144   by (asm_simp_tac (impl_ss addsimps (Impl.inv2_def::transitions)) 3);
       
   145   (* 6 *)
       
   146   by(forward_tac [rewrite_rule [Impl.inv1_def]
       
   147                                (inv1 RS invariantE) RS conjunct1] 1);
       
   148   by (asm_full_simp_tac (impl_ss addsimps [leq_imp_leq_suc,Impl.inv2_def]
       
   149                                  addsimps transitions) 1);
       
   150   (* 5 *)
       
   151   by (asm_full_simp_tac (impl_ss addsimps [leq_imp_leq_suc,Impl.inv2_def]
       
   152                                  addsimps transitions) 1);
       
   153   (* 4 *)
       
   154   by (forward_tac [rewrite_rule [Impl.inv1_def]
       
   155                                 (inv1 RS invariantE) RS conjunct1] 1);
       
   156   by (asm_full_simp_tac (impl_ss addsimps [Impl.inv2_def] 
       
   157                                  addsimps transitions) 1);
       
   158   by (fast_tac (HOL_cs addDs [plus_leD1,leD]) 1);
       
   159 
       
   160   (* 3 *)
       
   161   by (forward_tac [rewrite_rule [Impl.inv1_def] (inv1 RS invariantE)] 1);
       
   162 
       
   163   by (asm_full_simp_tac (impl_ss addsimps 
       
   164                          (Impl.inv2_def::transitions)) 1);
       
   165   by (fold_tac [rewrite_rule [Packet.hdr_def]Impl.hdr_sum_def]);
       
   166   by (fast_tac (HOL_cs addDs [plus_leD1,leD]) 1);
       
   167 
       
   168   (* 2 *)
       
   169   by (asm_full_simp_tac (impl_ss addsimps 
       
   170                          (Impl.inv2_def::transitions)) 1);
       
   171   by(forward_tac [rewrite_rule [Impl.inv1_def]
       
   172                                (inv1 RS invariantE) RS conjunct1] 1);
       
   173   by (rtac impI 1);
       
   174   by (rtac impI 1);
       
   175   by (REPEAT (etac conjE 1));
       
   176   by (dres_inst_tac [("k","count(rsch(s), ~ sbit(sen(s)))")] 
       
   177                      (standard(leq_add_leq RS mp)) 1);
       
   178   by (asm_full_simp_tac HOL_ss 1);
       
   179 
       
   180   (* 1 *)
       
   181   by (asm_full_simp_tac (impl_ss addsimps 
       
   182                          (Impl.inv2_def::transitions)) 1);
       
   183   by(forward_tac [rewrite_rule [Impl.inv1_def]
       
   184                                (inv1 RS invariantE) RS conjunct2] 1);
       
   185   by (rtac impI 1);
       
   186   by (rtac impI 1);
       
   187   by (REPEAT (etac conjE 1));
       
   188   by (fold_tac  [rewrite_rule[Packet.hdr_def]Impl.hdr_sum_def]);
       
   189   by (dres_inst_tac [("k","hdr_sum(srch(s), sbit(sen(s)))")] 
       
   190                      (standard(leq_add_leq RS mp)) 1);
       
   191   by (asm_full_simp_tac HOL_ss 1);
       
   192 val inv2 = result();
       
   193 
       
   194 
       
   195 (* INVARIANT 3 *)
       
   196 goal Impl.thy "invariant(impl_ioa, inv3)";
       
   197 
       
   198   by (rtac invariantI 1); 
       
   199   (* Base case *)
       
   200   by (asm_full_simp_tac (impl_ss addsimps 
       
   201                     (Impl.inv3_def :: (receiver_projections 
       
   202                                        @ sender_projections @ impl_ioas))) 1);
       
   203 
       
   204   by (asm_simp_tac (impl_ss addsimps impl_ioas) 1);
       
   205   by (Action.action.induct_tac "a" 1);
       
   206 
       
   207   (* 10 *)
       
   208   by (asm_full_simp_tac (impl_ss addsimps 
       
   209                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   210                   setloop (split_tac [expand_if])) 1);
       
   211 
       
   212   (* 9 *)
       
   213   by (asm_full_simp_tac (impl_ss addsimps 
       
   214                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   215                   setloop (split_tac [expand_if])) 1);
       
   216 
       
   217   (* 8 *)
       
   218   by (asm_full_simp_tac (impl_ss addsimps 
       
   219                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   220                   setloop (split_tac [expand_if])) 1);
       
   221   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   222   by (asm_full_simp_tac (HOL_ss addsimps [cons_imp_not_null]) 1);
       
   223   by (hyp_subst_tac 1);
       
   224   by (etac exE 1);
       
   225   by (asm_full_simp_tac list_ss 1);
       
   226 
       
   227   (* 7 *)
       
   228   by (asm_full_simp_tac (impl_ss addsimps 
       
   229       (Suc_pred_lemma::append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   230                   setloop (split_tac [expand_if])) 1);
       
   231   by (fast_tac HOL_cs 1);
       
   232 
       
   233   (* 6 *)
       
   234   by (asm_full_simp_tac (impl_ss addsimps 
       
   235                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   236                   setloop (split_tac [expand_if])) 1);
       
   237   (* 5 *)
       
   238   by (asm_full_simp_tac (impl_ss addsimps 
       
   239                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   240                   setloop (split_tac [expand_if])) 1);
       
   241 
       
   242   (* 4 *)
       
   243   by (asm_full_simp_tac (impl_ss addsimps 
       
   244                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   245                   setloop (split_tac [expand_if])) 1);
       
   246 
       
   247   (* 3 *)
       
   248   by (asm_full_simp_tac (impl_ss addsimps 
       
   249                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   250                   setloop (split_tac [expand_if])) 1);
       
   251 
       
   252   (* 2 *)
       
   253   by (asm_full_simp_tac (impl_ss addsimps transitions) 1);
       
   254   by (simp_tac (HOL_ss addsimps [Impl.inv3_def]) 1);
       
   255   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   256   by (rtac (imp_or_lem RS iffD2) 1);
       
   257   by (rtac impI 1);
       
   258   by (forward_tac [rewrite_rule [Impl.inv2_def] (inv2 RS invariantE)] 1);
       
   259   by (asm_full_simp_tac list_ss 1);
       
   260   by (REPEAT (etac conjE 1));
       
   261   by (res_inst_tac [("j","count(ssent(sen(s)),~ sbit(sen(s)))"),
       
   262                     ("k","count(rsent(rec(s)), sbit(sen(s)))")] le_trans 1);
       
   263   by (forward_tac [rewrite_rule [Impl.inv1_def]
       
   264                                 (inv1 RS invariantE) RS conjunct2] 1);
       
   265   by (asm_full_simp_tac (list_ss addsimps 
       
   266                          [Impl.hdr_sum_def, Multiset.count_def]) 1);
       
   267   by (rtac (less_eq_add_cong RS mp RS mp) 1);
       
   268   by (rtac countm_props 1);
       
   269   by (simp_tac (list_ss addsimps [Packet.hdr_def]) 1);
       
   270   by (rtac countm_props 1);
       
   271   by (simp_tac (list_ss addsimps [Packet.hdr_def]) 1);
       
   272   by (assume_tac 1);
       
   273 
       
   274 
       
   275   (* 1 *)
       
   276   by (asm_full_simp_tac (impl_ss addsimps 
       
   277                   (append_cons::not_hd_append::Impl.inv3_def::transitions)
       
   278                   setloop (split_tac [expand_if])) 1);
       
   279   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   280   by (rtac (imp_or_lem RS iffD2) 1);
       
   281   by (rtac impI 1);
       
   282   by (forward_tac [rewrite_rule [Impl.inv2_def] (inv2 RS invariantE)] 1);
       
   283   by (asm_full_simp_tac list_ss 1);
       
   284   by (REPEAT (etac conjE 1));
       
   285   by (dtac mp 1);
       
   286   by (assume_tac 1);
       
   287   by (etac allE 1);
       
   288   by (dtac (imp_or_lem RS iffD1) 1);
       
   289   by (dtac mp 1);
       
   290   by (assume_tac 1);
       
   291   by (assume_tac 1);
       
   292 val inv3 = result();
       
   293 
       
   294 
       
   295 
       
   296 (* INVARIANT 4 *)
       
   297 
       
   298 goal Impl.thy "invariant(impl_ioa, inv4)";
       
   299 
       
   300   by (rtac invariantI 1); 
       
   301   (* Base case *)
       
   302   by (asm_full_simp_tac (impl_ss addsimps 
       
   303                     (Impl.inv4_def :: (receiver_projections 
       
   304                                        @ sender_projections @ impl_ioas))) 1);
       
   305 
       
   306   by (asm_simp_tac (impl_ss addsimps impl_ioas) 1);
       
   307   by (Action.action.induct_tac "a" 1);
       
   308 
       
   309   (* 10 *)
       
   310   by (asm_full_simp_tac (impl_ss addsimps 
       
   311                          (append_cons::Impl.inv4_def::transitions)
       
   312                   setloop (split_tac [expand_if])) 1);
       
   313 
       
   314   (* 9 *)
       
   315   by (asm_full_simp_tac (impl_ss addsimps 
       
   316                          (append_cons::Impl.inv4_def::transitions)
       
   317                   setloop (split_tac [expand_if])) 1);
       
   318 
       
   319   (* 8 *)
       
   320   by (asm_full_simp_tac (impl_ss addsimps 
       
   321                          (append_cons::Impl.inv4_def::transitions)
       
   322                   setloop (split_tac [expand_if])) 1);
       
   323   (* 7 *)
       
   324   by (asm_full_simp_tac (impl_ss addsimps 
       
   325                          (append_cons::Impl.inv4_def::transitions)
       
   326                   setloop (split_tac [expand_if])) 1);
       
   327 
       
   328   (* 6 *)
       
   329   by (asm_full_simp_tac (impl_ss addsimps 
       
   330                          (append_cons::Impl.inv4_def::transitions)
       
   331                   setloop (split_tac [expand_if])) 1);
       
   332 
       
   333   (* 5 *)
       
   334   by (asm_full_simp_tac (impl_ss addsimps 
       
   335                          (append_cons::Impl.inv4_def::transitions)
       
   336                   setloop (split_tac [expand_if])) 1);
       
   337 
       
   338   (* 4 *)
       
   339   by (asm_full_simp_tac (impl_ss addsimps 
       
   340                          (append_cons::Impl.inv4_def::transitions)
       
   341                   setloop (split_tac [expand_if])) 1);
       
   342 
       
   343   (* 3 *)
       
   344   by (asm_full_simp_tac (impl_ss addsimps 
       
   345                          (append_cons::Impl.inv4_def::transitions)
       
   346                   setloop (split_tac [expand_if])) 1);
       
   347 
       
   348   (* 2 *)
       
   349   by (asm_full_simp_tac (impl_ss addsimps 
       
   350                          (append_cons::Impl.inv4_def::transitions)
       
   351                   setloop (split_tac [expand_if])) 1);
       
   352 
       
   353   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   354   by(forward_tac [rewrite_rule [Impl.inv2_def]
       
   355                                (inv2 RS invariantE)] 1);
       
   356  
       
   357   by (asm_full_simp_tac list_ss 1);
       
   358 
       
   359   (* 1 *)
       
   360   by (asm_full_simp_tac (impl_ss addsimps 
       
   361                          (append_cons::Impl.inv4_def::transitions)
       
   362                   setloop (split_tac [expand_if])) 1);
       
   363   by (strip_tac  1 THEN REPEAT (etac conjE 1));
       
   364   by (rtac ccontr 1);
       
   365   by(forward_tac [rewrite_rule [Impl.inv2_def]
       
   366                                (inv2 RS invariantE)] 1);
       
   367   by(forward_tac [rewrite_rule [Impl.inv3_def]
       
   368                                (inv3 RS invariantE)] 1);
       
   369   by (asm_full_simp_tac list_ss 1);
       
   370   by (eres_inst_tac [("x","m")] allE 1);
       
   371   by (dtac (less_leq_less RS mp RS mp) 1);
       
   372   by (dtac (left_add_leq RS mp) 1);
       
   373   by (asm_full_simp_tac list_ss 1);
       
   374   by (asm_full_simp_tac arith_ss 1);
       
   375 val inv4 = result();
       
   376 
       
   377 
       
   378 
       
   379 (* rebind them *)
       
   380 
       
   381 val inv1 = rewrite_rule [Impl.inv1_def] (inv1 RS invariantE);
       
   382 val inv2 = rewrite_rule [Impl.inv2_def] (inv2 RS invariantE);
       
   383 val inv3 = rewrite_rule [Impl.inv3_def] (inv3 RS invariantE);
       
   384 val inv4 = rewrite_rule [Impl.inv4_def] (inv4 RS invariantE);
       
   385