Theory Refinement
section ‹Models, Invariants and Refinements›
theory Refinement imports Infra
begin
subsection ‹Specifications, reachability, and behaviours.›
text ‹Transition systems are multi-pointed graphs.›
record 's TS =
init :: "'s set"
trans :: "('s × 's) set"
text ‹The inductive set of reachable states.›
inductive_set
reach :: "('s, 'a) TS_scheme ⇒ 's set"
for T :: "('s, 'a) TS_scheme"
where
r_init [intro]: "s ∈ init T ⟹ s ∈ reach T"
| r_trans [intro]: "⟦ (s, t) ∈ trans T; s ∈ reach T ⟧ ⟹ t ∈ reach T"
subsubsection ‹Finite behaviours›
text ‹Note that behaviours grow at the head of the list, i.e., the initial
state is at the end.›
inductive_set
beh :: "('s, 'a) TS_scheme ⇒ ('s list) set"
for T :: "('s, 'a) TS_scheme"
where
b_empty [iff]: "[] ∈ beh T"
| b_init [intro]: "s ∈ init T ⟹ [s] ∈ beh T"
| b_trans [intro]: "⟦ s # b ∈ beh T; (s, t) ∈ trans T ⟧ ⟹ t # s # b ∈ beh T"
inductive_cases beh_non_empty: "s # b ∈ beh T"
text ‹Behaviours are prefix closed.›
lemma beh_immediate_prefix_closed:
"s # b ∈ beh T ⟹ b ∈ beh T"
by (erule beh_non_empty, auto)
lemma beh_prefix_closed:
"c @ b ∈ beh T ⟹ b ∈ beh T"
by (induct c, auto dest!: beh_immediate_prefix_closed)
text ‹States in behaviours are exactly reachable.›
lemma beh_in_reach [rule_format]:
"b ∈ beh T ⟹ (∀s ∈ set b. s ∈ reach T)"
by (erule beh.induct) (auto)
lemma reach_in_beh:
assumes "s ∈ reach T" shows "∃b ∈ beh T. s ∈ set b"
using assms
proof (induction s rule: reach.induct)
case (r_init s)
hence "s ∈ set [s]" and "[s] ∈ beh T" by auto
thus ?case by fastforce
next
case (r_trans s t)
then obtain b where "b ∈ beh T" and "s ∈ set b" by blast
from ‹s ∈ set b› obtain b1 b2 where "b = b2 @ s # b1" by (blast dest: split_list)
with ‹b ∈ beh T› have "s # b1 ∈ beh T" by (blast intro: beh_prefix_closed)
with ‹(s, t) ∈ trans T› have "t # s # b1 ∈ beh T" by blast
thus ?case by force
qed
lemma reach_equiv_beh_states: "reach T = ⋃ (set`(beh T))"
by (auto intro!: reach_in_beh beh_in_reach)
subsubsection ‹Specifications, observability, and implementation›
text ‹Specifications add an observer function to transition systems.›
record ('s, 'o) spec = "'s TS" +
obs :: "'s ⇒ 'o"
lemma beh_obs_upd [simp]: "beh (S(| obs := x |)) = beh S"
by (safe) (erule beh.induct, auto)+
lemma reach_obs_upd [simp]: "reach (S(| obs := x |)) = reach S"
by (safe) (erule reach.induct, auto)+
text ‹Observable behaviour and reachability.›
definition
obeh :: "('s, 'o) spec ⇒ ('o list) set" where
"obeh S ≡ (map (obs S))`(beh S)"
definition
oreach :: "('s, 'o) spec ⇒ 'o set" where
"oreach S ≡ (obs S)`(reach S)"
lemma oreach_equiv_obeh_states:
"oreach S = ⋃ (set`(obeh S))"
by (auto simp add: reach_equiv_beh_states oreach_def obeh_def)
lemma obeh_pi_translation:
"(map pi)`(obeh S) = obeh (S(| obs := pi o (obs S) |))"
by (auto simp add: obeh_def image_comp)
lemma oreach_pi_translation:
"pi`(oreach S) = oreach (S(| obs := pi o (obs S) |))"
by (auto simp add: oreach_def)
text ‹A predicate $P$ on the states of a specification is \emph{observable}
if it cannot distinguish between states yielding the same observation.
Equivalently, $P$ is observable if it is the inverse image under the
observation function of a predicate on observations.›
definition
observable :: "['s ⇒ 'o, 's set] ⇒ bool"
where
"observable ob P ≡ ∀s s'. ob s = ob s' ⟶ s' ∈ P ⟶ s ∈ P"
definition
observable2 :: "['s ⇒ 'o, 's set] ⇒ bool"
where
"observable2 ob P ≡ ∃Q. P = ob-`Q"
definition
observable3 :: "['s ⇒ 'o, 's set] ⇒ bool"
where
"observable3 ob P ≡ ob-`ob`P ⊆ P"
lemma observableE [elim]:
"⟦observable ob P; ob s = ob s'; s' ∈ P⟧ ⟹ s ∈ P"
by (unfold observable_def) (fast)
lemma observable2_equiv_observable: "observable2 ob P = observable ob P"
by (unfold observable_def observable2_def) (auto)
lemma observable3_equiv_observable2: "observable3 ob P = observable2 ob P"
by (unfold observable3_def observable2_def) (auto)
lemma observable_id [simp]: "observable id P"
by (simp add: observable_def)
text ‹The set extension of a function @{term "ob"} is the left adjoint of
a Galois connection on the powerset lattices over domain and range of @{term "ob"}
where the right adjoint is the inverse image function.›
lemma image_vimage_adjoints: "(ob`P ⊆ Q) = (P ⊆ ob-`Q)"
by auto
declare image_vimage_subset [simp, intro]
declare vimage_image_subset [simp, intro]
text ‹Similar but "reversed" (wrt to adjointness) relationships only hold under
additional conditions.›
lemma image_r_vimage_l: "⟦ Q ⊆ ob`P; observable ob P ⟧ ⟹ ob-`Q ⊆ P"
by (auto)
lemma vimage_l_image_r: "⟦ ob-`Q ⊆ P; Q ⊆ range ob ⟧ ⟹ Q ⊆ ob`P"
by (drule image_mono [where f=ob], auto)
text ‹Internal and external invariants›
lemma external_from_internal_invariant:
"⟦ reach S ⊆ P; (obs S)`P ⊆ Q ⟧
⟹ oreach S ⊆ Q"
by (auto simp add: oreach_def)
lemma external_from_internal_invariant_vimage:
"⟦ reach S ⊆ P; P ⊆ (obs S)-`Q ⟧
⟹ oreach S ⊆ Q"
by (erule external_from_internal_invariant) (auto)
lemma external_to_internal_invariant_vimage:
"⟦ oreach S ⊆ Q; (obs S)-`Q ⊆ P ⟧
⟹ reach S ⊆ P"
by (auto simp add: oreach_def)
lemma external_to_internal_invariant:
"⟦ oreach S ⊆ Q; Q ⊆ (obs S)`P; observable (obs S) P ⟧
⟹ reach S ⊆ P"
by (erule external_to_internal_invariant_vimage) (auto)
lemma external_equiv_internal_invariant_vimage:
"⟦ P = (obs S)-`Q ⟧
⟹ (oreach S ⊆ Q) = (reach S ⊆ P)"
by (fast intro: external_from_internal_invariant_vimage
external_to_internal_invariant_vimage
del: subsetI)
lemma external_equiv_internal_invariant:
"⟦ (obs S)`P = Q; observable (obs S) P ⟧
⟹ (oreach S ⊆ Q) = (reach S ⊆ P)"
by (rule external_equiv_internal_invariant_vimage) (auto)
text ‹Our notion of implementation is inclusion of observable behaviours.›
definition
implements :: "['p ⇒ 'o, ('s, 'o) spec, ('t, 'p) spec] ⇒ bool" where
"implements pi Sa Sc ≡ (map pi)`(obeh Sc) ⊆ obeh Sa"
text ‹Reflexivity and transitivity›
lemma implements_refl: "implements id S S"
by (auto simp add: implements_def)
lemma implements_trans:
"⟦ implements pi1 S1 S2; implements pi2 S2 S3 ⟧
⟹ implements (pi1 o pi2) S1 S3"
by (fastforce simp add: implements_def image_subset_iff)
text ‹Preservation of external invariants›
lemma implements_oreach:
"implements pi Sa Sc ⟹ pi`(oreach Sc) ⊆ oreach Sa"
by (auto simp add: implements_def oreach_equiv_obeh_states dest!: subsetD)
lemma external_invariant_preservation:
"⟦ oreach Sa ⊆ Q; implements pi Sa Sc ⟧
⟹ pi`(oreach Sc) ⊆ Q"
by (rule subset_trans [OF implements_oreach]) (auto)
lemma external_invariant_translation:
"⟦ oreach Sa ⊆ Q; pi-`Q ⊆ P; implements pi Sa Sc ⟧
⟹ oreach Sc ⊆ P"
apply (rule subset_trans [OF vimage_image_subset, of pi])
apply (rule subset_trans [where B="pi-`Q"])
apply (intro vimage_mono external_invariant_preservation, auto)
done
text ‹Preservation of internal invariants›
lemma internal_invariant_translation:
"⟦ reach Sa ⊆ Pa; Pa ⊆ obs Sa -` Qa; pi -` Qa ⊆ Q; obs S -` Q ⊆ P;
implements pi Sa S ⟧
⟹ reach S ⊆ P"
by (rule external_from_internal_invariant_vimage [
THEN external_invariant_translation,
THEN external_to_internal_invariant_vimage])
subsection ‹Invariants›
text ‹First we define Hoare triples over transition relations and then
we derive proof rules to establish invariants.›
subsubsection ‹Hoare triples›
definition
PO_hoare :: "['s set, ('s × 's) set, 's set] ⇒ bool"
("(3{_} _ {> _})" [0, 0, 0] 90)
where
"{pre} R {> post} ≡ R``pre ⊆ post"
lemmas PO_hoare_defs = PO_hoare_def Image_def
lemma "{P} R {> Q} = (∀s t. s ∈ P ⟶ (s, t) ∈ R ⟶ t ∈ Q)"
by (auto simp add: PO_hoare_defs)
text ‹Some essential facts about Hoare triples.›
lemma hoare_conseq_left [intro]:
"⟦ {P'} R {> Q}; P ⊆ P' ⟧
⟹ {P} R {> Q}"
by (auto simp add: PO_hoare_defs)
lemma hoare_conseq_right:
"⟦ {P} R {> Q'}; Q' ⊆ Q ⟧
⟹ {P} R {> Q}"
by (auto simp add: PO_hoare_defs)
lemma hoare_false_left [simp]:
"{{}} R {> Q}"
by (auto simp add: PO_hoare_defs)
lemma hoare_true_right [simp]:
"{P} R {> UNIV}"
by (auto simp add: PO_hoare_defs)
lemma hoare_conj_right [intro!]:
"⟦ {P} R {> Q1}; {P} R {> Q2} ⟧
⟹ {P} R {> Q1 ∩ Q2}"
by (auto simp add: PO_hoare_defs)
text ‹Special transition relations.›
lemma hoare_stop [simp, intro!]:
"{P} {} {> Q}"
by (auto simp add: PO_hoare_defs)
lemma hoare_skip [simp, intro!]:
"P ⊆ Q ⟹ {P} Id {> Q}"
by (auto simp add: PO_hoare_defs)
lemma hoare_trans_Un [iff]:
"{P} R1 ∪ R2 {> Q} = ({P} R1 {> Q} ∧ {P} R2 {> Q})"
by (auto simp add: PO_hoare_defs)
lemma hoare_trans_UN [iff]:
"{P} ⋃ x. R x {> Q} = (∀x. {P} R x {> Q})"
by (auto simp add: PO_hoare_defs)
subsubsection ‹Characterization of reachability›
lemma reach_init: "reach T ⊆ I ⟹ init T ⊆ I"
by (auto dest: subsetD)
lemma reach_trans: "reach T ⊆ I ⟹ {reach T} trans T {> I}"
by (auto simp add: PO_hoare_defs)
text ‹Useful consequences.›
corollary init_reach [iff]: "init T ⊆ reach T"
by (rule reach_init, simp)
corollary trans_reach [iff]: "{reach T} trans T {> reach T}"
by (rule reach_trans, simp)
subsubsection ‹Invariant proof rules›
text ‹Basic proof rule for invariants.›
lemma inv_rule_basic:
"⟦ init T ⊆ P; {P} (trans T) {> P} ⟧
⟹ reach T ⊆ P"
by (safe, erule reach.induct, auto simp add: PO_hoare_def)
text ‹General invariant proof rule. This rule is complete (set
@{term "I = reach T"}).›
lemma inv_rule:
"⟦ init T ⊆ I; I ⊆ P; {I} (trans T) {> I} ⟧
⟹ reach T ⊆ P"
apply (rule subset_trans, auto)
apply (erule reach.induct, auto simp add: PO_hoare_def)
done
text ‹The following rule is equivalent to the previous one.›
lemma INV_rule:
"⟦ init T ⊆ I; {I ∩ reach T} (trans T) {> I} ⟧
⟹ reach T ⊆ I"
by (safe, erule reach.induct, auto simp add: PO_hoare_defs)
text ‹Proof of equivalence.›
lemma inv_rule_from_INV_rule:
"⟦ init T ⊆ I; I ⊆ P; {I} (trans T) {> I} ⟧
⟹ reach T ⊆ P"
apply (rule subset_trans, auto del: subsetI)
apply (rule INV_rule, auto)
done
lemma INV_rule_from_inv_rule:
"⟦ init T ⊆ I; {I ∩ reach T} (trans T) {> I} ⟧
⟹ reach T ⊆ I"
by (rule_tac I="I ∩ reach T" in inv_rule, auto)
text ‹Incremental proof rule for invariants using auxiliary invariant(s).
This rule might have become obsolete by addition of $INV\_rule$.›
lemma inv_rule_incr:
"⟦ init T ⊆ I; {I ∩ J} (trans T) {> I}; reach T ⊆ J ⟧
⟹ reach T ⊆ I"
by (rule INV_rule, auto)
subsection ‹Refinement›
text ‹Our notion of refinement is simulation. We first define a general
notion of relational Hoare tuple, which we then use to define the refinement
proof obligation. Finally, we show that observation-consistent refinement
of specifications implies the implementation relation between them.›
subsubsection ‹Relational Hoare tuples›
text ‹Relational Hoare tuples formalize the following generalized simulation
diagram:
\begin{verbatim}
o -- Ra --> o
| |
pre post
| |
v V
o -- Rc --> o
\end{verbatim}
Here, $Ra$ and $Rc$ are the abstract and concrete transition relations,
and $pre$ and $post$ are the pre- and post-relations.
(In the definiton below, the operator @{term "(O)"} stands for relational
composition, which is defined as follows: @{thm relcomp_def [no_vars]}.)›
definition
PO_rhoare ::
"[('s × 't) set, ('s × 's) set, ('t × 't) set, ('s × 't) set] ⇒ bool"
("(4{_} _, _ {> _})" [0, 0, 0] 90)
where
"{pre} Ra, Rc {> post} ≡ pre O Rc ⊆ Ra O post"
lemmas PO_rhoare_defs = PO_rhoare_def relcomp_unfold
text ‹Facts about relational Hoare tuples.›
lemma relhoare_conseq_left [intro]:
"⟦ {pre'} Ra, Rc {> post}; pre ⊆ pre' ⟧
⟹ {pre} Ra, Rc {> post}"
by (auto simp add: PO_rhoare_defs dest!: subsetD)
lemma relhoare_conseq_right:
"⟦ {pre} Ra, Rc {> post'}; post' ⊆ post ⟧
⟹ {pre} Ra, Rc {> post}"
by (auto simp add: PO_rhoare_defs)
lemma relhoare_false_left [simp]:
"{ {} } Ra, Rc {> post}"
by (auto simp add: PO_rhoare_defs)
lemma relhoare_true_right [simp]:
"{pre} Ra, Rc {> UNIV} = (Domain (pre O Rc) ⊆ Domain Ra)"
by (auto simp add: PO_rhoare_defs)
lemma Domain_rel_comp [intro]:
"Domain pre ⊆ R ⟹ Domain (pre O Rc) ⊆ R"
by (auto simp add: Domain_def)
lemma rel_hoare_skip [iff]: "{R} Id, Id {> R}"
by (auto simp add: PO_rhoare_def)
text ‹Reflexivity and transitivity.›
lemma relhoare_refl [simp]: "{Id} R, R {> Id}"
by (auto simp add: PO_rhoare_defs)
lemma rhoare_trans:
"⟦ {R1} T1, T2 {> R1}; {R2} T2, T3 {> R2} ⟧
⟹ {R1 O R2} T1, T3 {> R1 O R2}"
apply (auto simp add: PO_rhoare_def del: subsetI)
apply (drule subset_refl [THEN relcomp_mono, where r=R1])
apply (drule subset_refl [THEN [2] relcomp_mono, where s=R2])
apply (auto simp add: O_assoc del: subsetI)
done
text ‹Conjunction in the post-relation cannot be split in general. However,
here are two useful special cases. In the first case the abstract transtition
relation is deterministic and in the second case one conjunct is a cartesian
product of two state predicates.›
lemma relhoare_conj_right_det:
"⟦ {pre} Ra, Rc {> post1}; {pre} Ra, Rc {> post2};
single_valued Ra ⟧
⟹ {pre} Ra, Rc {> post1 ∩ post2}"
by (auto simp add: PO_rhoare_defs dest: single_valuedD dest!: subsetD)
lemma relhoare_conj_right_cartesian [intro]:
"⟦ {Domain pre} Ra {> I}; {Range pre} Rc {> J};
{pre} Ra, Rc {> post} ⟧
⟹ {pre} Ra, Rc {> post ∩ I × J}"
by (force simp add: PO_rhoare_defs PO_hoare_defs Domain_def Range_def)
text ‹Separate rule for cartesian products.›
corollary relhoare_cartesian:
"⟦ {Domain pre} Ra {> I}; {Range pre} Rc {> J};
{pre} Ra, Rc {> post} ⟧
⟹ {pre} Ra, Rc {> I × J}"
by (auto intro: relhoare_conseq_right)
text ‹Unions of transition relations.›
lemma relhoare_concrete_Un [simp]:
"{pre} Ra, Rc1 ∪ Rc2 {> post}
= ({pre} Ra, Rc1 {> post} ∧ {pre} Ra, Rc2 {> post})"
apply (auto simp add: PO_rhoare_defs)
apply (auto dest!: subsetD)
done
lemma relhoare_concrete_UN [simp]:
"{pre} Ra, ⋃x. Rc x {> post} = (∀x. {pre} Ra, Rc x {> post})"
apply (auto simp add: PO_rhoare_defs)
apply (auto dest!: subsetD)
done
lemma relhoare_abstract_Un_left [intro]:
"⟦ {pre} Ra1, Rc {> post} ⟧
⟹ {pre} Ra1 ∪ Ra2, Rc {> post}"
by (auto simp add: PO_rhoare_defs)
lemma relhoare_abstract_Un_right [intro]:
"⟦ {pre} Ra2, Rc {> post} ⟧
⟹ {pre} Ra1 ∪ Ra2, Rc {> post}"
by (auto simp add: PO_rhoare_defs)
lemma relhoare_abstract_UN [intro!]:
"⟦ {pre} Ra x, Rc {> post} ⟧
⟹ {pre} ⋃x. Ra x, Rc {> post}"
apply (auto simp add: PO_rhoare_defs)
apply (auto dest!: subsetD)
done
subsubsection ‹Refinement proof obligations›
text ‹A transition system refines another one if the initial states and the
transitions are refined.
Initial state refinement means that for each concrete initial state there is
a related abstract one. Transition refinement means that the simulation
relation is preserved (as expressed by a relational Hoare tuple).
›
definition
PO_refines ::
"[('s × 't) set, ('s, 'a) TS_scheme, ('t, 'b) TS_scheme] ⇒ bool"
where
"PO_refines R Ta Tc ≡ (
init Tc ⊆ R``(init Ta)
∧ {R} (trans Ta), (trans Tc) {> R}
)"
lemma PO_refinesI:
"⟦ init Tc ⊆ R``(init Ta); {R} (trans Ta), (trans Tc) {> R} ⟧ ⟹ PO_refines R Ta Tc"
by (simp add: PO_refines_def)
lemma PO_refinesE [elim]:
"⟦ PO_refines R Ta Tc; ⟦ init Tc ⊆ R``(init Ta); {R} (trans Ta), (trans Tc) {> R} ⟧ ⟹ P ⟧
⟹ P"
by (simp add: PO_refines_def)
text ‹Basic refinement rule. This is just an introduction rule for the
definition.›
lemma refine_basic:
"⟦ init Tc ⊆ R``(init Ta); {R} (trans Ta), (trans Tc) {> R} ⟧
⟹ PO_refines R Ta Tc"
by (simp add: PO_refines_def)
text ‹The following proof rule uses individual invariants @{term "I"}
and @{term "J"} of the concrete and abstract systems to strengthen the
simulation relation $R$.
The hypotheses state that these state predicates are indeed invariants.
Note that the pre-condition of the invariant preservation hypotheses for
@{term "I"} and @{term "J"} are strengthened by adding the predicates
@{term "Domain (R ∩ UNIV × J)"} and @{term "Range (R ∩ I × UNIV)"},
respectively. In particular, the latter predicate may be essential, if a
concrete invariant depends on the simulation relation and an abstract invariant,
i.e. to "transport" abstract invariants to the concrete system.›
lemma refine_init_using_invariants:
"⟦ init Tc ⊆ R``(init Ta); init Ta ⊆ I; init Tc ⊆ J ⟧
⟹ init Tc ⊆ (R ∩ I × J)``(init Ta)"
by (auto simp add: Image_def dest!: bspec subsetD)
lemma refine_trans_using_invariants:
"⟦ {R ∩ I × J} (trans Ta), (trans Tc) {> R};
{I ∩ Domain (R ∩ UNIV × J)} (trans Ta) {> I};
{J ∩ Range (R ∩ I × UNIV)} (trans Tc) {> J} ⟧
⟹ {R ∩ I × J} (trans Ta), (trans Tc) {> R ∩ I × J}"
by (rule relhoare_conj_right_cartesian) (auto)
text ‹This is our main rule for refinements.›
lemma refine_using_invariants:
"⟦ {R ∩ I × J} (trans Ta), (trans Tc) {> R};
{I ∩ Domain (R ∩ UNIV × J)} (trans Ta) {> I};
{J ∩ Range (R ∩ I × UNIV)} (trans Tc) {> J};
init Tc ⊆ R``(init Ta);
init Ta ⊆ I; init Tc ⊆ J ⟧
⟹ PO_refines (R ∩ I × J) Ta Tc"
by (unfold PO_refines_def)
(intro refine_init_using_invariants refine_trans_using_invariants conjI)
subsubsection ‹Deriving invariants from refinements›
text ‹Some invariants can only be proved after the simulation has been
established, because they depend on the simulation relation and some abstract
invariants. Here is a rule to derive invariant theorems from the refinement.›
lemma PO_refines_implies_Range_init:
"PO_refines R Ta Tc ⟹ init Tc ⊆ Range R"
by (auto simp add: PO_refines_def)
lemma PO_refines_implies_Range_trans:
"PO_refines R Ta Tc ⟹ {Range R} trans Tc {> Range R}"
by (auto simp add: PO_refines_def PO_rhoare_def PO_hoare_def)
lemma PO_refines_implies_Range_invariant:
"PO_refines R Ta Tc ⟹ reach Tc ⊆ Range R"
by (rule INV_rule)
(auto intro!: PO_refines_implies_Range_init
PO_refines_implies_Range_trans)
text ‹The following rules are more useful in proofs.›
corollary INV_init_from_refinement:
"⟦ PO_refines R Ta Tc; Range R ⊆ I ⟧
⟹ init Tc ⊆ I"
by (drule PO_refines_implies_Range_init, auto)
corollary INV_trans_from_refinement:
"⟦ PO_refines R Ta Tc; K ⊆ Range R; Range R ⊆ I ⟧
⟹ {K} trans Tc {> I}"
apply (drule PO_refines_implies_Range_trans)
apply (auto intro: hoare_conseq_right)
done
corollary INV_from_refinement:
"⟦ PO_refines R Ta Tc; Range R ⊆ I ⟧
⟹ reach Tc ⊆ I"
by (drule PO_refines_implies_Range_invariant, fast)
subsubsection ‹Refinement of specifications›
text ‹Lift relation membership to finite sequences›
inductive_set
seq_lift :: "('s × 't) set ⇒ ('s list × 't list) set"
for R :: "('s × 't) set"
where
sl_nil [iff]: "([], []) ∈ seq_lift R"
| sl_cons [intro]:
"⟦ (xs, ys) ∈ seq_lift R; (x, y) ∈ R ⟧ ⟹ (x#xs, y#ys) ∈ seq_lift R"
inductive_cases sl_cons_right_invert: "(ba', t # bc) ∈ seq_lift R"
text ‹For each concrete behaviour there is a related abstract one.›
lemma behaviour_refinement:
"⟦ PO_refines R Ta Tc; bc ∈ beh Tc⟧
⟹ ∃ba ∈ beh Ta. (ba, bc) ∈ seq_lift R"
apply (erule beh.induct, auto)
apply (clarsimp simp add: PO_refines_def Image_def)
apply (drule subsetD, auto)
apply (erule sl_cons_right_invert, clarsimp)
apply (rename_tac s bc s' ba t)
apply (auto simp add: PO_refines_def PO_rhoare_def)
apply (thin_tac "X ⊆ Y" for X Y)
apply (drule subsetD, auto)
done
text ‹Observation consistency of a relation is defined using a mediator
function @{term "pi"} to abstract the concrete observation. This allows us
to also refine the observables as we move down a refinement branch.
›
definition
obs_consistent ::
"[('s × 't) set, 'p ⇒ 'o, ('s, 'o) spec, ('t, 'p) spec] ⇒ bool"
where
"obs_consistent R pi Sa Sc ≡ (∀s t. (s, t) ∈ R ⟶ pi (obs Sc t) = obs Sa s)"
lemma obs_consistent_refl [iff]: "obs_consistent Id id S S"
by (simp add: obs_consistent_def)
lemma obs_consistent_trans [intro]:
"⟦ obs_consistent R1 pi1 S1 S2; obs_consistent R2 pi2 S2 S3 ⟧
⟹ obs_consistent (R1 O R2) (pi1 o pi2) S1 S3"
by (auto simp add: obs_consistent_def)
lemma obs_consistent_empty: "obs_consistent {} pi Sa Sc"
by (auto simp add: obs_consistent_def)
lemma obs_consistent_conj1 [intro]:
"obs_consistent R pi Sa Sc ⟹ obs_consistent (R ∩ R') pi Sa Sc"
by (auto simp add: obs_consistent_def)
lemma obs_consistent_conj2 [intro]:
"obs_consistent R pi Sa Sc ⟹ obs_consistent (R' ∩ R) pi Sa Sc"
by (auto simp add: obs_consistent_def)
lemma obs_consistent_behaviours:
"⟦ obs_consistent R pi Sa Sc; bc ∈ beh Sc; ba ∈ beh Sa; (ba, bc) ∈ seq_lift R⟧
⟹ map pi (map (obs Sc) bc) = map (obs Sa) ba"
by (erule seq_lift.induct) (auto simp add: obs_consistent_def)
text ‹Definition of refinement proof obligations.›
definition
refines ::
"[('s × 't) set, 'p ⇒ 'o, ('s, 'o) spec, ('t, 'p) spec] ⇒ bool"
where
"refines R pi Sa Sc ≡ obs_consistent R pi Sa Sc ∧ PO_refines R Sa Sc"
lemmas refines_defs =
refines_def PO_refines_def
lemma refinesI:
"⟦ PO_refines R Sa Sc; obs_consistent R pi Sa Sc ⟧
⟹ refines R pi Sa Sc"
by (simp add: refines_def)
lemma refinesE [elim]:
"⟦ refines R pi Sa Sc; ⟦ PO_refines R Sa Sc; obs_consistent R pi Sa Sc ⟧ ⟹ P ⟧
⟹ P"
by (simp add: refines_def)
text ‹Reflexivity and transitivity of refinement.›
lemma refinement_reflexive: "refines Id id S S"
by (auto simp add: refines_defs)
lemma refinement_transitive:
"⟦ refines R1 pi1 S1 S2; refines R2 pi2 S2 S3 ⟧
⟹ refines (R1 O R2) (pi1 o pi2) S1 S3"
apply (auto simp add: refines_defs del: subsetI
intro: rhoare_trans)
apply (fastforce dest: Image_mono)
done
text ‹Soundness of refinement for proving implementation›
lemma observable_behaviour_refinement:
"⟦ refines R pi Sa Sc; bc ∈ obeh Sc ⟧ ⟹ map pi bc ∈ obeh Sa"
by (auto simp add: refines_def obeh_def image_def
dest!: behaviour_refinement obs_consistent_behaviours)
theorem refinement_soundness:
"refines R pi Sa Sc ⟹ implements pi Sa Sc"
by (auto simp add: implements_def
elim!: observable_behaviour_refinement)
text ‹Extended versions of refinement proof rules including observations›
lemmas Refinement_basic = refine_basic [THEN refinesI]
lemmas Refinement_using_invariants = refine_using_invariants [THEN refinesI]
lemma refines_reachable_strengthening:
"refines R pi Sa Sc ⟹ refines (R ∩ reach Sa × reach Sc) pi Sa Sc"
by (auto intro!: Refinement_using_invariants)
text ‹Inhertitance of internal invariants through refinements›
lemma INV_init_from_Refinement:
"⟦refines R pi Sa Sc; Range R ⊆ I⟧ ⟹ init Sc ⊆ I"
by (blast intro: INV_init_from_refinement)
lemma INV_trans_from_Refinement:
" ⟦refines R pi Sa Sc; K ⊆ Range R; Range R ⊆ I⟧ ⟹ {K} TS.trans Sc {> I}"
by (blast intro: INV_trans_from_refinement)
lemma INV_from_Refinement_basic:
"⟦ refines R pi Sa Sc; Range R ⊆ I ⟧ ⟹ reach Sc ⊆ I"
by (rule INV_from_refinement) blast
lemma INV_from_Refinement_using_invariants:
assumes "refines R pi Sa Sc" "Range (R ∩ I × J) ⊆ K"
"reach Sa ⊆ I" "reach Sc ⊆ J"
shows "reach Sc ⊆ K"
proof (rule INV_from_Refinement_basic)
show "refines (R ∩ reach Sa × reach Sc) pi Sa Sc" using assms(1)
by (rule refines_reachable_strengthening)
next
show "Range (R ∩ reach Sa × reach Sc) ⊆ K" using assms(2-4) by blast
qed
end