From Undecidability.TM Require Import ProgrammingTools.
From Undecidability Require Import CaseNat CaseList List.App.

From Undecidability.L.Complexity Require Import UpToCNary.

Local Arguments skipn { A } !n !l.
Local Arguments plus : simpl never.
Local Arguments mult : simpl never.

Local Arguments Encode_list : simpl never.
Local Arguments Encode_nat : simpl never.

Section bla.
  Import FinTypes.

  Polymorphic Lemma leUpToC_clean domain (f F: Rtuple domain ):
    Fun' f c Fun' F
     f c F.
  Proof.
    prove_nary id.
  Qed.

End bla.

Module ConcatRepeat.
  Section concatRepeat.

    Variable sig sigX : finType.
    Variable (X : Type) (cX : codable sigX X).
    Variable ( : Retract (sigList sigX) sig) ( : Retract sigNat sig).


    Definition Rel__step : pRel sig^+ (option unit) 3 :=
       tin '(yout,tout)
         (cs : list X) (n:) ( xs : list X) ,
          tin[@] cs
          tin[@] n
          tin[@] xs
          match yout, n with
          | (Some tt), 0
            tout[@] cs
            isVoid tout[@]
            tout[@] xs
          | None, S n
            tout[@] cs
            tout[@] n
            tout[@] csxs
          | _, _ False
          end.

    Definition M__step : pTM sig^+ (option unit) 3 :=
      If (LiftTapes (ChangeAlphabet CaseNat _) [||])
         (Return (LiftTapes (ChangeAlphabet (App' _) _) [|; |]) None)
         (Return (LiftTapes (Reset _) [||]) (Some tt)).

    Lemma Realise__step : M__step Rel__step .
    Proof.
      eapply Realise_monotone.
      {unfold M__step. TM_Correct. now apply App'_Realise. }
      intros t (y,t') H. cbn.
      intros cs n xs Hcs Hn Hxs. cbn in H.
      destruct H as [H|H].
      -destruct H as (?&(HCase&HCaseRem)&&&HApp&HAppRem). clear .
       modpon HCase. destruct n. easy. TMSimp. modpon HApp.
       repeat eapply conj. 1-3:now contains_ext.
      -destruct H as (?&(HCase&HCaseRem)&&&HReset&HResetRem). clear .
       modpon HCase. destruct n. 2:easy. TMSimp. modpon HReset.
       repeat eapply conj. 1,3:now contains_ext. now isVoid_mono.
    Qed.


    Definition Rel : pRel sig^+ unit 3 :=
      ignoreParam (
           tin tout
             (cs : list X) (n:) ( xs : list X) ,
              tin[@] cs
              tin[@] n
              tin[@] xs
              tout[@] cs
              isVoid tout[@]
              tout[@] concat (repeat cs n)xs
        ).

    Definition M : pTM sig^+ unit 3 := While M__step.

    Lemma Realise : M Rel .
    Proof.
      eapply Realise_monotone.
      {unfold M. TM_Correct. apply Realise__step. }
      eapply WhileInduction;intros;hnf;intros cs n xs Hcs Hn Hxs.
      -hnf in HLastStep. modpon HLastStep. destruct yout,n. 2:easy.
       TMSimp. easy.
      -hnf in HStar. modpon HStar. destruct n. easy. TMSimp.
       modpon HLastStep. replace (cs concat (repeat cs n)) with (concat (repeat cs (n+1))).
       +rewrite repeat_add_app, concat_app;cbn;autorewrite with list;cbn. easy.
       +rewrite Nat.add_comm. reflexivity.
    Qed.


    Definition Ter__step time : tRel sig^+ 3 :=
       tin k (cs : list X) (n:) ( xs : list X) ,
          tin[@] cs tin[@] n tin[@] xs time (n,size cs) k.

    Lemma _Terminates__step : { time : UpToC ( '(n,l) (if n =? 0 then 0 else l) + 1)
                                     & M__step Ter__step time}.
    Proof.
      eexists_UpToC time. [time]:refine ( '(n,l) if n =? 0 then _ else _).
      eapply TerminatesIn_monotone.
      { unfold M__step. TM_Correct. now apply App'_Terminates with (cX:=cX). }
      {
        intros tin k (cs&n&xs&Hcs&Hn&Hxs&Hk). cbn. eexists. eexists (if n =? 0 then _ else _).
        infTer 3. rewrite Hk.
        2:{ clear Hk time.
            intros tout b (HCaseNat&HRem). modpon HCaseNat.
            destruct b. all:destruct (Nat.eqb_spec n 0) as [|]. all:try now (destruct n;exfalso).
            2:{rewrite in HCaseNat. exists 0. split. 2:reflexivity. simpl_surject. TMSimp. contains_ext. }
            unfold ;cbn. eexists cs, xs. TMSimp. split. 2:split. 3:reflexivity.
            all:simpl_surject. all:contains_ext.
        }
        unfold CaseNat_steps,Reset_steps.
        unfold time.
        destruct Nat.eqb. reflexivity.
        rewrite (correct__leUpToC (App'_steps_nice _)).
        set (size cs) as l. reflexivity.
      }
      unfold time.
      apply leUpToC_finCases with
          (cases := (b:bool) match b return (if b then * else ) _ with
                                false l (0,l)
                              | true '(n,l) (S n, l)
                              end).
      { intros [[ |n] l]. now (exists false; eauto). now (exists true;eexists (_,_);eauto). }
      intros []. nary apply leUpToC_clean.
      2:{ cbn [fst snd]. rewrite Nat.eqb_refl. smpl_upToC_solve. }
      cbn [Nat.eqb]. smpl_upToC_solve.
    Qed.


    Definition Terminates__step := _Terminates__step.

    Definition Ter time : tRel sig^+ 3 :=
       tin k (cs : list X) (n:) ( xs : list X) ,
          tin[@] cs tin[@] n tin[@] xs time (size cs,n) k.

    Lemma _Terminates : {time : UpToC ( '(l,n) n * l + 1)
                                & M Ter time }.
    Proof.
      evar ( : ). evar ( : ).
      exists_UpToC ( '(l,n) n * l * + ). unfold M.
      eapply TerminatesIn_monotone.
      -TM_Correct. now apply Realise__step. now apply Terminates__step.
      -apply WhileCoInduction. unfold Ter.
       intros tin k (cs&n&xs&Hcs&Hn&Hxs&Hk).
       eexists. unfold Ter__step. split.
       { exists cs,n,xs. repeat simple apply conj. 1-3:eassumption. rewrite UpToC_le. reflexivity. }
       unfold Rel__step. intros ymid tmid Hstep. modpon Hstep.
       destruct ymid as [[]| ],n. all:try now exfalso.
       +rewrite Nat.eqb_refl. rewrite Hk, Nat.mul_0_l, Nat.mul_1_r,!Nat.add_0_l. unfold . reflexivity.
       +TMSimp.
        eexists. split.
        *repeat simple eapply ex_intro. repeat simple apply conj. 1-3:contains_ext. reflexivity.
        * rewrite Hk. ring_simplify. set (c' := c__leUpToC).
          assert (size cs > 0). 1:{rewrite Encode_list_hasSize. destruct cs;cbn;nia. }
          replace with (2 + 2*c'). 2:now unfold c',. nia.
      - smpl_upToC_solve.
    Qed.


    Definition Terminates := _Terminates.

  End concatRepeat.

End ConcatRepeat.