Mercurial > hg > Members > kono > Proof > galois
view FLComm.agda @ 224:71e08656273b
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 07 Dec 2020 22:30:29 +0900 |
parents | 9da456c2bfe3 |
children | 08a99237e56e |
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 import Algebra open Group (Symmetric n) hiding (refl) open Solvable (Symmetric n) open _∧_ -- open import Relation.Nary using (⌊_⌋) open import Relation.Nullary.Decidable hiding (⌊_⌋) -- Flist : {n : ℕ } (i : ℕ) → i < suc n → FList n → FList n → FList (suc n) -- Flist zero i<n [] _ = [] -- Flist zero i<n (a ∷# x ) z = FLinsert ( zero :: a ) (Flist zero i<n x z ) -- Flist (suc i) (s≤s i<n) [] z = Flist i (<-trans i<n a<sa) z z -- Flist (suc i) i<n (a ∷# x ) z = FLinsert ((fromℕ< i<n ) :: a ) (Flist (suc i) i<n x z ) -- -- ∀Flist : {n : ℕ } → FL n → FList n -- ∀Flist {zero} f0 = f0 ∷# [] -- Flist {suc n} (x :: y) = Flist n a<sa (∀Flist y) (∀Flist y) -- all FL record AnyFL (n : ℕ) (p : FL n) : Set where field anyList : FList n anyP : (x : FL n) → p f≤ x → Any (x ≡_ ) anyList open import fin open AnyFL anyFL : (n : ℕ ) → AnyFL n FL0 anyFL zero = record { anyList = f0 ∷# [] ; anyP = any00 } where any00 : (p : FL zero) → FL0 f≤ p → Any (p ≡_) (f0 ∷# []) any00 f0 (case1 refl) = here refl anyFL (suc n) = any01 n (anyList (anyFL n)) (anyP (anyFL n) FL0 (case1 refl) ) {!!} where -- any03 : AnyFL (suc n) (fromℕ< a<sa :: fmax) → AnyFL (suc n) FL0 -- loop on i any02 : (i : ℕ ) → (i<n : i < suc n ) → (a : FL n) → AnyFL (suc n) (fromℕ< i<n :: a) → AnyFL (suc n) (zero :: a) any02 zero (s≤s z≤n) a any = any any02 (suc i) (s≤s i<n) a any = any02 i (<-trans i<n a<sa) a record { anyList = cons ((fromℕ< (s≤s i<n )) :: a) (anyList any) any07 ; anyP = any08 } where any07 : fresh (FL (suc n)) ⌊ _f<?_ ⌋ (fromℕ< (s≤s i<n) :: a) (anyList any) any07 = {!!} any08 : (x : FL (suc n)) → (fromℕ< (<-trans i<n a<sa) :: a) f≤ x → Any (x ≡_) (cons (fromℕ< (s≤s i<n) :: a) (anyList any) any07 ) any08 = {!!} -- loop on a any03 : (L : FList n) → (a : FL n) → fresh (FL n) ⌊ _f<?_ ⌋ a L → AnyFL (suc n) (fromℕ< a<sa :: a ) → AnyFL (suc n) FL0 any03 [] a ar any = {!!} -- any02 n a<sa a any any03 (cons b L br) a ( Data.Product._,_ (Level.lift a<b)_) any = any03 L b br record { anyList = anyList any04 ; anyP = any05 } where any04 : AnyFL (suc n) (zero :: a) any04 = any02 n a<sa a any any05 : (x : FL (suc n)) → (fromℕ< a<sa :: b) f≤ x → Any (x ≡_) (anyList any04) -- 0<fmax : zero Data.Fin.< fromℕ< a<sa any05 x mb≤x = anyP any04 x (any06 a b x (toWitness a<b) mb≤x) where any06 : {n : ℕ } → (a b : FL n) → (x : FL (suc n)) → a f< b → (fromℕ< {n} a<sa :: b) f≤ x → (zero :: a) f≤ x any06 {suc n} a b x a<b (case1 refl) = case2 (f<n 0<fmax) any06 {suc n} a b x a<b (case2 mb<x) = case2 (f<-trans (f<n 0<fmax) mb<x) any01 : (i : ℕ ) → (L : FList n) → Any (FL0 ≡_) L → AnyFL (suc n) fmax → AnyFL (suc n) FL0 any01 zero [] () any01 (suc i) [] () any01 zero (cons a L x) _ any = {!!} any01 (suc i) (cons .FL0 L x) (here refl) any = any01 i L {!!} {!!} -- can't happen any01 (suc i) (cons a L ar) (there b) any = any03 L a ar {!!} 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) -- all comm cobmbination in P and Q record AnyComm (p0 q0 : FL n) (P Q : FList n) : Set where field commList : FList n commAny : (p q : FL n) → Any ( p ≡_ ) P → Any ( q ≡_ ) Q → p0 f≤ p → q0 f≤ q → Any (perm→FL [ FL→perm p , FL→perm q ] ≡_) commList open AnyComm -- q₂ anyComm : (P Q : FList n) → AnyComm FL0 FL0 P Q anyComm [] [] = record { commList = [] ; commAny = λ _ _ () } anyComm [] (cons q Q qr) = record { commList = [] ; commAny = λ _ _ () } anyComm (cons p P pr) [] = record { commList = [] ; commAny = λ _ _ _ () } anyComm (cons p P pr) (cons q Q qr) = anyc0n (cons q Q qr) where anyc0n : (Q1 : FList n) → AnyComm {!!} {!!} (cons p P pr) (cons q Q qr) anyc0n [] = record { commList = FLinsert (perm→FL [ FL→perm p , FL→perm q ]) (commList (anyComm P (cons q Q qr))) ; commAny = anyc03 } where anyc03 : (p₁ q₁ : FL n) → Any ( p₁ ≡_) (cons p P pr) → Any (q₁ ≡_) (cons q Q qr) → {!!} → {!!} → Any ((perm→FL [ FL→perm p₁ , FL→perm q₁ ]) ≡_) (FLinsert (perm→FL [ FL→perm p , FL→perm q ]) (commList (anyComm P (cons q Q qr)))) anyc03 p₁ q₁ (there anyp) (here x) _ _ = insAny _ (commAny (anyComm P (cons q Q qr)) p₁ q₁ anyp (here x) {!!} {!!} ) anyc03 p₁ q₁ (there anyp) (there anyq) _ _ = insAny _ (commAny (anyComm P (cons q Q qr)) p₁ q₁ anyp (there anyq) {!!} {!!} ) anyc03 p₁ q₁ (here refl) (there anyq) _ _ = insAny _ (commAny (anyComm P (cons q Q qr)) p₁ q₁ {!!} (there anyq) {!!} {!!} ) anyc03 p₁ q₁ (here refl) (here x) _ _ = {!!} anyc0n (cons q₂ Q1 qr₂ ) = record { commList = FLinsert (perm→FL [ FL→perm p , FL→perm q₂ ]) (commList (anyc0n Q1)) ; commAny = anyc01 Q1 q₂ qr₂ } where anyc01 : (Q1 : FList n) (q₂ : FL n) → (qr₂ : fresh (FL n) ⌊ _f<?_ ⌋ q₂ Q1 ) → (p₁ q₁ : FL n) → Any (p₁ ≡_) (cons p P pr) → Any (q₁ ≡_) (cons q Q qr) → {!!} → {!!} → Any (perm→FL [ FL→perm p₁ , FL→perm q₁ ] ≡_) (FLinsert (perm→FL [ FL→perm p , FL→perm q₂ ]) (commList (anyc0n Q1)) ) anyc01 Q1 q₂ qr₂ p₁ q₁ (here refl) (there anyq) _ _ = insAny _ (commAny (anyc0n Q1) p₁ q₁ (here refl) (there anyq) {!!} {!!} ) anyc01 Q1 q₂ qr₂ p₁ q₁ (there anyp) (here refl) _ _ = insAny _ (commAny (anyc0n Q1) p₁ q₁ (there anyp) (here refl) {!!} {!!} ) anyc01 Q1 q₂ qr₂ p₁ q₁ (there anyp) (there anyq) _ _ = insAny _ (commAny (anyc0n Q1) p₁ q₁ (there anyp) (there anyq) {!!} {!!} ) anyc01 Q1 q₂ qr₂ p₁ q₁ (here refl) (here refl) _ _ with x∈FLins (perm→FL [ FL→perm p , FL→perm q₂ ]) (commList (anyc0n Q1)) ... | t = {!!} -- t : p₁ q₂ → p₁ q₁ -- {-# TERMINATING #-} 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 -- tl3 case commc : (L3 L1 : FList n) → Any ((perm→FL [ FL→perm G , FL→perm H ]) ≡_) L3 → Any ((perm→FL [ FL→perm G , FL→perm H ]) ≡_) (tl3 G L1 L3) commc L3 [] any = any commc L3 (cons a L1 _) any = commc (FLinsert (perm→FL [ FL→perm G , FL→perm a ]) L3) L1 (insAny _ any) comm6 : perm→FL [ FL→perm G , FL→perm H ] ≡ perm→FL [ g , h ] comm6 = pcong-pF (comm-resp (FL←iso _) (FL←iso _)) comm3 : (L1 : FList n) → Any (H ≡_) L1 → (L3 : FList n) → Any ((perm→FL [ g , h ]) ≡_) (tl3 G L1 L3) comm3 (H ∷# []) (here refl) L3 = subst (λ k → Any (k ≡_) (FLinsert (perm→FL [ FL→perm G , FL→perm H ]) L3 ) ) comm6 (x∈FLins ( perm→FL [ FL→perm G , FL→perm H ] ) L3 ) comm3 (cons H L1 x₁) (here refl) L3 = subst (λ k → Any (k ≡_) (tl3 G L1 (FLinsert (perm→FL [ FL→perm G , FL→perm H ]) L3))) comm6 (commc (FLinsert (perm→FL [ FL→perm G , FL→perm H ]) L3 ) L1 (x∈FLins ( perm→FL [ FL→perm G , FL→perm H ] ) L3)) comm3 (cons a L _) (there b) L3 = comm3 L b (FLinsert (perm→FL [ FL→perm G , FL→perm a ]) L3) -- tl2 case 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 = comm7 L L3 where comm8 : (L L4 L2 : FList n) → (a : FL n) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 L2) → Any ((perm→FL [ g , h ]) ≡_ ) (tl2 L4 L1 (tl3 a L L2)) comm8← : (L L4 L2 : FList n) → (a : FL n) → ¬ ( a ≡ perm→FL g ) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 (tl3 a L L2)) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 L2) comm9← : (L4 L2 : FList n) → (a a₁ : FL n) → ¬ ( a ≡ perm→FL g ) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2)) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 L2) -- Any (_≡ (perm→FL [ g , h ])) (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2) → Any ((perm→FL [ g , h ]) ≡_) L2 comm9← [] L2 a a₁ not any = {!!} comm9← (cons a₂ L4 x) L2 a a₁ not any = comm8 L1 L4 L2 a₂ (comm9← L4 L2 a a₁ not (comm8← L1 L4 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2 ) a₂ {!!} any)) -- Any (_≡ (perm→FL [ g , h ])) (tl2 L4 L1 (tl3 a₂ L1 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2))) → -- Any (_≡ (perm→FL [ g , h ])) (tl2 L4 L1 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2)) -- Any (_≡ (perm→FL [ g , h ])) (tl2 L4 L1 L2) -- Any (_≡ (perm→FL [ g , h ])) (tl2 L4 L1 (tl3 a₂ L1 L2)) comm8← [] L4 L2 a _ any = any comm8← (cons a₁ L x) L4 L2 a not any = comm8← L L4 L2 a not (comm9← L4 (tl3 a L L2 ) a a₁ not (subst (λ k → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 k )) {!!} any )) -- Any (_≡ (perm→FL [ g , h ])) (tl2 L4 L1 (tl3 a L (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2))) → -- Any (_≡ (perm→FL [ g , h ])) (tl2 L4 L1 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) (tl3 a L L2))) comm9 : (L4 L2 : FList n) → (a a₁ : FL n) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 L2) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L4 L1 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2)) comm8 [] L4 L2 a any = any comm8 (cons a₁ L x) L4 L2 a any = comm8 L L4 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2) a (comm9 L4 L2 a a₁ any) comm9 [] L2 a a₁ any = insAny _ any comm9 (cons a₂ L4 x) L2 a a₁ any = comm8 L1 L4 (FLinsert (perm→FL [ FL→perm a , FL→perm a₁ ]) L2) a₂ (comm9 L4 L2 a a₁ (comm8← L1 L4 L2 a₂ {!!} any)) -- found g, we have to walk thru till the end comm7 : (L L3 : FList n) → Any ((perm→FL [ g , h ]) ≡_) (tl2 L L1 (tl3 G L1 L3)) -- at the end, find h comm7 [] L3 = comm3 L1 b L3 -- add n path comm7 (cons a L4 xr) L3 = comm8 L1 L4 (tl3 G L1 L3) a (comm7 L4 L3) -- accumerate tl3 comm2 (cons x L xr) L1 (there a) b L3 = comm2 L L1 a b (tl3 x L1 L3) 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