mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Cleanup the Pauli group example
This commit is contained in:
		@@ -1,5 +1,5 @@
 | 
				
			|||||||
Require Import HoTT HitTactics.
 | 
					Require Import HoTT HitTactics.
 | 
				
			||||||
Require Export FSets lattice_examples.
 | 
					Require Export FSets lattice_examples k_finite.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Section quantifiers.
 | 
					Section quantifiers.
 | 
				
			||||||
  Context {A : Type} `{Univalence}.
 | 
					  Context {A : Type} `{Univalence}.
 | 
				
			||||||
@@ -132,9 +132,8 @@ Section simple_example.
 | 
				
			|||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
End simple_example.
 | 
					End simple_example.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Require Import k_finite.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Section pauli.
 | 
					Section pauli.
 | 
				
			||||||
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Inductive Pauli : Type :=
 | 
					  Inductive Pauli : Type :=
 | 
				
			||||||
  | XP : Pauli
 | 
					  | XP : Pauli
 | 
				
			||||||
@@ -193,32 +192,25 @@ Section pauli.
 | 
				
			|||||||
    apply _.
 | 
					    apply _.
 | 
				
			||||||
  Defined.  
 | 
					  Defined.  
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Context `{Univalence}.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition Pauli_list : FSet Pauli := {|XP|} ∪ {|YP|} ∪ {|ZP|} ∪ {|IP|}.
 | 
					  Definition Pauli_list : FSet Pauli := {|XP|} ∪ {|YP|} ∪ {|ZP|} ∪ {|IP|}.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
 | 
					  Ltac solve_in_list :=
 | 
				
			||||||
 | 
					    apply tr; try apply idpath;
 | 
				
			||||||
 | 
					    first [ apply inl ; solve_in_list | apply inr ; solve_in_list ].
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Theorem Pauli_finite : Kf Pauli.
 | 
					  Theorem Pauli_finite : Kf Pauli.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    unfold Kf, Kf_sub, Kf_sub_intern.
 | 
					    apply Kf_unfold.
 | 
				
			||||||
    exists Pauli_list.
 | 
					    exists Pauli_list.
 | 
				
			||||||
    apply path_forall.
 | 
					 | 
				
			||||||
    unfold map.
 | 
					    unfold map.
 | 
				
			||||||
    intros [ | | | ] ; rewrite ?union_isIn ; apply path_iff_hprop ; try constructor ; intros [].
 | 
					    intros [ | | | ]; rewrite !union_isIn; solve_in_list.
 | 
				
			||||||
    - apply (tr(inl(tr idpath))).
 | 
					 | 
				
			||||||
    - apply (tr(inr(tr(inl(tr idpath))))).
 | 
					 | 
				
			||||||
    - apply (tr(inr(tr(inr(tr(inl(tr idpath))))))).
 | 
					 | 
				
			||||||
    - refine (tr(inr(tr(inr(tr(inr(tr idpath))))))).
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Theorem Pauli_all (P : Pauli -> hProp) : all P Pauli_list -> forall (x : Pauli), P x.
 | 
					  Theorem Pauli_all (P : Pauli -> hProp) : all P Pauli_list -> forall (x : Pauli), P x.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros HP x.
 | 
					    intros HP x.
 | 
				
			||||||
    refine (all_elim P Pauli_list x HP _).
 | 
					    refine (all_elim P Pauli_list x HP _).
 | 
				
			||||||
    destruct x ; rewrite ?union_isIn ; try constructor.
 | 
					    destruct x ; rewrite ?union_isIn; solve_in_list.
 | 
				
			||||||
    - apply (tr(inl(tr idpath))).
 | 
					 | 
				
			||||||
    - apply (tr(inr(tr(inl(tr idpath))))).
 | 
					 | 
				
			||||||
    - apply (tr(inr(tr(inr(tr(inl(tr idpath))))))).
 | 
					 | 
				
			||||||
    - refine (tr(inr(tr(inr(tr(inr(tr idpath))))))).
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition comm x y : hProp := BuildhProp(Pauli_mult x y = Pauli_mult y x).
 | 
					  Definition comm x y : hProp := BuildhProp(Pauli_mult x y = Pauli_mult y x).
 | 
				
			||||||
 
 | 
				
			|||||||
		Reference in New Issue
	
	Block a user