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