Port the codebase to HottClasses

Initial work + use the latest version of HoTT
This commit is contained in:
Dan Frumin 2017-11-01 17:47:41 +01:00
parent c6f756a856
commit 4fafbf175e
17 changed files with 389 additions and 495 deletions

View File

@ -62,7 +62,7 @@ Section ListToSet.
Definition empty_empty : list_to_set A = := idpath. Definition empty_empty : list_to_set A = := idpath.
Lemma filter_comprehension (ϕ : A -> Bool) (l : list A) : Lemma filter_comprehension (ϕ : A -> Bool) (l : list A) :
list_to_set A (filter ϕ l) = {| list_to_set A l & ϕ |}. list_to_set A (filter ϕ l) = {| list_to_set A l | ϕ |}.
Proof. Proof.
induction l ; cbn in *. induction l ; cbn in *.
- reflexivity. - reflexivity.
@ -170,11 +170,11 @@ Section refinement_examples.
Defined. Defined.
Lemma exist_elim (X : list A) (ϕ : A -> hProp) Lemma exist_elim (X : list A) (ϕ : A -> hProp)
: list_exist ϕ X -> hexists (fun a => a X * ϕ a). : list_exist ϕ X -> hexists (fun a => a X * ϕ a)%type.
Proof. Proof.
rewrite list_exist_set. rewrite list_exist_set.
assert (hexists (fun a : A => a (list_to_set A X) * ϕ a) assert (hexists (fun a : A => a (list_to_set A X) * ϕ a)
-> hexists (fun a : A => a X * ϕ a)) -> hexists (fun a : A => a X * ϕ a))%type
as H2. as H2.
{ {
intros H1. intros H1.

View File

@ -1,5 +1,5 @@
(** Some examples of lattices. *) (** Some examples of lattices. *)
Require Import HoTT lattice_interface. Require Export HoTT lattice_interface.
(** [Bool] is a lattice. *) (** [Bool] is a lattice. *)
Section BoolLattice. Section BoolLattice.
@ -10,14 +10,18 @@ Section BoolLattice.
; auto ; auto
; try contradiction. ; try contradiction.
Instance maximum_bool : maximum Bool := orb. Instance maximum_bool : Join Bool := orb.
Instance minimum_bool : minimum Bool := andb. Instance minimum_bool : Meet Bool := andb.
Instance bottom_bool : bottom Bool := false. Instance bottom_bool : Bottom Bool := false.
Global Instance boundedjoinsemilattice_bool : BoundedJoinSemiLattice Bool.
Proof. repeat (split ; (apply _ || solve_bool)). Defined.
Global Instance meetsemilattice_bool : MeetSemiLattice Bool.
Proof. repeat (split ; (apply _ || solve_bool)). Defined.
Global Instance boundedmeetsemilattice_bool : @BoundedSemiLattice Bool () true.
Proof. repeat (split ; (apply _ || solve_bool)). Defined.
Global Instance lattice_bool : Lattice Bool. Global Instance lattice_bool : Lattice Bool.
Proof. Proof. split ; (apply _ || solve_bool). Defined.
split ; solve_bool.
Defined.
Definition and_true : forall b, andb b true = b. Definition and_true : forall b, andb b true = b.
Proof. Proof.
@ -50,12 +54,12 @@ End BoolLattice.
Create HintDb bool_lattice_hints. Create HintDb bool_lattice_hints.
Hint Resolve associativity : bool_lattice_hints. Hint Resolve associativity : bool_lattice_hints.
Hint Resolve (associativity _ _ _)^ : bool_lattice_hints. (* Hint Resolve (associativity _ _ _)^ : bool_lattice_hints. *)
Hint Resolve commutativity : bool_lattice_hints. Hint Resolve commutativity : bool_lattice_hints.
Hint Resolve absorb : bool_lattice_hints. Hint Resolve absorption : bool_lattice_hints.
Hint Resolve idempotency : bool_lattice_hints. Hint Resolve idempotency : bool_lattice_hints.
Hint Resolve neutralityL : bool_lattice_hints. Hint Resolve left_identity : bool_lattice_hints.
Hint Resolve neutralityR : bool_lattice_hints. Hint Resolve right_identity : bool_lattice_hints.
Hint Resolve Hint Resolve
associativity associativity
@ -66,25 +70,51 @@ Hint Resolve
(** If [B] is a lattice, then [A -> B] is a lattice. *) (** If [B] is a lattice, then [A -> B] is a lattice. *)
Section fun_lattice. Section fun_lattice.
Context {A B : Type}. Context {A B : Type}.
Context `{Lattice B}. Context `{BJoin : Join B}.
Context `{BMeet : Meet B}.
Context `{@Lattice B BJoin BMeet}.
Context `{Funext}. Context `{Funext}.
Context `{BBottom : Bottom B}.
Global Instance max_fun : maximum (A -> B) := Global Instance bot_fun : Bottom (A -> B)
fun (f g : A -> B) (a : A) => max_L0 (f a) (g a). := fun _ => .
Global Instance min_fun : minimum (A -> B) := Global Instance join_fun : Join (A -> B) :=
fun (f g : A -> B) (a : A) => min_L0 (f a) (g a). fun (f g : A -> B) (a : A) => (f a) (g a).
Global Instance bot_fun : bottom (A -> B) Global Instance meet_fun : Meet (A -> B) :=
:= fun _ => empty_L. fun (f g : A -> B) (a : A) => (f a) (g a).
Ltac solve_fun := Ltac solve_fun :=
compute ; intros ; apply path_forall ; intro ; compute ; intros ; apply path_forall ; intro ;
eauto with lattice_hints typeclass_instances. eauto with lattice_hints typeclass_instances.
Create HintDb test1.
Lemma associativity_lat `{Lattice A} (x y z : A) :
x (y z) = x y z.
Proof. apply associativity. Defined.
Hint Resolve associativity : test1.
Hint Resolve associativity_lat : test1.
Global Instance lattice_fun : Lattice (A -> B). Global Instance lattice_fun : Lattice (A -> B).
Proof. Proof.
split ; solve_fun. repeat (split; try (apply _)).
eauto with test1.
(* TODO *)
all: solve_fun.
apply associativity.
apply commutativity.
apply idempotency. apply _.
apply associativity.
apply commutativity.
apply idempotency. apply _.
Defined.
Global Instance boundedjoinsemilattice_fun
`{@BoundedJoinSemiLattice B BJoin BBottom} :
BoundedJoinSemiLattice (A -> B).
Proof.
repeat split; try apply _; try solve_fun.
Defined. Defined.
End fun_lattice. End fun_lattice.
@ -92,24 +122,26 @@ End fun_lattice.
Section sub_lattice. Section sub_lattice.
Context {A : Type} {P : A -> hProp}. Context {A : Type} {P : A -> hProp}.
Context `{Lattice A}. Context `{Lattice A}.
Context {Hmax : forall x y, P x -> P y -> P (max_L0 x y)}. Context `{Bottom A}.
Context {Hmin : forall x y, P x -> P y -> P (min_L0 x y)}. Context {Hmax : forall x y, P x -> P y -> P (x y)}.
Context {Hbot : P empty_L}. Context {Hmin : forall x y, P x -> P y -> P (x y)}.
Context {Hbot : P }.
Definition AP : Type := sig P. Definition AP : Type := sig P.
Instance botAP : bottom AP := (empty_L ; Hbot). Instance botAP : Bottom AP.
Proof. refine ( _). apply Hbot. Defined.
Instance maxAP : maximum AP := Instance maxAP : Join AP :=
fun x y => fun x y =>
match x, y with match x, y with
| (a ; pa), (b ; pb) => (max_L0 a b ; Hmax a b pa pb) | (a ; pa), (b ; pb) => (a b ; Hmax a b pa pb)
end. end.
Instance minAP : minimum AP := Instance minAP : Meet AP :=
fun x y => fun x y =>
match x, y with match x, y with
| (a ; pa), (b ; pb) => (min_L0 a b ; Hmin a b pa pb) | (a ; pa), (b ; pb) => (a b ; Hmin a b pa pb)
end. end.
Instance hprop_sub : forall c, IsHProp (P c). Instance hprop_sub : forall c, IsHProp (P c).
@ -127,19 +159,25 @@ Section sub_lattice.
Global Instance lattice_sub : Lattice AP. Global Instance lattice_sub : Lattice AP.
Proof. Proof.
split ; solve_sub. repeat (split ; try (apply _ || solve_sub)).
apply associativity.
apply commutativity.
apply idempotency. apply _.
apply associativity.
apply commutativity.
apply idempotency. apply _.
Defined. Defined.
End sub_lattice. End sub_lattice.
Instance lor : maximum hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)). Instance lor : Join hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)).
Delimit Scope logic_scope with L. Delimit Scope logic_scope with L.
Notation "A B" := (lor A B) (at level 20, right associativity) : logic_scope. Notation "A B" := (lor A B) (at level 20, right associativity) : logic_scope.
Arguments lor _%L _%L. Arguments lor _%L _%L.
Open Scope logic_scope. Open Scope logic_scope.
Instance land : minimum hProp := fun X Y => BuildhProp (prod X Y). Instance land : Meet hProp := fun X Y => BuildhProp (prod X Y).
Instance lfalse : bottom hProp := False_hp. Instance lfalse : Bottom hProp := False_hp.
Notation "A ∧ B" := (land A B) (at level 20, right associativity) : logic_scope. Notation "A ∧ B" := (land A B) (at level 20, right associativity) : logic_scope.
Arguments land _%L _%L. Arguments land _%L _%L.
@ -181,17 +219,18 @@ Section hPropLattice.
intros X Y Z. intros X Y Z.
symmetry. symmetry.
apply path_hprop. apply path_hprop.
symmetry.
apply equiv_prod_assoc. apply equiv_prod_assoc.
Defined. Defined.
Instance lor_idempotent : Idempotent lor. Instance lor_idempotent : BinaryIdempotent lor.
Proof. Proof.
intros X. intros X.
apply path_iff_hprop ; lor_intros apply path_iff_hprop ; lor_intros
; try(refine (tr(inl _))) ; auto. ; try(refine (tr(inl _))) ; auto.
Defined. Defined.
Instance land_idempotent : Idempotent land. Instance land_idempotent : BinaryIdempotent land.
Proof. Proof.
intros X. intros X.
apply path_iff_hprop ; cbn. apply path_iff_hprop ; cbn.
@ -199,14 +238,14 @@ Section hPropLattice.
- intros a ; apply (pair a a). - intros a ; apply (pair a a).
Defined. Defined.
Instance lor_neutrall : NeutralL lor lfalse. Instance lor_neutrall : LeftIdentity lor lfalse.
Proof. Proof.
intros X. intros X.
apply path_iff_hprop ; lor_intros ; try contradiction apply path_iff_hprop ; lor_intros ; try contradiction
; try (refine (tr(inr _))) ; auto. ; try (refine (tr(inr _))) ; auto.
Defined. Defined.
Instance lor_neutralr : NeutralR lor lfalse. Instance lor_neutralr : RightIdentity lor lfalse.
Proof. Proof.
intros X. intros X.
apply path_iff_hprop ; lor_intros ; try contradiction apply path_iff_hprop ; lor_intros ; try contradiction
@ -235,16 +274,17 @@ Section hPropLattice.
* apply (tr (inl X)). * apply (tr (inl X)).
Defined. Defined.
Global Instance lattice_hprop : Lattice hProp := Global Instance lattice_hprop : Lattice hProp.
{ commutative_min := _ ; Proof. repeat (split ; try apply _). Defined.
commutative_max := _ ;
associative_min := _ ; Global Instance bounded_jsl_hprop : BoundedJoinSemiLattice hProp.
associative_max := _ ; Proof. repeat (split ; try apply _). Qed.
idempotent_min := _ ;
idempotent_max := _ ; Global Instance top_hprop : Top hProp := Unit_hp.
neutralL_max := _ ; Global Instance bounded_msl_hprop : @BoundedSemiLattice hProp () .
neutralR_max := _ ; Proof.
absorption_min_max := _ ; repeat (split; try apply _); cbv.
absorption_max_min := _ - intros X. apply path_trunctype ; apply prod_unit_l.
}. - intros X. apply path_trunctype ; apply prod_unit_r.
Defined.
End hPropLattice. End hPropLattice.

View File

@ -1,119 +1,60 @@
(** Interface for lattices and join semilattices. *) (** Interface for lattices and join semilattices. *)
Require Import HoTT. Require Import HoTT.
From HoTT.Classes.interfaces Require Export abstract_algebra canonical_names.
From HoTT.Classes Require Export theory.lattices.
(** Some preliminary notions to define lattices. *) (* (** Join semilattices as a typeclass. They only have a join operator. *) *)
Section binary_operation. (* Section JoinSemiLattice. *)
Definition operation (A : Type) := A -> A -> A. (* Variable A : Type. *)
(* Context {max_L : Join A} {empty_L : Bottom A}. *)
Variable (A : Type) (* Class JoinSemiLattice := *)
(f : operation A). (* { *)
(* commutative_max_js :> Commutative max_L ; *)
(* associative_max_js :> Associative max_L ; *)
(* idempotent_max_js :> BinaryIdempotent max_L ; *)
(* neutralL_max_js :> LeftIdentity max_L empty_L ; *)
(* neutralR_max_js :> RightIdentity max_L empty_L ; *)
(* }. *)
(* End JoinSemiLattice. *)
Class Commutative := (* Arguments JoinSemiLattice _ {_} {_}. *)
commutativity : forall x y, f x y = f y x.
Class Associative := (* Create HintDb joinsemilattice_hints. *)
associativity : forall x y z, f (f x y) z = f x (f y z). (* Hint Resolve associativity : joinsemilattice_hints. *)
(* Hint Resolve (associativity _ _ _)^ : joinsemilattice_hints. *)
(* Hint Resolve commutativity : joinsemilattice_hints. *)
(* Hint Resolve idempotency : joinsemilattice_hints. *)
(* Hint Resolve neutralityL : joinsemilattice_hints. *)
(* Hint Resolve neutralityR : joinsemilattice_hints. *)
Class Idempotent := (* (** Lattices as a typeclass which have both a join and a meet. *) *)
idempotency : forall x, f x x = x. (* Section Lattice. *)
(* Variable A : Type. *)
(* Context {max_L : maximum A} {min_L : minimum A} {empty_L : bottom A}. *)
Variable g : operation A. (* Class Lattice := *)
(* { *)
(* commutative_min :> Commutative min_L ; *)
(* commutative_max :> Commutative max_L ; *)
(* associative_min :> Associative min_L ; *)
(* associative_max :> Associative max_L ; *)
(* idempotent_min :> Idempotent min_L ; *)
(* idempotent_max :> Idempotent max_L ; *)
(* neutralL_max :> NeutralL max_L empty_L ; *)
(* neutralR_max :> NeutralR max_L empty_L ; *)
(* absorption_min_max :> Absorption min_L max_L ; *)
(* absorption_max_min :> Absorption max_L min_L *)
(* }. *)
(* End Lattice. *)
Class Absorption := (* Arguments Lattice _ {_} {_} {_}. *)
absorb : forall x y, f x (g x y) = x.
Variable (n : A).
Class NeutralL :=
neutralityL : forall x, f n x = x.
Class NeutralR :=
neutralityR : forall x, f x n = x.
End binary_operation.
Arguments Commutative {_} _.
Arguments Associative {_} _.
Arguments Idempotent {_} _.
Arguments NeutralL {_} _ _.
Arguments NeutralR {_} _ _.
Arguments Absorption {_} _ _.
Arguments commutativity {_} {_} {_} _ _.
Arguments associativity {_} {_} {_} _ _ _.
Arguments idempotency {_} {_} {_} _.
Arguments neutralityL {_} {_} {_} {_} _.
Arguments neutralityR {_} {_} {_} {_} _.
Arguments absorb {_} {_} {_} {_} _ _.
(** The operations in a lattice. *)
Section lattice_operations.
Variable (A : Type).
Class maximum :=
max_L : operation A.
Class minimum :=
min_L : operation A.
Class bottom :=
empty : A.
End lattice_operations.
Arguments max_L {_} {_} _.
Arguments min_L {_} {_} _.
Arguments empty {_}.
(** Join semilattices as a typeclass. They only have a join operator. *)
Section JoinSemiLattice.
Variable A : Type.
Context {max_L : maximum A} {empty_L : bottom A}.
Class JoinSemiLattice :=
{
commutative_max_js :> Commutative max_L ;
associative_max_js :> Associative max_L ;
idempotent_max_js :> Idempotent max_L ;
neutralL_max_js :> NeutralL max_L empty_L ;
neutralR_max_js :> NeutralR max_L empty_L ;
}.
End JoinSemiLattice.
Arguments JoinSemiLattice _ {_} {_}.
Create HintDb joinsemilattice_hints.
Hint Resolve associativity : joinsemilattice_hints.
Hint Resolve (associativity _ _ _)^ : joinsemilattice_hints.
Hint Resolve commutativity : joinsemilattice_hints.
Hint Resolve idempotency : joinsemilattice_hints.
Hint Resolve neutralityL : joinsemilattice_hints.
Hint Resolve neutralityR : joinsemilattice_hints.
(** Lattices as a typeclass which have both a join and a meet. *)
Section Lattice.
Variable A : Type.
Context {max_L : maximum A} {min_L : minimum A} {empty_L : bottom A}.
Class Lattice :=
{
commutative_min :> Commutative min_L ;
commutative_max :> Commutative max_L ;
associative_min :> Associative min_L ;
associative_max :> Associative max_L ;
idempotent_min :> Idempotent min_L ;
idempotent_max :> Idempotent max_L ;
neutralL_max :> NeutralL max_L empty_L ;
neutralR_max :> NeutralR max_L empty_L ;
absorption_min_max :> Absorption min_L max_L ;
absorption_max_min :> Absorption max_L min_L
}.
End Lattice.
Arguments Lattice _ {_} {_} {_}.
Create HintDb lattice_hints. Create HintDb lattice_hints.
Hint Resolve associativity : lattice_hints. Hint Resolve associativity : lattice_hints.
Hint Resolve (associativity _ _ _)^ : lattice_hints. (* Hint Resolve (associativity _ _ _)^ : lattice_hints. *)
Hint Resolve commutativity : lattice_hints. Hint Resolve commutativity : lattice_hints.
Hint Resolve absorb : lattice_hints. Hint Resolve absorption : lattice_hints.
Hint Resolve idempotency : lattice_hints. Hint Resolve idempotency : lattice_hints.
Hint Resolve neutralityL : lattice_hints. Hint Resolve left_identity : lattice_hints.
Hint Resolve neutralityR : lattice_hints. Hint Resolve right_identity : lattice_hints.

View File

@ -1,5 +1,7 @@
(** The structure of a monad. *) (** The structure of a monad. *)
(* TODO REMOVE THIS *)
Require Import HoTT. Require Import HoTT.
Require Export HoTT.Classes.interfaces.monad.
Section functor. Section functor.
Variable (F : Type -> Type). Variable (F : Type -> Type).
@ -12,34 +14,8 @@ Section functor.
}. }.
End functor. End functor.
Section monad_operations. Section monad_join.
Variable (F : Type -> Type). Context `{Monad M}.
Class hasPure := Definition mjoin {A} (m : M (M A)) : M A := bind m id.
{ End monad_join.
pure : forall (A : Type), A -> F A
}.
Class hasBind :=
{
bind : forall (A : Type), F(F A) -> F A
}.
End monad_operations.
Arguments pure {_} {_} _ _.
Arguments bind {_} {_} _ _.
Section monad.
Variable (F : Type -> Type).
Context `{Functor F} `{hasPure F} `{hasBind F}.
Class Monad :=
{
bind_assoc : forall {A : Type},
bind A o bind (F A) = bind A o fmap F (bind A) ;
bind_neutral_l : forall {A : Type},
bind A o pure (F A) = idmap ;
bind_neutral_r : forall {A : Type},
bind A o fmap F (pure A) = idmap
}.
End monad.

View File

@ -16,7 +16,7 @@ Section interface.
f_empty : forall A, f A = ; f_empty : forall A, f A = ;
f_singleton : forall A a, f A (singleton a) = {|a|}; f_singleton : forall A a, f A (singleton a) = {|a|};
f_union : forall A X Y, f A (union X Y) = (f A X) (f A Y); f_union : forall A X Y, f A (union X Y) = (f A X) (f A Y);
f_filter : forall A φ X, f A (filter φ X) = {| f A X & φ |}; f_filter : forall A φ X, f A (filter φ X) = {| f A X | φ |};
f_member : forall A a X, member a X = a (f A X) f_member : forall A a X, member a X = a (f A X)
}. }.
@ -175,12 +175,12 @@ Section quotient_properties.
Proof. Proof.
intros ϕ X. intros ϕ X.
apply (quotient_iso _)^-1. apply (quotient_iso _)^-1.
simple refine ({|_ & ϕ|}). simple refine ({|_ | ϕ|}).
apply (quotient_iso (f A) X). apply (quotient_iso (f A) X).
Defined. Defined.
Definition well_defined_filter (A : Type) (ϕ : A -> Bool) (X : T A) : Definition well_defined_filter (A : Type) (ϕ : A -> Bool) (X : T A) :
{|class_of _ X & ϕ|} = class_of _ {|X & ϕ|}. {|class_of _ X | ϕ|} = class_of _ {|X | ϕ|}.
Proof. Proof.
rewrite <- (eissect (quotient_iso _)). rewrite <- (eissect (quotient_iso _)).
simpl. simpl.
@ -224,30 +224,35 @@ Section quotient_properties.
apply (ap _ HXY). apply (ap _ HXY).
Defined. Defined.
Instance View_max A : maximum (View A). Instance join_view A : Join (View A).
Proof. Proof.
apply view_union. apply view_union.
Defined. Defined.
Hint Unfold Commutative Associative Idempotent NeutralL NeutralR View_max view_union. Local Hint Unfold Commutative Associative HeteroAssociative Idempotent BinaryIdempotent LeftIdentity RightIdentity join_view view_union sg_op join_is_sg_op meet_is_sg_op.
Instance bottom_view A : bottom (View A). Instance bottom_view A : Bottom (View A).
Proof. Proof.
apply View_empty. apply View_empty.
Defined. Defined.
Ltac sq1 := autounfold ; intros ; try f_ap Ltac sq1 := autounfold; intros;
; rewrite ?(eisretr (quotient_iso _)) unfold view_union; try f_ap;
; eauto with lattice_hints typeclass_instances. rewrite ?(eisretr (quotient_iso _)).
Ltac sq2 := autounfold ; intros Ltac sq2 := autounfold; intros;
; rewrite <- (eissect (quotient_iso _)), ?(eisretr (quotient_iso _)) unfold view_union;
; f_ap ; simpl rewrite <- (eissect (quotient_iso _)), ?(eisretr (quotient_iso _));
; reduce T ; eauto with lattice_hints typeclass_instances. f_ap; simpl.
Global Instance view_lattice A : JoinSemiLattice (View A). Global Instance view_lattice A : BoundedJoinSemiLattice (View A).
Proof. Proof.
split ; try sq1 ; try sq2. repeat split; try apply _.
- sq1. apply associativity.
- sq2. apply left_identity.
- sq2. apply right_identity.
- sq1. apply commutativity.
- sq2. apply binary_idempotent.
Defined. Defined.
End quotient_properties. End quotient_properties.
@ -316,7 +321,8 @@ Section properties.
Ltac via_quotient := intros ; apply reflect_f_eq ; simpl Ltac via_quotient := intros ; apply reflect_f_eq ; simpl
; rewrite <- ?(well_defined_union T _), <- ?(well_defined_empty T _) ; rewrite <- ?(well_defined_union T _), <- ?(well_defined_empty T _)
; eauto with lattice_hints typeclass_instances. ; try (apply commutativity || apply associativity || apply binary_idempotent || apply left_identity || apply right_identity).
(* TODO ; eauto with lattice_hints typeclass_instances. *)
Lemma union_comm : forall A (X Y : T A), Lemma union_comm : forall A (X Y : T A),
set_eq f (X Y) (Y X). set_eq f (X Y) (Y X).
@ -328,6 +334,7 @@ Section properties.
set_eq f ((X Y) Z) (X (Y Z)). set_eq f ((X Y) Z) (X (Y Z)).
Proof. Proof.
via_quotient. via_quotient.
symmetry. apply associativity.
Defined. Defined.
Lemma union_idem : forall A (X : T A), Lemma union_idem : forall A (X : T A),

View File

@ -53,7 +53,7 @@ Infix "∩" := intersection (at level 8, right associativity).
Notation "( ∩ )" := intersection (only parsing). Notation "( ∩ )" := intersection (only parsing).
Notation "( X ∩ )" := (intersection X) (only parsing). Notation "( X ∩ )" := (intersection X) (only parsing).
Notation "( ∩ Y )" := (fun X => X Y) (only parsing). Notation "( ∩ Y )" := (fun X => X Y) (only parsing).
Notation "{| X & ϕ |}" := (filter ϕ X). Notation "{| X | ϕ |}" := (filter ϕ X).
Infix "" := member (at level 9, right associativity). Infix "" := member (at level 9, right associativity).
Infix "" := subset (at level 10, right associativity). Infix "" := subset (at level 10, right associativity).
Infix "∈_d" := member_dec (at level 9, right associativity). Infix "∈_d" := member_dec (at level 9, right associativity).

View File

@ -139,11 +139,17 @@ Section relations.
- apply False_hp. - apply False_hp.
- apply (fun a' => merely(a = a')). - apply (fun a' => merely(a = a')).
- apply lor. - apply lor.
(* TODO *)
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply associativity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply commutativity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply left_identity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply right_identity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
intros. simpl. apply binary_idempotent.
Defined. Defined.
(** Subset relation of finite sets. *) (** Subset relation of finite sets. *)
@ -153,15 +159,15 @@ Section relations.
hrecursion X. hrecursion X.
- apply Unit_hp. - apply Unit_hp.
- apply (fun a => a Y). - apply (fun a => a Y).
- intros X1 X2. - apply land.
exists (prod X1 X2). (* TODO *)
exact _.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply associativity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
- intros. apply commutativity.
apply path_trunctype ; apply prod_unit_l. - apply left_identity.
- intros. - apply right_identity.
apply path_trunctype ; apply prod_unit_r.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
intros. apply binary_idempotent.
Defined. Defined.
End relations. End relations.

View File

@ -6,7 +6,7 @@ Section length.
Definition length : FSet A -> nat. Definition length : FSet A -> nat.
simple refine (FSet_cons_rec _ _ _ _ _ _). simple refine (FSet_cons_rec _ _ _ _ _ _).
- apply 0. - apply 0%nat.
- intros a X n. - intros a X n.
apply (if a _d X then n else (S n)). apply (if a _d X then n else (S n)).
- intros X a n. - intros X a n.
@ -295,7 +295,7 @@ End length_sum.
Section length_zero_one. Section length_zero_one.
Context {A : Type} `{Univalence} `{MerelyDecidablePaths A}. Context {A : Type} `{Univalence} `{MerelyDecidablePaths A}.
Lemma Z_not_S n : S n = 0 -> Empty. Lemma Z_not_S n : S n = 0%nat -> Empty.
Proof. Proof.
refine (@equiv_path_nat (n.+1) 0)^-1. refine (@equiv_path_nat (n.+1) 0)^-1.
Defined. Defined.
@ -311,7 +311,7 @@ Section length_zero_one.
apply ((equiv_path_nat)^-1 X). apply ((equiv_path_nat)^-1 X).
Defined. Defined.
Theorem length_zero : forall (X : FSet A) (HX : length X = 0), X = . Theorem length_zero : forall (X : FSet A) (HX : length X = 0%nat), X = .
Proof. Proof.
simple refine (FSet_cons_ind (fun Z => _) _ _ _ _ _) simple refine (FSet_cons_ind (fun Z => _) _ _ _ _ _)
; try (intros ; apply path_ishprop) ; simpl. ; try (intros ; apply path_ishprop) ; simpl.
@ -326,7 +326,7 @@ Section length_zero_one.
* contradiction (Z_not_S _ HaX). * contradiction (Z_not_S _ HaX).
Defined. Defined.
Theorem length_one : forall (X : FSet A) (HX : length X = 1), hexists (fun a => X = {|a|}). Theorem length_one : forall (X : FSet A) (HX : length X = 1%nat), hexists (fun a => X = {|a|}).
Proof. Proof.
simple refine (FSet_cons_ind (fun Z => _) _ _ _ _ _) simple refine (FSet_cons_ind (fun Z => _) _ _ _ _ _)
; try (intros ; apply path_ishprop) ; simpl. ; try (intros ; apply path_ishprop) ; simpl.

View File

@ -19,25 +19,20 @@ Section operations.
- apply (idem o f). - apply (idem o f).
Defined. Defined.
Global Instance fset_pure : hasPure FSet. Global Instance return_fset : Return FSet := fun _ a => {|a|}.
Proof.
split.
apply (fun _ a => {|a|}).
Defined.
Global Instance fset_bind : hasBind FSet. Global Instance bind_fset : Bind FSet.
Proof. Proof.
split. intros A B m f.
intros A. hrecursion m.
hrecursion.
- exact . - exact .
- exact idmap. - exact f.
- apply (). - apply ().
- apply assoc. - apply assoc.
- apply comm. - apply comm.
- apply nl. - apply nl.
- apply nr. - apply nr.
- apply union_idem. - intros; apply union_idem.
Defined. Defined.
(** Set-theoretical operations. *) (** Set-theoretical operations. *)
@ -110,8 +105,8 @@ Section operations.
- intros a f. - intros a f.
apply {|f (a; tr idpath)|}. apply {|f (a; tr idpath)|}.
- intros X1 X2 HX1 HX2 f. - intros X1 X2 HX1 HX2 f.
pose (fX1 := fun Z : {a : A & a X1} => f (pr1 Z;tr (inl (pr2 Z)))). pose (fX1 := fun Z : {a : A | a X1} => f (pr1 Z;tr (inl (pr2 Z)))).
pose (fX2 := fun Z : {a : A & a X2} => f (pr1 Z;tr (inr (pr2 Z)))). pose (fX2 := fun Z : {a : A | a X2} => f (pr1 Z;tr (inr (pr2 Z)))).
specialize (HX1 fX1). specialize (HX1 fX1).
specialize (HX2 fX2). specialize (HX2 fX2).
apply (HX1 HX2). apply (HX1 HX2).
@ -201,9 +196,9 @@ Section operations_decidable.
Defined. Defined.
Global Instance fset_intersection : hasIntersection (FSet A) Global Instance fset_intersection : hasIntersection (FSet A)
:= fun X Y => {|X & (fun a => a _d Y)|}. := fun X Y => {|X | fun a => a _d Y |}.
Definition difference := fun X Y => {|X & (fun a => negb a _d Y)|}. Definition difference := fun X Y => {|X | fun a => negb a _d Y|}.
End operations_decidable. End operations_decidable.
Section FSet_cons_rec. Section FSet_cons_rec.

View File

@ -6,7 +6,7 @@ Require Import lattice_interface lattice_examples monad_interface.
Section monad_fset. Section monad_fset.
Context `{Funext}. Context `{Funext}.
Global Instance fset_functorish : Functorish FSet. Global Instance functorish_fset : Functorish FSet.
Proof. Proof.
simple refine (Build_Functorish _ _ _). simple refine (Build_Functorish _ _ _).
- intros ? ? f. - intros ? ? f.
@ -19,7 +19,7 @@ Section monad_fset.
; try (intros ; apply path_ishprop). ; try (intros ; apply path_ishprop).
Defined. Defined.
Global Instance fset_functor : Functor FSet. Global Instance functor_fset : Functor FSet.
Proof. Proof.
split. split.
intros. intros.
@ -30,24 +30,18 @@ Section monad_fset.
; try (intros ; apply path_ishprop). ; try (intros ; apply path_ishprop).
Defined. Defined.
Global Instance fset_monad : Monad FSet. Global Instance monad_fset : Monad FSet.
Proof. Proof.
split. split.
- intros. - intros. reflexivity.
apply path_forall. - intros A X.
intro X. hrecursion X;
hrecursion X ; try (intros; f_ap) ; try (intros; compute[bind ret bind_fset return_fset]; simpl; f_ap);
try (intros; apply path_ishprop). try (intros; apply path_ishprop); try reflexivity.
- intros. - intros A B C X f g.
apply path_forall. hrecursion X;
intro X. try (intros; compute[bind ret bind_fset return_fset]; simpl; f_ap);
hrecursion X ; try (intros; f_ap) ; try (intros; apply path_ishprop); try reflexivity.
try (intros; apply path_ishprop).
- intros.
apply path_forall.
intro X.
hrecursion X ; try (intros; f_ap) ;
try (intros; apply path_ishprop).
Defined. Defined.
Lemma fmap_isIn `{Univalence} {A B : Type} (f : A -> B) (a : A) (X : FSet A) : Lemma fmap_isIn `{Univalence} {A B : Type} (f : A -> B) (a : A) (X : FSet A) :
@ -64,7 +58,7 @@ Section monad_fset.
+ right. by apply HX1. + right. by apply HX1.
Defined. Defined.
Instance fmap_Surjection `{Univalence} {A B : Type} (f : A -> B) Instance surjection_fmap `{Univalence} {A B : Type} (f : A -> B)
{Hsurj : IsSurjection f} : IsSurjection (fmap FSet f). {Hsurj : IsSurjection f} : IsSurjection (fmap FSet f).
Proof. Proof.
apply BuildIsSurjection. apply BuildIsSurjection.
@ -87,8 +81,8 @@ Section monad_fset.
f_ap. f_ap.
Defined. Defined.
Lemma bind_isIn `{Univalence} {A : Type} (X : FSet (FSet A)) (a : A) : Lemma mjoin_isIn `{Univalence} {A : Type} (X : FSet (FSet A)) (a : A) :
(exists x, x X * a x) -> a (bind _ X). (exists x, prod (x X) (a x)) -> a (mjoin X).
Proof. Proof.
hinduction X; hinduction X;
try (intros; apply path_forall; intro; apply path_ishprop). try (intros; apply path_forall; intro; apply path_ishprop).
@ -253,13 +247,13 @@ Section properties_membership.
: a X Y = a X a Y := idpath. : a X Y = a X a Y := idpath.
Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A, Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A,
{|X Y & ϕ|} = {|X & ϕ|} {|Y & ϕ|}. {| (X Y) | ϕ|} = {|X | ϕ|} {|Y | ϕ|}.
Proof. Proof.
reflexivity. reflexivity.
Defined. Defined.
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A, Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
a {|X & ϕ|} = if ϕ a then a X else False_hp. a {|X | ϕ|} = if ϕ a then a X else False_hp.
Proof. Proof.
hinduction ; try (intros ; apply set_path2). hinduction ; try (intros ; apply set_path2).
- destruct (ϕ a) ; reflexivity. - destruct (ϕ a) ; reflexivity.
@ -386,8 +380,8 @@ Section properties_membership.
** apply p. ** apply p.
** apply path_ishprop. ** apply path_ishprop.
- intros X1 X2 HX1 HX2 f b. - intros X1 X2 HX1 HX2 f b.
pose (fX1 := fun Z : {a : A & a X1} => f (pr1 Z;tr (inl (pr2 Z)))). pose (fX1 := fun Z : {a : A | a X1} => f (pr1 Z;tr (inl (pr2 Z)))).
pose (fX2 := fun Z : {a : A & a X2} => f (pr1 Z;tr (inr (pr2 Z)))). pose (fX2 := fun Z : {a : A | a X2} => f (pr1 Z;tr (inr (pr2 Z)))).
specialize (HX1 fX1 b). specialize (HX1 fX1 b).
specialize (HX2 fX2 b). specialize (HX2 fX2 b).
apply path_iff_hprop. apply path_iff_hprop.
@ -461,8 +455,8 @@ Section properties_membership.
apply tr. apply tr.
unfold IsEmbedding, hfiber in *. unfold IsEmbedding, hfiber in *.
specialize (f_inj (f a)). specialize (f_inj (f a)).
pose ((a;idpath (f a)) : {x : A & f x = f a}) as p1. pose ((a;idpath (f a)) : {x : A | f x = f a}) as p1.
pose ((b;Hfa^) : {x : A & f x = f a}) as p2. pose ((b;Hfa^) : {x : A | f x = f a}) as p2.
assert (p1 = p2) as Hp1p2. assert (p1 = p2) as Hp1p2.
{ apply f_inj. } { apply f_inj. }
apply (ap (fun x => x.1) Hp1p2). apply (ap (fun x => x.1) Hp1p2).
@ -485,31 +479,18 @@ Ltac toHProp :=
Section fset_join_semilattice. Section fset_join_semilattice.
Context {A : Type}. Context {A : Type}.
Instance: bottom(FSet A). Global Instance bottom_fset : Bottom (FSet A) := .
Proof.
unfold bottom.
apply .
Defined.
Instance: maximum(FSet A). Global Instance join_fset : Join (FSet A) := fun x y => x y.
Proof.
intros x y.
apply (x y).
Defined.
Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A). Global Instance boundedjoinsemilattice_fset : BoundedJoinSemiLattice (FSet A).
Proof. Proof.
split. repeat split; try apply _; cbv.
- intros ? ?. - apply assoc.
apply comm. - apply nl.
- intros ? ? ?. - apply nr.
apply (assoc _ _ _)^. - apply comm.
- intros ?. - apply union_idem.
apply union_idem.
- intros x.
apply nl.
- intros ?.
apply nr.
Defined. Defined.
End fset_join_semilattice. End fset_join_semilattice.
@ -569,10 +550,10 @@ Section properties_membership_decidable.
Defined. Defined.
Lemma comprehension_isIn_d (Y : FSet A) (ϕ : A -> Bool) (a : A) : Lemma comprehension_isIn_d (Y : FSet A) (ϕ : A -> Bool) (a : A) :
a _d {|Y & ϕ|} = andb (a _d Y) (ϕ a). a _d {|Y | ϕ|} = andb (a _d Y) (ϕ a).
Proof. Proof.
unfold member_dec, fset_member_bool, dec ; simpl. unfold member_dec, fset_member_bool, dec ; simpl.
destruct (isIn_decidable a {|Y & ϕ|}) as [t | t] destruct (isIn_decidable a {|Y | ϕ|}) as [t | t]
; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t ; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
; destruct (ϕ a) ; try reflexivity ; try contradiction ; destruct (ϕ a) ; try reflexivity ; try contradiction
; try (contradiction (n t)) ; contradiction (t n). ; try (contradiction (n t)) ; contradiction (t n).
@ -692,27 +673,14 @@ Section set_lattice.
Context `{MerelyDecidablePaths A}. Context `{MerelyDecidablePaths A}.
Context `{Univalence}. Context `{Univalence}.
Instance fset_max : maximum (FSet A). Global Instance meet_fset : Meet (FSet A) := intersection.
Proof.
intros x y.
apply (x y).
Defined.
Instance fset_min : minimum (FSet A). Global Instance lattice_fset : Lattice (FSet A).
Proof. Proof.
intros x y. repeat split; try apply _;
apply (x y). compute[sg_op meet_is_sg_op meet_fset];
Defined. toBool.
apply associativity.
Instance fset_bot : bottom (FSet A).
Proof.
unfold bottom.
apply .
Defined.
Instance lattice_fset : Lattice (FSet A).
Proof.
split ; toBool.
Defined. Defined.
End set_lattice. End set_lattice.
@ -720,7 +688,7 @@ End set_lattice.
Section dec_eq. Section dec_eq.
Context {A : Type} `{DecidablePaths A} `{Univalence}. Context {A : Type} `{DecidablePaths A} `{Univalence}.
Instance fsets_dec_eq : DecidablePaths (FSet A). Global Instance dec_eq_fset : DecidablePaths (FSet A).
Proof. Proof.
intros X Y. intros X Y.
apply (decidable_equiv' ((Y X) * (X Y)) (eq_subset X Y)^-1). apply (decidable_equiv' ((Y X) * (X Y)) (eq_subset X Y)^-1).
@ -732,29 +700,32 @@ End dec_eq.
Section comprehension_properties. Section comprehension_properties.
Context {A : Type} `{Univalence}. Context {A : Type} `{Univalence}.
Lemma comprehension_false : forall (X : FSet A), {|X & fun _ => false|} = . Lemma comprehension_false : forall (X : FSet A), {|X | fun _ => false|} = .
Proof. Proof.
toHProp. toHProp.
Defined. Defined.
Lemma comprehension_subset : forall ϕ (X : FSet A), Lemma comprehension_subset : forall ϕ (X : FSet A),
{|X & ϕ|} X = X. {|X | ϕ|} X = X.
Proof. Proof.
toHProp. toHProp.
destruct (ϕ a) ; eauto with lattice_hints typeclass_instances. destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
- apply binary_idempotent.
- apply left_identity.
Defined. Defined.
Lemma comprehension_or : forall ϕ ψ (X : FSet A), Lemma comprehension_or : forall ϕ ψ (X : FSet A),
{|X & (fun a => orb (ϕ a) (ψ a))|} {|X | (fun a => orb (ϕ a) (ψ a))|}
= {|X & ϕ|} {|X & ψ|}. = {|X | ϕ|} {|X | ψ|}.
Proof. Proof.
toHProp. toHProp.
symmetry ; destruct (ϕ a) ; destruct (ψ a) symmetry ; destruct (ϕ a) ; destruct (ψ a)
; eauto with lattice_hints typeclass_instances. (* ; eauto with lattice_hints typeclass_instances; *)
; simpl; (apply binary_idempotent || apply left_identity || apply right_identity).
Defined. Defined.
Lemma comprehension_all : forall (X : FSet A), Lemma comprehension_all : forall (X : FSet A),
{|X & fun _ => true|} = X. {|X | fun _ => true|} = X.
Proof. Proof.
toHProp. toHProp.
Defined. Defined.

View File

@ -13,12 +13,16 @@ Section quantifiers.
- intros X Y. - intros X Y.
apply (BuildhProp (X * Y)). apply (BuildhProp (X * Y)).
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
(* TODO eauto hints *)
apply associativity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply commutativity.
- intros. - intros.
apply path_trunctype ; apply prod_unit_l. apply path_trunctype ; apply prod_unit_l.
- intros. - intros.
apply path_trunctype ; apply prod_unit_r. apply path_trunctype ; apply prod_unit_r.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
intros. apply binary_idempotent.
Defined. Defined.
Lemma all_intro X : forall (HX : forall x, x X -> P x), all X. Lemma all_intro X : forall (HX : forall x, x X -> P x), all X.
@ -62,10 +66,16 @@ Section quantifiers.
- apply P. - apply P.
- apply lor. - apply lor.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
(* TODO eauto with .. *)
apply associativity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply commutativity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply left_identity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
apply right_identity.
- eauto with lattice_hints typeclass_instances. - eauto with lattice_hints typeclass_instances.
intros. apply binary_idempotent.
Defined. Defined.
Lemma exist_intro X a : a X -> P a -> exist X. Lemma exist_intro X a : a X -> P a -> exist X.
@ -84,7 +94,7 @@ Section quantifiers.
* apply (inr (HX2 t Pa)). * apply (inr (HX2 t Pa)).
Defined. Defined.
Lemma exist_elim X : exist X -> hexists (fun a => a X * P a). Lemma exist_elim X : exist X -> hexists (fun a => a X * P a)%type.
Proof. Proof.
hinduction X ; try (intros ; apply path_ishprop). hinduction X ; try (intros ; apply path_ishprop).
- contradiction. - contradiction.
@ -132,7 +142,7 @@ Section simple_example.
Context `{Univalence}. Context `{Univalence}.
Definition P : nat -> hProp := fun n => BuildhProp(n = n). Definition P : nat -> hProp := fun n => BuildhProp(n = n).
Definition X : FSet nat := {|0|} {|1|}. Definition X : FSet nat := {|0%nat|} {|1%nat|}.
Definition simple_example : all P X. Definition simple_example : all P X.
Proof. Proof.

View File

@ -66,12 +66,12 @@ Section length.
Context {A : Type} `{Univalence}. Context {A : Type} `{Univalence}.
Variable (length : FSet A -> nat) Variable (length : FSet A -> nat)
(length_singleton : forall (a : A), length {|a|} = 1) (length_singleton : forall (a : A), length {|a|} = 1%nat)
(length_one : forall (X : FSet A), length X = 1 -> hexists (fun a => X = {|a|})). (length_one : forall (X : FSet A), length X = 1%nat -> hexists (fun a => X = {|a|})).
Theorem dec_length (a b : A) : Decidable(merely(a = b)). Theorem dec_length (a b : A) : Decidable(merely(a = b)).
Proof. Proof.
destruct (dec (length ({|a|} {|b|}) = 1)). destruct (dec (length ({|a|} {|b|}) = 1%nat)).
- refine (inl _). - refine (inl _).
pose (length_one _ p) as Hp. pose (length_one _ p) as Hp.
simple refine (Trunc_rec _ Hp). simple refine (Trunc_rec _ Hp).

View File

@ -1,84 +1,54 @@
(** If [A] has a total order, then we can pick the minimum of finite sets. *) (** If [A] has a total order, then we can pick the minimum of finite sets. *)
Require Import HoTT HitTactics. Require Import HoTT HitTactics.
Require Import HoTT.Classes.interfaces.orders.
Require Import kuratowski.kuratowski_sets kuratowski.operations kuratowski.properties. Require Import kuratowski.kuratowski_sets kuratowski.operations kuratowski.properties.
Definition relation A := A -> A -> Type.
Section TotalOrder.
Class IsTop (A : Type) (R : relation A) (a : A) :=
top_max : forall x, R x a.
Class LessThan (A : Type) :=
leq : relation A.
Class Antisymmetric {A} (R : relation A) :=
antisymmetry : forall x y, R x y -> R y x -> x = y.
Class Total {A} (R : relation A) :=
total : forall x y, x = y \/ R x y \/ R y x.
Class TotalOrder (A : Type) {R : LessThan A} :=
{ TotalOrder_Reflexive :> Reflexive R | 2 ;
TotalOrder_Antisymmetric :> Antisymmetric R | 2;
TotalOrder_Transitive :> Transitive R | 2;
TotalOrder_Total :> Total R | 2; }.
End TotalOrder.
Section minimum. Section minimum.
Context {A : Type}. Context {A : Type}.
Context `{TotalOrder A}. Context (Ale : Le A).
Context `{@TotalOrder A Ale}.
Class IsTop `{Top A} := is_top : forall (x : A), x .
Definition min (x y : A) : A. Definition min (x y : A) : A.
Proof. Proof.
destruct (@total _ R _ x y). destruct (total _ x y).
- apply x. - apply x.
- destruct s as [s | s]. - apply y.
* apply x.
* apply y.
Defined. Defined.
Lemma min_spec1 x y : R (min x y) x. Lemma min_spec1 x y : (min x y) x.
Proof. Proof.
unfold min. unfold min.
destruct (total x y) ; simpl. destruct (total _ x y) ; simpl.
- reflexivity. - reflexivity.
- destruct s as [ | t]. - assumption.
* reflexivity.
* apply t.
Defined. Defined.
Lemma min_spec2 x y z : R z x -> R z y -> R z (min x y). Lemma min_spec2 x y z : z x -> z y -> z (min x y).
Proof. Proof.
intros. intros.
unfold min. unfold min.
destruct (total x y) as [ | s]. destruct (total _ x y); simpl; assumption.
* assumption.
* try (destruct s) ; assumption.
Defined. Defined.
Lemma min_comm x y : min x y = min y x. Lemma min_comm x y : min x y = min y x.
Proof. Proof.
unfold min. unfold min.
destruct (total x y) ; destruct (total y x) ; simpl. destruct (total _ x y) ; destruct (total _ y x) ; simpl; try reflexivity.
- assumption. + apply antisymmetry with (); [apply _|assumption..].
- destruct s as [s | s] ; auto. + apply antisymmetry with (); [apply _|assumption..].
- destruct s as [s | s] ; symmetry ; auto.
- destruct s as [s | s] ; destruct s0 as [s0 | s0] ; try reflexivity.
* apply (@antisymmetry _ R _ _) ; assumption.
* apply (@antisymmetry _ R _ _) ; assumption.
Defined. Defined.
Lemma min_idem x : min x x = x. Lemma min_idem x : min x x = x.
Proof. Proof.
unfold min. unfold min.
destruct (total x x) ; simpl. destruct (total _ x x) ; simpl; reflexivity.
- reflexivity.
- destruct s ; reflexivity.
Defined. Defined.
Lemma min_assoc x y z : min (min x y) z = min x (min y z). Lemma min_assoc x y z : min (min x y) z = min x (min y z).
Proof. Proof.
apply (@antisymmetry _ R _ _). apply antisymmetry with (). apply _.
- apply min_spec2. - apply min_spec2.
* etransitivity ; apply min_spec1. * etransitivity ; apply min_spec1.
* apply min_spec2. * apply min_spec2.
@ -96,43 +66,36 @@ Section minimum.
; rewrite min_comm ; apply min_spec1. ; rewrite min_comm ; apply min_spec1.
Defined. Defined.
Variable (top : A). Lemma min_nr `{IsTop} x : min x = x.
Context `{IsTop A R top}.
Lemma min_nr x : min x top = x.
Proof. Proof.
intros. intros.
unfold min. unfold min.
destruct (total x top). destruct (total _ x ).
- reflexivity. - reflexivity.
- destruct s. - apply antisymmetry with (). apply _.
* reflexivity. + assumption.
* apply (@antisymmetry _ R _ _). + apply is_top.
** assumption.
** refine (top_max _). apply _.
Defined. Defined.
Lemma min_nl x : min top x = x. Lemma min_nl `{IsTop} x : min x = x.
Proof. Proof.
rewrite min_comm. rewrite min_comm.
apply min_nr. apply min_nr.
Defined. Defined.
Lemma min_top_l x y : min x y = top -> x = top. Lemma min_top_l `{IsTop} x y : min x y = -> x = .
Proof. Proof.
unfold min. unfold min.
destruct (total x y). destruct (total _ x y) as [?|s].
- apply idmap. - apply idmap.
- destruct s as [s | s]. - intros X.
* apply idmap. rewrite X in s.
* intros X. apply antisymmetry with (). apply _.
rewrite X in s. * apply is_top.
apply (@antisymmetry _ R _ _). * assumption.
** apply top_max.
** assumption.
Defined. Defined.
Lemma min_top_r x y : min x y = top -> y = top. Lemma min_top_r `{IsTop} x y : min x y = -> y = .
Proof. Proof.
rewrite min_comm. rewrite min_comm.
apply min_top_l. apply min_top_l.
@ -144,72 +107,60 @@ Section add_top.
Variable (A : Type). Variable (A : Type).
Context `{TotalOrder A}. Context `{TotalOrder A}.
Definition Top := A + Unit. Definition Topped := (A + Unit)%type.
Definition top : Top := inr tt. Global Instance top_topped : Top Topped := inr tt.
Global Instance RTop : LessThan Top. Global Instance RTop : Le Topped.
Proof. Proof.
unfold relation. intros [a1 | ?] [a2 | ?].
induction 1 as [a1 | ] ; induction 1 as [a2 | ]. - apply (a1 a2).
- apply (R a1 a2).
- apply Unit_hp. - apply Unit_hp.
- apply False_hp. - apply False_hp.
- apply Unit_hp. - apply Unit_hp.
Defined. Defined.
Global Instance rtop_hprop : Instance rtop_hprop :
is_mere_relation A R -> is_mere_relation Top RTop. is_mere_relation A R -> is_mere_relation Topped RTop.
Proof. Proof.
intros P a b. intros ? P a b.
destruct a ; destruct b ; apply _. destruct a ; destruct b ; simpl; apply _.
Defined. Defined.
Global Instance RTopOrder : TotalOrder Top. Global Instance RTopOrder : TotalOrder RTop.
Proof. Proof.
split. repeat split; try apply _.
- intros x ; induction x ; unfold RTop ; simpl. - intros [?|?] ; cbv.
* reflexivity. * reflexivity.
* apply tt. * apply tt.
- intros x y ; induction x as [a1 | ] ; induction y as [a2 | ] ; unfold RTop ; simpl - intros [a1 | []] [a2 | []] [a3 | []]; cbv
; try contradiction. ; try contradiction; try (by intros; apply tt).
* intros ; f_ap. apply transitivity.
apply (@antisymmetry _ R _ _) ; assumption. - intros [a1 |[]] [a2 |[]]; cbv; try contradiction.
* intros ; induction b ; induction b0. + intros. f_ap. apply antisymmetry with Ale; [apply _|assumption..].
reflexivity. + intros; reflexivity.
- intros x y z ; induction x as [a1 | b1] ; induction y as [a2 | b2] - intros [a1|[]] [a2|[]]; cbv.
; induction z as [a3 | b3] ; unfold RTop ; simpl * apply total. apply _.
; try contradiction ; intros ; try (apply tt). * apply (inl tt).
transitivity a2 ; assumption. * apply (inr tt).
- intros x y. * apply (inr tt).
unfold RTop ; simpl.
induction x as [a1 | b1] ; induction y as [a2 | b2] ; try (apply (inl idpath)).
* destruct (TotalOrder_Total a1 a2).
** left ; f_ap ; assumption.
** right ; assumption.
* apply (inr(inl tt)).
* apply (inr(inr tt)).
* left ; induction b1 ; induction b2 ; reflexivity.
Defined. Defined.
Global Instance top_a_top : IsTop Top RTop top. Global Instance is_top_topped : IsTop RTop.
Proof. Proof. intros [?|?]; cbv; apply tt. Defined.
intro x ; destruct x ; apply tt.
Defined.
End add_top. End add_top.
(** If [A] has a total order, then a nonempty finite set has a minimum element. *) (** If [A] has a total order, then a nonempty finite set has a minimum element. *)
Section min_set. Section min_set.
Variable (A : Type). Variable (A : Type).
Context `{TotalOrder A}. Context `{TotalOrder A}.
Context `{is_mere_relation A R}.
Context `{Univalence} `{IsHSet A}. Context `{Univalence} `{IsHSet A}.
Definition min_set : FSet A -> Top A. Definition min_set : FSet A -> Topped A.
Proof. Proof.
hrecursion. hrecursion.
- apply (top A). - apply .
- apply inl. - apply inl.
- apply min. - apply min with (RTop A). apply _.
- intros ; symmetry ; apply min_assoc. - intros ; symmetry ; apply min_assoc.
- apply min_comm. - apply min_comm.
- apply min_nl. apply _. - apply min_nl. apply _.
@ -217,7 +168,7 @@ Section min_set.
- intros ; apply min_idem. - intros ; apply min_idem.
Defined. Defined.
Definition empty_min : forall (X : FSet A), min_set X = top A -> X = . Definition empty_min : forall (X : FSet A), min_set X = -> X = .
Proof. Proof.
simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _) simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _)
; try (intros ; apply path_forall ; intro q ; apply set_path2) ; try (intros ; apply path_forall ; intro q ; apply set_path2)
@ -230,21 +181,19 @@ Section min_set.
refine (not_is_inl_and_inr' (inl a) _ _). refine (not_is_inl_and_inr' (inl a) _ _).
* apply tt. * apply tt.
* rewrite X ; apply tt. * rewrite X ; apply tt.
- intros. - intros x y ? ? ?.
assert (min_set x = top A). assert (min_set x = ).
{ { simple refine (min_top_l _ _ (min_set y) _) ; assumption. }
simple refine (min_top_l _ _ (min_set y) _) ; assumption.
}
rewrite (X X2). rewrite (X X2).
rewrite nl. rewrite nl.
assert (min_set y = top A). assert (min_set y = ).
{ simple refine (min_top_r _ (min_set x) _ _) ; assumption. } { simple refine (min_top_r _ (min_set x) _ _) ; assumption. }
rewrite (X0 X3). rewrite (X0 X3).
reflexivity. reflexivity.
Defined. Defined.
Definition min_set_spec (a : A) : forall (X : FSet A), Definition min_set_spec (a : A) : forall (X : FSet A),
a X -> RTop A (min_set X) (inl a). a X -> min_set X inl a.
Proof. Proof.
simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _) simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _)
; try (intros ; apply path_ishprop) ; try (intros ; apply path_ishprop)
@ -259,13 +208,13 @@ Section min_set.
unfold member in X, X0. unfold member in X, X0.
destruct X1. destruct X1.
* specialize (X t). * specialize (X t).
assert (RTop A (min (min_set x) (min_set y)) (min_set x)) as X1. assert (RTop A (min _ (min_set x) (min_set y)) (min_set x)) as X1.
{ apply min_spec1. } { apply min_spec1. }
etransitivity. etransitivity.
{ apply X1. } { apply X1. }
assumption. assumption.
* specialize (X0 t). * specialize (X0 t).
assert (RTop A (min (min_set x) (min_set y)) (min_set y)) as X1. assert (RTop A (min _ (min_set x) (min_set y)) (min_set y)) as X1.
{ rewrite min_comm ; apply min_spec1. } { rewrite min_comm ; apply min_spec1. }
etransitivity. etransitivity.
{ apply X1. } { apply X1. }

View File

@ -8,9 +8,9 @@ Section finite_hott.
Context `{Univalence}. Context `{Univalence}.
(* A subobject is B-finite if its extension is B-finite as a type *) (* A subobject is B-finite if its extension is B-finite as a type *)
Definition Bfin (X : Sub A) : hProp := BuildhProp (Finite {a : A & a X}). Definition Bfin (X : Sub A) : hProp := BuildhProp (Finite {a : A | a X}).
Global Instance singleton_contr a `{IsHSet A} : Contr {b : A & b {|a|}}. Global Instance singleton_contr a `{IsHSet A} : Contr {b : A | b {|a|}}.
Proof. Proof.
exists (a; tr idpath). exists (a; tr idpath).
intros [b p]. intros [b p].
@ -20,7 +20,7 @@ Section finite_hott.
apply p^. apply p^.
Defined. Defined.
Definition singleton_fin_equiv' a : Fin 1 -> {b : A & b {|a|}}. Definition singleton_fin_equiv' a : Fin 1 -> {b : A | b {|a|}}.
Proof. Proof.
intros _. apply (a; tr idpath). intros _. apply (a; tr idpath).
Defined. Defined.
@ -66,7 +66,8 @@ Section finite_hott.
: DecidablePaths A. : DecidablePaths A.
Proof. Proof.
intros a b. intros a b.
destruct (U {|a|} {|b|} (singleton a) (singleton b)) as [n pn]. specialize (U {|a|} {|b|} (singleton a) (singleton b)).
destruct U as [n pn].
strip_truncations. strip_truncations.
unfold Sect in *. unfold Sect in *.
destruct pn as [f [g fg gf _]], n as [ | [ | n]]. destruct pn as [f [g fg gf _]], n as [ | [ | n]].
@ -75,9 +76,9 @@ Section finite_hott.
apply (tr(inl(tr idpath))). apply (tr(inl(tr idpath))).
- refine (inl _). - refine (inl _).
pose (s1 := (a;tr(inl(tr idpath))) pose (s1 := (a;tr(inl(tr idpath)))
: {c : A & Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}). : {c : A | Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}).
pose (s2 := (b;tr(inr(tr idpath))) pose (s2 := (b;tr(inr(tr idpath)))
: {c : A & Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}). : {c : A | Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}).
refine (ap (fun x => x.1) (gf s1)^ @ _ @ (ap (fun x => x.1) (gf s2))). refine (ap (fun x => x.1) (gf s1)^ @ _ @ (ap (fun x => x.1) (gf s2))).
assert (fs_eq : f s1 = f s2). assert (fs_eq : f s1 = f s2).
{ by apply path_ishprop. } { by apply path_ishprop. }
@ -116,9 +117,9 @@ Section singleton_set.
Variable (A : Type). Variable (A : Type).
Context `{Univalence}. Context `{Univalence}.
Variable (HA : forall a, {b : A & b {|a|}} <~> Fin 1). Variable (HA : forall a, {b : A | b {|a|}} <~> Fin 1).
Definition el x : {b : A & b {|x|}} := (x;tr idpath). Definition el x : {b : A | b {|x|}} := (x;tr idpath).
Theorem single_bfin_set : forall (x : A) (p : x = x), p = idpath. Theorem single_bfin_set : forall (x : A) (p : x = x), p = idpath.
Proof. Proof.
@ -132,7 +133,7 @@ Section singleton_set.
rewrite <- ap_compose. rewrite <- ap_compose.
enough(forall r, enough(forall r,
(eissect HA X)^ (eissect HA X)^
@ ap (fun x0 : {b : A & b {|x|}} => HA^-1 (HA x0)) r @ ap (fun x0 : {b : A | b {|x|}} => HA^-1 (HA x0)) r
@ eissect HA X = r @ eissect HA X = r
) as H2. ) as H2.
{ {
@ -166,7 +167,7 @@ End singleton_set.
Section empty. Section empty.
Variable (A : Type). Variable (A : Type).
Variable (X : A -> hProp) Variable (X : A -> hProp)
(Xequiv : {a : A & a X} <~> Fin 0). (Xequiv : {a : A | a X} <~> Fin 0).
Context `{Univalence}. Context `{Univalence}.
Lemma X_empty : X = . Lemma X_empty : X = .
@ -185,10 +186,10 @@ Section split.
Variable (A : Type). Variable (A : Type).
Variable (P : A -> hProp) Variable (P : A -> hProp)
(n : nat) (n : nat)
(f : {a : A & P a } <~> Fin n + Unit). (f : {a : A | P a } <~> Fin n + Unit).
Definition split : exists P' : Sub A, exists b : A, Definition split : exists P' : Sub A, exists b : A,
({a : A & P' a} <~> Fin n) * (forall x, P x = (P' x merely (x = b))). prod ({a : A | P' a} <~> Fin n) (forall x, P x = (P' x merely (x = b))).
Proof. Proof.
pose (fun x : A => sig (fun y : Fin n => x = (f^-1 (inl y)).1)) as P'. pose (fun x : A => sig (fun y : Fin n => x = (f^-1 (inl y)).1)) as P'.
assert (forall x, IsHProp (P' x)). assert (forall x, IsHProp (P' x)).
@ -267,7 +268,7 @@ Arguments Bfin {_} _.
Arguments split {_} {_} _ _ _. Arguments split {_} {_} _ _ _.
Section Bfin_no_singletons. Section Bfin_no_singletons.
Definition S1toSig (x : S1) : {x : S1 & merely(x = base)}. Definition S1toSig (x : S1) : {x : S1 | merely(x = base)}.
Proof. Proof.
exists x. exists x.
simple refine (S1_ind (fun z => merely(z = base)) (tr idpath) _ x). simple refine (S1_ind (fun z => merely(z = base)) (tr idpath) _ x).
@ -382,7 +383,7 @@ Section kfin_bfin.
Lemma notIn_ext_union_singleton (b : A) (Y : Sub A) : Lemma notIn_ext_union_singleton (b : A) (Y : Sub A) :
~ (b Y) -> ~ (b Y) ->
{a : A & a ({|b|} Y)} <~> {a : A & a Y} + Unit. {a : A | a ({|b|} Y)} <~> {a : A | a Y} + Unit.
Proof. Proof.
intros HYb. intros HYb.
unshelve eapply BuildEquiv. unshelve eapply BuildEquiv.
@ -422,7 +423,7 @@ Section kfin_bfin.
strip_truncations. strip_truncations.
revert fX. revert X. revert fX. revert X.
induction n; intros X fX. induction n; intros X fX.
- rewrite (X_empty _ X fX), (neutralL_max (Sub A)). - rewrite (X_empty _ X fX). rewrite left_identity.
apply HY. apply HY.
- destruct (split X n fX) as - destruct (split X n fX) as
(X' & b & HX' & HX). (X' & b & HX' & HX).
@ -431,10 +432,10 @@ Section kfin_bfin.
+ cut (X Y = X' Y). + cut (X Y = X' Y).
{ intros HXY. rewrite HXY. assumption. } { intros HXY. rewrite HXY. assumption. }
apply path_forall. intro a. apply path_forall. intro a.
unfold union, sub_union, max_fun. unfold union, sub_union, join, join_fun.
rewrite HX. rewrite HX.
rewrite (commutativity (X' a)). rewrite (commutativity (X' a)).
rewrite (associativity _ (X' a)). rewrite <- (associativity _ (X' a)).
apply path_iff_hprop. apply path_iff_hprop.
* intros Ha. * intros Ha.
strip_truncations. strip_truncations.
@ -448,15 +449,14 @@ Section kfin_bfin.
strip_truncations. strip_truncations.
exists (n'.+1). exists (n'.+1).
apply tr. apply tr.
transitivity ({a : A & a (fun a => merely (a = b)) (X' Y)}). transitivity ({a : A | a (fun a => merely (a = b)) (X' Y)}).
* apply equiv_functor_sigma_id. * apply equiv_functor_sigma_id.
intro a. intro a.
rewrite <- (associative_max (Sub A)). rewrite (associativity (fun a => merely (a = b)) X' Y).
assert (X = X' (fun a => merely (a = b))) as HX_. assert (X = X' (fun a => merely (a = b))) as HX_.
** apply path_forall. intros ?. ** apply path_forall. intros ?.
unfold union, sub_union, max_fun.
apply HX. apply HX.
** rewrite HX_, <- (commutative_max (Sub A) X'). ** rewrite HX_, <- (commutativity X').
reflexivity. reflexivity.
* etransitivity. * etransitivity.
{ apply (notIn_ext_union_singleton b _ HX'Yb). } { apply (notIn_ext_union_singleton b _ HX'Yb). }
@ -466,7 +466,7 @@ Section kfin_bfin.
Definition FSet_to_Bfin : forall (X : FSet A), Bfin (map X). Definition FSet_to_Bfin : forall (X : FSet A), Bfin (map X).
Proof. Proof.
hinduction; try (intros; apply path_ishprop). hinduction; try (intros; apply path_ishprop).
- exists 0. - exists 0%nat.
apply tr. apply tr.
simple refine (BuildEquiv _ _ _ _). simple refine (BuildEquiv _ _ _ _).
destruct 1 as [? []]. destruct 1 as [? []].

View File

@ -300,7 +300,7 @@ Section subobjects.
simpl. apply path_ishprop. simpl. apply path_ishprop.
Defined. Defined.
Fixpoint weaken_list_r (P Q : Sub A) (ls : list (sigT Q)) : list (sigT (max_L P Q)). Fixpoint weaken_list_r (P Q : Sub A) (ls : list (sigT Q)) : list (sigT (P Q)).
Proof. Proof.
destruct ls as [|[x Hx] ls]. destruct ls as [|[x Hx] ls].
- exact nil. - exact nil.
@ -323,7 +323,7 @@ Section subobjects.
Defined. Defined.
Fixpoint concatD {P Q : Sub A} Fixpoint concatD {P Q : Sub A}
(ls : list (sigT P)) (ls' : list (sigT Q)) : list (sigT (max_L P Q)). (ls : list (sigT P)) (ls' : list (sigT Q)) : list (sigT (P Q)).
Proof. Proof.
destruct ls as [|[y Hy] ls]. destruct ls as [|[y Hy] ls].
- apply weaken_list_r. exact ls'. - apply weaken_list_r. exact ls'.
@ -356,7 +356,7 @@ Section subobjects.
Defined. Defined.
Lemma enumeratedS_union (P Q : Sub A) : Lemma enumeratedS_union (P Q : Sub A) :
enumeratedS P -> enumeratedS Q -> enumeratedS (max_L P Q). enumeratedS P -> enumeratedS Q -> enumeratedS (P Q).
Proof. Proof.
intros HP HQ. intros HP HQ.
strip_truncations; apply tr. strip_truncations; apply tr.
@ -388,7 +388,7 @@ Section subobjects.
Transparent enumeratedS. Transparent enumeratedS.
Instance hprop_sub_fset (P : Sub A) : Instance hprop_sub_fset (P : Sub A) :
IsHProp {X : FSet A & k_finite.map X = P}. IsHProp {X : FSet A | k_finite.map X = P}.
Proof. Proof.
apply hprop_allpath. intros [X HX] [Y HY]. apply hprop_allpath. intros [X HX] [Y HY].
assert (X = Y) as HXY. assert (X = Y) as HXY.
@ -433,7 +433,7 @@ Section subobjects.
Definition enumeratedS_to_FSet : Definition enumeratedS_to_FSet :
forall (P : Sub A), enumeratedS P -> forall (P : Sub A), enumeratedS P ->
{X : FSet A & k_finite.map X = P}. {X : FSet A | k_finite.map X = P}.
Proof. Proof.
intros P HP. intros P HP.
strip_truncations. strip_truncations.

View File

@ -33,7 +33,7 @@ Section k_finite.
Definition Kf : hProp := Kf_sub (fun x => True). Definition Kf : hProp := Kf_sub (fun x => True).
Instance: IsHProp {X : FSet A & forall a : A, map X a}. Instance: IsHProp {X : FSet A | forall a : A, map X a}.
Proof. Proof.
apply hprop_allpath. apply hprop_allpath.
intros [X PX] [Y PY]. intros [X PX] [Y PY].
@ -77,19 +77,16 @@ Section structure_k_finite.
Context (A : Type). Context (A : Type).
Context `{Univalence}. Context `{Univalence}.
Lemma map_union : forall X Y : FSet A, map (X Y) = max_fun (map X) (map Y). Lemma map_union : forall X Y : FSet A, map (X Y) = (map X) (map Y).
Proof. Proof.
intros. intros.
unfold map, max_fun.
reflexivity. reflexivity.
Defined. Defined.
Lemma k_finite_union : closedUnion (Kf_sub A). Lemma k_finite_union : closedUnion (Kf_sub A).
Proof. Proof.
unfold closedUnion, Kf_sub, Kf_sub_intern. unfold closedUnion, Kf_sub, Kf_sub_intern.
intros. intros X Y [SX XP] [SY YP].
destruct X0 as [SX XP].
destruct X1 as [SY YP].
exists (SX SY). exists (SX SY).
rewrite map_union. rewrite map_union.
rewrite XP, YP. rewrite XP, YP.
@ -179,9 +176,9 @@ Section k_properties.
| inl a => ({|a|} : FSet A) | inl a => ({|a|} : FSet A)
| inr b => | inr b =>
end). end).
exists (bind _ (fset_fmap f X)). exists (mjoin (fset_fmap f X)).
intro a. intro a.
apply bind_isIn. apply mjoin_isIn.
specialize (HX (inl a)). specialize (HX (inl a)).
exists {|a|}. split; [ | apply tr; reflexivity ]. exists {|a|}. split; [ | apply tr; reflexivity ].
apply (fmap_isIn f (inl a) X). apply (fmap_isIn f (inl a) X).
@ -259,33 +256,35 @@ Section alternative_definition.
Local Ltac help_solve := Local Ltac help_solve :=
repeat (let x := fresh in intro x ; destruct x) ; intros repeat (let x := fresh in intro x ; destruct x) ; intros
; try (simple refine (path_sigma _ _ _ _ _)) ; try (apply path_ishprop) ; simpl ; try (simple refine (path_sigma _ _ _ _ _)) ; try (apply path_ishprop) ; simpl
; unfold union, sub_union, max_fun ; unfold union, sub_union, join, join_fun
; apply path_forall ; apply path_forall
; intro z ; intro z
; eauto with lattice_hints typeclass_instances. ; eauto with lattice_hints typeclass_instances.
Definition fset_to_k : FSet A -> {P : A -> hProp & kf_sub P}. Definition fset_to_k : FSet A -> {P : A -> hProp | kf_sub P}.
Proof. Proof.
hinduction. assert (IsHSet {P : A -> hProp | kf_sub P}) as Hs.
{ apply trunc_sigma. }
simple refine (FSet_rec A {P : A -> hProp | kf_sub P} Hs _ _ _ _ _ _ _ _).
- exists . - exists .
auto. simpl. auto.
- intros a. - intros a.
exists {|a|}. exists {|a|}.
auto. simpl. auto.
- intros [P1 HP1] [P2 HP2]. - intros [P1 HP1] [P2 HP2].
exists (P1 P2). exists (P1 P2).
intros ? ? ? HP. intros ? ? ? HP.
apply HP. apply HP.
* apply HP1 ; assumption. * apply HP1 ; assumption.
* apply HP2 ; assumption. * apply HP2 ; assumption.
- help_solve. - help_solve. (* TODO: eauto *) apply associativity.
- help_solve. - help_solve. apply commutativity.
- help_solve. - help_solve. apply left_identity.
- help_solve. - help_solve. apply right_identity.
- help_solve. - help_solve. apply binary_idempotent.
Defined. Defined.
Definition k_to_fset : {P : A -> hProp & kf_sub P} -> FSet A. Definition k_to_fset : {P : A -> hProp | kf_sub P} -> FSet A.
Proof. Proof.
intros [P HP]. intros [P HP].
destruct (HP (Kf_sub _) (k_finite_empty _) (k_finite_singleton _) (k_finite_union _)). destruct (HP (Kf_sub _) (k_finite_empty _) (k_finite_singleton _) (k_finite_union _)).
@ -299,7 +298,7 @@ Section alternative_definition.
refine ((ap (fun z => _ z) HX2^)^ @ (ap (fun z => z X2) HX1^)^). refine ((ap (fun z => _ z) HX2^)^ @ (ap (fun z => z X2) HX1^)^).
Defined. Defined.
Lemma k_to_fset_to_k (X : {P : A -> hProp & kf_sub P}) : fset_to_k(k_to_fset X) = X. Lemma k_to_fset_to_k (X : {P : A -> hProp | kf_sub P}) : fset_to_k(k_to_fset X) = X.
Proof. Proof.
simple refine (path_sigma _ _ _ _ _) ; try (apply path_ishprop). simple refine (path_sigma _ _ _ _ _) ; try (apply path_ishprop).
apply path_forall. apply path_forall.

View File

@ -6,9 +6,9 @@ Section subobjects.
Definition Sub := A -> hProp. Definition Sub := A -> hProp.
Global Instance sub_empty : hasEmpty Sub := fun _ => False_hp. Global Instance sub_empty : hasEmpty Sub := fun _ => .
Global Instance sub_union : hasUnion Sub := max_fun. Global Instance sub_union : hasUnion Sub := join.
Global Instance sub_intersection : hasIntersection Sub := min_fun. Global Instance sub_intersection : hasIntersection Sub := meet.
Global Instance sub_singleton : hasSingleton Sub A Global Instance sub_singleton : hasSingleton Sub A
:= fun a b => BuildhProp (Trunc (-1) (b = a)). := fun a b => BuildhProp (Trunc (-1) (b = a)).
Global Instance sub_membership : hasMembership Sub A := fun a X => X a. Global Instance sub_membership : hasMembership Sub A := fun a X => X a.