Skip to content
Snippets Groups Projects
Commit c55672d8 authored by Ralf Jung's avatar Ralf Jung
Browse files

prove that we actually completely characterized inclusion of the STS RA; derive a simpler form

parent 9bf9afb4
No related branches found
No related tags found
No related merge requests found
...@@ -228,23 +228,44 @@ Proof. ...@@ -228,23 +228,44 @@ Proof.
repeat (done || constructor). repeat (done || constructor).
Qed. Qed.
Lemma sts_frag_included S1 S2 T1 T2 Tdf : Lemma sts_frag_included S1 S2 T1 T2 :
sts.closed R tok S1 T1 sts.closed R tok S2 T2 sts.closed R tok S2 T2
T2 T1 Tdf T1 Tdf sts_frag S1 T1 sts_frag S2 T2
S2 (S1 sts.up_set R tok S2 Tdf) (sts.closed R tok S1 T1 Tf, T2 T1 Tf T1 Tf S2 (S1 sts.up_set R tok S2 Tf)).
sts_frag S1 T1 sts_frag S2 T2.
Proof. Proof.
move=>Hcl1 Hcl2 Htk Hdf Hs. exists (sts_frag (sts.up_set R tok S2 Tdf) Tdf). move=>Hcl2. split.
split; first split; simpl. - intros [xf EQ]. destruct xf as [xf vf Hvf]. destruct xf as [Sf Tf|Sf Tf].
- intros _. split_ands. { exfalso. inversion_clear EQ. apply H0 in Hcl2. simpl in Hcl2.
+ done. inversion Hcl2. }
+ apply sts.closed_up_set. inversion_clear EQ.
* move=>s Hs2. move:(sts.closed_disjoint _ _ _ _ Hcl2 _ Hs2). move:(H0 Hcl2)=>{H0} H0. inversion_clear H0.
destruct H as [H _]. move:(H Hcl2)=>{H} [/= Hcl1 [Hclf Hdisj]].
apply Hvf in Hclf. simpl in Hclf. clear Hvf.
inversion_clear Hdisj. split; last (exists Tf; split_ands); [done..|].
apply (anti_symm ()).
+ move=>s HS2. apply elem_of_intersection. split; first by apply H2.
by apply sts.subseteq_up_set.
+ move=>s /elem_of_intersection [HS1 Hscl]. apply H2. split; first done.
destruct Hscl as [s' [Hsup Hs']].
eapply sts.closed_steps; last (hnf in Hsup; eexact Hsup); first done.
solve_elem_of +H2 Hs'.
- intros (Hcl1 & Tf & Htk & Hf & Hs). exists (sts_frag (sts.up_set R tok S2 Tf) Tf).
split; first split; simpl;[|done|].
+ intros _. split_ands; first done.
* apply sts.closed_up_set; last by eapply sts.closed_ne.
move=>s Hs2. move:(sts.closed_disjoint _ _ _ _ Hcl2 _ Hs2).
solve_elem_of +Htk. solve_elem_of +Htk.
* by eapply sts.closed_ne. * constructor; last done. rewrite -Hs. by eapply sts.closed_ne.
+ constructor; last done. rewrite -Hs. by eapply sts.closed_ne. + intros _. constructor; [ solve_elem_of +Htk | done].
- done. Qed.
- intros _. constructor; [ solve_elem_of +Htk | done].
Qed. Lemma sts_frag_included' S1 S2 T :
sts.closed R tok S2 T sts.closed R tok S1 T
S2 (S1 sts.up_set R tok S2 )
sts_frag S1 T sts_frag S2 T.
Proof.
intros. apply sts_frag_included; first done.
split; first done. exists ∅. split_ands; done || solve_elem_of+.
Qed.
End stsRA. End stsRA.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment