Theory Correctness

Up to index of Isabelle/HOLCF/IOA/NTP

theory Correctness
imports Impl Spec
begin

(*  Title:      HOL/IOA/NTP/Correctness.thy
    ID:         $Id: Correctness.thy,v 1.9 2008/03/19 21:47:36 wenzelm Exp $
    Author:     Tobias Nipkow & Konrad Slind
*)

header {* The main correctness proof: Impl implements Spec *}

theory Correctness
imports Impl Spec
begin

definition
  hom :: "'m impl_state => 'm list" where
  "hom s = rq(rec(s)) @ (if rbit(rec s) = sbit(sen s) then sq(sen s)
                         else tl(sq(sen s)))"

declaration {* fn _ =>
  (* repeated from Traces.ML *)
  Classical.map_cs (fn cs => cs delSWrapper "split_all_tac")
*}

lemmas hom_ioas = Spec.ioa_def Spec.trans_def sender_trans_def receiver_trans_def impl_ioas
  and impl_asigs = sender_asig_def receiver_asig_def srch_asig_def rsch_asig_def

declare split_paired_All [simp del]


text {*
  A lemma about restricting the action signature of the implementation
  to that of the specification.
*}

lemma externals_lemma: 
 "a:externals(asig_of(Automata.restrict impl_ioa (externals spec_sig))) =  
  (case a of                   
      S_msg(m) => True         
    | R_msg(m) => True         
    | S_pkt(pkt) => False   
    | R_pkt(pkt) => False   
    | S_ack(b) => False     
    | R_ack(b) => False     
    | C_m_s => False           
    | C_m_r => False           
    | C_r_s => False           
    | C_r_r(m) => False)"
 apply (simp (no_asm) add: externals_def restrict_def restrict_asig_def Spec.sig_def asig_projections)

  apply (induct_tac "a")
  apply (simp_all (no_asm) add: actions_def asig_projections)
  txt {* 2 *}
  apply (simp (no_asm) add: impl_ioas)
  apply (simp (no_asm) add: impl_asigs)
  apply (simp (no_asm) add: asig_of_par asig_comp_def asig_projections)
  apply (simp (no_asm) add: "transitions" unfold_renaming)
  txt {* 1 *}
  apply (simp (no_asm) add: impl_ioas)
  apply (simp (no_asm) add: impl_asigs)
  apply (simp (no_asm) add: asig_of_par asig_comp_def asig_projections)
  done

lemmas sels = sbit_def sq_def ssending_def rbit_def rq_def rsending_def


text {* Proof of correctness *}
lemma ntp_correct:
  "is_weak_ref_map hom (Automata.restrict impl_ioa (externals spec_sig)) spec_ioa"
apply (unfold Spec.ioa_def is_weak_ref_map_def)
apply (simp (no_asm) cong del: if_weak_cong split del: split_if add: Correctness.hom_def
  cancel_restrict externals_lemma)
apply (rule conjI)
 apply (simp (no_asm) add: hom_ioas)
 apply (simp (no_asm_simp) add: sels)
apply (rule allI)+
apply (rule imp_conj_lemma)

apply (induct_tac "a")
apply (simp_all (no_asm_simp) add: hom_ioas)
apply (frule inv4)
apply force

apply (frule inv4)
apply (frule inv2)
apply (erule disjE)
apply (simp (no_asm_simp))
apply force

apply (frule inv2)
apply (erule disjE)

apply (frule inv3)
apply (case_tac "sq (sen (s))=[]")

apply (simp add: hom_ioas)
apply (blast dest!: add_leD1 [THEN leD])

apply (case_tac "m = hd (sq (sen (s)))")

apply force

apply simp
apply (blast dest!: add_leD1 [THEN leD])

apply simp
done

end

lemma hom_ioas:

  spec_ioa == (spec_sig, {[]}, spec_trans, {}, {})
  spec_trans ==
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of S_msg m => t = s @ [m] | R_msg m => s = m # t
          | _ => False}
  sender_trans ==
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of
          S_msg m =>
            sq t = sq s @ [m] ∧
            ssent t = ssent s ∧
            srcvd t = srcvd s ∧ sbit t = sbit s ∧ ssending t = ssending s
          | S_pkt pkt =>
              hdr pkt = sbit s ∧
              (∃Q. sq s = msg pkt # Q) ∧
              sq t = sq s ∧
              ssent t = addm (ssent s) (sbit s) ∧
              srcvd t = srcvd s ∧ sbit t = sbit s ∧ ssending s ∧ ssending t
          | R_ack b =>
              sq t = sq s ∧
              ssent t = ssent s ∧
              srcvd t = addm (srcvd s) b ∧
              sbit t = sbit s ∧ ssending t = ssending s
          | C_m_s =>
              count (ssent s) (¬ sbit s) < count (srcvd s) (¬ sbit s) ∧
              sq t = sq s ∧
              ssent t = ssent s ∧
              srcvd t = srcvd s ∧ sbit t = sbit s ∧ ssending s ∧ ¬ ssending t
          | C_r_s =>
              count (ssent s) (sbit s)  count (srcvd s) (¬ sbit s) ∧
              sq t = tl (sq s) ∧
              ssent t = ssent s ∧
              srcvd t = srcvd s ∧ sbit t = (¬ sbit s) ∧ ¬ ssending s ∧ ssending t
          | _ => False}
  receiver_trans ==
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of
          R_msg m =>
            rq s = m # rq t ∧
            rsent t = rsent s ∧
            rrcvd t = rrcvd s ∧ rbit t = rbit s ∧ rsending t = rsending s
          | R_pkt pkt =>
              rq t = rq s ∧
              rsent t = rsent s ∧
              rrcvd t = addm (rrcvd s) pkt ∧
              rbit t = rbit s ∧ rsending t = rsending s
          | S_ack b =>
              b = rbit s ∧
              rq t = rq s ∧
              rsent t = addm (rsent s) (rbit s) ∧
              rrcvd t = rrcvd s ∧ rbit t = rbit s ∧ rsending s ∧ rsending t
          | C_m_r =>
              count (rsent s) (¬ rbit s) < countm (rrcvd s) (λy. hdr y = rbit s) ∧
              rq t = rq s ∧
              rsent t = rsent s ∧
              rrcvd t = rrcvd s ∧ rbit t = rbit s ∧ rsending s ∧ ¬ rsending t
          | C_r_r m =>
              count (rsent s) (rbit s)  countm (rrcvd s) (λy. hdr y = rbit s) ∧
              count (rsent s) (¬ rbit s) < count (rrcvd s) (rbit s, m) ∧
              rq t = rq s @ [m] ∧
              rsent t = rsent s ∧
              rrcvd t = rrcvd s ∧ rbit t = (¬ rbit s) ∧ ¬ rsending s ∧ rsending t
          | _ => False}
  impl_ioa == sender_ioa || receiver_ioa || srch_ioa || rsch_ioa
  sender_ioa == (sender_asig, {([], {|}, {|}, False, True)}, sender_trans, {}, {})
  receiver_ioa ==
  (receiver_asig, {([], {|}, {|}, False, False)}, receiver_trans, {}, {})
  srch_ioa == (srch_asig, {{|}}, srch_trans, srch_wfair, srch_sfair)
  rsch_ioa == (rsch_asig, {{|}}, rsch_trans, rsch_wfair, rsch_sfair)

and impl_asigs:

  sender_asig ==
  ((UN m. {S_msg m}) ∪ (UN b. {R_ack b}), UN pkt. {S_pkt pkt}, {C_m_s, C_r_s})
  receiver_asig ==
  (UN pkt. {R_pkt pkt}, (UN m. {R_msg m}) ∪ (UN b. {S_ack b}),
   insert C_m_r (UN m. {C_r_r m}))
  srch_asig == asig_of srch_ioa
  rsch_asig == asig_of rsch_ioa

lemma externals_lemma:

  (aext (restrict impl_ioa (externals spec_sig))) =
  (case a of S_msg m => True | R_msg m => True | _ => False)

lemma sels:

  sbit == fst o snd o snd o snd
  sq == fst
  ssending == snd o snd o snd o snd
  rbit == fst o snd o snd o snd
  rq == fst
  rsending == snd o snd o snd o snd

lemma ntp_correct:

  is_weak_ref_map hom (restrict impl_ioa (externals spec_sig)) spec_ioa