| author | wenzelm | 
| Wed, 26 Feb 2014 14:59:24 +0100 | |
| changeset 55767 | 96ddf9bf12ac | 
| parent 51703 | f2e92fc0c8aa | 
| child 58270 | 16648edf16e3 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/NTP/Correctness.thy | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 2 | Author: Tobias Nipkow & Konrad Slind | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | *) | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | |
| 17244 | 5 | header {* The main correctness proof: Impl implements Spec *}
 | 
| 6 | ||
| 7 | theory Correctness | |
| 8 | imports Impl Spec | |
| 9 | begin | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 10 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19739diff
changeset | 11 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19739diff
changeset | 12 | hom :: "'m impl_state => 'm list" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19739diff
changeset | 13 | "hom s = rq(rec(s)) @ (if rbit(rec s) = sbit(sen s) then sq(sen s) | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19739diff
changeset | 14 | else tl(sq(sen s)))" | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 15 | |
| 51703 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 wenzelm parents: 
42151diff
changeset | 16 | setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *}
 | 
| 19739 | 17 | |
| 18 | lemmas hom_ioas = Spec.ioa_def Spec.trans_def sender_trans_def receiver_trans_def impl_ioas | |
| 19 | and impl_asigs = sender_asig_def receiver_asig_def srch_asig_def rsch_asig_def | |
| 20 | ||
| 21 | declare split_paired_All [simp del] | |
| 22 | ||
| 23 | ||
| 24 | text {*
 | |
| 25 | A lemma about restricting the action signature of the implementation | |
| 26 | to that of the specification. | |
| 27 | *} | |
| 28 | ||
| 29 | lemma externals_lemma: | |
| 30 | "a:externals(asig_of(Automata.restrict impl_ioa (externals spec_sig))) = | |
| 31 | (case a of | |
| 32 | S_msg(m) => True | |
| 33 | | R_msg(m) => True | |
| 34 | | S_pkt(pkt) => False | |
| 35 | | R_pkt(pkt) => False | |
| 36 | | S_ack(b) => False | |
| 37 | | R_ack(b) => False | |
| 38 | | C_m_s => False | |
| 39 | | C_m_r => False | |
| 40 | | C_r_s => False | |
| 41 | | C_r_r(m) => False)" | |
| 42 | apply (simp (no_asm) add: externals_def restrict_def restrict_asig_def Spec.sig_def asig_projections) | |
| 43 | ||
| 44 | apply (induct_tac "a") | |
| 45 | apply (simp_all (no_asm) add: actions_def asig_projections) | |
| 46 |   txt {* 2 *}
 | |
| 47 | apply (simp (no_asm) add: impl_ioas) | |
| 48 | apply (simp (no_asm) add: impl_asigs) | |
| 49 | apply (simp (no_asm) add: asig_of_par asig_comp_def asig_projections) | |
| 35215 
a03462cbf86f
get rid of warnings about duplicate simp rules in all HOLCF theories
 huffman parents: 
35174diff
changeset | 50 | apply (simp (no_asm) add: "transitions"(1) unfold_renaming) | 
| 19739 | 51 |   txt {* 1 *}
 | 
| 52 | apply (simp (no_asm) add: impl_ioas) | |
| 53 | apply (simp (no_asm) add: impl_asigs) | |
| 54 | apply (simp (no_asm) add: asig_of_par asig_comp_def asig_projections) | |
| 55 | done | |
| 56 | ||
| 57 | lemmas sels = sbit_def sq_def ssending_def rbit_def rq_def rsending_def | |
| 58 | ||
| 59 | ||
| 60 | text {* Proof of correctness *}
 | |
| 61 | lemma ntp_correct: | |
| 62 | "is_weak_ref_map hom (Automata.restrict impl_ioa (externals spec_sig)) spec_ioa" | |
| 63 | apply (unfold Spec.ioa_def is_weak_ref_map_def) | |
| 64 | apply (simp (no_asm) cong del: if_weak_cong split del: split_if add: Correctness.hom_def | |
| 65 | cancel_restrict externals_lemma) | |
| 66 | apply (rule conjI) | |
| 67 | apply (simp (no_asm) add: hom_ioas) | |
| 68 | apply (simp (no_asm_simp) add: sels) | |
| 69 | apply (rule allI)+ | |
| 70 | apply (rule imp_conj_lemma) | |
| 71 | ||
| 72 | apply (induct_tac "a") | |
| 73 | apply (simp_all (no_asm_simp) add: hom_ioas) | |
| 74 | apply (frule inv4) | |
| 75 | apply force | |
| 76 | ||
| 77 | apply (frule inv4) | |
| 78 | apply (frule inv2) | |
| 79 | apply (erule disjE) | |
| 80 | apply (simp (no_asm_simp)) | |
| 81 | apply force | |
| 82 | ||
| 83 | apply (frule inv2) | |
| 84 | apply (erule disjE) | |
| 85 | ||
| 86 | apply (frule inv3) | |
| 87 | apply (case_tac "sq (sen (s))=[]") | |
| 88 | ||
| 89 | apply (simp add: hom_ioas) | |
| 90 | apply (blast dest!: add_leD1 [THEN leD]) | |
| 91 | ||
| 92 | apply (case_tac "m = hd (sq (sen (s)))") | |
| 93 | ||
| 94 | apply force | |
| 95 | ||
| 96 | apply simp | |
| 97 | apply (blast dest!: add_leD1 [THEN leD]) | |
| 98 | ||
| 99 | apply simp | |
| 100 | done | |
| 17244 | 101 | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 102 | end |