Require Import List Arith Lia.

From Undecidability.Shared.Libs.DLW
  Require Import utils_tac utils_list utils_nat gcd rel_iter pos vec.

Require Import Undecidability.FRACTRAN.FRACTRAN.

Set Implicit Arguments.

Section fractran_utils.

  Implicit Type l : list (*).

  Fact fractran_step_nil_inv x y : nil /F/ x y False.
  Proof. split; inversion 1. Qed.

  Fact fractran_step_cons_inv p q l x y :
     (p,q)::l /F/ x y q*y = p*x divides q (p*x) l /F/ x y.
  Proof.
    split.
    + inversion 1; auto.
    + intros [ H | ( & ) ]; [ constructor 1 | constructor 2 ]; auto.
  Qed.


  Fact mul_pos_inj_l q x y : q 0 q*x = q*y x = y.
  Proof.
    intros .
    destruct q; try .
    apply le_antisym; apply mult_S_le_reg_l with q; .
  Qed.


  Lemma fractran_step_inv P x y :
            P /F/ x y
          l p q r, P = l(p,q)::r
                          ( u v, In (u,v) l divides v (u*x))
                          q*y=p*x.
  Proof.
    induction 1 as [ p q P x y H | u v P x y ].
    + exists nil, p, q, P; simpl; tauto.
    + destruct as (l & p & q & r & & & ).
      exists ((u,v)::l), p, q, r; simpl; msplit 2; auto.
      intros ? ? [ H | H ]; auto; inversion H; subst; auto.
  Qed.



  Lemma fractran_step_fun l x y1 y2 :
           fractran_regular l
         l /F/ x
         l /F/ x = .
  Proof.
    intros ; revert .
    induction 1 as [ p q l x | p q l x ];
      intros ; rewrite fractran_step_cons_inv in ;
      destruct as [ | ( & ) ].
    + apply Forall_cons_inv, proj1 in .
      simpl in .
      revert ; rewrite .
      apply mul_pos_inj_l; auto.
    + destruct ; exists ; rewrite ; ring.
    + destruct ; exists ; rewrite ; ring.
    + apply Forall_cons_inv, proj2 in ; auto.
  Qed.



  Lemma fractran_step_bound l :
          fractran_regular l
        { k | x y, l /F/ x y y k*x }.
  Proof.
    unfold fractran_regular.
    induction l as [ | (p,q) l IHl ].
    + intros _; exists 1; auto.
      intros ? ? H; exfalso; revert H; apply fractran_step_nil_inv.
    + intros H; rewrite Forall_cons_inv in H; simpl in H.
      destruct H as (Hq & H).
      destruct (IHl H) as (k & ).
      exists (p+k).
      intros x y Hxy.
      apply fractran_step_cons_inv in Hxy.
      destruct Hxy as [ Hxy | (_ & Hxy) ].
      * rewrite Nat.mul_add_distr_r, Hxy.
        destruct q; simpl; try .
      * apply le_trans with (1 := _ _ Hxy).
        apply mult_le_compat; .
  Qed.


  Fact fractan_stop_nil_inv x : fractran_stop nil x True.
  Proof.
    split; try tauto; intros _ z; rewrite fractran_step_nil_inv; tauto.
  Qed.


  Fact fractan_stop_cons_inv p q l x :
       fractran_stop ((p,q)::l) x divides q (p*x) fractran_stop l x.
  Proof.
    split.
      * intros H.
        assert ( divides q (p*x)) as H'.
        { intros (z & Hz); apply (H z); constructor; rewrite Hz; ring. }
        split; auto.
        intros z Hz; apply (H z); constructor 2; auto.
      * intros ( & ) z Hz.
        apply fractran_step_cons_inv in Hz.
        destruct Hz as [ | ( & ) ].
        - apply ; exists z; rewrite ; ring.
        - apply ( _ ).
  Qed.


  Fact fractran_step_dec l x : { y | l /F/ x y } + { fractran_stop l x }.
  Proof.
    revert x; induction l as [ | (a,b) l IH ]; intros x.
    + right; rewrite fractan_stop_nil_inv; auto.
    + destruct x as [ | x ].
      { left; exists 0; constructor 1; ring. }
      destruct a as [ | a ].
      { left; exists 0; constructor 1; ring. }
      destruct b as [ | b ].
      * assert ( divides 0 (S a*S x)) as C.
        { intros C; apply divides_0_inv, mult_is_O in C; . }
        destruct (IH (S x)) as [ (y & Hy) | Hx ].
        - left; exists y; constructor 2; auto.
        - right; rewrite fractan_stop_cons_inv; split; auto.
      * generalize (div_rem_spec1 (S a*S x) (S b)).
        case_eq (rem (S a * S x) (S b)).
        - intros ; left; exists (div (S a*S x) (S b)); constructor 1; rewrite at 2; ring.
        - intros d Hd _.
          assert ( divides (S b) (S a*S x)) as C.
          { rewrite divides_rem_eq, Hd; discriminate. }
          destruct (IH (S x)) as [ (y & Hy) | Hx ].
          ++ left; exists y; constructor 2; auto.
          ++ right; rewrite fractan_stop_cons_inv; split; auto.
  Qed.



  Let remove_zero_den l := filter ( c if eq_nat_dec (snd c) 0 then false else true) l.

  Let remove_zero_den_Forall l : fractran_regular (remove_zero_den l).
  Proof.
    unfold fractran_regular.
    induction l as [ | (p,q) ]; simpl; auto.
    destruct (eq_nat_dec q 0); auto.
  Qed.


  Section zero_cases.

    Fact fractran_zero_num_step l :
            Exists ( c fst c = 0) l
          x, y, l /F/ x y.
    Proof.
      induction 1 as [ (p,q) l Hl | (p,q) l Hl IHl ]; simpl in Hl.
      + intros x; exists 0; subst; constructor; .
      + intros x.
        destruct (divides_dec (p*x) q) as [ (y & Hy) | C ].
        * exists y; constructor; rewrite Hy, mult_comm; auto.
        * destruct (IHl x) as (y & Hy); exists y; constructor 2; auto.
    Qed.


    Lemma FRACTRAN_HALTING_zero_num l x :
             Exists ( c fst c = 0) l
           FRACTRAN_HALTING (l,x) False.
    Proof.
      intros H; split; try tauto.
      intros (y & _ & ).
      destruct fractran_zero_num_step with (1 := H) (x := y) as (z & Hz).
      apply with (1 := Hz).
    Qed.


    Fact fractran_step_head_not_zero p q l y : q 0 (p,q)::l /F/ 0 y y = 0.
    Proof.
      intros .
      apply fractran_step_cons_inv in .
      destruct as [ | ( & _) ].
      + rewrite Nat.mul_0_r in ; apply mult_is_O in ; .
      + destruct ; exists 0; ring.
    Qed.


    Fact fractran_rt_head_not_zero p q l n y : q 0 fractran_steps ((p,q)::l) n 0 y y = 0.
    Proof.
      intros .
      induction n as [ | n IHn ].
      + simpl; auto.
      + intros (a & & ).
        apply fractran_step_head_not_zero in ; subst; auto.
    Qed.


    Lemma FRACTRAN_HALTING_on_zero_first_no_zero_den p q l : q 0 FRACTRAN_HALTING ((p,q)::l,0) False.
    Proof.
      intros Hq; split; try tauto.
      intros (y & (k & ) & ).
      apply fractran_rt_head_not_zero in ; auto.
      subst y; apply ( 0); constructor 1; ring.
    Qed.


    Fact fractran_step_no_zero_den l x y : fractran_regular l l /F/ x y x = 0 y = 0.
    Proof.
      unfold fractran_regular.
      intros ; revert .
      induction 1 as [ p q l x y H | p q l x y ]; rewrite Forall_cons_inv; simpl; intros ( & ) ?; subst.
      + rewrite Nat.mul_0_r in H; apply mult_is_O in H; .
      + destruct ; exists 0; ring.
    Qed.


    Fact fractran_step_no_zero_num l : Forall ( c fst c 0) l x y, l /F/ x y y = 0 x = 0.
    Proof.
      intros x y ; revert .
      induction 1 as [ p q l x y | p q l x y ]; intros ; rewrite Forall_cons_inv in ; simpl in ; destruct as ( & ); auto.
      intros; subst y; rewrite Nat.mul_0_r in ; symmetry in ; apply mult_is_O in ; .
    Qed.


    Fact fractran_rt_no_zero_den l n y : fractran_regular l fractran_steps l n 0 y y = 0.
    Proof.
      intros H; induction n as [ | n IHn ].
      + simpl; auto.
      + intros (x & & ).
        apply fractran_step_no_zero_den with (1 := H) in ; subst; auto.
    Qed.


    Fact fractran_rt_no_zero_num l : Forall ( c fst c 0) l n x, fractran_steps l n x 0 x = 0.
    Proof.
      intros H; induction n as [ | n IHn ]; intros x; simpl; auto.
      intros (y & & ).
      apply IHn in .
      revert ; apply fractran_step_no_zero_num; auto.
    Qed.


    Fact fractran_zero_num l x : Exists ( c fst c = 0) l y, l /F/ x y.
    Proof.
      induction 1 as [ (p,q) l | (p,q) l Hl IHl ]; simpl in *.
      + exists 0; constructor 1; subst; ring.
      + destruct (divides_dec (p*x) q) as [ (y & Hy) | H ].
        * exists y; constructor 1; rewrite Hy; ring.
        * destruct IHl as (y & Hy); exists y; constructor 2; auto.
    Qed.


    Corollary FRACTRAN_HALTING_0_num l x : Exists ( c fst c = 0) l FRACTRAN_HALTING (l,x) False.
    Proof.
      intros ; split; try tauto; intros (y & & ).
      destruct (fractran_zero_num y ) as (z & ?).
      apply ( z); auto.
    Qed.


    Lemma fractran_step_zero l :
             Forall ( c fst c 0) l
           x y, x 0 l /F/ x y
                      remove_zero_den l /F/ x y.
    Proof.
      induction 1 as [ | (p,q) l ]; intros x y Hxy; simpl in *.
      * split; simpl; inversion 1.
      * simpl; destruct (eq_nat_dec q 0) as [ Hq | Hq ].
        - rewrite ; auto; subst; split; intros H.
          + apply fractran_step_cons_inv in H; destruct H as [ H | ( & ) ]; auto.
            simpl in H; symmetry in H; apply mult_is_O in H; .
          + constructor 2; auto; intros H'.
            apply divides_0_inv, mult_is_O in H'; .
        - split; intros H.
          + apply fractran_step_cons_inv in H; destruct H as [ H | ( & ) ]; auto.
            -- constructor 1; auto.
            -- constructor 2; auto; apply ; auto.
          + apply fractran_step_cons_inv in H; destruct H as [ H | ( & ) ]; auto.
            -- constructor 1; auto.
            -- constructor 2; auto; apply ; auto.
    Qed.


    Lemma fractran_rt_no_zero_den_0_0 l :
           l nil fractran_regular l fractran_step l 0 0.
    Proof.
      destruct l as [ | (p,q) l ]; intros H.
      + destruct ; auto.
      + clear ; apply Forall_cons_inv, proj1 in H.
        constructor 1; ring.
    Qed.


    Corollary FRACTRAN_HALTING_l_0_no_zero_den l :
            l nil fractran_regular l FRACTRAN_HALTING (l,0) False.
    Proof.
      intros ; split; try tauto.
      generalize (fractran_rt_no_zero_den_0_0 ); intros .
      intros (y & (n & ) & ).
      apply fractran_rt_no_zero_den in ; auto; subst.
      apply ( 0); auto.
    Qed.


    Lemma FRACTRAN_HALTING_nil_x x : FRACTRAN_HALTING (nil,x) True.
    Proof.
      split; try tauto; intros _.
      exists x; split.
      + exists 0; simpl; auto.
      + inversion 1.
    Qed.


    Lemma FRACTRAN_HALTING_l_1_no_zero_den l x :
           l nil
         x 0
         Forall ( c fst c 0) l
         FRACTRAN_HALTING (l,x) FRACTRAN_HALTING (remove_zero_den l,x).
    Proof.
      intros .
      generalize (fractran_step_zero ); intros .
      assert ( n y, fractran_steps l n x y y 0) as .
      { intros n y H ?; subst; apply ; revert H; apply fractran_rt_no_zero_num; auto. }
      assert ( n y, fractran_steps (remove_zero_den l) n x y y 0) as .
      { intros n y H ?; subst; apply ; revert H; apply fractran_rt_no_zero_num; auto; apply Forall_filter; auto. }
      assert ( n y, rel_iter (fractran_step l) n x y rel_iter (fractran_step (remove_zero_den l)) n x y) as .
      { induction n as [ | n IHn ]; intros y.
        + simpl; try tauto.
        + do 2 rewrite rel_iter_S; split.
          * intros (a & & ); exists a; split.
            - apply IHn; auto.
            - apply ; auto.
              apply with (1 := ).
          * intros (a & & ); exists a; split.
            - apply IHn; auto.
            - apply ; auto.
              apply with (1 := ). }
      split; intros (y & (n & ) & ); exists y; split.
      + exists n; apply ; auto.
      + intros z Hz; apply ( z); revert Hz; apply , with n; auto.
      + exists n; apply ; auto.
      + intros z Hz; apply ( z); revert Hz; apply , with n; auto.
    Qed.


    Lemma FRACTRAN_HALTING_hard p l :
           p 0
         FRACTRAN_HALTING ((p,0)::l,0)
        x, x 0 FRACTRAN_HALTING ((p,0)::l,x).
    Proof.
      intros H; split.
      + intros (y & (n & ) & ).
        assert (y 0) as Hy.
        { intro; subst; apply ( 0); constructor; auto. }
        unfold fractran_steps in ; rewrite rel_iter_sequence in .
        destruct as (f & & & ).
        destruct (first_non_zero f n) as (i & & & ); subst; auto.
        exists (f (i+1)); split; auto.
        exists (f n); split; auto.
        exists (n-i-1); red; rewrite rel_iter_sequence.
        exists ( j f (j+i+1)); split; auto.
        split; [ f_equal; | ].
        intros j Hj; apply ; .
      + intros (x & & (y & (n & Hn) & )).
        exists y; split; auto.
        exists (S n), x; split; auto.
        constructor 1; ring.
    Qed.


  End zero_cases.


  Fact FRACTAN_cases ll : { Exists ( c fst c = 0) ll }
                         + { Forall ( c snd c 0) ll }
                         + { p : & { mm | Forall ( c fst c 0) ll Exists ( c snd c = 0) ll
                                            ll = (p,0):: mm p 0 } }
                         + { p : & { q : & { mm | Forall ( c fst c 0) ll Exists ( c snd c = 0) ll
                                           ll = (p,q):: mm q 0 p 0 } } }.
  Proof.
    destruct (Forall_Exists_dec ( c : * snd c 0)) with (l := ll) as [ ? | ]; auto.
    { intros (p, q); destruct (eq_nat_dec q 0); simpl; subst; [ right | left]; . }
    assert (Exists ( c snd c = 0) ll) as .
    { revert ; induction 1; [ constructor 1 | constructor 2 ]; auto; . }
    clear .
    destruct (Forall_Exists_dec ( c : * fst c 0)) with (l := ll) as [ | ].
    { intros (p, q); destruct (eq_nat_dec p 0); simpl; subst; [ right | left ]; . }
    2: { do 3 left; clear ; revert ; induction 1; [ constructor 1 | constructor 2 ]; auto; . }
    case_eq ll.
    { intro; subst; exfalso; inversion . }
    intros (p,q) mm Hll.
    destruct (eq_nat_dec p 0) as [ Hp | Hp ].
    { subst; rewrite Forall_cons_inv in ; simpl in ; . }
    destruct q.
    + left; right; exists p, mm; subst; auto.
    + right; exists p, (S q), mm; subst; repeat (split; auto).
  Qed.


End fractran_utils.