mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Added decidable quantification
This commit is contained in:
		@@ -26,3 +26,4 @@ misc/dec_lem.v
 | 
				
			|||||||
misc/ordered.v
 | 
					misc/ordered.v
 | 
				
			||||||
misc/projective.v
 | 
					misc/projective.v
 | 
				
			||||||
misc/dec_kuratowski.v
 | 
					misc/dec_kuratowski.v
 | 
				
			||||||
 | 
					misc/dec_fset.v
 | 
				
			||||||
@@ -117,3 +117,17 @@ Section quantifiers.
 | 
				
			|||||||
    hinduction ; try (apply _) ; try (intros ; apply path_ishprop).
 | 
					    hinduction ; try (apply _) ; try (intros ; apply path_ishprop).
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
End quantifiers.
 | 
					End quantifiers.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Section simple_example.
 | 
				
			||||||
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Definition P : nat -> hProp := fun n => BuildhProp(n = n).
 | 
				
			||||||
 | 
					  Definition X : FSet nat := {|0|} ∪ {|1|}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Definition simple_example : all P X.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    refine (from_squash (all P X)).
 | 
				
			||||||
 | 
					    compute.
 | 
				
			||||||
 | 
					    apply tt.
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					End simple_example.
 | 
				
			||||||
@@ -1,6 +1,20 @@
 | 
				
			|||||||
(** Some general prerequisities in homotopy type theory. *)
 | 
					(** Some general prerequisities in homotopy type theory. *)
 | 
				
			||||||
Require Import HoTT.
 | 
					Require Import HoTT.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Definition squash (A : Type) `{Decidable A} : Type :=
 | 
				
			||||||
 | 
					  match dec A with
 | 
				
			||||||
 | 
					  | inl _ => Unit
 | 
				
			||||||
 | 
					  | inr _ => Empty
 | 
				
			||||||
 | 
					  end.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Definition from_squash (A : Type) `{Decidable A} {x : squash A} : A.
 | 
				
			||||||
 | 
					Proof.
 | 
				
			||||||
 | 
					  unfold squash in *.
 | 
				
			||||||
 | 
					  destruct (dec A).
 | 
				
			||||||
 | 
					  - apply a.
 | 
				
			||||||
 | 
					  - contradiction.
 | 
				
			||||||
 | 
					Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Lemma ap_inl_path_sum_inl {A B} (x y : A) (p : inl x = inl y) :
 | 
					Lemma ap_inl_path_sum_inl {A B} (x y : A) (p : inl x = inl y) :
 | 
				
			||||||
  ap inl (path_sum_inl B p) = p.
 | 
					  ap inl (path_sum_inl B p) = p.
 | 
				
			||||||
Proof.
 | 
					Proof.
 | 
				
			||||||
 
 | 
				
			|||||||
		Reference in New Issue
	
	Block a user