Skip to content
Snippets Groups Projects
Commit 3a262d02 authored by Johannes Kloos's avatar Johannes Kloos
Browse files

Proofs of infinity and Fresh instances.

We prove that various types are infinite, notably:
- nat, N, positive and Z;
- string (using pretty-printing of nat);
- option, with an infinite element type;
- list, with an inhabited element type.

Furthermore, we instantiate Fresh for strings.
parent a8d02255
No related branches found
No related tags found
1 merge request!13Provide an Infinite typeclass and a generic implementation of Fresh.
...@@ -39,4 +39,5 @@ theories/list.v ...@@ -39,4 +39,5 @@ theories/list.v
theories/functions.v theories/functions.v
theories/hlist.v theories/hlist.v
theories/sorting.v theories/sorting.v
theories/infinite.v
(* 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
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment