Set Implicit Arguments.
Require Import RelationClasses Morphisms Init.Wf List Lia Init.Nat Setoid.
From Undecidability.HOU Require Import calculus.calculus unification.unification.
From Undecidability.HOU.second_order Require Export diophantine_equations goldfarb.encoding goldfarb.multiplication.
Import ListNotations.

Set Default Proof Using "Type".

Section EquationEquivalences.

  Variable (: fin exp ag).
  Hypothesis (N: x, normal ( x)).

  Section Variables.
    Lemma forward_vars x n: encodes ( (F x)) n fst (varEQ x) snd (varEQ x).
    Proof.
      intros H. unfold varEQ; cbn; unfold funcomp.
      change shift with (add 1); rewrite ren_plus_combine, !H; cbn; unfold var_zero.
      change (g a (var 1)) with (enc 1 (var 1)); now simplify.
    Qed.


    Lemma backward_vars x: fst (varEQ x) snd (varEQ x) n, encodes ( (F x)) n.
    Proof using N.
      cbn; simplify.
      unfold funcomp; change shift with (add 1); rewrite ren_plus_combine; unfold var_zero; cbn.
      intros H; eapply normal_forms_encodes in H; (eauto 2).
    Qed.


  End Variables.

  Section Constants.
      Variable (n: ) (x: ).
      Hypothesis (Ex: encodes ( (F x)) n).

      Lemma forward_consts: n = 1 fst (constEQ x) snd (constEQ x).
      Proof using Ex.
        cbn; unfold funcomp; change shift with (add 1); rewrite ren_plus_combine; unfold var_zero; cbn.
        rewrite Ex; asimpl; now intros .
      Qed.


      Lemma backward_consts: fst (constEQ x) snd (constEQ x) n = 1.
      Proof using Ex.
        cbn.
        unfold funcomp; change shift with (add 1); rewrite ren_plus_combine; unfold var_zero; cbn.
        rewrite Ex. asimpl. intros ? % equiv_lam_elim % equiv_lam_elim.
        eapply enc_injective;[| |eauto]; (eauto 2).
      Qed.


  End Constants.

  Section Addition.
    Variable (m n p: ) (x y z: ).
    Hypothesis (Ex: encodes ( (F x)) m) (Ey: encodes ( (F y)) n) (Ez: encodes ( (F z)) p).

    Lemma forward_add: m + n = p fst (addEQ x y z) snd (addEQ x y z).
    Proof using Ex Ey Ez.
      intros ; cbn.
      unfold funcomp; change shift with (add 1); rewrite !ren_plus_combine; unfold var_zero; cbn.
      erewrite Ex, Ey, Ez; now simplify.
    Qed.


    Lemma backward_add:
       fst (addEQ x y z) snd (addEQ x y z) m + n = p.
    Proof using Ex Ey Ez.
      cbn.
      unfold funcomp; change shift with (add 1); rewrite !ren_plus_combine; unfold var_zero; cbn.
      rewrite Ex, Ey, Ez; simplify; intros H % equiv_lam_elim % equiv_lam_elim.
      rewrite enc_app in H.
      eapply enc_injective;[| |eauto]; (eauto 2).
    Qed.


  End Addition.

  Section Multiplication.

    Variable (m n p: ) (x y z: ).
    Hypothesis
      (Ex: encodes ( (F x)) m) (Ey: encodes ( (F y)) n) (Ez: encodes ( (F z)) p).

    Lemma forward_mul:
        m * n = p (G x y z) = T n m fst (mulEQ x y z) snd (mulEQ x y z).
    Proof using Ex Ey Ez.
      intros . unfold mulEQ; cbn. unfold funcomp.
      rewrite . change shift with (add 1); rewrite !ren_plus_combine; cbn [plus].
      rewrite Ex, Ey, Ez. subst p. eapply G_forward.
    Qed.


    Lemma backward_mult:
       fst (mulEQ x y z) snd (mulEQ x y z) m * n = p (G x y z) = T n m.
    Proof using N Ex Ey Ez.
      unfold mulEQ; cbn. unfold funcomp.
      change shift with (add 1); rewrite !ren_plus_combine; cbn [plus].
      rewrite Ex, Ey, Ez. intros H. apply G_iff in H; (eauto 2).
    Qed.


  End Multiplication.

End EquationEquivalences.

Section Forward.

  Variables (E: list deq).

  Definition gf n := enc n (var 0).

  Lemma tab_typing {X} n (f: exp X) g k Gamma:
    ( i, ⊢(n) f i : g i) ⊢₊(n) tab f k : tab g k.
  Proof.
    intros; induction k; cbn; (eauto 2).
    eapply orderlisttyping_app; (eauto 2).
  Qed.


  Lemma gf_typing Gamma n: ⊢(2) gf n : .
  Proof.
    econstructor. eapply enc_typing. eauto.
  Qed.


  Lemma GT_typing Gamma m n: ⊢(2) T n m : .
  Proof.
    do 3 econstructor. eapply lin_typing; [|eauto].
    rewrite repeated_tab. simplify.
    eapply tab_typing.
    intros; econstructor; cbn. econstructor.
    all: eauto using enc_typing.
  Qed.


  Definition enc_sol (sigma: ) (x: fin) :=
    match partition_F_G x with
    | inl (exist _ x _) gf ( x)
    | inr (exist _ (x,y,z) _) T ( y) ( x)
    end.

  Lemma enc_sol_encodes theta h: encodes (enc_sol (F h)) ( h).
  Proof.
    unfold enc_sol.
    destruct partition_F_G as [[x ?]|[((x,y),z) ?]].
    - eapply F_injective in e as .
      intros t . unfold gf. cbn; rewrite stepBeta; asimpl; (eauto 2).
    - exfalso; eapply disjoint_F_G; (eauto 2).
  Qed.


  Lemma enc_sol_T theta x y z: (enc_sol (G x y z)) = T ( y) ( x).
  Proof.
    unfold enc_sol.
    destruct partition_F_G as [[x' ?]|[((x',y'),z') ?]].
    - exfalso; eapply disjoint_F_G; (eauto 2).
    - now eapply G_injective in as ( & & ).
  Qed.


  Lemma typing_forward sigma:
    ( E) [] ⊩(2) enc_sol : Gamma__deq E.
  Proof.
    eauto; intros ????; unfold enc_sol.
    destruct partition_F_G as [[]|[((x & y) & z) ?]]; subst.
    + eapply nth_Gamma__deq_F in as . eapply gf_typing.
    + eapply nth_Gamma__deq_G in as . eapply GT_typing.
  Qed.


  Lemma H10_SU: E SOU ag 2 (H10_to_SOU E).
  Proof.
    intros [ H].
    exists []. exists (enc_sol ).
    split. now eapply typing_forward.
    intros s t; change s with (fst (s,t)) at 2; change t with (snd (s,t)) at 3.
    remember (s,t) as e. clear Heqe s t.
    intros H'; cbn in *. eapply in_Equations in H' as (d & ? & ?).
      destruct d; cbn in *; intuition idtac; subst.
      all: try eapply forward_add. all: try eapply forward_consts.
      all: try eapply forward_mul.
      all: try eapply forward_vars.
      all: try eapply enc_sol_encodes.
      all: eapply H in ; inv ; (eauto 2).
      apply enc_sol_T.
    Qed.


End Forward.

Section Backward.

  Definition decode_subst (sigma: fin exp ag) (N: x, normal ( x)) (x: fin) :=
    match dec_enc (N (F x)) with
    | inl (exist _ n _) n
    | inr _ 0
    end.

  Lemma decode_subst_encodes (sigma: fin exp ag) N x n:
    encodes ( (F x)) n decode_subst N x = n.
  Proof.
    intros H; unfold decode_subst.
    destruct dec_enc as [[m ]|].
    - specialize (H a id); asimpl in H. rewrite H in . eapply enc_injective in as []; (eauto 2).
    - exfalso. eapply . exists n. specialize (H a id); asimpl in H. now rewrite H.
  Qed.


  Lemma SU_H10 E: SOU ag 2 (H10_to_SOU E) E.
  Proof.
    rewrite SOU_NSOU; (eauto 2). intros ( & & T & EQ & N).
    exists (decode_subst N). intros e H; pose (Q := eqs e).
    assert ( p, p Q fst p snd p) as EQs; [|clear EQ].
    - intros [s t] G. eapply EQ.
      eapply in_Equations. eauto.
    - destruct e; econstructor; cbn in Q, EQs.
      all: specialize (EQs (varEQ x)) as EQx; mp EQx; intuition idtac; eapply backward_vars in EQx as [n EQx]; (eauto 2).
      2 - 3: specialize (EQs (varEQ y)) as EQy; mp EQy; intuition idtac; eapply backward_vars in EQy as [m EQy]; (eauto 2).
      2 - 3: specialize (EQs (varEQ z)) as EQz; mp EQz; intuition idtac; eapply backward_vars in EQz as [p EQz]; (eauto 2).
      all: repeat (erewrite decode_subst_encodes;[|eauto]).
      + eapply backward_consts; (eauto 4).
      + eapply backward_add; (eauto 1); eapply EQs; (eauto 5).
      + eapply backward_mult; (eauto 1); eapply EQs; intuition.
  Qed.

End Backward.

Lemma Goldfarb': SOU ag 2.
Proof.
  exists H10_to_SOU; intros x; split.
  - eapply H10_SU.
  - eapply SU_H10.
Qed.


Definition nileq: eq ag := ( a, a).
Definition conseqs e1 e2 :=
  match , with
  | ( , ), ( , ) ( g , g )
  | _, _ nileq
  end.

Notation foldeqs := (fold_right conseqs nileq).

Lemma conseqs_combine sigma s1 s2 t1 t2:
   ( ) ( ) ( ) ( ) ( g ) ( g ).
Proof.
  cbn; split.
  - intros [ ]; do 2 (eapply equiv_lam_elim in ; eapply equiv_lam_elim in ). now rewrite , .
  - intros ; do 2 eapply equiv_lam_elim in . Injection . Injection H. now rewrite , .
Qed.


Lemma foldeqs_lambda_lambda E:
   s t, foldeqs E = ( s, t).
Proof.
  induction E as [|[s t]]; cbn; (eauto 2).
  - do 2 eexists; reflexivity.
  - destruct IHE as (s'&t'&IH).
    do 2 (destruct s; try solve [do 2 eexists; reflexivity]).
    do 2 (destruct t; try solve [do 2 eexists; reflexivity]).
    rewrite IH; do 2 eexists; reflexivity.
Qed.


Lemma foldeqs_correct sigma E:
  ( s1 s2, (, ) E ( t, = t) ( t, = t))
  ( •₊ left_side E) ≡₊ ( •₊ right_side E)
  ( fst (foldeqs E) snd (foldeqs E)).
Proof.
  induction E as [| [s t]]; simplify.
  - cbn; firstorder.
  - intros H. apply id in H as H'. specialize (H' s t) as ([ ]&[ ]); lauto; subst.
    cbn. specialize (foldeqs_lambda_lambda E) as (u&v&H'); rewrite H' in *.
    mp IHE; (eauto 2). intros; eapply H; firstorder. cbn in IHE.
    cbn; unfold left_side, right_side in *.
    split; [intros [ ] % equiv_lstep_cons_inv| intros ]; do 2 eapply equiv_lam_elim in .
    + destruct IHE as [IHE _]; rewrite ; mp IHE; (eauto 1);
        do 2 eapply equiv_lam_elim in IHE; now rewrite IHE.
    + Injection . Injection . rewrite .
      destruct IHE as [_ ]; (eauto 2). now rewrite .
Qed.


Lemma foldeqs_typing Gamma E:
   ⊢₊₊(2) E : repeat ( ) (length E) (2) (foldeqs E) : .
Proof.
  remember (repeat ( ) (| E |)) as L. induction 1 in HeqL |- *; cbn.
  - repeat econstructor.
  - do 2 (destruct s; try solve [repeat econstructor]).
    do 2 (destruct t; try solve [repeat econstructor]).
    specialize (foldeqs_lambda_lambda E) as (u&v&H'); rewrite H' in *.
    cbn in HeqL. injection HeqL as . mp IHeqs_ordertyping; (eauto 2).
    inv . inv . inv H. inv . destruct IHeqs_ordertyping as [ ]. inv . inv .
    inv . inv . repeat econstructor; (eauto 2).
Qed.


Lemma Goldfarb : OU 2 ag.
Proof.
  unshelve eexists.
  - intros E. unshelve econstructor.
    exact (Gamma__deq E). 1: eapply fst. 2: eapply snd.
    1 - 2: eapply foldeqs, (H10_to_SOU E).
    exact ( ).
    all: eapply foldeqs_typing; cbn -[Gamma__deq].
    all: eapply H10_to_SOU.
  - intros E. split.
    + intros H % H10_SU. destruct H as ( & & & ).
      exists . exists . intuition. eapply foldeqs_correct, equiv_pointwise_eqs; (eauto 2).
      cbn. intros (d&&) % in_flat_map.
      destruct d; cbn in ; intuition.
      all: change with (fst (, )); rewrite ?H, ?.
      all: change with (snd (, )); rewrite ?H, ?.
      all: cbn; (eauto 2).
    + intros ( & & T & EQ). eapply SU_H10. exists , .
      intuition. eapply equiv_eqs_pointwise; (eauto 1); eapply foldeqs_correct; (eauto 2).
      cbn. intros (d&&) % in_flat_map.
      destruct d; cbn in ; intuition.
      all: change with (fst (, )); rewrite ?, ?.
      all: change with (snd (, )); rewrite ?, ?.
      all: cbn; (eauto 2).
Qed.