Require Import List Arith Lia.

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

From Undecidability.StackMachines.BSM
  Require Import bsm_defs.

From Undecidability.MinskyMachines.MM
  Require Import mm_defs mm_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 "I '/BSM/' s -1> t" := (bsm_sss I s t) (at level 70, no associativity).
Local Notation "P '/BSM/' s -+> t" := (sss_progress (@bsm_sss _) P s t) (at level 70, no associativity).
Local Notation "P '/BSM/' s > t" := (sss_compute (@bsm_sss _) P s t) (at level 70, no associativity).
Local Notation "P '/BSM/' s ~~> t" := (sss_output (@bsm_sss _) P s t) (at level 70, no associativity).
Local Notation "P '/BSM/' s ↓" := (sss_terminates (@bsm_sss _) P s)(at level 70, no associativity).

Local Notation "P '/MM/' s -+> t" := (sss_progress (@mm_sss _) P s t) (at level 70, no associativity).
Local Notation "P '/MM/' s > t" := (sss_compute (@mm_sss _) P s t) (at level 70, no associativity).
Local Notation "P '/MM/' s '~~>' t" := (sss_output (@mm_sss _) P s t) (at level 70, no associativity).
Local Notation "P '/MM/' s ↓" := (sss_terminates (@mm_sss _) P s)(at level 70, no associativity).

Section simulator.

  Ltac dest x y := destruct (pos_eq_dec x y) as [ | ]; [ subst x | ]; rew vec.

  Variables (m : ).


  Let n := 2+m.
  Local Definition : pos n := pos0.
  Local Definition : pos n := pos1.
  Local Definition reg p: pos n := pos_nxt (pos_nxt p).

  Local Lemma Hv12 : . Proof. discriminate. Qed.

  Local Lemma Hvr1 : p, reg p . Proof. discriminate. Qed.
  Local Lemma Hvr2 : p, reg p . Proof. discriminate. Qed.

  Let Hreg : p q, reg p = reg q p = q.
  Proof. intros; do 2 apply pos_nxt_inj; apply H. Qed.

  Definition bsm_state_enc (v : vec (list bool) m) w :=
            w#> = 0
          w#> = 0
          p, w#>(reg p) = stack_enc (v#>p).


  Definition bsm_instr_compile lnk i ii :=
    match ii with
      | PUSH s Zero mm_push_Zero (reg s) (lnk i)
      | PUSH s One mm_push_One (reg s) (lnk i)
      | POP s j k mm_pop (reg s) (lnk i) (lnk j) (lnk (1+i)) (lnk k)
    end.

  Definition bsm_instr_compile_length (ii : bsm_instr m) :=
    match ii with
      | PUSH _ Zero 7
      | PUSH _ One 8
      | POP _ _ _ 16
    end.

  Fact bsm_instr_compile_length_eq lnk i ii : length (bsm_instr_compile lnk i ii) = bsm_instr_compile_length ii.
  Proof. destruct ii as [ | ? [] ]; simpl; auto. Qed.

  Fact bsm_instr_compile_length_geq ii : 1 bsm_instr_compile_length ii.
  Proof. destruct ii as [ | ? [] ]; simpl; auto; . Qed.

  Hint Resolve bsm_instr_compile_length_eq bsm_instr_compile_length_geq : core.


  Lemma bsm_instr_compile_sound : instruction_compiler_sound bsm_instr_compile (@bsm_sss _) (@mm_sss _) bsm_state_enc.
  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 p j k v Hv
                      | i p j k v ll Hll
                      | i p j k v ll Hll
                      | i p [] v
                      ]; simpl; intros H; generalize H; intros ( & & ).

    + exists ; split; auto.
      apply mm_pop_void_progress; auto using Hv12, Hvr1, Hvr2.
      rewrite , Hv; auto.

    + exists ([(stack_enc ll)/reg p]); repeat split; auto; rew vec.
      * apply mm_pop_Zero_progress; auto using Hv12, Hvr1, Hvr2.
        rewrite , Hll; auto.
      * intros q; dest p q.
        assert (reg p reg q); rew vec.

    + exists ([(stack_enc ll)/reg p]); repeat split; auto; rew vec.
      * apply mm_pop_One_progress; auto using Hv12, Hvr1, Hvr2.
        rewrite , Hll; auto.
      * intros q; dest p q.
        assert (reg p reg q); rew vec.

    + exists ([(stack_enc (One::v#>p))/reg p]); repeat split; auto; rew vec.
      rewrite ; apply mm_push_One_progress; auto using Hv12, Hvr1, Hvr2.
      intros q; dest p q.
      assert (reg p reg q); rew vec.

    + exists ([(stack_enc (Zero::v#>p))/reg p]); repeat split; auto; rew vec.
      rewrite ; apply mm_push_Zero_progress; auto using Hv12, Hvr1, Hvr2.
      intros q; dest p q.
      assert (reg p reg q); rew vec.
  Qed.


  Hint Resolve bsm_instr_compile_sound : core.

  Section bsm_sim.

    Variable (iP : ) (cP : list (bsm_instr m)).

    Local Definition lnk_Q_pair := @gen_compiler_correction _ _ _ _ bsm_instr_compile_length_eq _ _ _ _ (@bsm_sss_total' _)
                     (@mm_sss_fun _) _ bsm_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 "⋈" := bsm_state_enc (at level 70, no associativity).

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

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

    Variable v : vec (list bool) m.

    Local Definition w := 0##0##vec_map stack_enc v.

    Let w_prop : bsm_state_enc v w.
    Proof.
      red; unfold w, , ; repeat split; rew vec.
      intros p; unfold reg; simpl.
      rewrite vec_pos_map; trivial.
    Qed.



    Local Lemma Q_spec1 : (iP,cP) /BSM/ (iP,v) w', Q /MM/ (1,w) ~~> (code_end Q, w') w'#> = 0 w'#> = 0.
    Proof.
      intros ((,) & ).
      destruct HQ1 with (1 := conj w_prop ) as (w' & & ).
      rewrite (proj2 (proj2 Hlnk) ), (proj1 (proj2 Hlnk)).
      * exists w'; split; auto; red in ; tauto.
      * apply .
    Qed.


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


    Local Definition bsm_mm_sim := snd Q.

    Theorem bsm_mm_sim_spec : (iP,cP) /BSM/ (iP,v) (1,bsm_mm_sim) /MM/ (1,w) .
    Proof.
      rewrite (proj1 Hlnk) at 1.
      rewrite surjective_pairing.
      split; auto using Q_spec2.
      intros H.
      destruct (Q_spec1 H) as (w' & & _).
      exists (code_end Q, w'); auto.
    Qed.


    Local Definition iE := code_end Q.


    Local Definition cN := mm_nullify iE (map ( p pos_nxt (pos_nxt p)) (pos_list m)).
    Local Definition cE := cN DEC 0 :: nil.

    Local Lemma E_spec w' : w'#> = 0 w'#> = 0 (iE,cE) /MM/ (iE,w') -+> (0,vec_zero).
    Proof.
      intros .
      unfold cE.
      apply sss_compute_progress_trans with (length cN+iE,vec_zero).
      + apply subcode_sss_compute with (P := (iE,cN)); auto.
        apply mm_nullify_compute; auto.
        * intros p Hp.
          apply in_map_iff in Hp.
          destruct Hp as (x & & ); subst; discriminate.
        * intros p Hp.
          apply in_map_iff in Hp.
          destruct Hp as (x & & ); subst; apply vec_zero_spec.
        * intros p Hp.
          unfold n, , in *; simpl in p.
          pos_inv p; auto.
          pos_inv p; auto.
          destruct Hp; apply in_map_iff; exists p; split; auto.
          apply pos_list_prop.
      + apply subcode_sss_progress with (P := (length cN+iE,DEC 0::nil)); auto.
        mm sss DEC 0 with 0.
        apply subcode_refl.
        mm sss stop.
    Qed.


    Definition bsm_mm := snd Q cE.

    Let cQ_sim : Q <sc (1,bsm_mm).
    Proof.
      pose proof Hlnk as Hlnk.
      unfold bsm_mm; destruct Q as (iQ,cQ); simpl in Hlnk.
      simpl snd; rewrite (proj1 Hlnk); auto.
    Qed.


    Let cE_sim : (iE,cE) <sc (1,bsm_mm).
    Proof.
      unfold iE, bsm_mm; subcode_tac; solve list eq.
      rewrite (proj1 Hlnk); auto.
    Qed.



    Theorem bsm_mm_spec : (iP,cP) /BSM/ (iP,v) (1,bsm_mm) /MM/ (1,w) ~~> (0,vec_zero).
    Proof.
      split.
      * intros .
        apply Q_spec1 in .
        destruct as (w' & ( & ) & & ).
        split.
        2: simpl; .
        apply sss_compute_trans with ( := (iE,w')).
        revert ; apply subcode_sss_compute; auto.
        apply sss_progress_compute.
        generalize (E_spec _ ); apply subcode_sss_progress; auto.
      * intros .
        apply Q_spec2.
        apply subcode_sss_terminates with (1 := cQ_sim).
        exists (0,vec_zero); auto.
    Qed.


  End bsm_sim.

End simulator.

Theorem bsm_mm_compiler_1 n i (P : list (bsm_instr n)) :
  { Q : list (mm_instr (pos (2+n))) | v, (i,P) /BSM/ (i,v) (1,Q) /MM/ (1,0##0##vec_map stack_enc v) }.
Proof. exists (bsm_mm_sim i P); apply bsm_mm_sim_spec. Qed.

Theorem bsm_mm_compiler_2 n i (P : list (bsm_instr n)) :
  { Q : list (mm_instr (pos (2+n))) | v, (i,P) /BSM/ (i,v) (1,Q) /MM/ (1,0##0##vec_map stack_enc v) ~~> (0,vec_zero) }.
Proof. exists (bsm_mm i P); apply bsm_mm_spec. Qed.