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
Simon Spies
stdpp
Commits
3887c4f7
Commit
3887c4f7
authored
Jun 26, 2019
by
Simon Spies
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
seqZ: rename, by, cbn -> simpl
parent
bfebb1e6
Pipeline
#17999
canceled with stage
Changes
1
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
with
26 additions
and
28 deletions
+26
-28
theories/list.v
theories/list.v
+26
-28
No files found.
theories/list.v
View file @
3887c4f7
...
...
@@ -396,11 +396,9 @@ Definition positives_unflatten (p : positive) : option (list positive) :=
(** [seqZ m n] generates the sequence [m], [m + 1], ..., [m + n - 1] over integers, provided [n >= 0]. If n < 0, then the range is empty. **)
Definition seqZ (m len: Z) : list Z := (λ
x
: nat, Z.add
x
m) <$> (seq 0 (Z.to_nat len)).
Definition seqZ (m len: Z) : list Z := (λ
i
: nat, Z.add
i
m) <$> (seq 0 (Z.to_nat len)).
Arguments seqZ : simpl never.
(** * Basic tactics on lists *)
(** The tactic [discriminate_list] discharges a goal if it submseteq
a list equality involving [(::)] and [(++)] of two lists that have a different
...
...
@@ -1470,63 +1468,63 @@ Qed.
(** ** Properties of the [seqZ] function *)
Section seqZ.
Implicit Types (m n : Z) (i j: nat).
Open Scope Z.
Local
Open Scope Z.
Lemma
empty_
seqZ m n: n ≤ 0 → seqZ m n = [].
Proof. intros H; destruct n;
cbn
; eauto; lia. Qed.
Lemma seqZ
_nil
m n: n ≤ 0 → seqZ m n = [].
Proof. intros H; destruct n;
simpl
; eauto; lia. Qed.
Lemma s
tep_seqZ
m n: n > 0 → seqZ m n = m :: seqZ (m + 1) (n - 1).
Lemma s
eqZ_cons
m n: n > 0 → seqZ m n = m :: seqZ (m + 1) (n - 1).
Proof.
intros H. unfold seqZ.
replace (Z.to_nat n) with (S (Z.to_nat (n - 1))) by
(rewrite <-Z2Nat.inj_succ; [f_equal|]; lia).
cbn
; f_equal; try lia.
simpl
; f_equal; try lia.
erewrite <-fmap_seq, map_map, map_ext; eauto.
intros; lia.
Qed.
Lemma length
_seqZ
m n: length (seqZ m n) = Z.to_nat n.
Lemma
seqZ_
length m n: length (seqZ m n) = Z.to_nat n.
Proof.
unfold seqZ;
now
rewrite map_length, seq_length.
unfold seqZ;
by
rewrite map_length, seq_length.
Qed.
Lemma
fmap_seqZ
m m' n:
(
Z.add m
)
<$> seqZ m' n = seqZ (m + m') n.
Lemma
seqZ_fmap
m m' n: Z.add m <$> seqZ m' n = seqZ (m + m') n.
Proof.
assert (0 ≤ n ∨ n < 0) as [H|H] by lia.
- revert m'. pattern n. eapply natlike_ind; auto; clear n H.
intros n H1 IH j. rewrite s
tep_seqZ
; try lia.
symmetry. rewrite s
tep_seqZ
; try lia.
intros n H1 IH j. rewrite s
eqZ_cons
; try lia.
symmetry. rewrite s
eqZ_cons
; try lia.
replace (Z.succ n - 1) with n by lia.
cbn
; rewrite IH.
simpl
; rewrite IH.
f_equal; try lia; f_equal; lia.
- rewrite !
empty_
seqZ; auto; lia.
- rewrite !seqZ
_nil
; auto; lia.
Qed.
Lemma lookup
_seqZ
m n i: i < n → seqZ m n !! i = Some (m + i).
Lemma
seqZ_
lookup m n i: i < n → seqZ m n !! i = Some (m + i).
Proof.
assert (0 ≤ n ∨ n < 0) as [H|H] by lia.
- revert m i. pattern n. eapply natlike_ind; auto; clear n H.
intros; lia.
intros n H1 IH. intros j [|i] ?; rewrite s
tep_seqZ
. 2, 4: lia.
cbn
; f_equal; lia. replace (Z.succ n - 1) with n by lia.
cbn
; rewrite IH; f_equal; lia.
- rewrite !
empty_
seqZ; auto; lia.
intros n H1 IH. intros j [|i] ?; rewrite s
eqZ_cons
. 2, 4: lia.
simpl
; f_equal; lia. replace (Z.succ n - 1) with n by lia.
simpl
; rewrite IH; f_equal; lia.
- rewrite !seqZ
_nil
; auto; lia.
Qed.
Lemma lookup_
seqZ_
ge m n i: n ≤ i → seqZ m n !! i = None.
Lemma
seqZ_
lookup_ge m n i: n ≤ i → seqZ m n !! i = None.
Proof.
assert (0 ≤ n ∨ n < 0) as [H|H] by lia.
- revert m i. pattern n. eapply natlike_ind; auto; clear n H.
intros n H1 IH. intros j [|i] ?; rewrite s
tep_seqZ
. 2, 4: lia.
cbn
; f_equal; lia. replace (Z.succ n - 1) with n by lia.
cbn
; rewrite IH; f_equal; lia.
- rewrite !
empty_
seqZ; auto; lia.
intros n H1 IH. intros j [|i] ?; rewrite s
eqZ_cons
. 2, 4: lia.
simpl
; f_equal; lia. replace (Z.succ n - 1) with n by lia.
simpl
; rewrite IH; f_equal; lia.
- rewrite !seqZ
_nil
; auto; lia.
Qed.
Lemma lookup_
seqZ_
inv m n i m' : seqZ m n !! i = Some m' → m' = m + i ∧ i < n.
Lemma
seqZ_
lookup_inv m n i m' : seqZ m n !! i = Some m' → m' = m + i ∧ i < n.
Proof.
destruct (Z_le_gt_dec n i); [by rewrite lookup_
seqZ_
ge|].
rewrite lookup
_seqZ
by lia. intuition; [congruence|lia].
destruct (Z_le_gt_dec n i); [by rewrite
seqZ_
lookup_ge|].
rewrite
seqZ_
lookup by lia. intuition; [congruence|lia].
Qed.
End seqZ.
...
...
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