Require Import ADuplicateSymb.
Require Import AGraph.
Require Import APolyInt_MA.
Require Import ATrs.
Require Import List.
Require Import LogicUtil.
Require Import MonotonePolynom.
Require Import Polynom.
Require Import SN.
Require Import VecUtil.

Open Scope nat_scope.
(* termination problem *)

Module M.
  Inductive symb : Type :=
  | _plus__plus__1 : symb
  | a : symb
  | b : symb
  | rev : symb.
End M.

Lemma eq_symb_dec : forall f g : M.symb, {f=g}+{~f=g}.

Proof.
decide equality.
Defined.

Open Scope nat_scope.
Definition ar (s : M.symb) : nat :=
  match s with
  | M._plus__plus__1 => 2
  | M.a => 0
  | M.b => 0
  | M.rev => 1
  end.

Definition s0 := ASignature.mkSignature ar eq_symb_dec.
Definition s0_p := s0.
Definition V0 := @ATerm.Var s0.
Definition F0 := @ATerm.Fun s0.
Definition R0 := @ATrs.mkRule s0.

Module S0.
  Definition _plus__plus__1 x2 x1 := F0 M._plus__plus__1 (Vcons x2 (Vcons x1 Vnil)).
  Definition a := F0 M.a Vnil.
  Definition b := F0 M.b Vnil.
  Definition rev x1 := F0 M.rev (Vcons x1 Vnil).
End S0.

Definition E :=
   @nil (@ATrs.rule s0).

Definition R :=
   R0 (S0.rev S0.a)
      S0.a
:: R0 (S0.rev S0.b)
      S0.b
:: R0 (S0.rev (S0._plus__plus__1 (V0 0) (V0 1)))
      (S0._plus__plus__1 (S0.rev (V0 1)) (S0.rev (V0 0)))
:: R0 (S0.rev (S0._plus__plus__1 (V0 0) (V0 0)))
      (S0.rev (V0 0))
:: @nil (@ATrs.rule s0).

Definition rel := ATrs.red_mod E R.

(* symbol marking *)

Definition s1 := dup_sig s0.
Definition s1_p := s0.
Definition V1 := @ATerm.Var s1.
Definition F1 := @ATerm.Fun s1.
Definition R1 := @ATrs.mkRule s1.

Module S1.
  Definition h_plus__plus__1 x2 x1 := F1 (hd_symb s1_p M._plus__plus__1) (Vcons x2 (Vcons x1 Vnil)).
  Definition _plus__plus__1 x2 x1 := F1 (int_symb s1_p M._plus__plus__1) (Vcons x2 (Vcons x1 Vnil)).
  Definition ha := F1 (hd_symb s1_p M.a) Vnil.
  Definition a := F1 (int_symb s1_p M.a) Vnil.
  Definition hb := F1 (hd_symb s1_p M.b) Vnil.
  Definition b := F1 (int_symb s1_p M.b) Vnil.
  Definition hrev x1 := F1 (hd_symb s1_p M.rev) (Vcons x1 Vnil).
  Definition rev x1 := F1 (int_symb s1_p M.rev) (Vcons x1 Vnil).
End S1.

(* polynomial interpretation 1 *)

Module PIS1 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.rev) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    | (int_symb M.rev) =>
         (3%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.a) =>
         nil
    | (int_symb M.a) =>
         nil
    | (hd_symb M.b) =>
         nil
    | (int_symb M.b) =>
         (3%Z, Vnil)
      :: nil
    | (hd_symb M._plus__plus__1) =>
         nil
    | (int_symb M._plus__plus__1) =>
         (2%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (3%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (3%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    end.

  Lemma trsInt_wm : forall f, pweak_monotone (trsInt f).
  Proof.
    pmonotone.
  Qed.

End PIS1.

Module PI1 := PolyInt PIS1.

(* termination proof *)

Lemma termination : WF rel.

Proof.
unfold rel.
dp_trans.
mark.
PI1.prove_termination.
termination_trivial.
Qed.