From Undecidability.TM Require Import TM Util.TM_facts.

Definition Assert (sig : Type) (n : ) : Type := tapes sig n Prop.


Definition Triple_Rel {sig : finType} {n : } {F : Type} (P : Assert sig n) (Q : F Assert sig n) : pRel sig F n :=
   tin '(yout, tout) P tin Q yout tout.

Inductive Triple {sig : finType} {n : } {F : Type} P (pM : pTM sig F n) Q :=
  TripleI : pM Triple_Rel P Q Triple P pM Q.

Lemma TripleE {sig : finType} {n : } {F : Type} P (pM : pTM sig F n) Q:
  Triple P pM Q pM Triple_Rel P Q.
Proof. now inversion 1. Qed.

Lemma Triple_iff {sig : finType} {n : } {F : Type} P (pM : pTM sig F n) Q:
  Triple P pM Q pM Triple_Rel P Q.
Proof. split;eauto using TripleE,TripleI. Qed.


Definition Triple_TRel {sig : finType} {n : } (P : Assert sig n) (k : ) : tRel sig n :=
   tin k' P tin k k'.

Inductive TripleT {sig : finType} {n : } {F : Type} (P : Assert sig n) (k : ) (pM : pTM sig F n) (Q : F Assert sig n) :=
 TripleTI : Triple P pM Q pM Triple_TRel P k TripleT P k pM Q.

 Lemma TripleTE {sig : finType} {n : } {F : Type} P k (pM : pTM sig F n) Q:
  TripleT P k pM Q Triple P pM Q pM Triple_TRel P k.
Proof. now inversion 1. Qed.

Lemma TripleT_iff {sig : finType} {n : } {F : Type} P k (pM : pTM sig F n) Q:
  TripleT P k pM Q (Triple P pM Q pM Triple_TRel P k).
Proof. split;firstorder eauto using TripleTE,TripleTI. Qed.

Lemma TripleT_Triple {sig : finType} {n : } {F : Type} P k (pM : pTM sig F n) Q :
  TripleT P k pM Q
  Triple P pM Q.
Proof. now intros ?%TripleTE. Qed.

#[export] Hint Resolve TripleT_Triple : core.

Lemma Triple_False {sig : finType} {n : } {F : Type} (pM : pTM sig F n) Q :
  Triple ( _ False) pM Q.
Proof. hnf. firstorder. Qed.

#[export] Hint Resolve Triple_False : core.

Lemma TripleT_False {sig : finType} {n : } {F : Type} k (pM : pTM sig F n) Q :
  TripleT ( _ False) k pM Q.
Proof. hnf. firstorder. Qed.

#[export] Hint Resolve TripleT_False : core.

Lemma Triple_True {sig : finType} {n : } {F : Type} (pM : pTM sig F n) P :
  Triple P pM ( _ _ True).
Proof. hnf. firstorder. Qed.

#[export] Hint Resolve Triple_True : core.


Lemma Realise_Triple {sig : finType} {n : } {F : Type} (P : Assert sig n) (pM : pTM sig F n) (Q : F Assert sig n) (R : pRel sig F n) :
  pM R
  ( tin yout tout, R tin (yout, tout) P tin Q yout tout)
  Triple P pM Q.
Proof.
  intros . apply TripleI. unfold Triple_Rel. eapply Realise_monotone.
  + eapply .
  + intros tin (yout, tout). firstorder.
Qed.


Lemma Realise_TripleT {sig : finType} {n : } {F : Type} (P : Assert sig n) (k : ) (pM : pTM sig F n) (Q : F Assert sig n) (R : pRel sig F n) (T : tRel sig n) :
  pM R
   pM T
  ( tin yout tout, R tin (yout, tout) P tin Q yout tout)
  ( tin k', P tin k k' T tin k')
  TripleT P k pM Q.
Proof.
  intros . split.
  - eapply Realise_Triple; eauto.
  - unfold Triple_TRel. eapply TerminatesIn_monotone; eauto.
    intros tin k' (?&Hk). firstorder.
Qed.


Lemma Triple_Realise {sig : finType} {n : } {F : Type} (P : Assert sig n) (pM : pTM sig F n) (Q : F Assert sig n) :
  Triple P pM Q
  pM ( tin '(yout,tout) P tin Q yout tout).
Proof.
  intros ?%TripleE. now unfold Triple_Rel in *.
Qed.


Lemma TripleT_Realise {sig : finType} {n : } {F : Type} (P : Assert sig n) k (pM : pTM sig F n) (Q : F Assert sig n) :
  TripleT P k pM Q
  pM ( tin '(yout,tout) P tin Q yout tout).
Proof.
  intros ?%TripleTE. now apply Triple_Realise .
Qed.


Lemma TripleT_TerminatesIn {sig : finType} {n : } {F : Type} (P : Assert sig n) (k : ) (pM : pTM sig F n) (Q : F Assert sig n):
  TripleT P k pM Q
   pM ( tin k' P tin k k').
Proof.
  intros ?%TripleTE. unfold Triple_TRel in *. tauto.
Qed.


Lemma RealiseIn_TripleT {sig : finType} {n : } {F : Type} (P : Assert sig n) k (pM : pTM sig F n) (Q : F Assert sig n) (R : pRel sig F n) :
  pM c(k) R
  ( tin yout tout, R tin (yout, tout) P tin Q yout tout)
  TripleT P k pM Q.
Proof.
  intros . split.
  - eapply Realise_Triple.
    + eapply RealiseIn_Realise; eauto.
    + firstorder.
  - unfold Triple_TRel. eapply TerminatesIn_monotone.
    + eapply Realise_total. apply .
    + now intros tin k' (H&Hk).
Qed.


Inductive Entails (sig : Type) (n : ) (P1 P2 : Assert sig n) :=
  EntailsI : ( tin, tin tin) Entails .

Lemma EntailsE (sig : Type) (n : ) (P1 P2 : Assert sig n):
Entails ( tin, tin tin).
Proof. now inversion 1. Qed.

Arguments Entails {_ _} _ _.

Lemma Entails_iff (sig : Type) (n : ) (P1 P2 : Assert sig n):
Entails ( tin, tin tin).
Proof. split;firstorder eauto using EntailsE,EntailsI. Qed.

#[export] Hint Resolve EntailsI : core.

Instance Entails_PO (sig : Type) (n : ): PreOrder (@Entails sig n).
Proof. split;hnf. all:setoid_rewrite Entails_iff. all:eauto. Qed.


Lemma Consequence (sig : finType) (n : ) (F : Type) (P1 P2 : Assert sig n) (Q1 Q2 : F Assert sig n) (pM : pTM sig F n) :
  Triple pM
  Entails
  ( y, Entails ( y) ( y))
  Triple pM .
Proof.
  intros %TripleE . apply TripleI. setoid_rewrite Entails_iff in . setoid_rewrite Entails_iff in .
  eapply Realise_monotone.
  { apply . }
  { intros tin (yout, tout) . cbn in . cbn. intros . auto.
  }
Qed.


Lemma asPointwise A X (R: A A Prop) f g: ( (x:X), R (f x) (g x)) pointwise_relation X R f g.
Proof. now cbv. Qed.

Instance Triple_Entails_Proper sig n F: Proper (Entails --> eq pointwise_relation F Entails Basics.impl) (@Triple sig n F).
Proof. repeat intro. subst. eapply Consequence;eauto. Qed.

Lemma Consequence_pre (sig : finType) (n : ) (F : Type) (P1 P2 : Assert sig n) (Q : F Assert sig n) (pM : pTM sig F n) :
  Triple pM Q
  Entails
  Triple pM Q.
Proof. now intros ? . Qed.

Lemma Consequence_post (sig : finType) (n : ) (F : Type) (P : Assert sig n) (Q1 Q2 : F Assert sig n) (pM : pTM sig F n) :
  Triple P pM
  ( y, Entails ( y) ( y))
  Triple P pM .
Proof. now intros ? %asPointwise. Qed.

Lemma ConsequenceT (sig : finType) (n : ) (F : Type) (P1 P2 : Assert sig n) (k1 k2 : ) (Q1 Q2 : F Assert sig n) (pM : pTM sig F n) :
  TripleT pM
  Entails
  ( y, Entails ( y) ( y))
  
  TripleT pM .
Proof.
  intros (&)%TripleTE %asPointwise .
  split. now rewrite , .
  - eapply TerminatesIn_monotone.
    + apply .
    + unfold Triple_TRel. intros tin k (H&H'). split. 2: . setoid_rewrite Entails_iff in . eauto.
Qed.


Instance TripleT_Entails_Proper sig n F: Proper (Entails --> le eq pointwise_relation F Entails Basics.impl) (@TripleT sig n F).
Proof. repeat intro. subst. eapply ConsequenceT;eauto. Qed.

Lemma ConsequenceT_post (sig : finType) (n : ) (F : Type) (P : Assert sig n) (k : ) (Q1 Q2 : F Assert sig n) (pM : pTM sig F n) :
  TripleT P k pM
  ( y, Entails ( y) ( y))
  TripleT P k pM .
Proof. intros. eapply ConsequenceT; eauto. Qed.

Lemma ConsequenceT_pre (sig : finType) (n : ) (F : Type) (P1 P2 : Assert sig n) (k1 k2 : ) (Q : F Assert sig n) (pM : pTM sig F n) :
  TripleT pM Q
  Entails
  
  TripleT pM Q.
Proof. intros. eapply ConsequenceT; eauto. Qed.



Lemma Triple_exists_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : X Assert sig n) (Q : F Assert sig n) :
  ( (x : X), Triple (P x) pM Q)
  Triple ( tin x : X, P x tin) pM (Q).
Proof. setoid_rewrite Triple_iff. unfold Triple_Rel. firstorder. Qed.

Lemma Triple_exists_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : Assert sig n) (Q : X F Assert sig n) :
  ( (x : X), Triple P pM (Q x))
  Triple P pM ( yout tout x : X, Q x yout tout).
Proof. setoid_rewrite Triple_iff. unfold Triple_Rel. firstorder. Qed.

Lemma Triple_forall_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : X Assert sig n) (Q : F Assert sig n) :
  ( (x : X), Triple (P x) pM Q)
  Triple ( tin x : X, P x tin) pM (Q).
Proof. setoid_rewrite Triple_iff. unfold Triple_Rel. firstorder. Qed.

Lemma Triple_forall_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : Assert sig n) (Q : X F Assert sig n) :
  ( (x : X), Triple P pM (Q x))
  Triple P pM ( yout tout x : X, Q x yout tout).
Proof. setoid_rewrite Triple_iff. unfold Triple_Rel. firstorder. Qed.

Lemma Triple_eta_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q : F Assert sig n) :
  Triple P pM Q
  Triple ( tin P tin) pM Q.
Proof. exact id. Qed.

Lemma Triple_eta_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q : F Assert sig n) :
  Triple P pM Q
  Triple P pM ( yout Q yout).
Proof. exact id. Qed.

Lemma Triple_eta_con2 {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q : F Assert sig n) :
  Triple P pM Q
  Triple P pM ( yout tout Q yout tout).
Proof. exact id. Qed.

Lemma Triple_and_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P1 : Assert sig n) (P2 : Prop) (Q : F Assert sig n) :
  ( Triple pM Q)
  Triple ( tin tin) pM Q.
Proof. setoid_rewrite Triple_iff. unfold Triple_Rel. firstorder. Qed.

Lemma Triple_and_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q1 : F Assert sig n) (Q2 : Prop) :
  Triple P pM
  
  Triple P pM ( yout tout yout tout).
Proof. setoid_rewrite Triple_iff. unfold Triple_Rel. firstorder. Qed.


Lemma TripleT_exists_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : X Assert sig n) (Q : F Assert sig n) k :
  ( (x : X), TripleT (P x) k pM Q)
  TripleT ( tin x : X, P x tin) k pM (Q).
Proof. setoid_rewrite TripleT_iff;setoid_rewrite Triple_iff. unfold Triple_Rel,Triple_TRel. firstorder. Qed.

Lemma TripleT_exists_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : Assert sig n) (Q : X F Assert sig n) k :
  ( (x : X), TripleT P k pM (Q x))
  TripleT P k pM ( yout tout x : X, Q x yout tout).
Proof. setoid_rewrite TripleT_iff;setoid_rewrite Triple_iff. unfold Triple_Rel,Triple_TRel. firstorder. Qed.

Lemma TripleT_forall_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : X Assert sig n) (Q : F Assert sig n) k :
  ( (x : X), TripleT (P x) k pM Q)
  TripleT ( tin x : X, P x tin) k pM (Q).
Proof. setoid_rewrite TripleT_iff;setoid_rewrite Triple_iff. unfold Triple_Rel,Triple_TRel. firstorder. Qed.

Lemma TripleT_forall_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (X : Type) (P : Assert sig n) (Q : X F Assert sig n) k {inhX: inhabitedC X} :
  ( (x : X), TripleT P k pM (Q x))
  TripleT P k pM ( yout tout x : X, Q x yout tout).
Proof.
  intros H; split.
  - apply Triple_forall_con. apply H.
  - apply H. apply default.
Qed.


Lemma TripleT_eta_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q : F Assert sig n) k :
  TripleT P k pM Q
  TripleT ( tin P tin) k pM Q.
Proof. exact id. Qed.

Lemma TripleT_eta_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q : F Assert sig n) k :
  TripleT P k pM Q
  TripleT P k pM ( yout Q yout).
Proof. exact id. Qed.

Lemma TripleT_eta_con2 {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q : F Assert sig n) k :
  TripleT P k pM Q
  TripleT P k pM ( yout tout Q yout tout).
Proof. exact id. Qed.

Lemma TripleT_and_pre {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P1 : Assert sig n) (P2 : Prop) (Q : F Assert sig n) k :
  ( TripleT k pM Q)
  TripleT ( tin tin) k pM Q.
Proof. setoid_rewrite TripleT_iff;setoid_rewrite Triple_iff. unfold Triple_Rel,Triple_TRel. firstorder. Qed.

Lemma TripleT_and_con {sig : finType} {n : } {F : Type} (pM : pTM sig F n)
      (P : Assert sig n) (Q1 : F Assert sig n) (Q2 : Prop) k :
  TripleT P k pM
  
  TripleT P k pM ( yout tout yout tout).
Proof. setoid_rewrite TripleT_iff;setoid_rewrite Triple_iff. unfold Triple_Rel,Triple_TRel. firstorder. Qed.