diff --git a/_CoqProject b/_CoqProject index 9c07436c78fa8550e83214ca67b3e82958e477a5..2fdb39c70834286abb38f3a3de0ad68773e99ee1 100644 --- a/_CoqProject +++ b/_CoqProject @@ -39,4 +39,5 @@ theories/list.v theories/functions.v theories/hlist.v theories/sorting.v +theories/infinite.v diff --git a/theories/infinite.v b/theories/infinite.v new file mode 100644 index 0000000000000000000000000000000000000000..79739fbc40e151c0db5a804559f79c08b6052dbf --- /dev/null +++ b/theories/infinite.v @@ -0,0 +1,24 @@ +(* Copyright (c) 2012-2017, Coq-std++ developers. *) +(* This file is distributed under the terms of the BSD license. *) +From stdpp Require Import pretty fin_collections. + +Instance string_infinite: Infinite string := {| inject := λ x, "~" +:+ pretty x |}. +Instance nat_infinite: Infinite nat := {| inject := id |}. +Instance N_infinite: Infinite N := {| inject_injective := Nat2N.inj |}. +Instance pos_infinite: Infinite positive := {| inject_injective := SuccNat2Pos.inj |}. +Instance Z_infinite: Infinite Z := {| inject_injective := Nat2Z.inj |}. +Instance option_infinite `{Infinite A}: Infinite (option A) := {| inject := Some ∘ inject |}. +Program Instance list_infinite `{Inhabited A}: Infinite (list A) := + {| inject := λ i, replicate i inhabitant |}. +Next Obligation. +Proof. + intros * i j eqrep%(f_equal length). + rewrite !replicate_length in eqrep; done. +Qed. + +(** Derive fresh instances. *) +Section StringFresh. + Context `{FinCollection string C, ∀ (x: string) (s: C), Decision (x ∈ s)}. + Global Instance string_fresh: Fresh string C := fresh_generic. + Global Instance string_fresh_spec: FreshSpec string C := _. +End StringFresh. \ No newline at end of file