diff --git a/theories/base.v b/theories/base.v index 51dcdad2f233e0fa907ac48a321b8dd648e4044d..859c6f009c545ee34b790e35bda610872105c3d9 100644 --- a/theories/base.v +++ b/theories/base.v @@ -9,6 +9,7 @@ Global Set Automatic Coercions Import. Global Set Asymmetric Patterns. Global Unset Transparent Obligations. From Coq Require Export Morphisms RelationClasses List Bool Utf8 Setoid. +Set Default Proof Using "Type*". Export ListNotations. From Coq.Program Require Export Basics Syntax. Obligation Tactic := idtac. diff --git a/theories/bset.v b/theories/bset.v index 9c8fe4753c8c47a3a3eda31e1693ef2521f5429c..91501e33703fdda2297814a3d9f5ea63124cf7d9 100644 --- a/theories/bset.v +++ b/theories/bset.v @@ -2,6 +2,7 @@ (* This file is distributed under the terms of the BSD license. *) (** This file implements bsets as functions into Prop. *) From stdpp Require Export prelude. +Set Default Proof Using "Type*". Record bset (A : Type) : Type := mkBSet { bset_car : A → bool }. Arguments mkBSet {_} _. diff --git a/theories/coPset.v b/theories/coPset.v index c1212ffd1d8fc5abd16ce2a627115f2bcbaf401e..49e323349baf4cb6a68ec7123834a6a5fe16e337 100644 --- a/theories/coPset.v +++ b/theories/coPset.v @@ -13,6 +13,7 @@ Since [positive]s are bitstrings, we encode [coPset]s as trees that correspond to the decision function that map bitstrings to bools. *) From stdpp Require Export collections. From stdpp Require Import pmap gmap mapset. +Set Default Proof Using "Type*". Local Open Scope positive_scope. (** * The tree data structure *) diff --git a/theories/collections.v b/theories/collections.v index 126dc59e1ea18bbf692c658facd74c651da09160..672eac8c8c208092ef2e31727dadc150a79cc6e9 100644 --- a/theories/collections.v +++ b/theories/collections.v @@ -4,6 +4,7 @@ importantly, it implements some tactics to automatically solve goals involving collections. *) From stdpp Require Export orders list. +Set Default Proof Using "Type*". Instance collection_equiv `{ElemOf A C} : Equiv C := λ X Y, ∀ x, x ∈ X ↔ x ∈ Y. diff --git a/theories/countable.v b/theories/countable.v index 00e84a9eace9ba34d950be6b6b002791733d8687..aa7dad648869048d62f144a4b52af7922ec4b758 100644 --- a/theories/countable.v +++ b/theories/countable.v @@ -1,6 +1,7 @@ (* Copyright (c) 2012-2015, Robbert Krebbers. *) (* This file is distributed under the terms of the BSD license. *) From stdpp Require Export list. +Set Default Proof Using "Type*". Local Open Scope positive. Class Countable A `{EqDecision A} := { diff --git a/theories/decidable.v b/theories/decidable.v index 4748a82201b7922bbf347084767bc86003bc085b..57dc031f3a62f74df47f68f9dcc6032070f33fa8 100644 --- a/theories/decidable.v +++ b/theories/decidable.v @@ -4,6 +4,7 @@ with a decidable equality. Such propositions are collected by the [Decision] type class. *) From stdpp Require Export proof_irrel. +Set Default Proof Using "Type*". Hint Extern 200 (Decision _) => progress (lazy beta) : typeclass_instances. diff --git a/theories/fin_collections.v b/theories/fin_collections.v index ed094d36eb2b7aa8599d03e60e3b3aab891ab775..fd7a4d38ea2ce08041a439e7d68288e2c4f71cce 100644 --- a/theories/fin_collections.v +++ b/theories/fin_collections.v @@ -6,6 +6,7 @@ principles on finite collections . *) From Coq Require Import Permutation. From stdpp Require Import relations listset. From stdpp Require Export numbers collections. +Set Default Proof Using "Type*". Instance collection_size `{Elements A C} : Size C := length ∘ elements. Definition collection_fold `{Elements A C} {B} diff --git a/theories/fin_map_dom.v b/theories/fin_map_dom.v index 5856b682d344e88cbb409c0c9e5d4334cf9f727d..cbc6c1a1015b601cf87f18fc5ced9811e58ed91e 100644 --- a/theories/fin_map_dom.v +++ b/theories/fin_map_dom.v @@ -4,6 +4,7 @@ maps. We provide such an axiomatization, instead of implementing the domain function in a generic way, to allow more efficient implementations. *) From stdpp Require Export collections fin_maps. +Set Default Proof Using "Type*". Class FinMapDom K M D `{FMap M, ∀ A, Lookup K A (M A), ∀ A, Empty (M A), ∀ A, PartialAlter K A (M A), diff --git a/theories/fin_maps.v b/theories/fin_maps.v index 2c8634da993b7667842c6fa07f753f6a9db829d0..cc85921f2f41792e5d4c8f6dd0b91d3d968f6764 100644 --- a/theories/fin_maps.v +++ b/theories/fin_maps.v @@ -6,6 +6,7 @@ induction principles for finite maps and implements the tactic [simplify_map_eq] to simplify goals involving finite maps. *) From Coq Require Import Permutation. From stdpp Require Export relations orders vector. +Set Default Proof Using "Type*". (** * Axiomatization of finite maps *) (** We require Leibniz equality to be extensional on finite maps. This of diff --git a/theories/finite.v b/theories/finite.v index eb8b868eec3e4188c48c13adaa7ffdc92a2f73a3..d15fa39e85ad3aabfef227b1384c56f309aec63f 100644 --- a/theories/finite.v +++ b/theories/finite.v @@ -1,6 +1,7 @@ (* Copyright (c) 2012-2015, Robbert Krebbers. *) (* This file is distributed under the terms of the BSD license. *) From stdpp Require Export countable vector. +Set Default Proof Using "Type*". Class Finite A `{EqDecision A} := { enum : list A; diff --git a/theories/functions.v b/theories/functions.v index 44da98377fcf8b828a7de9743ecb462ad47df188..4c37187d2a8146e1bb0dacd986d8ffab3a0b50f8 100644 --- a/theories/functions.v +++ b/theories/functions.v @@ -1,4 +1,5 @@ From stdpp Require Export base tactics. +Set Default Proof Using "Type*". Section definitions. Context {A T : Type} `{EqDecision A}. diff --git a/theories/gmap.v b/theories/gmap.v index 3f4c395836b916d21cf58cb4ca4f987c20288cd9..a1bf157f471bca4c9c295a82c720bbbaa68d94a2 100644 --- a/theories/gmap.v +++ b/theories/gmap.v @@ -4,6 +4,7 @@ type. The implementation is based on [Pmap]s, radix-2 search trees. *) From stdpp Require Export countable fin_maps fin_map_dom. From stdpp Require Import pmap mapset set. +Set Default Proof Using "Type*". (** * The data structure *) (** We pack a [Pmap] together with a proof that ensures that all keys correspond diff --git a/theories/gmultiset.v b/theories/gmultiset.v index 2c9f85356f587d24d268df36167eb3b3e633214c..7f1f83eb762e96661a72fd215a1bc8f52f142abe 100644 --- a/theories/gmultiset.v +++ b/theories/gmultiset.v @@ -1,6 +1,7 @@ (* Copyright (c) 2012-2016, Robbert Krebbers. *) (* This file is distributed under the terms of the BSD license. *) From stdpp Require Import gmap. +Set Default Proof Using "Type*". Record gmultiset A `{Countable A} := GMultiSet { gmultiset_car : gmap A nat }. Arguments GMultiSet {_ _ _} _. diff --git a/theories/hashset.v b/theories/hashset.v index e93e5214f50507a1debfe6dd72f6c52e0b3a51c4..ffed2bd15a4a66b36bbd27d1ddce94f93d57be80 100644 --- a/theories/hashset.v +++ b/theories/hashset.v @@ -5,6 +5,7 @@ using radix-2 search trees. Each hash bucket is thus indexed using an binary integer of type [Z], and contains an unordered list without duplicates. *) From stdpp Require Export fin_maps listset. From stdpp Require Import zmap. +Set Default Proof Using "Type*". Record hashset {A} (hash : A → Z) := Hashset { hashset_car : Zmap (list A); diff --git a/theories/hlist.v b/theories/hlist.v index 6f4c5227b7b1d9468d3791bbcd9e965fd9825ce5..39ad98053bf35cf7dd8ca1633b567eb5b0c53134 100644 --- a/theories/hlist.v +++ b/theories/hlist.v @@ -1,4 +1,5 @@ From stdpp Require Import tactics. +Set Default Proof Using "Type*". Local Set Universe Polymorphism. (* Not using [list Type] in order to avoid universe inconsistencies *) diff --git a/theories/lexico.v b/theories/lexico.v index 42d6b31903ead119b0a1ebfc444a9c700707ab68..f5402ebd63331014aecad124d372c68d0ffe73d7 100644 --- a/theories/lexico.v +++ b/theories/lexico.v @@ -3,6 +3,7 @@ (** This files defines a lexicographic order on various common data structures and proves that it is a partial order having a strong variant of trichotomy. *) From stdpp Require Import numbers. +Set Default Proof Using "Type*". Notation cast_trichotomy T := match T with diff --git a/theories/list.v b/theories/list.v index 5847a12428d30fca31c48ecfadca66a0834d862d..a9f260c4432ac29dbc0000b70d2ba8019ae285d5 100644 --- a/theories/list.v +++ b/theories/list.v @@ -4,6 +4,7 @@ are not in the Coq standard library. *) From Coq Require Export Permutation. From stdpp Require Export numbers base option. +Set Default Proof Using "Type*". Arguments length {_} _. Arguments cons {_} _ _. diff --git a/theories/listset.v b/theories/listset.v index 8ba410eaa1b1b48eca1f35b7aebc83cee8955438..8966cc2e86df5251be0505163e81bdb8e1b3dd77 100644 --- a/theories/listset.v +++ b/theories/listset.v @@ -3,6 +3,7 @@ (** This file implements finite set as unordered lists without duplicates removed. This implementation forms a monad. *) From stdpp Require Export collections list. +Set Default Proof Using "Type*". Record listset A := Listset { listset_car: list A }. Arguments listset_car {_} _. diff --git a/theories/listset_nodup.v b/theories/listset_nodup.v index 73447c3519fd31ceae7421cac4ef63ae31498f07..e3b76d511b0bc4854471608d97d07dca6f08e2b2 100644 --- a/theories/listset_nodup.v +++ b/theories/listset_nodup.v @@ -4,6 +4,7 @@ Although this implementation is slow, it is very useful as decidable equality is the only constraint on the carrier set. *) From stdpp Require Export collections list. +Set Default Proof Using "Type*". Record listset_nodup A := ListsetNoDup { listset_nodup_car : list A; listset_nodup_prf : NoDup listset_nodup_car diff --git a/theories/mapset.v b/theories/mapset.v index ffaae8b4afd4e0b9204d05d0c9bd9dcf4efc19b7..1e921468a848f350ce99cf263fb6bcf12ae544bc 100644 --- a/theories/mapset.v +++ b/theories/mapset.v @@ -4,6 +4,7 @@ elements of the unit type. Since maps enjoy extensional equality, the constructed finite sets do so as well. *) From stdpp Require Export fin_map_dom. +(* FIXME: This file needs a 'Proof Using' hint. *) Record mapset (M : Type → Type) : Type := Mapset { mapset_car: M (unit : Type) }. diff --git a/theories/natmap.v b/theories/natmap.v index 0cd0a4c3c57f071d165fd6120e3cff4f4ff26745..d109590bbc65a0980f430860139a51b2db69f72a 100644 --- a/theories/natmap.v +++ b/theories/natmap.v @@ -4,6 +4,7 @@ over Coq's data type of unary natural numbers [nat]. The implementation equips a list with a proof of canonicity. *) From stdpp Require Import fin_maps mapset. +Set Default Proof Using "Type*". Notation natmap_raw A := (list (option A)). Definition natmap_wf {A} (l : natmap_raw A) := diff --git a/theories/nmap.v b/theories/nmap.v index 00c249b0996ee21b6140da250054e85df05c5994..12645efbaa269de49be4cc9c195d25001a6a5c5f 100644 --- a/theories/nmap.v +++ b/theories/nmap.v @@ -4,6 +4,7 @@ maps whose keys range over Coq's data type of binary naturals [N]. *) From stdpp Require Import pmap mapset. From stdpp Require Export prelude fin_maps. +Set Default Proof Using "Type*". Local Open Scope N_scope. diff --git a/theories/numbers.v b/theories/numbers.v index 38f5a86217647cb76f7d8d982041d359a4019c32..034a4ca14ce014e8e918fadf55eb19fdf28752d1 100644 --- a/theories/numbers.v +++ b/theories/numbers.v @@ -6,6 +6,7 @@ notations. *) From Coq Require Export EqdepFacts PArith NArith ZArith NPeano. From Coq Require Import QArith Qcanon. From stdpp Require Export base decidable option. +Set Default Proof Using "Type*". Open Scope nat_scope. Coercion Z.of_nat : nat >-> Z. diff --git a/theories/option.v b/theories/option.v index 73ae6eff1f69195554bac488c2b11b3e9d6bac4d..63378a36a1136b056b363daadcf614a88c4ed9bf 100644 --- a/theories/option.v +++ b/theories/option.v @@ -3,6 +3,7 @@ (** This file collects general purpose definitions and theorems on the option data type that are not in the Coq standard library. *) From stdpp Require Export tactics. +Set Default Proof Using "Type*". Inductive option_reflect {A} (P : A → Prop) (Q : Prop) : option A → Type := | ReflectSome x : P x → option_reflect P Q (Some x) diff --git a/theories/orders.v b/theories/orders.v index b0d84aeb1d297b30658a1935502e4a5a791d424f..bc301900b0f68d31fe893caaaba650a2e11194e6 100644 --- a/theories/orders.v +++ b/theories/orders.v @@ -3,6 +3,7 @@ (** Properties about arbitrary pre-, partial, and total orders. We do not use the relation [⊆] because we often have multiple orders on the same structure *) From stdpp Require Export tactics. +Set Default Proof Using "Type*". Section orders. Context {A} {R : relation A}. diff --git a/theories/pmap.v b/theories/pmap.v index 8237bdccef4cbdedff419336230c5459817dd990..a6966f39ef92ad83321ac192ca5704a3e310fafc 100644 --- a/theories/pmap.v +++ b/theories/pmap.v @@ -10,6 +10,7 @@ Leibniz equality to become extensional. *) From Coq Require Import PArith. From stdpp Require Import mapset countable. From stdpp Require Export fin_maps. +Set Default Proof Using "Type*". Local Open Scope positive_scope. Local Hint Extern 0 (@eq positive _ _) => congruence. diff --git a/theories/pretty.v b/theories/pretty.v index 58fd38347cb830e869560976e93956b39a6fa73f..afd2b6899b6b6f2925c00aa6a49d34f92450315d 100644 --- a/theories/pretty.v +++ b/theories/pretty.v @@ -3,6 +3,7 @@ From stdpp Require Export strings. From stdpp Require Import relations. From Coq Require Import Ascii. +Set Default Proof Using "Type*". Class Pretty A := pretty : A → string. Definition pretty_N_char (x : N) : ascii := diff --git a/theories/proof_irrel.v b/theories/proof_irrel.v index 2af33406524bb5cc3d32589f5c6c838e88cb2b76..4d9c6d7d83bf459912342375e766e2b18dde9d97 100644 --- a/theories/proof_irrel.v +++ b/theories/proof_irrel.v @@ -2,6 +2,7 @@ (* This file is distributed under the terms of the BSD license. *) (** This file collects facts on proof irrelevant types/propositions. *) From stdpp Require Export base. +Set Default Proof Using "Type*". Hint Extern 200 (ProofIrrel _) => progress (lazy beta) : typeclass_instances. diff --git a/theories/relations.v b/theories/relations.v index dfedb109cb087a415b71c3be9113b5f6d2ecb804..3152cc7888f61f35ecc4a89339d569fdb575976c 100644 --- a/theories/relations.v +++ b/theories/relations.v @@ -6,6 +6,7 @@ small step semantics. This file defines a hint database [ars] containing some theorems on abstract rewriting systems. *) From Coq Require Import Wf_nat. From stdpp Require Export tactics base. +Set Default Proof Using "Type*". (** * Definitions *) Section definitions. diff --git a/theories/set.v b/theories/set.v index 84ef65ff4e20fa25607515aca50d49a644563437..27a979751c390c47c7cf780e27ce6913858e6f38 100644 --- a/theories/set.v +++ b/theories/set.v @@ -2,6 +2,7 @@ (* This file is distributed under the terms of the BSD license. *) (** This file implements sets as functions into Prop. *) From stdpp Require Export collections. +Set Default Proof Using "Type*". Record set (A : Type) : Type := mkSet { set_car : A → Prop }. Add Printing Constructor set. @@ -51,4 +52,4 @@ Proof. intros HPQ. constructor. apply HPQ. Qed. Global Opaque set_elem_of set_top set_empty set_singleton. Global Opaque set_union set_intersection set_difference. -Global Opaque set_ret set_bind set_fmap set_join. \ No newline at end of file +Global Opaque set_ret set_bind set_fmap set_join. diff --git a/theories/sorting.v b/theories/sorting.v index 63b361831c3d7967ab4b3a2762206f427975c442..36fff7fb136f70f4645e5a4c99f4de242148e6fd 100644 --- a/theories/sorting.v +++ b/theories/sorting.v @@ -4,6 +4,7 @@ standard library, but without using the module system. *) From Coq Require Export Sorted. From stdpp Require Export orders list. +Set Default Proof Using "Type*". Section merge_sort. Context {A} (R : relation A) `{∀ x y, Decision (R x y)}. diff --git a/theories/streams.v b/theories/streams.v index 216a426b515f745754dd12667b4908c50c09c0f6..725a59158e64a09efab7bd04928ecd819535e789 100644 --- a/theories/streams.v +++ b/theories/streams.v @@ -1,6 +1,7 @@ (* Copyright (c) 2012-2015, Robbert Krebbers. *) (* This file is distributed under the terms of the BSD license. *) From stdpp Require Export tactics. +Set Default Proof Using "Type*". CoInductive stream (A : Type) : Type := scons : A → stream A → stream A. Arguments scons {_} _ _. diff --git a/theories/stringmap.v b/theories/stringmap.v index 426d392bfe0f3fa51b11f8f4d7d810f7333c61be..9de316cae7b019ad4195f5458ef3cf5d8b25a63a 100644 --- a/theories/stringmap.v +++ b/theories/stringmap.v @@ -6,6 +6,7 @@ search trees (uncompressed Patricia trees) as implemented in the file [pmap] and guarantees logarithmic-time operations. *) From stdpp Require Export fin_maps pretty. From stdpp Require Import gmap. +Set Default Proof Using "Type*". Notation stringmap := (gmap string). Notation stringset := (gset string). diff --git a/theories/strings.v b/theories/strings.v index f08e9b0a089a8258777297e9cc70bbffa3391443..93718899c741b6b8987ec498d3406cd415cbf25f 100644 --- a/theories/strings.v +++ b/theories/strings.v @@ -4,6 +4,7 @@ From Coq Require Import Ascii. From Coq Require Export String. From stdpp Require Export list. From stdpp Require Import countable. +Set Default Proof Using "Type*". (* To avoid randomly ending up with String.length because this module is imported hereditarily somewhere. *) diff --git a/theories/tactics.v b/theories/tactics.v index f8747a7eec416bfbcfb2cb4f85a01789748bd29c..f104f381e9c200d323e5d9a0eb578bb50ee9f8c0 100644 --- a/theories/tactics.v +++ b/theories/tactics.v @@ -5,6 +5,7 @@ the development. *) From Coq Require Import Omega. From Coq Require Export Lia. From stdpp Require Export decidable. +Set Default Proof Using "Type*". Lemma f_equal_dep {A B} (f g : ∀ x : A, B x) x : f = g → f x = g x. Proof. intros ->; reflexivity. Qed. diff --git a/theories/vector.v b/theories/vector.v index 2d663f43cbb4bc49802946a82774970af879f1d8..b0d6d326916c6085de6d98a3b7e08407950e60c0 100644 --- a/theories/vector.v +++ b/theories/vector.v @@ -6,6 +6,7 @@ definitions from the standard library, but renames or changes their notations, so that it becomes more consistent with the naming conventions in this development. *) From stdpp Require Export list. +Set Default Proof Using "Type*". Open Scope vector_scope. (** * The fin type *) diff --git a/theories/zmap.v b/theories/zmap.v index c3f5364722555614475136c5f9f2b653c888d19b..32689e5859a04a323fb65af130b9a75e51f293e9 100644 --- a/theories/zmap.v +++ b/theories/zmap.v @@ -4,6 +4,7 @@ maps whose keys range over Coq's data type of binary naturals [Z]. *) From stdpp Require Import pmap mapset. From stdpp Require Export prelude fin_maps. +Set Default Proof Using "Type*". Local Open Scope Z_scope. Record Zmap (A : Type) : Type :=