mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Added disjunction.
This commit is contained in:
		@@ -1,6 +1,7 @@
 | 
				
			|||||||
-R . "" COQC = hoqc COQDEP = hoqdep
 | 
					-R . "" COQC = hoqc COQDEP = hoqdep
 | 
				
			||||||
-R ../prelude ""
 | 
					-R ../prelude ""
 | 
				
			||||||
definition.v
 | 
					definition.v
 | 
				
			||||||
 | 
					disjunction.v
 | 
				
			||||||
operations.v
 | 
					operations.v
 | 
				
			||||||
Ext.v
 | 
					Ext.v
 | 
				
			||||||
properties.v
 | 
					properties.v
 | 
				
			||||||
 
 | 
				
			|||||||
							
								
								
									
										69
									
								
								FiniteSets/disjunction.v
									
									
									
									
									
										Normal file
									
								
							
							
						
						
									
										69
									
								
								FiniteSets/disjunction.v
									
									
									
									
									
										Normal file
									
								
							@@ -0,0 +1,69 @@
 | 
				
			|||||||
 | 
					Require Import HoTT.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Definition lor (X Y : hProp) : hProp := BuildhProp (Trunc (-1) (sum X Y)).
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Infix "\/" := lor.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Section lor_props.
 | 
				
			||||||
 | 
					  Variable X Y Z : hProp.
 | 
				
			||||||
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Theorem lor_assoc : (X \/ (Y \/ Z)) = ((X \/ Y) \/ Z).
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    apply path_iff_hprop ; cbn.
 | 
				
			||||||
 | 
					    * simple refine (Trunc_ind _ _).
 | 
				
			||||||
 | 
					      intros [x | yz] ; cbn.
 | 
				
			||||||
 | 
					      + apply (tr (inl (tr (inl x)))).
 | 
				
			||||||
 | 
					      + simple refine (Trunc_ind _ _ yz).
 | 
				
			||||||
 | 
					        intros [y | z].
 | 
				
			||||||
 | 
					        ++ apply (tr (inl (tr (inr y)))). 
 | 
				
			||||||
 | 
					        ++ apply (tr (inr z)).
 | 
				
			||||||
 | 
					    * simple refine (Trunc_ind _ _).
 | 
				
			||||||
 | 
					      intros [xy | z] ; cbn.
 | 
				
			||||||
 | 
					      + simple refine (Trunc_ind _ _ xy).
 | 
				
			||||||
 | 
					        intros [x | y].
 | 
				
			||||||
 | 
					        ++ apply (tr (inl x)). 
 | 
				
			||||||
 | 
					        ++ apply (tr (inr (tr (inl y)))).
 | 
				
			||||||
 | 
					      + apply (tr (inr (tr (inr z)))).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Theorem lor_comm : (X \/ Y) = (Y \/ X).
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    apply path_iff_hprop ; cbn.
 | 
				
			||||||
 | 
					    * simple refine (Trunc_ind _ _).
 | 
				
			||||||
 | 
					      intros [x | y].
 | 
				
			||||||
 | 
					      + apply (tr (inr x)).
 | 
				
			||||||
 | 
					      + apply (tr (inl y)).
 | 
				
			||||||
 | 
					    * simple refine (Trunc_ind _ _).
 | 
				
			||||||
 | 
					      intros [y | x].
 | 
				
			||||||
 | 
					      + apply (tr (inr y)).
 | 
				
			||||||
 | 
					      + apply (tr (inl x)).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Theorem lor_nl : (False_hp \/ X) = X.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    apply path_iff_hprop ; cbn.
 | 
				
			||||||
 | 
					    * simple refine (Trunc_ind _ _).
 | 
				
			||||||
 | 
					      intros [ | x].
 | 
				
			||||||
 | 
					      + apply Empty_rec.
 | 
				
			||||||
 | 
					      + apply x.
 | 
				
			||||||
 | 
					    * apply (fun x => tr (inr x)).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Theorem lor_nr : (X \/ False_hp) = X.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    apply path_iff_hprop ; cbn.
 | 
				
			||||||
 | 
					    * simple refine (Trunc_ind _ _).
 | 
				
			||||||
 | 
					      intros [x | ].
 | 
				
			||||||
 | 
					      + apply x.
 | 
				
			||||||
 | 
					      + apply Empty_rec.
 | 
				
			||||||
 | 
					    * apply (fun x => tr (inl x)).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Theorem lor_idem : (X \/ X) = X.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    apply path_iff_hprop ; cbn.
 | 
				
			||||||
 | 
					    - simple refine (Trunc_ind _ _).
 | 
				
			||||||
 | 
					      intros [x | x] ; apply x.
 | 
				
			||||||
 | 
					    - apply (fun x => tr (inl x)).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
		Reference in New Issue
	
	Block a user