1
0
mirror of https://github.com/nmvdw/HITs-Examples synced 2025-11-03 15:13:51 +01:00

A negligible change in the structure

This commit is contained in:
Niels
2017-09-07 15:19:48 +02:00
parent 4ab70ae1fe
commit 474c9324ca
29 changed files with 2082 additions and 1459 deletions

62
FiniteSets/misc/T.v Normal file
View File

@@ -0,0 +1,62 @@
(** Merely decidable equality implies LEM *)
Require Import HoTT HitTactics prelude.
Section TR.
Context `{Univalence}.
Variable A : hProp.
Definition T := Unit + Unit.
Definition R (x y : T) : hProp :=
match x, y with
| inl tt, inl tt => Unit_hp
| inl tt, inr tt => A
| inr tt, inl tt => A
| inr tt, inr tt => Unit_hp
end.
Global Instance R_refl : Reflexive R.
Proof.
intro x ; destruct x as [[ ] | [ ]] ; apply tt.
Defined.
Global Instance R_sym : Symmetric R.
Proof.
repeat (let x := fresh in intro x ; destruct x as [[ ] | [ ]])
; auto ; apply tt.
Defined.
Global Instance R_trans : Transitive R.
Proof.
repeat (let x := fresh in intro x ; destruct x as [[ ] | [ ]]) ; intros
; auto ; apply tt.
Defined.
Definition TR : Type := quotient R.
Definition TR_zero : TR := class_of R (inl tt).
Definition TR_one : TR := class_of R (inr tt).
Definition equiv_pathspace_T : (TR_zero = TR_one) = (R (inl tt) (inr tt))
:= path_universe (sets_exact R (inl tt) (inr tt)).
Global Instance quotientB_recursion : HitRecursion TR :=
{
indTy := _;
recTy :=
forall (P : Type) (HP: IsHSet P) (u : T -> P),
(forall x y : T, R x y -> u x = u y) -> TR -> P;
H_inductor := quotient_ind R ;
H_recursor := @quotient_rec _ R _
}.
End TR.
Theorem merely_dec `{Univalence} : (forall (A : Type) (a b : A), hor (a = b) (~a = b))
->
forall (A : hProp), A + (~A).
Proof.
intros X A.
specialize (X (TR A) (TR_zero A) (TR_one A)).
rewrite equiv_pathspace_T in X.
strip_truncations.
apply X.
Defined.

288
FiniteSets/misc/bad.v Normal file
View File

@@ -0,0 +1,288 @@
(* This is a /bad/ definition of FSets, without the 0-truncation.
Here we show that the resulting type is not an h-set. *)
Require Import HoTT HitTactics.
Require Import set_names.
Module Export FSet.
Section FSet.
Private Inductive FSet (A : Type): Type :=
| E : FSet A
| L : A -> FSet A
| U : FSet A -> FSet A -> FSet A.
Global Instance fset_empty : forall A, hasEmpty (FSet A) := E.
Global Instance fset_singleton : forall A, hasSingleton (FSet A) A := L.
Global Instance fset_union : forall A, hasUnion (FSet A) := U.
Variable A : Type.
Axiom assoc : forall (x y z : FSet A),
x (y z) = (x y) z.
Axiom comm : forall (x y : FSet A),
x y = y x.
Axiom nl : forall (x : FSet A),
x = x.
Axiom nr : forall (x : FSet A),
x = x.
Axiom idem : forall (x : A),
{|x|} {|x|} = {|x|}.
End FSet.
Arguments assoc {_} _ _ _.
Arguments comm {_} _ _.
Arguments nl {_} _.
Arguments nr {_} _.
Arguments idem {_} _.
Section FSet_induction.
Variable A: Type.
Variable (P : FSet A -> Type).
Variable (eP : P ).
Variable (lP : forall a: A, P {|a |}).
Variable (uP : forall (x y: FSet A), P x -> P y -> P (x y)).
Variable (assocP : forall (x y z : FSet A)
(px: P x) (py: P y) (pz: P z),
assoc x y z #
(uP x (y z) px (uP y z py pz))
=
(uP (x y) z (uP x y px py) pz)).
Variable (commP : forall (x y: FSet A) (px: P x) (py: P y),
comm x y #
uP x y px py = uP y x py px).
Variable (nlP : forall (x : FSet A) (px: P x),
nl x # uP x eP px = px).
Variable (nrP : forall (x : FSet A) (px: P x),
nr x # uP x px eP = px).
Variable (idemP : forall (x : A),
idem x # uP {|x|} {|x|} (lP x) (lP x) = lP x).
(* Induction principle *)
Fixpoint FSet_ind
(x : FSet A)
{struct x}
: P x
:= (match x return _ -> _ -> _ -> _ -> _ -> P x with
| E => fun _ _ _ _ _ => eP
| L a => fun _ _ _ _ _ => lP a
| U y z => fun _ _ _ _ _ => uP y z (FSet_ind y) (FSet_ind z)
end) assocP commP nlP nrP idemP.
Axiom FSet_ind_beta_assoc : forall (x y z : FSet A),
apD FSet_ind (assoc x y z) =
(assocP x y z (FSet_ind x) (FSet_ind y) (FSet_ind z)).
Axiom FSet_ind_beta_comm : forall (x y : FSet A),
apD FSet_ind (comm x y) = commP x y (FSet_ind x) (FSet_ind y).
Axiom FSet_ind_beta_nl : forall (x : FSet A),
apD FSet_ind (nl x) = nlP x (FSet_ind x).
Axiom FSet_ind_beta_nr : forall (x : FSet A),
apD FSet_ind (nr x) = nrP x (FSet_ind x).
Axiom FSet_ind_beta_idem : forall (x : A), apD FSet_ind (idem x) = idemP x.
End FSet_induction.
Section FSet_recursion.
Variable A : Type.
Variable P : Type.
Variable e : P.
Variable l : A -> P.
Variable u : P -> P -> P.
Variable assocP : forall (x y z : P), u x (u y z) = u (u x y) z.
Variable commP : forall (x y : P), u x y = u y x.
Variable nlP : forall (x : P), u e x = x.
Variable nrP : forall (x : P), u x e = x.
Variable idemP : forall (x : A), u (l x) (l x) = l x.
Definition FSet_rec : FSet A -> P.
Proof.
simple refine (FSet_ind A _ _ _ _ _ _ _ _ _) ; try (intros ; simple refine ((transport_const _ _) @ _)) ; cbn.
- apply e.
- apply l.
- intros x y ; apply u.
- apply assocP.
- apply commP.
- apply nlP.
- apply nrP.
- apply idemP.
Defined.
Definition FSet_rec_beta_assoc : forall (x y z : FSet A),
ap FSet_rec (assoc x y z)
=
assocP (FSet_rec x) (FSet_rec y) (FSet_rec z).
Proof.
intros.
unfold FSet_rec.
eapply (cancelL (transport_const (assoc x y z) _)).
simple refine ((apD_const _ _)^ @ _).
apply FSet_ind_beta_assoc.
Defined.
Definition FSet_rec_beta_comm : forall (x y : FSet A),
ap FSet_rec (comm x y)
=
commP (FSet_rec x) (FSet_rec y).
Proof.
intros.
unfold FSet_rec.
eapply (cancelL (transport_const (comm x y) _)).
simple refine ((apD_const _ _)^ @ _).
apply FSet_ind_beta_comm.
Defined.
Definition FSet_rec_beta_nl : forall (x : FSet A),
ap FSet_rec (nl x)
=
nlP (FSet_rec x).
Proof.
intros.
unfold FSet_rec.
eapply (cancelL (transport_const (nl x) _)).
simple refine ((apD_const _ _)^ @ _).
apply FSet_ind_beta_nl.
Defined.
Definition FSet_rec_beta_nr : forall (x : FSet A),
ap FSet_rec (nr x)
=
nrP (FSet_rec x).
Proof.
intros.
unfold FSet_rec.
eapply (cancelL (transport_const (nr x) _)).
simple refine ((apD_const _ _)^ @ _).
apply FSet_ind_beta_nr.
Defined.
Definition FSet_rec_beta_idem : forall (a : A),
ap FSet_rec (idem a)
=
idemP a.
Proof.
intros.
unfold FSet_rec.
eapply (cancelL (transport_const (idem a) _)).
simple refine ((apD_const _ _)^ @ _).
apply FSet_ind_beta_idem.
Defined.
End FSet_recursion.
Instance FSet_recursion A : HitRecursion (FSet A) :=
{
indTy := _; recTy := _;
H_inductor := FSet_ind A; H_recursor := FSet_rec A
}.
End FSet.
Section set_sphere.
From HoTT.HIT Require Import Circle.
Context `{Univalence}.
Instance S1_recursion : HitRecursion S1 :=
{
indTy := _; recTy := _;
H_inductor := S1_ind; H_recursor := S1_rec
}.
Variable A : Type.
Definition f (x : S1) : x = x.
Proof.
hrecursion x.
- exact loop.
- refine (transport_paths_FlFr _ _ @ _).
hott_simpl.
Defined.
Definition S1op (x y : S1) : S1.
Proof.
hrecursion y.
- exact x. (* x + base = x *)
- apply f.
Defined.
Lemma S1op_nr (x : S1) : S1op x base = x.
Proof. reflexivity. Defined.
Lemma S1op_nl (x : S1) : S1op base x = x.
Proof.
hrecursion x.
- exact loop.
- refine (transport_paths_FlFr loop _ @ _).
hott_simpl.
apply moveR_pM. apply moveR_pM. hott_simpl.
refine (ap_V _ _ @ _).
f_ap. apply S1_rec_beta_loop.
Defined.
Lemma S1op_assoc (x y z : S1) : S1op x (S1op y z) = S1op (S1op x y) z.
Proof.
hrecursion z.
- reflexivity.
- refine (transport_paths_FlFr loop _ @ _).
hott_simpl.
apply moveR_Mp. hott_simpl.
rewrite S1_rec_beta_loop.
rewrite ap_compose.
rewrite S1_rec_beta_loop.
hrecursion y.
+ symmetry. apply S1_rec_beta_loop.
+ apply is1type_S1.
Qed.
Lemma S1op_comm (x y : S1) : S1op x y = S1op y x.
Proof.
hrecursion x.
- apply S1op_nl.
- hrecursion y.
+ rewrite transport_paths_FlFr. hott_simpl.
rewrite S1_rec_beta_loop. reflexivity.
+ apply is1type_S1.
Defined.
Definition FSet_to_S : FSet A -> S1.
Proof.
hrecursion.
- exact base.
- intro a. exact base.
- exact S1op.
- apply S1op_assoc.
- apply S1op_comm.
- apply S1op_nl.
- apply S1op_nr.
- simpl. reflexivity.
Defined.
Lemma FSet_S_ap : (nl (E A)) = (nr ) -> idpath = loop.
Proof.
intros H1.
enough (ap FSet_to_S (nl ) = ap FSet_to_S (nr )) as H'.
- rewrite FSet_rec_beta_nl in H'.
rewrite FSet_rec_beta_nr in H'.
simpl in H'. unfold S1op_nr in H'.
exact H'^.
- f_ap.
Defined.
Lemma FSet_not_hset : IsHSet (FSet A) -> False.
Proof.
intros H1.
enough (idpath = loop).
- assert (S1_encode _ idpath = S1_encode _ (loopexp loop (pos Int.one))) as H' by f_ap.
rewrite S1_encode_loopexp in H'. simpl in H'. symmetry in H'.
apply (pos_neq_zero H').
- apply FSet_S_ap.
apply set_path2.
Qed.
End set_sphere.

274
FiniteSets/misc/ordered.v Normal file
View File

@@ -0,0 +1,274 @@
(** If [A] has a total order, then we can pick the minimum of finite sets. *)
Require Import HoTT HitTactics.
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}.
Definition min (x y : A) : A.
Proof.
destruct (@total _ R _ x y).
- apply x.
- destruct s as [s | s].
* apply x.
* apply y.
Defined.
Lemma min_spec1 x y : R (min x y) x.
Proof.
unfold min.
destruct (total x y) ; simpl.
- reflexivity.
- destruct s as [ | t].
* reflexivity.
* apply t.
Defined.
Lemma min_spec2 x y z : R z x -> R z y -> R z (min x y).
Proof.
intros.
unfold min.
destruct (total x y) as [ | s].
* assumption.
* try (destruct s) ; 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.
Defined.
Lemma min_idem x : min x x = x.
Proof.
unfold min.
destruct (total x x) ; simpl.
- reflexivity.
- destruct s ; reflexivity.
Defined.
Lemma min_assoc x y z : min (min x y) z = min x (min y z).
Proof.
apply (@antisymmetry _ R _ _).
- apply min_spec2.
* etransitivity ; apply min_spec1.
* apply min_spec2.
** etransitivity ; try (apply min_spec1).
simpl.
rewrite min_comm ; apply min_spec1.
** rewrite min_comm ; apply min_spec1.
- apply min_spec2.
* apply min_spec2.
** apply min_spec1.
** etransitivity.
{ rewrite min_comm ; apply min_spec1. }
apply min_spec1.
* transitivity (min y z); simpl
; rewrite min_comm ; apply min_spec1.
Defined.
Variable (top : A).
Context `{IsTop A R top}.
Lemma min_nr x : min x top = x.
Proof.
intros.
unfold min.
destruct (total x top).
- reflexivity.
- destruct s.
* reflexivity.
* apply (@antisymmetry _ R _ _).
** assumption.
** refine (top_max _). apply _.
Defined.
Lemma min_nl x : min top x = x.
Proof.
rewrite min_comm.
apply min_nr.
Defined.
Lemma min_top_l x y : min x y = top -> x = top.
Proof.
unfold min.
destruct (total x y).
- apply idmap.
- destruct s as [s | s].
* apply idmap.
* intros X.
rewrite X in s.
apply (@antisymmetry _ R _ _).
** apply top_max.
** assumption.
Defined.
Lemma min_top_r x y : min x y = top -> y = top.
Proof.
rewrite min_comm.
apply min_top_l.
Defined.
End minimum.
Section add_top.
Variable (A : Type).
Context `{TotalOrder A}.
Definition Top := A + Unit.
Definition top : Top := inr tt.
Global Instance RTop : LessThan Top.
Proof.
unfold relation.
induction 1 as [a1 | ] ; induction 1 as [a2 | ].
- apply (R 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.
Proof.
intros P a b.
destruct a ; destruct b ; apply _.
Defined.
Global Instance RTopOrder : TotalOrder Top.
Proof.
split.
- intros x ; induction x ; unfold RTop ; simpl.
* 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.
Defined.
Global Instance top_a_top : IsTop Top RTop top.
Proof.
intro x ; destruct x ; 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.
Proof.
hrecursion.
- apply (top A).
- apply inl.
- apply min.
- intros ; symmetry ; apply min_assoc.
- apply min_comm.
- apply min_nl. apply _.
- apply min_nr. apply _.
- intros ; apply min_idem.
Defined.
Definition empty_min : forall (X : FSet A), min_set X = top A -> X = .
Proof.
simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _)
; try (intros ; apply path_forall ; intro q ; apply set_path2)
; simpl.
- intros ; reflexivity.
- intros.
unfold top in X.
enough Empty.
{ contradiction. }
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.
}
rewrite (X X2).
rewrite nl.
assert (min_set y = top A).
{ 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).
Proof.
simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _)
; try (intros ; apply path_ishprop)
; simpl.
- contradiction.
- intros.
strip_truncations.
rewrite X.
reflexivity.
- intros.
strip_truncations.
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.
{ 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.
{ rewrite min_comm ; apply min_spec1. }
etransitivity.
{ apply X1. }
assumption.
Defined.
End min_set.

View File

@@ -0,0 +1,144 @@
Require Import HoTT HitTactics.
Require Import subobjects.k_finite subobjects.b_finite FSets.
Require Import misc.T.
Class IsProjective (X : Type) :=
projective : forall {P Q : Type} (p : P -> Q) (f : X -> Q),
IsSurjection p -> hexists (fun (g : X -> P) => p o g = f).
Instance IsProjective_IsHProp `{Univalence} X : IsHProp (IsProjective X).
Proof. unfold IsProjective. apply _. Defined.
Instance Unit_Projective `{Univalence} : IsProjective Unit.
Proof.
intros P Q p f Hsurj.
pose (x' := center (merely (hfiber p (f tt)))).
simple refine (@Trunc_rec (-1) (hfiber p (f tt)) _ _ _ x'). clear x'; intro x.
simple refine (tr (fun _ => x.1;_)). simpl.
apply path_forall; intros [].
apply x.2.
Defined.
Instance Empty_Projective `{Univalence} : IsProjective Empty.
Proof.
intros P Q p f Hsurj.
apply tr. exists Empty_rec.
apply path_forall. intros [].
Defined.
Instance Sum_Projective `{Univalence} {A B: Type} `{IsProjective A} `{IsProjective B} :
IsProjective (A + B).
Proof.
intros P Q p f Hsurj.
pose (f1 := fun a => f (inl a)).
pose (f2 := fun b => f (inr b)).
pose (g1' := projective p f1 Hsurj).
pose (g2' := projective p f2 Hsurj).
simple refine (Trunc_rec _ g1') ; intros [g1 pg1].
simple refine (Trunc_rec _ g2') ; intros [g2 pg2].
simple refine (tr (_;_)).
- intros [a | b].
+ apply (g1 a).
+ apply (g2 b).
- apply path_forall; intros [a | b]; simpl.
+ apply (ap (fun h => h a) pg1).
+ apply (ap (fun h => h b) pg2).
Defined.
(* All Bishop-finite sets are projective *)
Section b_fin_projective.
Context `{Univalence}.
Global Instance bishop_projective (X : Type) (Hfin : Finite X) : IsProjective X.
Proof.
simple refine (finite_ind_hprop (fun X _ => IsProjective X) _ _ X);
simpl; apply _.
Defined.
End b_fin_projective.
Section k_fin_lemoo_projective.
Context `{Univalence}.
Context {LEMoo : forall (P : Type), Decidable P}.
Global Instance kuratowski_projective_oo (X : Type) (Hfin : Kf X) : IsProjective X.
Proof.
assert (Finite X).
{ eapply Kf_to_Bf; auto.
intros pp qq. apply LEMoo. }
apply _.
Defined.
End k_fin_lemoo_projective.
Section k_fin_lem_projective.
Context `{Univalence}.
Context {LEM : forall (P : Type) {Hprop : IsHProp P}, Decidable P}.
Variable (X : Type).
Context `{IsHSet X}.
Global Instance kuratowski_projective (Hfin : Kf X) : IsProjective X.
Proof.
assert (Finite X).
{ eapply Kf_to_Bf; auto.
intros pp qq. apply LEM. apply _. }
apply _.
Defined.
End k_fin_lem_projective.
Section k_fin_projective_lem.
Context `{Univalence}.
Variable (P : Type).
Context `{IsHProp P}.
Definition X : Type := TR (BuildhProp P).
Instance X_set : IsHSet X.
Proof.
apply _.
Defined.
Definition X_fin : Kf X.
Proof.
apply Kf_unfold.
exists ({|TR_zero _|} {|TR_one _|}).
hinduction.
- destruct x as [ [ ] | [ ] ].
* apply (tr (inl (tr idpath))).
* apply (tr (inr (tr idpath))).
- intros.
apply path_ishprop.
Defined.
Definition p (a : Unit + Unit) : X :=
match a with
| inl _ => TR_zero _
| inr _ => TR_one _
end.
Instance p_surj : IsSurjection p.
Proof.
apply BuildIsSurjection.
hinduction.
- destruct x as [[ ] | [ ]].
* apply tr. exists (inl tt). reflexivity.
* apply tr. exists (inr tt). reflexivity.
- intros.
apply path_ishprop.
Defined.
Lemma LEM `{IsProjective X} : P + ~P.
Proof.
pose (k := projective p idmap _).
unfold hexists in k.
simple refine (Trunc_rec _ k); clear k; intros [g Hg].
destruct (dec (g (TR_zero _) = g (TR_one _))) as [Hℵ | Hℵ].
- left.
assert (TR_zero (BuildhProp P) = TR_one _) as Hbc.
{ pose (ap p Hℵ) as Hα.
rewrite (ap (fun h => h (TR_zero _)) Hg) in Hα.
rewrite (ap (fun h => h (TR_one _)) Hg) in Hα.
assumption. }
refine (classes_eq_related _ _ _ Hbc).
- right. intros HP.
apply Hℵ.
refine (ap g (related_classes_eq _ _)).
apply HP.
Defined.
End k_fin_projective_lem.