Require Import List Arith Lia.

From Undecidability.Shared.Libs.DLW
  Require Import utils pos vec subcode sss compiler_correction.

From Undecidability.MinskyMachines.MMA
  Require Import mma_defs mma_utils.

Set Implicit Arguments.

Set Default Proof Using "Type".

Tactic Notation "rew" "length" := autorewrite with length_db.

Local Notation "e #> x" := (vec_pos e x).
Local Notation "e [ v / x ]" := (vec_change e x v).

Local Notation "P // s -+> t" := (sss_progress (@mma_sss _) P s t).
Local Notation "P // s > t" := (sss_compute (@mma_sss _) P s t).
Local Notation "P // s ~~> t" := (sss_output (@mma_sss _) P s t).
Local Notation "P // s ↓" := (sss_terminates (@mma_sss _) P s).


Section mma_sim.

  Variables (n : ).


  Definition mma_instr_compile lnk (_ : ) (ii : mm_instr (pos n)) :=
    match ii with
      | INC k INC k :: nil
      | DEC k j DEC k (lnk j) :: nil
    end.

  Definition mma_instr_compile_length (ii : mm_instr (pos n)) := 1.

  Fact mma_instr_compile_length_eq lnk i ii : length (mma_instr_compile lnk i ii) = mma_instr_compile_length ii.
  Proof. destruct ii; simpl; auto. Qed.


  Fact mma_instr_compile_length_geq ii : 1 mma_instr_compile_length ii.
  Proof. cbv; . Qed.

  Hint Resolve mma_instr_compile_length_eq mma_instr_compile_length_geq : core.
  Hint Resolve subcode_refl : core.


  Lemma mma_instr_compile_sound : instruction_compiler_sound mma_instr_compile (@mma_sss _) (@mma_sss _) eq.
  Proof.
    intros lnk I H; revert H .
    change with (snd (,)) at 2.
    change with (fst (,)) at 2 3 4 6 7 8.
    change with (snd (,)) at 2.
    change with (fst (,)) at 2.
    generalize (,) (,); clear .
    induction 1
      as [ i x k | i x k v H | i x k v u H ];
      simpl; intros .
    + exists ( [(S (#>x))/x]); split; auto.
      mma sss INC with x.
      mma sss stop; now f_equal.
    + exists ; split; auto.
      mma sss DEC 0 with x (lnk k).
      mma sss stop; now f_equal.
    + exists ([u/x]); split; auto.
      mma sss DEC S with x (lnk k) u.
      mma sss stop.
  Qed.


  Hint Resolve mma_instr_compile_sound : core.

  Section mma_sim.

    Variable (iP : ) (cP : list (mm_instr (pos n))).

    Local Definition lnk_Q_pair := @gen_compiler_correction _ _ _ _ mma_instr_compile_length_eq _ _ _ _ (@mma_sss_total_ni _)
                     (@mma_sss_fun _) _ mma_instr_compile_sound (iP,cP) 1.

    Local Definition lnk := lnk_Q_pair.
    Local Definition Q := proj1_sig ( lnk_Q_pair).

    Local Lemma Hlnk : fst Q = 1 lnk iP = 1 i, out_code i (iP,cP) lnk i = code_end Q.
    Proof.
      repeat split; apply (proj2_sig ( lnk_Q_pair)).
    Qed.


    Infix "⋈" := (@eq (vec n)) (at level 70, no associativity).

    Local Lemma HQ1 : i1 v1 w1 i2 v2, (iP,cP) // (,) ~~> (,)
                     w2, Q // (lnk ,) ~~> (lnk ,).
    Proof. apply (proj2_sig ( lnk_Q_pair)). Qed.

    Local Lemma HQ1' i1 v1 i2 v2 : (iP,cP) // (,) ~~> (,)
                         Q // (lnk ,) ~~> (lnk ,).
    Proof.
      intros H; destruct (@HQ1 ) as ( & & ?); auto.
    Qed.


    Local Lemma HQ2 : i1 v1 w1 j2 w2, Q // (lnk ,) ~~> (,)
                     i2 v2, (iP,cP) // (,) ~~> (,) = lnk .
    Proof. apply (proj2_sig ( lnk_Q_pair)). Qed.

    Local Lemma HQ2' i1 v1 j2 v2 : Q // (lnk ,) ~~> (,)
                i2, (iP,cP) // (,) ~~> (,) = lnk .
    Proof.
      intros H; destruct (@HQ2 ) as ( & ? & & ? & ); auto.
      exists ; auto.
    Qed.


    Variable v : vec n.


    Local Lemma Q_spec1 : (iP,cP) // (iP,v) w', Q // (1,v) ~~> (code_end Q, w').
    Proof.
      intros ((,) & ).
      exists .
      rewrite (proj2 (proj2 Hlnk) ), (proj1 (proj2 Hlnk)); auto using HQ1'.
      destruct ; auto.
    Qed.


    Local Lemma Q_spec2 : Q // (1,v) (iP,cP) // (iP,v) .
    Proof.
      intros ((j,) & ).
      rewrite (proj1 (proj2 Hlnk)) in .
      destruct HQ2' with (1 := ) as ( & ? & ).
      exists (,); auto.
    Qed.


    Definition mma_sim := snd Q.
    Definition mma_sim_end := code_end Q.

    Theorem mma_sim_spec : ((iP,cP) // (iP,v) w', (1,mma_sim) // (1,v) ~~> (1+length mma_sim, w'))
                         ((1,mma_sim) // (1,v) (iP,cP) // (iP,v) ).
    Proof.
      replace (1+length mma_sim) with (code_end Q).
      replace (1,mma_sim) with Q.
      + split; (auto using Q_spec1, Q_spec2).
      + rewrite (surjective_pairing Q); f_equal; auto.
        apply (proj1 Hlnk).
      + unfold code_end; f_equal.
        apply (proj1 Hlnk).
    Qed.


  End mma_sim.

End mma_sim.

Section mma2_simul.


  Variable (iP : ) (cP : list (mm_instr (pos 2))).

  Let Q := mma_sim iP cP.
  Local Definition eQ := 1+length Q.

  Local Definition cN : list (mm_instr (pos 2)) := mma_null pos0 eQ mma_null pos1 (1+eQ) mma_jump 0 pos0.

  Local Definition := @mma_null_length 2 pos0 eQ.
  Local Definition := @mma_null_length 2 pos1 (1+eQ).

  Let N_spec v : (eQ,cN) // (eQ,v) -+> (0,vec_zero).
  Proof.
    unfold cN.
    apply sss_progress_trans with (1+eQ,v[0/pos0]).
    1: { apply subcode_sss_progress with (P := (eQ,mma_null pos0 eQ)); auto.
         apply mma_null_progress; auto. }
    apply sss_progress_trans with (2+eQ,(v[0/pos0])[0/pos1]).
    1: { pose proof as .
         apply subcode_sss_progress with (P := (1+eQ,mma_null pos1 (1+eQ))); auto.
         apply mma_null_progress; auto. }
    replace ((v[0/pos0])[0/pos1]) with (@vec_zero 2).
    2: { apply vec_pos_ext; intros p.
         repeat (invert pos p; rew vec). }
    pose proof as . pose proof as .
    apply subcode_sss_progress with (P := (2+eQ,mma_jump 0 pos0)); auto.
    apply mma_jump_progress.
  Qed.


  Definition mma2_simul := Q cN.

  Let cQ_sim : (1,Q) <sc (1,mma2_simul).
  Proof. unfold mma2_simul; auto. Qed.

  Let cN_sim : (eQ,cN) <sc (1,mma2_simul).
  Proof. unfold mma2_simul, eQ; auto. Qed.

  Theorem mma2_simul_spec v : (iP,cP) // (iP,v) (1,mma2_simul) // (1,v) ~~> (0,vec_zero).
  Proof.
    split.
    * intros H; apply mma_sim_spec in H; revert H.
      intros (w & Hw & _).
      split; try (simpl; ).
      apply sss_compute_trans with (eQ,w).
      + revert Hw; apply subcode_sss_compute; auto.
      + apply sss_progress_compute.
        generalize (N_spec w).
        apply subcode_sss_progress; auto.
    * intros .
      apply mma_sim_spec; fold Q.
      apply subcode_sss_terminates with (1 := cQ_sim).
      now exists (0,vec_zero).
  Qed.


End mma2_simul.


Theorem mma2_simulator i (P : list (mm_instr (pos 2))) :
  { Q | v, (i,P) // (i,v) (1,Q) // (1,v) ~~> (0,vec_zero) }.
Proof. exists (mma2_simul i P); apply mma2_simul_spec. Qed.