mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Strengthened mere choice
This commit is contained in:
		@@ -177,30 +177,47 @@ Section properties.
 | 
				
			|||||||
    toHProp.
 | 
					    toHProp.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
 | 
					  Local Ltac solve_mc :=
 | 
				
			||||||
 | 
					    repeat (let x := fresh in intro x ; try(destruct x))
 | 
				
			||||||
 | 
					    ; simpl
 | 
				
			||||||
 | 
					    ; rewrite transport_sum
 | 
				
			||||||
 | 
					    ; try (f_ap ; apply path_ishprop)
 | 
				
			||||||
 | 
					    ; try (f_ap ; apply set_path2).
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Lemma merely_choice : forall X : FSet A, (X = ∅) + (hexists (fun a => a ∈ X)).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
 | 
					    hinduction.
 | 
				
			||||||
    - apply (tr (inl idpath)).
 | 
					    - apply (inl idpath).
 | 
				
			||||||
    - intro a.
 | 
					    - intro a.
 | 
				
			||||||
      refine (tr (inr (tr (a ; tr idpath)))).
 | 
					      refine (inr (tr (a ; tr idpath))).
 | 
				
			||||||
    - intros X Y TX TY.
 | 
					    - intros X Y TX TY.
 | 
				
			||||||
      strip_truncations.
 | 
					      destruct TX as [XE | HX] ; destruct TY as [YE | HY].
 | 
				
			||||||
      destruct TX as [XE | HX] ; destruct TY as [YE | HY] ; try(strip_truncations ; apply tr).
 | 
					      * refine (inl _).
 | 
				
			||||||
      * refine (tr (inl _)).
 | 
					 | 
				
			||||||
        rewrite XE, YE.
 | 
					        rewrite XE, YE.
 | 
				
			||||||
        apply (union_idem ∅).
 | 
					        apply (union_idem ∅).
 | 
				
			||||||
      * destruct HY as [a Ya].
 | 
					      * right.
 | 
				
			||||||
        refine (inr (tr _)).
 | 
					        strip_truncations.
 | 
				
			||||||
 | 
					        destruct HY as [a Ya].
 | 
				
			||||||
 | 
					        refine (tr _).
 | 
				
			||||||
        exists a.
 | 
					        exists a.
 | 
				
			||||||
        apply (tr (inr Ya)).
 | 
					        apply (tr (inr Ya)).
 | 
				
			||||||
      * destruct HX as [a Xa].
 | 
					      * right.
 | 
				
			||||||
        refine (inr (tr _)).
 | 
					        strip_truncations.
 | 
				
			||||||
 | 
					        destruct HX as [a Xa].
 | 
				
			||||||
 | 
					        refine (tr _).
 | 
				
			||||||
        exists a.
 | 
					        exists a.
 | 
				
			||||||
        apply (tr (inl Xa)).
 | 
					        apply (tr (inl Xa)).
 | 
				
			||||||
      * destruct (HX, HY) as [[a Xa] [b Yb]].
 | 
					      * right.
 | 
				
			||||||
        refine (inr (tr _)).
 | 
					        strip_truncations.
 | 
				
			||||||
 | 
					        destruct (HX, HY) as [[a Xa] [b Yb]].
 | 
				
			||||||
 | 
					        refine (tr _).
 | 
				
			||||||
        exists a.
 | 
					        exists a.
 | 
				
			||||||
        apply (tr (inl Xa)).
 | 
					        apply (tr (inl Xa)).
 | 
				
			||||||
 | 
					    - solve_mc.
 | 
				
			||||||
 | 
					    - solve_mc.
 | 
				
			||||||
 | 
					    - solve_mc.
 | 
				
			||||||
 | 
					    - solve_mc.
 | 
				
			||||||
 | 
					    - solve_mc.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma separation_isIn (B : Type) : forall (X : FSet A) (f : {a | a ∈ X} -> B) (b : B),
 | 
					  Lemma separation_isIn (B : Type) : forall (X : FSet A) (f : {a | a ∈ X} -> B) (b : B),
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -110,8 +110,7 @@ Section structure_k_finite.
 | 
				
			|||||||
    intros.
 | 
					    intros.
 | 
				
			||||||
    destruct X0 as [SX EX].
 | 
					    destruct X0 as [SX EX].
 | 
				
			||||||
    rewrite EX.
 | 
					    rewrite EX.
 | 
				
			||||||
    simple refine (Trunc_ind _ _ (merely_choice SX)).
 | 
					    destruct (merely_choice SX) as [SXE | H1].
 | 
				
			||||||
    intros [SXE | H1].
 | 
					 | 
				
			||||||
    - rewrite SXE.
 | 
					    - rewrite SXE.
 | 
				
			||||||
      apply (tr (inl idpath)).
 | 
					      apply (tr (inl idpath)).
 | 
				
			||||||
    - apply (tr (inr H1)).
 | 
					    - apply (tr (inr H1)).
 | 
				
			||||||
 
 | 
				
			|||||||
		Reference in New Issue
	
	Block a user