Require Import ADPUnif.
Require Import ADecomp.
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 :=
  | _0_1 : symb
  | _if_2 : symb
  | active : symb
  | add : symb
  | and : symb
  | cons : symb
  | false : symb
  | first : symb
  | from : symb
  | mark : symb
  | nil : symb
  | s : symb
  | true : 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._0_1 => 0
  | M._if_2 => 3
  | M.active => 1
  | M.add => 2
  | M.and => 2
  | M.cons => 2
  | M.false => 0
  | M.first => 2
  | M.from => 1
  | M.mark => 1
  | M.nil => 0
  | M.s => 1
  | M.true => 0
  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 _0_1 := F0 M._0_1 Vnil.
  Definition _if_2 x3 x2 x1 := F0 M._if_2 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))).
  Definition active x1 := F0 M.active (Vcons x1 Vnil).
  Definition add x2 x1 := F0 M.add (Vcons x2 (Vcons x1 Vnil)).
  Definition and x2 x1 := F0 M.and (Vcons x2 (Vcons x1 Vnil)).
  Definition cons x2 x1 := F0 M.cons (Vcons x2 (Vcons x1 Vnil)).
  Definition false := F0 M.false Vnil.
  Definition first x2 x1 := F0 M.first (Vcons x2 (Vcons x1 Vnil)).
  Definition from x1 := F0 M.from (Vcons x1 Vnil).
  Definition mark x1 := F0 M.mark (Vcons x1 Vnil).
  Definition nil := F0 M.nil Vnil.
  Definition s x1 := F0 M.s (Vcons x1 Vnil).
  Definition true := F0 M.true Vnil.
End S0.

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

Definition R :=
   R0 (S0.active (S0.and S0.true (V0 0)))
      (S0.mark (V0 0))
:: R0 (S0.active (S0.and S0.false (V0 0)))
      (S0.mark S0.false)
:: R0 (S0.active (S0._if_2 S0.true (V0 0) (V0 1)))
      (S0.mark (V0 0))
:: R0 (S0.active (S0._if_2 S0.false (V0 0) (V0 1)))
      (S0.mark (V0 1))
:: R0 (S0.active (S0.add S0._0_1 (V0 0)))
      (S0.mark (V0 0))
:: R0 (S0.active (S0.add (S0.s (V0 0)) (V0 1)))
      (S0.mark (S0.s (S0.add (V0 0) (V0 1))))
:: R0 (S0.active (S0.first S0._0_1 (V0 0)))
      (S0.mark S0.nil)
:: R0 (S0.active (S0.first (S0.s (V0 0)) (S0.cons (V0 1) (V0 2))))
      (S0.mark (S0.cons (V0 1) (S0.first (V0 0) (V0 2))))
:: R0 (S0.active (S0.from (V0 0)))
      (S0.mark (S0.cons (V0 0) (S0.from (S0.s (V0 0)))))
:: R0 (S0.mark (S0.and (V0 0) (V0 1)))
      (S0.active (S0.and (S0.mark (V0 0)) (V0 1)))
:: R0 (S0.mark S0.true)
      (S0.active S0.true)
:: R0 (S0.mark S0.false)
      (S0.active S0.false)
:: R0 (S0.mark (S0._if_2 (V0 0) (V0 1) (V0 2)))
      (S0.active (S0._if_2 (S0.mark (V0 0)) (V0 1) (V0 2)))
:: R0 (S0.mark (S0.add (V0 0) (V0 1)))
      (S0.active (S0.add (S0.mark (V0 0)) (V0 1)))
:: R0 (S0.mark S0._0_1)
      (S0.active S0._0_1)
:: R0 (S0.mark (S0.s (V0 0)))
      (S0.active (S0.s (V0 0)))
:: R0 (S0.mark (S0.first (V0 0) (V0 1)))
      (S0.active (S0.first (S0.mark (V0 0)) (S0.mark (V0 1))))
:: R0 (S0.mark S0.nil)
      (S0.active S0.nil)
:: R0 (S0.mark (S0.cons (V0 0) (V0 1)))
      (S0.active (S0.cons (V0 0) (V0 1)))
:: R0 (S0.mark (S0.from (V0 0)))
      (S0.active (S0.from (V0 0)))
:: R0 (S0.and (S0.mark (V0 0)) (V0 1))
      (S0.and (V0 0) (V0 1))
:: R0 (S0.and (V0 0) (S0.mark (V0 1)))
      (S0.and (V0 0) (V0 1))
:: R0 (S0.and (S0.active (V0 0)) (V0 1))
      (S0.and (V0 0) (V0 1))
:: R0 (S0.and (V0 0) (S0.active (V0 1)))
      (S0.and (V0 0) (V0 1))
:: R0 (S0._if_2 (S0.mark (V0 0)) (V0 1) (V0 2))
      (S0._if_2 (V0 0) (V0 1) (V0 2))
:: R0 (S0._if_2 (V0 0) (S0.mark (V0 1)) (V0 2))
      (S0._if_2 (V0 0) (V0 1) (V0 2))
:: R0 (S0._if_2 (V0 0) (V0 1) (S0.mark (V0 2)))
      (S0._if_2 (V0 0) (V0 1) (V0 2))
:: R0 (S0._if_2 (S0.active (V0 0)) (V0 1) (V0 2))
      (S0._if_2 (V0 0) (V0 1) (V0 2))
:: R0 (S0._if_2 (V0 0) (S0.active (V0 1)) (V0 2))
      (S0._if_2 (V0 0) (V0 1) (V0 2))
:: R0 (S0._if_2 (V0 0) (V0 1) (S0.active (V0 2)))
      (S0._if_2 (V0 0) (V0 1) (V0 2))
:: R0 (S0.add (S0.mark (V0 0)) (V0 1))
      (S0.add (V0 0) (V0 1))
:: R0 (S0.add (V0 0) (S0.mark (V0 1)))
      (S0.add (V0 0) (V0 1))
:: R0 (S0.add (S0.active (V0 0)) (V0 1))
      (S0.add (V0 0) (V0 1))
:: R0 (S0.add (V0 0) (S0.active (V0 1)))
      (S0.add (V0 0) (V0 1))
:: R0 (S0.s (S0.mark (V0 0)))
      (S0.s (V0 0))
:: R0 (S0.s (S0.active (V0 0)))
      (S0.s (V0 0))
:: R0 (S0.first (S0.mark (V0 0)) (V0 1))
      (S0.first (V0 0) (V0 1))
:: R0 (S0.first (V0 0) (S0.mark (V0 1)))
      (S0.first (V0 0) (V0 1))
:: R0 (S0.first (S0.active (V0 0)) (V0 1))
      (S0.first (V0 0) (V0 1))
:: R0 (S0.first (V0 0) (S0.active (V0 1)))
      (S0.first (V0 0) (V0 1))
:: R0 (S0.cons (S0.mark (V0 0)) (V0 1))
      (S0.cons (V0 0) (V0 1))
:: R0 (S0.cons (V0 0) (S0.mark (V0 1)))
      (S0.cons (V0 0) (V0 1))
:: R0 (S0.cons (S0.active (V0 0)) (V0 1))
      (S0.cons (V0 0) (V0 1))
:: R0 (S0.cons (V0 0) (S0.active (V0 1)))
      (S0.cons (V0 0) (V0 1))
:: R0 (S0.from (S0.mark (V0 0)))
      (S0.from (V0 0))
:: R0 (S0.from (S0.active (V0 0)))
      (S0.from (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_0_1 := F1 (hd_symb s1_p M._0_1) Vnil.
  Definition _0_1 := F1 (int_symb s1_p M._0_1) Vnil.
  Definition h_if_2 x3 x2 x1 := F1 (hd_symb s1_p M._if_2) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))).
  Definition _if_2 x3 x2 x1 := F1 (int_symb s1_p M._if_2) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))).
  Definition hactive x1 := F1 (hd_symb s1_p M.active) (Vcons x1 Vnil).
  Definition active x1 := F1 (int_symb s1_p M.active) (Vcons x1 Vnil).
  Definition hadd x2 x1 := F1 (hd_symb s1_p M.add) (Vcons x2 (Vcons x1 Vnil)).
  Definition add x2 x1 := F1 (int_symb s1_p M.add) (Vcons x2 (Vcons x1 Vnil)).
  Definition hand x2 x1 := F1 (hd_symb s1_p M.and) (Vcons x2 (Vcons x1 Vnil)).
  Definition and x2 x1 := F1 (int_symb s1_p M.and) (Vcons x2 (Vcons x1 Vnil)).
  Definition hcons x2 x1 := F1 (hd_symb s1_p M.cons) (Vcons x2 (Vcons x1 Vnil)).
  Definition cons x2 x1 := F1 (int_symb s1_p M.cons) (Vcons x2 (Vcons x1 Vnil)).
  Definition hfalse := F1 (hd_symb s1_p M.false) Vnil.
  Definition false := F1 (int_symb s1_p M.false) Vnil.
  Definition hfirst x2 x1 := F1 (hd_symb s1_p M.first) (Vcons x2 (Vcons x1 Vnil)).
  Definition first x2 x1 := F1 (int_symb s1_p M.first) (Vcons x2 (Vcons x1 Vnil)).
  Definition hfrom x1 := F1 (hd_symb s1_p M.from) (Vcons x1 Vnil).
  Definition from x1 := F1 (int_symb s1_p M.from) (Vcons x1 Vnil).
  Definition hmark x1 := F1 (hd_symb s1_p M.mark) (Vcons x1 Vnil).
  Definition mark x1 := F1 (int_symb s1_p M.mark) (Vcons x1 Vnil).
  Definition hnil := F1 (hd_symb s1_p M.nil) Vnil.
  Definition nil := F1 (int_symb s1_p M.nil) Vnil.
  Definition hs x1 := F1 (hd_symb s1_p M.s) (Vcons x1 Vnil).
  Definition s x1 := F1 (int_symb s1_p M.s) (Vcons x1 Vnil).
  Definition htrue := F1 (hd_symb s1_p M.true) Vnil.
  Definition true := F1 (int_symb s1_p M.true) Vnil.
End S1.

(* graph decomposition 1 *)

Definition cs1 : list (list (@ATrs.rule s1)) :=

   (  R1 (S1.hfrom (S1.active (V1 0)))
         (S1.hfrom (V1 0))
   :: R1 (S1.hfrom (S1.mark (V1 0)))
         (S1.hfrom (V1 0))
   :: nil)

:: (  R1 (S1.hcons (V1 0) (S1.mark (V1 1)))
         (S1.hcons (V1 0) (V1 1))
   :: R1 (S1.hcons (S1.mark (V1 0)) (V1 1))
         (S1.hcons (V1 0) (V1 1))
   :: R1 (S1.hcons (S1.active (V1 0)) (V1 1))
         (S1.hcons (V1 0) (V1 1))
   :: R1 (S1.hcons (V1 0) (S1.active (V1 1)))
         (S1.hcons (V1 0) (V1 1))
   :: nil)

:: (  R1 (S1.hfirst (V1 0) (S1.mark (V1 1)))
         (S1.hfirst (V1 0) (V1 1))
   :: R1 (S1.hfirst (S1.mark (V1 0)) (V1 1))
         (S1.hfirst (V1 0) (V1 1))
   :: R1 (S1.hfirst (S1.active (V1 0)) (V1 1))
         (S1.hfirst (V1 0) (V1 1))
   :: R1 (S1.hfirst (V1 0) (S1.active (V1 1)))
         (S1.hfirst (V1 0) (V1 1))
   :: nil)

:: (  R1 (S1.hs (S1.active (V1 0)))
         (S1.hs (V1 0))
   :: R1 (S1.hs (S1.mark (V1 0)))
         (S1.hs (V1 0))
   :: nil)

:: (  R1 (S1.hadd (V1 0) (S1.mark (V1 1)))
         (S1.hadd (V1 0) (V1 1))
   :: R1 (S1.hadd (S1.mark (V1 0)) (V1 1))
         (S1.hadd (V1 0) (V1 1))
   :: R1 (S1.hadd (S1.active (V1 0)) (V1 1))
         (S1.hadd (V1 0) (V1 1))
   :: R1 (S1.hadd (V1 0) (S1.active (V1 1)))
         (S1.hadd (V1 0) (V1 1))
   :: nil)

:: (  R1 (S1.h_if_2 (V1 0) (S1.mark (V1 1)) (V1 2))
         (S1.h_if_2 (V1 0) (V1 1) (V1 2))
   :: R1 (S1.h_if_2 (S1.mark (V1 0)) (V1 1) (V1 2))
         (S1.h_if_2 (V1 0) (V1 1) (V1 2))
   :: R1 (S1.h_if_2 (V1 0) (V1 1) (S1.mark (V1 2)))
         (S1.h_if_2 (V1 0) (V1 1) (V1 2))
   :: R1 (S1.h_if_2 (S1.active (V1 0)) (V1 1) (V1 2))
         (S1.h_if_2 (V1 0) (V1 1) (V1 2))
   :: R1 (S1.h_if_2 (V1 0) (S1.active (V1 1)) (V1 2))
         (S1.h_if_2 (V1 0) (V1 1) (V1 2))
   :: R1 (S1.h_if_2 (V1 0) (V1 1) (S1.active (V1 2)))
         (S1.h_if_2 (V1 0) (V1 1) (V1 2))
   :: nil)

:: (  R1 (S1.hand (V1 0) (S1.mark (V1 1)))
         (S1.hand (V1 0) (V1 1))
   :: R1 (S1.hand (S1.mark (V1 0)) (V1 1))
         (S1.hand (V1 0) (V1 1))
   :: R1 (S1.hand (S1.active (V1 0)) (V1 1))
         (S1.hand (V1 0) (V1 1))
   :: R1 (S1.hand (V1 0) (S1.active (V1 1)))
         (S1.hand (V1 0) (V1 1))
   :: nil)

:: (  R1 (S1.hmark (S1.from (V1 0)))
         (S1.hfrom (V1 0))
   :: nil)

:: (  R1 (S1.hmark (S1.cons (V1 0) (V1 1)))
         (S1.hcons (V1 0) (V1 1))
   :: nil)

:: (  R1 (S1.hmark (S1.first (V1 0) (V1 1)))
         (S1.hfirst (S1.mark (V1 0)) (S1.mark (V1 1)))
   :: nil)

:: (  R1 (S1.hmark (S1.s (V1 0)))
         (S1.hs (V1 0))
   :: nil)

:: (  R1 (S1.hmark (S1.add (V1 0) (V1 1)))
         (S1.hadd (S1.mark (V1 0)) (V1 1))
   :: nil)

:: (  R1 (S1.hmark (S1._if_2 (V1 0) (V1 1) (V1 2)))
         (S1.h_if_2 (S1.mark (V1 0)) (V1 1) (V1 2))
   :: nil)

:: (  R1 (S1.hmark (S1.and (V1 0) (V1 1)))
         (S1.hand (S1.mark (V1 0)) (V1 1))
   :: nil)

:: (  R1 (S1.hactive (S1.from (V1 0)))
         (S1.hs (V1 0))
   :: nil)

:: (  R1 (S1.hactive (S1.from (V1 0)))
         (S1.hfrom (S1.s (V1 0)))
   :: nil)

:: (  R1 (S1.hactive (S1.from (V1 0)))
         (S1.hcons (V1 0) (S1.from (S1.s (V1 0))))
   :: nil)

:: (  R1 (S1.hactive (S1.first (S1.s (V1 0)) (S1.cons (V1 1) (V1 2))))
         (S1.hfirst (V1 0) (V1 2))
   :: nil)

:: (  R1 (S1.hactive (S1.first (S1.s (V1 0)) (S1.cons (V1 1) (V1 2))))
         (S1.hcons (V1 1) (S1.first (V1 0) (V1 2)))
   :: nil)

:: (  R1 (S1.hactive (S1.add (S1.s (V1 0)) (V1 1)))
         (S1.hadd (V1 0) (V1 1))
   :: nil)

:: (  R1 (S1.hactive (S1.add (S1.s (V1 0)) (V1 1)))
         (S1.hs (S1.add (V1 0) (V1 1)))
   :: nil)

:: (  R1 (S1.hmark (S1.nil))
         (S1.hactive (S1.nil))
   :: nil)

:: (  R1 (S1.hactive (S1.first (S1._0_1) (V1 0)))
         (S1.hmark (S1.nil))
   :: nil)

:: (  R1 (S1.hmark (S1._0_1))
         (S1.hactive (S1._0_1))
   :: nil)

:: (  R1 (S1.hmark (S1.false))
         (S1.hactive (S1.false))
   :: nil)

:: (  R1 (S1.hmark (S1.true))
         (S1.hactive (S1.true))
   :: nil)

:: (  R1 (S1.hactive (S1.and (S1.false) (V1 0)))
         (S1.hmark (S1.false))
   :: nil)

:: (  R1 (S1.hmark (S1.and (V1 0) (V1 1)))
         (S1.hactive (S1.and (S1.mark (V1 0)) (V1 1)))
   :: R1 (S1.hactive (S1.and (S1.true) (V1 0)))
         (S1.hmark (V1 0))
   :: R1 (S1.hmark (S1.and (V1 0) (V1 1)))
         (S1.hmark (V1 0))
   :: R1 (S1.hmark (S1._if_2 (V1 0) (V1 1) (V1 2)))
         (S1.hactive (S1._if_2 (S1.mark (V1 0)) (V1 1) (V1 2)))
   :: R1 (S1.hactive (S1._if_2 (S1.true) (V1 0) (V1 1)))
         (S1.hmark (V1 0))
   :: R1 (S1.hmark (S1._if_2 (V1 0) (V1 1) (V1 2)))
         (S1.hmark (V1 0))
   :: R1 (S1.hmark (S1.add (V1 0) (V1 1)))
         (S1.hactive (S1.add (S1.mark (V1 0)) (V1 1)))
   :: R1 (S1.hactive (S1._if_2 (S1.false) (V1 0) (V1 1)))
         (S1.hmark (V1 1))
   :: R1 (S1.hmark (S1.add (V1 0) (V1 1)))
         (S1.hmark (V1 0))
   :: R1 (S1.hmark (S1.s (V1 0)))
         (S1.hactive (S1.s (V1 0)))
   :: R1 (S1.hactive (S1.add (S1._0_1) (V1 0)))
         (S1.hmark (V1 0))
   :: R1 (S1.hmark (S1.first (V1 0) (V1 1)))
         (S1.hactive (S1.first (S1.mark (V1 0)) (S1.mark (V1 1))))
   :: R1 (S1.hactive (S1.add (S1.s (V1 0)) (V1 1)))
         (S1.hmark (S1.s (S1.add (V1 0) (V1 1))))
   :: R1 (S1.hmark (S1.first (V1 0) (V1 1)))
         (S1.hmark (V1 0))
   :: R1 (S1.hmark (S1.first (V1 0) (V1 1)))
         (S1.hmark (V1 1))
   :: R1 (S1.hmark (S1.cons (V1 0) (V1 1)))
         (S1.hactive (S1.cons (V1 0) (V1 1)))
   :: R1 (S1.hactive (S1.first (S1.s (V1 0)) (S1.cons (V1 1) (V1 2))))
         (S1.hmark (S1.cons (V1 1) (S1.first (V1 0) (V1 2))))
   :: R1 (S1.hmark (S1.from (V1 0)))
         (S1.hactive (S1.from (V1 0)))
   :: R1 (S1.hactive (S1.from (V1 0)))
         (S1.hmark (S1.cons (V1 0) (S1.from (S1.s (V1 0)))))
   :: nil)

:: nil.

(* 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.active) =>
         nil
    | (int_symb M.active) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         (3%Z, Vnil)
      :: nil
    | (hd_symb M.mark) =>
         nil
    | (int_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         (1%Z, (Vcons 0 Vnil))
      :: nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         nil
    | (hd_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    | (int_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS1.

Module PI1 := PolyInt PIS1.

(* polynomial interpretation 2 *)

Module PIS2 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         nil
    | (int_symb M.active) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         nil
    | (hd_symb M.mark) =>
         nil
    | (int_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         (3%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (int_symb M.cons) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS2.

Module PI2 := PolyInt PIS2.

(* polynomial interpretation 3 *)

Module PIS3 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         nil
    | (int_symb M.active) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         nil
    | (hd_symb M.mark) =>
         nil
    | (int_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         nil
    | (hd_symb M.first) =>
         (3%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (int_symb M.first) =>
         nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS3.

Module PI3 := PolyInt PIS3.

(* polynomial interpretation 4 *)

Module PIS4 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         nil
    | (int_symb M.active) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         (3%Z, Vnil)
      :: nil
    | (hd_symb M.mark) =>
         nil
    | (int_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    | (int_symb M.s) =>
         (1%Z, (Vcons 0 Vnil))
      :: nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS4.

Module PI4 := PolyInt PIS4.

(* polynomial interpretation 5 *)

Module PIS5 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         nil
    | (int_symb M.active) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         nil
    | (hd_symb M.mark) =>
         nil
    | (int_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         (3%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (int_symb M.add) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS5.

Module PI5 := PolyInt PIS5.

(* polynomial interpretation 6 *)

Module PIS6 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         nil
    | (int_symb M.active) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         nil
    | (hd_symb M.mark) =>
         nil
    | (int_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         (3%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (3%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (int_symb M._if_2) =>
         (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS6.

Module PI6 := PolyInt PIS6.

(* polynomial interpretation 7 *)

Module PIS7 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         nil
    | (int_symb M.active) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         (3%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (int_symb M.and) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         nil
    | (hd_symb M.mark) =>
         nil
    | (int_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS7.

Module PI7 := PolyInt PIS7.

(* polynomial interpretation 8 *)

Module PIS8 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         (2%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (int_symb M.active) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         nil
    | (hd_symb M.mark) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (int_symb M.mark) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil))))
      :: (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (3%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         (2%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (3%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    end.

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

End PIS8.

Module PI8 := PolyInt PIS8.

(* polynomial interpretation 9 *)

Module PIS9 (*<: TPolyInt*).

  Definition sig := s1.

  Definition trsInt f :=
    match f as f return poly (@ASignature.arity s1 f) with
    | (hd_symb M.active) =>
         nil
    | (int_symb M.active) =>
         (1%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.and) =>
         nil
    | (int_symb M.and) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.true) =>
         nil
    | (int_symb M.true) =>
         (2%Z, Vnil)
      :: nil
    | (hd_symb M.mark) =>
         (2%Z, (Vcons 1 Vnil))
      :: nil
    | (int_symb M.mark) =>
         (1%Z, (Vcons 0 Vnil))
      :: (2%Z, (Vcons 1 Vnil))
      :: nil
    | (hd_symb M.false) =>
         nil
    | (int_symb M.false) =>
         nil
    | (hd_symb M._if_2) =>
         nil
    | (int_symb M._if_2) =>
         (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil))))
      :: (3%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil))))
      :: (3%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))
      :: nil
    | (hd_symb M.add) =>
         nil
    | (int_symb M.add) =>
         (1%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M._0_1) =>
         nil
    | (int_symb M._0_1) =>
         (3%Z, Vnil)
      :: nil
    | (hd_symb M.s) =>
         nil
    | (int_symb M.s) =>
         nil
    | (hd_symb M.first) =>
         nil
    | (int_symb M.first) =>
         (2%Z, (Vcons 0 (Vcons 0 Vnil)))
      :: (1%Z, (Vcons 1 (Vcons 0 Vnil)))
      :: (2%Z, (Vcons 0 (Vcons 1 Vnil)))
      :: nil
    | (hd_symb M.nil) =>
         nil
    | (int_symb M.nil) =>
         (2%Z, Vnil)
      :: nil
    | (hd_symb M.cons) =>
         nil
    | (int_symb M.cons) =>
         nil
    | (hd_symb M.from) =>
         nil
    | (int_symb M.from) =>
         (2%Z, (Vcons 0 Vnil))
      :: nil
    end.

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

End PIS9.

Module PI9 := PolyInt PIS9.

(* termination proof *)

Lemma termination : WF rel.

Proof.
unfold rel.
dp_trans.
mark.
let D := fresh "D" in let R := fresh "R" in set_rules_to D; set_mod_rules_to R;
graph_decomp (dpg_unif_N 100 R D) cs1; subst D; subst R.
dpg_unif_N_correct.
right. PI1.prove_termination.
termination_trivial.
right. PI2.prove_termination.
termination_trivial.
right. PI3.prove_termination.
termination_trivial.
right. PI4.prove_termination.
termination_trivial.
right. PI5.prove_termination.
termination_trivial.
right. PI6.prove_termination.
termination_trivial.
right. PI7.prove_termination.
termination_trivial.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
left. co_scc.
right. PI8.prove_termination.
PI9.prove_termination.
termination_trivial.
Qed.