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_4 : symb | _1_5 : symb | _if_6 : symb | _minus__3 : symb | _plus__2 : symb | _sharp__1 : symb | and : symb | bs : symb | false : symb | ge : symb | l : symb | max : symb | min : symb | n : symb | not : symb | size : symb | true : symb | val : symb | wb : 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_4 => 1 | M._1_5 => 1 | M._if_6 => 3 | M._minus__3 => 2 | M._plus__2 => 2 | M._sharp__1 => 0 | M.and => 2 | M.bs => 1 | M.false => 0 | M.ge => 2 | M.l => 1 | M.max => 1 | M.min => 1 | M.n => 3 | M.not => 1 | M.size => 1 | M.true => 0 | M.val => 1 | M.wb => 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 _0_4 x1 := F0 M._0_4 (Vcons x1 Vnil). Definition _1_5 x1 := F0 M._1_5 (Vcons x1 Vnil). Definition _if_6 x3 x2 x1 := F0 M._if_6 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition _minus__3 x2 x1 := F0 M._minus__3 (Vcons x2 (Vcons x1 Vnil)). Definition _plus__2 x2 x1 := F0 M._plus__2 (Vcons x2 (Vcons x1 Vnil)). Definition _sharp__1 := F0 M._sharp__1 Vnil. Definition and x2 x1 := F0 M.and (Vcons x2 (Vcons x1 Vnil)). Definition bs x1 := F0 M.bs (Vcons x1 Vnil). Definition false := F0 M.false Vnil. Definition ge x2 x1 := F0 M.ge (Vcons x2 (Vcons x1 Vnil)). Definition l x1 := F0 M.l (Vcons x1 Vnil). Definition max x1 := F0 M.max (Vcons x1 Vnil). Definition min x1 := F0 M.min (Vcons x1 Vnil). Definition n x3 x2 x1 := F0 M.n (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition not x1 := F0 M.not (Vcons x1 Vnil). Definition size x1 := F0 M.size (Vcons x1 Vnil). Definition true := F0 M.true Vnil. Definition val x1 := F0 M.val (Vcons x1 Vnil). Definition wb x1 := F0 M.wb (Vcons x1 Vnil). End S0. Definition E := @nil (@ATrs.rule s0). Definition R := R0 (S0._0_4 S0._sharp__1) S0._sharp__1 :: R0 (S0._plus__2 (V0 0) S0._sharp__1) (V0 0) :: R0 (S0._plus__2 S0._sharp__1 (V0 0)) (V0 0) :: R0 (S0._plus__2 (S0._0_4 (V0 0)) (S0._0_4 (V0 1))) (S0._0_4 (S0._plus__2 (V0 0) (V0 1))) :: R0 (S0._plus__2 (S0._0_4 (V0 0)) (S0._1_5 (V0 1))) (S0._1_5 (S0._plus__2 (V0 0) (V0 1))) :: R0 (S0._plus__2 (S0._1_5 (V0 0)) (S0._0_4 (V0 1))) (S0._1_5 (S0._plus__2 (V0 0) (V0 1))) :: R0 (S0._plus__2 (S0._1_5 (V0 0)) (S0._1_5 (V0 1))) (S0._0_4 (S0._plus__2 (S0._plus__2 (V0 0) (V0 1)) (S0._1_5 S0._sharp__1))) :: R0 (S0._plus__2 (V0 0) (S0._plus__2 (V0 1) (V0 2))) (S0._plus__2 (S0._plus__2 (V0 0) (V0 1)) (V0 2)) :: R0 (S0._minus__3 (V0 0) S0._sharp__1) (V0 0) :: R0 (S0._minus__3 S0._sharp__1 (V0 0)) S0._sharp__1 :: R0 (S0._minus__3 (S0._0_4 (V0 0)) (S0._0_4 (V0 1))) (S0._0_4 (S0._minus__3 (V0 0) (V0 1))) :: R0 (S0._minus__3 (S0._0_4 (V0 0)) (S0._1_5 (V0 1))) (S0._1_5 (S0._minus__3 (S0._minus__3 (V0 0) (V0 1)) (S0._1_5 S0._sharp__1))) :: R0 (S0._minus__3 (S0._1_5 (V0 0)) (S0._0_4 (V0 1))) (S0._1_5 (S0._minus__3 (V0 0) (V0 1))) :: R0 (S0._minus__3 (S0._1_5 (V0 0)) (S0._1_5 (V0 1))) (S0._0_4 (S0._minus__3 (V0 0) (V0 1))) :: R0 (S0.not S0.false) S0.true :: R0 (S0.not S0.true) S0.false :: R0 (S0.and (V0 0) S0.true) (V0 0) :: R0 (S0.and (V0 0) S0.false) S0.false :: R0 (S0._if_6 S0.true (V0 0) (V0 1)) (V0 0) :: R0 (S0._if_6 S0.false (V0 0) (V0 1)) (V0 1) :: R0 (S0.ge (S0._0_4 (V0 0)) (S0._0_4 (V0 1))) (S0.ge (V0 0) (V0 1)) :: R0 (S0.ge (S0._0_4 (V0 0)) (S0._1_5 (V0 1))) (S0.not (S0.ge (V0 1) (V0 0))) :: R0 (S0.ge (S0._1_5 (V0 0)) (S0._0_4 (V0 1))) (S0.ge (V0 0) (V0 1)) :: R0 (S0.ge (S0._1_5 (V0 0)) (S0._1_5 (V0 1))) (S0.ge (V0 0) (V0 1)) :: R0 (S0.ge (V0 0) S0._sharp__1) S0.true :: R0 (S0.ge S0._sharp__1 (S0._1_5 (V0 0))) S0.false :: R0 (S0.ge S0._sharp__1 (S0._0_4 (V0 0))) (S0.ge S0._sharp__1 (V0 0)) :: R0 (S0.val (S0.l (V0 0))) (V0 0) :: R0 (S0.val (S0.n (V0 0) (V0 1) (V0 2))) (V0 0) :: R0 (S0.min (S0.l (V0 0))) (V0 0) :: R0 (S0.min (S0.n (V0 0) (V0 1) (V0 2))) (S0.min (V0 1)) :: R0 (S0.max (S0.l (V0 0))) (V0 0) :: R0 (S0.max (S0.n (V0 0) (V0 1) (V0 2))) (S0.max (V0 2)) :: R0 (S0.bs (S0.l (V0 0))) S0.true :: R0 (S0.bs (S0.n (V0 0) (V0 1) (V0 2))) (S0.and (S0.and (S0.ge (V0 0) (S0.max (V0 1))) (S0.ge (S0.min (V0 2)) (V0 0))) (S0.and (S0.bs (V0 1)) (S0.bs (V0 2)))) :: R0 (S0.size (S0.l (V0 0))) (S0._1_5 S0._sharp__1) :: R0 (S0.size (S0.n (V0 0) (V0 1) (V0 2))) (S0._plus__2 (S0._plus__2 (S0.size (V0 0)) (S0.size (V0 1))) (S0._1_5 S0._sharp__1)) :: R0 (S0.wb (S0.l (V0 0))) S0.true :: R0 (S0.wb (S0.n (V0 0) (V0 1) (V0 2))) (S0.and (S0._if_6 (S0.ge (S0.size (V0 1)) (S0.size (V0 2))) (S0.ge (S0._1_5 S0._sharp__1) (S0._minus__3 (S0.size (V0 1)) (S0.size (V0 2)))) (S0.ge (S0._1_5 S0._sharp__1) (S0._minus__3 (S0.size (V0 2)) (S0.size (V0 1))))) (S0.and (S0.wb (V0 1)) (S0.wb (V0 2)))) :: @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_4 x1 := F1 (hd_symb s1_p M._0_4) (Vcons x1 Vnil). Definition _0_4 x1 := F1 (int_symb s1_p M._0_4) (Vcons x1 Vnil). Definition h_1_5 x1 := F1 (hd_symb s1_p M._1_5) (Vcons x1 Vnil). Definition _1_5 x1 := F1 (int_symb s1_p M._1_5) (Vcons x1 Vnil). Definition h_if_6 x3 x2 x1 := F1 (hd_symb s1_p M._if_6) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition _if_6 x3 x2 x1 := F1 (int_symb s1_p M._if_6) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition h_minus__3 x2 x1 := F1 (hd_symb s1_p M._minus__3) (Vcons x2 (Vcons x1 Vnil)). Definition _minus__3 x2 x1 := F1 (int_symb s1_p M._minus__3) (Vcons x2 (Vcons x1 Vnil)). Definition h_plus__2 x2 x1 := F1 (hd_symb s1_p M._plus__2) (Vcons x2 (Vcons x1 Vnil)). Definition _plus__2 x2 x1 := F1 (int_symb s1_p M._plus__2) (Vcons x2 (Vcons x1 Vnil)). Definition h_sharp__1 := F1 (hd_symb s1_p M._sharp__1) Vnil. Definition _sharp__1 := F1 (int_symb s1_p M._sharp__1) 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 hbs x1 := F1 (hd_symb s1_p M.bs) (Vcons x1 Vnil). Definition bs x1 := F1 (int_symb s1_p M.bs) (Vcons x1 Vnil). Definition hfalse := F1 (hd_symb s1_p M.false) Vnil. Definition false := F1 (int_symb s1_p M.false) Vnil. Definition hge x2 x1 := F1 (hd_symb s1_p M.ge) (Vcons x2 (Vcons x1 Vnil)). Definition ge x2 x1 := F1 (int_symb s1_p M.ge) (Vcons x2 (Vcons x1 Vnil)). Definition hl x1 := F1 (hd_symb s1_p M.l) (Vcons x1 Vnil). Definition l x1 := F1 (int_symb s1_p M.l) (Vcons x1 Vnil). Definition hmax x1 := F1 (hd_symb s1_p M.max) (Vcons x1 Vnil). Definition max x1 := F1 (int_symb s1_p M.max) (Vcons x1 Vnil). Definition hmin x1 := F1 (hd_symb s1_p M.min) (Vcons x1 Vnil). Definition min x1 := F1 (int_symb s1_p M.min) (Vcons x1 Vnil). Definition hn x3 x2 x1 := F1 (hd_symb s1_p M.n) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition n x3 x2 x1 := F1 (int_symb s1_p M.n) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition hnot x1 := F1 (hd_symb s1_p M.not) (Vcons x1 Vnil). Definition not x1 := F1 (int_symb s1_p M.not) (Vcons x1 Vnil). Definition hsize x1 := F1 (hd_symb s1_p M.size) (Vcons x1 Vnil). Definition size x1 := F1 (int_symb s1_p M.size) (Vcons x1 Vnil). Definition htrue := F1 (hd_symb s1_p M.true) Vnil. Definition true := F1 (int_symb s1_p M.true) Vnil. Definition hval x1 := F1 (hd_symb s1_p M.val) (Vcons x1 Vnil). Definition val x1 := F1 (int_symb s1_p M.val) (Vcons x1 Vnil). Definition hwb x1 := F1 (hd_symb s1_p M.wb) (Vcons x1 Vnil). Definition wb x1 := F1 (int_symb s1_p M.wb) (Vcons x1 Vnil). End S1. (* graph decomposition 1 *) Definition cs1 : list (list (@ATrs.rule s1)) := ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hand (S1.wb (V1 1)) (S1.wb (V1 2))) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.h_if_6 (S1.ge (S1.size (V1 1)) (S1.size (V1 2))) (S1.ge (S1._1_5 (S1._sharp__1)) (S1._minus__3 (S1.size (V1 1)) (S1.size (V1 2)))) (S1.ge (S1._1_5 (S1._sharp__1)) (S1._minus__3 (S1.size (V1 2)) (S1.size (V1 1))))) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hand (S1._if_6 (S1.ge (S1.size (V1 1)) (S1.size (V1 2))) (S1.ge (S1._1_5 (S1._sharp__1)) (S1._minus__3 (S1.size (V1 1)) (S1.size (V1 2)))) (S1.ge (S1._1_5 (S1._sharp__1)) (S1._minus__3 (S1.size (V1 2)) (S1.size (V1 1))))) (S1.and (S1.wb (V1 1)) (S1.wb (V1 2)))) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hand (S1.bs (V1 1)) (S1.bs (V1 2))) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hand (S1.ge (V1 0) (S1.max (V1 1))) (S1.ge (S1.min (V1 2)) (V1 0))) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hand (S1.and (S1.ge (V1 0) (S1.max (V1 1))) (S1.ge (S1.min (V1 2)) (V1 0))) (S1.and (S1.bs (V1 1)) (S1.bs (V1 2)))) :: nil) :: ( R1 (S1.hmax (S1.n (V1 0) (V1 1) (V1 2))) (S1.hmax (V1 2)) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hmax (V1 1)) :: nil) :: ( R1 (S1.hmin (S1.n (V1 0) (V1 1) (V1 2))) (S1.hmin (V1 1)) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hmin (V1 2)) :: nil) :: ( R1 (S1.hge (S1._0_4 (V1 0)) (S1._1_5 (V1 1))) (S1.hnot (S1.ge (V1 1) (V1 0))) :: nil) :: ( R1 (S1.hge (S1._sharp__1) (S1._0_4 (V1 0))) (S1.hge (S1._sharp__1) (V1 0)) :: nil) :: ( R1 (S1.hge (S1._0_4 (V1 0)) (S1._1_5 (V1 1))) (S1.hge (V1 1) (V1 0)) :: R1 (S1.hge (S1._0_4 (V1 0)) (S1._0_4 (V1 1))) (S1.hge (V1 0) (V1 1)) :: R1 (S1.hge (S1._1_5 (V1 0)) (S1._0_4 (V1 1))) (S1.hge (V1 0) (V1 1)) :: R1 (S1.hge (S1._1_5 (V1 0)) (S1._1_5 (V1 1))) (S1.hge (V1 0) (V1 1)) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hge (S1._1_5 (S1._sharp__1)) (S1._minus__3 (S1.size (V1 2)) (S1.size (V1 1)))) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hge (S1._1_5 (S1._sharp__1)) (S1._minus__3 (S1.size (V1 1)) (S1.size (V1 2)))) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hge (S1.size (V1 1)) (S1.size (V1 2))) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hge (S1.min (V1 2)) (V1 0)) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hge (V1 0) (S1.max (V1 1))) :: nil) :: ( R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hbs (V1 2)) :: R1 (S1.hbs (S1.n (V1 0) (V1 1) (V1 2))) (S1.hbs (V1 1)) :: nil) :: ( R1 (S1.h_minus__3 (S1._1_5 (V1 0)) (S1._1_5 (V1 1))) (S1.h_0_4 (S1._minus__3 (V1 0) (V1 1))) :: nil) :: ( R1 (S1.h_minus__3 (S1._0_4 (V1 0)) (S1._0_4 (V1 1))) (S1.h_0_4 (S1._minus__3 (V1 0) (V1 1))) :: nil) :: ( R1 (S1.h_minus__3 (S1._0_4 (V1 0)) (S1._1_5 (V1 1))) (S1.h_minus__3 (S1._minus__3 (V1 0) (V1 1)) (S1._1_5 (S1._sharp__1))) :: R1 (S1.h_minus__3 (S1._0_4 (V1 0)) (S1._1_5 (V1 1))) (S1.h_minus__3 (V1 0) (V1 1)) :: R1 (S1.h_minus__3 (S1._0_4 (V1 0)) (S1._0_4 (V1 1))) (S1.h_minus__3 (V1 0) (V1 1)) :: R1 (S1.h_minus__3 (S1._1_5 (V1 0)) (S1._0_4 (V1 1))) (S1.h_minus__3 (V1 0) (V1 1)) :: R1 (S1.h_minus__3 (S1._1_5 (V1 0)) (S1._1_5 (V1 1))) (S1.h_minus__3 (V1 0) (V1 1)) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.h_minus__3 (S1.size (V1 2)) (S1.size (V1 1))) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.h_minus__3 (S1.size (V1 1)) (S1.size (V1 2))) :: nil) :: ( R1 (S1.h_plus__2 (S1._1_5 (V1 0)) (S1._1_5 (V1 1))) (S1.h_0_4 (S1._plus__2 (S1._plus__2 (V1 0) (V1 1)) (S1._1_5 (S1._sharp__1)))) :: nil) :: ( R1 (S1.h_plus__2 (S1._0_4 (V1 0)) (S1._0_4 (V1 1))) (S1.h_0_4 (S1._plus__2 (V1 0) (V1 1))) :: nil) :: ( R1 (S1.h_plus__2 (S1._0_4 (V1 0)) (S1._1_5 (V1 1))) (S1.h_plus__2 (V1 0) (V1 1)) :: R1 (S1.h_plus__2 (S1._0_4 (V1 0)) (S1._0_4 (V1 1))) (S1.h_plus__2 (V1 0) (V1 1)) :: R1 (S1.h_plus__2 (S1._1_5 (V1 0)) (S1._0_4 (V1 1))) (S1.h_plus__2 (V1 0) (V1 1)) :: R1 (S1.h_plus__2 (S1._1_5 (V1 0)) (S1._1_5 (V1 1))) (S1.h_plus__2 (S1._plus__2 (V1 0) (V1 1)) (S1._1_5 (S1._sharp__1))) :: R1 (S1.h_plus__2 (S1._1_5 (V1 0)) (S1._1_5 (V1 1))) (S1.h_plus__2 (V1 0) (V1 1)) :: R1 (S1.h_plus__2 (V1 0) (S1._plus__2 (V1 1) (V1 2))) (S1.h_plus__2 (S1._plus__2 (V1 0) (V1 1)) (V1 2)) :: R1 (S1.h_plus__2 (V1 0) (S1._plus__2 (V1 1) (V1 2))) (S1.h_plus__2 (V1 0) (V1 1)) :: nil) :: ( R1 (S1.hsize (S1.n (V1 0) (V1 1) (V1 2))) (S1.h_plus__2 (S1.size (V1 0)) (S1.size (V1 1))) :: nil) :: ( R1 (S1.hsize (S1.n (V1 0) (V1 1) (V1 2))) (S1.h_plus__2 (S1._plus__2 (S1.size (V1 0)) (S1.size (V1 1))) (S1._1_5 (S1._sharp__1))) :: nil) :: ( R1 (S1.hsize (S1.n (V1 0) (V1 1) (V1 2))) (S1.hsize (V1 1)) :: R1 (S1.hsize (S1.n (V1 0) (V1 1) (V1 2))) (S1.hsize (V1 0)) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hsize (V1 2)) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hsize (V1 1)) :: nil) :: ( R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hwb (V1 2)) :: R1 (S1.hwb (S1.n (V1 0) (V1 1) (V1 2))) (S1.hwb (V1 1)) :: 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._0_4) => nil | (int_symb M._0_4) => nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (1%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)))) :: (3%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.max) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => nil | (hd_symb M.size) => nil | (int_symb M.size) => nil | (hd_symb M.wb) => nil | (int_symb M.wb) => 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._0_4) => nil | (int_symb M._0_4) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (3%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => (3%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.min) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 0 Vnil)) :: (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (2%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => (2%Z, (Vcons 0 Vnil)) :: (2%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._0_4) => nil | (int_symb M._0_4) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (3%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => 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._0_4) => nil | (int_symb M._0_4) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => (3%Z, (Vcons 1 (Vcons 0 Vnil))) :: (3%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (3%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (1%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => (1%Z, (Vcons 0 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._0_4) => nil | (int_symb M._0_4) => nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (2%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => (1%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => (1%Z, Vnil) :: nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (3%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (3%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.bs) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => (2%Z, (Vcons 0 Vnil)) :: (2%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._0_4) => nil | (int_symb M._0_4) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._minus__3) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (1%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => (2%Z, (Vcons 0 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._0_4) => nil | (int_symb M._0_4) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (3%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (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.min) => nil | (int_symb M.min) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => nil | (hd_symb M.size) => nil | (int_symb M.size) => (2%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS7. Module PI7 := PolyInt PIS7. (* graph decomposition 2 *) Definition cs2 : list (list (@ATrs.rule s1)) := ( R1 (S1.h_plus__2 (S1._1_5 (V1 0)) (S1._1_5 (V1 1))) (S1.h_plus__2 (S1._plus__2 (V1 0) (V1 1)) (S1._1_5 (S1._sharp__1))) :: nil) :: ( R1 (S1.h_plus__2 (V1 0) (S1._plus__2 (V1 1) (V1 2))) (S1.h_plus__2 (V1 0) (V1 1)) :: R1 (S1.h_plus__2 (V1 0) (S1._plus__2 (V1 1) (V1 2))) (S1.h_plus__2 (S1._plus__2 (V1 0) (V1 1)) (V1 2)) :: nil) :: nil. (* 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._0_4) => nil | (int_symb M._0_4) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (3%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => (1%Z, (Vcons 0 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._0_4) => nil | (int_symb M._0_4) => (1%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M._plus__2) => (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._1_5) => nil | (int_symb M._1_5) => nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (1%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => (1%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS9. Module PI9 := PolyInt PIS9. (* polynomial interpretation 10 *) Module PIS10 (*<: TPolyInt*). Definition sig := s1. Definition trsInt f := match f as f return poly (@ASignature.arity s1 f) with | (hd_symb M._0_4) => nil | (int_symb M._0_4) => (2%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (2%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (3%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (3%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (3%Z, (Vcons 0 Vnil)) :: (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.size) => (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.size) => (2%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => nil | (int_symb M.wb) => (1%Z, (Vcons 0 Vnil)) :: nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS10. Module PI10 := PolyInt PIS10. (* polynomial interpretation 11 *) Module PIS11 (*<: TPolyInt*). Definition sig := s1. Definition trsInt f := match f as f return poly (@ASignature.arity s1 f) with | (hd_symb M._0_4) => nil | (int_symb M._0_4) => nil | (hd_symb M._sharp__1) => nil | (int_symb M._sharp__1) => nil | (hd_symb M._plus__2) => nil | (int_symb M._plus__2) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._1_5) => nil | (int_symb M._1_5) => (2%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M._minus__3) => nil | (int_symb M._minus__3) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.not) => nil | (int_symb M.not) => (1%Z, (Vcons 0 Vnil)) :: nil | (hd_symb M.false) => nil | (int_symb M.false) => nil | (hd_symb M.true) => nil | (int_symb M.true) => (1%Z, Vnil) :: nil | (hd_symb M.and) => nil | (int_symb M.and) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (hd_symb M._if_6) => nil | (int_symb M._if_6) => (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.ge) => nil | (int_symb M.ge) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: nil | (hd_symb M.val) => nil | (int_symb M.val) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.l) => nil | (int_symb M.l) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.n) => nil | (int_symb M.n) => (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (3%Z, (Vcons 1 (Vcons 0 (Vcons 0 Vnil)))) :: (3%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.min) => nil | (int_symb M.min) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.max) => nil | (int_symb M.max) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.bs) => nil | (int_symb M.bs) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.size) => nil | (int_symb M.size) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.wb) => (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.wb) => (2%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS11. Module PI11 := PolyInt PIS11. (* 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. left. co_scc. left. co_scc. left. co_scc. left. co_scc. left. co_scc. left. co_scc. right. PI1.prove_termination. termination_trivial. left. co_scc. right. PI2.prove_termination. termination_trivial. left. co_scc. left. co_scc. right. PI3.prove_termination. termination_trivial. right. PI4.prove_termination. termination_trivial. left. co_scc. left. co_scc. left. co_scc. left. co_scc. left. co_scc. right. PI5.prove_termination. termination_trivial. left. co_scc. left. co_scc. right. PI6.prove_termination. termination_trivial. left. co_scc. left. co_scc. left. co_scc. left. co_scc. right. PI7.prove_termination. 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) cs2; subst D; subst R. dpg_unif_N_correct. right. PI8.prove_termination. termination_trivial. right. PI9.prove_termination. termination_trivial. left. co_scc. left. co_scc. right. PI10.prove_termination. termination_trivial. left. co_scc. left. co_scc. right. PI11.prove_termination. termination_trivial. Qed.