Skip to content
Snippets Groups Projects

Compare revisions

Changes are shown as if the source revision was being merged into the target revision. Learn more about comparing revisions.

Source

Select target project
No results found

Target

Select target project
  • iris/stdpp
  • johannes/stdpp
  • proux1/stdpp
  • dosualdo/stdpp
  • benoit/coq-stdpp
  • dfrumin/coq-stdpp
  • haidang/stdpp
  • amintimany/coq-stdpp
  • swasey/coq-stdpp
  • simongregersen/stdpp
  • proux/stdpp
  • janno/coq-stdpp
  • amaurremi/coq-stdpp
  • msammler/stdpp
  • tchajed/stdpp
  • YaZko/stdpp
  • maximedenes/stdpp
  • jakobbotsch/stdpp
  • Blaisorblade/stdpp
  • simonspies/stdpp
  • lepigre/stdpp
  • devilhena/stdpp
  • simonfv/stdpp
  • jihgfee/stdpp
  • snyke7/stdpp
  • Armael/stdpp
  • gmalecha/stdpp
  • olaure01/stdpp
  • sarahzrf/stdpp
  • atrieu/stdpp
  • herbelin/stdpp
  • arthuraa/stdpp
  • lgaeher/stdpp
  • mrhaandi/stdpp
  • mattam82/stdpp
  • Quarkbeast/stdpp
  • aa755/stdpp
  • gmevel/stdpp
  • lstefane/stdpp
  • jung/stdpp
  • vsiles/stdpp
  • dlesbre/stdpp
  • bergwerf/stdpp
  • marijnvanwezel/stdpp
  • ivanbakel/stdpp
  • tperami/stdpp
  • adamAndMath/stdpp
  • Villetaneuse/stdpp
  • sanjit/stdpp
  • yiyunliu/stdpp
  • thomas-lamiaux/stdpp
  • Tragicus/stdpp
  • kbedarka/stdpp
53 results
Show changes
Commits on Source (24)
...@@ -7,12 +7,25 @@ API-breaking change is listed. ...@@ -7,12 +7,25 @@ API-breaking change is listed.
- Add lemma `lookup_total_fmap`. - Add lemma `lookup_total_fmap`.
- Add lemmas about `last` and `head`: `last_app_Some`, `last_app_None`, - Add lemmas about `last` and `head`: `last_app_Some`, `last_app_None`,
`head_app_Some`, `head_app_None` and `head_app`. (by Marijn van Wezel) `head_app_Some`, `head_app_None` and `head_app`. (by Marijn van Wezel)
- Add tactic `notypeclasses apply` that works like `notypeclasses refine` but
automatically adds `_` until the given lemma takes no more arguments.
- Rename `map_filter_empty_iff` to `map_empty_filter` and add - Rename `map_filter_empty_iff` to `map_empty_filter` and add
`map_empty_filter_1` and `map_empty_filter_2`. (by Michael Sammler) `map_empty_filter_1` and `map_empty_filter_2`. (by Michael Sammler)
- Add lemma about `zip_with`: `lookup_zip_with_None` and add lemmas for `zip`: - Add lemma about `zip_with`: `lookup_zip_with_None` and add lemmas for `zip`:
`length_zip`, `zip_nil_inv`, `lookup_zip_Some`,`lookup_zip_None`. (by Kimaya Bedarkar) `length_zip`, `zip_nil_inv`, `lookup_zip_Some`,`lookup_zip_None`. (by Kimaya Bedarkar)
- Add `elem_of_seq` and `seq_nil`. (by Kimaya Bedarkar)
- Add lemmas `StronglySorted_app`, `StronglySorted_cons` and
`StronglySorted_app_2`. Rename lemmas
`elem_of_StronglySorted_app``StronglySorted_app_1_elem_of`,
`StronglySorted_app_inv_l``StronglySorted_app_1_l`
`StronglySorted_app_inv_r``StronglySorted_app_1_r`. (by Marijn van Wezel)
- Add `SetUnfoldElemOf` instance for `dom` on `gmultiset`. (by Marijn van Wezel)
- Split up `stdpp.list` into smaller files. Users should keep importing
`stdpp.list`; the organization into smaller modules is considered an
implementation detail. `stdpp.list_numbers` is now re-exported by `stdpp.list`
and also considered part of the list module, so existing imports should be
updated.
- Remove `list_remove` and `list_remove_list`. There were no lemmas about these
functions; they existed solely to facilitate the proofs of decidability of
`submseteq` and `≡ₚ`, which have been refactored.
The following `sed` script should perform most of the renaming The following `sed` script should perform most of the renaming
(on macOS, replace `sed` by `gsed`, installed via e.g. `brew install gnu-sed`). (on macOS, replace `sed` by `gsed`, installed via e.g. `brew install gnu-sed`).
...@@ -21,6 +34,9 @@ Note that the script is not idempotent, do not run it twice. ...@@ -21,6 +34,9 @@ Note that the script is not idempotent, do not run it twice.
sed -i -E -f- $(find theories -name "*.v") <<EOF sed -i -E -f- $(find theories -name "*.v") <<EOF
# length # length
s/\bmap_filter_empty_iff\b/map_empty_filter/g s/\bmap_filter_empty_iff\b/map_empty_filter/g
# StronglySorted
s/\belem_of_StronglySorted_app\b/StronglySorted_app_1_elem_of/g
s/\bStronglySorted_app_inv_(l|r)\b/StronglySorted_app_1_\1/g
EOF EOF
``` ```
......
...@@ -52,6 +52,11 @@ stdpp/lexico.v ...@@ -52,6 +52,11 @@ stdpp/lexico.v
stdpp/propset.v stdpp/propset.v
stdpp/decidable.v stdpp/decidable.v
stdpp/list.v stdpp/list.v
stdpp/list_basics.v
stdpp/list_relations.v
stdpp/list_monad.v
stdpp/list_misc.v
stdpp/list_tactics.v
stdpp/list_numbers.v stdpp/list_numbers.v
stdpp/functions.v stdpp/functions.v
stdpp/hlist.v stdpp/hlist.v
......
...@@ -159,6 +159,13 @@ Section basic_lemmas. ...@@ -159,6 +159,13 @@ Section basic_lemmas.
Global Instance gmultiset_elem_of_dec : RelDecision (∈@{gmultiset A}). Global Instance gmultiset_elem_of_dec : RelDecision (∈@{gmultiset A}).
Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined. Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined.
Lemma gmultiset_elem_of_dom x X : x dom X x X.
Proof.
unfold dom, gmultiset_dom, elem_of at 2, gmultiset_elem_of, multiplicity.
destruct X as [X]; simpl; rewrite elem_of_dom, <-not_eq_None_Some.
destruct (X !! x); naive_solver lia.
Qed.
End basic_lemmas. End basic_lemmas.
(** * A solver for multisets *) (** * A solver for multisets *)
...@@ -299,6 +306,9 @@ Section multiset_unfold. ...@@ -299,6 +306,9 @@ Section multiset_unfold.
intros ??; constructor. rewrite gmultiset_elem_of_intersection. intros ??; constructor. rewrite gmultiset_elem_of_intersection.
by rewrite (set_unfold_elem_of x X P), (set_unfold_elem_of x Y Q). by rewrite (set_unfold_elem_of x X P), (set_unfold_elem_of x Y Q).
Qed. Qed.
Global Instance set_unfold_gmultiset_dom x X :
SetUnfoldElemOf x (dom X) (x X).
Proof. constructor. apply gmultiset_elem_of_dom. Qed.
End multiset_unfold. End multiset_unfold.
(** Step 3: instantiate hypotheses *) (** Step 3: instantiate hypotheses *)
...@@ -554,12 +564,6 @@ Section more_lemmas. ...@@ -554,12 +564,6 @@ Section more_lemmas.
exists (x,n); split; [|by apply elem_of_map_to_list]. exists (x,n); split; [|by apply elem_of_map_to_list].
apply elem_of_replicate; auto with lia. apply elem_of_replicate; auto with lia.
Qed. Qed.
Lemma gmultiset_elem_of_dom x X : x dom X x X.
Proof.
unfold dom, gmultiset_dom, elem_of at 2, gmultiset_elem_of, multiplicity.
destruct X as [X]; simpl; rewrite elem_of_dom, <-not_eq_None_Some.
destruct (X !! x); naive_solver lia.
Qed.
(** Properties of the set_fold operation *) (** Properties of the set_fold operation *)
Lemma gmultiset_set_fold_empty {B} (f : A B B) (b : B) : Lemma gmultiset_set_fold_empty {B} (f : A B B) (b : B) :
......
This diff is collapsed.
This diff is collapsed.
This diff is collapsed.
This diff is collapsed.
(** This file collects general purpose definitions and theorems on (** This file collects general purpose definitions and theorems on
lists of numbers that are not in the Coq standard library. *) lists of numbers that are not in the Coq standard library. *)
From stdpp Require Export list. From stdpp Require Export list_basics list_monad list_misc list_tactics.
From stdpp Require Import options. From stdpp Require Import options.
(** * Definitions *) (** * Definitions *)
...@@ -98,6 +98,12 @@ Section seq. ...@@ -98,6 +98,12 @@ Section seq.
k seq j n j k < j + n. k seq j n j k < j + n.
Proof. rewrite elem_of_list_In, in_seq. done. Qed. Proof. rewrite elem_of_list_In, in_seq. done. Qed.
Lemma seq_nil n m : seq n m = [] m = 0.
Proof. by destruct m. Qed.
Lemma seq_subseteq_mono m n1 n2 : n1 n2 seq m n1 seq m n2.
Proof. by intros Hle i Hi%elem_of_seq; apply elem_of_seq; lia. Qed.
Lemma Forall_seq (P : nat Prop) i n : Lemma Forall_seq (P : nat Prop) i n :
Forall P (seq i n) j, i j < i + n P j. Forall P (seq i n) j, i j < i + n P j.
Proof. rewrite Forall_forall. setoid_rewrite elem_of_seq. auto with lia. Qed. Proof. rewrite Forall_forall. setoid_rewrite elem_of_seq. auto with lia. Qed.
...@@ -118,6 +124,7 @@ Section seq. ...@@ -118,6 +124,7 @@ Section seq.
- rewrite take_nil. replace (m `min` 0) with 0 by lia. done. - rewrite take_nil. replace (m `min` 0) with 0 by lia. done.
- destruct m; simpl; auto with f_equal. - destruct m; simpl; auto with f_equal.
Qed. Qed.
End seq. End seq.
(** ** Properties of the [seqZ] function *) (** ** Properties of the [seqZ] function *)
......
This diff is collapsed.
From Coq Require Export Permutation.
From stdpp Require Export numbers base option list_basics list_relations list_monad.
From stdpp Require Import options.
(** * Reflection over lists *)
(** We define a simple data structure [rlist] to capture a syntactic
representation of lists consisting of constants, applications and the nil list.
Note that we represent [(x ::.)] as [rapp (rnode [x])]. For now, we abstract
over the type of constants, but later we use [nat]s and a list representing
a corresponding environment. *)
Inductive rlist (A : Type) :=
rnil : rlist A | rnode : A rlist A | rapp : rlist A rlist A rlist A.
Global Arguments rnil {_} : assert.
Global Arguments rnode {_} _ : assert.
Global Arguments rapp {_} _ _ : assert.
Module rlist.
Fixpoint to_list {A} (t : rlist A) : list A :=
match t with
| rnil => [] | rnode l => [l] | rapp t1 t2 => to_list t1 ++ to_list t2
end.
Notation env A := (list (list A)) (only parsing).
Definition eval {A} (E : env A) : rlist nat list A :=
fix go t :=
match t with
| rnil => []
| rnode i => default [] (E !! i)
| rapp t1 t2 => go t1 ++ go t2
end.
(** A simple quoting mechanism using type classes. [QuoteLookup E1 E2 x i]
means: starting in environment [E1], look up the index [i] corresponding to the
constant [x]. In case [x] has a corresponding index [i] in [E1], the original
environment is given back as [E2]. Otherwise, the environment [E2] is extended
with a binding [i] for [x]. *)
Section quote_lookup.
Context {A : Type}.
Class QuoteLookup (E1 E2 : list A) (x : A) (i : nat) := {}.
Global Instance quote_lookup_here E x : QuoteLookup (x :: E) (x :: E) x 0 := {}.
Global Instance quote_lookup_end x : QuoteLookup [] [x] x 0 := {}.
Global Instance quote_lookup_further E1 E2 x i y :
QuoteLookup E1 E2 x i QuoteLookup (y :: E1) (y :: E2) x (S i) | 1000 := {}.
End quote_lookup.
Section quote.
Context {A : Type}.
Class Quote (E1 E2 : env A) (l : list A) (t : rlist nat) := {}.
Global Instance quote_nil E1 : Quote E1 E1 [] rnil := {}.
Global Instance quote_node E1 E2 l i:
QuoteLookup E1 E2 l i Quote E1 E2 l (rnode i) | 1000 := {}.
Global Instance quote_cons E1 E2 E3 x l i t :
QuoteLookup E1 E2 [x] i
Quote E2 E3 l t Quote E1 E3 (x :: l) (rapp (rnode i) t) := {}.
Global Instance quote_app E1 E2 E3 l1 l2 t1 t2 :
Quote E1 E2 l1 t1 Quote E2 E3 l2 t2 Quote E1 E3 (l1 ++ l2) (rapp t1 t2) := {}.
End quote.
Section eval.
Context {A} (E : env A).
Lemma eval_alt t : eval E t = to_list t ≫= default [] (E !!.).
Proof.
induction t; csimpl.
- done.
- by rewrite (right_id_L [] (++)).
- rewrite bind_app. by f_equal.
Qed.
Lemma eval_eq t1 t2 : to_list t1 = to_list t2 eval E t1 = eval E t2.
Proof. intros Ht. by rewrite !eval_alt, Ht. Qed.
Lemma eval_Permutation t1 t2 :
to_list t1 to_list t2 eval E t1 eval E t2.
Proof. intros Ht. by rewrite !eval_alt, Ht. Qed.
Lemma eval_submseteq t1 t2 :
to_list t1 ⊆+ to_list t2 eval E t1 ⊆+ eval E t2.
Proof. intros Ht. by rewrite !eval_alt, Ht. Qed.
End eval.
End rlist.
(** * Tactics *)
Ltac quote_Permutation :=
match goal with
| |- ?l1 ?l2 =>
match type of (_ : rlist.Quote [] _ l1 _) with rlist.Quote _ ?E2 _ ?t1 =>
match type of (_ : rlist.Quote E2 _ l2 _) with rlist.Quote _ ?E3 _ ?t2 =>
change (rlist.eval E3 t1 rlist.eval E3 t2)
end end
end.
Ltac solve_Permutation :=
quote_Permutation; apply rlist.eval_Permutation;
compute_done.
Ltac quote_submseteq :=
match goal with
| |- ?l1 ⊆+ ?l2 =>
match type of (_ : rlist.Quote [] _ l1 _) with rlist.Quote _ ?E2 _ ?t1 =>
match type of (_ : rlist.Quote E2 _ l2 _) with rlist.Quote _ ?E3 _ ?t2 =>
change (rlist.eval E3 t1 ⊆+ rlist.eval E3 t2)
end end
end.
Ltac solve_submseteq :=
quote_submseteq; apply rlist.eval_submseteq;
compute_done.
Ltac decompose_elem_of_list := repeat
match goal with
| H : ?x [] |- _ => by destruct (not_elem_of_nil x)
| H : _ _ :: _ |- _ => apply elem_of_cons in H; destruct H
| H : _ _ ++ _ |- _ => apply elem_of_app in H; destruct H
end.
Ltac solve_length :=
simplify_eq/=;
repeat (rewrite length_fmap || rewrite length_app);
repeat match goal with
| H : _ =@{list _} _ |- _ => apply (f_equal length) in H
| H : Forall2 _ _ _ |- _ => apply Forall2_length in H
| H : context[length (_ <$> _)] |- _ => rewrite length_fmap in H
end; done || congruence.
Ltac simplify_list_eq ::= repeat
match goal with
| _ => progress simplify_eq/=
| H : [?x] !! ?i = Some ?y |- _ =>
destruct i; [change (Some x = Some y) in H | discriminate]
| H : _ <$> _ = [] |- _ => apply fmap_nil_inv in H
| H : [] = _ <$> _ |- _ => symmetry in H; apply fmap_nil_inv in H
| H : zip_with _ _ _ = [] |- _ => apply zip_with_nil_inv in H; destruct H
| H : [] = zip_with _ _ _ |- _ => symmetry in H
| |- context [(_ ++ _) ++ _] => rewrite <-(assoc_L (++))
| H : context [(_ ++ _) ++ _] |- _ => rewrite <-(assoc_L (++)) in H
| H : context [_ <$> (_ ++ _)] |- _ => rewrite fmap_app in H
| |- context [_ <$> (_ ++ _)] => rewrite fmap_app
| |- context [_ ++ []] => rewrite (right_id_L [] (++))
| H : context [_ ++ []] |- _ => rewrite (right_id_L [] (++)) in H
| |- context [take _ (_ <$> _)] => rewrite <-fmap_take
| H : context [take _ (_ <$> _)] |- _ => rewrite <-fmap_take in H
| |- context [drop _ (_ <$> _)] => rewrite <-fmap_drop
| H : context [drop _ (_ <$> _)] |- _ => rewrite <-fmap_drop in H
| H : _ ++ _ = _ ++ _ |- _ =>
repeat (rewrite <-app_comm_cons in H || rewrite <-(assoc_L (++)) in H);
apply app_inj_1 in H; [destruct H|solve_length]
| H : _ ++ _ = _ ++ _ |- _ =>
repeat (rewrite app_comm_cons in H || rewrite (assoc_L (++)) in H);
apply app_inj_2 in H; [destruct H|solve_length]
| |- context [zip_with _ (_ ++ _) (_ ++ _)] =>
rewrite zip_with_app by solve_length
| |- context [take _ (_ ++ _)] => rewrite take_app_length' by solve_length
| |- context [drop _ (_ ++ _)] => rewrite drop_app_length' by solve_length
| H : context [zip_with _ (_ ++ _) (_ ++ _)] |- _ =>
rewrite zip_with_app in H by solve_length
| H : context [take _ (_ ++ _)] |- _ =>
rewrite take_app_length' in H by solve_length
| H : context [drop _ (_ ++ _)] |- _ =>
rewrite drop_app_length' in H by solve_length
| H : ?l !! ?i = _, H2 : context [(_ <$> ?l) !! ?i] |- _ =>
rewrite list_lookup_fmap, H in H2
end.
Ltac decompose_Forall_hyps :=
repeat match goal with
| H : Forall _ [] |- _ => clear H
| H : Forall _ (_ :: _) |- _ => rewrite Forall_cons in H; destruct H
| H : Forall _ (_ ++ _) |- _ => rewrite Forall_app in H; destruct H
| H : Forall2 _ [] [] |- _ => clear H
| H : Forall2 _ (_ :: _) [] |- _ => destruct (Forall2_cons_nil_inv _ _ _ H)
| H : Forall2 _ [] (_ :: _) |- _ => destruct (Forall2_nil_cons_inv _ _ _ H)
| H : Forall2 _ [] ?k |- _ => apply Forall2_nil_inv_l in H
| H : Forall2 _ ?l [] |- _ => apply Forall2_nil_inv_r in H
| H : Forall2 _ (_ :: _) (_ :: _) |- _ =>
apply Forall2_cons_1 in H; destruct H
| H : Forall2 _ (_ :: _) ?k |- _ =>
let k_hd := fresh k "_hd" in let k_tl := fresh k "_tl" in
apply Forall2_cons_inv_l in H; destruct H as (k_hd&k_tl&?&?&->);
rename k_tl into k
| H : Forall2 _ ?l (_ :: _) |- _ =>
let l_hd := fresh l "_hd" in let l_tl := fresh l "_tl" in
apply Forall2_cons_inv_r in H; destruct H as (l_hd&l_tl&?&?&->);
rename l_tl into l
| H : Forall2 _ (_ ++ _) ?k |- _ =>
let k1 := fresh k "_1" in let k2 := fresh k "_2" in
apply Forall2_app_inv_l in H; destruct H as (k1&k2&?&?&->)
| H : Forall2 _ ?l (_ ++ _) |- _ =>
let l1 := fresh l "_1" in let l2 := fresh l "_2" in
apply Forall2_app_inv_r in H; destruct H as (l1&l2&?&?&->)
| _ => progress simplify_eq/=
| H : Forall3 _ _ (_ :: _) _ |- _ =>
apply Forall3_cons_inv_m in H; destruct H as (?&?&?&?&?&?&?&?)
| H : Forall2 _ (_ :: _) ?k |- _ =>
apply Forall2_cons_inv_l in H; destruct H as (?&?&?&?&?)
| H : Forall2 _ ?l (_ :: _) |- _ =>
apply Forall2_cons_inv_r in H; destruct H as (?&?&?&?&?)
| H : Forall2 _ (_ ++ _) (_ ++ _) |- _ =>
apply Forall2_app_inv in H; [destruct H|solve_length]
| H : Forall2 _ ?l (_ ++ _) |- _ =>
apply Forall2_app_inv_r in H; destruct H as (?&?&?&?&?)
| H : Forall2 _ (_ ++ _) ?k |- _ =>
apply Forall2_app_inv_l in H; destruct H as (?&?&?&?&?)
| H : Forall3 _ _ (_ ++ _) _ |- _ =>
apply Forall3_app_inv_m in H; destruct H as (?&?&?&?&?&?&?&?)
| H : Forall ?P ?l, H1 : ?l !! _ = Some ?x |- _ =>
(* to avoid some stupid loops, not fool proof *)
unless (P x) by auto using Forall_app_2, Forall_nil_2;
let E := fresh in
assert (P x) as E by (apply (Forall_lookup_1 P _ _ _ H H1)); lazy beta in E
| H : Forall2 ?P ?l ?k |- _ =>
match goal with
| H1 : l !! ?i = Some ?x, H2 : k !! ?i = Some ?y |- _ =>
unless (P x y) by done; let E := fresh in
assert (P x y) as E by (by apply (Forall2_lookup_lr P l k i x y));
lazy beta in E
| H1 : l !! ?i = Some ?x |- _ =>
try (match goal with _ : k !! i = Some _ |- _ => fail 2 end);
destruct (Forall2_lookup_l P _ _ _ _ H H1) as (?&?&?)
| H2 : k !! ?i = Some ?y |- _ =>
try (match goal with _ : l !! i = Some _ |- _ => fail 2 end);
destruct (Forall2_lookup_r P _ _ _ _ H H2) as (?&?&?)
end
| H : Forall3 ?P ?l ?l' ?k |- _ =>
lazymatch goal with
| H1:l !! ?i = Some ?x, H2:l' !! ?i = Some ?y, H3:k !! ?i = Some ?z |- _ =>
unless (P x y z) by done; let E := fresh in
assert (P x y z) as E by (by apply (Forall3_lookup_lmr P l l' k i x y z));
lazy beta in E
| H1 : l !! _ = Some ?x |- _ =>
destruct (Forall3_lookup_l P _ _ _ _ _ H H1) as (?&?&?&?&?)
| H2 : l' !! _ = Some ?y |- _ =>
destruct (Forall3_lookup_m P _ _ _ _ _ H H2) as (?&?&?&?&?)
| H3 : k !! _ = Some ?z |- _ =>
destruct (Forall3_lookup_r P _ _ _ _ _ H H3) as (?&?&?&?&?)
end
end.
Ltac list_simplifier :=
simplify_eq/=;
repeat match goal with
| _ => progress decompose_Forall_hyps
| _ => progress simplify_list_eq
| H : _ <$> _ = _ :: _ |- _ =>
apply fmap_cons_inv in H; destruct H as (?&?&?&?&?)
| H : _ :: _ = _ <$> _ |- _ => symmetry in H
| H : _ <$> _ = _ ++ _ |- _ =>
apply fmap_app_inv in H; destruct H as (?&?&?&?&?)
| H : _ ++ _ = _ <$> _ |- _ => symmetry in H
| H : zip_with _ _ _ = _ :: _ |- _ =>
apply zip_with_cons_inv in H; destruct H as (?&?&?&?&?&?&?&?)
| H : _ :: _ = zip_with _ _ _ |- _ => symmetry in H
| H : zip_with _ _ _ = _ ++ _ |- _ =>
apply zip_with_app_inv in H; destruct H as (?&?&?&?&?&?&?&?&?)
| H : _ ++ _ = zip_with _ _ _ |- _ => symmetry in H
end.
Ltac decompose_Forall := repeat
match goal with
| |- Forall _ _ => by apply Forall_true
| |- Forall _ [] => constructor
| |- Forall _ (_ :: _) => constructor
| |- Forall _ (_ ++ _) => apply Forall_app_2
| |- Forall _ (_ <$> _) => apply Forall_fmap
| |- Forall _ (_ ≫= _) => apply Forall_bind
| |- Forall2 _ _ _ => apply Forall_Forall2_diag
| |- Forall2 _ [] [] => constructor
| |- Forall2 _ (_ :: _) (_ :: _) => constructor
| |- Forall2 _ (_ ++ _) (_ ++ _) => first
[ apply Forall2_app; [by decompose_Forall |]
| apply Forall2_app; [| by decompose_Forall]]
| |- Forall2 _ (_ <$> _) _ => apply Forall2_fmap_l
| |- Forall2 _ _ (_ <$> _) => apply Forall2_fmap_r
| _ => progress decompose_Forall_hyps
| H : Forall _ (_ <$> _) |- _ => rewrite Forall_fmap in H
| H : Forall _ (_ ≫= _) |- _ => rewrite Forall_bind in H
| |- Forall _ _ =>
apply Forall_lookup_2; intros ???; progress decompose_Forall_hyps
| |- Forall2 _ _ _ =>
apply Forall2_same_length_lookup_2; [solve_length|];
intros ?????; progress decompose_Forall_hyps
end.
(** The [simplify_suffix] tactic removes [suffix] hypotheses that are
tautologies, and simplifies [suffix] hypotheses involving [(::)] and
[(++)]. *)
Ltac simplify_suffix := repeat
match goal with
| H : suffix (_ :: _) _ |- _ => destruct (suffix_cons_not _ _ H)
| H : suffix (_ :: _) [] |- _ => apply suffix_nil_inv in H
| H : suffix (_ ++ _) (_ ++ _) |- _ => apply suffix_app_inv in H
| H : suffix (_ :: _) (_ :: _) |- _ =>
destruct (suffix_cons_inv _ _ _ _ H); clear H
| H : suffix ?x ?x |- _ => clear H
| H : suffix ?x (_ :: ?x) |- _ => clear H
| H : suffix ?x (_ ++ ?x) |- _ => clear H
| _ => progress simplify_eq/=
end.
(** The [solve_suffix] tactic tries to solve goals involving [suffix]. It
uses [simplify_suffix] to simplify hypotheses and tries to solve [suffix]
conclusions. This tactic either fails or proves the goal. *)
Ltac solve_suffix := by intuition (repeat
match goal with
| _ => done
| _ => progress simplify_suffix
| |- suffix [] _ => apply suffix_nil
| |- suffix _ _ => reflexivity
| |- suffix _ (_ :: _) => apply suffix_cons_r
| |- suffix _ (_ ++ _) => apply suffix_app_r
| H : suffix _ _ False |- _ => destruct H
end).
...@@ -2,6 +2,7 @@ ...@@ -2,6 +2,7 @@
standard library, but without using the module system. *) standard library, but without using the module system. *)
From Coq Require Export Sorted. From Coq Require Export Sorted.
From stdpp Require Export orders list. From stdpp Require Export orders list.
From stdpp Require Import sets.
From stdpp Require Import options. From stdpp Require Import options.
Section merge_sort. Section merge_sort.
...@@ -48,25 +49,36 @@ Inductive TlRel {A} (R : relation A) (a : A) : list A → Prop := ...@@ -48,25 +49,36 @@ Inductive TlRel {A} (R : relation A) (a : A) : list A → Prop :=
Section sorted. Section sorted.
Context {A} (R : relation A). Context {A} (R : relation A).
Lemma elem_of_StronglySorted_app l1 l2 x1 x2 : Lemma StronglySorted_cons l x :
StronglySorted R (l1 ++ l2) x1 l1 x2 l2 R x1 x2. StronglySorted R (x :: l)
Forall (R x) l StronglySorted R l.
Proof. split; [inv 1|constructor]; naive_solver. Qed.
Lemma StronglySorted_app l1 l2 :
StronglySorted R (l1 ++ l2)
( x1 x2, x1 l1 x2 l2 R x1 x2)
StronglySorted R l1
StronglySorted R l2.
Proof. Proof.
induction l1 as [|x1' l1 IH]; simpl; [by rewrite elem_of_nil|]. induction l1 as [|x1' l1 IH]; simpl.
intros [? Hall]%StronglySorted_inv [->|?]%elem_of_cons ?; [|by auto]. - set_solver by eauto using SSorted_nil.
rewrite Forall_app, !Forall_forall in Hall. naive_solver. - rewrite !StronglySorted_cons, IH, !Forall_forall. set_solver.
Qed. Qed.
Lemma StronglySorted_app_inv_l l1 l2 : Lemma StronglySorted_app_2 l1 l2 :
( x1 x2, x1 l1 x2 l2 R x1 x2)
StronglySorted R l1
StronglySorted R l2
StronglySorted R (l1 ++ l2).
Proof. by rewrite StronglySorted_app. Qed.
Lemma StronglySorted_app_1_elem_of l1 l2 x1 x2 :
StronglySorted R (l1 ++ l2) x1 l1 x2 l2 R x1 x2.
Proof. rewrite StronglySorted_app. naive_solver. Qed.
Lemma StronglySorted_app_1_l l1 l2 :
StronglySorted R (l1 ++ l2) StronglySorted R l1. StronglySorted R (l1 ++ l2) StronglySorted R l1.
Proof. Proof. rewrite StronglySorted_app. naive_solver. Qed.
induction l1 as [|x1' l1 IH]; simpl; Lemma StronglySorted_app_1_r l1 l2 :
[|inv 1]; decompose_Forall; constructor; auto.
Qed.
Lemma StronglySorted_app_inv_r l1 l2 :
StronglySorted R (l1 ++ l2) StronglySorted R l2. StronglySorted R (l1 ++ l2) StronglySorted R l2.
Proof. Proof. rewrite StronglySorted_app. naive_solver. Qed.
induction l1 as [|x1' l1 IH]; simpl;
[|inv 1]; decompose_Forall; auto.
Qed.
Lemma Sorted_StronglySorted `{!Transitive R} l : Lemma Sorted_StronglySorted `{!Transitive R} l :
Sorted R l StronglySorted R l. Sorted R l StronglySorted R l.
......
...@@ -738,19 +738,6 @@ Tactic Notation "ospecialize" "*" uconstr(p) := ...@@ -738,19 +738,6 @@ Tactic Notation "ospecialize" "*" uconstr(p) :=
pose proof p as H'; clear H; rename H' into H pose proof p as H'; clear H; rename H' into H
)). )).
(** Tactic that works like [notypeclasses refine] but automatically determines
the right number of [_]. This is *not* goal-directed, it will add [_] until the
given term no longer has a function type (determined without any unfolding). *)
Tactic Notation "notypeclasses" "apply" uconstr(p) :=
opose_core p ltac:(fun p => evar_foralls p () ltac:(fun t => t)
ltac:(fun p => notypeclasses refine p ||
let T := type of p in
lazymatch goal with |- ?G =>
fail "the given term has type" T "while it is expected to have type" G
end
)
).
(** The block definitions are taken from [Coq.Program.Equality] and can be used (** The block definitions are taken from [Coq.Program.Equality] and can be used
by tactics to separate their goal from hypotheses they generalize over. *) by tactics to separate their goal from hypotheses they generalize over. *)
Definition block {A : Type} (a : A) := a. Definition block {A : Type} (a : A) := a.
......
...@@ -755,7 +755,7 @@ Global Hint Resolve bv_wrap_simplify_mul | 10 : bv_wrap_simplify_db. ...@@ -755,7 +755,7 @@ Global Hint Resolve bv_wrap_simplify_mul | 10 : bv_wrap_simplify_db.
tries to simplify them by removing any [bv_wrap] inside z1. *) tries to simplify them by removing any [bv_wrap] inside z1. *)
Ltac bv_wrap_simplify_left := Ltac bv_wrap_simplify_left :=
lazymatch goal with |- bv_wrap _ _ = _ => idtac end; lazymatch goal with |- bv_wrap _ _ = _ => idtac end;
etrans; [ notypeclasses apply bv_wrap_simplify_proof; etrans; [ notypeclasses refine (bv_wrap_simplify_proof _ _ _);
typeclasses eauto with bv_wrap_simplify_db | ] typeclasses eauto with bv_wrap_simplify_db | ]
. .
......
...@@ -58,6 +58,9 @@ Section test. ...@@ -58,6 +58,9 @@ Section test.
Lemma test_elem_of_6 x y X : {[+ x; y +]} X x X y X. Lemma test_elem_of_6 x y X : {[+ x; y +]} X x X y X.
Proof. multiset_solver. Qed. Proof. multiset_solver. Qed.
Lemma test_elem_of_dom x X : x dom X x X.
Proof. multiset_solver. Qed.
(** Tests where the goals do not involve the multiset connectives *) (** Tests where the goals do not involve the multiset connectives *)
Lemma test_goal_1 x y X : {[+ x +]} X {[+ y +]} x = y. Lemma test_goal_1 x y X : {[+ x +]} X {[+ y +]} x = y.
Proof. multiset_solver. Qed. Proof. multiset_solver. Qed.
......
...@@ -65,22 +65,9 @@ use opose proof instead. ...@@ -65,22 +65,9 @@ use opose proof instead.
H' : even n H' : even n
============================ ============================
even n even n
"notypeclasses_apply_test"
: string
1 goal
n : nat
============================
Decision (n = 0)
1 goal
n : nat
============================
Decision (n = 0)
"notypeclasses_apply_error" "notypeclasses_apply_error"
: string : string
The command has indeed failed with message: The command has indeed failed with message:
Tactic failure: the given term has type Q while it is expected to have type The reference notypeclasses was not found in the current environment.
P.
The command has indeed failed with message: The command has indeed failed with message:
No such assumption. No such assumption.
...@@ -234,16 +234,6 @@ Proof. ...@@ -234,16 +234,6 @@ Proof.
done. done.
Qed. Qed.
Check "notypeclasses_apply_test".
Lemma notypeclasses_apply_test n : {n = 0} + {¬ n = 0}.
Proof.
(* Partially applied lemma with typeclass in one of the trailing [_]. *)
notypeclasses apply (@decide (n = 0)). Show.
Restart. Proof.
(* Partially applied lemma with typeclass in one of the explicit [_]. *)
notypeclasses apply (@decide (n = 0) _). Show.
Abort.
Check "notypeclasses_apply_error". Check "notypeclasses_apply_error".
Lemma notypeclasses_apply_error (P Q : Prop) : (P Q) Q P. Lemma notypeclasses_apply_error (P Q : Prop) : (P Q) Q P.
Proof. Proof.
......