Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Pierre-Marie Pédrot
Iris
Commits
8bc4cec0
Commit
8bc4cec0
authored
Feb 13, 2016
by
Ralf Jung
Browse files
create a hint database for solving some simple uPred's
parent
bd7ebdec
Changes
2
Show whitespace changes
Inline
Side-by-side
algebra/upred.v
View file @
8bc4cec0
...
...
@@ -980,4 +980,14 @@ Proof. by rewrite -(always_always Q); apply always_entails_l'. Qed.
Lemma
always_entails_r
P
Q
`
{!
AlwaysStable
Q
}
:
(
P
⊑
Q
)
→
P
⊑
(
P
★
Q
).
Proof
.
by
rewrite
-(
always_always
Q
)
;
apply
always_entails_r'
.
Qed
.
End
uPred_logic
.
End
uPred
.
End
uPred_logic
.
(* uPred Hint DB *)
Create
HintDb
uPred
.
Hint
Resolve
or_elim
or_intro_l'
or_intro_r'
:
uPred
.
Hint
Resolve
and_intro
and_elim_l'
and_elim_r'
:
uPred
.
Hint
Resolve
always_mono
:
uPred
.
Hint
Resolve
sep_elim_l'
sep_elim_r'
sep_mono
:
uPred
.
Hint
Immediate
True_intro
False_elim
:
uPred
.
End
uPred
.
program_logic/viewshifts.v
View file @
8bc4cec0
...
...
@@ -92,7 +92,7 @@ Lemma vs_open_close N E P Q R :
Proof
.
intros
;
apply
(
always_intro
_
_
),
impl_intro_l
.
rewrite
always_and_sep_r
assoc
[(
P
★
_
)%
I
]
comm
-
assoc
.
eapply
pvs_open_close
;
[
by
eauto
using
sep_elim_l'
..|].
eapply
pvs_open_close
;
[
by
eauto
with
uPred
..|].
rewrite
sep_elim_r
.
apply
wand_intro_l
.
(* Oh wow, this is annyoing... *)
rewrite
assoc
-
always_and_sep_r'
.
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment