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 := | Cons : symb | False : symb | Nil : symb | Pair : symb | True : symb | append : symb | match_0 : symb | match_1 : symb | match_2 : symb | match_3 : symb | match_4 : symb | match_5 : symb | part : symb | quick : symb | test : 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.Cons => 2 | M.False => 0 | M.Nil => 0 | M.Pair => 2 | M.True => 0 | M.append => 2 | M.match_0 => 3 | M.match_1 => 3 | M.match_2 => 5 | M.match_3 => 7 | M.match_4 => 2 | M.match_5 => 4 | M.part => 2 | M.quick => 1 | M.test => 2 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 Cons x2 x1 := F0 M.Cons (Vcons x2 (Vcons x1 Vnil)). Definition False := F0 M.False Vnil. Definition Nil := F0 M.Nil Vnil. Definition Pair x2 x1 := F0 M.Pair (Vcons x2 (Vcons x1 Vnil)). Definition True := F0 M.True Vnil. Definition append x2 x1 := F0 M.append (Vcons x2 (Vcons x1 Vnil)). Definition match_0 x3 x2 x1 := F0 M.match_0 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition match_1 x3 x2 x1 := F0 M.match_1 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition match_2 x5 x4 x3 x2 x1 := F0 M.match_2 (Vcons x5 (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))))). Definition match_3 x7 x6 x5 x4 x3 x2 x1 := F0 M.match_3 (Vcons x7 (Vcons x6 (Vcons x5 (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))))))). Definition match_4 x2 x1 := F0 M.match_4 (Vcons x2 (Vcons x1 Vnil)). Definition match_5 x4 x3 x2 x1 := F0 M.match_5 (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil)))). Definition part x2 x1 := F0 M.part (Vcons x2 (Vcons x1 Vnil)). Definition quick x1 := F0 M.quick (Vcons x1 Vnil). Definition test x2 x1 := F0 M.test (Vcons x2 (Vcons x1 Vnil)). End S0. Definition E := @nil (@ATrs.rule s0). Definition R := R0 (S0.test (V0 0) (V0 1)) S0.True :: R0 (S0.test (V0 0) (V0 1)) S0.False :: R0 (S0.append (V0 0) (V0 1)) (S0.match_0 (V0 0) (V0 1) (V0 0)) :: R0 (S0.match_0 (V0 0) (V0 1) S0.Nil) (V0 1) :: R0 (S0.match_0 (V0 0) (V0 1) (S0.Cons (V0 2) (V0 3))) (S0.Cons (V0 2) (S0.append (V0 3) (V0 1))) :: R0 (S0.part (V0 0) (V0 1)) (S0.match_1 (V0 0) (V0 1) (V0 1)) :: R0 (S0.match_1 (V0 0) (V0 1) S0.Nil) (S0.Pair S0.Nil S0.Nil) :: R0 (S0.match_1 (V0 0) (V0 1) (S0.Cons (V0 2) (V0 3))) (S0.match_2 (V0 2) (V0 3) (V0 0) (V0 1) (S0.part (V0 0) (V0 3))) :: R0 (S0.match_2 (V0 0) (V0 1) (V0 2) (V0 3) (S0.Pair (V0 4) (V0 5))) (S0.match_3 (V0 4) (V0 5) (V0 0) (V0 1) (V0 2) (V0 3) (S0.test (V0 2) (V0 0))) :: R0 (S0.match_3 (V0 0) (V0 1) (V0 2) (V0 3) (V0 4) (V0 5) S0.False) (S0.Pair (S0.Cons (V0 2) (V0 0)) (V0 1)) :: R0 (S0.match_3 (V0 0) (V0 1) (V0 2) (V0 3) (V0 4) (V0 5) S0.True) (S0.Pair (V0 0) (S0.Cons (V0 2) (V0 1))) :: R0 (S0.quick (V0 0)) (S0.match_4 (V0 0) (V0 0)) :: R0 (S0.match_4 (V0 0) S0.Nil) S0.Nil :: R0 (S0.match_4 (V0 0) (S0.Cons (V0 1) (V0 2))) (S0.match_5 (V0 1) (V0 2) (V0 0) (S0.part (V0 1) (V0 2))) :: R0 (S0.match_5 (V0 0) (V0 1) (V0 2) (S0.Pair (V0 3) (V0 4))) (S0.append (S0.quick (V0 3)) (S0.Cons (V0 0) (S0.quick (V0 4)))) :: @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 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 hNil := F1 (hd_symb s1_p M.Nil) Vnil. Definition Nil := F1 (int_symb s1_p M.Nil) Vnil. Definition hPair x2 x1 := F1 (hd_symb s1_p M.Pair) (Vcons x2 (Vcons x1 Vnil)). Definition Pair x2 x1 := F1 (int_symb s1_p M.Pair) (Vcons x2 (Vcons x1 Vnil)). Definition hTrue := F1 (hd_symb s1_p M.True) Vnil. Definition True := F1 (int_symb s1_p M.True) Vnil. Definition happend x2 x1 := F1 (hd_symb s1_p M.append) (Vcons x2 (Vcons x1 Vnil)). Definition append x2 x1 := F1 (int_symb s1_p M.append) (Vcons x2 (Vcons x1 Vnil)). Definition hmatch_0 x3 x2 x1 := F1 (hd_symb s1_p M.match_0) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition match_0 x3 x2 x1 := F1 (int_symb s1_p M.match_0) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition hmatch_1 x3 x2 x1 := F1 (hd_symb s1_p M.match_1) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition match_1 x3 x2 x1 := F1 (int_symb s1_p M.match_1) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition hmatch_2 x5 x4 x3 x2 x1 := F1 (hd_symb s1_p M.match_2) (Vcons x5 (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))))). Definition match_2 x5 x4 x3 x2 x1 := F1 (int_symb s1_p M.match_2) (Vcons x5 (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))))). Definition hmatch_3 x7 x6 x5 x4 x3 x2 x1 := F1 (hd_symb s1_p M.match_3) (Vcons x7 (Vcons x6 (Vcons x5 (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))))))). Definition match_3 x7 x6 x5 x4 x3 x2 x1 := F1 (int_symb s1_p M.match_3) (Vcons x7 (Vcons x6 (Vcons x5 (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))))))). Definition hmatch_4 x2 x1 := F1 (hd_symb s1_p M.match_4) (Vcons x2 (Vcons x1 Vnil)). Definition match_4 x2 x1 := F1 (int_symb s1_p M.match_4) (Vcons x2 (Vcons x1 Vnil)). Definition hmatch_5 x4 x3 x2 x1 := F1 (hd_symb s1_p M.match_5) (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil)))). Definition match_5 x4 x3 x2 x1 := F1 (int_symb s1_p M.match_5) (Vcons x4 (Vcons x3 (Vcons x2 (Vcons x1 Vnil)))). Definition hpart x2 x1 := F1 (hd_symb s1_p M.part) (Vcons x2 (Vcons x1 Vnil)). Definition part x2 x1 := F1 (int_symb s1_p M.part) (Vcons x2 (Vcons x1 Vnil)). Definition hquick x1 := F1 (hd_symb s1_p M.quick) (Vcons x1 Vnil). Definition quick x1 := F1 (int_symb s1_p M.quick) (Vcons x1 Vnil). Definition htest x2 x1 := F1 (hd_symb s1_p M.test) (Vcons x2 (Vcons x1 Vnil)). Definition test x2 x1 := F1 (int_symb s1_p M.test) (Vcons x2 (Vcons x1 Vnil)). End S1. (* graph decomposition 1 *) Definition cs1 : list (list (@ATrs.rule s1)) := ( R1 (S1.hmatch_2 (V1 0) (V1 1) (V1 2) (V1 3) (S1.Pair (V1 4) (V1 5))) (S1.htest (V1 2) (V1 0)) :: nil) :: ( R1 (S1.hmatch_2 (V1 0) (V1 1) (V1 2) (V1 3) (S1.Pair (V1 4) (V1 5))) (S1.hmatch_3 (V1 4) (V1 5) (V1 0) (V1 1) (V1 2) (V1 3) (S1.test (V1 2) (V1 0))) :: nil) :: ( R1 (S1.hmatch_1 (V1 0) (V1 1) (S1.Cons (V1 2) (V1 3))) (S1.hmatch_2 (V1 2) (V1 3) (V1 0) (V1 1) (S1.part (V1 0) (V1 3))) :: nil) :: ( R1 (S1.hmatch_1 (V1 0) (V1 1) (S1.Cons (V1 2) (V1 3))) (S1.hpart (V1 0) (V1 3)) :: R1 (S1.hpart (V1 0) (V1 1)) (S1.hmatch_1 (V1 0) (V1 1) (V1 1)) :: nil) :: ( R1 (S1.hmatch_4 (V1 0) (S1.Cons (V1 1) (V1 2))) (S1.hpart (V1 1) (V1 2)) :: nil) :: ( R1 (S1.hmatch_0 (V1 0) (V1 1) (S1.Cons (V1 2) (V1 3))) (S1.happend (V1 3) (V1 1)) :: R1 (S1.happend (V1 0) (V1 1)) (S1.hmatch_0 (V1 0) (V1 1) (V1 0)) :: nil) :: ( R1 (S1.hmatch_5 (V1 0) (V1 1) (V1 2) (S1.Pair (V1 3) (V1 4))) (S1.happend (S1.quick (V1 3)) (S1.Cons (V1 0) (S1.quick (V1 4)))) :: nil) :: ( R1 (S1.hmatch_5 (V1 0) (V1 1) (V1 2) (S1.Pair (V1 3) (V1 4))) (S1.hquick (V1 3)) :: R1 (S1.hquick (V1 0)) (S1.hmatch_4 (V1 0) (V1 0)) :: R1 (S1.hmatch_4 (V1 0) (S1.Cons (V1 1) (V1 2))) (S1.hmatch_5 (V1 1) (V1 2) (V1 0) (S1.part (V1 1) (V1 2))) :: R1 (S1.hmatch_5 (V1 0) (V1 1) (V1 2) (S1.Pair (V1 3) (V1 4))) (S1.hquick (V1 4)) :: 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.test) => nil | (int_symb M.test) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: nil | (hd_symb M.True) => nil | (int_symb M.True) => (1%Z, Vnil) :: nil | (hd_symb M.False) => nil | (int_symb M.False) => (1%Z, Vnil) :: nil | (hd_symb M.append) => nil | (int_symb M.append) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_0) => nil | (int_symb M.match_0) => (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (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) => (2%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.part) => (3%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.part) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_1) => (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (int_symb M.match_1) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.Pair) => nil | (int_symb M.Pair) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_2) => nil | (int_symb M.match_2) => (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))))) :: nil | (hd_symb M.match_3) => nil | (int_symb M.match_3) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))))))) :: nil | (hd_symb M.quick) => nil | (int_symb M.quick) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.match_4) => nil | (int_symb M.match_4) => (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_5) => nil | (int_symb M.match_5) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil))))) :: (1%Z, (Vcons 0 (Vcons 0 (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. (* 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.test) => nil | (int_symb M.test) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: nil | (hd_symb M.True) => nil | (int_symb M.True) => (1%Z, Vnil) :: nil | (hd_symb M.False) => nil | (int_symb M.False) => (1%Z, Vnil) :: nil | (hd_symb M.append) => (3%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (int_symb M.append) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_0) => (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (int_symb M.match_0) => (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (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) => (2%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.part) => nil | (int_symb M.part) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_1) => nil | (int_symb M.match_1) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.Pair) => nil | (int_symb M.Pair) => (1%Z, (Vcons 0 (Vcons 0 Vnil))) :: (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_2) => nil | (int_symb M.match_2) => (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))))) :: nil | (hd_symb M.match_3) => nil | (int_symb M.match_3) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))))))) :: nil | (hd_symb M.quick) => nil | (int_symb M.quick) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.match_4) => nil | (int_symb M.match_4) => (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_5) => nil | (int_symb M.match_5) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil))))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (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.test) => nil | (int_symb M.test) => nil | (hd_symb M.True) => nil | (int_symb M.True) => nil | (hd_symb M.False) => nil | (int_symb M.False) => nil | (hd_symb M.append) => nil | (int_symb M.append) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_0) => nil | (int_symb M.match_0) => (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 Vnil)))) :: (1%Z, (Vcons 0 (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))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.part) => nil | (int_symb M.part) => (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_1) => nil | (int_symb M.match_1) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))) :: nil | (hd_symb M.Pair) => nil | (int_symb M.Pair) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_2) => nil | (int_symb M.match_2) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 1 Vnil)))))) :: nil | (hd_symb M.match_3) => nil | (int_symb M.match_3) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (1%Z, (Vcons 1 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: (1%Z, (Vcons 0 (Vcons 1 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))))))) :: nil | (hd_symb M.quick) => (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.quick) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.match_4) => (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.match_4) => (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.match_5) => (2%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))) :: nil | (int_symb M.match_5) => (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 0 Vnil))))) :: (1%Z, (Vcons 0 (Vcons 0 (Vcons 0 (Vcons 1 Vnil))))) :: nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS3. Module PI3 := PolyInt PIS3. (* graph decomposition 2 *) Definition cs2 : list (list (@ATrs.rule s1)) := ( R1 (S1.hquick (V1 0)) (S1.hmatch_4 (V1 0) (V1 0)) :: nil) :: ( R1 (S1.hmatch_5 (V1 0) (V1 1) (V1 2) (S1.Pair (V1 3) (V1 4))) (S1.hquick (V1 4)) :: nil) :: ( R1 (S1.hmatch_5 (V1 0) (V1 1) (V1 2) (S1.Pair (V1 3) (V1 4))) (S1.hquick (V1 3)) :: nil) :: nil. (* 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. right. PI1.prove_termination. termination_trivial. left. co_scc. right. PI2.prove_termination. termination_trivial. left. co_scc. right. PI3.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. left. co_scc. left. co_scc. left. co_scc. Qed.