diff --git a/FiniteSets/implementations/lists.v b/FiniteSets/implementations/lists.v index 09ef949..018e144 100644 --- a/FiniteSets/implementations/lists.v +++ b/FiniteSets/implementations/lists.v @@ -62,7 +62,7 @@ Section ListToSet. Definition empty_empty : list_to_set A ∅ = ∅ := idpath. 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. induction l ; cbn in *. - reflexivity. @@ -170,11 +170,11 @@ Section refinement_examples. Defined. 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. rewrite list_exist_set. 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. { intros H1. @@ -240,4 +240,4 @@ Section refinement_examples. rewrite length_compute, fset_list_memb. reflexivity. Defined. -End refinement_examples. \ No newline at end of file +End refinement_examples. diff --git a/FiniteSets/interfaces/lattice_examples.v b/FiniteSets/interfaces/lattice_examples.v index 647f937..c613f3e 100644 --- a/FiniteSets/interfaces/lattice_examples.v +++ b/FiniteSets/interfaces/lattice_examples.v @@ -1,5 +1,5 @@ (** Some examples of lattices. *) -Require Import HoTT lattice_interface. +Require Export HoTT lattice_interface. (** [Bool] is a lattice. *) Section BoolLattice. @@ -10,14 +10,18 @@ Section BoolLattice. ; auto ; try contradiction. - Instance maximum_bool : maximum Bool := orb. - Instance minimum_bool : minimum Bool := andb. - Instance bottom_bool : bottom Bool := false. + Instance maximum_bool : Join Bool := orb. + Instance minimum_bool : Meet Bool := andb. + 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. - Proof. - split ; solve_bool. - Defined. + Proof. split ; (apply _ || solve_bool). Defined. Definition and_true : forall b, andb b true = b. Proof. @@ -50,12 +54,12 @@ End BoolLattice. 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 commutativity : bool_lattice_hints. -Hint Resolve absorb : bool_lattice_hints. +Hint Resolve absorption : bool_lattice_hints. Hint Resolve idempotency : bool_lattice_hints. -Hint Resolve neutralityL : bool_lattice_hints. -Hint Resolve neutralityR : bool_lattice_hints. +Hint Resolve left_identity : bool_lattice_hints. +Hint Resolve right_identity : bool_lattice_hints. Hint Resolve associativity @@ -66,25 +70,51 @@ Hint Resolve (** If [B] is a lattice, then [A -> B] is a lattice. *) Section fun_lattice. Context {A B : Type}. - Context `{Lattice B}. + Context `{BJoin : Join B}. + Context `{BMeet : Meet B}. + Context `{@Lattice B BJoin BMeet}. Context `{Funext}. + Context `{BBottom : Bottom B}. - Global Instance max_fun : maximum (A -> B) := - fun (f g : A -> B) (a : A) => max_L0 (f a) (g a). + Global Instance bot_fun : Bottom (A -> B) + := fun _ => ⊥. - Global Instance min_fun : minimum (A -> B) := - fun (f g : A -> B) (a : A) => min_L0 (f a) (g a). + Global Instance join_fun : Join (A -> B) := + fun (f g : A -> B) (a : A) => (f a) ⊔ (g a). - Global Instance bot_fun : bottom (A -> B) - := fun _ => empty_L. + Global Instance meet_fun : Meet (A -> B) := + fun (f g : A -> B) (a : A) => (f a) ⊓ (g a). Ltac solve_fun := compute ; intros ; apply path_forall ; intro ; 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). 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. End fun_lattice. @@ -92,24 +122,26 @@ End fun_lattice. Section sub_lattice. Context {A : Type} {P : A -> hProp}. Context `{Lattice A}. - Context {Hmax : forall x y, P x -> P y -> P (max_L0 x y)}. - Context {Hmin : forall x y, P x -> P y -> P (min_L0 x y)}. - Context {Hbot : P empty_L}. + Context `{Bottom A}. + Context {Hmax : forall x y, P x -> P y -> P (x ⊔ y)}. + Context {Hmin : forall x y, P x -> P y -> P (x ⊓ y)}. + Context {Hbot : 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 => 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. - Instance minAP : minimum AP := + Instance minAP : Meet AP := fun x y => 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. Instance hprop_sub : forall c, IsHProp (P c). @@ -127,19 +159,25 @@ Section sub_lattice. Global Instance lattice_sub : Lattice AP. 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. 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. Notation "A ∨ B" := (lor A B) (at level 20, right associativity) : logic_scope. Arguments lor _%L _%L. Open Scope logic_scope. -Instance land : minimum hProp := fun X Y => BuildhProp (prod X Y). -Instance lfalse : bottom hProp := False_hp. +Instance land : Meet hProp := fun X Y => BuildhProp (prod X Y). +Instance lfalse : Bottom hProp := False_hp. Notation "A ∧ B" := (land A B) (at level 20, right associativity) : logic_scope. Arguments land _%L _%L. @@ -181,17 +219,18 @@ Section hPropLattice. intros X Y Z. symmetry. apply path_hprop. + symmetry. apply equiv_prod_assoc. Defined. - Instance lor_idempotent : Idempotent lor. + Instance lor_idempotent : BinaryIdempotent lor. Proof. intros X. apply path_iff_hprop ; lor_intros ; try(refine (tr(inl _))) ; auto. Defined. - Instance land_idempotent : Idempotent land. + Instance land_idempotent : BinaryIdempotent land. Proof. intros X. apply path_iff_hprop ; cbn. @@ -199,14 +238,14 @@ Section hPropLattice. - intros a ; apply (pair a a). Defined. - Instance lor_neutrall : NeutralL lor lfalse. + Instance lor_neutrall : LeftIdentity lor lfalse. Proof. intros X. apply path_iff_hprop ; lor_intros ; try contradiction ; try (refine (tr(inr _))) ; auto. Defined. - Instance lor_neutralr : NeutralR lor lfalse. + Instance lor_neutralr : RightIdentity lor lfalse. Proof. intros X. apply path_iff_hprop ; lor_intros ; try contradiction @@ -235,16 +274,17 @@ Section hPropLattice. * apply (tr (inl X)). Defined. - Global Instance lattice_hprop : Lattice hProp := - { commutative_min := _ ; - commutative_max := _ ; - associative_min := _ ; - associative_max := _ ; - idempotent_min := _ ; - idempotent_max := _ ; - neutralL_max := _ ; - neutralR_max := _ ; - absorption_min_max := _ ; - absorption_max_min := _ - }. + Global Instance lattice_hprop : Lattice hProp. + Proof. repeat (split ; try apply _). Defined. + + Global Instance bounded_jsl_hprop : BoundedJoinSemiLattice hProp. + Proof. repeat (split ; try apply _). Qed. + + Global Instance top_hprop : Top hProp := Unit_hp. + Global Instance bounded_msl_hprop : @BoundedSemiLattice hProp (⊓) ⊤. + Proof. + repeat (split; try apply _); cbv. + - intros X. apply path_trunctype ; apply prod_unit_l. + - intros X. apply path_trunctype ; apply prod_unit_r. + Defined. End hPropLattice. diff --git a/FiniteSets/interfaces/lattice_interface.v b/FiniteSets/interfaces/lattice_interface.v index 2084d1c..7997a88 100644 --- a/FiniteSets/interfaces/lattice_interface.v +++ b/FiniteSets/interfaces/lattice_interface.v @@ -1,119 +1,60 @@ (** Interface for lattices and join semilattices. *) 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. *) -Section binary_operation. - Definition operation (A : Type) := A -> A -> A. - - Variable (A : Type) - (f : operation A). +(* (** Join semilattices as a typeclass. They only have a join operator. *) *) +(* Section JoinSemiLattice. *) +(* Variable A : Type. *) +(* Context {max_L : Join A} {empty_L : Bottom A}. *) - Class Commutative := - commutativity : forall x y, f x y = f y x. +(* Class JoinSemiLattice := *) +(* { *) +(* 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 Associative := - associativity : forall x y z, f (f x y) z = f x (f y z). +(* Arguments JoinSemiLattice _ {_} {_}. *) - Class Idempotent := - idempotency : forall x, f x x = x. +(* 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. *) - Variable g : operation A. +(* (** 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 Absorption := - absorb : forall x y, f x (g x y) = x. +(* 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. *) - 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 _ {_} {_} {_}. +(* Arguments Lattice _ {_} {_} {_}. *) Create HintDb lattice_hints. Hint Resolve associativity : lattice_hints. -Hint Resolve (associativity _ _ _)^ : lattice_hints. +(* Hint Resolve (associativity _ _ _)^ : lattice_hints. *) Hint Resolve commutativity : lattice_hints. -Hint Resolve absorb : lattice_hints. +Hint Resolve absorption : lattice_hints. Hint Resolve idempotency : lattice_hints. -Hint Resolve neutralityL : lattice_hints. -Hint Resolve neutralityR : lattice_hints. +Hint Resolve left_identity : lattice_hints. +Hint Resolve right_identity : lattice_hints. diff --git a/FiniteSets/interfaces/monad_interface.v b/FiniteSets/interfaces/monad_interface.v index 7d748e1..b62ffd9 100644 --- a/FiniteSets/interfaces/monad_interface.v +++ b/FiniteSets/interfaces/monad_interface.v @@ -1,5 +1,7 @@ (** The structure of a monad. *) +(* TODO REMOVE THIS *) Require Import HoTT. +Require Export HoTT.Classes.interfaces.monad. Section functor. Variable (F : Type -> Type). @@ -11,35 +13,9 @@ Section functor. fmap F (g o f) = fmap F g o fmap F f }. End functor. - -Section monad_operations. - Variable (F : Type -> Type). - Class hasPure := - { - pure : forall (A : Type), A -> F A - }. +Section monad_join. + Context `{Monad M}. - 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. \ No newline at end of file + Definition mjoin {A} (m : M (M A)) : M A := bind m id. +End monad_join. diff --git a/FiniteSets/interfaces/set_interface.v b/FiniteSets/interfaces/set_interface.v index bf8519e..17901c2 100644 --- a/FiniteSets/interfaces/set_interface.v +++ b/FiniteSets/interfaces/set_interface.v @@ -16,7 +16,7 @@ Section interface. f_empty : forall A, f 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_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) }. @@ -175,12 +175,12 @@ Section quotient_properties. Proof. intros ϕ X. apply (quotient_iso _)^-1. - simple refine ({|_ & ϕ|}). + simple refine ({|_ | ϕ|}). apply (quotient_iso (f A) X). Defined. Definition well_defined_filter (A : Type) (ϕ : A -> Bool) (X : T A) : - {|class_of _ X & ϕ|} = class_of _ {|X & ϕ|}. + {|class_of _ X | ϕ|} = class_of _ {|X | ϕ|}. Proof. rewrite <- (eissect (quotient_iso _)). simpl. @@ -224,30 +224,35 @@ Section quotient_properties. apply (ap _ HXY). Defined. - Instance View_max A : maximum (View A). + Instance join_view A : Join (View A). Proof. apply view_union. 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. apply View_empty. Defined. - Ltac sq1 := autounfold ; intros ; try f_ap - ; rewrite ?(eisretr (quotient_iso _)) - ; eauto with lattice_hints typeclass_instances. + Ltac sq1 := autounfold; intros; + unfold view_union; try f_ap; + rewrite ?(eisretr (quotient_iso _)). + + Ltac sq2 := autounfold; intros; + unfold view_union; + rewrite <- (eissect (quotient_iso _)), ?(eisretr (quotient_iso _)); + f_ap; simpl. - Ltac sq2 := autounfold ; intros - ; rewrite <- (eissect (quotient_iso _)), ?(eisretr (quotient_iso _)) - ; f_ap ; simpl - ; reduce T ; eauto with lattice_hints typeclass_instances. - - Global Instance view_lattice A : JoinSemiLattice (View A). + Global Instance view_lattice A : BoundedJoinSemiLattice (View A). 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. End quotient_properties. @@ -316,7 +321,8 @@ Section properties. Ltac via_quotient := intros ; apply reflect_f_eq ; simpl ; 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), set_eq f (X ∪ Y) (Y ∪ X). @@ -328,6 +334,7 @@ Section properties. set_eq f ((X ∪ Y) ∪ Z) (X ∪ (Y ∪ Z)). Proof. via_quotient. + symmetry. apply associativity. Defined. Lemma union_idem : forall A (X : T A), @@ -376,4 +383,4 @@ Section refinement. (h : FSet A -> B) : T A -> B := fun X => h(quotient_iso (f A) (class_of _ X)). -End refinement. \ No newline at end of file +End refinement. diff --git a/FiniteSets/interfaces/set_names.v b/FiniteSets/interfaces/set_names.v index 5f9c480..07674d3 100644 --- a/FiniteSets/interfaces/set_names.v +++ b/FiniteSets/interfaces/set_names.v @@ -53,8 +53,8 @@ Infix "∩" := intersection (at level 8, right associativity). Notation "( ∩ )" := intersection (only parsing). Notation "( X ∩ )" := (intersection X) (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 "⊆" := subset (at level 10, right associativity). Infix "∈_d" := member_dec (at level 9, right associativity). -Infix "⊆_d" := subset_dec (at level 10, right associativity). \ No newline at end of file +Infix "⊆_d" := subset_dec (at level 10, right associativity). diff --git a/FiniteSets/kuratowski/kuratowski_sets.v b/FiniteSets/kuratowski/kuratowski_sets.v index 32aa2da..4ad7e6d 100644 --- a/FiniteSets/kuratowski/kuratowski_sets.v +++ b/FiniteSets/kuratowski/kuratowski_sets.v @@ -139,11 +139,17 @@ Section relations. - apply False_hp. - apply (fun a' => merely(a = a')). - apply lor. + (* TODO *) - eauto with lattice_hints typeclass_instances. + apply associativity. - eauto with lattice_hints typeclass_instances. + apply commutativity. - eauto with lattice_hints typeclass_instances. + apply left_identity. - eauto with lattice_hints typeclass_instances. + apply right_identity. - eauto with lattice_hints typeclass_instances. + intros. simpl. apply binary_idempotent. Defined. (** Subset relation of finite sets. *) @@ -153,15 +159,15 @@ Section relations. hrecursion X. - apply Unit_hp. - apply (fun a => a ∈ Y). - - intros X1 X2. - exists (prod X1 X2). - exact _. + - apply land. + (* TODO *) - eauto with lattice_hints typeclass_instances. + apply associativity. - eauto with lattice_hints typeclass_instances. - - intros. - apply path_trunctype ; apply prod_unit_l. - - intros. - apply path_trunctype ; apply prod_unit_r. + apply commutativity. + - apply left_identity. + - apply right_identity. - eauto with lattice_hints typeclass_instances. + intros. apply binary_idempotent. Defined. -End relations. \ No newline at end of file +End relations. diff --git a/FiniteSets/kuratowski/length.v b/FiniteSets/kuratowski/length.v index 7f5845d..892258c 100644 --- a/FiniteSets/kuratowski/length.v +++ b/FiniteSets/kuratowski/length.v @@ -6,7 +6,7 @@ Section length. Definition length : FSet A -> nat. simple refine (FSet_cons_rec _ _ _ _ _ _). - - apply 0. + - apply 0%nat. - intros a X n. apply (if a ∈_d X then n else (S n)). - intros X a n. @@ -295,7 +295,7 @@ End length_sum. Section length_zero_one. 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. refine (@equiv_path_nat (n.+1) 0)^-1. Defined. @@ -311,7 +311,7 @@ Section length_zero_one. apply ((equiv_path_nat)^-1 X). 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. simple refine (FSet_cons_ind (fun Z => _) _ _ _ _ _) ; try (intros ; apply path_ishprop) ; simpl. @@ -326,7 +326,7 @@ Section length_zero_one. * contradiction (Z_not_S _ HaX). 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. simple refine (FSet_cons_ind (fun Z => _) _ _ _ _ _) ; try (intros ; apply path_ishprop) ; simpl. @@ -351,4 +351,4 @@ Section length_zero_one. * rewrite (length_zero _ (remove_S _ _ HaX)). apply nr. Defined. -End length_zero_one. \ No newline at end of file +End length_zero_one. diff --git a/FiniteSets/kuratowski/operations.v b/FiniteSets/kuratowski/operations.v index dd726d8..b7f570c 100644 --- a/FiniteSets/kuratowski/operations.v +++ b/FiniteSets/kuratowski/operations.v @@ -19,25 +19,20 @@ Section operations. - apply (idem o f). Defined. - Global Instance fset_pure : hasPure FSet. - Proof. - split. - apply (fun _ a => {|a|}). - Defined. + Global Instance return_fset : Return FSet := fun _ a => {|a|}. - Global Instance fset_bind : hasBind FSet. + Global Instance bind_fset : Bind FSet. Proof. - split. - intros A. - hrecursion. + intros A B m f. + hrecursion m. - exact ∅. - - exact idmap. + - exact f. - apply (∪). - apply assoc. - apply comm. - apply nl. - apply nr. - - apply union_idem. + - intros; apply union_idem. Defined. (** Set-theoretical operations. *) @@ -110,8 +105,8 @@ Section operations. - intros a f. apply {|f (a; tr idpath)|}. - intros X1 X2 HX1 HX2 f. - 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 (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)))). specialize (HX1 fX1). specialize (HX2 fX2). apply (HX1 ∪ HX2). @@ -201,9 +196,9 @@ Section operations_decidable. Defined. 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. Section FSet_cons_rec. diff --git a/FiniteSets/kuratowski/properties.v b/FiniteSets/kuratowski/properties.v index 43d7a21..1474321 100644 --- a/FiniteSets/kuratowski/properties.v +++ b/FiniteSets/kuratowski/properties.v @@ -6,7 +6,7 @@ Require Import lattice_interface lattice_examples monad_interface. Section monad_fset. Context `{Funext}. - Global Instance fset_functorish : Functorish FSet. + Global Instance functorish_fset : Functorish FSet. Proof. simple refine (Build_Functorish _ _ _). - intros ? ? f. @@ -19,7 +19,7 @@ Section monad_fset. ; try (intros ; apply path_ishprop). Defined. - Global Instance fset_functor : Functor FSet. + Global Instance functor_fset : Functor FSet. Proof. split. intros. @@ -30,24 +30,18 @@ Section monad_fset. ; try (intros ; apply path_ishprop). Defined. - Global Instance fset_monad : Monad FSet. + Global Instance monad_fset : Monad FSet. Proof. split. - - intros. - apply path_forall. - intro X. - hrecursion X ; try (intros; f_ap) ; - try (intros; apply path_ishprop). - - intros. - apply path_forall. - intro X. - hrecursion X ; try (intros; f_ap) ; - try (intros; apply path_ishprop). - - intros. - apply path_forall. - intro X. - hrecursion X ; try (intros; f_ap) ; - try (intros; apply path_ishprop). + - intros. reflexivity. + - intros A X. + hrecursion X; + try (intros; compute[bind ret bind_fset return_fset]; simpl; f_ap); + try (intros; apply path_ishprop); try reflexivity. + - intros A B C X f g. + hrecursion X; + try (intros; compute[bind ret bind_fset return_fset]; simpl; f_ap); + try (intros; apply path_ishprop); try reflexivity. Defined. 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. 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). Proof. apply BuildIsSurjection. @@ -87,8 +81,8 @@ Section monad_fset. f_ap. Defined. - Lemma bind_isIn `{Univalence} {A : Type} (X : FSet (FSet A)) (a : A) : - (exists x, x ∈ X * a ∈ x) -> a ∈ (bind _ X). + Lemma mjoin_isIn `{Univalence} {A : Type} (X : FSet (FSet A)) (a : A) : + (exists x, prod (x ∈ X) (a ∈ x)) -> a ∈ (mjoin X). Proof. hinduction X; try (intros; apply path_forall; intro; apply path_ishprop). @@ -253,13 +247,13 @@ Section properties_membership. : a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath. Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A, - {|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}. + {| (X ∪ Y) | ϕ|} = {|X | ϕ|} ∪ {|Y | ϕ|}. Proof. reflexivity. Defined. 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. hinduction ; try (intros ; apply set_path2). - destruct (ϕ a) ; reflexivity. @@ -386,8 +380,8 @@ Section properties_membership. ** apply p. ** apply path_ishprop. - intros X1 X2 HX1 HX2 f b. - 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 (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)))). specialize (HX1 fX1 b). specialize (HX2 fX2 b). apply path_iff_hprop. @@ -461,8 +455,8 @@ Section properties_membership. apply tr. unfold IsEmbedding, hfiber in *. specialize (f_inj (f a)). - 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 ((a;idpath (f a)) : {x : A | f x = f a}) as p1. + pose ((b;Hfa^) : {x : A | f x = f a}) as p2. assert (p1 = p2) as Hp1p2. { apply f_inj. } apply (ap (fun x => x.1) Hp1p2). @@ -485,32 +479,19 @@ Ltac toHProp := Section fset_join_semilattice. Context {A : Type}. - Instance: bottom(FSet A). - Proof. - unfold bottom. - apply ∅. - Defined. + Global Instance bottom_fset : Bottom (FSet A) := ∅. - Instance: maximum(FSet A). - Proof. - intros x y. - apply (x ∪ y). - Defined. + Global Instance join_fset : Join (FSet A) := fun x y => x ∪ y. - Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A). + Global Instance boundedjoinsemilattice_fset : BoundedJoinSemiLattice (FSet A). Proof. - split. - - intros ? ?. - apply comm. - - intros ? ? ?. - apply (assoc _ _ _)^. - - intros ?. - apply union_idem. - - intros x. - apply nl. - - intros ?. - apply nr. - Defined. + repeat split; try apply _; cbv. + - apply assoc. + - apply nl. + - apply nr. + - apply comm. + - apply union_idem. + Defined. End fset_join_semilattice. (* Lemmas relating operations to the membership predicate *) @@ -569,10 +550,10 @@ Section properties_membership_decidable. Defined. 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. 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 (ϕ a) ; try reflexivity ; try contradiction ; try (contradiction (n t)) ; contradiction (t n). @@ -692,27 +673,14 @@ Section set_lattice. Context `{MerelyDecidablePaths A}. Context `{Univalence}. - Instance fset_max : maximum (FSet A). - Proof. - intros x y. - apply (x ∪ y). - Defined. + Global Instance meet_fset : Meet (FSet A) := intersection. - Instance fset_min : minimum (FSet A). + Global Instance lattice_fset : Lattice (FSet A). Proof. - intros x y. - apply (x ∩ y). - Defined. - - Instance fset_bot : bottom (FSet A). - Proof. - unfold bottom. - apply ∅. - Defined. - - Instance lattice_fset : Lattice (FSet A). - Proof. - split ; toBool. + repeat split; try apply _; + compute[sg_op meet_is_sg_op meet_fset]; + toBool. + apply associativity. Defined. End set_lattice. @@ -720,7 +688,7 @@ End set_lattice. Section dec_eq. Context {A : Type} `{DecidablePaths A} `{Univalence}. - Instance fsets_dec_eq : DecidablePaths (FSet A). + Global Instance dec_eq_fset : DecidablePaths (FSet A). Proof. intros X Y. apply (decidable_equiv' ((Y ⊆ X) * (X ⊆ Y)) (eq_subset X Y)^-1). @@ -732,29 +700,32 @@ End dec_eq. Section comprehension_properties. 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. toHProp. Defined. Lemma comprehension_subset : forall ϕ (X : FSet A), - {|X & ϕ|} ∪ X = X. + {|X | ϕ|} ∪ X = X. Proof. toHProp. destruct (ϕ a) ; eauto with lattice_hints typeclass_instances. + - apply binary_idempotent. + - apply left_identity. Defined. Lemma comprehension_or : forall ϕ ψ (X : FSet A), - {|X & (fun a => orb (ϕ a) (ψ a))|} - = {|X & ϕ|} ∪ {|X & ψ|}. + {|X | (fun a => orb (ϕ a) (ψ a))|} + = {|X | ϕ|} ∪ {|X | ψ|}. Proof. toHProp. 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. Lemma comprehension_all : forall (X : FSet A), - {|X & fun _ => true|} = X. + {|X | fun _ => true|} = X. Proof. toHProp. Defined. diff --git a/FiniteSets/misc/dec_fset.v b/FiniteSets/misc/dec_fset.v index 4d608d7..ce57d27 100644 --- a/FiniteSets/misc/dec_fset.v +++ b/FiniteSets/misc/dec_fset.v @@ -12,13 +12,17 @@ Section quantifiers. - apply P. - intros 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. + apply commutativity. - intros. apply path_trunctype ; apply prod_unit_l. - intros. apply path_trunctype ; apply prod_unit_r. - eauto with lattice_hints typeclass_instances. + intros. apply binary_idempotent. Defined. Lemma all_intro X : forall (HX : forall x, x ∈ X -> P x), all X. @@ -61,11 +65,17 @@ Section quantifiers. - apply False_hp. - apply P. - apply lor. + - eauto with lattice_hints typeclass_instances. + (* TODO eauto with .. *) + apply associativity. + - eauto with lattice_hints typeclass_instances. + apply commutativity. + - eauto with lattice_hints typeclass_instances. + apply left_identity. - eauto with lattice_hints typeclass_instances. + apply right_identity. - eauto with lattice_hints typeclass_instances. - - eauto with lattice_hints typeclass_instances. - - eauto with lattice_hints typeclass_instances. - - eauto with lattice_hints typeclass_instances. + intros. apply binary_idempotent. Defined. Lemma exist_intro X a : a ∈ X -> P a -> exist X. @@ -84,7 +94,7 @@ Section quantifiers. * apply (inr (HX2 t Pa)). 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. hinduction X ; try (intros ; apply path_ishprop). - contradiction. @@ -132,7 +142,7 @@ Section simple_example. Context `{Univalence}. 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. Proof. diff --git a/FiniteSets/misc/dec_kuratowski.v b/FiniteSets/misc/dec_kuratowski.v index a06e7bf..6d9ced2 100644 --- a/FiniteSets/misc/dec_kuratowski.v +++ b/FiniteSets/misc/dec_kuratowski.v @@ -66,12 +66,12 @@ Section length. Context {A : Type} `{Univalence}. Variable (length : FSet A -> nat) - (length_singleton : forall (a : A), length {|a|} = 1) - (length_one : forall (X : FSet A), length X = 1 -> hexists (fun a => X = {|a|})). + (length_singleton : forall (a : A), length {|a|} = 1%nat) + (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)). Proof. - destruct (dec (length ({|a|} ∪ {|b|}) = 1)). + destruct (dec (length ({|a|} ∪ {|b|}) = 1%nat)). - refine (inl _). pose (length_one _ p) as Hp. simple refine (Trunc_rec _ Hp). @@ -91,4 +91,4 @@ Section length. rewrite p, idem in n. apply (n (length_singleton b)). Defined. -End length. \ No newline at end of file +End length. diff --git a/FiniteSets/misc/ordered.v b/FiniteSets/misc/ordered.v index 7dde386..46ab730 100644 --- a/FiniteSets/misc/ordered.v +++ b/FiniteSets/misc/ordered.v @@ -1,84 +1,54 @@ (** If [A] has a total order, then we can pick the minimum of finite sets. *) Require Import HoTT HitTactics. +Require Import HoTT.Classes.interfaces.orders. 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. 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. Proof. - destruct (@total _ R _ x y). + destruct (total _ x y). - apply x. - - destruct s as [s | s]. - * apply x. - * apply y. + - apply y. Defined. - Lemma min_spec1 x y : R (min x y) x. + Lemma min_spec1 x y : (min x y) ≤ x. Proof. unfold min. - destruct (total x y) ; simpl. + destruct (total _ x y) ; simpl. - reflexivity. - - destruct s as [ | t]. - * reflexivity. - * apply t. + - assumption. 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. intros. unfold min. - destruct (total x y) as [ | s]. - * assumption. - * try (destruct s) ; assumption. + destruct (total _ x y); simpl; assumption. Defined. Lemma min_comm x y : min x y = min y x. Proof. unfold min. - destruct (total x y) ; destruct (total y x) ; simpl. - - assumption. - - destruct s as [s | s] ; auto. - - 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. + destruct (total _ x y) ; destruct (total _ y x) ; simpl; try reflexivity. + + apply antisymmetry with (≤); [apply _|assumption..]. + + apply antisymmetry with (≤); [apply _|assumption..]. Defined. Lemma min_idem x : min x x = x. Proof. unfold min. - destruct (total x x) ; simpl. - - reflexivity. - - destruct s ; reflexivity. + destruct (total _ x x) ; simpl; reflexivity. Defined. Lemma min_assoc x y z : min (min x y) z = min x (min y z). Proof. - apply (@antisymmetry _ R _ _). + apply antisymmetry with (≤). apply _. - apply min_spec2. * etransitivity ; apply min_spec1. * apply min_spec2. @@ -96,43 +66,36 @@ Section minimum. ; rewrite min_comm ; apply min_spec1. Defined. - Variable (top : A). - Context `{IsTop A R top}. - - Lemma min_nr x : min x top = x. + Lemma min_nr `{IsTop} x : min x ⊤ = x. Proof. intros. unfold min. - destruct (total x top). + destruct (total _ x ⊤). - reflexivity. - - destruct s. - * reflexivity. - * apply (@antisymmetry _ R _ _). - ** assumption. - ** refine (top_max _). apply _. + - apply antisymmetry with (≤). apply _. + + assumption. + + apply is_top. Defined. - Lemma min_nl x : min top x = x. + Lemma min_nl `{IsTop} x : min ⊤ x = x. Proof. rewrite min_comm. apply min_nr. 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. unfold min. - destruct (total x y). + destruct (total _ x y) as [?|s]. - apply idmap. - - destruct s as [s | s]. - * apply idmap. - * intros X. - rewrite X in s. - apply (@antisymmetry _ R _ _). - ** apply top_max. - ** assumption. + - intros X. + rewrite X in s. + apply antisymmetry with (≤). apply _. + * apply is_top. + * assumption. 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. rewrite min_comm. apply min_top_l. @@ -144,72 +107,60 @@ Section add_top. Variable (A : Type). Context `{TotalOrder A}. - Definition Top := A + Unit. - Definition top : Top := inr tt. + Definition Topped := (A + Unit)%type. + Global Instance top_topped : Top Topped := inr tt. - Global Instance RTop : LessThan Top. + Global Instance RTop : Le Topped. Proof. - unfold relation. - induction 1 as [a1 | ] ; induction 1 as [a2 | ]. - - apply (R a1 a2). + intros [a1 | ?] [a2 | ?]. + - apply (a1 ≤ a2). - apply Unit_hp. - apply False_hp. - apply Unit_hp. Defined. - Global Instance rtop_hprop : - is_mere_relation A R -> is_mere_relation Top RTop. + Instance rtop_hprop : + is_mere_relation A R -> is_mere_relation Topped RTop. Proof. - intros P a b. - destruct a ; destruct b ; apply _. + intros ? P a b. + destruct a ; destruct b ; simpl; apply _. Defined. - Global Instance RTopOrder : TotalOrder Top. + Global Instance RTopOrder : TotalOrder RTop. Proof. - split. - - intros x ; induction x ; unfold RTop ; simpl. + repeat split; try apply _. + - intros [?|?] ; cbv. * reflexivity. * apply tt. - - intros x y ; induction x as [a1 | ] ; induction y as [a2 | ] ; unfold RTop ; simpl - ; try contradiction. - * intros ; f_ap. - apply (@antisymmetry _ R _ _) ; assumption. - * intros ; induction b ; induction b0. - reflexivity. - - intros x y z ; induction x as [a1 | b1] ; induction y as [a2 | b2] - ; induction z as [a3 | b3] ; unfold RTop ; simpl - ; try contradiction ; intros ; try (apply tt). - transitivity a2 ; assumption. - - intros x y. - 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. + - intros [a1 | []] [a2 | []] [a3 | []]; cbv + ; try contradiction; try (by intros; apply tt). + apply transitivity. + - intros [a1 |[]] [a2 |[]]; cbv; try contradiction. + + intros. f_ap. apply antisymmetry with Ale; [apply _|assumption..]. + + intros; reflexivity. + - intros [a1|[]] [a2|[]]; cbv. + * apply total. apply _. + * apply (inl tt). + * apply (inr tt). + * apply (inr tt). Defined. - Global Instance top_a_top : IsTop Top RTop top. - Proof. - intro x ; destruct x ; apply tt. - Defined. -End add_top. + Global Instance is_top_topped : IsTop RTop. + Proof. intros [?|?]; cbv; apply tt. Defined. +End add_top. (** If [A] has a total order, then a nonempty finite set has a minimum element. *) Section min_set. Variable (A : Type). Context `{TotalOrder A}. - Context `{is_mere_relation A R}. Context `{Univalence} `{IsHSet A}. - Definition min_set : FSet A -> Top A. + Definition min_set : FSet A -> Topped A. Proof. hrecursion. - - apply (top A). + - apply ⊤. - apply inl. - - apply min. + - apply min with (RTop A). apply _. - intros ; symmetry ; apply min_assoc. - apply min_comm. - apply min_nl. apply _. @@ -217,7 +168,7 @@ Section min_set. - intros ; apply min_idem. 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. simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _) ; 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) _ _). * apply tt. * rewrite X ; apply tt. - - intros. - assert (min_set x = top A). - { - simple refine (min_top_l _ _ (min_set y) _) ; assumption. - } + - intros x y ? ? ?. + assert (min_set x = ⊤). + { simple refine (min_top_l _ _ (min_set y) _) ; assumption. } rewrite (X X2). rewrite nl. - assert (min_set y = top A). + assert (min_set y = ⊤). { simple refine (min_top_r _ (min_set x) _ _) ; assumption. } rewrite (X0 X3). reflexivity. Defined. 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. simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _) ; try (intros ; apply path_ishprop) @@ -259,16 +208,16 @@ Section min_set. unfold member in X, X0. destruct X1. * 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. } etransitivity. { apply X1. } assumption. * 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. } etransitivity. { apply X1. } assumption. Defined. -End min_set. \ No newline at end of file +End min_set. diff --git a/FiniteSets/subobjects/b_finite.v b/FiniteSets/subobjects/b_finite.v index ca7eda3..bc4e496 100644 --- a/FiniteSets/subobjects/b_finite.v +++ b/FiniteSets/subobjects/b_finite.v @@ -8,9 +8,9 @@ Section finite_hott. Context `{Univalence}. (* 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. exists (a; tr idpath). intros [b p]. @@ -20,7 +20,7 @@ Section finite_hott. apply p^. Defined. - Definition singleton_fin_equiv' a : Fin 1 -> {b : A & b ∈ {|a|}}. + Definition singleton_fin_equiv' a : Fin 1 -> {b : A | b ∈ {|a|}}. Proof. intros _. apply (a; tr idpath). Defined. @@ -66,7 +66,8 @@ Section finite_hott. : DecidablePaths A. Proof. 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. unfold Sect in *. destruct pn as [f [g fg gf _]], n as [ | [ | n]]. @@ -75,9 +76,9 @@ Section finite_hott. apply (tr(inl(tr idpath))). - refine (inl _). 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))) - : {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))). assert (fs_eq : f s1 = f s2). { by apply path_ishprop. } @@ -116,9 +117,9 @@ Section singleton_set. Variable (A : Type). 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. Proof. @@ -132,7 +133,7 @@ Section singleton_set. rewrite <- ap_compose. enough(forall r, (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 ) as H2. { @@ -166,7 +167,7 @@ End singleton_set. Section empty. Variable (A : Type). Variable (X : A -> hProp) - (Xequiv : {a : A & a ∈ X} <~> Fin 0). + (Xequiv : {a : A | a ∈ X} <~> Fin 0). Context `{Univalence}. Lemma X_empty : X = ∅. @@ -185,10 +186,10 @@ Section split. Variable (A : Type). Variable (P : A -> hProp) (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, - ({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. pose (fun x : A => sig (fun y : Fin n => x = (f^-1 (inl y)).1)) as P'. assert (forall x, IsHProp (P' x)). @@ -267,7 +268,7 @@ Arguments Bfin {_} _. Arguments split {_} {_} _ _ _. Section Bfin_no_singletons. - Definition S1toSig (x : S1) : {x : S1 & merely(x = base)}. + Definition S1toSig (x : S1) : {x : S1 | merely(x = base)}. Proof. exists 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) : ~ (b ∈ Y) -> - {a : A & a ∈ ({|b|} ∪ Y)} <~> {a : A & a ∈ Y} + Unit. + {a : A | a ∈ ({|b|} ∪ Y)} <~> {a : A | a ∈ Y} + Unit. Proof. intros HYb. unshelve eapply BuildEquiv. @@ -422,7 +423,7 @@ Section kfin_bfin. strip_truncations. revert fX. revert X. induction n; intros X fX. - - rewrite (X_empty _ X fX), (neutralL_max (Sub A)). + - rewrite (X_empty _ X fX). rewrite left_identity. apply HY. - destruct (split X n fX) as (X' & b & HX' & HX). @@ -431,10 +432,10 @@ Section kfin_bfin. + cut (X ∪ Y = X' ∪ Y). { intros HXY. rewrite HXY. assumption. } apply path_forall. intro a. - unfold union, sub_union, max_fun. + unfold union, sub_union, join, join_fun. rewrite HX. rewrite (commutativity (X' a)). - rewrite (associativity _ (X' a)). + rewrite <- (associativity _ (X' a)). apply path_iff_hprop. * intros Ha. strip_truncations. @@ -448,15 +449,14 @@ Section kfin_bfin. strip_truncations. exists (n'.+1). 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. 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_. ** apply path_forall. intros ?. - unfold union, sub_union, max_fun. apply HX. - ** rewrite HX_, <- (commutative_max (Sub A) X'). + ** rewrite HX_, <- (commutativity X'). reflexivity. * etransitivity. { 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). Proof. hinduction; try (intros; apply path_ishprop). - - exists 0. + - exists 0%nat. apply tr. simple refine (BuildEquiv _ _ _ _). destruct 1 as [? []]. @@ -490,4 +490,4 @@ Proof. apply path_ishprop. * refine (fun a => (a;HY a);fun _ => _). reflexivity. -Defined. \ No newline at end of file +Defined. diff --git a/FiniteSets/subobjects/enumerated.v b/FiniteSets/subobjects/enumerated.v index 6cf45bd..cf832d6 100644 --- a/FiniteSets/subobjects/enumerated.v +++ b/FiniteSets/subobjects/enumerated.v @@ -300,7 +300,7 @@ Section subobjects. simpl. apply path_ishprop. 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. destruct ls as [|[x Hx] ls]. - exact nil. @@ -323,7 +323,7 @@ Section subobjects. Defined. 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. destruct ls as [|[y Hy] ls]. - apply weaken_list_r. exact ls'. @@ -356,7 +356,7 @@ Section subobjects. Defined. Lemma enumeratedS_union (P Q : Sub A) : - enumeratedS P -> enumeratedS Q -> enumeratedS (max_L P Q). + enumeratedS P -> enumeratedS Q -> enumeratedS (P ⊔ Q). Proof. intros HP HQ. strip_truncations; apply tr. @@ -388,7 +388,7 @@ Section subobjects. Transparent enumeratedS. 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. apply hprop_allpath. intros [X HX] [Y HY]. assert (X = Y) as HXY. @@ -433,7 +433,7 @@ Section subobjects. Definition enumeratedS_to_FSet : forall (P : Sub A), enumeratedS P -> - {X : FSet A & k_finite.map X = P}. + {X : FSet A | k_finite.map X = P}. Proof. intros P HP. strip_truncations. diff --git a/FiniteSets/subobjects/k_finite.v b/FiniteSets/subobjects/k_finite.v index ba36ffc..e501b57 100644 --- a/FiniteSets/subobjects/k_finite.v +++ b/FiniteSets/subobjects/k_finite.v @@ -33,7 +33,7 @@ Section k_finite. 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. apply hprop_allpath. intros [X PX] [Y PY]. @@ -77,19 +77,16 @@ Section structure_k_finite. Context (A : Type). 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. - intros. - unfold map, max_fun. - reflexivity. + intros. + reflexivity. Defined. Lemma k_finite_union : closedUnion (Kf_sub A). Proof. unfold closedUnion, Kf_sub, Kf_sub_intern. - intros. - destruct X0 as [SX XP]. - destruct X1 as [SY YP]. + intros X Y [SX XP] [SY YP]. exists (SX ∪ SY). rewrite map_union. rewrite XP, YP. @@ -179,9 +176,9 @@ Section k_properties. | inl a => ({|a|} : FSet A) | inr b => ∅ end). - exists (bind _ (fset_fmap f X)). + exists (mjoin (fset_fmap f X)). intro a. - apply bind_isIn. + apply mjoin_isIn. specialize (HX (inl a)). exists {|a|}. split; [ | apply tr; reflexivity ]. apply (fmap_isIn f (inl a) X). @@ -259,33 +256,35 @@ Section alternative_definition. Local Ltac help_solve := repeat (let x := fresh in intro x ; destruct x) ; intros ; 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 ; intro z ; 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. - 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 ∅. - auto. + simpl. auto. - intros a. exists {|a|}. - auto. + simpl. auto. - intros [P1 HP1] [P2 HP2]. exists (P1 ∪ P2). intros ? ? ? HP. apply HP. * apply HP1 ; assumption. * apply HP2 ; assumption. - - help_solve. - - help_solve. - - help_solve. - - help_solve. - - help_solve. + - help_solve. (* TODO: eauto *) apply associativity. + - help_solve. apply commutativity. + - help_solve. apply left_identity. + - help_solve. apply right_identity. + - help_solve. apply binary_idempotent. 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. intros [P HP]. 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^)^). 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. simple refine (path_sigma _ _ _ _ _) ; try (apply path_ishprop). apply path_forall. diff --git a/FiniteSets/subobjects/sub.v b/FiniteSets/subobjects/sub.v index 74c959f..6bd7920 100644 --- a/FiniteSets/subobjects/sub.v +++ b/FiniteSets/subobjects/sub.v @@ -6,9 +6,9 @@ Section subobjects. Definition Sub := A -> hProp. - Global Instance sub_empty : hasEmpty Sub := fun _ => False_hp. - Global Instance sub_union : hasUnion Sub := max_fun. - Global Instance sub_intersection : hasIntersection Sub := min_fun. + Global Instance sub_empty : hasEmpty Sub := fun _ => ⊥. + Global Instance sub_union : hasUnion Sub := join. + Global Instance sub_intersection : hasIntersection Sub := meet. Global Instance sub_singleton : hasSingleton Sub A := fun a b => BuildhProp (Trunc (-1) (b = a)). Global Instance sub_membership : hasMembership Sub A := fun a X => X a.