Mercurial > hg > Members > kono > Proof > galois
view FLComm.agda @ 188:9e0d946d35b6
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sat, 28 Nov 2020 21:50:16 +0900 |
parents | c22ef5bc695a |
children | d0b678eec506 |
line wrap: on
line source
{-# OPTIONS --allow-unsolved-metas #-} open import Data.Nat -- using (ℕ; suc; zero; s≤s ; z≤n ) module FLComm (n : ℕ) where open import Level renaming ( suc to Suc ; zero to Zero ) hiding (lift) open import Data.Fin hiding ( _<_ ; _≤_ ; _-_ ; _+_ ; _≟_) open import Data.Fin.Properties hiding ( <-trans ; ≤-refl ; ≤-trans ; ≤-irrelevant ; _≟_ ) renaming ( <-cmp to <-fcmp ) open import Data.Fin.Permutation open import Data.Nat.Properties open import Relation.Binary.PropositionalEquality hiding ( [_] ) renaming ( sym to ≡-sym ) open import Data.List using (List; []; _∷_ ; length ; _++_ ; tail ) renaming (reverse to rev ) open import Data.Product hiding (_,_ ) open import Relation.Nullary open import Data.Unit open import Data.Empty open import Relation.Binary.Core open import Relation.Binary.Definitions hiding (Symmetric ) open import logic open import nat open import FLutil open import Putil import Solvable open import Symmetric -- infixr 100 _::_ open import Relation.Nary using (⌊_⌋) open import Data.List.Fresh hiding ([_]) open import Data.List.Fresh.Relation.Unary.Any open Solvable (Symmetric n) tl3 : (FL n) → ( z : FList n) → FList n → FList n tl3 h [] w = w tl3 h (x ∷# z) w = tl3 h z (FLinsert ( perm→FL [ FL→perm h , FL→perm x ] ) w ) tl2 : ( x z : FList n) → FList n → FList n tl2 [] _ x = x tl2 (h ∷# x) z w = tl2 x z (tl3 h z w) CommFList : FList n → FList n CommFList x = tl2 x x [] CommFListN : ℕ → FList n CommFListN 0 = ∀Flist fmax CommFListN (suc i) = CommFList (CommFListN i) open import Algebra open Group (Symmetric n) CommStage→ : (i : ℕ) → (x : Permutation n n ) → deriving i x → Any (perm→FL x ≡_) ( CommFListN i ) CommStage→ zero x (Level.lift tt) = AnyFList (perm→FL x) CommStage→ (suc i) .( [ g , h ] ) (comm {g} {h} p q) = comm2 (CommFListN i) (CommFListN i) (CommStage→ i g p) (CommStage→ i h q) [] where G = perm→FL g H = perm→FL h comm3 : (L L1 : FList n) → (xr : fresh (FL n) ⌊ _f<?_ ⌋ G L ) → Any (H ≡_) L1 → (L3 : FList n) → Any (_≡_ (perm→FL [ g , h ])) (tl2 (cons G L xr) L1 L3) comm3 [] L1 _ b L3 = {!!} comm3 (cons a L x) L1 _ b L3 = {!!} comm2 : (L L1 : FList n) → Any (G ≡_) L → Any (H ≡_) L1 → (L3 : FList n) → Any (perm→FL [ g , h ] ≡_) (tl2 L L1 L3) comm2 (cons G L xr) L1 (here refl) b L3 = comm3 L L1 xr b L3 comm2 (cons x L xr) L1 (there a) b L3 = comm2 L L1 a b {!!} CommStage→ (suc i) x (ccong {f} {x} eq p) = subst (λ k → Any (k ≡_) (tl2 (CommFListN i) (CommFListN i) [])) (comm4 eq) (CommStage→ (suc i) f p ) where comm4 : f =p= x → perm→FL f ≡ perm→FL x comm4 = pcong-pF CommSolved : (x : Permutation n n) → (y : FList n) → y ≡ FL0 ∷# [] → (FL→perm (FL0 {n}) =p= pid ) → Any (perm→FL x ≡_) y → x =p= pid CommSolved x .(cons FL0 [] (Level.lift tt)) refl eq0 (here eq) = FLpid _ eq eq0