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

Merge branch 'robbert/big_andL_orL' into 'master'

Add some missing lemmas for `big_andL` and `big_orL`.

See merge request iris/iris!720
parents 26ebf1ee 060ab725
No related branches found
No related tags found
No related merge requests found
......@@ -974,12 +974,14 @@ Section and_list.
( k x, l !! k = Some x Φ k x) [ list] kx l, Φ k x.
Proof. rewrite big_andL_forall //. Qed.
Global Instance big_andL_nil_persistent Φ :
Persistent ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_andL_persistent Φ l :
( k x, Persistent (Φ k x)) Persistent ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
Lemma big_andL_impl Φ Ψ m :
([ list] kx m, Φ k x)
( k x, m !! k = Some x Φ k x Ψ k x) -∗
[ list] kx m, Ψ k x.
Proof.
rewrite -big_andL_forall -big_andL_and.
by setoid_rewrite bi.impl_elim_r.
Qed.
Lemma big_andL_pure_1 (φ : nat A Prop) l :
([ list] kx l, φ k x) ⊢@{PROP} ⌜∀ k x, l !! k = Some x φ k x⌝.
......@@ -1005,6 +1007,26 @@ Section and_list.
apply big_andL_pure_2.
Qed.
Lemma big_andL_later Φ l :
([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_andL_laterN Φ n l :
▷^n ([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, ▷^n Φ k x).
Proof. apply (big_opL_commute _). Qed.
Global Instance big_andL_nil_persistent Φ :
Persistent ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_andL_persistent Φ l :
( k x, Persistent (Φ k x)) Persistent ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
Global Instance big_andL_nil_timeless Φ :
Timeless ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_andL_timeless Φ l :
( k x, Timeless (Φ k x)) Timeless ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
End and_list.
Section or_list.
......@@ -1122,12 +1144,28 @@ Section or_list.
([ list] kx l, Φ k x) Q ⊣⊢ ([ list] kx l, Φ k x Q).
Proof. setoid_rewrite (comm bi_sep). apply big_orL_sep_l. Qed.
Lemma big_orL_later Φ l :
l []
([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, Φ k x).
Proof. apply (big_opL_commute1 _). Qed.
Lemma big_orL_laterN Φ n l :
l []
▷^n ([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, ▷^n Φ k x).
Proof. apply (big_opL_commute1 _). Qed.
Global Instance big_orL_nil_persistent Φ :
Persistent ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_orL_persistent Φ l :
( k x, Persistent (Φ k x)) Persistent ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
Global Instance big_orL_nil_timeless Φ :
Timeless ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_orL_timeless Φ l :
( k x, Timeless (Φ k x)) Timeless ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
End or_list.
(** ** Big ops over finite maps *)
......
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