https://github.com/EasyCrypt/easycrypt
Revision 9f6a2f9c698e6c47fa841a0b040bc45d82a601c5 authored by Benjamin Gregoire on 07 July 2016, 06:32:45 UTC, committed by Benjamin Gregoire on 07 July 2016, 06:32:45 UTC
1 parent 776af38
Raw File
Tip revision: 9f6a2f9c698e6c47fa841a0b040bc45d82a601c5 authored by Benjamin Gregoire on 07 July 2016, 06:32:45 UTC
add make
Tip revision: 9f6a2f9
FinType.eca
(* --------------------------------------------------------------------
 * Copyright (c) - 2012--2016 - IMDEA Software Institute
 * Copyright (c) - 2012--2016 - Inria
 *
 * Distributed under the terms of the CeCILL-B-V1 license
 * -------------------------------------------------------------------- *)

(* -------------------------------------------------------------------- *)
require import Option Pred Int IntExtra List.

(* -------------------------------------------------------------------- *)
type t.

op enum : t list.

op card : int = size enum.

axiom enum_spec : forall x, count (pred1 x) enum = 1.

(* -------------------------------------------------------------------- *)
lemma enumP : forall x, mem enum x.
proof.
move=> x; have: 0 < count (pred1 x) enum by rewrite enum_spec.
by move/has_count/hasP; case=> y [h @/pred1 <-].
qed.

lemma enum_uniq : uniq enum.
proof. by apply/count_mem_uniq=> x; rewrite enumP enum_spec. qed.

lemma card_gt0 : 0 < card.
proof.
rewrite /card; have: mem enum witness by rewrite enumP.
by case: enum=> //= x s _; rewrite addzC ltzS size_ge0.
qed.
back to top