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 | a__adx : symb | a__hd : symb | a__incr : symb | a__nats : symb | a__tl : symb | a__zeros : symb | adx : symb | cons : symb | hd : symb | incr : symb | mark : symb | nats : symb | s : symb | tl : symb | zeros : 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.a__adx => 1 | M.a__hd => 1 | M.a__incr => 1 | M.a__nats => 0 | M.a__tl => 1 | M.a__zeros => 0 | M.adx => 1 | M.cons => 2 | M.hd => 1 | M.incr => 1 | M.mark => 1 | M.nats => 0 | M.s => 1 | M.tl => 1 | M.zeros => 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 a__adx x1 := F0 M.a__adx (Vcons x1 Vnil). Definition a__hd x1 := F0 M.a__hd (Vcons x1 Vnil). Definition a__incr x1 := F0 M.a__incr (Vcons x1 Vnil). Definition a__nats := F0 M.a__nats Vnil. Definition a__tl x1 := F0 M.a__tl (Vcons x1 Vnil). Definition a__zeros := F0 M.a__zeros Vnil. Definition adx x1 := F0 M.adx (Vcons x1 Vnil). Definition cons x2 x1 := F0 M.cons (Vcons x2 (Vcons x1 Vnil)). Definition hd x1 := F0 M.hd (Vcons x1 Vnil). Definition incr x1 := F0 M.incr (Vcons x1 Vnil). Definition mark x1 := F0 M.mark (Vcons x1 Vnil). Definition nats := F0 M.nats Vnil. Definition s x1 := F0 M.s (Vcons x1 Vnil). Definition tl x1 := F0 M.tl (Vcons x1 Vnil). Definition zeros := F0 M.zeros Vnil. End S0. Definition E := @nil (@ATrs.rule s0). Definition R := R0 S0.a__nats (S0.a__adx S0.a__zeros) :: R0 S0.a__zeros (S0.cons S0._0_1 S0.zeros) :: R0 (S0.a__incr (S0.cons (V0 0) (V0 1))) (S0.cons (S0.s (V0 0)) (S0.incr (V0 1))) :: R0 (S0.a__adx (S0.cons (V0 0) (V0 1))) (S0.a__incr (S0.cons (V0 0) (S0.adx (V0 1)))) :: R0 (S0.a__hd (S0.cons (V0 0) (V0 1))) (S0.mark (V0 0)) :: R0 (S0.a__tl (S0.cons (V0 0) (V0 1))) (S0.mark (V0 1)) :: R0 (S0.mark S0.nats) S0.a__nats :: R0 (S0.mark (S0.adx (V0 0))) (S0.a__adx (S0.mark (V0 0))) :: R0 (S0.mark S0.zeros) S0.a__zeros :: R0 (S0.mark (S0.incr (V0 0))) (S0.a__incr (S0.mark (V0 0))) :: R0 (S0.mark (S0.hd (V0 0))) (S0.a__hd (S0.mark (V0 0))) :: R0 (S0.mark (S0.tl (V0 0))) (S0.a__tl (S0.mark (V0 0))) :: R0 (S0.mark (S0.cons (V0 0) (V0 1))) (S0.cons (V0 0) (V0 1)) :: R0 (S0.mark S0._0_1) S0._0_1 :: R0 (S0.mark (S0.s (V0 0))) (S0.s (V0 0)) :: R0 S0.a__nats S0.nats :: R0 (S0.a__adx (V0 0)) (S0.adx (V0 0)) :: R0 S0.a__zeros S0.zeros :: R0 (S0.a__incr (V0 0)) (S0.incr (V0 0)) :: R0 (S0.a__hd (V0 0)) (S0.hd (V0 0)) :: R0 (S0.a__tl (V0 0)) (S0.tl (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 ha__adx x1 := F1 (hd_symb s1_p M.a__adx) (Vcons x1 Vnil). Definition a__adx x1 := F1 (int_symb s1_p M.a__adx) (Vcons x1 Vnil). Definition ha__hd x1 := F1 (hd_symb s1_p M.a__hd) (Vcons x1 Vnil). Definition a__hd x1 := F1 (int_symb s1_p M.a__hd) (Vcons x1 Vnil). Definition ha__incr x1 := F1 (hd_symb s1_p M.a__incr) (Vcons x1 Vnil). Definition a__incr x1 := F1 (int_symb s1_p M.a__incr) (Vcons x1 Vnil). Definition ha__nats := F1 (hd_symb s1_p M.a__nats) Vnil. Definition a__nats := F1 (int_symb s1_p M.a__nats) Vnil. Definition ha__tl x1 := F1 (hd_symb s1_p M.a__tl) (Vcons x1 Vnil). Definition a__tl x1 := F1 (int_symb s1_p M.a__tl) (Vcons x1 Vnil). Definition ha__zeros := F1 (hd_symb s1_p M.a__zeros) Vnil. Definition a__zeros := F1 (int_symb s1_p M.a__zeros) Vnil. Definition hadx x1 := F1 (hd_symb s1_p M.adx) (Vcons x1 Vnil). Definition adx x1 := F1 (int_symb s1_p M.adx) (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 hhd x1 := F1 (hd_symb s1_p M.hd) (Vcons x1 Vnil). Definition hd x1 := F1 (int_symb s1_p M.hd) (Vcons x1 Vnil). Definition hincr x1 := F1 (hd_symb s1_p M.incr) (Vcons x1 Vnil). Definition incr x1 := F1 (int_symb s1_p M.incr) (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 hnats := F1 (hd_symb s1_p M.nats) Vnil. Definition nats := F1 (int_symb s1_p M.nats) 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 htl x1 := F1 (hd_symb s1_p M.tl) (Vcons x1 Vnil). Definition tl x1 := F1 (int_symb s1_p M.tl) (Vcons x1 Vnil). Definition hzeros := F1 (hd_symb s1_p M.zeros) Vnil. Definition zeros := F1 (int_symb s1_p M.zeros) Vnil. End S1. (* graph decomposition 1 *) Definition cs1 : list (list (@ATrs.rule s1)) := ( R1 (S1.hmark (S1.incr (V1 0))) (S1.ha__incr (S1.mark (V1 0))) :: nil) :: ( R1 (S1.hmark (S1.zeros)) (S1.ha__zeros) :: nil) :: ( R1 (S1.ha__adx (S1.cons (V1 0) (V1 1))) (S1.ha__incr (S1.cons (V1 0) (S1.adx (V1 1)))) :: nil) :: ( R1 (S1.hmark (S1.adx (V1 0))) (S1.ha__adx (S1.mark (V1 0))) :: nil) :: ( R1 (S1.ha__nats) (S1.ha__zeros) :: nil) :: ( R1 (S1.ha__nats) (S1.ha__adx (S1.a__zeros)) :: nil) :: ( R1 (S1.hmark (S1.nats)) (S1.ha__nats) :: nil) :: ( R1 (S1.hmark (S1.adx (V1 0))) (S1.hmark (V1 0)) :: R1 (S1.hmark (S1.incr (V1 0))) (S1.hmark (V1 0)) :: R1 (S1.hmark (S1.hd (V1 0))) (S1.ha__hd (S1.mark (V1 0))) :: R1 (S1.ha__hd (S1.cons (V1 0) (V1 1))) (S1.hmark (V1 0)) :: R1 (S1.hmark (S1.hd (V1 0))) (S1.hmark (V1 0)) :: R1 (S1.hmark (S1.tl (V1 0))) (S1.ha__tl (S1.mark (V1 0))) :: R1 (S1.ha__tl (S1.cons (V1 0) (V1 1))) (S1.hmark (V1 1)) :: R1 (S1.hmark (S1.tl (V1 0))) (S1.hmark (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.a__nats) => nil | (int_symb M.a__nats) => nil | (hd_symb M.a__adx) => nil | (int_symb M.a__adx) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__zeros) => nil | (int_symb M.a__zeros) => nil | (hd_symb M.cons) => nil | (int_symb M.cons) => (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.zeros) => nil | (int_symb M.zeros) => nil | (hd_symb M.a__incr) => nil | (int_symb M.a__incr) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.s) => nil | (int_symb M.s) => nil | (hd_symb M.incr) => nil | (int_symb M.incr) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.adx) => nil | (int_symb M.adx) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__hd) => (3%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.a__hd) => (3%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.mark) => (2%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.mark) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__tl) => (3%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.a__tl) => (2%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.nats) => nil | (int_symb M.nats) => nil | (hd_symb M.hd) => nil | (int_symb M.hd) => (3%Z, (Vcons 0 Vnil)) :: (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.tl) => nil | (int_symb M.tl) => (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 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.a__nats) => nil | (int_symb M.a__nats) => (1%Z, Vnil) :: nil | (hd_symb M.a__adx) => nil | (int_symb M.a__adx) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__zeros) => nil | (int_symb M.a__zeros) => nil | (hd_symb M.cons) => nil | (int_symb M.cons) => (2%Z, (Vcons 1 (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.zeros) => nil | (int_symb M.zeros) => nil | (hd_symb M.a__incr) => nil | (int_symb M.a__incr) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.s) => nil | (int_symb M.s) => nil | (hd_symb M.incr) => nil | (int_symb M.incr) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.adx) => nil | (int_symb M.adx) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__hd) => nil | (int_symb M.a__hd) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.mark) => (2%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.mark) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__tl) => nil | (int_symb M.a__tl) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.nats) => nil | (int_symb M.nats) => (1%Z, Vnil) :: nil | (hd_symb M.hd) => nil | (int_symb M.hd) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.tl) => nil | (int_symb M.tl) => (3%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.a__nats) => nil | (int_symb M.a__nats) => (3%Z, Vnil) :: nil | (hd_symb M.a__adx) => nil | (int_symb M.a__adx) => (3%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__zeros) => nil | (int_symb M.a__zeros) => nil | (hd_symb M.cons) => nil | (int_symb M.cons) => (2%Z, (Vcons 1 (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.zeros) => nil | (int_symb M.zeros) => nil | (hd_symb M.a__incr) => nil | (int_symb M.a__incr) => (2%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.s) => nil | (int_symb M.s) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.incr) => nil | (int_symb M.incr) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.adx) => nil | (int_symb M.adx) => (1%Z, (Vcons 0 Vnil)) :: (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__hd) => nil | (int_symb M.a__hd) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.mark) => (1%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.mark) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.a__tl) => nil | (int_symb M.a__tl) => (2%Z, (Vcons 0 Vnil)) :: (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.nats) => nil | (int_symb M.nats) => (1%Z, Vnil) :: nil | (hd_symb M.hd) => nil | (int_symb M.hd) => (2%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.tl) => nil | (int_symb M.tl) => (2%Z, (Vcons 0 Vnil)) :: (3%Z, (Vcons 1 Vnil)) :: nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS3. Module PI3 := PolyInt PIS3. (* 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. left. co_scc. right. PI1.prove_termination. PI2.prove_termination. PI3.prove_termination. termination_trivial. Qed.