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 := | app : symb | compose : symb | cons : symb | hd : symb | init : symb | last : symb | nil : symb | reverse : symb | reverse2 : symb | tl : 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.app => 2 | M.compose => 0 | M.cons => 0 | M.hd => 0 | M.init => 0 | M.last => 0 | M.nil => 0 | M.reverse => 0 | M.reverse2 => 0 | M.tl => 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 app x2 x1 := F0 M.app (Vcons x2 (Vcons x1 Vnil)). Definition compose := F0 M.compose Vnil. Definition cons := F0 M.cons Vnil. Definition hd := F0 M.hd Vnil. Definition init := F0 M.init Vnil. Definition last := F0 M.last Vnil. Definition nil := F0 M.nil Vnil. Definition reverse := F0 M.reverse Vnil. Definition reverse2 := F0 M.reverse2 Vnil. Definition tl := F0 M.tl Vnil. End S0. Definition E := @nil (@ATrs.rule s0). Definition R := R0 (S0.app (S0.app (S0.app S0.compose (V0 0)) (V0 1)) (V0 2)) (S0.app (V0 1) (S0.app (V0 0) (V0 2))) :: R0 (S0.app S0.reverse (V0 0)) (S0.app (S0.app S0.reverse2 (V0 0)) S0.nil) :: R0 (S0.app (S0.app S0.reverse2 S0.nil) (V0 0)) (V0 0) :: R0 (S0.app (S0.app S0.reverse2 (S0.app (S0.app S0.cons (V0 0)) (V0 1))) (V0 2)) (S0.app (S0.app S0.reverse2 (V0 1)) (S0.app (S0.app S0.cons (V0 0)) (V0 2))) :: R0 (S0.app S0.hd (S0.app (S0.app S0.cons (V0 0)) (V0 1))) (V0 0) :: R0 (S0.app S0.tl (S0.app (S0.app S0.cons (V0 0)) (V0 1))) (V0 1) :: R0 S0.last (S0.app (S0.app S0.compose S0.hd) S0.reverse) :: R0 S0.init (S0.app (S0.app S0.compose S0.reverse) (S0.app (S0.app S0.compose S0.tl) S0.reverse)) :: @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 happ x2 x1 := F1 (hd_symb s1_p M.app) (Vcons x2 (Vcons x1 Vnil)). Definition app x2 x1 := F1 (int_symb s1_p M.app) (Vcons x2 (Vcons x1 Vnil)). Definition hcompose := F1 (hd_symb s1_p M.compose) Vnil. Definition compose := F1 (int_symb s1_p M.compose) Vnil. Definition hcons := F1 (hd_symb s1_p M.cons) Vnil. Definition cons := F1 (int_symb s1_p M.cons) Vnil. Definition hhd := F1 (hd_symb s1_p M.hd) Vnil. Definition hd := F1 (int_symb s1_p M.hd) Vnil. Definition hinit := F1 (hd_symb s1_p M.init) Vnil. Definition init := F1 (int_symb s1_p M.init) Vnil. Definition hlast := F1 (hd_symb s1_p M.last) Vnil. Definition last := F1 (int_symb s1_p M.last) Vnil. Definition hnil := F1 (hd_symb s1_p M.nil) Vnil. Definition nil := F1 (int_symb s1_p M.nil) Vnil. Definition hreverse := F1 (hd_symb s1_p M.reverse) Vnil. Definition reverse := F1 (int_symb s1_p M.reverse) Vnil. Definition hreverse2 := F1 (hd_symb s1_p M.reverse2) Vnil. Definition reverse2 := F1 (int_symb s1_p M.reverse2) Vnil. Definition htl := F1 (hd_symb s1_p M.tl) Vnil. Definition tl := F1 (int_symb s1_p M.tl) Vnil. End S1. (* graph decomposition 1 *) Definition cs1 : list (list (@ATrs.rule s1)) := ( R1 (S1.happ (S1.app (S1.reverse2) (S1.app (S1.app (S1.cons) (V1 0)) (V1 1))) (V1 2)) (S1.happ (S1.cons) (V1 0)) :: nil) :: ( R1 (S1.happ (S1.app (S1.reverse2) (S1.app (S1.app (S1.cons) (V1 0)) (V1 1))) (V1 2)) (S1.happ (S1.reverse2) (V1 1)) :: nil) :: ( R1 (S1.happ (S1.reverse) (V1 0)) (S1.happ (S1.reverse2) (V1 0)) :: nil) :: ( R1 (S1.happ (S1.app (S1.app (S1.compose) (V1 0)) (V1 1)) (V1 2)) (S1.happ (V1 0) (V1 2)) :: R1 (S1.happ (S1.app (S1.app (S1.compose) (V1 0)) (V1 1)) (V1 2)) (S1.happ (V1 1) (S1.app (V1 0) (V1 2))) :: R1 (S1.happ (S1.reverse) (V1 0)) (S1.happ (S1.app (S1.reverse2) (V1 0)) (S1.nil)) :: R1 (S1.happ (S1.app (S1.reverse2) (S1.app (S1.app (S1.cons) (V1 0)) (V1 1))) (V1 2)) (S1.happ (S1.app (S1.reverse2) (V1 1)) (S1.app (S1.app (S1.cons) (V1 0)) (V1 2))) :: R1 (S1.happ (S1.app (S1.reverse2) (S1.app (S1.app (S1.cons) (V1 0)) (V1 1))) (V1 2)) (S1.happ (S1.app (S1.cons) (V1 0)) (V1 2)) :: nil) :: ( R1 (S1.hinit) (S1.happ (S1.compose) (S1.tl)) :: nil) :: ( R1 (S1.hinit) (S1.happ (S1.app (S1.compose) (S1.tl)) (S1.reverse)) :: nil) :: ( R1 (S1.hinit) (S1.happ (S1.compose) (S1.reverse)) :: nil) :: ( R1 (S1.hinit) (S1.happ (S1.app (S1.compose) (S1.reverse)) (S1.app (S1.app (S1.compose) (S1.tl)) (S1.reverse))) :: nil) :: ( R1 (S1.hlast) (S1.happ (S1.compose) (S1.hd)) :: nil) :: ( R1 (S1.hlast) (S1.happ (S1.app (S1.compose) (S1.hd)) (S1.reverse)) :: 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.app) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.app) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.compose) => nil | (int_symb M.compose) => nil | (hd_symb M.reverse) => nil | (int_symb M.reverse) => (1%Z, Vnil) :: nil | (hd_symb M.reverse2) => nil | (int_symb M.reverse2) => nil | (hd_symb M.nil) => nil | (int_symb M.nil) => nil | (hd_symb M.cons) => nil | (int_symb M.cons) => nil | (hd_symb M.hd) => nil | (int_symb M.hd) => nil | (hd_symb M.tl) => nil | (int_symb M.tl) => nil | (hd_symb M.last) => nil | (int_symb M.last) => (1%Z, Vnil) :: nil | (hd_symb M.init) => nil | (int_symb M.init) => (3%Z, 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.app) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.app) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.compose) => nil | (int_symb M.compose) => (1%Z, Vnil) :: nil | (hd_symb M.reverse) => nil | (int_symb M.reverse) => nil | (hd_symb M.reverse2) => nil | (int_symb M.reverse2) => nil | (hd_symb M.nil) => nil | (int_symb M.nil) => nil | (hd_symb M.cons) => nil | (int_symb M.cons) => (1%Z, Vnil) :: nil | (hd_symb M.hd) => nil | (int_symb M.hd) => nil | (hd_symb M.tl) => nil | (int_symb M.tl) => nil | (hd_symb M.last) => nil | (int_symb M.last) => (1%Z, Vnil) :: nil | (hd_symb M.init) => nil | (int_symb M.init) => (2%Z, 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.app) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.app) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.compose) => nil | (int_symb M.compose) => nil | (hd_symb M.reverse) => nil | (int_symb M.reverse) => nil | (hd_symb M.reverse2) => nil | (int_symb M.reverse2) => nil | (hd_symb M.nil) => nil | (int_symb M.nil) => nil | (hd_symb M.cons) => nil | (int_symb M.cons) => (1%Z, Vnil) :: nil | (hd_symb M.hd) => nil | (int_symb M.hd) => nil | (hd_symb M.tl) => nil | (int_symb M.tl) => (1%Z, Vnil) :: nil | (hd_symb M.last) => nil | (int_symb M.last) => (3%Z, Vnil) :: nil | (hd_symb M.init) => nil | (int_symb M.init) => (2%Z, 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.app) => (2%Z, (Vcons 1 (Vcons 0 Vnil))) :: nil | (int_symb M.app) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.compose) => nil | (int_symb M.compose) => nil | (hd_symb M.reverse) => nil | (int_symb M.reverse) => (1%Z, Vnil) :: nil | (hd_symb M.reverse2) => nil | (int_symb M.reverse2) => (1%Z, Vnil) :: nil | (hd_symb M.nil) => nil | (int_symb M.nil) => nil | (hd_symb M.cons) => nil | (int_symb M.cons) => nil | (hd_symb M.hd) => nil | (int_symb M.hd) => nil | (hd_symb M.tl) => nil | (int_symb M.tl) => (1%Z, Vnil) :: nil | (hd_symb M.last) => nil | (int_symb M.last) => (2%Z, Vnil) :: nil | (hd_symb M.init) => nil | (int_symb M.init) => (3%Z, Vnil) :: nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS4. Module PI4 := PolyInt PIS4. (* 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. PI2.prove_termination. PI3.prove_termination. PI4.prove_termination. termination_trivial. left. co_scc. left. co_scc. left. co_scc. left. co_scc. left. co_scc. left. co_scc. Qed.