From Coq Require Import Strings.String.
Require Import terms.
Require Import subst.
Require Import equiv.
Require Import subtype.
Require Import context.

(* Given a context, there is a morphism path from τ to τ' *)
Reserved Notation "Γ '|-' σ '~>' τ" (at level 101, σ at next level, τ at next level).

Open Scope ladder_expr_scope.

Inductive morphism_path : context -> type_term -> type_term -> Prop :=
  | M_Sub : forall Γ τ τ',
    (τ :<= τ') ->
    (Γ |- τ ~> τ')

  | M_Single : forall Γ h τ τ',
    (context_contains Γ h [< τ ->morph τ' >]) ->
    (Γ |- τ ~> τ')

  | M_Chain : forall Γ τ τ' τ'',
    (Γ |- τ ~> τ') ->
    (Γ |- τ' ~> τ'') ->
    (Γ |- τ ~> τ'')

  | M_Lift : forall Γ σ τ τ',
    (Γ |- τ ~> τ') ->
    (Γ |- [< σ ~ τ >] ~> [< σ ~ τ' >])

  | M_MapSeq : forall Γ τ τ',
    (Γ |- τ ~> τ') ->
    (Γ |- [< [τ] >] ~> [< [τ'] >])

where "Γ '|-' s '~>' t" := (morphism_path Γ s t).

Lemma id_morphism_path : forall Γ τ, Γ |- τ ~> τ.
Proof.
  intros.
  apply M_Sub, TSubRepr_Refl, TEq_Refl.
Qed.

Inductive translate_morphism_path : context -> type_term -> type_term -> expr_term -> Prop :=
  | Translate_Descend : forall Γ τ τ',
    (τ :<= τ') ->
    (translate_morphism_path Γ τ τ'
      (expr_morph "x" τ [{ %"x"% des τ' }]))

  | Translate_Lift : forall Γ σ τ τ' m,
    (Γ |- τ ~> τ') ->
    (translate_morphism_path Γ τ τ' m) ->
    (translate_morphism_path Γ [< σ ~ τ >] [< σ ~ τ' >]
      (expr_morph "x" [< σ ~ τ >] [{ (m (%"x"% des τ)) as σ }]))

  | Translate_Single : forall Γ h τ τ',
    (context_contains Γ h [< τ ->morph τ' >]) ->
    (translate_morphism_path Γ τ τ' [{ %h% }])

  | Translate_Chain : forall Γ τ τ' τ'' m1 m2,
    (translate_morphism_path Γ τ τ' m1) -> 
    (translate_morphism_path Γ τ' τ'' m2) ->
    (translate_morphism_path Γ τ τ''
      (expr_morph "x" τ [{ m2 (m1 %"x"%) }]))

  | Translate_MapSeq : forall Γ τ τ' m,
    (translate_morphism_path Γ τ τ' m) ->
    (translate_morphism_path Γ [< [τ] >] [< [τ'] >]
      [{
        λ"xs",[τ] ↦morph (%"map"% # τ # τ' m %"xs"%)
      }])
.

Example morphism_paths :
  (ctx_assign "degrees-to-turns" [< $"Angle"$~$"Degrees"$~$"ℝ"$ ->morph $"Angle"$~$"Turns"$~$"ℝ"$ >]
    (ctx_assign "turns-to-radians" [< $"Angle"$~$"Turns"$~$"ℝ"$ ->morph $"Angle"$~$"Radians"$~$"ℝ"$ >]
      ctx_empty))

  |- [< [ $"Hue"$~$"Angle"$~$"Degrees"$~$"ℝ"$ ] >]
  ~> [< [ $"Hue"$~$"Angle"$~$"Radians"$~$"ℝ"$ ] >]
.
Proof.
  intros.
  apply M_MapSeq.
  apply M_Lift.
  apply M_Chain with (τ':=[<$"Angle"$~$"Turns"$~$"ℝ"$>]).
  apply M_Single with (h:="degrees-to-turns"%string).
  apply C_take.
  
  apply M_Single with (h:="turns-to-radians"%string).
  apply C_shuffle.
  apply C_take.
Qed.