From Undecidability.TM.Code Require Import ProgrammingTools.
From Undecidability Require Import TM.Code.CaseNat.


Definition Subtract_Rel : pRel (sigNat^+) bool 2 :=
     tin '(b, tout)
     (x y : ),
        tin [@] x
         tin [@] y
         tout [@] x - y
         isVoid tout [@]
         b = (y <=? x).

Definition Subtract :pTM (sigNat^+) bool 2:=
    While (
        If(CaseNat @ [||])
          (If (CaseNat @ [||])
              (Return Nop None)
              (Return (Reset _ @ [||]) (Some false)) )
          (Return (Reset _ @ [||]) (Some true))
    ).

Lemma Subtract_Realise : Subtract Subtract_Rel.
Proof.
  eapply Realise_monotone.
  {
    unfold Subtract. TM_Correct. all: (try (notypeclasses refine (@Reset_Realise _ _ _ _ _);shelve)).
  }
  {
    apply WhileInduction.
    - intros tin b tout [(t0_&&[|[t1_ ]]) |?] x y;TMSimp;[ | ].
      + modpon H. modpon . destruct y. easy. destruct x. 2:easy.
        modpon . rewrite Nat.sub_0_l. cbn ["<=?"]. repeat split. contains_ext. isVoid_mono.
      + modpon H. destruct y. 2:easy.
        modpon . rewrite Nat.sub_0_r. cbn ["<=?"]. repeat split. contains_ext. isVoid_mono.
    - intros tin tmid tout b [(t0_&&[|[t1_ ]]) |?] H' x y;TMSimp;[].
      modpon H. modpon . destruct x,y;try contradiction;[].
      modpon H'. rewrite Nat.sub_succ. cbn ["<=?"]. easy.
  }
Qed.


From Undecidability Require Import UpToC Hoare.

Theorem Subtract_SpecT:
{ f : UpToC ( x x + 1) & (x y : ),
  TripleT ≃≃([],[|Contains _ x;Contains _ y|])
  (f (y)) Subtract
  ( b ≃≃([b = (y <=? x)],
    ([|Contains _ (x-y);Void|])))}.
Proof.
  eexists_UpToC f. intros x y.
  unfold Subtract.
  eapply ConsequenceT.
  refine (While_SpecTReg (PRE := '(x,y) (_,_))
  (INV := '(x,y) b ([b = match y with 0 Some true | _ match x with 0 Some false | _ None end end]
    ,match y with 0 [|Contains _ x;Void|] | S y' [|Contains _ (pred x);match x with 0 Void | _ Contains _ y' end|] end))
    (POST := '(x,y) b (_,_)) (f__loop := '(x,y) _)
    (f__step := '(x,y) match y with 0 _ : | S y' _ end ) _ _ ((x,y))); clear x y; intros (x,y).
  3:{ cbn. reflexivity. }
  3:{ cbn. reflexivity. }
  3:{ cbn. reflexivity. }
  -cbn. hstep.
    +now hsteps_cbn.
    +refine (_ : TripleT _ match y with 0 _ | S _ match x with 0 _ | _ _ end end _ _).
      destruct y;hintros Hy. nia. hstep.
      *now hsteps_cbn.
      * cbn. destruct x;hintros ?. nia. hsteps_cbn. tspec_ext.
      *cbn. hintros . cbn. hsteps_cbn. tspec_ext.
      *cbn. intros b Hb. destruct b,x. 1,4:exfalso;nia. all:reflexivity.
    +cbn. hintros . hsteps_cbn. tspec_ext.
    +cbn - ["+" "*"]. intros b Hb.
     destruct y,b. 1,4:exfalso;nia. all:reflexivity.
  - cbn. split.
    + intros b Hb. destruct y.
      { inv Hb. cbn - ["+"]. rewrite Nat.sub_0_r. split. now tspec_ext. shelve. }
      split. 2:shelve.
      destruct x. all:inv Hb. cbn. tspec_ext.
    + destruct y. easy. destruct x. easy.
      eexists (_,_). split. cbn. reflexivity.
      split. shelve. cbn. reflexivity.
  - Unshelve.
    4,5,6:unfold Reset_steps;rewrite ?Encode_nat_hasSize;ring_simplify.
    [f]:exact ( y y * 30 + CaseNat_steps + 30).
    2:exact 0.
    1:now subst f;smpl_upToC_solve.
    all:unfold f.
    2:destruct x.
    all:nia.
Qed.


Ltac hstep_Subtract :=
  lazymatch goal with
  | [ |- TripleT ?P ?k Subtract ?Q ] notypeclasses refine ( Subtract_SpecT _ _);shelve
  end.

Smpl Add hstep_Subtract : hstep_Spec.