Set Implicit Arguments.
Require Import Lia List.
From Undecidability.HOU Require Import calculus.calculus concon.conservativity_constants
  unification.higher_order_unification unification.systemunification
  unification.nth_order_unification.
Import ListNotations.

Set Default Proof Using "Type".

Hint Rewrite @consts_Lam @consts_AppL @consts_AppR : simplify.

Section InhabitingTypes.

  Variable (X: Const).
  Implicit Types (s t: exp X) ( : ctx) ( : fin exp X).

  Definition inhab (x: fin) (A: type) := Lambda (arity A) (@var X (arity A + x)).

  Lemma inhab_ren delta x A: ren (inhab x A) = inhab ( x) A.
  Proof.
    unfold inhab. asimpl. f_equal. f_equal.
    rewrite it_up_ren_ge; simplify; eauto.
  Qed.


  Lemma inhab_typing Delta x A:
    nth x = Some (target A) ⊢(1) inhab x A : A.
  Proof.
    intros . unfold inhab.
    destruct (type_decompose A) as (L & & ); simplify in ; cbn in .
    specialize (arity_decomposed L ) as .
    rewrite . eapply Lambda_ordertyping; eauto.
    econstructor; cbn; eauto.
    rewrite nth_error_app2; simplify; eauto; .
  Qed.


  Lemma inhab_typing_inv Gamma x A:
     ⊢(1) inhab x A : A nth x = Some (target A).
  Proof.
    intros H; unfold inhab in H.
    eapply Lambda_ordertyping_inv in H as (L & B & ? & ? & ?); subst.
    inv H. rewrite in .
    rewrite nth_error_app2 in ; simplify in *; eauto.
    destruct B; cbn in ; try . now rewrite .
  Qed.


  Lemma inhab_app Gamma Delta x A:
       ⊢(1) inhab x A : A ⊢(1) inhab (length + x) A : A.
  Proof.
    intros H % inhab_typing_inv. eapply inhab_typing.
    rewrite nth_error_app2; simplify; eauto.
  Qed.


  Lemma inhab_typing' Gamma x A:
    nth x = Some A target' ⊢(1) inhab x A : A.
  Proof.
    intros H; eapply inhab_typing.
    unfold target'; erewrite map_nth_error; now eauto.
  Qed.


  Lemma consts_inhab x A: consts (inhab x A) nil.
  Proof.
    unfold inhab; rewrite consts_Lam; reflexivity.
  Qed.


  Lemma normal_inhab x A: normal (inhab x A).
  Proof.
    unfold inhab. generalize (arity A + x) as k.
    induction (arity A); cbn; eauto using normal_lam_intro.
  Qed.


End InhabitingTypes.

#[export] Hint Resolve inhab_app inhab_typing inhab_typing' ord_target' : core.

Section Conservativity.
  Variable (X: Const).

  Section DowncastLemmas.

    Implicit Types (s t: exp X) (: ctx) (A: type).
    Variable (n: ).
    Hypothesis (Leq: 1 n).

    Lemma downcast_variables Gamma s t Delta sigma:
       : s t
       Sigma tau, : s t ord' 1.
    Proof.
      intros.
      pose ( x := match nth x with
                  | Some A inhab X x A
                  | None var x
                  end).
      exists (target' ). exists ( >> subst_exp ). split; [|split].
      - intros ???; eapply preservation_under_substitution; eauto.
        unfold ; intros ?? EQ; rewrite EQ.
        eauto using ordertyping_soundness.
      - erewrite compSubstSubst_exp; eauto. symmetry.
        erewrite compSubstSubst_exp; eauto. symmetry.
        now rewrite .
      - eauto.
    Qed.


    Lemma downcast_constants_ordered m Gamma s t Delta sigma:
      1 m ⊩(m) : s t
       Sigma tau, ⊩(m) : s t
      ord' 1 ( x c, c consts ( x) c Consts [s; t]).
    Proof.
      intros L' T EQ.
      pose ( := Consts [s; t]).
      pose (C := Consts (map (nats (length )))).
      pose ( c := match find c with
                  | Some n const c
                  | None
                    match find c C with
                    | Some x inhab X (length + x) (ctype X c)
                    | None var 0
                    end
                  end).
      exists (target' (map (ctype X) C)). exists ( >> subst_consts ). intuition.
      - intros ???. eapply ordertyping_preservation_consts.
        now eapply weakening_ordertyping_app, T.
        intros c . unfold .
        destruct (find c ); eauto.
        + econstructor; eapply typing_constants; eauto.
        + eapply Consts_consts
            with (S := (map (nats (length )))) in as .
          unfold C; eapply find_in in as [y ]. rewrite .
          eapply ordertyping_monotone, inhab_app, inhab_typing'; eauto.
          erewrite map_nth_error; eauto. now eapply find_Some.
          eapply in_map, lt_nats, nth_error_Some_lt; eauto.
      - rewrite !subst_const_comm_id. now rewrite EQ.
        all: eapply subst_consts_ident; intros x;
          rewrite Consts_consts with (S := [s; t]); intuition.
        all: now unfold , ; eapply find_in in H as [? ].
      - eapply consts_in_subst_consts in H as [d []].
        unfold in *. destruct find eqn: .
        + cbn in ; intuition; subst.
          eapply find_Some, nth_error_In in ; exact .
        + destruct (find d C) eqn: .
          all: rewrite ?consts_inhab in ; cbn in ; intuition.
    Qed.


    Lemma downcast_constants Gamma s t Delta sigma:
       : s t
       Sigma tau, : s t
      ord' 1 ( x c, c consts ( x) c Consts [s; t]).
    Proof using n Leq.
      intros [m H] % ordertypingSubst_complete EQ.
      eapply ordertypingSubst_monotone with ( := S m) in H; eauto.
      eapply downcast_constants_ordered in H as ( & & ?); [||eauto].
      exists ; exists ; intuition. eapply ordertypingSubst_soundness; eauto.
    Qed.


    Lemma ordertyping_from_basetyping Sigma u B:
      ( c, c consts u ord (ctype X c) S n)
       u : B normal u ord B S n ord' n
       ⊢(n) u : B.
    Proof.
      induction 2; [eauto|eauto| |].
      - simplify; intros; econstructor; eapply IHtyping; intuition.
        eauto using normal_lam_elim. cbn; simplify; intuition.
      - intros; enough (ord A n).
        + econstructor; [eapply | eapply ].
          1, 5: intros; eapply H; cbn; intuition.
          1, 4: eauto using normal_app_r, normal_app_l.
          2, 4: eauto.
          1, 2: simplify; eauto; intuition.
        + eapply head_atom in ; eauto; cbn in .
          destruct (head_decompose s) as [T].
          rewrite in H0_.
          eapply AppR_typing_inv in H0_ as (? & ? & ?).
          enough (ord (Arr (rev x) (A B)) S n)
            as by (simplify in ; intuition).
          destruct (head s); cbn in ; eauto.
          all: inv .
          eapply ord'_elements in ; eauto.
          eapply H; cbn; simplify; cbn; intuition.
    Qed.


  Lemma ordertyping_weak_ordertyping Gamma Delta sigma (s t: exp X):
    ( x A, nth x = Some A x vars s vars t ⊢(n) x : A) s t
    ( Sigma tau, ⊩(n) : s t).
  Proof using Leq.
    intros T EQ.
    pose ( x := if x (vars s vars t)
                then x
                else match nth x with
                      | Some A inhab X (length + x) A
                      | None var x
                      end).
    exists ( target' ). exists . split.
    - intros x B H'. unfold . rewrite H'; destruct dec_in.
      eapply weakening_ordertyping_app, T; eauto.
      eapply ordertyping_monotone; eauto.
    - rewrite !(subst_extensional) with ( := ) ( := ); eauto.
      all: intros; unfold ; destruct dec_in; eauto; simplify in *.
      all: exfalso; eauto.
  Qed.


  Lemma unification_downcast sigma Gamma s t Delta A:
     : ⊢(n) s : A ⊢(n) t : A s t
     Sigma tau, ⊩(n) : s t.
  Proof using Leq.
    intros T H.
    pose (P x := x vars s vars t).
    edestruct (downcast_variables) as ( & & T' & H' & O');
      eauto; clear T H .
    edestruct (downcast_constants) as ( & & T'' & H'' & O'' & C'');
      eauto; clear T' H' .
    edestruct (normalise_subst) as ( & R & H & T); eauto.
    eapply ordertyping_weak_ordertyping with ( := ).
    - intros. eapply ordertyping_from_basetyping.
      + intros ?; erewrite consts_subset_steps; eauto.
        intros % C''; cbn in ; simplify in .
        destruct ; eapply typing_constants; [|eauto| |eauto]; eauto.
      + eauto.
      + domin ; eauto.
      + simplify in ; destruct ; eapply vars_ordertyping_nth; eauto.
      + simplify; rewrite O', O''; cbn; eauto.
    - rewrite !subst_pointwise_equiv with ( := ) ( := ); eauto.
  Qed.


  Lemma linearize_consts (S: list (exp X)):
    consts (linearize_terms S) Consts S.
  Proof.
    unfold linearize_terms. cbn; simplify.
    split; unfold Consts; intros ? [x] % in_flat_map; eapply in_flat_map.
    all: intuition; try mapinj. eexists; intuition; eauto.
    now rewrite consts_ren in .
    exists (ren shift x); intuition. now rewrite consts_ren.
  Qed.


  Lemma unification_downcast_eqs sigma Delta (E: eqs X) Gamma L:
    ( ⊢₊₊(n) E : L)
    ( : ) ( •₊ left_side E) ≡₊ ( •₊ right_side E)
     Sigma tau, ⊩(n) : ( •₊ left_side E) ≡₊ ( •₊ right_side E).
  Proof using Leq.
    intros H.
    pose (P x := x Vars' E). pose (s := linearize_terms (left_side E)).
    pose (t := linearize_terms (right_side E)).
    edestruct (downcast_variables) with (s := s) (t := t) as ( & & T' & H' & O'); eauto.
    1: unfold s, t; now rewrite !linearize_terms_subst, linearize_terms_equiv.
    clear H .
    edestruct (downcast_constants) with (s := s) (t := t)
      as ( & & T'' & H'' & O'' & C''); eauto; clear T' H' .
    edestruct (normalise_subst) as ( & R & H & T); eauto.
    edestruct ordertyping_weak_ordertyping with (s := s) (t := t) ( := ).
    - intros. eapply ordertyping_from_basetyping.
      + intros ?; erewrite consts_subset_steps; eauto.
        intros % C''. cbn [Consts flat_map] in .
        simplify in . unfold s, t in .
        rewrite !linearize_consts in .
        destruct ; eapply typing_Consts.
        eapply left_ordertyping; eauto. eauto.
        eapply right_ordertyping; eauto. eauto.
      + eauto.
      + domin ; eauto.
      + simplify in . unfold s, t in ; rewrite !linearize_vars in .
        destruct as [|]; eapply Vars_listordertyping in .
        2, 4: eauto using left_ordertyping, right_ordertyping.
        all: destruct as [? [ ]]; rewrite in ;
          injection as ; eauto.
      + simplify; rewrite O', O''; cbn; eauto.
    - rewrite !subst_pointwise_equiv with ( := ) ( := ); eauto.
    - destruct as []. exists x. exists . destruct ; split; eauto.
      unfold s, t in ; now rewrite linearize_terms_equiv, !linearize_terms_subst.
  Qed.


  End DowncastLemmas.

  Program Instance unification_monotone {n} (D: orduni n X): orduni (S n) X :=
    { Gamma₀ := Gamma₀; s₀ := s₀ ; t₀ := t₀; A₀ := A₀; }.

  Lemma unification_monotone_forward n (I: orduni n X):
    OU n X I OU (S n) X (unification_monotone I).
  Proof.
    intros ( & & T & H). exists ; exists . intuition.
    eapply ordertypingSubst_monotone; eauto.
  Qed.


  Lemma unification_monotone_backward n (I: orduni n X):
    1 n OU (S n) X (unification_monotone I) OU n X I.
  Proof.
    intros Leq ( & & T & H).
    unfold OU. eapply unification_downcast; eauto.
    intros ???; eapply ordertyping_soundness; eauto.
  Qed.


  Program Instance unification_conservative {n} (D: orduni n X): uni X :=
    { Gamma := Gamma₀; s := s₀ ; t := t₀; A := A₀; }.
  Next Obligation.
    eapply ordertyping_soundness; eauto.
  Qed.
  Next Obligation.
    eapply ordertyping_soundness; eauto.
  Qed.

  Lemma unification_conservative_forward n (I: orduni n X):
    OU n X I U X (unification_conservative I).
  Proof.
    intros ( & & T & H). exists ; exists . intuition.
    eapply ordertypingSubst_soundness; eauto.
  Qed.


  Lemma unification_conservative_backward n (I: orduni n X):
    1 n U X (unification_conservative I) OU n X I.
  Proof.
    intros Leq ( & & T & H).
    unfold OU. eapply unification_downcast; eauto.
  Qed.


  Program Instance SU_monotone {n} (I: ordsysuni X n): ordsysuni X (S n) :=
    { Gamma₀' := @Gamma₀' _ _ I; E₀' := @E₀' _ _ I ; L₀' := @L₀' _ _ I; }.
  Next Obligation.
    eapply eqs_ordertyping_step, @H₀'.
  Qed.

  Lemma SU_monotone_forward n (I: ordsysuni X n):
    SOU X n I SOU X (S n) (SU_monotone I).
  Proof.
    intros ( & & T & H). exists ; exists . intuition.
    eapply ordertypingSubst_monotone; eauto.
  Qed.


  Lemma SU_monotone_backward n (I: ordsysuni X n):
    1 n SOU X (S n) (SU_monotone I) SOU X n I.
  Proof.
    intros Leq ( & & T & H).
    unfold OU.
    edestruct unification_downcast_eqs as ( & & ); eauto using equiv_pointwise_eqs, @H₀'.
    intros ???; eapply ordertyping_soundness; eauto.
    exists , ; intuition; eauto using equiv_eqs_pointwise.
  Qed.


  Program Instance SU_conservative {n} (I: ordsysuni X n): sysuni X :=
    { Gammaᵤ' := @Gamma₀' _ _ I; Eᵤ' := @E₀' _ _ I ; Lᵤ' := @L₀' _ _ I; }.
  Next Obligation.
    eapply eqs_ordertyping_soundness, H₀'.
  Qed.

  Lemma SU_conservative_forward n (I: ordsysuni X n):
    SOU X n I SU X (SU_conservative I).
  Proof.
    intros ( & & T & H). exists ; exists . intuition.
    eapply ordertypingSubst_soundness; eauto.
  Qed.


  Lemma SU_conservative_backward n (I: ordsysuni X n):
    1 n SU X (SU_conservative I) SOU X n I.
  Proof.
    intros Leq ( & & T & H).
    edestruct unification_downcast_eqs as ( & & ); eauto using equiv_pointwise_eqs, @H₀'.
    exists , ; intuition; eauto using equiv_eqs_pointwise.
  Qed.


  Theorem unification_step n: 1 n OU n X OU (S n) X.
  Proof.
    intros H; exists unification_monotone; split;
      eauto using unification_monotone_forward, unification_monotone_backward.
  Qed.


  Theorem unification_steps n m: 1 n m OU n X OU m X.
  Proof.
    intros [ ]; induction ; eauto using unification_step, Arith.Le.le_trans.
  Qed.


  Theorem unification_conserve n: 1 n OU n X U X.
  Proof.
    intros H; exists unification_conservative; split;
      eauto using unification_conservative_forward, unification_conservative_backward.
  Qed.


  Theorem systemunification_step n: 1 n @SOU X n @SOU X (S n).
  Proof.
    intros H; exists SU_monotone; split;
      eauto using SU_monotone_forward, SU_monotone_backward.
  Qed.


  Theorem systemunification_steps n m: 1 n m @SOU X n @SOU X m.
  Proof.
    intros [ ]; induction ; eauto using systemunification_step, Arith.Le.le_trans.
  Qed.


  Theorem systemunification_conserve n: 1 n @SOU X n @SU X.
  Proof.
    intros H; exists SU_conservative; split;
      eauto using SU_conservative_forward, SU_conservative_backward.
  Qed.


End Conservativity.