-
Notifications
You must be signed in to change notification settings - Fork 194
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
Merge pull request #1428 from andreaslyn/infinitary-universal-algebra
Algebra operations
Showing
8 changed files
with
475 additions
and
18 deletions.
There are no files selected for viewing
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,220 @@ | ||
(** This file continues the development of algebra [Operation]. It | ||
gives a way to construct operations using (conventional) curried | ||
functions, and shows that such curried operations are equivalent | ||
to the uncurried operations [Operation]. *) | ||
|
||
Require Export HoTT.Algebra.Universal.Algebra. | ||
|
||
Require Import | ||
HoTT.Basics | ||
HoTT.Types | ||
HoTT.DProp | ||
HoTT.Spaces.Finite | ||
HoTT.Spaces.Nat | ||
HoTT.Spaces.Finite.FinSeq. | ||
|
||
Import notations_algebra. | ||
|
||
(** Functions [head_dom'] and [head_dom] are used to get the first | ||
element of a nonempty operation domain [a : forall i, A (ss i)]. *) | ||
|
||
Monomorphic Definition head_dom' {σ} (A : Carriers σ) (n : nat) | ||
: forall (N : n > 0) (ss : FinSeq n (Sort σ)) (a : forall i, A (ss i)), | ||
A (fshead' n N ss) | ||
:= match n with | ||
| 0 => fun N ss _ => Empty_rec N | ||
| n'.+1 => fun N ss a => a fin_zero | ||
end. | ||
|
||
Monomorphic Definition head_dom {σ} (A : Carriers σ) {n : nat} | ||
(ss : FinSeq n.+1 (Sort σ)) (a : forall i, A (ss i)) | ||
: A (fshead ss) | ||
:= head_dom' A n.+1 tt ss a. | ||
|
||
(** Functions [tail_dom'] and [tail_dom] are used to obtain the tail | ||
of an operation domain [a : forall i, A (ss i)]. *) | ||
|
||
Monomorphic Definition tail_dom' {σ} (A : Carriers σ) (n : nat) | ||
: forall (ss : FinSeq n (Sort σ)) (a : forall i, A (ss i)) (i : Fin (pred n)), | ||
A (fstail' n ss i) | ||
:= match n with | ||
| 0 => fun ss _ i => Empty_rec i | ||
| n'.+1 => fun ss a i => a (fsucc i) | ||
end. | ||
|
||
Monomorphic Definition tail_dom {σ} (A : Carriers σ) {n : nat} | ||
(ss : FinSeq n.+1 (Sort σ)) (a : forall i, A (ss i)) | ||
: forall i, A (fstail ss i) | ||
:= tail_dom' A n.+1 ss a. | ||
|
||
(** Functions [cons_dom'] and [cons_dom] to add an element to | ||
the front of a given domain [a : forall i, A (ss i)]. *) | ||
|
||
Monomorphic Definition cons_dom' {σ} (A : Carriers σ) {n : nat} | ||
: forall (i : Fin n) (ss : FinSeq n (Sort σ)) (N : n > 0), | ||
A (fshead' n N ss) -> (forall i, A (fstail' n ss i)) -> A (ss i) | ||
:= fin_ind | ||
(fun n i => | ||
forall (ss : Fin n -> Sort σ) (N : n > 0), | ||
A (fshead' n N ss) -> (forall i, A (fstail' n ss i)) -> A (ss i)) | ||
(fun n' => fun _ z => match z with tt => fun x _ => x end) | ||
(fun n' i' _ => fun _ _ _ xs => xs i'). | ||
|
||
Definition cons_dom {σ} (A : Carriers σ) | ||
{n : nat} (ss : FinSeq n.+1 (Sort σ)) | ||
(x : A (fshead ss)) (xs : forall i, A (fstail ss i)) | ||
: forall i : Fin n.+1, A (ss i) | ||
:= fun i => cons_dom' A i ss tt x xs. | ||
|
||
(** The empty domain: *) | ||
|
||
Definition nil_dom {σ} (A : Carriers σ) (ss : FinSeq 0 (Sort σ)) | ||
: forall i : Fin 0, A (ss i) | ||
:= Empty_ind (A o ss). | ||
|
||
(** A specialization of [Operation] to finite [Fin n] arity. *) | ||
|
||
Definition FiniteOperation {σ : Signature} (A : Carriers σ) | ||
{n : nat} (ss : FinSeq n (Sort σ)) (t : Sort σ) : Type | ||
:= Operation A {| Arity := Fin n; sorts_dom := ss; sort_cod := t |}. | ||
|
||
(** A type of curried operations | ||
<< | ||
CurriedOperation A [s1, ..., sn] t := A s1 -> ... -> A sn -> A t. | ||
>> *) | ||
|
||
Fixpoint CurriedOperation {σ} (A : Carriers σ) {n : nat} | ||
: (FinSeq n (Sort σ)) -> Sort σ -> Type | ||
:= match n with | ||
| 0 => fun ss t => A t | ||
| n'.+1 => | ||
fun ss t => A (fshead ss) -> CurriedOperation A (fstail ss) t | ||
end. | ||
|
||
(** Function [operation_uncurry] is used to uncurry an operation | ||
<< | ||
operation_uncurry A [s1, ..., sn] t (op : CurriedOperation A [s1, ..., sn] t) | ||
: FiniteOperation A [s1, ..., sn] t | ||
:= fun (x1 : A s1, ..., xn : A xn) => op x1 ... xn | ||
>> | ||
See [equiv_operation_curry] below. *) | ||
|
||
Fixpoint operation_uncurry {σ} (A : Carriers σ) {n : nat} | ||
: forall (ss : FinSeq n (Sort σ)) (t : Sort σ), | ||
CurriedOperation A ss t -> FiniteOperation A ss t | ||
:= match n with | ||
| 0 => fun ss t op _ => op | ||
| n'.+1 => | ||
fun ss t op a => | ||
operation_uncurry A (fstail ss) t (op (a fin_zero)) (a o fsucc) | ||
end. | ||
|
||
Local Example computation_example_operation_uncurry | ||
: forall | ||
(σ : Signature) (A : Carriers σ) (n : nat) (s1 s2 t : Sort σ) | ||
(ss := (fscons s1 (fscons s2 fsnil))) | ||
(op : CurriedOperation A ss t) (a : forall i, A (ss i)), | ||
operation_uncurry A ss t op | ||
= fun a => op (a fin_zero) (a (fsucc fin_zero)). | ||
Proof. | ||
reflexivity. | ||
Qed. | ||
|
||
(** Function [operation_curry] is used to curry an operation | ||
<< | ||
operation_curry A [s1, ..., sn] t (op : FiniteOperation A [s1, ..., sn] t) | ||
: CurriedOperation A [s1, ..., sn] t | ||
:= fun (x1 : A s1) ... (xn : A xn) => op (x1, ..., xn) | ||
>> | ||
See [equiv_operation_curry] below. *) | ||
|
||
Fixpoint operation_curry {σ} (A : Carriers σ) {n : nat} | ||
: forall (ss : FinSeq n (Sort σ)) (t : Sort σ), | ||
FiniteOperation A ss t -> CurriedOperation A ss t | ||
:= match n with | ||
| 0 => fun ss t op => op (Empty_ind _) | ||
| n'.+1 => | ||
fun ss t op x => | ||
operation_curry A (fstail ss) t (op o cons_dom A ss x) | ||
end. | ||
|
||
Local Example computation_example_operation_curry | ||
: forall | ||
(σ : Signature) (A : Carriers σ) (n : nat) (s1 s2 t : Sort σ) | ||
(ss := (fscons s1 (fscons s2 fsnil))) | ||
(op : FiniteOperation A ss t) | ||
(x1 : A s1) (x2 : A s2), | ||
operation_curry A ss t op | ||
= fun x1 x2 => op (cons_dom A ss x1 (cons_dom A _ x2 (nil_dom A _))). | ||
Proof. | ||
reflexivity. | ||
Qed. | ||
|
||
Lemma expand_cons_dom' {σ} (A : Carriers σ) (n : nat) | ||
: forall (i : Fin n) (ss : FinSeq n (Sort σ)) (N : n > 0) | ||
(a : forall i, A (ss i)), | ||
cons_dom' A i ss N (head_dom' A n N ss a) (tail_dom' A n ss a) = a i. | ||
Proof. | ||
intro i. | ||
induction i using fin_ind; intros ss N a. | ||
- unfold cons_dom'. | ||
rewrite compute_fin_ind_fin_zero. | ||
by destruct N. | ||
- unfold cons_dom'. | ||
by rewrite compute_fin_ind_fsucc. | ||
Qed. | ||
|
||
Lemma expand_cons_dom `{Funext} {σ} (A : Carriers σ) | ||
{n : nat} (ss : FinSeq n.+1 (Sort σ)) (a : forall i, A (ss i)) | ||
: cons_dom A ss (head_dom A ss a) (tail_dom A ss a) = a. | ||
Proof. | ||
funext i. | ||
apply expand_cons_dom'. | ||
Defined. | ||
|
||
Lemma path_operation_curry_to_cunurry `{Funext} {σ} (A : Carriers σ) | ||
{n : nat} (ss : FinSeq n (Sort σ)) (t : Sort σ) | ||
: operation_uncurry A ss t o operation_curry A ss t == idmap. | ||
Proof. | ||
intro a. | ||
induction n as [| n IHn]. | ||
- funext d. refine (ap a _). apply path_contr. | ||
- funext a'. | ||
refine (ap (fun x => x _) (IHn _ _) @ _). | ||
refine (ap a _). | ||
apply expand_cons_dom. | ||
Qed. | ||
|
||
Lemma path_operation_uncurry_to_curry `{Funext} {σ} (A : Carriers σ) | ||
{n : nat} (ss : FinSeq n (Sort σ)) (t : Sort σ) | ||
: operation_curry A ss t o operation_uncurry A ss t == idmap. | ||
Proof. | ||
intro a. | ||
induction n; [reflexivity|]. | ||
funext x. | ||
refine (_ @ IHn (fstail ss) (a x)). | ||
refine (ap (operation_curry A (fstail ss) t) _). | ||
funext a'. | ||
simpl. | ||
unfold cons_dom, cons_dom'. | ||
rewrite compute_fin_ind_fin_zero. | ||
refine (ap (operation_uncurry A (fstail ss) t (a x)) _). | ||
funext i'. | ||
now rewrite compute_fin_ind_fsucc. | ||
Qed. | ||
|
||
Global Instance isequiv_operation_curry `{Funext} {σ} (A : Carriers σ) | ||
{n : nat} (ss : FinSeq n (Sort σ)) (t : Sort σ) | ||
: IsEquiv (operation_curry A ss t). | ||
Proof. | ||
srapply isequiv_adjointify. | ||
- apply operation_uncurry. | ||
- apply path_operation_uncurry_to_curry. | ||
- apply path_operation_curry_to_cunurry. | ||
Defined. | ||
|
||
Definition equiv_operation_curry `{Funext} {σ} (A : Carriers σ) | ||
{n : nat} (ss : FinSeq n (Sort σ)) (t : Sort σ) | ||
: FiniteOperation A ss t <~> CurriedOperation A ss t | ||
:= Build_Equiv _ _ (operation_curry A ss t) _. | ||
|
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -1,4 +1,5 @@ | ||
|
||
Require Export Finite.Fin. | ||
Require Export Finite.FinNat. | ||
Require Export Finite.FinInduction. | ||
Require Export Finite.Finite. | ||
Require Export Finite.Tactics. |
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,217 @@ | ||
Require Import | ||
HoTT.Basics | ||
HoTT.Types | ||
HoTT.HSet | ||
HoTT.DProp | ||
HoTT.Spaces.Finite.Fin | ||
HoTT.Spaces.Finite.FinInduction | ||
HoTT.Spaces.Nat. | ||
|
||
(** Finite-dimensional sequence. It is often referred to as vector, | ||
but we call it finite sequence [FinSeq] to avoid confusion with | ||
vector from linear algebra. | ||
Note that the induction principle [finseq_]*) | ||
|
||
Definition FinSeq@{u} (n : nat) (A : Type@{u}) : Type@{u} := Fin n -> A. | ||
|
||
(** The empty finite sequence. *) | ||
|
||
Definition fsnil {A : Type} : FinSeq 0 A := Empty_rec. | ||
|
||
Definition path_fsnil `{Funext} {A : Type} (v : FinSeq 0 A) : fsnil = v. | ||
Proof. | ||
apply path_contr. | ||
Defined. | ||
|
||
(** Add an element in the end of a finite sequence, [fscons'] and [fscons]. *) | ||
|
||
Definition fscons' {A : Type} (n : nat) (a : A) (v : FinSeq (pred n) A) | ||
: FinSeq n A | ||
:= fun i => fin_rec (fun n => FinSeq (pred n) A -> A) | ||
(fun _ _ => a) (fun n' i _ v => v i) i v. | ||
|
||
Definition fscons {A : Type} {n : nat} : A -> FinSeq n A -> FinSeq n.+1 A | ||
:= fscons' n.+1. | ||
|
||
(** Take the first element of a non-empty finite sequence, | ||
[fshead'] and [fshead]. *) | ||
|
||
Definition fshead' {A} (n : nat) : n > 0 -> FinSeq n A -> A | ||
:= match n with | ||
| 0 => fun N _ => Empty_rec N | ||
| n'.+1 => fun _ v => v fin_zero | ||
end. | ||
|
||
Definition fshead {A} {n : nat} : FinSeq n.+1 A -> A := fshead' n.+1 tt. | ||
|
||
Definition compute_fshead' {A} n (N : n > 0) (a : A) (v : FinSeq (pred n) A) | ||
: fshead' n N (fscons' n a v) = a. | ||
Proof. | ||
destruct n; [elim N|]. | ||
exact (apD10 (compute_fin_rec_fin_zero _ _ _ _) v). | ||
Defined. | ||
|
||
Definition compute_fshead {A} {n} (a : A) (v : FinSeq n A) | ||
: fshead (fscons a v) = a. | ||
Proof. | ||
apply compute_fshead'. | ||
Defined. | ||
|
||
(** If the sequence is non-empty, then remove the first element. *) | ||
|
||
Definition fstail' {A} (n : nat) : FinSeq n A -> FinSeq (pred n) A | ||
:= match n with | ||
| 0 => fun _ => Empty_rec | ||
| n'.+1 => fun v i => v (fsucc i) | ||
end. | ||
|
||
(** Remove the first element from a non-empty sequence. *) | ||
|
||
Definition fstail {A} {n : nat} : FinSeq n.+1 A -> FinSeq n A := fstail' n.+1. | ||
|
||
Definition compute_fstail' {A} n (a : A) (v : FinSeq (pred n) A) | ||
: fstail' n (fscons' n a v) == v. | ||
Proof. | ||
intro i. | ||
destruct n; [elim i|]. | ||
exact (apD10 (compute_fin_rec_fsucc _ _ _ _) v). | ||
Defined. | ||
|
||
Definition compute_fstail `{Funext} {A} {n} (a : A) (v : FinSeq n A) | ||
: fstail (fscons a v) = v. | ||
Proof. | ||
funext i. | ||
apply compute_fstail'. | ||
Defined. | ||
|
||
(** A non-empty finite sequence is equal to [fscons] of head and tail, | ||
[path_expand_fscons'] and [path_expand_fscons]. *) | ||
|
||
Lemma path_expand_fscons' {A : Type} (n : nat) | ||
(i : Fin n) (N : n > 0) (v : FinSeq n A) | ||
: fscons' n (fshead' n N v) (fstail' n v) i = v i. | ||
Proof. | ||
induction i using fin_ind. | ||
- apply compute_fshead. | ||
- apply (compute_fstail' n.+1 (fshead v) (fstail v)). | ||
Defined. | ||
|
||
Lemma path_expand_fscons `{Funext} {A} {n} (v : FinSeq n.+1 A) | ||
: fscons (fshead v) (fstail v) = v. | ||
Proof. | ||
funext i. | ||
apply path_expand_fscons'. | ||
Defined. | ||
|
||
(** The following [path_fscons'] and [path_fscons] gives a way to construct | ||
a path between [fscons] finite sequences. They cooperate nicely with | ||
[path_expand_fscons'] and [path_expand_fscons]. *) | ||
|
||
Definition path_fscons' {A} n {a1 a2 : A} {v1 v2 : FinSeq (pred n) A} | ||
(p : a1 = a2) (q : forall i, v1 i = v2 i) (i : Fin n) | ||
: fscons' n a1 v1 i = fscons' n a2 v2 i. | ||
Proof. | ||
induction i using fin_ind. | ||
- exact (compute_fshead _ _ @ p @ (compute_fshead _ _)^). | ||
- refine (_ @ (compute_fstail' n.+1 a2 v2 i)^). | ||
exact (compute_fstail' n.+1 a1 v1 i @ q i). | ||
Defined. | ||
|
||
Definition compute_path_fscons' {A} (n : nat) | ||
(a : A) (v : FinSeq (pred n) A) (i : Fin n) | ||
: path_fscons' n (idpath a) (fun j => idpath (v j)) i = idpath. | ||
Proof. | ||
induction i using fin_ind; unfold path_fscons'. | ||
- rewrite compute_fin_ind_fin_zero. | ||
refine (ap (fun p => p @ _) (concat_p1 _) @ _). | ||
apply concat_pV. | ||
- rewrite compute_fin_ind_fsucc. | ||
refine (ap (fun p => p @ _) (concat_p1 _) @ _). | ||
apply concat_pV. | ||
Qed. | ||
|
||
Definition path_fscons `{Funext} {A} {n} {a1 a2 : A} (p : a1 = a2) | ||
{v1 v2 : FinSeq n A} (q : v1 = v2) | ||
: fscons a1 v1 = fscons a2 v2. | ||
Proof. | ||
funext i. apply path_fscons'. | ||
- assumption. | ||
- intro j. exact (apD10 q j). | ||
Defined. | ||
|
||
Lemma compute_path_fscons `{Funext} {A} {n} (a : A) (v : FinSeq n A) | ||
: path_fscons (idpath a) (idpath v) = idpath. | ||
Proof. | ||
refine (ap (path_forall _ _) _ @ eta_path_forall _ _ _). | ||
funext i. exact (compute_path_fscons' n.+1 a v i). | ||
Defined. | ||
|
||
(** The lemmas [path_expand_fscons_fscons'] and [path_expand_fscons_fscons] | ||
identify [path_expand_fscons'] with [path_fscons'] and | ||
[path_expand_fscons] with [path_fscons]. *) | ||
|
||
Lemma path_expand_fscons_fscons' {A : Type} (n : nat) | ||
(N : n > 0) (a : A) (v : FinSeq (pred n) A) (i : Fin n) | ||
: path_expand_fscons' n i N (fscons' n a v) = | ||
path_fscons' n (compute_fshead' n N a v) (compute_fstail' n a v) i. | ||
Proof. | ||
induction i using fin_ind; unfold path_fscons', path_expand_fscons'. | ||
- do 2 rewrite compute_fin_ind_fin_zero. | ||
refine (_ @ concat_p_pp _ _ _). | ||
refine (_ @ (ap (fun p => _ @ p) (concat_pV _))^). | ||
exact (concat_p1 _)^. | ||
- do 2 rewrite compute_fin_ind_fsucc. | ||
refine (_ @ concat_p_pp _ _ _). | ||
refine (_ @ (ap (fun p => _ @ p) (concat_pV _))^). | ||
exact (concat_p1 _)^. | ||
Qed. | ||
|
||
Lemma path_expand_fscons_fscons `{Funext} | ||
{A : Type} {n : nat} (a : A) (v : FinSeq n A) | ||
: path_expand_fscons (fscons a v) = | ||
path_fscons (compute_fshead a v) (compute_fstail a v). | ||
Proof. | ||
refine (ap (path_forall _ _) _). | ||
funext i. | ||
pose (p := eisretr apD10 (compute_fstail' n.+1 a v)). | ||
refine (_ @ (ap (fun f => _ f i) p)^). | ||
exact (path_expand_fscons_fscons' n.+1 tt a v i). | ||
Defined. | ||
|
||
(** The induction principle for finite sequence, [finseq_ind]. | ||
Note that it uses funext and does not compute. *) | ||
|
||
Lemma finseq_ind `{Funext} {A : Type} (P : forall n, FinSeq n A -> Type) | ||
(z : P 0 fsnil) (s : forall n a (v : FinSeq n A), P n v -> P n.+1 (fscons a v)) | ||
{n : nat} (v : FinSeq n A) | ||
: P n v. | ||
Proof. | ||
induction n. | ||
- exact (transport (P 0) (path_fsnil v) z). | ||
- refine (transport (P n.+1) (path_expand_fscons v) _). | ||
apply s. apply IHn. | ||
Defined. | ||
|
||
Lemma compute_finseq_ind_fsnil `{Funext} {A : Type} | ||
(P : forall n, FinSeq n A -> Type) (z : P 0 fsnil) | ||
(s : forall (n : nat) (a : A) (v : FinSeq n A), P n v -> P n.+1 (fscons a v)) | ||
: finseq_ind P z s fsnil = z. | ||
Proof. | ||
exact (ap (fun x => _ x z) (hset_path2 1 (path_fsnil fsnil)))^. | ||
Defined. | ||
|
||
Lemma compute_finseq_ind_fscons `{Funext} {A : Type} | ||
(P : forall n, FinSeq n A -> Type) (z : P 0 fsnil) | ||
(s : forall (n : nat) (a : A) (v : FinSeq n A), P n v -> P n.+1 (fscons a v)) | ||
{n : nat} (a : A) (v : FinSeq n A) | ||
: finseq_ind P z s (fscons a v) = s n a v (finseq_ind P z s v). | ||
Proof. | ||
simpl. | ||
induction (path_expand_fscons_fscons a v)^. | ||
set (p1 := compute_fshead a v). | ||
set (p2 := compute_fstail a v). | ||
induction p1, p2. | ||
exact (ap (fun p => transport _ p _) (compute_path_fscons _ _)). | ||
Defined. | ||
|
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters