Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
10
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Iris
stdpp
Commits
be4eb648
Commit
be4eb648
authored
Nov 26, 2018
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Versions of `elem_of_list_split` that give first or last element.
parent
f8f6d0a9
Pipeline
#13053
passed with stage
in 18 minutes and 36 seconds
Changes
1
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
with
22 additions
and
0 deletions
+22
-0
theories/list.v
theories/list.v
+22
-0
No files found.
theories/list.v
View file @
be4eb648
...
...
@@ -647,6 +647,28 @@ Proof.
induction 1 as [x l|x y l ? [l1 [l2 ->]]]; [by eexists [], l|].
by exists (y :: l1), l2.
Qed.
Lemma elem_of_list_split_l `{EqDecision A} l x :
x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2 ∧ x ∉ l1.
Proof.
induction 1 as [x l|x y l ? IH].
{ exists [], l. rewrite elem_of_nil. naive_solver. }
destruct (decide (x = y)) as [->|?].
- exists [], l. rewrite elem_of_nil. naive_solver.
- destruct IH as (l1 & l2 & -> & ?).
exists (y :: l1), l2. rewrite elem_of_cons. naive_solver.
Qed.
Lemma elem_of_list_split_r `{EqDecision A} l x :
x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2 ∧ x ∉ l2.
Proof.
induction l as [|y l IH] using rev_ind.
{ by rewrite elem_of_nil. }
destruct (decide (x = y)) as [->|].
- exists l, []. rewrite elem_of_nil. naive_solver.
- rewrite elem_of_app, elem_of_list_singleton. intros [?| ->]; try done.
destruct IH as (l1 & l2 & -> & ?); auto.
exists l1, (l2 ++ [y]).
rewrite elem_of_app, elem_of_list_singleton, <-(assoc_L (++)). naive_solver.
Qed.
Lemma elem_of_list_lookup_1 l x : x ∈ l → ∃ i, l !! i = Some x.
Proof.
induction 1 as [|???? IH]; [by exists 0 |].
...
...
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