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 := | _dot__1 : symb | _eq__2 : symb | _if_3 : symb | nil : symb | purge : symb | remove : 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._dot__1 => 2 | M._eq__2 => 2 | M._if_3 => 3 | M.nil => 0 | M.purge => 1 | M.remove => 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 _dot__1 x2 x1 := F0 M._dot__1 (Vcons x2 (Vcons x1 Vnil)). Definition _eq__2 x2 x1 := F0 M._eq__2 (Vcons x2 (Vcons x1 Vnil)). Definition _if_3 x3 x2 x1 := F0 M._if_3 (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition nil := F0 M.nil Vnil. Definition purge x1 := F0 M.purge (Vcons x1 Vnil). Definition remove x2 x1 := F0 M.remove (Vcons x2 (Vcons x1 Vnil)). End S0. Definition E := @nil (@ATrs.rule s0). Definition R := R0 (S0.purge S0.nil) S0.nil :: R0 (S0.purge (S0._dot__1 (V0 0) (V0 1))) (S0._dot__1 (V0 0) (S0.purge (S0.remove (V0 0) (V0 1)))) :: R0 (S0.remove (V0 0) S0.nil) S0.nil :: R0 (S0.remove (V0 0) (S0._dot__1 (V0 1) (V0 2))) (S0._if_3 (S0._eq__2 (V0 0) (V0 1)) (S0.remove (V0 0) (V0 2)) (S0._dot__1 (V0 1) (S0.remove (V0 0) (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_dot__1 x2 x1 := F1 (hd_symb s1_p M._dot__1) (Vcons x2 (Vcons x1 Vnil)). Definition _dot__1 x2 x1 := F1 (int_symb s1_p M._dot__1) (Vcons x2 (Vcons x1 Vnil)). Definition h_eq__2 x2 x1 := F1 (hd_symb s1_p M._eq__2) (Vcons x2 (Vcons x1 Vnil)). Definition _eq__2 x2 x1 := F1 (int_symb s1_p M._eq__2) (Vcons x2 (Vcons x1 Vnil)). Definition h_if_3 x3 x2 x1 := F1 (hd_symb s1_p M._if_3) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition _if_3 x3 x2 x1 := F1 (int_symb s1_p M._if_3) (Vcons x3 (Vcons x2 (Vcons x1 Vnil))). Definition hnil := F1 (hd_symb s1_p M.nil) Vnil. Definition nil := F1 (int_symb s1_p M.nil) Vnil. Definition hpurge x1 := F1 (hd_symb s1_p M.purge) (Vcons x1 Vnil). Definition purge x1 := F1 (int_symb s1_p M.purge) (Vcons x1 Vnil). Definition hremove x2 x1 := F1 (hd_symb s1_p M.remove) (Vcons x2 (Vcons x1 Vnil)). Definition remove x2 x1 := F1 (int_symb s1_p M.remove) (Vcons x2 (Vcons x1 Vnil)). End S1. (* graph decomposition 1 *) Definition cs1 : list (list (@ATrs.rule s1)) := ( R1 (S1.hremove (V1 0) (S1._dot__1 (V1 1) (V1 2))) (S1.hremove (V1 0) (V1 2)) :: nil) :: ( R1 (S1.hpurge (S1._dot__1 (V1 0) (V1 1))) (S1.hremove (V1 0) (V1 1)) :: nil) :: ( R1 (S1.hpurge (S1._dot__1 (V1 0) (V1 1))) (S1.hpurge (S1.remove (V1 0) (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.purge) => nil | (int_symb M.purge) => (3%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.nil) => nil | (int_symb M.nil) => (3%Z, Vnil) :: nil | (hd_symb M._dot__1) => nil | (int_symb M._dot__1) => (3%Z, (Vcons 0 (Vcons 0 Vnil))) :: (3%Z, (Vcons 1 (Vcons 0 Vnil))) :: (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M.remove) => (2%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (int_symb M.remove) => (1%Z, (Vcons 1 (Vcons 0 Vnil))) :: (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._if_3) => nil | (int_symb M._if_3) => (3%Z, (Vcons 0 (Vcons 0 (Vcons 0 Vnil)))) :: nil | (hd_symb M._eq__2) => nil | (int_symb M._eq__2) => 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.purge) => (3%Z, (Vcons 1 Vnil)) :: nil | (int_symb M.purge) => (1%Z, (Vcons 1 Vnil)) :: nil | (hd_symb M.nil) => nil | (int_symb M.nil) => nil | (hd_symb M._dot__1) => nil | (int_symb M._dot__1) => (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.remove) => nil | (int_symb M.remove) => (1%Z, (Vcons 0 (Vcons 1 Vnil))) :: nil | (hd_symb M._if_3) => nil | (int_symb M._if_3) => nil | (hd_symb M._eq__2) => nil | (int_symb M._eq__2) => nil end. Lemma trsInt_wm : forall f, pweak_monotone (trsInt f). Proof. pmonotone. Qed. End PIS2. Module PI2 := PolyInt PIS2. (* 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. left. co_scc. right. PI2.prove_termination. termination_trivial. Qed.