Theory Correctness

Up to index of Isabelle/HOLCF/IOA/ABP

theory Correctness
imports Env Impl Impl_finite
begin

(*  Title:      HOLCF/IOA/ABP/Correctness.thy
    ID:         $Id: Correctness.thy,v 1.9 2007/10/21 12:21:50 wenzelm Exp $
    Author:     Olaf Müller
*)

header {* The main correctness proof: System_fin implements System *}

theory Correctness
imports IOA Env Impl Impl_finite
begin

consts
  reduce         :: "'a list => 'a list"
primrec
  reduce_Nil:  "reduce [] = []"
  reduce_Cons: "reduce(x#xs) =
                 (case xs of
                     [] => [x]
               |   y#ys => (if (x=y)
                              then reduce xs
                              else (x#(reduce xs))))"

definition
  abs where
    "abs  =
      (%p.(fst(p),(fst(snd(p)),(fst(snd(snd(p))),
       (reduce(fst(snd(snd(snd(p))))),reduce(snd(snd(snd(snd(p))))))))))"

definition
  system_ioa :: "('m action, bool * 'm impl_state)ioa" where
  "system_ioa = (env_ioa || impl_ioa)"

definition
  system_fin_ioa :: "('m action, bool * 'm impl_state)ioa" where
  "system_fin_ioa = (env_ioa || impl_fin_ioa)"


axiomatization where
  sys_IOA: "IOA system_ioa" and
  sys_fin_IOA: "IOA system_fin_ioa"



declare split_paired_All [simp del] Collect_empty_eq [simp del]

lemmas [simp] =
  srch_asig_def rsch_asig_def rsch_ioa_def srch_ioa_def ch_ioa_def
  ch_asig_def srch_actions_def rsch_actions_def rename_def rename_set_def asig_of_def
  actions_def exis_elim srch_trans_def rsch_trans_def ch_trans_def
  trans_of_def asig_projections set_lemmas

lemmas abschannel_fin [simp] =
  srch_fin_asig_def rsch_fin_asig_def
  rsch_fin_ioa_def srch_fin_ioa_def
  ch_fin_ioa_def ch_fin_trans_def ch_fin_asig_def

lemmas impl_ioas = sender_ioa_def receiver_ioa_def
  and impl_trans = sender_trans_def receiver_trans_def
  and impl_asigs = sender_asig_def receiver_asig_def

declare let_weak_cong [cong]
declare Let_def [simp] ioa_triple_proj [simp] starts_of_par [simp]

lemmas env_ioas = env_ioa_def env_asig_def env_trans_def
lemmas hom_ioas [simp] = env_ioas impl_ioas impl_trans impl_asigs asig_projections set_lemmas


subsection {* lemmas about reduce *}

lemma l_iff_red_nil: "(reduce l = []) = (l = [])"
  by (induct l) (auto split: list.split)

lemma hd_is_reduce_hd: "s ~= [] --> hd s = hd (reduce s)"
  by (induct s) (auto split: list.split)

text {* to be used in the following Lemma *}
lemma rev_red_not_nil [rule_format]:
    "l ~= [] --> reverse (reduce l) ~= []"
  by (induct l) (auto split: list.split)

text {* shows applicability of the induction hypothesis of the following Lemma 1 *}
lemma last_ind_on_first:
    "l ~= [] ==> hd (reverse (reduce (a # l))) = hd (reverse (reduce l))"
  apply simp
  apply (tactic {* auto_tac (claset(),
    HOL_ss addsplits [thm"list.split"]
    addsimps (thms "reverse.simps" @ [thm "hd_append", thm "rev_red_not_nil"])) *})
  done

text {* Main Lemma 1 for @{text "S_pkt"} in showing that reduce is refinement. *}
lemma reduce_hd:
   "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then
       reduce(l@[x])=reduce(l) else
       reduce(l@[x])=reduce(l)@[x]"
apply (simplesubst split_if)
apply (rule conjI)
txt {* @{text "-->"} *}
apply (induct_tac "l")
apply (simp (no_asm))
apply (case_tac "list=[]")
 apply (simp add: reverse.simps)
 apply (rule impI)
apply (simp (no_asm))
apply (cut_tac l = "list" in cons_not_nil)
 apply (simp del: reduce_Cons)
 apply (erule exE)+
 apply hypsubst
apply (simp del: reduce_Cons add: last_ind_on_first l_iff_red_nil)
txt {* @{text "<--"} *}
apply (simp (no_asm) add: and_de_morgan_and_absorbe l_iff_red_nil)
apply (induct_tac "l")
apply (simp (no_asm))
apply (case_tac "list=[]")
apply (cut_tac [2] l = "list" in cons_not_nil)
apply simp
apply (auto simp del: reduce_Cons simp add: last_ind_on_first l_iff_red_nil split: split_if)
apply simp
done


text {* Main Lemma 2 for R_pkt in showing that reduce is refinement. *}
lemma reduce_tl: "s~=[] ==>
     if hd(s)=hd(tl(s)) & tl(s)~=[] then
       reduce(tl(s))=reduce(s) else
       reduce(tl(s))=tl(reduce(s))"
apply (cut_tac l = "s" in cons_not_nil)
apply simp
apply (erule exE)+
apply (auto split: list.split)
done


subsection {* Channel Abstraction *}

declare split_if [split del]

lemma channel_abstraction: "is_weak_ref_map reduce ch_ioa ch_fin_ioa"
apply (simp (no_asm) add: is_weak_ref_map_def)
txt {* main-part *}
apply (rule allI)+
apply (rule imp_conj_lemma)
apply (induct_tac "a")
txt {* 2 cases *}
apply (simp_all (no_asm) cong del: if_weak_cong add: externals_def)
txt {* fst case *}
 apply (rule impI)
 apply (rule disjI2)
apply (rule reduce_hd)
txt {* snd case *}
 apply (rule impI)
 apply (erule conjE)+
 apply (erule disjE)
apply (simp add: l_iff_red_nil)
apply (erule hd_is_reduce_hd [THEN mp])
apply (simp add: l_iff_red_nil)
apply (rule conjI)
apply (erule hd_is_reduce_hd [THEN mp])
apply (rule bool_if_impl_or [THEN mp])
apply (erule reduce_tl)
done

declare split_if [split]

lemma sender_abstraction: "is_weak_ref_map reduce srch_ioa srch_fin_ioa"
apply (tactic {*
  simp_tac (HOL_ss addsimps [thm "srch_fin_ioa_def", thm "rsch_fin_ioa_def",
    thm "srch_ioa_def", thm "rsch_ioa_def", thm "rename_through_pmap",
    thm "channel_abstraction"]) 1 *})
done

lemma receiver_abstraction: "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa"
apply (tactic {*
  simp_tac (HOL_ss addsimps [thm "srch_fin_ioa_def", thm "rsch_fin_ioa_def",
    thm "srch_ioa_def", thm "rsch_ioa_def", thm "rename_through_pmap",
    thm "channel_abstraction"]) 1 *})
done


text {* 3 thms that do not hold generally! The lucky restriction here is
   the absence of internal actions. *}
lemma sender_unchanged: "is_weak_ref_map (%id. id) sender_ioa sender_ioa"
apply (simp (no_asm) add: is_weak_ref_map_def)
txt {* main-part *}
apply (rule allI)+
apply (induct_tac a)
txt {* 7 cases *}
apply (simp_all (no_asm) add: externals_def)
done

text {* 2 copies of before *}
lemma receiver_unchanged: "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa"
apply (simp (no_asm) add: is_weak_ref_map_def)
txt {* main-part *}
apply (rule allI)+
apply (induct_tac a)
txt {* 7 cases *}
apply (simp_all (no_asm) add: externals_def)
done

lemma env_unchanged: "is_weak_ref_map (%id. id) env_ioa env_ioa"
apply (simp (no_asm) add: is_weak_ref_map_def)
txt {* main-part *}
apply (rule allI)+
apply (induct_tac a)
txt {* 7 cases *}
apply (simp_all (no_asm) add: externals_def)
done


lemma compat_single_ch: "compatible srch_ioa rsch_ioa"
apply (simp add: compatible_def Int_def)
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done

text {* totally the same as before *}
lemma compat_single_fin_ch: "compatible srch_fin_ioa rsch_fin_ioa"
apply (simp add: compatible_def Int_def)
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done

lemmas del_simps = trans_of_def srch_asig_def rsch_asig_def
  asig_of_def actions_def srch_trans_def rsch_trans_def srch_ioa_def
  srch_fin_ioa_def rsch_fin_ioa_def rsch_ioa_def sender_trans_def
  receiver_trans_def set_lemmas

lemma compat_rec: "compatible receiver_ioa (srch_ioa || rsch_ioa)"
apply (simp del: del_simps
  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
apply simp
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done

text {* 5 proofs totally the same as before *}
lemma compat_rec_fin: "compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)"
apply (simp del: del_simps
  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
apply simp
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done

lemma compat_sen: "compatible sender_ioa
       (receiver_ioa || srch_ioa || rsch_ioa)"
apply (simp del: del_simps
  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
apply simp
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done

lemma compat_sen_fin: "compatible sender_ioa
       (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)"
apply (simp del: del_simps
  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
apply simp
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done

lemma compat_env: "compatible env_ioa
       (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)"
apply (simp del: del_simps
  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
apply simp
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done

lemma compat_env_fin: "compatible env_ioa
       (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)"
apply (simp del: del_simps
  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
apply simp
apply (rule set_ext)
apply (induct_tac x)
apply simp_all
done


text {* lemmata about externals of channels *}
lemma ext_single_ch: "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) &
    externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))"
  by (simp add: externals_def)


subsection {* Soundness of Abstraction *}

lemmas ext_simps = externals_of_par ext_single_ch
  and compat_simps = compat_single_ch compat_single_fin_ch compat_rec
    compat_rec_fin compat_sen compat_sen_fin compat_env compat_env_fin
  and abstractions = env_unchanged sender_unchanged
    receiver_unchanged sender_abstraction receiver_abstraction


(* FIX: this proof should be done with compositionality on trace level, not on
        weak_ref_map level, as done here with fxg_is_weak_ref_map_of_product_IOA

Goal "is_weak_ref_map  abs  system_ioa  system_fin_ioa"

by (simp_tac (impl_ss delsimps ([srch_ioa_def, rsch_ioa_def, srch_fin_ioa_def,
                                 rsch_fin_ioa_def] @ env_ioas @ impl_ioas)
                      addsimps [system_def, system_fin_def, abs_def,
                                impl_ioa_def, impl_fin_ioa_def, sys_IOA,
                                sys_fin_IOA]) 1);

by (REPEAT (EVERY[rtac fxg_is_weak_ref_map_of_product_IOA 1,
                  simp_tac (ss addsimps abstractions) 1,
                  rtac conjI 1]));

by (ALLGOALS (simp_tac (ss addsimps ext_ss @ compat_ss)));

qed "system_refinement";
*)

end

lemma

  srch_asig = asig_of srch_ioa
  rsch_asig = asig_of rsch_ioa
  rsch_ioa = rename ch_ioa rsch_actions
  srch_ioa = rename ch_ioa srch_actions
  ch_ioa = (ch_asig, {[]}, ch_trans, {}, {})
  ch_asig = (UN b. {S b}, UN b. {R b}, {})
  srch_actions akt =
  action_case None empty emptyp. Some (S p)) (λp. Some (R p)) empty empty akt
  rsch_actions akt =
  action_case None empty empty empty emptyb. Some (S b)) (λb. Some (R b)) akt
  rename ioa ren ==
  ((rename_set (inp ioa) ren, rename_set (out ioa) ren,
    rename_set (Automata.int ioa) ren),
   starts_of ioa,
   {tr. let s = fst tr; a = fst (snd tr); t = snd (snd tr)
        in ∃x. Some x = ren as -x--ioa-> t},
   {rename_set s ren |s. s ∈ wfair_of ioa},
   {rename_set s ren |s. s ∈ sfair_of ioa})
  rename_set A ren == {b. ∃x. Some x = ren bxA}
  asig_of == fst
  actions asig = inputs asigoutputs asiginternals asig
  (∃x. x = PQ x) = Q P
  srch_trans = trans_of srch_ioa
  rsch_trans = trans_of rsch_ioa
  ch_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of S b => t = st = s @ [b]
          | R b => s  [] ∧ b = hd s ∧ (t = st = tl s)}
  trans_of == fst o snd o snd
  inputs = fst
  outputs = fst o snd
  internals = snd o snd
  f x ∈ (UN x. {f x})
  f x y ∈ (UN x y. {f x y})
  x. a  f x ==> a  (UN x. {f x})
  x y. a  f x y ==> a  (UN x y. {f x y})

lemma abschannel_fin:

  srch_fin_asig = asig_of srch_fin_ioa
  rsch_fin_asig = asig_of rsch_fin_ioa
  rsch_fin_ioa = rename ch_fin_ioa rsch_actions
  srch_fin_ioa = rename ch_fin_ioa srch_actions
  ch_fin_ioa = (ch_fin_asig, {[]}, ch_fin_trans, {}, {})
  ch_fin_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of
          S b =>
            t = s ∨ (if b = hd (reverse s) ∧ s  [] then t = s else t = s @ [b])
          | R b => s  [] ∧ b = hd s ∧ (t = st = tl s)}
  ch_fin_asig = ch_asig

lemma impl_ioas:

  sender_ioa = (sender_asig, {([], True)}, sender_trans, {}, {})
  receiver_ioa = (receiver_asig, {([], False)}, receiver_trans, {}, {})

and impl_trans:

  sender_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of Next => if sq s = [] then t = s else False
          | S_msg m => sq t = sq s @ [m] ∧ sbit t = sbit s
          | S_pkt pkt =>
              sq s  [] ∧
              hdr pkt = sbit smsg pkt = hd (sq s) ∧ sq t = sq ssbit t = sbit s
          | R_ack b =>
              if b = sbit s then sq t = tl (sq s) ∧ sbit t = (¬ sbit s)
              else sq t = sq ssbit t = sbit s
          | _ => False}
  receiver_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of
          R_msg m =>
            rq s  [] ∧ m = hd (rq s) ∧ rq t = tl (rq s) ∧ rbit t = rbit s
          | R_pkt pkt =>
              if hdr pkt  rbit srq s = []
              then rq t = rq s @ [msg pkt] ∧ rbit t = (¬ rbit s)
              else rq t = rq srbit t = rbit s
          | S_ack b => b = rbit srq t = rq srbit t = rbit s | _ => False}

and impl_asigs:

  sender_asig = ((UN m. {S_msg m}) ∪ (UN b. {R_ack b}), UN pkt. {S_pkt pkt}, {})
  receiver_asig = (UN pkt. {R_pkt pkt}, (UN m. {R_msg m}) ∪ (UN b. {S_ack b}), {})

lemma env_ioas:

  env_ioa = (env_asig, {True}, env_trans, {}, {})
  env_asig == ({Next}, UN m. {S_msg m}, {})
  env_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of Next => t = True | S_msg m => s = True ∧ t = False
          | _ => False}

lemma hom_ioas:

  env_ioa = (env_asig, {True}, env_trans, {}, {})
  env_asig == ({Next}, UN m. {S_msg m}, {})
  env_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of Next => t = True | S_msg m => s = True ∧ t = False
          | _ => False}
  sender_ioa = (sender_asig, {([], True)}, sender_trans, {}, {})
  receiver_ioa = (receiver_asig, {([], False)}, receiver_trans, {}, {})
  sender_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of Next => if sq s = [] then t = s else False
          | S_msg m => sq t = sq s @ [m] ∧ sbit t = sbit s
          | S_pkt pkt =>
              sq s  [] ∧
              hdr pkt = sbit smsg pkt = hd (sq s) ∧ sq t = sq ssbit t = sbit s
          | R_ack b =>
              if b = sbit s then sq t = tl (sq s) ∧ sbit t = (¬ sbit s)
              else sq t = sq ssbit t = sbit s
          | _ => False}
  receiver_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of
          R_msg m =>
            rq s  [] ∧ m = hd (rq s) ∧ rq t = tl (rq s) ∧ rbit t = rbit s
          | R_pkt pkt =>
              if hdr pkt  rbit srq s = []
              then rq t = rq s @ [msg pkt] ∧ rbit t = (¬ rbit s)
              else rq t = rq srbit t = rbit s
          | S_ack b => b = rbit srq t = rq srbit t = rbit s | _ => False}
  sender_asig = ((UN m. {S_msg m}) ∪ (UN b. {R_ack b}), UN pkt. {S_pkt pkt}, {})
  receiver_asig = (UN pkt. {R_pkt pkt}, (UN m. {R_msg m}) ∪ (UN b. {S_ack b}), {})
  inputs = fst
  outputs = fst o snd
  internals = snd o snd
  f x ∈ (UN x. {f x})
  f x y ∈ (UN x y. {f x y})
  x. a  f x ==> a  (UN x. {f x})
  x y. a  f x y ==> a  (UN x y. {f x y})

lemmas about reduce

lemma l_iff_red_nil:

  (reduce l = []) = (l = [])

lemma hd_is_reduce_hd:

  s  [] --> hd s = hd (reduce s)

lemma rev_red_not_nil:

  l  [] ==> reverse (reduce l)  []

lemma last_ind_on_first:

  l  [] ==> hd (reverse (reduce (a # l))) = hd (reverse (reduce l))

lemma reduce_hd:

  if x = hd (reverse (reduce l)) ∧ reduce l  [] then reduce (l @ [x]) = reduce l
  else reduce (l @ [x]) = reduce l @ [x]

lemma reduce_tl:

  s  []
  ==> if hd s = hd (tl s) ∧ tl s  [] then reduce (tl s) = reduce s
      else reduce (tl s) = tl (reduce s)

Channel Abstraction

lemma channel_abstraction:

  is_weak_ref_map reduce ch_ioa ch_fin_ioa

lemma sender_abstraction:

  is_weak_ref_map reduce srch_ioa srch_fin_ioa

lemma receiver_abstraction:

  is_weak_ref_map reduce rsch_ioa rsch_fin_ioa

lemma sender_unchanged:

  is_weak_ref_mapid. id) sender_ioa sender_ioa

lemma receiver_unchanged:

  is_weak_ref_mapid. id) receiver_ioa receiver_ioa

lemma env_unchanged:

  is_weak_ref_mapid. id) env_ioa env_ioa

lemma compat_single_ch:

  compatible srch_ioa rsch_ioa

lemma compat_single_fin_ch:

  compatible srch_fin_ioa rsch_fin_ioa

lemma del_simps:

  trans_of == fst o snd o snd
  srch_asig = asig_of srch_ioa
  rsch_asig = asig_of rsch_ioa
  asig_of == fst
  actions asig = inputs asigoutputs asiginternals asig
  srch_trans = trans_of srch_ioa
  rsch_trans = trans_of rsch_ioa
  srch_ioa = rename ch_ioa srch_actions
  srch_fin_ioa = rename ch_fin_ioa srch_actions
  rsch_fin_ioa = rename ch_fin_ioa rsch_actions
  rsch_ioa = rename ch_ioa rsch_actions
  sender_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of Next => if sq s = [] then t = s else False
          | S_msg m => sq t = sq s @ [m] ∧ sbit t = sbit s
          | S_pkt pkt =>
              sq s  [] ∧
              hdr pkt = sbit smsg pkt = hd (sq s) ∧ sq t = sq ssbit t = sbit s
          | R_ack b =>
              if b = sbit s then sq t = tl (sq s) ∧ sbit t = (¬ sbit s)
              else sq t = sq ssbit t = sbit s
          | _ => False}
  receiver_trans =
  {tr. let s = fst tr; t = snd (snd tr)
       in case fst (snd tr) of
          R_msg m =>
            rq s  [] ∧ m = hd (rq s) ∧ rq t = tl (rq s) ∧ rbit t = rbit s
          | R_pkt pkt =>
              if hdr pkt  rbit srq s = []
              then rq t = rq s @ [msg pkt] ∧ rbit t = (¬ rbit s)
              else rq t = rq srbit t = rbit s
          | S_ack b => b = rbit srq t = rq srbit t = rbit s | _ => False}
  f x ∈ (UN x. {f x})
  f x y ∈ (UN x y. {f x y})
  x. a  f x ==> a  (UN x. {f x})
  x y. a  f x y ==> a  (UN x y. {f x y})

lemma compat_rec:

  compatible receiver_ioa (srch_ioa || rsch_ioa)

lemma compat_rec_fin:

  compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)

lemma compat_sen:

  compatible sender_ioa (receiver_ioa || srch_ioa || rsch_ioa)

lemma compat_sen_fin:

  compatible sender_ioa (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)

lemma compat_env:

  compatible env_ioa (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)

lemma compat_env_fin:

  compatible env_ioa (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)

lemma ext_single_ch:

  ext srch_fin_ioa = ext srch_ioaext rsch_fin_ioa = ext rsch_ioa

Soundness of Abstraction

lemma ext_simps:

  ext (A1.0 || A2.0) = ext A1.0ext A2.0
  ext srch_fin_ioa = ext srch_ioaext rsch_fin_ioa = ext rsch_ioa

and compat_simps:

  compatible srch_ioa rsch_ioa
  compatible srch_fin_ioa rsch_fin_ioa
  compatible receiver_ioa (srch_ioa || rsch_ioa)
  compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)
  compatible sender_ioa (receiver_ioa || srch_ioa || rsch_ioa)
  compatible sender_ioa (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)
  compatible env_ioa (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)
  compatible env_ioa (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)

and abstractions:

  is_weak_ref_mapid. id) env_ioa env_ioa
  is_weak_ref_mapid. id) sender_ioa sender_ioa
  is_weak_ref_mapid. id) receiver_ioa receiver_ioa
  is_weak_ref_map reduce srch_ioa srch_fin_ioa
  is_weak_ref_map reduce rsch_ioa rsch_fin_ioa